For variables of type α
and formulas of type β
, Entails.eval a f
is meant to determine whether
a formula f
is true under assignment a
.
Instances
a ⊨ f
reads formula f
is true under assignment a
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
a ⊭ f
reads formula f
is false under assignment a
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
def
Std.Tactic.BVDecide.LRAT.Internal.Unsatisfiable
(α : Type u)
{σ : Type v}
[Entails α σ]
(f : σ)
:
f
is not true under any assignment.
Equations
Instances For
def
Std.Tactic.BVDecide.LRAT.Internal.Liff
(α : Type u)
{σ1 : Type v}
{σ2 : Type w}
[Entails α σ1]
[Entails α σ2]
(f1 : σ1)
(f2 : σ2)
:
f1
and f2
are logically equivalent
Equations
- Std.Tactic.BVDecide.LRAT.Internal.Liff α f1 f2 = ∀ (a : α → Bool), Std.Tactic.BVDecide.LRAT.Internal.Entails.eval a f1 ↔ Std.Tactic.BVDecide.LRAT.Internal.Entails.eval a f2
Instances For
def
Std.Tactic.BVDecide.LRAT.Internal.Limplies
(α : Type u)
{σ1 : Type v}
{σ2 : Type w}
[Entails α σ1]
[Entails α σ2]
(f1 : σ1)
(f2 : σ2)
:
f1
logically implies f2
Equations
- Std.Tactic.BVDecide.LRAT.Internal.Limplies α f1 f2 = ∀ (a : α → Bool), Std.Tactic.BVDecide.LRAT.Internal.Entails.eval a f1 → Std.Tactic.BVDecide.LRAT.Internal.Entails.eval a f2
Instances For
theorem
Std.Tactic.BVDecide.LRAT.Internal.liff_unsat
{α : Type u}
{σ1 : Type v}
{σ2 : Type w}
[Entails α σ1]
[Entails α σ2]
(f1 : σ1)
(f2 : σ2)
(h : Liff α f1 f2)
:
Unsatisfiable α f1 ↔ Unsatisfiable α f2
theorem
Std.Tactic.BVDecide.LRAT.Internal.limplies_unsat
{α : Type u}
{σ1 : Type v}
{σ2 : Type w}
[Entails α σ1]
[Entails α σ2]
(f1 : σ1)
(f2 : σ2)
(h : Limplies α f2 f1)
:
Unsatisfiable α f1 → Unsatisfiable α f2
theorem
Std.Tactic.BVDecide.LRAT.Internal.incompatible_of_unsat
(α : Type u)
{σ1 : Type v}
{σ2 : Type w}
[Entails α σ1]
[Entails α σ2]
(f1 : σ1)
(f2 : σ2)
:
Unsatisfiable α f1 → Incompatible α f1 f2
theorem
Std.Tactic.BVDecide.LRAT.Internal.unsat_of_limplies_and_incompatible
(α : Type u)
{σ1 : Type v}
{σ2 : Type w}
[Entails α σ1]
[Entails α σ2]
(f1 : σ1)
(f2 : σ2)
:
Limplies α f1 f2 → Incompatible α f1 f2 → Unsatisfiable α f1
theorem
Std.Tactic.BVDecide.LRAT.Internal.Incompatible.symm
{α : Type u}
{σ1 : Type v}
{σ2 : Type w}
[Entails α σ1]
[Entails α σ2]
(f1 : σ1)
(f2 : σ2)
:
Incompatible α f1 f2 ↔ Incompatible α f2 f1