Skip to content

Commit e984473

Browse files
authored
fix: markNestedProofs preprocessor in grind (#8412)
This PR fixes the `markNestedProofs` preprocessor used in `grind`. There was a missing case (e.g., `Expr.mdata`)
1 parent 88f6439 commit e984473

File tree

2 files changed

+9
-1
lines changed

2 files changed

+9
-1
lines changed

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

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -47,7 +47,7 @@ where
4747
return e'
4848
-- Remark: we have to process `Expr.proj` since we only
4949
-- fold projections later during term internalization
50-
unless e.isApp || e.isForall || e.isProj do
50+
unless e.isApp || e.isForall || e.isProj || e.isMData do
5151
return e
5252
-- Check whether it is cached
5353
if let some r := (← get).find? e then
@@ -68,6 +68,8 @@ where
6868
pure e
6969
| .proj _ _ b =>
7070
pure <| e.updateProj! (← visit b)
71+
| .mdata _ b =>
72+
pure <| e.updateMData! (← visit b)
7173
| .forallE _ d b _ =>
7274
-- Recall that we have `ForallProp.lean`.
7375
let d' ← visit d

tests/lean/run/grind_mark_nested_proofs_bug.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -26,3 +26,9 @@ example (as bs cs : Array α) (v : α)
2626
(h₆ : j < as.size)
2727
: cs[j] = as[j] := by
2828
grind only [= Array.getElem_set_ne_abstracted, = Array.size_set] -- should work
29+
30+
opaque p : (i : Nat) → i ≠ 10Prop
31+
32+
example (h : ∀ i, (¬i > 0) ∨ ∀ h : i ≠ 10, p i h) : p 5 (by decide) := by
33+
have := h 5; clear h
34+
grind

0 commit comments

Comments
 (0)