Skip to content

Commit

Permalink
chore: adaptations for leanprover/lean4#3084 (#9452)
Browse files Browse the repository at this point in the history
These are the changes required to adapt to @digama0's leanprover/lean4#3084. 

Co-authored-by: @digama0 



Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
  • Loading branch information
semorrison and semorrison committed Jan 7, 2024
1 parent 26ccd1e commit 2c20845
Show file tree
Hide file tree
Showing 6 changed files with 20 additions and 14 deletions.
2 changes: 1 addition & 1 deletion Mathlib/CategoryTheory/Sites/CoverPreserving.lean
Original file line number Diff line number Diff line change
Expand Up @@ -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
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/MeasureTheory/Integral/SetToL1.lean
Original file line number Diff line number Diff line change
Expand Up @@ -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,
Expand Down
22 changes: 14 additions & 8 deletions Mathlib/Tactic/Cases.lean
Original file line number Diff line number Diff line change
Expand Up @@ -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 := #[]
Expand All @@ -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

Expand All @@ -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)
Expand All @@ -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)
Expand All @@ -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
4 changes: 2 additions & 2 deletions lake-manifest.json
Original file line number Diff line number Diff line change
Expand Up @@ -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",
Expand Down
2 changes: 1 addition & 1 deletion lakefile.lean
Original file line number Diff line number Diff line change
Expand Up @@ -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"
Expand Down
2 changes: 1 addition & 1 deletion lean-toolchain
Original file line number Diff line number Diff line change
@@ -1 +1 @@
leanprover/lean4:v4.5.0-rc1
leanprover/lean4:nightly-2023-12-22

0 comments on commit 2c20845

Please sign in to comment.