Skip to content

Commit 093be2a

Browse files
kim-emclaude
andcommitted
feat: add grind +simp mode
This PR adds `grind +simp`, which includes all `@[simp]` lemmas as E-matching theorems. Patterns are generated dynamically at tactic execution time. If performance becomes an issue with large simp databases, the solution is to hook into the `@[simp]` attribute handler to pre-compute and cache patterns at registration time. 🤖 Generated with [Claude Code](https://claude.com/claude-code) Co-Authored-By: Claude Opus 4.5 <[email protected]>
1 parent 19e1fe5 commit 093be2a

File tree

3 files changed

+110
-3
lines changed

3 files changed

+110
-3
lines changed

src/Init/Grind/Config.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -21,6 +21,13 @@ structure Config where
2121
/-- If `suggestions` is `true`, `grind` will invoke the currently configured library suggestion engine on the current goal,
2222
and add attempt to use the resulting suggestions as additional parameters to the `grind` tactic. -/
2323
suggestions : Bool := false
24+
/--
25+
If `simp` is `true`, `grind` will use all `@[simp]` lemmas as E-matching theorems.
26+
This generates grind patterns dynamically at tactic execution time, which may be slow
27+
with large simp databases. If performance becomes an issue, the proper solution is to
28+
hook into the `@[simp]` attribute handler to generate and cache patterns at registration time.
29+
-/
30+
simp : Bool := false
2431
/-- Maximum number of case-splits in a proof search branch. It does not include splits performed during normalization. -/
2532
splits : Nat := 9
2633
/-- Maximum number of E-matching (aka heuristic theorem instantiation) rounds before each case split. -/

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

Lines changed: 48 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -13,6 +13,7 @@ public import Lean.LibrarySuggestions.Basic
1313
import Lean.Meta.Tactic.Grind.SimpUtil
1414
import Lean.Meta.Tactic.Grind.Util
1515
import Lean.Meta.Tactic.Grind.EMatchTheoremParam
16+
import Lean.Meta.Tactic.Simp.Attr
1617
import Lean.Elab.Tactic.Grind.Basic
1718
import Lean.Elab.Tactic.Grind.Param
1819
import Lean.Meta.Tactic.Grind.Action
@@ -198,6 +199,46 @@ def elabGrindSuggestions
198199
throwError "unexpected modifier {p.flag}"
199200
return params
200201

202+
/--
203+
Check if the conclusion of a type (after stripping foralls) is an equality-like proposition
204+
(Eq, Iff, or HEq). Returns true if so.
205+
-/
206+
private def isEqLike (type : Expr) : Bool :=
207+
let conclusion := type.getForallBody
208+
conclusion.isAppOf ``Eq || conclusion.isAppOf ``Iff || conclusion.isAppOf ``HEq
209+
210+
/--
211+
Add simp theorems as E-matching theorems for grind.
212+
213+
**Performance note**: This function iterates over all simp lemmas and generates
214+
grind patterns dynamically at tactic execution time. This is potentially slow
215+
with large simp databases. If this becomes a bottleneck, the proper solution is
216+
to hook into the `@[simp]` attribute handler (in `Lean.Meta.Tactic.Simp.Attr`)
217+
to generate and cache grind patterns at attribute registration time.
218+
-/
219+
def elabGrindSimpTheorems (params : Grind.Params) : MetaM Grind.Params := do
220+
let simpTheorems ← getSimpTheorems
221+
let symPrios ← Grind.getGlobalSymbolPriorities
222+
let mut params := params
223+
-- Iterate over all simp lemma origins (includes both fwd and inv directions)
224+
for origin in simpTheorems.lemmaNames.toList do
225+
match origin with
226+
| .decl declName _ inv =>
227+
try
228+
let info ← getConstInfo declName
229+
-- Use eqLhs/eqRhs for equality-like theorems (faster pattern extraction),
230+
-- fall back to .default for other propositions like `P 37`
231+
let kind : Grind.EMatchTheoremKind :=
232+
if isEqLike info.type then
233+
if inv then .eqRhs false else .eqLhs false
234+
else
235+
.default false
236+
let thm ← Grind.mkEMatchTheoremForDecl declName kind symPrios
237+
params := { params with ematch := params.ematch.insert thm }
238+
catch _ => pure () -- Silently skip theorems that fail
239+
| _ => pure () -- Skip non-declaration origins (fvars, etc.)
240+
return params
241+
201242
open LibrarySuggestions in
202243
def elabGrindParamsAndSuggestions
203244
(params : Grind.Params)
@@ -225,6 +266,8 @@ def mkGrindParams
225266
else
226267
pure #[]
227268
let mut params ← elabGrindParamsAndSuggestions params ps suggestions (only := only) (lax := config.lax)
269+
if config.simp then
270+
params ← elabGrindSimpTheorems params
228271
trace[grind.debug.inj] "{params.inj.getOrigins.map (·.pp)}"
229272
if params.anchorRefs?.isSome then
230273
/-
@@ -297,17 +340,19 @@ def setGrindParams (stx : TSyntax `tactic) (params : Array Syntax) : TSyntax `ta
297340
def getGrindParams (stx : TSyntax `tactic) : Array Syntax :=
298341
stx.raw[grindParamsPos][1].getSepArgs
299342

300-
/-- Filter out `+suggestions` from the config syntax -/
343+
/-- Filter out `+suggestions` and `+simp` from the config syntax for `grind?` suggestions. -/
301344
def filterSuggestionsFromGrindConfig (config : TSyntax ``Lean.Parser.Tactic.optConfig) :
302345
TSyntax ``Lean.Parser.Tactic.optConfig :=
303346
let configItems := config.raw.getArgs
304347
let filteredItems := configItems.filter fun item =>
305-
-- Keep all items except +suggestions
348+
-- Keep all items except +suggestions and +simp
306349
-- Structure: null node -> configItem -> posConfigItem -> ["+", ident]
307350
match item[0]? with
308351
| some configItem => match configItem[0]? with
309352
| some posConfigItem => match posConfigItem[1]? with
310-
| some ident => !(posConfigItem.getKind == ``Lean.Parser.Tactic.posConfigItem && ident.getId == `suggestions)
353+
| some ident =>
354+
let id := ident.getId
355+
!(posConfigItem.getKind == ``Lean.Parser.Tactic.posConfigItem && (id == `suggestions || id == `simp))
311356
| none => true
312357
| none => true
313358
| none => true

tests/lean/run/grind_simp.lean

Lines changed: 55 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,55 @@
1+
/-!
2+
# Tests for `grind +simp`
3+
4+
This file tests the `grind +simp` feature, which allows grind to use all `@[simp]` lemmas
5+
as E-matching theorems.
6+
-/
7+
8+
-- Test basic simp lemma usage
9+
def foo (n : Nat) : Nat := n + 1
10+
def bar (n : Nat) : Nat := 1 + n
11+
12+
@[simp] theorem foo_eq_bar : foo n = bar n := by simp [foo, bar, Nat.add_comm]
13+
14+
-- Verify plain grind fails but grind +simp succeeds
15+
example : foo 5 = bar 5 := by
16+
fail_if_success grind
17+
grind +simp
18+
19+
-- Verify grind? +simp suggests the simp lemma
20+
/--
21+
info: Try these:
22+
[apply] grind only [= foo_eq_bar]
23+
[apply] grind => instantiate only [= foo_eq_bar]
24+
-/
25+
#guard_msgs in
26+
example : foo 5 = bar 5 := by grind? +simp
27+
28+
-- Test that +simp can be combined with hypotheses
29+
def qux (n : Nat) : Nat := n
30+
31+
@[simp] theorem qux_id : qux n = n := rfl
32+
33+
example (h : a = b) : qux a = b := by
34+
fail_if_success grind
35+
grind +simp
36+
37+
opaque P : Nat → Prop
38+
@[simp] axiom p : P 37
39+
40+
example : P 37 := by
41+
fail_if_success grind
42+
grind +simp
43+
44+
-- Test that grind +simp works with standard library simp lemmas
45+
-- String.length_append is currently @[simp] but not @[grind]
46+
-- (It's fine to delete this without replacement once we add `grind` annotations to String lemmas.)
47+
/--
48+
info: Try these:
49+
[apply] grind only [= String.length_append]
50+
[apply] grind => instantiate only [= String.length_append]
51+
-/
52+
#guard_msgs in
53+
example (s t : String) : (s ++ t).length = s.length + t.length := by
54+
fail_if_success grind
55+
grind? +simp

0 commit comments

Comments
 (0)