Skip to content
Open
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
5 changes: 5 additions & 0 deletions src/Init/Tactics.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1704,6 +1704,9 @@ structure LibrarySearchConfig where
like `[*]`) are searched as a fallback when no concrete-keyed lemmas are found.
Use `-star` to disable this fallback. -/
star : Bool := true
/-- If true, collect all successful lemmas instead of stopping at the first complete solution.
Use `+all` to enable this behavior. -/
all : Bool := false

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

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

Expand Down
32 changes: 24 additions & 8 deletions src/Lean/Elab/Tactic/LibrarySearch.lean
Original file line number Diff line number Diff line change
Expand Up @@ -42,21 +42,37 @@ def exact? (ref : Syntax) (config : Parser.Tactic.LibrarySearchConfig)
let allowFailure := fun g => do
let g ← g.withContext (instantiateMVars (.mvar g))
return required.all fun e => e.occurs g
match (← librarySearch goal tactic allowFailure (includeStar := config.star)) with
-- Found goal that closed problem
match (← librarySearch goal tactic allowFailure (includeStar := config.star)
(collectAll := config.all)) with
-- Found goal that closed problem (only when collectAll = false)
| none =>
addExactSuggestion ref (← instantiateMVars (mkMVar mvar)).headBeta (checkState? := initialState)
-- Found suggestions
-- Found suggestions (includes complete solutions when collectAll = true)
| some suggestions =>
if requireClose then
-- Separate complete solutions (empty remaining goals) from incomplete ones
let (complete, incomplete) := suggestions.partition (·.1.isEmpty)
if requireClose && !config.all then
let hint := if suggestions.isEmpty then "" else " Try `apply?` to see partial suggestions."
throwError "`exact?` could not close the goal.{hint}"
if requireClose && config.all && complete.isEmpty then
let hint := if incomplete.isEmpty then "" else " Try `apply?` to see partial suggestions."
throwError "`exact?` could not close the goal.{hint}"
reportOutOfHeartbeats `apply? ref
for (_, suggestionMCtx) in suggestions do
-- Collect complete solutions and show as a single "Try these:" message
let completeExprs ← complete.mapM fun (_, suggestionMCtx) =>
withMCtx suggestionMCtx do
addExactSuggestion ref (← instantiateMVars (mkMVar mvar)).headBeta
(checkState? := initialState) (addSubgoalsMsg := true) (tacticErrorAsInfo := true)
if suggestions.isEmpty then logError "apply? didn't find any relevant lemmas"
return (← instantiateMVars (mkMVar mvar)).headBeta
if !completeExprs.isEmpty then
addExactSuggestions ref completeExprs (checkState? := initialState)
-- Show incomplete solutions only if not requireClose (i.e., for apply?)
-- Note: we must call addExactSuggestion inside withMCtx because incomplete
-- solutions have unassigned metavariables that are only valid in that context
if !requireClose then
for (_, suggestionMCtx) in incomplete do
withMCtx suggestionMCtx do
addExactSuggestion ref (← instantiateMVars (mkMVar mvar)).headBeta
(checkState? := initialState) (addSubgoalsMsg := true) (tacticErrorAsInfo := true)
if suggestions.isEmpty then logError "apply? didn't find any relevant lemmas"
admitGoal goal (synthetic := false)

@[builtin_tactic Lean.Parser.Tactic.exact?]
Expand Down
25 changes: 19 additions & 6 deletions src/Lean/Meta/LazyDiscrTree.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1042,10 +1042,12 @@ def findImportMatches
profileitM Exception "lazy discriminator import initialization" (←getOptions) $ do
let t ← createImportedDiscrTree (createTreeCtx cctx) cNGen (←getEnv) addEntry
(constantsPerTask := constantsPerTask)
-- If a reference is provided, extract and store dropped entries
-- If a reference is provided, extract and append dropped entries
if let some droppedRef := droppedEntriesRef then
let (extracted, t) ← extractKeys t droppedKeys
droppedRef.set (some extracted)
-- Append to existing dropped entries (e.g., from module tree)
let existing := (← droppedRef.get).getD #[]
droppedRef.set (some (existing ++ extracted))
pure t
else
dropKeys t droppedKeys
Expand Down Expand Up @@ -1077,12 +1079,23 @@ def createModuleDiscrTree

/--
Creates reference for lazy discriminator tree that only contains this module's definitions.
If `droppedEntriesRef` is provided, dropped entries (e.g., star-indexed lemmas) are extracted
and appended to the array in the reference.
-/
def createModuleTreeRef (entriesForConst : Name → ConstantInfo → MetaM (Array (InitEntry α)))
(droppedKeys : List (List LazyDiscrTree.Key)) : MetaM (ModuleDiscrTreeRef α) := do
(droppedKeys : List (List LazyDiscrTree.Key))
(droppedEntriesRef : Option (IO.Ref (Option (Array α))) := none) :
MetaM (ModuleDiscrTreeRef α) := do
profileitM Exception "build module discriminator tree" (←getOptions) $ do
let t ← createModuleDiscrTree entriesForConst
let t ← dropKeys t droppedKeys
let t ← if let some droppedRef := droppedEntriesRef then
let (extracted, t) ← extractKeys t droppedKeys
-- Append to existing dropped entries (if any)
let existing := (← droppedRef.get).getD #[]
droppedRef.set (some (existing ++ extracted))
pure t
else
dropKeys t droppedKeys
pure { ref := ← IO.mkRef t }

/--
Expand Down Expand Up @@ -1147,7 +1160,7 @@ def findMatches (ext : EnvExtension (IO.Ref (Option (LazyDiscrTree α))))
(constantsPerTask : Nat := 1000)
(droppedEntriesRef : Option (IO.Ref (Option (Array α))) := none)
(ty : Expr) : MetaM (Array α) := do

let moduleTreeRef ← createModuleTreeRef addEntry droppedKeys
-- Pass droppedEntriesRef to also capture star-indexed lemmas from the current module
let moduleTreeRef ← createModuleTreeRef addEntry droppedKeys droppedEntriesRef
let incPrio _ v := v
findMatchesExt moduleTreeRef ext addEntry droppedKeys constantsPerTask droppedEntriesRef incPrio ty
31 changes: 19 additions & 12 deletions src/Lean/Meta/Tactic/LibrarySearch.lean
Original file line number Diff line number Diff line change
Expand Up @@ -334,15 +334,17 @@ Sequentially invokes a tactic `act` on each value in candidates on the current s

The tactic `act` should return a list of meta-variables that still need to be resolved.
If this list is empty, then no variables remain to be solved, and `tryOnEach` returns
`none` with the environment set so each goal is resolved.
`none` with the environment set so each goal is resolved (unless `collectAll` is true,
in which case it continues searching and collects complete solutions in the array).

If the action throws an internal exception with the `abortSpeculationId` id then
further computation is stopped and intermediate results returned. If any other
exception is thrown, then it is silently discarded.
-/
def tryOnEach
(act : Candidate → MetaM (List MVarId))
(candidates : Array Candidate) :
(candidates : Array Candidate)
(collectAll : Bool := false) :
MetaM (Option (Array (List MVarId × MetavarContext))) := do
let mut a := #[]
let s ← saveState
Expand All @@ -353,7 +355,7 @@ def tryOnEach
if isAbortSpeculation e then
break
| .ok remaining =>
if remaining.isEmpty then
if remaining.isEmpty && !collectAll then
return none
let ctx ← getMCtx
restoreState s
Expand All @@ -364,7 +366,8 @@ private def librarySearch' (goal : MVarId)
(tactic : List MVarId → MetaM (List MVarId))
(allowFailure : MVarId → MetaM Bool)
(leavePercentHeartbeats : Nat)
(includeStar : Bool := true) :
(includeStar : Bool := true)
(collectAll : Bool := false) :
MetaM (Option (Array (List MVarId × MetavarContext))) := do
withTraceNode `Tactic.librarySearch (return m!"{librarySearchEmoji ·} {← goal.getType}") do
profileitM Exception "librarySearch" (← getOptions) do
Expand All @@ -376,32 +379,35 @@ private def librarySearch' (goal : MVarId)
librarySearchLemma cfg tactic allowFailure cand
-- First pass: search with droppedKeys (excludes star-indexed lemmas)
let candidates ← librarySearchSymm libSearchFindDecls goal
match ← tryOnEach act candidates with
| none => return none -- Found a complete solution
match ← tryOnEach act candidates collectAll with
| none => return none -- Found a complete solution (only when collectAll = false)
| some results =>
-- Only do star fallback if:
-- 1. No results from primary search
-- 1. No complete solutions from primary search (when collectAll, check for empty remaining)
-- 2. includeStar is true
if !results.isEmpty || !includeStar then
let hasCompleteSolution := results.any (·.1.isEmpty)
if hasCompleteSolution || !results.isEmpty || !includeStar then
return some results
-- Second pass: try star-indexed lemmas (those with [*] or [Eq,*,*,*] keys)
-- No need for librarySearchSymm since getStarLemmas ignores the goal type
let starLemmas ← getStarLemmas
if starLemmas.isEmpty then return some results
let mctx ← getMCtx
let starCandidates := starLemmas.map ((goal, mctx), ·)
tryOnEach act starCandidates
tryOnEach act starCandidates collectAll

