Skip to content

Commit e5393cf

Browse files
authored
fix: cases tactic to handle non-atomic eliminator well (#8361)
This PR fixes a bug in the `cases` tacic introduced in #3188 that arises when cases (not induction) is used with a non-atomic expression in using and the argument indexing gets confused. This fixes #8360.
1 parent 3481f43 commit e5393cf

File tree

3 files changed

+18
-7
lines changed

3 files changed

+18
-7
lines changed

src/Lean/Elab/Tactic/Induction.lean

Lines changed: 10 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -128,10 +128,11 @@ private def getFType : M Expr := do
128128
pure fType
129129

130130
structure Result where
131-
elimApp : Expr
132-
motive : MVarId
133-
alts : Array Alt
134-
others : Array MVarId
131+
elimApp : Expr
132+
elimArgs : Array Expr
133+
motive : MVarId
134+
alts : Array Alt
135+
others : Array MVarId
135136
complexArgs : Array Expr
136137

137138
/--
@@ -209,7 +210,10 @@ partial def mkElimApp (elimInfo : ElimInfo) (targets : Array Expr) (tag : Name)
209210
unless targets.contains motiveArg do
210211
throwError "mkElimApp: Expected first {targets.size} arguments of motive in conclusion to be one of the targets:{indentExpr s.fType}"
211212
pure motiveArgs[targets.size:]
212-
return { elimApp := (← instantiateMVars s.f), alts, others, motive, complexArgs }
213+
let elimApp ← instantiateMVars s.f
214+
-- `elimArgs` is the argument list that the offsets in `elimInfo` work with
215+
let elimArgs := elimApp.getAppArgs[elimInfo.elimExpr.getAppNumArgs:]
216+
return { elimApp, elimArgs, alts, others, motive, complexArgs }
213217

214218
/--
215219
Given a goal `... targets ... |- C[targets, complexArgs]` associated with `mvarId`,
@@ -981,7 +985,7 @@ def evalCasesCore (stx : Syntax) (elimInfo : ElimInfo) (targets : Array Expr)
981985
let tag ← mvarId.getTag
982986
mvarId.withContext do
983987
let result ← withRef targetRef <| ElimApp.mkElimApp elimInfo targets tag
984-
let elimArgs := result.elimApp.getAppArgs
988+
let elimArgs := result.elimArgs
985989
let targets ← elimInfo.targetsPos.mapM fun i => instantiateMVars elimArgs[i]!
986990
let motiveType ← inferType elimArgs[elimInfo.motivePos]!
987991
let mvarId ← generalizeTargetsEq mvarId motiveType targets

src/Lean/Meta/Tactic/ElimInfo.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -62,7 +62,7 @@ def getElimExprInfo (elimExpr : Expr) (baseDeclName? : Option Name := none) : Me
6262
let motiveType ← inferType motive
6363
forallTelescopeReducing motiveType fun motiveParams motiveResultType => do
6464
unless motiveParams.size == motiveArgs.size do
65-
throwError "expecetd {motiveArgs.size} parameters at motive type, got {motiveParams.size}:{indentExpr motiveType}"
65+
throwError "expected {motiveArgs.size} parameters at motive type, got {motiveParams.size}:{indentExpr motiveType}"
6666
unless motiveResultType.isSort do
6767
throwError "motive result type must be a sort{indentExpr motiveType}"
6868
let some motivePos ← pure (xs.idxOf? motive) |

tests/lean/run/issue8360.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,7 @@
1+
axiom thm.{u_1, u_2} {α : Type u_1} {β : Type u_2} (f : α → β)
2+
(motive : List α → List β → Prop) (case1 : motive [] [])
3+
(case2 : ∀ (a : α) (as : List α), motive (a :: as) (f a :: List.map f as)) (l : List α) : motive l (List.map f l)
4+
5+
theorem map_isNil (f : α → β) (l : List α) :
6+
List.map f l = [] ↔ l = [] := by
7+
cases l using thm f <;> simp

0 commit comments

Comments
 (0)