We define discrimination trees for the purpose of unifying local expressions with library results.
This structure is based on Lean.Meta.DiscrTree.
I document here what features are not in the original:
- The keys - Key.lam,- Key.foralland- Key.bvarhave been introduced in order to allow for matching under lambda and forall binders.- Key.lamhas arity 1 and indexes the body.- Key.forallhas arity 2 and indexes the domain and the body. The reason for not indexing the domain of a lambda expression is that it is usually already determined, for example in- ∃ a : α, p, which is- @Exists α fun a : α => p, we don't want to index the domain- αtwice. In a forall expression it is necessary to index the domain, because in an implication- p → qwe need to index both- pand- q.- Key.bvarworks the same as- Key.fvar, but stores the De Bruijn index to identify it.- For example, this allows for more specific matching with the left hand side of - ∑ i ∈ range n, i = n * (n - 1) / 2, which is indexed by- [⟨Finset.sum, 5⟩, ⟨Nat, 0⟩, ⟨Nat, 0⟩, *0, ⟨Finset.Range, 1⟩, *1, λ, ⟨#0, 0⟩].
- The key - Key.startakes a- Natidentifier as an argument. For example, the library pattern- ?a + ?ais encoded as- [⟨HAdd.hAdd, 6⟩, *0, *0, *0, *1, *2, *2].- *0corresponds to the type of- a,- *1to the- HAddinstance, and- *2to- a. This means that it will only match an expression- x + yif- xis definitionally equal to- y. The matching algorithm requires that the same stars from the discrimination tree match with the same patterns in the lookup expression, and similarly requires that the same metavariables form the lookup expression match with the same pattern in the discrimination tree.
- The key - Key.opaquehas been introduced in order to index existential variables in lemmas like- Nat.exists_prime_and_dvd {n : ℕ} (hn : n ≠ 1) : ∃ p, Prime p ∧ p ∣ n, where the part- Prime pgets the pattern- [⟨Nat.Prime, 1⟩, ◾]. (◾ represents- Key.opaque) When matching,- Key.opaquecan only be matched by- Key.star.- Using the - WhnfCoreConfigargument, it is possible to disable β-reduction and ζ-reduction. As a result, we may get a lambda expression applied to an argument or a let-expression. Since there is no support for indexing these, they will be indexed by- Key.opaque.
- We keep track of the matching score of a unification. This score represents the number of keys that had to be the same for the unification to succeed. For example, matching - (1 + 2) + 3with- add_commgives a score of 2, since the pattern of commutativity is [⟨HAdd.hAdd, 6⟩, *0, *0, *0, *1, *2, *3], so matching- ⟨HAdd.hAdd, 6⟩gives 1 point, and matching- *0after its first appearence gives another point, but the third argument is an outParam, so this gets ignored. Similarly, matching it with- add_assocgives a score of 5.
- Patterns that have the potential to be η-reduced are put into the - RefinedDiscrTreeunder all possible reduced key sequences. This is for terms of the form- fun x => f (?m x₁ .. xₙ), where- ?mis a metavariable, and one of- x₁, .., xₙin- x. For example, the pattern- Continuous fun y => Real.exp (f y)])is indexed by both- [⟨Continuous, 5⟩, *0, ⟨Real, 0⟩, *1, *2, λ, ⟨Real.exp⟩, *3]and- [⟨Continuous, 5⟩, *0, ⟨Real, 0⟩, *1, *2, ⟨Real.exp⟩]so that it also comes up if you search with- Continuous Real.exp. Similarly,- Continuous fun x => f x + g xis indexed by both- [⟨Continuous, 1⟩, λ, ⟨HAdd.hAdd, 6⟩, *0, *0, *0, *1, *2, *3]and- [⟨Continuous, 1⟩, ⟨HAdd.hAdd, 5⟩, *0, *0, *0, *1, *2].
- For sub-expressions not at the root of the original expression we have some additional reductions: - Any combination of ofNat,Nat.zero,Nat.succand number literals is stored as just a number literal.
- The expression fun a : α => ais stored as@id α.- This makes lemmata such as continuous_id'redundant, which is the same ascontinuous_id, withidreplaced byfun x => x.
 
