Skip to content

Commit

Permalink
chore: adaptations for leanprover/lean4#3084
Browse files Browse the repository at this point in the history
  • Loading branch information
kim-em committed Jan 5, 2024
1 parent 26348e8 commit a85698d
Show file tree
Hide file tree
Showing 3 changed files with 17 additions and 11 deletions.
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 lakefile.lean
Original file line number Diff line number Diff line change
Expand Up @@ -26,9 +26,9 @@ 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" @ "bump/v4.6.0"
require Qq from git "https://github.com/leanprover-community/quote4" @ "master"
require aesop from git "https://github.com/leanprover-community/aesop" @ "master"
require aesop from git "https://github.com/leanprover-community/aesop" @ "bump/v4.6.0"
require proofwidgets from git "https://github.com/leanprover-community/ProofWidgets4" @ "v0.0.25"
require Cli from git "https://github.com/leanprover/lean4-cli" @ "main"

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 a85698d

Please sign in to comment.