The category of Heyting algebras.
- carrier : Type u_1
The underlying Heyting algebra.
- str : HeytingAlgebra ↑self
Instances For
Equations
- HeytAlg.instCoeSortType = { coe := HeytAlg.carrier }
@[reducible, inline]
Construct a bundled HeytAlg
from the underlying type and typeclass.
Equations
- HeytAlg.of X = HeytAlg.mk X
Instances For
instance
HeytAlg.instConcreteCategoryHeytingHomCarrier :
CategoryTheory.ConcreteCategory HeytAlg fun (x1 x2 : HeytAlg) => HeytingHom ↑x1 ↑x2
Equations
- One or more equations did not get rendered due to their size.
@[reducible, inline]
Typecheck a HeytingHom
as a morphism in HeytAlg
.
Equations
Instances For
Use the ConcreteCategory.hom
projection for @[simps]
lemmas.
Equations
- HeytAlg.Hom.Simps.hom X Y f = f.hom
Instances For
The results below duplicate the ConcreteCategory
simp lemmas, but we can keep them for dsimp
.
@[simp]
@[simp]
@[simp]
theorem
HeytAlg.ext
{X Y : HeytAlg}
{f g : X ⟶ Y}
(w : ∀ (x : ↑X), (CategoryTheory.ConcreteCategory.hom f) x = (CategoryTheory.ConcreteCategory.hom g) x)
:
@[simp]
@[simp]
theorem
HeytAlg.hom_ofHom
{X Y : Type u}
[HeytingAlgebra X]
[HeytingAlgebra Y]
(f : HeytingHom X Y)
:
@[simp]
@[simp]
theorem
HeytAlg.ofHom_comp
{X Y Z : Type u}
[HeytingAlgebra X]
[HeytingAlgebra Y]
[HeytingAlgebra Z]
(f : HeytingHom X Y)
(g : HeytingHom Y Z)
:
theorem
HeytAlg.ofHom_apply
{X Y : Type u}
[HeytingAlgebra X]
[HeytingAlgebra Y]
(f : HeytingHom X Y)
(x : X)
:
Equations
- HeytAlg.instInhabited = { default := HeytAlg.of PUnit.{?u.3 + 1} }
Equations
- One or more equations did not get rendered due to their size.
@[simp]
@[simp]
theorem
HeytAlg.hasForgetToLat_forget₂_map
{X✝ Y✝ : HeytAlg}
(f : X✝ ⟶ Y✝)
:
CategoryTheory.HasForget₂.forget₂.map f = BddDistLat.ofHom
(let __src := { toFun := ⇑(Hom.hom f), map_sup' := ⋯, map_inf' := ⋯ };
{ toFun := ⇑(Hom.hom f), map_sup' := ⋯, map_inf' := ⋯, map_top' := ⋯, map_bot' := ⋯ })
Constructs an isomorphism of Heyting algebras from an order isomorphism between them.
Equations
- One or more equations did not get rendered due to their size.