Skip to content

Commit ae24d01

Browse files
kim-emclaude
andcommitted
feat: add +all option to exact? and apply?
This PR adds a `+all` option to `exact?` and `apply?` that collects all successful lemmas instead of stopping at the first complete solution. This is useful when exploring what lemmas are available to prove a goal. 🤖 Generated with [Claude Code](https://claude.com/claude-code) Co-Authored-By: Claude <[email protected]>
1 parent 3c100ad commit ae24d01

File tree

4 files changed

+115
-19
lines changed

4 files changed

+115
-19
lines changed

src/Init/Tactics.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1704,6 +1704,9 @@ structure LibrarySearchConfig where
17041704
like `[*]`) are searched as a fallback when no concrete-keyed lemmas are found.
17051705
Use `-star` to disable this fallback. -/
17061706
star : Bool := true
1707+
/-- If true, collect all successful lemmas instead of stopping at the first complete solution.
1708+
Use `+all` to enable this behavior. -/
1709+
all : Bool := false
17071710

17081711
/--
17091712
Searches environment for definitions or theorems that can solve the goal using `exact`
@@ -1716,6 +1719,7 @@ ways to resolve the goal, and one wants to guide which lemma is used.
17161719
Use `+grind` to enable `grind` as a fallback discharger for subgoals.
17171720
Use `+try?` to enable `try?` as a fallback discharger for subgoals.
17181721
Use `-star` to disable fallback to star-indexed lemmas (like `Empty.elim`, `And.left`).
1722+
Use `+all` to collect all successful lemmas instead of stopping at the first.
17191723
-/
17201724
syntax (name := exact?) "exact?" optConfig (" using " (colGt ident),+)? : tactic
17211725

@@ -1729,6 +1733,7 @@ used when closing the goal.
17291733
Use `+grind` to enable `grind` as a fallback discharger for subgoals.
17301734
Use `+try?` to enable `try?` as a fallback discharger for subgoals.
17311735
Use `-star` to disable fallback to star-indexed lemmas.
1736+
Use `+all` to collect all successful lemmas instead of stopping at the first.
17321737
-/
17331738
syntax (name := apply?) "apply?" optConfig (" using " (colGt term),+)? : tactic
17341739

src/Lean/Elab/Tactic/LibrarySearch.lean

