Documentation

Mathlib.Lean.Meta.RefinedDiscrTree.Encode

Encoding an Expr as a sequence of Keys #

DTExpr is a simplified form of Expr. It is the intermediate step for converting from Expr to Array Key.

Instances For

    Return the size of the DTExpr. This is used for calculating the matching score when two expressions are equal. The score is not incremented at a lambda, which is so that the expressions ∀ x, p[x] and ∃ x, p[x] get the same size.

    Determine if two DTExprs are equivalent.

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

        Encoding an Expr #

        Given a DTExpr, return the linearized encoding in terms of Key, which is used for RefinedDiscrTree indexing.

        Equations
        Instances For

          Reduction procedure for the RefinedDiscrTree indexing.

          @[specialize #[]]
          partial def Lean.Meta.RefinedDiscrTree.lambdaTelescopeReduce {α : Type} {m : TypeType u_1} [Monad m] [MonadLiftT MetaM m] [MonadControlT MetaM m] [Nonempty α] (e : Expr) (fvars : List FVarId) (k : ExprList FVarIdm α) :
          m α

          Repeatedly apply reduce while stripping lambda binders and introducing their variables

          Check whether the expression is represented by Key.star and has arg as an argument.

          Equations
          Instances For

            Return true if e contains a loose bound variable.

            Equations
            Instances For

              Return for each argument whether it should be ignored.

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

                Return whether the argument should be ignored.

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

                  Return the encoding of e as a DTExpr. If root = false, then e is a strict sub expression of the original expression.

                  @[specialize #[]]

                  Return all pairs of body, bound variables that could possibly appear due to η-reduction

                  Equations
                  Instances For
                    @[specialize #[]]

                    run etaPossibilities, and cache the result if there are multiple possibilities.

                    Equations
                    Instances For

                      Return all encodings of e as a DTExpr, taking possible η-reductions into account. If root = false, then e is a strict sub expression of the original expression.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        def Lean.Meta.RefinedDiscrTree.mkDTExpr (e : Expr) (fvarInContext : FVarIdBool := fun (x : FVarId) => false) :

                        Return the encoding of e as a DTExpr.

                        Warning: to account for potential η-reductions of e, use mkDTExprs instead.

                        The argument fvarInContext allows you to specify which free variables in e will still be in the context when the RefinedDiscrTree is being used for lookup. It should return true only if the RefinedDiscrTree is built and used locally.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          def Lean.Meta.RefinedDiscrTree.mkDTExprs (e : Expr) (onlySpecific : Bool) (fvarInContext : FVarIdBool := fun (x : FVarId) => false) :

                          Similar to mkDTExpr. Return all encodings of e as a DTExpr, taking potential further η-reductions into account.

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