Skip to content

Commit d762aa1

Browse files
committed
include ground theorems in grind?
1 parent 725598e commit d762aa1

File tree

5 files changed

+37
-18
lines changed

5 files changed

+37
-18
lines changed

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

Lines changed: 10 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -188,7 +188,13 @@ def elabGrindSuggestions
188188
match attr with
189189
| .ematch kind =>
190190
try
191+
let oldSize := params.extraFacts.size
191192
params ← addEMatchTheorem params (mkIdent p.name) p.name kind false (warn := false)
193+
-- If the theorem went to extraFacts (ground theorem), synthesize syntax for it.
194+
-- Ground theorems bypass E-matching, so we need to track them separately for suggestions.
195+
if params.extraFacts.size > oldSize then
196+
let stx ← `(Parser.Tactic.grindParam| $(mkIdent p.name):ident)
197+
params := { params with extraFactsSyntax := params.extraFactsSyntax.push stx.raw }
192198
catch _ => pure () -- Don't worry if library suggestions gave bad theorems.
193199
| _ =>
194200
-- We could actually support arbitrary grind modifiers,
@@ -335,19 +341,12 @@ def evalGrindTraceCore (stx : Syntax) (trace := true) (verbose := true) (useSorr
335341
let config := { config with clean := false, trace, verbose, useSorry }
336342
let only := only.isSome
337343
let paramStxs := if let some params := params? then params.getElems else #[]
338-
-- Extract term parameters (non-ident params) to include in the suggestion.
339-
-- These are not tracked via E-matching, so we conservatively include them all.
340-
-- Ident params resolve to global declarations and are tracked via E-matching.
341-
-- Non-ident terms (like `show P by tac`) need to be preserved explicitly.
342-
let termParamStxs : Array Grind.TParam := paramStxs.filter fun p =>
343-
match p with
344-
| `(Parser.Tactic.grindParam| $[$_:grindMod]? $_:ident) => false
345-
| `(Parser.Tactic.grindParam| ! $[$_:grindMod]? $_:ident) => false
346-
| `(Parser.Tactic.grindParam| - $_:ident) => false
347-
| `(Parser.Tactic.grindParam| #$_:hexnum) => false
348-
| _ => true
349344
let mvarId ← getMainGoal
350345
let params ← mkGrindParams config only paramStxs mvarId
346+
-- Extract syntax of params that went to `extraFacts` (ground theorems and term arguments).
347+
-- These need to be included in `grind only` suggestions since they're not tracked via E-matching.
348+
let termParamStxs : Array Grind.TParam := params.extraFactsSyntax.filterMap fun s =>
349+
if s.isOfKind ``Parser.Tactic.grindParam then some ⟨s⟩ else none
351350
Grind.withProtectedMCtx config.abstractProof mvarId fun mvarId' => do
352351
let (tacs, _) ← Grind.GrindTacticM.runAtGoal mvarId' params do
353352
let finish ← Grind.Action.mkFinish

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

Lines changed: 9 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -125,7 +125,10 @@ def processParam (params : Grind.Params)
125125
for thm in thms do
126126
params := { params with extra := params.extra.push thm }
127127
| .ematch kind =>
128+
let oldSize := params.extraFacts.size
128129
params ← withRef p <| addEMatchTheorem params id declName kind minIndexable
130+
if params.extraFacts.size > oldSize then
131+
params := { params with extraFactsSyntax := params.extraFactsSyntax.push p.raw }
129132
| .cases eager =>
130133
if incremental then throwError "`cases` parameter are not supported here"
131134
ensureNoMinIndexable minIndexable
@@ -153,7 +156,10 @@ def processParam (params : Grind.Params)
153156
-- **Note**: We should not warn if `declName` is an inductive
154157
params ← withRef p <| addEMatchTheorem params id ctor (.default false) minIndexable (warn := False)
155158
else
159+
let oldSize := params.extraFacts.size
156160
params ← withRef p <| addEMatchTheorem params id declName (.default false) minIndexable (suggest := true)
161+
if params.extraFacts.size > oldSize then
162+
params := { params with extraFactsSyntax := params.extraFactsSyntax.push p.raw }
157163
| .symbol prio =>
158164
ensureNoMinIndexable minIndexable
159165
params := { params with symPrios := params.symPrios.insert declName prio }
@@ -219,7 +225,9 @@ def processTermParam (params : Grind.Params)
219225
throwError "invalid `grind` parameter, modifier is redundant since the parameter type is not a `forall`{indentExpr type}"
220226
unless levelParams.isEmpty do
221227
throwError "invalid `grind` parameter, parameter type is not a `forall` and is universe polymorphic{indentExpr type}"
222-
return { params with extraFacts := params.extraFacts.push proof }
228+
return { params with
229+
extraFacts := params.extraFacts.push proof,
230+
extraFactsSyntax := params.extraFactsSyntax.push p.raw }
223231

224232
/--
225233
Elaborates `grind` parameters.

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

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -40,6 +40,8 @@ structure Params where
4040
extra : PArray EMatchTheorem := {}
4141
extraInj : PArray InjectiveTheorem := {}
4242
extraFacts : PArray Expr := {}
43+
/-- Syntax of params that went to `extraFacts`, for `grind?` suggestions. -/
44+
extraFactsSyntax : Array Syntax := {}
4345
funCCs : NameSet := {}
4446
norm : Simp.Context
4547
normProcs : Array Simprocs

tests/lean/run/grind_ground_param.lean

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -6,3 +6,12 @@ example : π = 0 → False := by
66

77
example : 0 < 2 * π := by
88
grind [pi_pos]
9+
10+
-- Test that grind? includes ground theorems in suggestions
11+
/--
12+
info: Try this:
13+
[apply] grind only [pi_pos]
14+
-/
15+
#guard_msgs in
16+
example : π = 0 → False := by
17+
grind? [pi_pos]

tests/lean/run/try_library_suggestions.lean

Lines changed: 7 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -27,10 +27,11 @@ axiom special_7 : SpecialProperty 7
2727
set_library_suggestions (fun _ _ => pure #[{ name := `special_7, score := 1.0 }])
2828

2929
-- Expected: try? should find grind only [special_7]
30+
-- Note: Ground theorems (0 parameters) are added as facts during initialization,
31+
-- not via E-matching, so there's no instantiate script.
3032
/--
31-
info: Try these:
33+
info: Try this:
3234
[apply] grind only [special_7]
33-
[apply] grind => instantiate only [special_7]
3435
-/
3536
#guard_msgs in
3637
example : SpecialProperty 7 := by
@@ -78,11 +79,11 @@ set_library_suggestions (fun _ _ => pure #[
7879
{ name := `prop2_5, score := 0.7 }
7980
])
8081

81-
-- Expected: try? should use the best applicable one
82+
-- Note: Both ground theorems are included since we can't track which
83+
-- extraFacts were actually used. This is conservative but correct.
8284
/--
83-
info: Try these:
84-
[apply] grind only [prop1_5]
85-
[apply] grind => instantiate only [prop1_5]
85+
info: Try this:
86+
[apply] grind only [prop1_5, prop2_5]
8687
-/
8788
#guard_msgs in
8889
example : Property1 5 := by

0 commit comments

Comments
 (0)