Skip to content

Commit 69e1eae

Browse files
authored
feat: grind +lax ignores bad parameters (#10890)
This PR adds a `+lax` configuration option for `grind`, causing it to ignore parameters referring to non-existent theorems, or to theorems for which we can't generate a pattern. This allows throwing large sets of theorems (e.g. from a premise selection enginre) into `grind` to see what happens.
1 parent b1d4c9b commit 69e1eae

File tree

2 files changed

+24
-17
lines changed

2 files changed

+24
-17
lines changed

src/Init/Grind/Tactics.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -16,6 +16,9 @@ Passed to `grind` using, for example, the `grind (config := { matchEqs := true }
1616
structure Config where
1717
/-- If `trace` is `true`, `grind` records used E-matching theorems and case-splits. -/
1818
trace : Bool := false
19+
/-- If `lax` is `true`, `grind` will silently ignore any parameters referring to non-existent theorems
20+
or for which no patterns can be generated. -/
21+
lax : Bool := false
1922
/-- Maximum number of case-splits in a proof search branch. It does not include splits performed during normalization. -/
2023
splits : Nat := 9
2124
/-- Maximum number of E-matching (aka heuristic theorem instantiation) rounds before each case split. -/

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

Lines changed: 21 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -82,24 +82,28 @@ private def warnRedundantEMatchArg (s : Grind.EMatchTheorems) (declName : Name)
8282
pure m!"{ks}"
8383
logWarning m!"this parameter is redundant, environment already contains `{declName}` annotated with `{kinds}`"
8484

85-
def elabGrindParams (params : Grind.Params) (ps : TSyntaxArray ``Parser.Tactic.grindParam) (only : Bool) : MetaM Grind.Params := do
85+
def elabGrindParams (params : Grind.Params) (ps : TSyntaxArray ``Parser.Tactic.grindParam) (only : Bool) (lax : Bool := false) :
86+
MetaM Grind.Params := do
8687
let mut params := params
8788
for p in ps do
88-
match p with
89-
| `(Parser.Tactic.grindParam| - $id:ident) =>
90-
let declName ← realizeGlobalConstNoOverloadWithInfo id
91-
if let some declName ← Grind.isCasesAttrCandidate? declName false then
92-
Grind.ensureNotBuiltinCases declName
93-
params := { params with casesTypes := (← params.casesTypes.eraseDecl declName) }
94-
else if (← Grind.isInjectiveTheorem declName) then
95-
params := { params with inj := params.inj.erase (.decl declName) }
96-
else
97-
params := { params with ematch := (← params.ematch.eraseDecl declName) }
98-
| `(Parser.Tactic.grindParam| $[$mod?:grindMod]? $id:ident) =>
99-
params ← processParam params p mod? id (minIndexable := false)
100-
| `(Parser.Tactic.grindParam| ! $[$mod?:grindMod]? $id:ident) =>
101-
params ← processParam params p mod? id (minIndexable := true)
102-
| _ => throwError "unexpected `grind` parameter{indentD p}"
89+
try
90+
match p with
91+
| `(Parser.Tactic.grindParam| - $id:ident) =>
92+
let declName ← realizeGlobalConstNoOverloadWithInfo id
93+
if let some declName ← Grind.isCasesAttrCandidate? declName false then
94+
Grind.ensureNotBuiltinCases declName
95+
params := { params with casesTypes := (← params.casesTypes.eraseDecl declName) }
96+
else if (← Grind.isInjectiveTheorem declName) then
97+
params := { params with inj := params.inj.erase (.decl declName) }
98+
else
99+
params := { params with ematch := (← params.ematch.eraseDecl declName) }
100+
| `(Parser.Tactic.grindParam| $[$mod?:grindMod]? $id:ident) =>
101+
params ← processParam params p mod? id (minIndexable := false)
102+
| `(Parser.Tactic.grindParam| ! $[$mod?:grindMod]? $id:ident) =>
103+
params ← processParam params p mod? id (minIndexable := true)
104+
| _ => throwError "unexpected `grind` parameter{indentD p}"
105+
catch ex =>
106+
if !lax then throw ex
103107
return params
104108
where
105109
ensureNoMinIndexable (minIndexable : Bool) : MetaM Unit := do
@@ -208,7 +212,7 @@ def mkGrindParams (config : Grind.Config) (only : Bool) (ps : TSyntaxArray ``Pa
208212
let inj ← if only then pure default else Grind.getInjectiveTheorems
209213
let casesTypes ← if only then pure default else Grind.getCasesTypes
210214
let params := { params with ematch, casesTypes, inj }
211-
let params ← elabGrindParams params ps only
215+
let params ← elabGrindParams params ps only (lax := config.lax)
212216
trace[grind.debug.inj] "{params.inj.getOrigins.map (·.pp)}"
213217
return params
214218

0 commit comments

Comments
 (0)