Skip to content

Commit 519ccf5

Browse files
kim-emclaude
andauthored
feat: add solve_by_elim +suggestions (#11468)
This PR adds `+suggestions` support to `solve_by_elim`, following the pattern established by `grind +suggestions` and `simp_all +suggestions`. Gracefully handles invalid/nonexistent suggestions by filtering them out 🤖 Prepared with Claude Code Co-authored-by: Claude <[email protected]>
1 parent 1c1c534 commit 519ccf5

File tree

3 files changed

+98
-0
lines changed

3 files changed

+98
-0
lines changed

src/Lean/Elab/Tactic/SolveByElim.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ module
88
prelude
99
public import Lean.Meta.Tactic.SolveByElim
1010
public import Lean.Elab.Tactic.Config
11+
public import Lean.LibrarySuggestions.Basic
1112

1213
public section
1314

@@ -108,6 +109,21 @@ def evalSolveByElim : Tactic
108109
else
109110
pure [← getMainGoal]
110111
let cfg ← elabConfig cfg
112+
-- Add library suggestions if +suggestions is enabled
113+
let add ← if cfg.suggestions then
114+
let mainGoal ← getMainGoal
115+
let suggestions ← LibrarySuggestions.select mainGoal { caller := some "solve_by_elim" }
116+
let suggestionTerms ← suggestions.toList.filterMapM fun s => do
117+
-- Only include suggestions for constants that exist
118+
let env ← getEnv
119+
if env.contains s.name then
120+
let ident := mkCIdentFrom (← getRef) s.name (canonical := true)
121+
return some (⟨ident⟩ : Term)
122+
else
123+
return none
124+
pure (add ++ suggestionTerms)
125+
else
126+
pure add
111127
let [] ← processSyntax cfg o.isSome star add remove use goals |
112128
throwError "Internal error: `solve_by_elim` unexpectedly returned subgoals"
113129
pure ()

src/Lean/Meta/Tactic/SolveByElim.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -100,6 +100,12 @@ structure SolveByElimConfig extends ApplyRulesConfig where
100100
intro : Bool := true
101101
/-- Try calling `constructor` if no lemmas apply. -/
102102
constructor : Bool := true
103+
/--
104+
If `suggestions` is `true`, `solve_by_elim` will invoke the currently configured library
105+
suggestion engine on the current goal, and attempt to use the resulting suggestions as
106+
additional lemmas.
107+
-/
108+
suggestions : Bool := false
103109

104110
namespace SolveByElimConfig
105111

Lines changed: 76 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,76 @@
1+
import Lean.LibrarySuggestions.Basic
2+
import Lean.Elab.Tactic.SolveByElim
3+
4+
-- First test without any library suggestions set up
5+
/--
6+
error: No library suggestions engine registered. (Add `import Lean.LibrarySuggestions.Default` to use Lean's built-in engine, or use `set_library_suggestions` to configure a custom one.)
7+
-/
8+
#guard_msgs in
9+
example : True := by
10+
solve_by_elim +suggestions
11+
12+
-- Set up an empty library suggestion engine
13+
set_library_suggestions (fun _ _ => pure #[])
14+
15+
#guard_msgs in
16+
example : True := by
17+
solve_by_elim +suggestions
18+
19+
-- A custom proposition that solve_by_elim doesn't know about by default
20+
-- We use a custom non-defeq type so that `trivial` and `rfl` don't solve it
21+
inductive MyProp : Nat → Prop where
22+
| intro37 : MyProp 37
23+
| intro42 : MyProp 42
24+
25+
theorem myPropThm37 : MyProp 37 := MyProp.intro37
26+
27+
-- Without +suggestions (or explicit lemma), solve_by_elim should fail
28+
-- Note: need -constructor to prevent constructor from solving it
29+
/--
30+
error: failed
31+
-/
32+
#guard_msgs in
33+
example : MyProp 37 := by
34+
solve_by_elim -constructor
35+
36+
-- With explicit lemma, it works
37+
example : MyProp 37 := by
38+
solve_by_elim only [myPropThm37]
39+
40+
-- Set up library suggestion engine that suggests our theorem
41+
set_library_suggestions (fun _ _ => pure #[{ name := `myPropThm37, score := 1.0 }])
42+
43+
-- With +suggestions, it should work
44+
example : MyProp 37 := by
45+
solve_by_elim -constructor +suggestions
46+
47+
-- Test that suggestions work with local hypotheses
48+
example (h : MyProp 42) : MyProp 42 := by
49+
solve_by_elim +suggestions
50+
51+
-- Test with chain of applications
52+
inductive IsZero : Nat → Prop where
53+
| intro : IsZero 0
54+
55+
def fromZero (n : Nat) (h : n = 0) : IsZero n := by
56+
subst h
57+
exact IsZero.intro
58+
59+
theorem isZero_zero : IsZero 0 := IsZero.intro
60+
61+
set_library_suggestions (fun _ _ => pure #[
62+
{ name := `isZero_zero, score := 1.0 }
63+
])
64+
65+
example : IsZero 0 := by
66+
solve_by_elim +suggestions
67+
68+
-- Test with bad suggestions - they should be ignored silently
69+
set_library_suggestions (fun _ _ => pure #[
70+
{ name := `isZero_zero, score := 1.0 },
71+
{ name := `nonexistent_theorem, score := 0.5 }
72+
])
73+
74+
-- Bad suggestions should be filtered out, and the good one should work
75+
example : IsZero 0 := by
76+
solve_by_elim +suggestions

0 commit comments

Comments
 (0)