Skip to content

Commit ac376c8

Browse files
authored
Revert "fix: ground theorems as grind parameters (#11579)"
This reverts commit aa4aff2.
1 parent aa4aff2 commit ac376c8

File tree

5 files changed

+19
-60
lines changed

5 files changed

+19
-60
lines changed

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

Lines changed: 11 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -188,13 +188,7 @@ def elabGrindSuggestions
188188
match attr with
189189
| .ematch kind =>
190190
try
191-
let oldSize := params.extraFacts.size
192191
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 }
198192
catch _ => pure () -- Don't worry if library suggestions gave bad theorems.
199193
| _ =>
200194
-- We could actually support arbitrary grind modifiers,
@@ -341,12 +335,19 @@ def evalGrindTraceCore (stx : Syntax) (trace := true) (verbose := true) (useSorr
341335
let config := { config with clean := false, trace, verbose, useSorry }
342336
let only := only.isSome
343337
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
344349
let mvarId ← getMainGoal
345350
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
350351
Grind.withProtectedMCtx config.abstractProof mvarId fun mvarId' => do
351352
let (tacs, _) ← Grind.GrindTacticM.runAtGoal mvarId' params do
352353
let finish ← Grind.Action.mkFinish

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

Lines changed: 2 additions & 24 deletions
Original file line numberDiff line numberDiff line change
@@ -67,21 +67,7 @@ public def addEMatchTheorem (params : Grind.Params) (id : Ident) (declName : Nam
6767
Grind.mkEMatchTheoremForDecl declName kind params.symPrios (minIndexable := minIndexable)
6868
if warn && params.ematch.containsWithSamePatterns thm.origin thm.patterns thm.cnstrs then
6969
warnRedundantEMatchArg params.ematch declName
70-
if thm.numParams == 0 && kind matches .default _ then
71-
/-
72-
**Note**: ignores pattern and adds ground fact directly
73-
Motivation:
74-
```
75-
opaque π : Rat
76-
axiom pi_pos : 0 < π
77-
example : π = 0 → False := by
78-
grind [pi_pos]
79-
80-
```
81-
-/
82-
return { params with extraFacts := params.extraFacts.push thm.proof }
83-
else
84-
return { params with extra := params.extra.push thm }
70+
return { params with extra := params.extra.push thm }
8571
| .defn =>
8672
if (← isReducible declName) then
8773
throwError "`{.ofConstName declName}` is a reducible definition, `grind` automatically unfolds them"
@@ -125,10 +111,7 @@ def processParam (params : Grind.Params)
125111
for thm in thms do
126112
params := { params with extra := params.extra.push thm }
127113
| .ematch kind =>
128-
let oldSize := params.extraFacts.size
129114
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 }
132115
| .cases eager =>
133116
if incremental then throwError "`cases` parameter are not supported here"
134117
ensureNoMinIndexable minIndexable
@@ -156,10 +139,7 @@ def processParam (params : Grind.Params)
156139
-- **Note**: We should not warn if `declName` is an inductive
157140
params ← withRef p <| addEMatchTheorem params id ctor (.default false) minIndexable (warn := False)
158141
else
159-
let oldSize := params.extraFacts.size
160142
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 }
163143
| .symbol prio =>
164144
ensureNoMinIndexable minIndexable
165145
params := { params with symPrios := params.symPrios.insert declName prio }
@@ -225,9 +205,7 @@ def processTermParam (params : Grind.Params)
225205
throwError "invalid `grind` parameter, modifier is redundant since the parameter type is not a `forall`{indentExpr type}"
226206
unless levelParams.isEmpty do
227207
throwError "invalid `grind` parameter, parameter type is not a `forall` and is universe polymorphic{indentExpr type}"
228-
return { params with
229-
extraFacts := params.extraFacts.push proof,
230-
extraFactsSyntax := params.extraFactsSyntax.push p.raw }
208+
return { params with extraFacts := params.extraFacts.push proof }
231209

232210
/--
233211
Elaborates `grind` parameters.

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

Lines changed: 0 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -40,8 +40,6 @@ 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 := {}
4543
funCCs : NameSet := {}
4644
norm : Simp.Context
4745
normProcs : Array Simprocs

tests/lean/run/grind_ground_param.lean

Lines changed: 0 additions & 17 deletions
This file was deleted.

tests/lean/run/try_library_suggestions.lean

Lines changed: 6 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -27,11 +27,10 @@ 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.
3230
/--
33-
info: Try this:
31+
info: Try these:
3432
[apply] grind only [special_7]
33+
[apply] grind => instantiate only [special_7]
3534
-/
3635
#guard_msgs in
3736
example : SpecialProperty 7 := by
@@ -79,11 +78,11 @@ set_library_suggestions (fun _ _ => pure #[
7978
{ name := `prop2_5, score := 0.7 }
8079
])
8180

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

0 commit comments

Comments
 (0)