diff --git a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean index 6f29ac9836a05..88882611e28f7 100644 --- a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean +++ b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean @@ -120,7 +120,7 @@ theorem Presieve.FamilyOfElements.Compatible.functorPushforward : theorem CompatiblePreserving.apply_map {Y : C} {f : Y ⟶ Z} (hf : T f) : x.functorPushforward G (G.map f) (image_mem_functorPushforward G T hf) = x f hf := by unfold FamilyOfElements.functorPushforward - rcases e₁ : getFunctorPushforwardStructure (image_mem_functorPushforward G T hf) with + rcases getFunctorPushforwardStructure (image_mem_functorPushforward G T hf) with ⟨X, g, f', hg, eq⟩ simpa using hG.compatible ℱ h f' (𝟙 _) hg hf (by simp [eq]) #align category_theory.compatible_preserving.apply_map CategoryTheory.CompatiblePreserving.apply_map diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index a6468c2c05e87..dc7d4e079f130 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -638,7 +638,7 @@ theorem setToSimpleFunc_const' [Nonempty α] (T : Set α → F →L[ℝ] F') (x theorem setToSimpleFunc_const (T : Set α → F →L[ℝ] F') (hT_empty : T ∅ = 0) (x : F) {m : MeasurableSpace α} : SimpleFunc.setToSimpleFunc T (SimpleFunc.const α x) = T univ x := by - cases hα : isEmpty_or_nonempty α + cases isEmpty_or_nonempty α · have h_univ_empty : (univ : Set α) = ∅ := Subsingleton.elim _ _ rw [h_univ_empty, hT_empty] simp only [setToSimpleFunc, ContinuousLinearMap.zero_apply, sum_empty, diff --git a/Mathlib/Tactic/Cases.lean b/Mathlib/Tactic/Cases.lean index 425ea805c77c8..5f423d3e1d625 100644 --- a/Mathlib/Tactic/Cases.lean +++ b/Mathlib/Tactic/Cases.lean @@ -40,7 +40,8 @@ open Lean Meta Elab Elab.Tactic open private getAltNumFields in evalCases ElimApp.evalAlts.go in def ElimApp.evalNames (elimInfo : ElimInfo) (alts : Array ElimApp.Alt) (withArg : Syntax) - (numEqs := 0) (numGeneralized := 0) (toClear : Array FVarId := #[]) : + (numEqs := 0) (generalized : Array FVarId := #[]) (toClear : Array FVarId := #[]) + (toTag : Array (Ident × FVarId) := #[]) : TermElabM (Array MVarId) := do let mut names : List Syntax := withArg[1].getArgs |>.toList let mut subgoals := #[] @@ -50,10 +51,15 @@ def ElimApp.evalNames (elimInfo : ElimInfo) (alts : Array ElimApp.Alt) (withArg names := names' let (fvars, g) ← g.introN numFields <| altVarNames.map (getNameOfIdent' ·[0]) let some (g, subst) ← Cases.unifyEqs? numEqs g {} | pure () - let (_, g) ← g.introNP numGeneralized + let (introduced, g) ← g.introNP generalized.size + let subst := (generalized.zip introduced).foldl (init := subst) fun subst (a, b) => + subst.insert a (.fvar b) let g ← liftM $ toClear.foldlM (·.tryClear) g - for fvar in fvars, stx in altVarNames do - g.withContext <| (subst.apply <| .fvar fvar).addLocalVarInfoForBinderIdent ⟨stx⟩ + g.withContext do + for (stx, fvar) in toTag do + Term.addLocalVarInfo stx (subst.get fvar) + for fvar in fvars, stx in altVarNames do + (subst.get fvar).addLocalVarInfoForBinderIdent ⟨stx⟩ subgoals := subgoals.push g pure subgoals @@ -62,7 +68,7 @@ elab (name := induction') "induction' " tgts:(Parser.Tactic.casesTarget,+) usingArg:((" using " ident)?) withArg:((" with" (ppSpace colGt binderIdent)+)?) genArg:((" generalizing" (ppSpace colGt ident)+)?) : tactic => do - let targets ← elabCasesTargets tgts.1.getSepArgs + let (targets, toTag) ← elabCasesTargets tgts.1.getSepArgs let g :: gs ← getUnsolvedGoals | throwNoGoalsToBeSolved g.withContext do let elimInfo ← getElimNameInfo usingArg targets (induction := true) @@ -88,12 +94,12 @@ elab (name := induction') "induction' " tgts:(Parser.Tactic.casesTarget,+) ElimApp.setMotiveArg g elimArgs[elimInfo.motivePos]!.mvarId! targetFVarIds g.assign result.elimApp let subgoals ← ElimApp.evalNames elimInfo result.alts withArg - (numGeneralized := fvarIds.size) (toClear := targetFVarIds) + (generalized := fvarIds) (toClear := targetFVarIds) (toTag := toTag) setGoals <| (subgoals ++ result.others).toList ++ gs elab (name := cases') "cases' " tgts:(Parser.Tactic.casesTarget,+) usingArg:((" using " ident)?) withArg:((" with" (ppSpace colGt binderIdent)+)?) : tactic => do - let targets ← elabCasesTargets tgts.1.getSepArgs + let (targets, toTag) ← elabCasesTargets tgts.1.getSepArgs let g :: gs ← getUnsolvedGoals | throwNoGoalsToBeSolved g.withContext do let elimInfo ← getElimNameInfo usingArg targets (induction := false) @@ -108,5 +114,5 @@ elab (name := cases') "cases' " tgts:(Parser.Tactic.casesTarget,+) usingArg:((" ElimApp.setMotiveArg g motive.mvarId! targetsNew g.assign result.elimApp let subgoals ← ElimApp.evalNames elimInfo result.alts withArg - (numEqs := targets.size) (toClear := targetsNew) + (numEqs := targets.size) (toClear := targetsNew) (toTag := toTag) setGoals <| subgoals.toList ++ gs diff --git a/lake-manifest.json b/lake-manifest.json index 13c561aed7980..fe2644d28f3e5 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -4,10 +4,10 @@ [{"url": "https://github.com/leanprover/std4", "type": "git", "subDir": null, - "rev": "0f6bc5b32bf5b0498902d3b5f0806c75530539d5", + "rev": "4f632e6fac86f6c3b7d4ac127c0ce8b06ab86f63", "name": "std", "manifestFile": "lake-manifest.json", - "inputRev": "main", + "inputRev": "4f632e6fac86f6c3b7d4ac127c0ce8b06ab86f63", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/quote4", diff --git a/lakefile.lean b/lakefile.lean index db2b89928d405..2cde560b467f8 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -26,7 +26,7 @@ package mathlib where meta if get_config? doc = some "on" then -- do not download and build doc-gen4 by default require «doc-gen4» from git "https://github.com/leanprover/doc-gen4" @ "main" -require std from git "https://github.com/leanprover/std4" @ "main" +require std from git "https://github.com/leanprover/std4" @ "4f632e6fac86f6c3b7d4ac127c0ce8b06ab86f63" require Qq from git "https://github.com/leanprover-community/quote4" @ "master" require aesop from git "https://github.com/leanprover-community/aesop" @ "master" require proofwidgets from git "https://github.com/leanprover-community/ProofWidgets4" @ "v0.0.25" diff --git a/lean-toolchain b/lean-toolchain index 3f21e50bd46eb..0a830983958ad 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:v4.5.0-rc1 +leanprover/lean4:nightly-2023-12-22