Datatypes for cc #
Some of the data structures here are used in multiple parts of the tactic. We split them into their own file.
TODO #
This file is ported from C++ code, so many declarations lack documents.
Return true if e represents a constant value (numeral, character, or string).
Equations
- Mathlib.Tactic.CC.isValue e = (e.int?.isSome || e.isCharLit || e.isStringLit)
Instances For
Return true if e represents a value (nat/int numeral, character, or string).
In addition to the conditions in Mathlib.Tactic.CC.isValue, this also checks that
kernel computation can compare the values for equality.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given a reflexive relation R, and a proof H : a = b, build a proof for R a b
Equations
- One or more equations did not get rendered due to their size.
Instances For
Ordering on Expr.
Equations
- Mathlib.Tactic.CC.instOrdExpr_mathlib = { compare := fun (a b : Lean.Expr) => bif a.lt b then Ordering.lt else bif b.eqv a then Ordering.eq else Ordering.gt }
Instances For
Red-black maps whose keys are Exprs.
TODO: the choice between RBMap and HashMap is not obvious:
the current version follows the Lean 3 C++ implementation.
Once the cc tactic is used a lot in Mathlib, we should profile and see
if HashMap could be more optimal.
Equations
- Mathlib.Tactic.CC.RBExprMap α = Batteries.RBMap Lean.Expr α compare
Instances For
Red-black sets of Exprs.
TODO: the choice between RBSet and HashSet is not obvious:
the current version follows the Lean 3 C++ implementation.
Once the cc tactic is used a lot in Mathlib, we should profile and see
if HashSet could be more optimal.
Equations
Instances For
CongrTheorems equiped with additional infos used by congruence closure modules.
- type : Lean.Expr
- proof : Lean.Expr
- argKinds : Array Lean.Meta.CongrArgKind
- heqResult : BoolIf heqResultis true, then lemma is based on heterogeneous equality and the conclusion is a heterogeneous equality.
- hcongrTheorem : BoolIf hcongrTheoremis true, then lemma was created usingmkHCongrWithArity.
Instances For
Automatically generated congruence lemma based on heterogeneous equality.
This returns an annotated version of the result from Lean.Meta.mkHCongrWithArity.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Keys used to find corresponding CCCongrTheorems.
- fn : Lean.ExprThe function of the given CCCongrTheorem.
- nargs : NatThe number of arguments of fn.
Instances For
Configs used in congruence closure modules.
- ignoreInstances : BoolIf true, congruence closure will treat implicit instance arguments as constants.This means that setting ignoreInstances := falsewill fail to unify two definitionally equal instances of the same class.
- ac : BoolIf true, congruence closure modulo Associativity and Commutativity.
- If - hoFnsis- some fns, then full (and more expensive) support for higher-order functions is only considered for the functions in fns and local functions. The performance overhead is described in the paper "Congruence Closure in Intensional Type Theory". If- hoFnsis- none, then full support is provided for all constants.
- em : BoolIf true, then use excluded middle
- values : BoolIf true, we treat values as atomic symbols
Instances For
Equations
- Mathlib.Tactic.CC.instInhabitedCCConfig = { default := { ignoreInstances := default, ac := default, hoFns := default, em := default, values := default } }
An ACApps represents either just an Expr or applications of an associative and commutative
binary operator.
- ofExpr: Lean.Expr → Mathlib.Tactic.CC.ACAppsAn ACAppsof just anExpr.
- apps: Lean.Expr → Array Lean.Expr → Mathlib.Tactic.CC.ACAppsAn ACAppsof applications of a binary operator.argsare assumed to be sorted.See also ACApps.mkAppsifargsare not yet sorted.
Instances For
Equations
- Mathlib.Tactic.CC.instInhabitedACApps = { default := ↑default }
Equations
Return true iff e₁ is a "subset" of e₂.
Example: The result is true for e₁ := a*a*a*b*d and e₂ := a*a*a*a*b*b*c*d*d.
The result is also true for e₁ := a and e₂ := a*a*a*b*c.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Appends elements of the set difference e₁ \ e₂ to r.
Example: given e₁ := a*a*a*a*b*b*c*d*d*d and e₂ := a*a*a*b*b*d,
the result is #[a, c, d, d]
Precondition: e₂.isSubset e₁
Equations
- One or more equations did not get rendered due to their size.
Instances For
Appends arguments of e to r.
Equations
- Mathlib.Tactic.CC.ACApps.append op e r = match e with | Mathlib.Tactic.CC.ACApps.apps op' args => if (op' == op) = true then r ++ args else r | ↑e => Array.push r e
Instances For
Appends elements in the intersection of e₁ and e₂ to r.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Sorts args and applies them to ACApps.apps.
Equations
- Mathlib.Tactic.CC.ACApps.mkApps op args = Mathlib.Tactic.CC.ACApps.apps op (args.qsort Lean.Expr.lt 0)
Instances For
Flattens given two ACApps.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Converts an ACApps to an Expr. This returns none when the empty applications are given.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Red-black maps whose keys are ACAppses.
TODO: the choice between RBMap and HashMap is not obvious:
the current version follows the Lean 3 C++ implementation.
Once the cc tactic is used a lot in Mathlib, we should profile and see
if HashMap could be more optimal.
Equations
Instances For
Red-black sets of ACAppses.
TODO: the choice between RBSet and HashSet is not obvious:
the current version follows the Lean 3 C++ implementation.
Once the cc tactic is used a lot in Mathlib, we should profile and see
if HashSet could be more optimal.
Equations
Instances For
For proof terms generated by AC congruence closure modules, we want a placeholder as an equality
proof between given two terms which will be generated by non-AC congruence closure modules later.
DelayedExpr represents it using eqProof.
- ofExpr: Lean.Expr → Mathlib.Tactic.CC.DelayedExprA DelayedExprof just anExpr.
- eqProof: Lean.Expr → Lean.Expr → Mathlib.Tactic.CC.DelayedExprA placeholder as an equality proof between given two terms which will be generated by non-AC congruence closure modules later. 
- congrArg: Lean.Expr → Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.DelayedExprWill be applied to congr_arg.
- congrFun: Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.ACApps → Mathlib.Tactic.CC.DelayedExprWill be applied to congr_fun.
- eqSymm: Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.DelayedExprWill be applied to Eq.symm.
- eqSymmOpt: Mathlib.Tactic.CC.ACApps → Mathlib.Tactic.CC.ACApps → Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.DelayedExprWill be applied to Eq.symm.
- eqTrans: Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.DelayedExprWill be applied to Eq.trans.
- eqTransOpt: Mathlib.Tactic.CC.ACApps →
  Mathlib.Tactic.CC.ACApps →
    Mathlib.Tactic.CC.ACApps →
      Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.DelayedExprWill be applied to Eq.trans.
