Skip to content

Commit aaefa73

Browse files
committed
fix: finish? checks whether solver propagation steps are needed
This PR ensures that solver propagation steps are necessary in the generated tactic script to close the goal. It produces more compact proof scripts, but this is not just an optimization, if we include an unnecessary step, we may fail to replay the generated script when `cases` steps are pruned using non-chronological backtracking (NCB). For example, when executing `finish?`, we may have performed a `cases #<anchor>` step that enabled `ring` to propagate a new fact. If this fact is not used in the final proof, and the corresponding `cases #<anchor>` step is pruned by NCB, the `ring` step will fail during replay.
1 parent 53442d4 commit aaefa73

File tree

2 files changed

+56
-25
lines changed

2 files changed

+56
-25
lines changed

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

Lines changed: 34 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -257,22 +257,6 @@ def terminalAction (check : GoalM Bool) (mkTac : GrindM (TSyntax `grind)) : Acti
257257
else
258258
kna goal'
259259

260-
/--
261-
Helper action for satellite solvers that use `CheckResult`.
262-
-/
263-
def solverAction (check : GoalM CheckResult) (mkTac : GrindM (TSyntax `grind)) : Action := fun goal kna kp => do
264-
let (result, goal') ← GoalM.run goal check
265-
match result with
266-
| .none => kna goal'
267-
| .progress => kp goal'
268-
| .propagated =>
269-
let goal' ← GoalM.run' goal' processNewFacts
270-
if goal'.inconsistent then
271-
closeWith mkTac
272-
else
273-
concatTactic (← kp goal') mkTac
274-
| .closed => closeWith mkTac
275-
276260
def saveStateIfTracing : GrindM (Option SavedState) := do
277261
if (← getConfig).trace then
278262
return some (← saveState)
@@ -309,6 +293,40 @@ def checkTactic : Action := fun goal _ kp => do
309293
return r
310294
| _ => return r
311295

296+
/--
297+
Helper action for satellite solvers that use `CheckResult`.
298+
-/
299+
def solverAction (check : GoalM CheckResult) (mkTac : GrindM (TSyntax `grind)) : Action := fun goal kna kp => do
300+
let saved? ← saveStateIfTracing
301+
let (result, goal') ← GoalM.run goal check
302+
match result with
303+
| .none => kna goal'
304+
| .progress => kp goal'
305+
| .propagated =>
306+
let goal' ← GoalM.run' goal' processNewFacts
307+
if goal'.inconsistent then
308+
closeWith mkTac
309+
else if (← getConfig).trace then
310+
match (← kp goal') with
311+
| .closed seq =>
312+
/-
313+
**Note**: Check whether the progress made by this solver was actually used to close the goal.
314+
This is not just an optimization, if we include an unnecessary step, we may fail to replay
315+
the generated script when `cases` steps are pruned using non-chronological backtracking (NCB).
316+
For example, when executing `finish?`, we may have performed a `cases #<anchor>` step
317+
that enabled `ring` to propagate a new fact. If this fact is not used in the final proof,
318+
and the corresponding `cases #<anchor>` step is pruned by NCB, the `ring` step will fail during replay.
319+
-/
320+
if (← checkSeqAt saved? goal seq) then
321+
return .closed seq
322+
else
323+
let tac ← mkTac
324+
return .closed (tac :: seq)
325+
| r => return r
326+
else
327+
kp goal'
328+
| .closed => closeWith mkTac
329+
312330
section
313331
/-!
314332
Some sanity check properties.

tests/lean/run/grind_finish_trace.lean

Lines changed: 22 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -2,7 +2,7 @@ open Lean Grind
22

33
/--
44
info: Try this:
5-
[apply] cases #c4b6 <;> ring <;> cases #4c68 <;> ring
5+
[apply] cases #c4b6 <;> cases #4c68 <;> ring
66
-/
77
#guard_msgs in
88
example {α : Type} [CommRing α] (a b c d e : α) :
@@ -30,9 +30,7 @@ example (p : Nat → Prop) (x y z w : Int) :
3030

3131
/--
3232
info: Try this:
33-
[apply]
34-
ac
35-
cases #5c4b <;> cases #896f <;> ac
33+
[apply] cases #5c4b <;> cases #896f <;> ac
3634
-/
3735
#guard_msgs in
3836
example {α : Type} (op : α → α → α) [Std.Associative op] [Std.Commutative op] (a b c d e : α) :
@@ -68,12 +66,10 @@ info: Try this:
6866
[apply]
6967
cases #c4b6
7068
next =>
71-
ring
7269
cases #8c9f
7370
next => ring
7471
next => sorry
7572
next =>
76-
ring
7773
cases #8c9f
7874
next => ring
7975
next => sorry
@@ -91,11 +87,28 @@ info: Try this:
9187
instantiate only [= Nat.min_def]
9288
cases #7640
9389
next => sorry
94-
next =>
95-
ring
96-
lia
90+
next => lia
9791
-/
9892
#guard_msgs in
9993
example (as : Array α) (lo hi i j : Nat) (h₁ : lo ≤ i) (_ : i < j) (_ : j ≤ hi) (_ : j < as.size)
10094
(_ : ¬as.size = 0) : min lo (as.size - 1) < i := by
10195
grind => finish?
96+
97+
/--
98+
info: Try this:
99+
[apply]
100+
instantiate only [= getMsbD_setWidth']
101+
cases #aa9d
102+
-/
103+
#guard_msgs in
104+
open BitVec in
105+
example (ge : m ≥ n) (x : BitVec n) (i : Nat) :
106+
getMsbD (setWidth' ge x) i = (decide (m - n ≤ i) && getMsbD x (i + n - m)) := by
107+
grind => finish?
108+
109+
open BitVec in
110+
example (ge : m ≥ n) (x : BitVec n) (i : Nat) :
111+
getMsbD (setWidth' ge x) i = (decide (m - n ≤ i) && getMsbD x (i + n - m)) := by
112+
grind =>
113+
instantiate only [= getMsbD_setWidth']
114+
cases #aa9d

0 commit comments

Comments
 (0)