- This makes lemmata such as 
- Any expressions involving +,*,-,/or⁻¹is normalized to not have a lambda in front and to always have the default amount of arguments. e.g.(f + g) ais stored asf a + g aandfun x => f x + g xis stored asf + g.- This makes lemmata such as MeasureTheory.integral_integral_add'redundant, which is the same asMeasureTheory.integral_integral_add, withf a + g areplaced by(f + g) a
- it also means that a lemma like Continuous.mulcan be stated as talking aboutf * ginstead offun x => f x + g x.
 
- This makes lemmata such as 
 
- Any combination of 
I have also made some changes in the implementation:
- Instead of directly converting from ExprtoArray Keyduring insertion, and directly looking up from anExprduring lookup, I defined the intermediate structureDTExpr, which is a form ofExprthat only contains information relevant for the discrimination tree. EachExpris transformed into aDTExprbefore insertion or lookup. For insertion there could be multipleDTExprrepresentations due to potential η-reductions as mentioned above.
TODO:
- More thought could be put into the matching algorithm for non-trivial unifications. For example, when looking up the expression - ?a + ?a(for rewriting), there will only be results like- n + n = 2 * nor- a + b = b + a, but not like- n + 1 = n.succ, even though this would still unify.
- The reason why implicit arguments are not ignored by the discrimination tree is that they provide important type information. Because of this it seems more natural to index the types of expressions instead of indexing the implicit type arguments. Then each key would additionally index the type of that expression. So instead of indexing - ?a + ?bas- [⟨HAdd.hAdd, 6⟩, *0, *0, *0, *1, *2, *3], it would be indexed by something like- [(*0, ⟨HAdd.hAdd, 6⟩), _, _, _, _, (*0, *1), (*0, *2)]. The advantage of this would be that there will be less duplicate indexing of types, because many functions index the types of their arguments and their return type with implicit arguments, meaning that types unnecessarily get indexed multiple times. This modification can be explored, but it could very well not be an improvement.
Definitions #
Discrimination tree key.
- star: ℕ → Mathlib.Meta.FunProp.RefinedDiscrTree.KeyA metavariable. This key matches with anything. It stores an index. 
- opaque: Mathlib.Meta.FunProp.RefinedDiscrTree.KeyAn opaque variable. This key only matches with itself or Key.star.
- const: Lean.Name → ℕ → Mathlib.Meta.FunProp.RefinedDiscrTree.KeyA constant. It stores the name and the arity. 
- fvar: Lean.FVarId → ℕ → Mathlib.Meta.FunProp.RefinedDiscrTree.KeyA free variable. It stores the FVarIdand the arity.
- bvar: ℕ → ℕ → Mathlib.Meta.FunProp.RefinedDiscrTree.KeyA bound variable, from a lambda or forall binder. It stores the De Bruijn index and the arity. 
- lit: Lean.Literal → Mathlib.Meta.FunProp.RefinedDiscrTree.KeyA literal. 
- sort: Mathlib.Meta.FunProp.RefinedDiscrTree.KeyA sort. Universe levels are ignored. 
- lam: Mathlib.Meta.FunProp.RefinedDiscrTree.KeyA lambda function. 
- forall: Mathlib.Meta.FunProp.RefinedDiscrTree.KeyA dependent arrow. 
- proj: Lean.Name → ℕ → ℕ → Mathlib.Meta.FunProp.RefinedDiscrTree.KeyA projection. It stores the structure name, the projection index and the arity. 
Instances For
Constructor index used for ordering Key.
Note that the index of the star pattern is 0, so that when looking up in a Trie,
we can look at the start of the sorted array for all .star patterns.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Return the number of arguments that the Key takes.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Discrimination tree trie. See RefinedDiscrTree.
- node: {α : Type} → Array (Mathlib.Meta.FunProp.RefinedDiscrTree.Key × Mathlib.Meta.FunProp.RefinedDiscrTree.Trie α) → Mathlib.Meta.FunProp.RefinedDiscrTree.Trie α
- path: {α : Type} →
  Array Mathlib.Meta.FunProp.RefinedDiscrTree.Key →
    Mathlib.Meta.FunProp.RefinedDiscrTree.Trie α → Mathlib.Meta.FunProp.RefinedDiscrTree.Trie αSequence of nodes with only one child. keysis anArrayof size at least 1.
