@@ -12,20 +12,14 @@ import Lean.Meta.Tactic.Subst
1212
1313namespace Lean.Meta
1414
15- structure CaseValueSubgoal where
16- mvarId : MVarId
17- newH : FVarId
18- deriving Inhabited
19-
2015/--
21- Split goal `... |- C x` into two subgoals
22- `..., (h : x = value) |- C x`
23- `..., (h : x != value) |- C x`
24- where `fvarId` is `x`s id.
16+ Split goal `... |- C x`,, where `fvarId` is `x`s id, into two subgoals
17+ `..., |- (h : x = value) → C x`
18+ `..., |- (h : x != value) → C x`
2519 The type of `x` must have decidable equality.
2620 -/
2721def caseValue (mvarId : MVarId) (fvarId : FVarId) (value : Expr) (hName : Name := `h)
28- : MetaM (CaseValueSubgoal × CaseValueSubgoal ) :=
22+ : MetaM (MVarId × MVarId ) :=
2923 mvarId.withContext do
3024 let tag ← mvarId.getTag
3125 mvarId.checkNotAssigned `caseValue
@@ -38,15 +32,7 @@ def caseValue (mvarId : MVarId) (fvarId : FVarId) (value : Expr) (hName : Name :
3832 let elseMVar ← mkFreshExprSyntheticOpaqueMVar elseTarget tag
3933 let val ← mkAppOptM `dite #[none, xEqValue, none, thenMVar, elseMVar]
4034 mvarId.assign val
41- let (elseH, elseMVarId) ← elseMVar.mvarId!.intro1P
42- let elseSubgoal := { mvarId := elseMVarId, newH := elseH }
43- let (thenH, thenMVarId) ← thenMVar.mvarId!.intro1P
44- thenMVarId.withContext do
45- trace[Meta] "searching for decl"
46- let _ ← thenH.getDecl
47- trace[Meta] "found decl"
48- let thenSubgoal := { mvarId := thenMVarId, newH := thenH }
49- pure (thenSubgoal, elseSubgoal)
35+ return (thenMVar.mvarId!, elseMVar.mvarId!)
5036
5137public structure CaseValuesSubgoal where
5238 mvarId : MVarId
@@ -55,34 +41,44 @@ public structure CaseValuesSubgoal where
5541 deriving Inhabited
5642
5743/--
58- Split goal `... |- C x` into values.size + 1 subgoals
59- 1) `..., (h_1 : x = value[0]) |- C value[0]`
44+ Split goal `... |- C x`, where `fvarId` is `x`s id, into ` values.size + 1` subgoals
45+ 1) `..., (h_1 : x = value[0]) |- C value[0]`
6046 ...
61- n) `..., (h_n : x = value[n - 1]) |- C value[n - 1]`
47+ n) `..., (h_n : x = value[n - 1]) |- C value[n - 1]`
6248 n+1) `..., (h_1 : x != value[0]) ... (h_n : x != value[n-1]) |- C x`
6349 where `n = values.size`
64- where `fvarId` is `x`s id.
6550 The type of `x` must have decidable equality.
6651
6752 Remark: the last subgoal is for the "else" catchall case, and its `subst` is `{}`.
6853 Remark: the field `newHs` has size 1 forall but the last subgoal.
6954
70- If `substNewEqs = true`, then the new `h_i` equality hypotheses are substituted in the first `n` cases .
55+ If `needsHyps = false` then the else case comes without hypotheses .
7156-/
72- public def caseValues (mvarId : MVarId) (fvarId : FVarId) (values : Array Expr) (hNamePrefix := `h) : MetaM (Array CaseValuesSubgoal) :=
57+ public def caseValues (mvarId : MVarId) (fvarId : FVarId) (values : Array Expr) (hNamePrefix := `h)
58+ (needHyps := true ) : MetaM (Array CaseValuesSubgoal) :=
7359 let rec loop : Nat → MVarId → List Expr → Array FVarId → Array CaseValuesSubgoal → MetaM (Array CaseValuesSubgoal)
7460 | _, mvarId, [], _, _ => throwTacticEx `caseValues mvarId "list of values must not be empty"
7561 | i, mvarId, v::vs, hs, subgoals => do
76- let (thenSubgoal, elseSubgoal) ← caseValue mvarId fvarId v (hNamePrefix.appendIndexAfter i)
77- appendTagSuffix thenSubgoal.mvarId ((`case).appendIndexAfter i)
78- let thenMVarId ← thenSubgoal.mvarId.tryClearMany hs
79- let (subst, mvarId) ← substCore thenMVarId thenSubgoal.newH (symm := false ) {} (clearH := true )
80- let subgoals := subgoals.push { mvarId := mvarId, newHs := #[], subst := subst }
62+ let (thenMVarId, elseMVarId) ← caseValue mvarId fvarId v (hNamePrefix.appendIndexAfter i)
63+ appendTagSuffix thenMVarId ((`case).appendIndexAfter i)
64+ let thenMVarId ← thenMVarId.tryClearMany hs
65+ let (thenH, thenMVarId) ← thenMVarId.intro1P
66+ let (subst, thenMVarId) ← substCore thenMVarId thenH (symm := false ) {} (clearH := true )
67+ let subgoals := subgoals.push { mvarId := thenMVarId, newHs := #[], subst := subst }
68+ let (hs', elseMVarId) ←
69+ if needHyps then
70+ let (elseH, elseMVarId) ← elseMVarId.intro1P
71+ pure (hs.push elseH, elseMVarId)
72+ else
73+ let elseMVarId ← elseMVarId.intro1_
74+ pure (hs, elseMVarId)
8175 match vs with
8276 | [] => do
83- appendTagSuffix elseSubgoal.mvarId ((`case).appendIndexAfter (i+1 ))
84- pure $ subgoals.push { mvarId := elseSubgoal.mvarId, newHs := hs.push elseSubgoal.newH, subst := {} }
85- | vs => loop (i+1 ) elseSubgoal.mvarId vs (hs.push elseSubgoal.newH) subgoals
77+ appendTagSuffix elseMVarId ((`case).appendIndexAfter (i+1 ))
78+ pure $ subgoals.push { mvarId := elseMVarId, newHs := hs', subst := {} }
79+ | vs =>
80+ loop (i+1 ) elseMVarId vs hs' subgoals
81+
8682 loop 1 mvarId values.toList #[] #[]
8783
8884end Lean.Meta
0 commit comments