/--
Tries to solve the goal by applying a library lemma
then calling `tactic` on the resulting goals.

Typically here `tactic` is `solveByElim`.

If it successfully closes the goal, returns `none`.
If it successfully closes the goal, returns `none` (unless `collectAll` is true).
Otherwise, it returns `some a`, where `a : Array (List MVarId × MetavarContext)`,
with an entry for each library lemma which was successfully applied,
containing a list of the subsidiary goals, and the metavariable context after the application.
When `collectAll` is true, complete solutions (with empty remaining goals) are also included
in the array instead of returning early.

(Always succeeds, and the metavariable context stored in the monad is reverted,
unless the goal was completely solved.)
Expand All @@ -414,9 +420,10 @@ def librarySearch (goal : MVarId)
fun g => solveByElim [] (maxDepth := 6) (exfalso := false) g)
(allowFailure : MVarId → MetaM Bool := fun _ => pure true)
(leavePercentHeartbeats : Nat := 10)
(includeStar : Bool := true) :
(includeStar : Bool := true)
(collectAll : Bool := false) :
MetaM (Option (Array (List MVarId × MetavarContext))) := do
librarySearch' goal tactic allowFailure leavePercentHeartbeats includeStar
librarySearch' goal tactic allowFailure leavePercentHeartbeats includeStar collectAll

