Make proofs from a congruence closure #
The monad for the cc
tactic stores the current state of the tactic.
Instances For
Run a computation in the CCM
monad.
Equations
- x.run c = StateRefT'.run x c
Instances For
Update the cache
field of the state.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Read the cache
field of the state.
Equations
- Mathlib.Tactic.CC.CCM.getCache = do let __do_lift ← get pure __do_lift.cache
Instances For
Look up an entry associated with the given expression.
Equations
- Mathlib.Tactic.CC.CCM.getEntry e = do let __do_lift ← get pure (Batteries.RBMap.find? __do_lift.entries e)
Instances For
Use the normalizer to normalize e
.
If no normalizer was configured, returns e
itself.
Equations
- Mathlib.Tactic.CC.CCM.normalize e = do let __do_lift ← get match __do_lift.normalizer with | some normalizer => liftM (normalizer.normalize e) | x => pure e
Instances For
Return the root expression of the expression's congruence class.
Equations
- Mathlib.Tactic.CC.CCM.getRoot e = do let __do_lift ← get pure (__do_lift.root e)
Instances For
Is e
the root of its congruence class?
Equations
- Mathlib.Tactic.CC.CCM.isCgRoot e = do let __do_lift ← get pure (__do_lift.isCgRoot e)
Instances For
Return true iff the given function application are congruent
e₁
should have the form f a
and e₂
the form g b
.
See paper: Congruence Closure for Intensional Type Theory.
Try to find a congruence theorem for an application of fn
with nargs
arguments, with support
for HEq
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Try to find a congruence theorem for the expression e
with support for HEq
.
Equations
- Mathlib.Tactic.CC.CCM.mkCCCongrTheorem e = Mathlib.Tactic.CC.CCM.mkCCHCongrTheorem e.getAppFn e.getAppNumArgs
Instances For
Treat the entry associated with e
as a first-order function.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Update the modification time of the congruence class of e
.
Does the congruence class with root root
have any HEq
proofs?
Equations
- One or more equations did not get rendered due to their size.
Instances For
Apply symmetry to H
, which is an Eq
or a HEq
.
- If
heqProofs
is true, ensure the result is aHEq
(otherwise it is assumed to beEq
). - If
flipped
is true, applysymm
, otherwise keep the same direction.
Equations
- One or more equations did not get rendered due to their size.
Instances For
In a delayed way, apply symmetry to H
, which is an Eq
or a HEq
.
- If
heqProofs
is true, ensure the result is aHEq
(otherwise it is assumed to beEq
). - If
flipped
is true, applysymm
, otherwise keep the same direction.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Apply symmetry to H
, which is an Eq
or a HEq
.
- If
heqProofs
is true, ensure the result is aHEq
(otherwise it is assumed to beEq
). - If
flipped
is true, applysymm
, otherwise keep the same direction.
Equations
- One or more equations did not get rendered due to their size.
- Mathlib.Tactic.CC.CCM.flipProof (↑H_2) flipped heqProofs = Mathlib.Tactic.CC.EntryExpr.ofExpr <$> Mathlib.Tactic.CC.CCM.flipProofCore H_2 flipped heqProofs
- Mathlib.Tactic.CC.CCM.flipProof H flipped heqProofs = pure H
Instances For
Are e₁
and e₂
known to be in the same equivalence class?
Equations
- One or more equations did not get rendered due to their size.
Instances For
Is e₁ ≠ e₂
known to be true?
Note that this is stronger than not (isEqv e₁ e₂)
:
only if we can prove they are distinct this returns true
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Is the proposition e
known to be true?
Equations
Instances For
Is the proposition e
known to be false?
Equations
- Mathlib.Tactic.CC.CCM.isEqFalse e = Mathlib.Tactic.CC.CCM.isEqv e (Lean.Expr.const `False [])
Instances For
Apply transitivity to H₁
and H₂
, which are both Eq
or HEq
depending on heqProofs
.
Equations
- Mathlib.Tactic.CC.CCM.mkTrans H₁ H₂ heqProofs = if heqProofs = true then Lean.Meta.mkHEqTrans H₁ H₂ else Lean.Meta.mkEqTrans H₁ H₂
Instances For
Apply transitivity to H₁?
and H₂
, which are both Eq
or HEq
depending on heqProofs
.
If H₁?
is none
, return H₂
instead.
Equations
- Mathlib.Tactic.CC.CCM.mkTransOpt (some H₁) H₂ heqProofs = Mathlib.Tactic.CC.CCM.mkTrans H₁ H₂ heqProofs
- Mathlib.Tactic.CC.CCM.mkTransOpt none H₂ heqProofs = pure H₂
Instances For
Use congruence on arguments to prove lhs = rhs
.
That is, tries to prove that lhsFn lhsArgs[0] ... lhsArgs[n-1] = lhsFn rhsArgs[0] ... rhsArgs[n-1]
by showing that lhsArgs[i] = rhsArgs[i]
for all i
.
Fails if the head function of lhs
is not that of rhs
.
If e₁ : R lhs₁ rhs₁
, e₂ : R lhs₂ rhs₂
and lhs₁ = rhs₂
, where R
is a symmetric relation,
prove R lhs₁ rhs₁
is equivalent to R lhs₂ rhs₂
.
- if
lhs₁
is known to equallhs₂
, returnnone
- if
lhs₁
is not known to equalrhs₂
, fail.
Use congruence on arguments to prove e₁ = e₂
.
Special case: if e₁
and e₂
have the form R lhs₁ rhs₁
and R lhs₂ rhs₂
such that
R
is symmetric and lhs₁ = rhs₂
, then use those facts instead.
Turn a delayed proof into an actual proof term.
Use the format of H
to try and construct a proof or lhs = rhs
:
If asHEq
is true
, then build a proof for HEq e₁ e₂
.
Otherwise, build a proof for e₁ = e₂
.
The result is none
if e₁
and e₂
are not in the same equivalence class.
Build a proof for e₁ = e₂
.
The result is none
if e₁
and e₂
are not in the same equivalence class.
Build a proof for HEq e₁ e₂
.
The result is none
if e₁
and e₂
are not in the same equivalence class.
Build a proof for e = True
. Fails if e
is not known to be true.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Build a proof for e = False
. Fails if e
is not known to be false.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Build a proof for a = b
. Fails if a
and b
are not known to be equal.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given a
, a₁
and a₁NeB : a₁ ≠ b
, return a proof of a ≠ b
if a
and a₁
are in the
same equivalence class.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given aNeB₁ : a ≠ b₁
, b₁
and b
, return a proof of a ≠ b
if b
and b₁
are in the
same equivalence class.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Return the proof of e₁ = e₂
using ac_rfl
tactic.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given tr := t*r
sr := s*r
tEqs : t = s
, return a proof for tr = sr
We use a*b
to denote an AC application. That is, (a*b)*(c*a)
is the term a*a*b*c
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given e := lhs * r
and H : lhs = rhs
, return rhs * r
and the proof of e = rhs * r
.
Equations
Instances For
The single step of simplifyAC
.
Simplifies an expression e
by either simplifying one argument to the AC operator, or the whole
expression.
Equations
- One or more equations did not get rendered due to their size.
Instances For
If e
can be simplified by the AC module, return the simplified term and the proof term of the
equality.
Equations
- One or more equations did not get rendered due to their size.