- heqOfEq: Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.DelayedExprWill be applied to heq_of_eq.
- heqSymm: Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.DelayedExprWill be applied to HEq.symm.
Instances For
Equations
- Mathlib.Tactic.CC.instInhabitedDelayedExpr = { default := ↑default }
This is used as a proof term in Entrys instead of Expr.
- ofExpr: Lean.Expr → Mathlib.Tactic.CC.EntryExprAn EntryExprof just anExpr.
- congr: Mathlib.Tactic.CC.EntryExprdummy congruence proof, it is just a placeholder. 
- eqTrue: Mathlib.Tactic.CC.EntryExprdummy eq_true proof, it is just a placeholder 
- refl: Mathlib.Tactic.CC.EntryExprdummy refl proof, it is just a placeholder. 
- ofDExpr: Mathlib.Tactic.CC.DelayedExpr → Mathlib.Tactic.CC.EntryExprAn EntryExprof aDelayedExpr.
Instances For
Equations
- Mathlib.Tactic.CC.instInhabitedEntryExpr = { default := ↑default }
Equations
- One or more equations did not get rendered due to their size.
Equivalence class data associated with an expression e.
- next : Lean.Exprnext element in the equivalence class. 
- root : Lean.Exprroot (aka canonical) representative of the equivalence class. 
- cgRoot : Lean.Exprroot of the congruence class, it is meaningless if eis not an application.
- proof : Option Mathlib.Tactic.CC.EntryExpr
- Variable in the AC theory. 
- flipped : Boolproof has been flipped 
- interpreted : Booltrueif the node should be viewed as an abstract value
- constructor : Booltrueif head symbol is a constructor
- hasLambdas : Booltrueif equivalence class contains lambda expressions
- heqProofs : BoolheqProofs == trueiff some proofs in the equivalence class are based on heterogeneous equality. We represent equality and heterogeneous equality in a single equivalence class.
- fo : BoolIf fo == true, then the expression associated with this entry is an application, and we are using first-order approximation to encode it. That is, we ignore its partial applications.
- size : Natnumber of elements in the equivalence class, it is meaningless if e != root
- mt : NatThe field mtis used to implement the mod-time optimization introduce by the Simplify theorem prover. The basic idea is to introduce a counter gmt that records the number of heuristic instantiation that have occurred in the current branch. It is incremented after each round of heuristic instantiation. The fieldmtrecords the last time any proper descendant of of thie entry was involved in a merge.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Stores equivalence class data associated with an expression e.
Instances For
Equivalence class data associated with an expression e used by AC congruence closure
modules.
- idx : NatNatural number associated to an expression. 
- RLHSOccs : Mathlib.Tactic.CC.RBACAppsSetAC variables that occur on the left hand side of an equality which eoccurs as the left hand side of inCCState.acR.
- RRHSOccs : Mathlib.Tactic.CC.RBACAppsSetAC variables that occur on the left hand side of an equality which eoccurs as the right hand side of inCCState.acR. Don't confuse.
Instances For
Equations
- Mathlib.Tactic.CC.instInhabitedACEntry = { default := { idx := default, RLHSOccs := default, RRHSOccs := default } }
Returns the occurrences of this entry in either the LHS or RHS.
Instances For
Used to record when an expression processed by cc occurs in another expression.
- expr : Lean.Expr
- symmTable : BoolIf symmTableis true, then we should use thesymmCongruences, otherwisecongruences. Remark: this information is redundant, it can be inferred fromexpr. We use store it for performance reasons.
Instances For
Used to map an expression e to another expression that contains e.
When e is normalized, its parents should also change.
Instances For
- fo: Lean.Expr → Array Lean.Expr → Mathlib.Tactic.CC.CongruencesKeyfnis First-Order: we do not consider all partial applications.
- ho: Lean.Expr → Lean.Expr → Mathlib.Tactic.CC.CongruencesKeyfnis Higher-Order.
Instances For
Maps each expression (via mkCongruenceKey) to expressions it might be congruent to.
Equations
Instances For
The symmetric variant of Congruences.
The Name identifies which relation the congruence is considered for.
Note that this only works for two-argument relations: ModEq n and ModEq m are considered the
same.
Equations
Instances For
Stores the root representatives of subsingletons.
Instances For
Stores the root representatives of .instImplicit arguments.
Instances For
Instances For
Equations
Instances For
Congruence closure state.
This may be considered to be a set of expressions and an equivalence class over this set.
The equivalence class is generated by the equational rules that are added to the CCState and
congruence, that is, if a = b then f(a) = f(b) and so on.
- ignoreInstances : Bool
- ac : Bool
- em : Bool
- values : Bool
- entries : Mathlib.Tactic.CC.EntriesMaps known expressions to their equivalence class data. 
- parents : Mathlib.Tactic.CC.ParentsMaps an expression eto the expressionseoccurs in.
- congruences : Mathlib.Tactic.CC.CongruencesMaps each expression to a set of expressions it might be congruent to. 
- symmCongruences : Mathlib.Tactic.CC.SymmCongruencesMaps each expression to a set of expressions it might be congruent to, via the symmetrical relation. 
- subsingletonReprs : Mathlib.Tactic.CC.SubsingletonReprsStores the root representatives of subsingletons. 
- instImplicitReprs : Mathlib.Tactic.CC.InstImplicitReprsRecords which instances of the same class are defeq. 
- frozePartitions : BoolThe congruence closure module has a mode where the root of each equivalence class is marked as an interpreted/abstract value. Moreover, in this mode proof production is disabled. This capability is useful for heuristic instantiation. 
- canOps : Mathlib.Tactic.CC.RBExprMap Lean.ExprMapping from operators occurring in terms and their canonical representation in this module 
- opInfo : Mathlib.Tactic.CC.RBExprMap BoolWhether the canonical operator is suppoted by AC. 
- acEntries : Mathlib.Tactic.CC.RBExprMap Mathlib.Tactic.CC.ACEntryExtra Entryinformation used by the AC part of the tactic.
- Records equality between - ACApps.
- inconsistent : BoolReturns true if the CCStateis inconsistent. For example if it had botha = banda ≠ bin it.
- gmt : Nat"Global Modification Time". gmt is a number stored on the CCState, it is compared with the modification time of a cc_entry in e-matching. SeeCCState.mt.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Update the CCState by constructing and inserting a new Entry.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Get the root representative of the given expression.
Equations
- ccs.root e = match Batteries.RBMap.find? ccs.entries e with | some n => n.root | none => e
Instances For
Get the next element in the equivalence class.
Note that if the given Expr e is not in the graph then it will just return e.
Equations
- ccs.next e = match Batteries.RBMap.find? ccs.entries e with | some n => n.next | none => e
Instances For
Check if e is the root of the congruence class.
Equations
- ccs.isCgRoot e = match Batteries.RBMap.find? ccs.entries e with | some n => e == n.cgRoot | none => true
Instances For
"Modification Time". The field mt is used to implement the mod-time optimization introduced by the
Simplify theorem prover. The basic idea is to introduce a counter gmt that records the number of
heuristic instantiation that have occurred in the current branch. It is incremented after each round
of heuristic instantiation. The field mt records the last time any proper descendant of of thie
entry was involved in a merge.
Equations
- ccs.mt e = match Batteries.RBMap.find? ccs.entries e with | some n => n.mt | none => ccs.gmt
Instances For
Is the expression in an equivalence class with only one element (namely, itself)?
Equations
- ccs.inSingletonEqc e = match Batteries.RBMap.find? ccs.entries e with | some it => it.next == e | none => true
Instances For
Append to roots all the roots of equivalence classes in ccs.
If nonsingletonOnly is true, we skip all the singleton equivalence classes.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Check for integrity of the CCState.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Check for integrity of the CCState.
Equations
- ccs.checkInvariant = Batteries.RBMap.all ccs.entries fun (k : Lean.Expr) (n : Mathlib.Tactic.CC.Entry) => k != n.root || ccs.checkEqc k
Instances For
Equations
- ccs.getNumROccs e inLHS = match Batteries.RBMap.find? ccs.acEntries e with | some ent => Batteries.RBSet.size (ent.ROccs inLHS) | none => 0
Instances For
Search for the AC-variable (Entry.acVar) with the least occurrences in the state.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Instances For
Instances For
Pretty print the entry associated with the given expression.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Pretty print the entire cc graph.
If the nonSingleton argument is set to true then singleton equivalence classes will be
omitted.
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
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
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
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
Equations
- ccs.ppAC = (ccs.ppACDecls ++ Lean.MessageData.ofFormat (Std.Format.text "," ++ Lean.Format.line) ++ ccs.ppACR).sbracket
Instances For
The congruence closure module (optionally) uses a normalizer. The idea is to use it (if available) to normalize auxiliary expressions produced by internal propagation rules (e.g., subsingleton propagator).
- normalize : Lean.Expr → Lean.MetaM Lean.ExprThe congruence closure module (optionally) uses a normalizer. The idea is to use it (if available) to normalize auxiliary expressions produced by internal propagation rules (e.g., subsingleton propagator). 
Instances For
- propagated : Array Lean.Expr → Lean.MetaM Unit
- newAuxCCTerm : Lean.Expr → Lean.MetaM UnitCongruence closure module invokes the following method when a new auxiliary term is created during propagation. 
Instances For
CCStructure extends CCState (which records a set of facts derived by congruence closure)
by recording which steps still need to be taken to solve the goal.
- ignoreInstances : Bool
- ac : Bool
- em : Bool
- values : Bool
- entries : Mathlib.Tactic.CC.Entries
- parents : Mathlib.Tactic.CC.Parents
- congruences : Mathlib.Tactic.CC.Congruences
- symmCongruences : Mathlib.Tactic.CC.SymmCongruences
- subsingletonReprs : Mathlib.Tactic.CC.SubsingletonReprs
- instImplicitReprs : Mathlib.Tactic.CC.InstImplicitReprs
- frozePartitions : Bool
- canOps : Mathlib.Tactic.CC.RBExprMap Lean.Expr
- opInfo : Mathlib.Tactic.CC.RBExprMap Bool
- acEntries : Mathlib.Tactic.CC.RBExprMap Mathlib.Tactic.CC.ACEntry
- inconsistent : Bool
- gmt : Nat
- todo : Array Mathlib.Tactic.CC.TodoEntryEqualities that have been discovered but not processed. 
- acTodo : Array Mathlib.Tactic.CC.ACTodoEntryAC-equalities that have been discovered but not processed. 
- normalizer : Option Mathlib.Tactic.CC.CCNormalizer
- phandler : Option Mathlib.Tactic.CC.CCPropagationHandler
Instances For
Equations
- One or more equations did not get rendered due to their size.