end LibrarySearch

Expand Down
42 changes: 33 additions & 9 deletions tests/lean/librarySearch.lean
Original file line number Diff line number Diff line change
Expand Up @@ -90,16 +90,31 @@ info: Try this:
#guard_msgs in
example (X : Type) (P : Prop) (x : X) (h : ∀ x : X, x = x → P) : P := by show_term solve_by_elim

-- Could be any number of results (`fun x => x`, `id`, etc)
#guard_msgs (drop info) in
/--
info: Try this:
[apply] exact fun a => a
-/
#guard_msgs in
example (α : Prop) : α → α := by show_term solve_by_elim

-- These examples work via star-indexed fallback.
#guard_msgs (drop info) in
/--
info: Try this:
[apply] exact fun a => Classical.byContradiction a
-/
#guard_msgs in
example (p : Prop) : (¬¬p) → p := by apply?
#guard_msgs (drop info) in
/--
info: Try this:
[apply] exact h.left
-/
#guard_msgs in
example (a b : Prop) (h : a ∧ b) : a := by apply?
#guard_msgs (drop info) in
/--
info: Try this:
[apply] exact fun a a_1 => Classical.byContradiction fun a_2 => a a_2 a_1
-/
#guard_msgs in
example (P Q : Prop) : (¬ Q → ¬ P) → (P → Q) := by apply?

