Files
NNG/Game/Tactic/Cases.lean
2024-02-29 17:19:01 +01:00

69 lines
2.7 KiB
Lean4
Raw Blame History

This file contains ambiguous Unicode characters
This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.
import Game.MyNat.Definition
import Mathlib.Tactic.Cases
/-!
# Modified `cases` tactic
Modify `cases` tactic to always show `(0 : MyNat)` instead of `MyNat.zero` and
to support the lean3-style `with` keyword.
This is mainly copied and modified from the mathlib-tactic `cases'`.
-/
namespace MyNat
/-- Modified `casesOn` principal to use `0` instead of `MyNat.zero`. -/
def casesOn' {P : Sort u} (t : ) (zero : P 0)
(succ : (a : ) P (MyNat.succ a)) : P t := by
cases t with
| zero => assumption
| succ n =>
apply succ
open Lean Meta Elab Parser Tactic Mathlib.Tactic
open private getElimNameInfo from Lean.Elab.Tactic.Induction
/-- Modified `cases` tactic for this game.
Usage: `cases n with d` if `n : `; `cases h with h1 h2` if `h : P Q`; `cases h with c hc` if `h : a ≤ b`.
*(This implementation mimics the `cases'` from mathlib. The actual `cases` tactic in mathlib has a more complex syntax)*
-/
elab (name := cases) "cases " tgts:(Parser.Tactic.casesTarget,+) usingArg:((" using " ident)?)
withArg:((" with" (ppSpace colGt binderIdent)+)?) : tactic => do
let (targets, toTag) elabCasesTargets tgts.1.getSepArgs
let g :: gs getUnsolvedGoals | throwNoGoalsToBeSolved
g.withContext do
let elimInfo getElimNameInfo usingArg targets (induction := false)
-- Edit: If `MyNat.casesOn` is used, we want to use `MyNat.casesOn` instead.
let elimInfo match elimInfo.elimExpr.getAppFn.constName? with
| some `MyNat.casesOn =>
let modifiedUsingArgs : TSyntax Name.anonymous :=
match usingArg.raw with
| .node info kind #[] =>
-- TODO: How do you construct syntax in a semi-userfriendly way??
.node info kind #[.atom info "using", .ident info "MyNat.casesOn'".toSubstring `MyNat.casesOn' []]
| other => other
let newElimInfo getElimNameInfo modifiedUsingArgs targets (induction := false)
pure newElimInfo
| _ => pure elimInfo
let targets addImplicitTargets elimInfo targets
let result withRef tgts <| ElimApp.mkElimApp elimInfo targets ( g.getTag)
let elimArgs := result.elimApp.getAppArgs
let targets elimInfo.targetsPos.mapM (instantiateMVars elimArgs[·]!)
let motive := elimArgs[elimInfo.motivePos]!
let g generalizeTargetsEq g ( inferType motive) targets
let (targetsNew, g) g.introN targets.size
g.withContext do
ElimApp.setMotiveArg g motive.mvarId! targetsNew
g.assign result.elimApp
let subgoals ElimApp.evalNames elimInfo result.alts withArg
(numEqs := targets.size) (toClear := targetsNew) (toTag := toTag)
setGoals <| subgoals.toList ++ gs
end MyNat