Lines changed: 20 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -42,21 +42,34 @@ def exact? (ref : Syntax) (config : Parser.Tactic.LibrarySearchConfig)
4242
let allowFailure := fun g => do
4343
let g ← g.withContext (instantiateMVars (.mvar g))
4444
return required.all fun e => e.occurs g
45-
match (← librarySearch goal tactic allowFailure (includeStar := config.star)) with
46-
-- Found goal that closed problem
45+
match (← librarySearch goal tactic allowFailure (includeStar := config.star)
46+
(collectAll := config.all)) with
47+
-- Found goal that closed problem (only when collectAll = false)
4748
| none =>
4849
addExactSuggestion ref (← instantiateMVars (mkMVar mvar)).headBeta (checkState? := initialState)
49-
-- Found suggestions
50+
-- Found suggestions (includes complete solutions when collectAll = true)
5051
| some suggestions =>
51-
if requireClose then
52+
-- Separate complete solutions (empty remaining goals) from incomplete ones
53+
let (complete, incomplete) := suggestions.partition (·.1.isEmpty)
54+
if requireClose && !config.all then
5255
let hint := if suggestions.isEmpty then "" else " Try `apply?` to see partial suggestions."
5356
throwError "`exact?` could not close the goal.{hint}"
57+
if requireClose && config.all && complete.isEmpty then
58+
let hint := if incomplete.isEmpty then "" else " Try `apply?` to see partial suggestions."
59+
throwError "`exact?` could not close the goal.{hint}"
5460
reportOutOfHeartbeats `apply? ref
55-
for (_, suggestionMCtx) in suggestions do
61+
-- Show complete solutions (without subgoals message)
62+
for (_, suggestionMCtx) in complete do
5663
withMCtx suggestionMCtx do
5764
addExactSuggestion ref (← instantiateMVars (mkMVar mvar)).headBeta
58-
(checkState? := initialState) (addSubgoalsMsg := true) (tacticErrorAsInfo := true)
59-
if suggestions.isEmpty then logError "apply? didn't find any relevant lemmas"
65+
(checkState? := initialState) (tacticErrorAsInfo := true)
66+
-- Show incomplete solutions only if not requireClose (i.e., for apply?)
67+
if !requireClose then
68+
for (_, suggestionMCtx) in incomplete do
69+
withMCtx suggestionMCtx do
70+
addExactSuggestion ref (← instantiateMVars (mkMVar mvar)).headBeta
71+
(checkState? := initialState) (addSubgoalsMsg := true) (tacticErrorAsInfo := true)
72+
if suggestions.isEmpty then logError "apply? didn't find any relevant lemmas"
6073
admitGoal goal (synthetic := false)
6174

6275
@[builtin_tactic Lean.Parser.Tactic.exact?]

src/Lean/Meta/Tactic/LibrarySearch.lean

Lines changed: 19 additions & 12 deletions
Original file line numberDiff line numberDiff line change
@@ -334,15 +334,17 @@ Sequentially invokes a tactic `act` on each value in candidates on the current s
334334
335335
The tactic `act` should return a list of meta-variables that still need to be resolved.
336336
If this list is empty, then no variables remain to be solved, and `tryOnEach` returns
337-
`none` with the environment set so each goal is resolved.
337+
`none` with the environment set so each goal is resolved (unless `collectAll` is true,
338+
in which case it continues searching and collects complete solutions in the array).
338339
339340
If the action throws an internal exception with the `abortSpeculationId` id then
340341
further computation is stopped and intermediate results returned. If any other
341342
exception is thrown, then it is silently discarded.
342343
-/
343344
def tryOnEach
344345
(act : Candidate → MetaM (List MVarId))
345-
(candidates : Array Candidate) :
346+
(candidates : Array Candidate)
347+
(collectAll : Bool := false) :
346348
MetaM (Option (Array (List MVarId × MetavarContext))) := do
347349
let mut a := #[]
348350
let s ← saveState
@@ -353,7 +355,7 @@ def tryOnEach
353355
if isAbortSpeculation e then
354356
break
355357
| .ok remaining =>
356-
if remaining.isEmpty then
358+
if remaining.isEmpty && !collectAll then
357359
return none
358360
let ctx ← getMCtx
359361
restoreState s
@@ -364,7 +366,8 @@ private def librarySearch' (goal : MVarId)
364366
(tactic : List MVarId → MetaM (List MVarId))
365367
(allowFailure : MVarId → MetaM Bool)
366368
(leavePercentHeartbeats : Nat)
367-
(includeStar : Bool := true) :
369+
(includeStar : Bool := true)
370+
(collectAll : Bool := false) :
368371
MetaM (Option (Array (List MVarId × MetavarContext))) := do
369372
withTraceNode `Tactic.librarySearch (return m!"{librarySearchEmoji ·} {← goal.getType}") do
370373
profileitM Exception "librarySearch" (← getOptions) do
@@ -376,32 +379,35 @@ private def librarySearch' (goal : MVarId)
376379
librarySearchLemma cfg tactic allowFailure cand
377380
-- First pass: search with droppedKeys (excludes star-indexed lemmas)
378381
let candidates ← librarySearchSymm libSearchFindDecls goal
379-
match ← tryOnEach act candidates with
380-
| none => return none -- Found a complete solution
382+
match ← tryOnEach act candidates collectAll with
383+
| none => return none -- Found a complete solution (only when collectAll = false)
381384
| some results =>
382385
-- Only do star fallback if:
383-
-- 1. No results from primary search
386+
-- 1. No complete solutions from primary search (when collectAll, check for empty remaining)
384387
-- 2. includeStar is true
385-
if !results.isEmpty || !includeStar then
388+
let hasCompleteSolution := results.any (·.1.isEmpty)
389+
if hasCompleteSolution || !results.isEmpty || !includeStar then
386390
return some results
387391
-- Second pass: try star-indexed lemmas (those with [*] or [Eq,*,*,*] keys)
388392
-- No need for librarySearchSymm since getStarLemmas ignores the goal type
389393
let starLemmas ← getStarLemmas
390394
if starLemmas.isEmpty then return some results
391395
let mctx ← getMCtx
392396
let starCandidates := starLemmas.map ((goal, mctx), ·)
393-
tryOnEach act starCandidates
397+
tryOnEach act starCandidates collectAll
394398

