Documentation

Mathlib.Lean.Meta.RefinedDiscrTree.Lookup

Matching with a RefinedDiscrTree #

We use a very simple unification algorithm. For all star/metavariable patterns in the RefinedDiscrTree and in the target, we store the assignment, and when it is assigned again, we check that it is the same assignment.

If k is a key in children, return the corresponding Trie α. Otherwise return none.

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

    Return the possible Trie α that match with n metavariable.

    Return the possible Trie α that match with anything. We add 1 to the matching score when the key is .opaque, since this pattern is "harder" to match with.

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

      Return the possible Trie α that come from a Key.star, while keeping track of the Key.star assignments.

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

        Return the possible Trie α that match with e.

        @[specialize #[]]

        If e is not a metavariable, return the possible Trie α that exactly match with e.

        def Lean.Meta.RefinedDiscrTree.getMatchWithScore {α : Type} (d : RefinedDiscrTree α) (e : Expr) (unify : Bool) (allowRootStar : Bool := false) :

        Return the results from the RefinedDiscrTree that match the given expression, together with their matching scores, in decreasing order of score.

        Each entry of type Array α × Nat corresponds to one pattern.

        If unify := false, then metavariables in e are treated as opaque variables. This is for when you don't want to instantiate metavariables in e.

        If allowRootStar := false, then we don't allow e or the matched key in d to be a star pattern.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          def Lean.Meta.RefinedDiscrTree.getMatchWithScoreWithExtra {α : Type} (d : RefinedDiscrTree α) (e : Expr) (unify : Bool) (allowRootStar : Bool := false) :

          Similar to getMatchWithScore, but also returns matches with prefixes of e. We store the score, followed by the number of ignored arguments.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            partial def Lean.Meta.RefinedDiscrTree.getMatchWithScoreWithExtra.go {α : Type} (d : RefinedDiscrTree α) (unify : Bool) (allowRootStar : Bool := false) (e : Expr) (numIgnored : ) :

            go