Heyting algebras #
This file defines Heyting, co-Heyting and bi-Heyting algebras.
A Heyting algebra is a bounded distributive lattice with an implication operation ⇨
such that
a ≤ b ⇨ c ↔ a ⊓ b ≤ c
. It also comes with a pseudo-complement ᶜ
, such that aᶜ = a ⇨ ⊥
.
Co-Heyting algebras are dual to Heyting algebras. They have a difference \
and a negation ¬
such that a \ b ≤ c ↔ a ≤ b ⊔ c
and ¬a = ⊤ \ a
.
Bi-Heyting algebras are Heyting algebras that are also co-Heyting algebras.
From a logic standpoint, Heyting algebras precisely model intuitionistic logic, whereas boolean algebras model classical logic.
Heyting algebras are the order theoretic equivalent of cartesian-closed categories.
Main declarations #
GeneralizedHeytingAlgebra
: Heyting algebra without a top element (nor negation).GeneralizedCoheytingAlgebra
: Co-Heyting algebra without a bottom element (nor complement).HeytingAlgebra
: Heyting algebra.CoheytingAlgebra
: Co-Heyting algebra.BiheytingAlgebra
: bi-Heyting algebra.
References #
Tags #
Heyting, Brouwer, algebra, implication, negation, intuitionistic
Notation #
Equations
- Pi.instHImpForall = { himp := fun (a b : (i : ι) → π i) (i : ι) => a i ⇨ b i }
Equations
- Pi.instHNotForall = { hnot := fun (a : (i : ι) → π i) (i : ι) => ¬a i }
A generalized Heyting algebra is a lattice with an additional binary operation ⇨
called
Heyting implication such that (a ⇨ ·)
is right adjoint to (a ⊓ ·)
.
This generalizes HeytingAlgebra
by not requiring a bottom element.
Instances
A generalized co-Heyting algebra is a lattice with an additional binary
difference operation \
such that (· \ a)
is left adjoint to (· ⊔ a)
.
This generalizes CoheytingAlgebra
by not requiring a top element.
Instances
A Heyting algebra is a bounded lattice with an additional binary operation ⇨
called Heyting
implication such that (a ⇨ ·)
is right adjoint to (a ⊓ ·)
.
Instances
A co-Heyting algebra is a bounded lattice with an additional binary difference operation \
such that (· \ a)
is left adjoint to (· ⊔ a)
.
Instances
Equations
- HeytingAlgebra.toBoundedOrder = { toOrderTop := GeneralizedHeytingAlgebra.toOrderTop, toBot := OrderBot.toBot, bot_le := ⋯ }
Equations
- CoheytingAlgebra.toBoundedOrder = { toTop := OrderTop.toTop, le_top := ⋯, toBot := OrderBot.toBot, bot_le := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Construct a Heyting algebra from the lattice structure and Heyting implication alone.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Construct a Heyting algebra from the lattice structure and complement operator alone.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Construct a co-Heyting algebra from the lattice structure and the difference alone.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Construct a co-Heyting algebra from the difference and Heyting negation alone.
Equations
- One or more equations did not get rendered due to their size.
Instances For
In this section, we'll give interpretations of these results in the Heyting algebra model of
intuitionistic logic,- where ≤
can be interpreted as "validates", ⇨
as "implies", ⊓
as "and",
⊔
as "or", ⊥
as "false" and ⊤
as "true". Note that we confuse →
and ⊢
because those are
the same in this logic.
See also Prop.heytingAlgebra
.
p → q → r ↔ p ∧ q → r
p → q → r ↔ q ∧ p → r
p → q → r ↔ q → p → r
p → p → q ↔ p → q
(p → q) ∧ p → q
p ∧ (p → q) → q
(p → q) ∧ p ↔ q ∧ p
The deduction theorem in the Heyting algebra model of intuitionistic logic: an implication holds iff the conclusion follows from the hypothesis.
(q → r) → (p → q) → q → r
p → q → r ↔ q → p → r
See himp_le
for a stronger version in Boolean algebras.
Equations
- One or more equations did not get rendered due to their size.
Equations
- Prod.instGeneralizedHeytingAlgebra = { toLattice := Prod.instLattice α β, toOrderTop := Prod.instOrderTop α β, toHImp := Prod.instHImp α β, le_himp_iff := ⋯ }
Equations
- Pi.instGeneralizedHeytingAlgebra = { toLattice := Pi.instLattice, toOrderTop := Pi.instOrderTop, toHImp := Pi.instHImpForall, le_himp_iff := ⋯ }
Alias of sdiff_sup_self
.
Alias of sup_sdiff_self
.
See le_sdiff
for a stronger version in generalised Boolean algebras.
a version of sdiff_sup_sdiff_cancel
with more general hypotheses.
Equations
- GeneralizedCoheytingAlgebra.toDistribLattice = { toLattice := GeneralizedCoheytingAlgebra.toLattice, le_sup_inf := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- Prod.instGeneralizedCoheytingAlgebra = { toLattice := Prod.instLattice α β, toOrderBot := Prod.instOrderBot α β, toSDiff := Prod.instSDiff α β, sdiff_le_iff := ⋯ }
Equations
- Pi.instGeneralizedCoheytingAlgebra = { toLattice := Pi.instLattice, toOrderBot := Pi.instOrderBot, toSDiff := Pi.sdiff, sdiff_le_iff := ⋯ }
Alias of the reverse direction of le_compl_iff_disjoint_right
.
Alias of the reverse direction of le_compl_iff_disjoint_left
.
Alias of le_compl_comm
.
Alias of the forward direction of le_compl_comm
.
Equations
- One or more equations did not get rendered due to their size.
Equations
- Prod.instHeytingAlgebra = { toGeneralizedHeytingAlgebra := Prod.instGeneralizedHeytingAlgebra, toOrderBot := Prod.instOrderBot α β, toHasCompl := Prod.instHasCompl α β, himp_bot := ⋯ }
Equations
- Pi.instHeytingAlgebra = { toGeneralizedHeytingAlgebra := Pi.instGeneralizedHeytingAlgebra, toOrderBot := Pi.instOrderBot, toHasCompl := Pi.hasCompl, himp_bot := ⋯ }
Equations
- CoheytingAlgebra.toDistribLattice = { toLattice := GeneralizedCoheytingAlgebra.toLattice, le_sup_inf := ⋯ }
Alias of the reverse direction of hnot_le_iff_codisjoint_right
.
Alias of the reverse direction of hnot_le_iff_codisjoint_left
.
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- Pi.instCoheytingAlgebra = { toGeneralizedCoheytingAlgebra := Pi.instGeneralizedCoheytingAlgebra, toOrderTop := Pi.instOrderTop, toHNot := Pi.instHNotForall, top_sdiff := ⋯ }
Propositions form a Heyting algebra with implication as Heyting implication and negation as complement.
Equations
- One or more equations did not get rendered due to their size.
A bounded linear order is a bi-Heyting algebra by setting
a ⇨ b = ⊤
ifa ≤ b
anda ⇨ b = b
otherwise.a \ b = ⊥
ifa ≤ b
anda \ b = a
otherwise.
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.
Equations
- Prod.instBiheytingAlgebra = { toHeytingAlgebra := Prod.instHeytingAlgebra, toSDiff := GeneralizedCoheytingAlgebra.toSDiff, toHNot := CoheytingAlgebra.toHNot, sdiff_le_iff := ⋯, top_sdiff := ⋯ }
Equations
- Pi.instBiheytingAlgebra = { toHeytingAlgebra := Pi.instHeytingAlgebra, toSDiff := GeneralizedCoheytingAlgebra.toSDiff, toHNot := CoheytingAlgebra.toHNot, sdiff_le_iff := ⋯, top_sdiff := ⋯ }
Pullback a GeneralizedHeytingAlgebra
along an injection.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Pullback a GeneralizedCoheytingAlgebra
along an injection.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Pullback a HeytingAlgebra
along an injection.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Pullback a CoheytingAlgebra
along an injection.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Pullback a BiheytingAlgebra
along an injection.
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.