Skip to content

Commit cf22c36

Browse files
authored
feat: grind +premises (#10920)
This PR adds support for `grind +premises`, calling the currently configured premise selection algorithm and including the results as parameters to `grind`. (Recall that there is not currently a default premise selector provided by Lean4: you need a downstream premise selector to make use of this.)
1 parent 291d238 commit cf22c36

File tree

7 files changed

+155
-6
lines changed

7 files changed

+155
-6
lines changed

src/Init/Grind/Tactics.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -19,6 +19,9 @@ structure Config where
1919
/-- If `lax` is `true`, `grind` will silently ignore any parameters referring to non-existent theorems
2020
or for which no patterns can be generated. -/
2121
lax : Bool := false
22+
/-- If `premises` is `true`, `grind` will invoke the currently configured premise selecor on the current goal,
23+
and add attempt to use the resulting suggestions as premises to the `grind` tactic. -/
24+
premises : Bool := false
2225
/-- Maximum number of case-splits in a proof search branch. It does not include splits performed during normalization. -/
2326
splits : Nat := 9
2427
/-- Maximum number of E-matching (aka heuristic theorem instantiation) rounds before each case split. -/

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

Lines changed: 35 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -9,11 +9,13 @@ public import Lean.Meta.Tactic.Grind.Main
99
public import Lean.Meta.Tactic.TryThis
1010
public import Lean.Elab.Command
1111
public import Lean.Elab.Tactic.Config
12+
public import Lean.PremiseSelection.Basic
1213
import Lean.Meta.Tactic.Grind.SimpUtil
1314
import Lean.Meta.Tactic.Grind.EMatchTheoremParam
1415
import Lean.Elab.Tactic.Grind.Basic
1516
import Lean.Elab.MutualDef
1617
meta import Lean.Meta.Tactic.Grind.Parser
18+
1719
public section
1820
namespace Lean.Elab.Tactic
1921
open Meta
@@ -82,7 +84,16 @@ private def warnRedundantEMatchArg (s : Grind.EMatchTheorems) (declName : Name)
8284
pure m!"{ks}"
8385
logWarning m!"this parameter is redundant, environment already contains `{declName}` annotated with `{kinds}`"
8486

85-
def elabGrindParams (params : Grind.Params) (ps : TSyntaxArray ``Parser.Tactic.grindParam) (only : Bool) (lax : Bool := false) :
87+
private def parseModifier (s : String) : CoreM Grind.AttrKind := do
88+
let stx := Parser.runParserCategory (← getEnv) `Lean.Parser.Attr.grindMod s
89+
match stx with
90+
| .ok stx => Grind.getAttrKindCore stx
91+
| _ => throwError "unexpected modifier {s}"
92+
93+
open PremiseSelection in
94+
def elabGrindParams
95+
(params : Grind.Params) (ps : TSyntaxArray ``Parser.Tactic.grindParam) (only : Bool)
96+
(premises : Array Suggestion := #[]) (lax : Bool := false) :
8697
MetaM Grind.Params := do
8798
let mut params := params
8899
for p in ps do
@@ -104,6 +115,19 @@ def elabGrindParams (params : Grind.Params) (ps : TSyntaxArray ``Parser.Tactic.
104115
| _ => throwError "unexpected `grind` parameter{indentD p}"
105116
catch ex =>
106117
if !lax then throw ex
118+
for p in premises do
119+
let attr ← match p.flag with
120+
| some flag => parseModifier flag
121+
| none => pure <| .ematch (.default false)
122+
match attr with
123+
| .ematch kind =>
124+
params ← addEMatchTheorem params (mkIdent p.name) p.name kind false
125+
| _ =>
126+
-- We could actually support arbitrary grind modifiers,
127+
-- and call `processParam` rather than `addEMatchTheorem`,
128+
-- but this would require a larger refactor.
129+
-- Let's only do this if there is a prospect of a premise selector supprting this.
130+
throwError "unexpected modifier {p.flag}"
107131
return params
108132
where
109133
ensureNoMinIndexable (minIndexable : Bool) : MetaM Unit := do
@@ -206,13 +230,20 @@ where
206230
| _ =>
207231
throwError "invalid `grind` parameter, `{.ofConstName declName}` is not a theorem, definition, or inductive type"
208232

209-
def mkGrindParams (config : Grind.Config) (only : Bool) (ps : TSyntaxArray ``Parser.Tactic.grindParam) : MetaM Grind.Params := do
233+
def mkGrindParams
234+
(config : Grind.Config) (only : Bool) (ps : TSyntaxArray ``Parser.Tactic.grindParam) (mvarId : MVarId) :
235+
MetaM Grind.Params := do
210236
let params ← Grind.mkParams config
211237
let ematch ← if only then pure default else Grind.getEMatchTheorems
212238
let inj ← if only then pure default else Grind.getInjectiveTheorems
213239
let casesTypes ← if only then pure default else Grind.getCasesTypes
214240
let params := { params with ematch, casesTypes, inj }
215-
let params ← elabGrindParams params ps only (lax := config.lax)
241+
let premises ← if config.premises then
242+
let suggestions ← PremiseSelection.select mvarId
243+
pure suggestions
244+
else
245+
pure #[]
246+
let params ← elabGrindParams params ps only premises (lax := config.lax)
216247
trace[grind.debug.inj] "{params.inj.getOrigins.map (·.pp)}"
217248
return params
218249

@@ -226,7 +257,7 @@ def grind
226257
mvarId.admit
227258
return {}
228259
mvarId.withContext do
229-
let params ← mkGrindParams config only ps
260+
let params ← mkGrindParams config only ps mvarId
230261
let type ← mvarId.getType
231262
let mvar' ← mkFreshExprSyntheticOpaqueMVar type
232263
let finalize (result : Grind.Result) : TacticM Grind.Trace := do

src/Lean/PremiseSelection/Basic.lean

Lines changed: 5 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -195,7 +195,9 @@ builtin_initialize premiseSelectorExt : EnvExtension (Option Selector) ←
195195
def select (m : MVarId) (c : Config := {}) : MetaM (Array Suggestion) := do
196196
let some selector := premiseSelectorExt.getState (← getEnv) |
197197
throwError "No premise selector registered. \
198-
(Note the Lean does not provide a default premise selector, these must be installed by a downstream library.)"
198+
(Note that Lean does not provide a default premise selector, \
199+
these must be provided by a downstream library, \
200+
and configured using `set_premise_selector`.)"
199201
selector m c
200202

201203
/-!
@@ -216,6 +218,8 @@ open Lean Elab Command in
216218
@[builtin_command_elab setPremiseSelectorCmd, inherit_doc setPremiseSelectorCmd]
217219
def elabSetPremiseSelector : CommandElab
218220
| `(command| set_premise_selector $selector) => do
221+
if `Lean.PremiseSelection.Basic ∉ (← getEnv).header.moduleNames then
222+
logWarning "Add `import Lean.PremiseSelection.Basic` before using the `set_premise_selector` command."
219223
let selector ← liftTermElabM do
220224
try
221225
let selectorTerm ← Term.elabTermEnsuringType selector (some (Expr.const ``Selector []))

tests/lean/run/grind_lax.lean

Lines changed: 66 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,66 @@
1+
/-- error: Unknown constant `foo` -/
2+
#guard_msgs in
3+
example : False := by
4+
grind [foo]
5+
6+
/--
7+
error: `grind` failed
8+
case grind
9+
⊢ False
10+
[grind] Goal diagnostics
11+
[facts] Asserted facts
12+
-/
13+
#guard_msgs in
14+
example : False := by
15+
grind +lax [foo]
16+
17+
/-- error: Unknown constant `foo` -/
18+
#guard_msgs in
19+
example : False := by
20+
grind [-foo]
21+
22+
/--
23+
error: `grind` failed
24+
case grind
25+
⊢ False
26+
[grind] Goal diagnostics
27+
[facts] Asserted facts
28+
-/
29+
#guard_msgs in
30+
example : False := by
31+
grind +lax [-foo]
32+
33+
theorem foo : True := .intro
34+
35+
/--
36+
error: invalid E-matching equality theorem, conclusion must be an equality
37+
True
38+
-/
39+
#guard_msgs in
40+
example : False := by
41+
grind [= foo]
42+
43+
/--
44+
error: `grind` failed
45+
case grind
46+
⊢ False
47+
[grind] Goal diagnostics
48+
[facts] Asserted facts
49+
-/
50+
#guard_msgs in
51+
example : False := by
52+
grind +lax [= foo]
53+
54+
theorem bar : Nat = Int := sorry
55+
56+
/--
57+
error: invalid E-matching equality theorem, conclusion must be an equality
58+
True
59+
-/
60+
#guard_msgs in
61+
example : Int = Nat := by
62+
grind [= foo, = bar]
63+
64+
#guard_msgs in
65+
example : Int = Nat := by
66+
grind +lax [= foo, = bar]

tests/lean/run/grind_premises.lean

Lines changed: 40 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,40 @@
1+
import Lean.PremiseSelection.Basic
2+
3+
/--
4+
error: No premise selector registered. (Note that Lean does not provide a default premise selector, these must be provided by a downstream library, and configured using `set_premise_selector`.)
5+
-/
6+
#guard_msgs in
7+
example : True := by
8+
grind +premises
9+
10+
set_premise_selector (fun _ _ => pure #[])
11+
12+
#guard_msgs in
13+
example : True := by
14+
grind +premises
15+
16+
def P (_ : Nat) := True
17+
theorem p : P 7 := trivial
18+
19+
/--
20+
error: `grind` failed
21+
case grind
22+
h : ¬P 37
23+
⊢ False
24+
[grind] Goal diagnostics
25+
[facts] Asserted facts
26+
[prop] ¬P 37
27+
[eqc] False propositions
28+
[prop] P 37
29+
-/
30+
#guard_msgs in
31+
example : P 37 := by
32+
grind
33+
34+
example : P 7 := by
35+
grind [p]
36+
37+
set_premise_selector (fun _ _ => pure #[{ name := `p, score := 1.0 }])
38+
39+
example : P 7 := by
40+
grind +premises

tests/lean/run/premise_selection.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -15,7 +15,7 @@ error: Failed to elaborate Nat as a `MVarId → Config → MetaM (Array Suggesti
1515
set_premise_selector Nat
1616

1717
/--
18-
error: No premise selector registered. (Note the Lean does not provide a default premise selector, these must be installed by a downstream library.)
18+
error: No premise selector registered. (Note that Lean does not provide a default premise selector, these must be provided by a downstream library, and configured using `set_premise_selector`.)
1919
-/
2020
#guard_msgs in
2121
example : True := by
Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,5 @@
1+
/--
2+
warning: Add `import Lean.PremiseSelection.Basic` before using the `set_premise_selector` command.
3+
-/
4+
#guard_msgs (warning, drop error) in
5+
set_premise_selector (fun _ _ => pure #[])

0 commit comments

Comments
 (0)