@@ -13,6 +13,7 @@ public import Lean.LibrarySuggestions.Basic
1313import Lean.Meta.Tactic.Grind.SimpUtil
1414import Lean.Meta.Tactic.Grind.Util
1515import Lean.Meta.Tactic.Grind.EMatchTheoremParam
16+ import Lean.Meta.Tactic.Simp.Attr
1617import Lean.Elab.Tactic.Grind.Basic
1718import Lean.Elab.Tactic.Grind.Param
1819import 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+
201242open LibrarySuggestions in
202243def 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
297340def 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. -/
301344def 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
0 commit comments