Skip to content

Commit 206eb73

Browse files
authored
feat: finish? tactic for grind interactive mode (#10837)
This PR implements the `finish?` tactic for the `grind` interactive mode. When it successfully closes the goal, it produces a code action that allows the user to close the goal using explicit grind tactic steps, i.e., without any search. It also makes explicit which solvers have been used. This is just the first version, we will add many "bells and whistles" later. For example, `instantiate` steps will clearly show which theorems have been instantiated. Example: ```lean /-- info: Try this: [apply] ⏎ cases #b0f4 next => cases #50fc next => cases #50fc <;> lia -/ #guard_msgs in example (p : Nat → Prop) (x y z w : Int) : (x = 1 ∨ x = 2) → (w = 1 ∨ w = 4) → (y = 1 ∨ (∃ x : Nat, y = 3 - x ∧ p x)) → (z = 1 ∨ z = 0) → x + y ≤ 6 := by grind => finish? ``` The anchors in the generated script are based on stable hash codes. Moreover, users can hover over them to see the exact term used in the case split. `grind?` will also be implemented using the new framework.
1 parent 09f2220 commit 206eb73

File tree

6 files changed

+134
-2
lines changed

6 files changed

+134
-2
lines changed

src/Init/Grind/Interactive.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -100,6 +100,9 @@ syntax (name := done) "done" : grind
100100
/-- `finish` tries to close the current goal using `grind`'s default strategy -/
101101
syntax (name := finish) "finish" : grind
102102

103+
/-- `finish?` tries to close the current goal using `grind`'s default strategy and suggests a tactic script. -/
104+
syntax (name := finishTrace) "finish?" : grind
105+
103106
syntax (name := «have») "have" letDecl : grind
104107

105108
/-- Executes the given tactic block to close the current goal. -/

src/Lean/Elab/Tactic/Grind.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -10,3 +10,4 @@ public import Lean.Elab.Tactic.Grind.Basic
1010
public import Lean.Elab.Tactic.Grind.BuiltinTactic
1111
public import Lean.Elab.Tactic.Grind.ShowState
1212
public import Lean.Elab.Tactic.Grind.Have
13+
public import Lean.Elab.Tactic.Grind.Trace
Lines changed: 42 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,42 @@
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+
module
7+
prelude
8+
public import Lean.Elab.Tactic.Grind.Basic
9+
import Init.Grind.Interactive
10+
import Lean.Meta.Tactic.TryThis
11+
import Lean.Meta.Tactic.Grind.Action
12+
import Lean.Meta.Tactic.Grind.EMatchAction
13+
import Lean.Meta.Tactic.Grind.Split
14+
namespace Lean.Elab.Tactic.Grind
15+
open Meta
16+
open Meta.Grind
17+
18+
def withTracing (x : GrindTacticM α) : GrindTacticM α := do
19+
withReader (fun ctx => { ctx with ctx.config.trace := true }) x
20+
21+
def mkFinish (maxIterations : Nat) : IO Action := do
22+
let solvers ← Solvers.mkAction
23+
return Action.checkTactic >> (Action.done <|> solvers <|> Action.instantiate <|> Action.splitNext).loop maxIterations
24+
25+
def maxIterations := 1000 -- **TODO**: Add option
26+
27+
@[builtin_grind_tactic finishTrace] def evalFinishTrace : GrindTactic := fun stx => do
28+
let a ← mkFinish maxIterations
29+
let goal ← getMainGoal
30+
withTracing do
31+
match (← liftGrindM <| a.run goal) with
32+
| .closed seq =>
33+
replaceMainGoal []
34+
let seq := Action.mkGrindSeq seq
35+
Tactic.TryThis.addSuggestion stx { suggestion := .tsyntax seq }
36+
| .stuck gs =>
37+
let goal :: _ := gs | throwError "`finish?` failed, but resulting goal is not available"
38+
let params := (← read).params
39+
let result ← liftGrindM do mkResult params (some goal)
40+
throwError "`finish?` failed\n{← result.toMessageData}"
41+
42+
end Lean.Elab.Tactic.Grind

src/Lean/Meta/Tactic/Grind.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -43,9 +43,8 @@ public import Lean.Meta.Tactic.Grind.PropagateInj
4343
public import Lean.Meta.Tactic.Grind.Order
4444
public import Lean.Meta.Tactic.Grind.Anchor
4545
public import Lean.Meta.Tactic.Grind.Action
46-
46+
public import Lean.Meta.Tactic.Grind.EMatchAction
4747
public section
48-
4948
namespace Lean
5049

5150
/-! Trace options for `grind` users -/
Lines changed: 25 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,25 @@
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+
module
7+
prelude
8+
public import Lean.Meta.Tactic.Grind.Action
9+
public import Lean.Meta.Tactic.Grind.Intro
10+
import Lean.Meta.Tactic.Grind.EMatch
11+
public section
12+
namespace Lean.Meta.Grind.Action
13+
14+
def instantiate' : Action := fun goal kna kp => do
15+
let (progress, goal') ← GoalM.run goal ematch
16+
-- **TODO**: filter relevant instances
17+
if progress then
18+
concatTactic (← kp goal') `(grind| instantiate)
19+
else
20+
kna goal
21+
22+
def instantiate : Action :=
23+
instantiate' >> assertAll
24+
25+
end Lean.Meta.Grind.Action
Lines changed: 62 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,62 @@
1+
open Lean Grind
2+
3+
/--
4+
info: Try this:
5+
[apply] cases #c4b6 <;> ring <;> cases #4c68 <;> ring
6+
-/
7+
#guard_msgs in
8+
example {α : Type} [CommRing α] (a b c d e : α) :
9+
(a * a = b * c ∨ a^2 = c * b) →
10+
(a^2 = c * b ∨ e^2 = d * c) →
11+
(b^2 = d*c ∨ b^2 = c*d) →
12+
a*b*(b*a) = c^2*b*d := by
13+
grind => finish?
14+
15+
16+
/--
17+
info: Try this:
18+
[apply]
19+
cases #b0f4
20+
next => cases #50fc
21+
next => cases #50fc <;> lia
22+
-/
23+
#guard_msgs in
24+
example (p : Nat → Prop) (x y z w : Int) :
25+
(x = 1 ∨ x = 2) →
26+
(w = 1 ∨ w = 4) →
27+
(y = 1 ∨ (∃ x : Nat, y = 3 - x ∧ p x)) →
28+
(z = 1 ∨ z = 0) → x + y ≤ 6 := by
29+
grind => finish?
30+
31+
/--
32+
info: Try this:
33+
[apply]
34+
ac
35+
cases #5c4b <;> cases #896f <;> ac
36+
-/
37+
#guard_msgs in
38+
example {α : Type} (op : α → α → α) [Std.Associative op] [Std.Commutative op] (a b c d e : α) :
39+
(op a a = op b c ∨ op a a = op c b) →
40+
(op a a = op c b ∨ op e e = op d c) →
41+
(op b b = op d c ∨ op b b = op c d) →
42+
op (op a b) (op b a) = op (op c c) (op b d) := by
43+
grind => finish?
44+
45+
/--
46+
info: Try this:
47+
[apply]
48+
instantiate
49+
instantiate
50+
-/
51+
#guard_msgs in
52+
example (as bs cs : Array α) (v₁ v₂ : α)
53+
(i₁ i₂ j : Nat)
54+
(h₁ : i₁ < as.size)
55+
(h₂ : bs = as.set i₁ v₁)
56+
(h₃ : i₂ < bs.size)
57+
(h₃ : cs = bs.set i₂ v₂)
58+
(h₄ : i₁ ≠ j ∧ i₂ ≠ j)
59+
(h₅ : j < cs.size)
60+
(h₆ : j < as.size)
61+
: cs[j] = as[j] := by
62+
grind => finish?

0 commit comments

Comments
 (0)