Skip to content

Commit a0772dc

Browse files
authored
fix: grind internalization (#11318)
This PR fixes a local declaration internalization in `grind` that was exposed when using `grind -revert`. This bug was affecting a `grind` proof in Mathlib.
1 parent 90389a8 commit a0772dc

File tree

3 files changed

+56
-10
lines changed

3 files changed

+56
-10
lines changed

src/Lean/Meta/Tactic/Grind/Main.lean

Lines changed: 1 addition & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -215,16 +215,7 @@ private def addHypotheses (goal : Goal) : GrindM Goal := GoalM.run' goal do
215215
| none => add r.expr localDecl.toExpr
216216
| some h => add r.expr <| mkApp4 (mkConst ``Eq.mp [0]) type r.expr h localDecl.toExpr
217217
else
218-
/-
219-
**Note**: We must internalize local variables because their types may be empty, and may not be
220-
referenced anywhere else. Example:
221-
```
222-
example (a : { x : Int // x < 0 ∧ x > 0 }) : False := by grind
223-
```
224-
-/
225-
let e ← shareCommon localDecl.toExpr
226-
internalize e 0
227-
processNewFacts
218+
internalizeLocalDecl localDecl
228219

229220
private def initCore (mvarId : MVarId) (params : Params) : GrindM Goal := do
230221
/-

src/Lean/Meta/Tactic/Grind/Types.lean

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1288,6 +1288,23 @@ opaque internalize (e : Expr) (generation : Nat) (parent? : Option Expr := none)
12881288
@[extern "lean_grind_process_new_facts"]
12891289
opaque processNewFacts : GoalM Unit
12901290

1291+
/--
1292+
Internalizes a local declaration which is not a proposition.
1293+
**Note**: We must internalize local variables because their types may be empty, and may not be
1294+
referenced anywhere else. Example:
1295+
```
1296+
example (a : { x : Int // x < 0 ∧ x > 0 }) : False := by grind
1297+
```
1298+
`etaStruct` may also be applicable.
1299+
-/
1300+
def internalizeLocalDecl (localDecl : LocalDecl) : GoalM Unit := do
1301+
let e ← shareCommon localDecl.toExpr
1302+
internalize e 0
1303+
/-
1304+
**Note**: `internalize` may add new facts (e.g., `etaStruct` equality)
1305+
-/
1306+
processNewFacts
1307+
12911308
/--
12921309
Returns a proof that `a = b` if they have the same type. Otherwise, returns a proof of `a ≍ b`.
12931310
It assumes `a` and `b` are in the same equivalence class.
Lines changed: 38 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,38 @@
1+
namespace List
2+
3+
/-
4+
This example exposed a bug in the `grind` internalizer when `grind -revert`
5+
is used.
6+
-/
7+
8+
variable {α : Type} {β : α → Type}
9+
10+
def keys : List (Sigma β) → List α :=
11+
map Sigma.fst
12+
13+
@[grind =]
14+
theorem keys_cons {s} {l : List (Sigma β)} : (s :: l).keys = s.1 :: l.keys :=
15+
rfl
16+
17+
variable [DecidableEq α]
18+
19+
def dlookup (a : α) : List (Sigma β) → Option (β a)
20+
| [] => none
21+
| ⟨a', b⟩ :: l => if h : a' = a then some (Eq.recOn h b) else dlookup a l
22+
23+
@[grind =]
24+
theorem dlookup_cons_eq (l) (a : α) (b : β a) : dlookup a (⟨a, b⟩ :: l) = some b :=
25+
dif_pos rfl
26+
27+
@[grind =]
28+
theorem dlookup_cons_ne (l) {a} : ∀ s : Sigma β, a ≠ s.1 → dlookup a (s :: l) = dlookup a l
29+
| ⟨_, _⟩, h => dif_neg h.symm
30+
31+
@[grind =]
32+
theorem dlookup_isSome {a : α} {l : List (Sigma β)} : (dlookup a l).isSome ↔ a ∈ l.keys := by
33+
induction l with
34+
| nil => sorry
35+
| cons s _ _ =>
36+
by_cases a = s.fst
37+
· grind
38+
· sorry

0 commit comments

Comments
 (0)