Skip to content

Commit 47228b9

Browse files
authored
feat: arbitrary grind parameters (#11268)
This PR implements support for arbitrary `grind` parameters. The feature is similar to the one available in `simp`, where a proof term is treated as a local universe-polymorphic lemma. This feature relies on `grind -revert` (see #11248). For example, users can now write: ```lean def snd (p : α × β) : β := p.2 theorem snd_eq (a : α) (b : β) : snd (a, b) = b := rfl /-- trace: [grind.ematch.instance] snd_eq (a + 1): snd (a + 1, Type) = Type [grind.ematch.instance] snd_eq (a + 1): snd (a + 1, true) = true -/ #guard_msgs (trace) in set_option trace.grind.ematch.instance true in example (a : Nat) : (snd (a + 1, true), snd (a + 1, Type), snd (2, 2)) = (true, Type, snd (2, 2)) := by grind [snd_eq (a + 1)] ``` Note that in the example above, `snd_eq` is instantiated only twice, but with different universe parameters. As described in #11248, the new feature cannot be used with `grind +revert`.
1 parent 126fca1 commit 47228b9

File tree

4 files changed

+78
-7
lines changed

4 files changed

+78
-7
lines changed

src/Init/Grind/Interactive.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -11,19 +11,19 @@ namespace Lean.Parser.Tactic
1111

1212
syntax anchor := "#" noWs hexnum
1313

14-
syntax grindLemma := ppGroup((Attr.grindMod ppSpace)? ident)
14+
syntax grindLemma := ppGroup((Attr.grindMod ppSpace)? term)
1515
/--
1616
The `!` modifier instructs `grind` to consider only minimal indexable subexpressions
1717
when selecting patterns.
1818
-/
19-
syntax grindLemmaMin := ppGroup("!" (Attr.grindMod ppSpace)? ident)
19+
syntax grindLemmaMin := ppGroup("!" (Attr.grindMod ppSpace)? term)
2020

2121
syntax grindErase := "-" ident
2222
/--
2323
The `!` modifier instructs `grind` to consider only minimal indexable subexpressions
2424
when selecting patterns.
2525
-/
26-
syntax grindParam := grindErase <|> grindLemma <|> grindLemmaMin <|> anchor
26+
syntax grindParam := grindErase <|> grindLemmaMin <|> grindLemma <|> anchor
2727

2828
namespace Grind
2929
declare_syntax_cat grind_filter (behavior := both)
@@ -72,7 +72,7 @@ syntax (name := linarith) "linarith" : grind
7272
/-- The `sorry` tactic is a temporary placeholder for an incomplete tactic proof. -/
7373
syntax (name := «sorry») "sorry" : grind
7474

75-
syntax thm := anchor <|> grindLemma <|> grindLemmaMin
75+
syntax thm := anchor <|> grindLemmaMin <|> grindLemma
7676

7777
/--
7878
Instantiates theorems using E-matching.

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

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -146,13 +146,13 @@ def elabGrindParamsAndSuggestions
146146
(params : Grind.Params)
147147
(ps : TSyntaxArray ``Parser.Tactic.grindParam)
148148
(suggestions : Array Suggestion := #[])
149-
(only : Bool) (lax : Bool := false) : MetaM Grind.Params := do
149+
(only : Bool) (lax : Bool := false) : TermElabM Grind.Params := do
150150
let params ← elabGrindParams params ps (lax := lax) (only := only)
151151
elabGrindSuggestions params suggestions
152152

153153
def mkGrindParams
154154
(config : Grind.Config) (only : Bool) (ps : TSyntaxArray ``Parser.Tactic.grindParam) (mvarId : MVarId) :
155-
MetaM Grind.Params := do
155+
TermElabM Grind.Params := do
156156
let params ← Grind.mkParams config
157157
let ematch ← if only then pure default else Grind.getEMatchTheorems
158158
let inj ← if only then pure default else Grind.getInjectiveTheorems

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

Lines changed: 53 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -11,6 +11,7 @@ import Lean.Meta.Tactic.Grind.Internalize
1111
import Lean.Meta.Tactic.Grind.ForallProp
1212
import Lean.Elab.Tactic.Grind.Basic
1313
import Lean.Elab.Tactic.Grind.Anchor
14+
import Lean.Elab.SyntheticMVars
1415
namespace Lean.Elab.Tactic
1516
open Meta
1617

@@ -147,13 +148,60 @@ def processAnchor (params : Grind.Params) (val : TSyntax `hexnum) : CoreM Grind.
147148
let anchorRef ← Grind.elabAnchorRef val
148149
return { params with anchorRefs? := some <| anchorRefs.push anchorRef }
149150

151+
def checkNoRevert (params : Grind.Params) : CoreM Unit := do
152+
if params.config.revert then
153+
throwError "invalid `grind` parameter, only global declarations are allowed when `+revert` is used"
154+
155+
def processTermParam (params : Grind.Params)
156+
(p : TSyntax `Lean.Parser.Tactic.grindParam)
157+
(mod? : Option (TSyntax `Lean.Parser.Attr.grindMod))
158+
(term : Term)
159+
(minIndexable : Bool)
160+
: TermElabM Grind.Params := withRef p do
161+
checkNoRevert params
162+
let kind ← if let some mod := mod? then Grind.getAttrKindCore mod else pure .infer
163+
let kind ← match kind with
164+
| .ematch .user | .cases _ | .intro | .inj | .ext | .symbol _ =>
165+
throwError "invalid `grind` parameter, only global declarations are allowed with this kind of modifier"
166+
| .ematch kind => pure kind
167+
| .infer => pure <| .default false
168+
let thm? ← Term.withoutModifyingElabMetaStateWithInfo <| withRef p do
169+
let e ← Term.elabTerm term .none
170+
Term.synthesizeSyntheticMVars (postpone := .no) (ignoreStuckTC := true)
171+
let e ← instantiateMVars e
172+
if e.hasSyntheticSorry then
173+
return .none
174+
let e := e.eta
175+
if e.hasMVar then
176+
let r ← abstractMVars e
177+
return some (r.paramNames, r.expr)
178+
else
179+
return some (#[], e)
180+
let some (levelParams, proof) := thm? | return params
181+
unless (← isProof proof) do
182+
throwError "invalid `grind` parameter, proof term expected"
183+
let mkThm (kind : Grind.EMatchTheoremKind) (idx : Nat) : MetaM Grind.EMatchTheorem := do
184+
let id := ((`extra).appendIndexAfter idx)
185+
let some thm ← Grind.mkEMatchTheoremWithKind? (.stx id p) levelParams proof kind params.symPrios (minIndexable := minIndexable)
186+
| throwError "invalid `grind` parameter, failed to infer patterns"
187+
return thm
188+
let idx := params.extra.size
189+
match kind with
190+
| .eqBoth gen =>
191+
ensureNoMinIndexable minIndexable
192+
return { params with extra := params.extra.push (← mkThm (.eqLhs gen) idx) |>.push (← mkThm (.eqRhs gen) idx) }
193+
| _ =>
194+
if kind matches .eqLhs _ | .eqRhs _ then
195+
ensureNoMinIndexable minIndexable
196+
return { params with extra := params.extra.push (← mkThm kind idx) }
197+
150198
/--
151199
Elaborates `grind` parameters.
152200
`incremental = true` for tactics such as `finish`, in this case, we disable some kinds of parameters
153201
such as `- ident`.
154202
-/
155203
public def elabGrindParams (params : Grind.Params) (ps : TSyntaxArray ``Parser.Tactic.grindParam)
156-
(only : Bool) (lax : Bool := false) (incremental := false) : MetaM Grind.Params := do
204+
(only : Bool) (lax : Bool := false) (incremental := false) : TermElabM Grind.Params := do
157205
let mut params := params
158206
for p in ps do
159207
try
@@ -173,6 +221,10 @@ public def elabGrindParams (params : Grind.Params) (ps : TSyntaxArray ``Parser.T
173221
params ← processParam params p mod? id (minIndexable := false) (only := only) (incremental := incremental)
174222
| `(Parser.Tactic.grindParam| ! $[$mod?:grindMod]? $id:ident) =>
175223
params ← processParam params p mod? id (minIndexable := true) (only := only) (incremental := incremental)
224+
| `(Parser.Tactic.grindParam| $[$mod?:grindMod]? $e:term) =>
225+
params ← processTermParam params p mod? e (minIndexable := false)
226+
| `(Parser.Tactic.grindParam| ! $[$mod?:grindMod]? $e:term) =>
227+
params ← processTermParam params p mod? e (minIndexable := true)
176228
| `(Parser.Tactic.grindParam| #$anchor:hexnum) =>
177229
unless only do
178230
throwErrorAt anchor "invalid anchor, `only` modifier expected"
Lines changed: 19 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,19 @@
1+
def snd (p : α × β) : β :=
2+
p.2
3+
4+
theorem snd_eq (a : α) (b : β) : snd (a, b) = b :=
5+
rfl
6+
7+
/-- error: invalid `grind` parameter, only global declarations are allowed when `+revert` is used -/
8+
#guard_msgs in
9+
example (a : Nat) : (snd (a + 1, true), snd (a + 1, Type), snd (2, 2)) = (true, Type, snd (2, 2)) := by
10+
grind +revert [snd_eq (a + 1)]
11+
12+
/--
13+
trace: [grind.ematch.instance] snd_eq (a + 1): snd (a + 1, Type) = Type
14+
[grind.ematch.instance] snd_eq (a + 1): snd (a + 1, true) = true
15+
-/
16+
#guard_msgs (trace) in
17+
set_option trace.grind.ematch.instance true in
18+
example (a : Nat) : (snd (a + 1, true), snd (a + 1, Type), snd (2, 2)) = (true, Type, snd (2, 2)) := by
19+
grind only [snd_eq (a + 1)]

0 commit comments

Comments
 (0)