The category of boolean algebras.
- carrier : Type u_1
The underlying boolean algebra.
- str : BooleanAlgebra ↑self
Instances For
Equations
- BoolAlg.instCoeSortType = { coe := BoolAlg.carrier }
@[reducible, inline]
Construct a bundled BoolAlg
from the underlying type and typeclass.
Equations
- BoolAlg.of X = BoolAlg.mk X
Instances For
The type of morphisms in BoolAlg R
.
- hom' : BoundedLatticeHom ↑X ↑Y
The underlying
BoundedLatticeHom
.
Instances For
instance
BoolAlg.instConcreteCategoryBoundedLatticeHomCarrier :
CategoryTheory.ConcreteCategory BoolAlg fun (x1 x2 : BoolAlg) => BoundedLatticeHom ↑x1 ↑x2
Equations
- One or more equations did not get rendered due to their size.
@[reducible, inline]
Turn a morphism in BoolAlg
back into a BoundedLatticeHom
.
Equations
Instances For
@[reducible, inline]
abbrev
BoolAlg.ofHom
{X Y : Type u}
[BooleanAlgebra X]
[BooleanAlgebra Y]
(f : BoundedLatticeHom X Y)
:
Typecheck a BoundedLatticeHom
as a morphism in BoolAlg
.
Equations
Instances For
Use the ConcreteCategory.hom
projection for @[simps]
lemmas.
Equations
- BoolAlg.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
BoolAlg.ext
{X Y : BoolAlg}
{f g : X ⟶ Y}
(w : ∀ (x : ↑X), (CategoryTheory.ConcreteCategory.hom f) x = (CategoryTheory.ConcreteCategory.hom g) x)
:
@[simp]
@[simp]
theorem
BoolAlg.hom_ofHom
{X Y : Type u}
[BooleanAlgebra X]
[BooleanAlgebra Y]
(f : BoundedLatticeHom X Y)
:
@[simp]
@[simp]
theorem
BoolAlg.ofHom_comp
{X Y Z : Type u}
[BooleanAlgebra X]
[BooleanAlgebra Y]
[BooleanAlgebra Z]
(f : BoundedLatticeHom X Y)
(g : BoundedLatticeHom Y Z)
:
theorem
BoolAlg.ofHom_apply
{X Y : Type u}
[BooleanAlgebra X]
[BooleanAlgebra Y]
(f : BoundedLatticeHom X Y)
(x : X)
:
Equations
- BoolAlg.instInhabited = { default := BoolAlg.of PUnit.{?u.3 + 1} }
Turn a BoolAlg
into a BddDistLat
by forgetting its complement operation.
Equations
- X.toBddDistLat = BddDistLat.of ↑X
Instances For
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.
@[simp]
@[simp]
theorem
BoolAlg.hasForgetToHeytAlg_forget₂_map
{X Y : BoolAlg}
(f : X ⟶ Y)
:
CategoryTheory.HasForget₂.forget₂.map f = HeytAlg.ofHom { toFun := ⇑(Hom.hom f), map_sup' := ⋯, map_inf' := ⋯, map_bot' := ⋯, map_himp' := ⋯ }
Constructs an equivalence between Boolean algebras from an order isomorphism between them.
Equations
- One or more equations did not get rendered due to their size.
Instances For
OrderDual
as a functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
The powerset functor. Set
as a contravariant functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
theorem
typeToBoolAlgOp_map
{X Y : Type u}
(f : X ⟶ Y)
:
typeToBoolAlgOp.map f = (BoolAlg.ofHom
(let __src := { toFun := ⇑(CompleteLatticeHom.setPreimage f), map_sup' := ⋯, map_inf' := ⋯ };
{ toFun := ⇑(CompleteLatticeHom.setPreimage f), map_sup' := ⋯, map_inf' := ⋯, map_top' := ⋯, map_bot' := ⋯ })).op
@[simp]