Skip to content

Commit 64cf5e5

Browse files
authored
feat: improve grind search procedure (#6657)
This PR improves the `grind` search procedure, and adds the new configuration option: `failures`.
1 parent 127b3f9 commit 64cf5e5

File tree

6 files changed

+105
-13
lines changed

6 files changed

+105
-13
lines changed

src/Init/Grind/Tactics.lean

+2
Original file line numberDiff line numberDiff line change
@@ -45,6 +45,8 @@ structure Config where
4545
If `splitIndPred` is `true`, `grind` performs case-splitting on inductive predicates.
4646
Otherwise, it performs case-splitting only on types marked with `[grind_split]` attribute. -/
4747
splitIndPred : Bool := true
48+
/-- By default, `grind` halts as soon as it encounters a sub-goal where no further progress can be made. -/
49+
failures : Nat := 1
4850
deriving Inhabited, BEq
4951

5052
end Lean.Grind

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

-3
Original file line numberDiff line numberDiff line change
@@ -370,7 +370,4 @@ def ematchAndAssert : GrindTactic := fun goal => do
370370
return none
371371
assertAll goal
372372

373-
def ematchStar : GrindTactic :=
374-
ematchAndAssert.iterate
375-
376373
end Lean.Meta.Grind

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

+2-8
Original file line numberDiff line numberDiff line change
@@ -15,6 +15,7 @@ import Lean.Meta.Tactic.Grind.Inv
1515
import Lean.Meta.Tactic.Grind.Intro
1616
import Lean.Meta.Tactic.Grind.EMatch
1717
import Lean.Meta.Tactic.Grind.Split
18+
import Lean.Meta.Tactic.Grind.Solve
1819
import Lean.Meta.Tactic.Grind.SimpUtil
1920

2021
namespace Lean.Meta.Grind
@@ -68,17 +69,10 @@ private def initCore (mvarId : MVarId) : GrindM (List Goal) := do
6869
goals.forM (·.checkInvariants (expensive := true))
6970
return goals.filter fun goal => !goal.inconsistent
7071

71-
def all (goals : List Goal) (f : Goal → GrindM (List Goal)) : GrindM (List Goal) := do
72-
goals.foldlM (init := []) fun acc goal => return acc ++ (← f goal)
73-
74-
/-- A very simple strategy -/
75-
private def simple (goals : List Goal) : GrindM (List Goal) := do
76-
applyToAll (assertAll >> ematchStar >> (splitNext >> assertAll >> ematchStar).iterate) goals
77-
7872
def main (mvarId : MVarId) (config : Grind.Config) (mainDeclName : Name) (fallback : Fallback) : MetaM (List Goal) := do
7973
let go : GrindM (List Goal) := do
8074
let goals ← initCore mvarId
81-
let goals ← simple goals
75+
let goals ← solve goals
8276
let goals ← goals.filterMapM fun goal => do
8377
if goal.inconsistent then return none
8478
let goal ← GoalM.run' goal fallback

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

+92
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,92 @@
1+
/-
2+
Copyright (c) 2025 Amazon.com, Inc. or its affiliates. All Rights Reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Leonardo de Moura
5+
-/
6+
prelude
7+
import Lean.Meta.Tactic.Grind.Combinators
8+
import Lean.Meta.Tactic.Grind.Split
9+
import Lean.Meta.Tactic.Grind.EMatch
10+
11+
namespace Lean.Meta.Grind
12+
13+
namespace Solve
14+
15+
structure State where
16+
todo : List Goal
17+
failures : List Goal := []
18+
stop : Bool := false
19+
20+
private abbrev M := StateRefT State GrindM
21+
22+
def getNext? : M (Option Goal) := do
23+
let goal::todo := (← get).todo | return none
24+
modify fun s => { s with todo }
25+
return some goal
26+
27+
def pushGoal (goal : Goal) : M Unit :=
28+
modify fun s => { s with todo := goal :: s.todo }
29+
30+
def pushGoals (goals : List Goal) : M Unit :=
31+
modify fun s => { s with todo := goals ++ s.todo }
32+
33+
def pushFailure (goal : Goal) : M Unit := do
34+
modify fun s => { s with failures := goal :: s.failures }
35+
if (← get).failures.length ≥ (← getConfig).failures then
36+
modify fun s => { s with stop := true }
37+
38+
@[inline] def stepGuard (x : Goal → M Bool) (goal : Goal) : M Bool := do
39+
try
40+
x goal
41+
catch ex =>
42+
if ex.isMaxHeartbeat || ex.isMaxRecDepth then
43+
let goal ← goal.reportIssue ex.toMessageData
44+
pushFailure goal
45+
return true
46+
else
47+
throw ex
48+
49+
def applyTac (x : GrindTactic) (goal : Goal) : M Bool := do
50+
let go (goal : Goal) : M Bool := do
51+
let some goals ← x goal | return false
52+
pushGoals goals
53+
return true
54+
stepGuard go goal
55+
56+
def tryAssertNext : Goal → M Bool := applyTac assertNext
57+
58+
def tryEmatch : Goal → M Bool := applyTac ematchAndAssert
59+
60+
def trySplit : Goal → M Bool := applyTac splitNext
61+
62+
def maxNumFailuresReached : M Bool := do
63+
return (← get).failures.length ≥ (← getConfig).failures
64+
65+
partial def main : M Unit := do
66+
repeat do
67+
if (← get).stop then
68+
return ()
69+
let some goal ← getNext? |
70+
return ()
71+
if goal.inconsistent then
72+
continue
73+
if (← tryAssertNext goal) then
74+
continue
75+
if (← tryEmatch goal) then
76+
continue
77+
if (← trySplit goal) then
78+
continue
79+
pushFailure goal
80+
81+
end Solve
82+
83+
/--
84+
Try to solve/close the given goals, and returns the ones that could not be solved.
85+
-/
86+
def solve (goals : List Goal) : GrindM (List Goal) := do
87+
let (_, s) ← Solve.main.run { todo := goals }
88+
let todo ← s.todo.mapM fun goal => do
89+
goal.reportIssue m!"this goal was not fully processed due to previous failures, threshold: `(failures := {(← getConfig).failures})`"
90+
return s.failures.reverse ++ todo
91+
92+
end Lean.Meta.Grind

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

+7-2
Original file line numberDiff line numberDiff line change
@@ -416,14 +416,19 @@ def updateLastTag : GoalM Unit := do
416416
trace[grind] "working on goal `{currTag}`"
417417
modifyThe Grind.State fun s => { s with lastTag := currTag }
418418

419-
def reportIssue (msg : MessageData) : GoalM Unit := do
419+
def Goal.reportIssue (goal : Goal) (msg : MessageData) : MetaM Goal := do
420420
let msg ← addMessageContext msg
421-
modify fun s => { s with issues := .trace { cls := `issue } msg #[] :: s.issues }
421+
let goal := { goal with issues := .trace { cls := `issue } msg #[] :: goal.issues }
422422
/-
423423
We also add a trace message because we may want to know when
424424
an issue happened relative to other trace messages.
425425
-/
426426
trace[grind.issues] msg
427+
return goal
428+
429+
def reportIssue (msg : MessageData) : GoalM Unit := do
430+
let goal ← (← get).reportIssue msg
431+
set goal
427432

428433
/--
429434
Macro similar to `trace[...]`, but it includes the trace message `trace[grind] "working on <current goal>"`

tests/lean/run/grind_pre.lean

+2
Original file line numberDiff line numberDiff line change
@@ -198,6 +198,8 @@ h : p
198198
[eqc] False propositions
199199
[prop] ¬p
200200
[prop] ¬r
201+
[issues] Issues
202+
[issue] this goal was not fully processed due to previous failures, threshold: `(failures := 1)`
201203
-/
202204
#guard_msgs (error) in
203205
example (a : α) (p q r : Prop) : (h₁ : HEq p a) → (h₂ : HEq q a) → (h₃ : p = r) → False := by

0 commit comments

Comments
 (0)