Documentation

Mathlib.Tactic.Rewrites

The rewrites tactic. #

rw? tries to find a lemma which can rewrite the goal.

rw? should not be left in proofs; it is a search tool, like apply?.

Suggestions are printed as rw [h] or rw [← h].

Extract the lemma, with arguments, that was used to produce a RewriteResult.

Equations
Instances For

    Weight to multiply the "specificity" of a rewrite lemma by when rewriting forwards.

    Equations
    Instances For

      Weight to multiply the "specificity" of a rewrite lemma by when rewriting backwards.

      Equations
      Instances For

        Configuration for DiscrTree.

        Equations
        Instances For

          Prepare the discrimination tree entries for a lemma.

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

            Select = and local hypotheses.

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

              Construct the discrimination tree of all lemmas.

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

                  Data structure recording a potential rewrite to report from the rw? tactic.

                  • expr : Lean.Expr

                    The lemma we rewrote by. This is Expr, not just a Name, as it may be a local hypothesis.

                  • symm : Bool

                    True if we rewrote backwards (i.e. with rw [← h]).

                  • weight : Nat

                    The "weight" of the rewrite. This is calculated based on how specific the rewrite rule was.

                  • The result from the rw tactic.

                  • ppResult? : Option String

                    Pretty-printed result.

                  • rfl? : Option Bool

                    Can the new goal in result be closed by with_reducible rfl?

                  • The metavariable context after the rewrite. This needs to be stored as part of the result so we can backtrack the state.

                  Instances For

                    Update a RewriteResult by filling in the rfl? field if it is currently none, to reflect whether the remaining goal can be closed by with_reducible rfl.

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

                      Pretty print the result of the rewrite, and store it for later use.

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

                        Pretty print the result of the rewrite. If this will be done more than once you should use prepare_ppResult

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

                          Shortcut for calling solveByElim.

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

                            Should we try discharging side conditions? If so, using assumption, or solve_by_elim?

                            Instances For

                              Find lemmas which can rewrite the goal.

                              This core function returns a monadic list, to allow the caller to decide how long to search. See also rewrites for a more convenient interface.

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

                                Find lemmas which can rewrite the goal, and deduplicate based on pretty-printed results. Note that this builds a HashMap containing the results, and so may consume significant memory.

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

                                  Find lemmas which can rewrite the goal.

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

                                    Syntax for excluding some names, e.g. [-my_lemma, -my_theorem].

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

                                      rw? tries to find a lemma which can rewrite the goal.

                                      rw? should not be left in proofs; it is a search tool, like apply?.

                                      Suggestions are printed as rw [h] or rw [← h].

                                      You can use rw? [-my_lemma, -my_theorem] to prevent rw? using the named lemmas.

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