395399
/--
396400
Tries to solve the goal by applying a library lemma
397401
then calling `tactic` on the resulting goals.
398402
399403
Typically here `tactic` is `solveByElim`.
400404
401-
If it successfully closes the goal, returns `none`.
405+
If it successfully closes the goal, returns `none` (unless `collectAll` is true).
402406
Otherwise, it returns `some a`, where `a : Array (List MVarId × MetavarContext)`,
403407
with an entry for each library lemma which was successfully applied,
404408
containing a list of the subsidiary goals, and the metavariable context after the application.
409+
When `collectAll` is true, complete solutions (with empty remaining goals) are also included
410+
in the array instead of returning early.
405411
406412
(Always succeeds, and the metavariable context stored in the monad is reverted,
407413
unless the goal was completely solved.)
@@ -414,9 +420,10 @@ def librarySearch (goal : MVarId)
414420
fun g => solveByElim [] (maxDepth := 6) (exfalso := false) g)
415421
(allowFailure : MVarId → MetaM Bool := fun _ => pure true)
416422
(leavePercentHeartbeats : Nat := 10)
417-
(includeStar : Bool := true) :
423+
(includeStar : Bool := true)
424+
(collectAll : Bool := false) :
418425
MetaM (Option (Array (List MVarId × MetavarContext))) := do
419-
librarySearch' goal tactic allowFailure leavePercentHeartbeats includeStar
426+
librarySearch' goal tactic allowFailure leavePercentHeartbeats includeStar collectAll
420427

421428
end LibrarySearch
422429

Lines changed: 71 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,71 @@
1+
/-!
2+
# Tests for `exact? +all`
3+
4+
This tests the `+all` option which collects all successful lemmas
5+
instead of stopping at the first complete solution.
6+
-/
7+
8+
set_option linter.unusedVariables false
9+
10+
-- Create a custom inductive with limited lemmas for controlled testing
11+
inductive MyTrue : Prop where
12+
| intro
13+
14+
-- Define exactly three ways to prove MyTrue
15+
theorem myTrue1 : MyTrue := MyTrue.intro
16+
theorem myTrue2 : MyTrue := MyTrue.intro
17+
theorem myTrue3 : MyTrue := MyTrue.intro
18+
19+
-- Test: without +all, we get a single suggestion and the goal is solved
20+
/--
21+
info: Try this:
22+
[apply] exact myTrue3
23+
-/
24+
#guard_msgs in
25+
example : MyTrue := by exact?
26+
27+
-- Test: with +all, exact? finds all lemmas (including the constructor)
28+
-- The goal is admitted (sorry), and all suggestions are shown
29+
/--
30+
info: Try this:
31+
[apply] exact myTrue3
32+
---
33+
info: Try this:
34+
[apply] exact MyTrue.intro
35+
---
36+
info: Try this:
37+
[apply] exact myTrue1
38+
---
39+
info: Try this:
40+
[apply] exact myTrue2
41+
---
42+
warning: declaration uses 'sorry'
43+
-/
44+
#guard_msgs in
45+
example : MyTrue := by exact? +all
46+
47+
-- Test: apply? also supports +all
48+
/--
49+
info: Try this:
50+
[apply] exact myTrue3
51+
---
52+
info: Try this:
53+
[apply] exact MyTrue.intro
54+
---
55+
info: Try this:
56+
[apply] exact myTrue1
57+
---
58+
info: Try this:
59+
[apply] exact myTrue2
60+
---
61+
warning: declaration uses 'sorry'
62+
-/
63+
#guard_msgs in
64+
example : MyTrue := by apply? +all
65+
66+
-- Test: verify that +all without any matches still errors appropriately
67+
/--
68+
error: `exact?` could not close the goal. Try `apply?` to see partial suggestions.
69+
-/
70+
#guard_msgs in
71+
example (n : Nat) (h : n = n + 1) : False := by exact? +all

0 commit comments

Comments
 (0)