/--
Expand Down Expand Up @@ -264,10 +279,15 @@ info: Try this:
example {a b c : Nat} (h₁ : a ∣ c) (h₂ : a ∣ b + c) : a ∣ b := by apply?

-- These examples work via star-indexed fallback.
#guard_msgs (drop info) in
/--
info: Try this:
[apply] exact h.elim
-/
#guard_msgs in
example {α : Sort u} (h : Empty) : α := by apply?
#guard_msgs (drop info) in
example (f : A → C) (g : B → C) : (A ⊕ B) → C := by apply?

-- FIXME: this is timing out, what is it doing?
-- example (f : A → C) (g : B → C) : (A ⊕ B) → C := by apply?

opaque f : Nat → Nat
axiom F (a b : Nat) : f a ≤ f b ↔ a ≤ b
Expand Down Expand Up @@ -391,7 +411,11 @@ example : False := by apply?
example {α : Sort u} (h : Empty) : α := by apply? -star

-- With `+star`, we find `Empty.elim` via star-indexed lemma fallback.
#guard_msgs (drop info) in
/--
info: Try this:
[apply] exact h.elim
-/
#guard_msgs in
example {α : Sort u} (h : Empty) : α := by apply? +star

-- Verify that `-star` doesn't break normal (non-star-indexed) lemma search.
Expand Down
75 changes: 75 additions & 0 deletions tests/lean/run/library_search_all.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,75 @@
/-!
# Tests for `exact? +all`

This tests the `+all` option which collects all successful lemmas
instead of stopping at the first complete solution.
-/

set_option linter.unusedVariables false

-- Create a custom inductive with limited lemmas for controlled testing
inductive MyTrue : Prop where
| intro

-- Define exactly three ways to prove MyTrue
theorem myTrue1 : MyTrue := MyTrue.intro
theorem myTrue2 : MyTrue := MyTrue.intro
theorem myTrue3 : MyTrue := MyTrue.intro

-- Test: without +all, we get a single suggestion and the goal is solved
/--
info: Try this:
[apply] exact myTrue3
-/
#guard_msgs in
example : MyTrue := by exact?

-- Test: with +all, exact? finds all lemmas (including the constructor)
-- The goal is admitted (sorry), and all suggestions are shown in one message
/--
info: Try these:
[apply] exact myTrue3
[apply] exact MyTrue.intro
[apply] exact myTrue1
[apply] exact myTrue2
---
warning: declaration uses 'sorry'
-/
#guard_msgs in
example : MyTrue := by exact? +all

-- Test: apply? also supports +all
/--
info: Try these:
[apply] exact myTrue3
[apply] exact MyTrue.intro
[apply] exact myTrue1
[apply] exact myTrue2
---
warning: declaration uses 'sorry'
-/
#guard_msgs in
example : MyTrue := by apply? +all

-- Test: verify that +all without any matches still errors appropriately
/--
error: `exact?` could not close the goal. Try `apply?` to see partial suggestions.
-/
#guard_msgs in
example (n : Nat) (h : n = n + 1) : False := by exact? +all

def Empty.elim2 := @Empty.elim

/--
info: Try these:
[apply] exact h.elim
[apply] exact h.elim2
---
warning: declaration uses 'sorry'
-/
#guard_msgs in
example {α : Sort u} (h : Empty) : α := by exact? +all

/-- error: `exact?` could not close the goal. -/
#guard_msgs in
example {α : Sort u} (h : Empty) : α := by exact? +all -star
Loading