Documentation

Lean.Meta.Tactic.LibrarySearch

Library search #

This file defines tactics exact? and apply?, (formerly known as library_search) and a term elaborator exact?% that tries to find a lemma solving the current goal (subgoals are solved using solveByElim).

example : x < x + 1 := exact?%
example : Nat := by exact?

A discharger that tries grind on the goal. The proof is wrapped in Grind.Marker so that suggestions display (by grind) instead of the complex grind proof term. Returns some [] if grind succeeds, none otherwise (leaving the goal as a subgoal).

Equations
  • One or more equations did not get rendered due to their size.
Instances For

    A discharger that tries try? on the goal. The proof is wrapped in Try.Marker so that suggestions display (by try?) instead of the complex proof term. Returns some [] if try? succeeds, none otherwise (leaving the goal as a subgoal).

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      def Lean.Meta.LibrarySearch.solveByElim (required : List Expr) (exfalso : Bool) (goals : List MVarId) (maxDepth : Nat) (grind try? : Bool := false) :

      Wrapper for calling `Lean.Meta.SolveByElim.solveByElim with appropriate arguments for library search.

      If grind is true, grind will be used as a fallback discharger for subgoals that cannot be closed by other means. If try? is true, try? will be used as a fallback discharger (via grind internally).

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For

        A "modifier" for a declaration.

        • none indicates the original declaration,
        • mp indicates that (possibly after binders) the declaration is an , and we want to consider the forward direction,
        • mpr similarly, but for the backward direction.
        • none : DeclMod

          the original declaration

        • mp : DeclMod

          the forward direction of an iff

        • mpr : DeclMod

          the backward direction of an iff

        Instances For
          @[reducible]

          LibrarySearch has an extension mechanism for replacing the function used to find candidate lemmas.

          Equations
          Instances For

            We drop .star and Eq * * * from the discriminator trees because they match too much.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For

              Create function for finding relevant declarations. Also captures dropped entries in starLemmasExt for fallback search.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For

                Get star-indexed lemmas (lazily computed during tree initialization).

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For

                  Return an action that returns true when the remaining heartbeats is less than the currently remaining heartbeats * leavePercent / 100.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    def Lean.Meta.LibrarySearch.interleaveWith {α : Type u_1} {β : Type u_2} {γ : Type} (f : αγ) (x : Array α) (g : βγ) (y : Array β) :

                    Interleave x y interleaves the elements of x and y until one is empty and then returns final elements in other list.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      def Lean.Meta.LibrarySearch.abortSpeculation {m : Type u_1 → Type u_2} {α : Type u_1} [MonadExcept Exception m] :
                      m α

                      Called to abort speculative execution in library search.

                      Equations
                      Instances For
                        @[reducible]

                        A library search candidate using symmetry includes the goal to solve, the metavar context for that goal, and the name and orientation of a rule to try using with goal.

                        Equations
                        Instances For

                          Run searchFn on both the goal and symm applied to the goal.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Instances For

                            Create lemma from name and mod.

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For

                              Sequentially invokes a tactic act on each value in candidates on the current state.

                              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 (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.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                def Lean.Meta.LibrarySearch.librarySearch (goal : MVarId) (tactic : List MVarIdMetaM (List MVarId) := fun (g : List MVarId) => solveByElim [] false g 6) (allowFailure : MVarIdMetaM Bool := fun (x : MVarId) => pure true) (leavePercentHeartbeats : Nat := 10) (includeStar : Bool := true) (collectAll : Bool := false) :

                                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 (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.)

                                (Note that if solveByElim solves some but not all subsidiary goals, this is not currently tracked.)

                                Equations
                                • One or more equations did not get rendered due to their size.
                                Instances For