- values: {α : Type} → Array α → Mathlib.Meta.FunProp.RefinedDiscrTree.Trie α
Instances For
Equations
- Mathlib.Meta.FunProp.RefinedDiscrTree.instInhabitedTrie = { default := Mathlib.Meta.FunProp.RefinedDiscrTree.Trie.node #[] }
Trie.path constructor that only inserts the path if it is non-empty.
Equations
- Mathlib.Meta.FunProp.RefinedDiscrTree.Trie.mkPath keys child = if keys.isEmpty = true then child else Mathlib.Meta.FunProp.RefinedDiscrTree.Trie.path keys child
Instances For
Trie constructor for a single value, taking the keys starting at index i.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Trie.node constructor for combining two Key, Trie α pairs.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Return the values from a Trie α, assuming that it is a leaf
Equations
- One or more equations did not get rendered due to their size.
Instances For
Return the children of a Trie α, assuming that it is not a leaf.
The result is sorted by the Key's
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Mathlib.Meta.FunProp.RefinedDiscrTree.instToFormatTrie = { format := Mathlib.Meta.FunProp.RefinedDiscrTree.Trie.format }
Discrimination tree. It is an index from expressions to values of type α.
- root : Lean.PersistentHashMap Mathlib.Meta.FunProp.RefinedDiscrTree.Key (Mathlib.Meta.FunProp.RefinedDiscrTree.Trie α)The underlying PersistentHashMapof aRefinedDiscrTree.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- Mathlib.Meta.FunProp.RefinedDiscrTree.instToFormat = { format := Mathlib.Meta.FunProp.RefinedDiscrTree.format }
DTExpr is a simplified form of Expr.
It is the intermediate step for converting from Expr to Array Key.
- star: Option Lean.MVarId → Mathlib.Meta.FunProp.RefinedDiscrTree.DTExprA metavariable. It optionally stores an MVarId.
- opaque: Mathlib.Meta.FunProp.RefinedDiscrTree.DTExprAn opaque variable or a let-expression in the case WhnfCoreConfig.zeta := false.
- const: Lean.Name → Array Mathlib.Meta.FunProp.RefinedDiscrTree.DTExpr → Mathlib.Meta.FunProp.RefinedDiscrTree.DTExprA constant. It stores the name and the arguments. 
- fvar: Lean.FVarId → Array Mathlib.Meta.FunProp.RefinedDiscrTree.DTExpr → Mathlib.Meta.FunProp.RefinedDiscrTree.DTExprA free variable. It stores the FVarIdand the argumenst
- bvar: ℕ → Array Mathlib.Meta.FunProp.RefinedDiscrTree.DTExpr → Mathlib.Meta.FunProp.RefinedDiscrTree.DTExprA bound variable. It stores the De Bruijn index and the arguments 
- lit: Lean.Literal → Mathlib.Meta.FunProp.RefinedDiscrTree.DTExprA literal. 
- sort: Mathlib.Meta.FunProp.RefinedDiscrTree.DTExprA sort. 
- lam: Mathlib.Meta.FunProp.RefinedDiscrTree.DTExpr → Mathlib.Meta.FunProp.RefinedDiscrTree.DTExprA lambda function. It stores the body. 
- forall: Mathlib.Meta.FunProp.RefinedDiscrTree.DTExpr →
  Mathlib.Meta.FunProp.RefinedDiscrTree.DTExpr → Mathlib.Meta.FunProp.RefinedDiscrTree.DTExprA dependent arrow. It stores the domain and body. 
- proj: Lean.Name →
  ℕ →
    Mathlib.Meta.FunProp.RefinedDiscrTree.DTExpr →
      Array Mathlib.Meta.FunProp.RefinedDiscrTree.DTExpr → Mathlib.Meta.FunProp.RefinedDiscrTree.DTExprA projection. It stores the structure name, projection index, struct body and arguments. 
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.
Encoding an Expr #
Given a DTExpr, return the linearized encoding in terms of Key,
which is used for RefinedDiscrTree indexing.
Equations
- e.flatten initCapacity = StateT.run' (Mathlib.Meta.FunProp.RefinedDiscrTree.DTExpr.flattenAux (Array.mkEmpty initCapacity) e) { stars := #[] }
Instances For
Reduction procedure for the RefinedDiscrTree indexing.
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
- One or more equations did not get rendered due to their size.
- Mathlib.Meta.FunProp.RefinedDiscrTree.isStarWithArg arg x = false
Instances For
Return true if e contains a loose bound variable.
Equations
- e.hasLooseBVars = Mathlib.Meta.FunProp.RefinedDiscrTree.DTExpr.hasLooseBVarsAux 0 e
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
Introduce new lambdas by η-expansion.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Normalize an application of a heterogenous binary operator like HAdd.hAdd, using:
- f = fun x => f xto increase the arity to 6
- (f + g) a = f a + g ato decrease the arity to 6
- (fun x => f x + g x) = f + gto get rid of any lambdas in front
Equations
- One or more equations did not get rendered due to their size.
Instances For
use that (fun x => f x + g x) = f + g
Equations
- One or more equations did not get rendered due to their size.
- Mathlib.Meta.FunProp.RefinedDiscrTree.MkDTExpr.reduceHBinOpAux.distributeLambdas [] type lhs rhs = pure (type, lhs, rhs, [])
Instances For
Normalize an application if the head is  +, *, - or /.
Optionally return the (type, lhs, rhs, lambdas).
Equations
- One or more equations did not get rendered due to their size.
Instances For
Normalize an application of a unary operator like Inv.inv, using:
- f⁻¹ a = (f a)⁻¹to decrease the arity to 3
- (fun x => (f a)⁻¹) = f⁻¹to get rid of any lambdas in front
Equations
- One or more equations did not get rendered due to their size.
Instances For
use that (fun x => (f x)⁻¹) = f⁻¹
Equations
- One or more equations did not get rendered due to their size.
- Mathlib.Meta.FunProp.RefinedDiscrTree.MkDTExpr.reduceUnOpAux.distributeLambdas [] type arg = pure (type, arg, [])
Instances For
Normalize an application if the head is ⁻¹ or -.
Optionally return the (type, arg, lambdas).
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.
Equations
- One or more equations did not get rendered due to their size.
Return all pairs of body, bound variables that could possibly appear due to η-reduction
Equations
- One or more equations did not get rendered due to their size.
- Mathlib.Meta.FunProp.RefinedDiscrTree.MkDTExpr.etaPossibilities e lambdas k = HOrElse.hOrElse (k e lambdas) fun (x : Unit) => failure
Instances For
run etaPossibilities, and cache the result if there are multiple possibilities.
Equations
- One or more equations did not get rendered due to their size.
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
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
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
Inserting intro a RefinedDiscrTree #
Insert the value v at index keys : Array Key in a RefinedDiscrTree.
Warning: to accound for η-reduction, an entry may need to be added at multiple indexes,
so it is recommended to use RefinedDiscrTree.insert for insertion.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Insert the value v at index e : DTExpr in a RefinedDiscrTree.
Warning: to accound for η-reduction, an entry may need to be added at multiple indexes,
so it is recommended to use RefinedDiscrTree.insert for insertion.
Equations
- d.insertDTExpr e v = d.insertInRefinedDiscrTree e.flatten v
Instances For
Insert the value v at index e : Expr in a RefinedDiscrTree.
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.
if onlySpecific := true, then we filter out the patterns * and Eq * * *.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Insert the value vLhs at index lhs, and if rhs is indexed differently, then also
insert the value vRhs at index rhs.
Equations
- One or more equations did not get rendered due to their size.
Instances For
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.
If e is not a metavariable, return the possible Trie α that exactly match with e.
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
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
go
Apply a monadic function to the array of values at each node in a RefinedDiscrTree.
Apply a monadic function to the array of values at each node in a RefinedDiscrTree.
Equations
- d.mapArraysM f = do let __do_lift ← d.root.mapM fun (x : Mathlib.Meta.FunProp.RefinedDiscrTree.Trie α) => x.mapArraysM f pure { root := __do_lift }
Instances For
Apply a function to the array of values at each node in a RefinedDiscrTree.
Equations
- d.mapArrays f = d.mapArraysM f