The category of bounded distributive lattices #
This defines BddDistLat
, the category of bounded distributive lattices.
Note that this category is sometimes called DistLat
when
being a lattice is understood to entail having a bottom and a top element.
The category of bounded distributive lattices with bounded lattice morphisms.
- str : DistribLattice ↑self.toDistLat
- isBoundedOrder : BoundedOrder ↑self.toDistLat
Instances For
Equations
- BddDistLat.instCoeSortType = { coe := fun (X : BddDistLat) => ↑X.toDistLat }
Equations
Construct a bundled BddDistLat
from a BoundedOrder
DistribLattice
.
Equations
Instances For
The type of morphisms in BddDistLat R
.
- hom' : BoundedLatticeHom ↑X.toDistLat ↑Y.toDistLat
The underlying
BoundedLatticeHom
.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Turn a morphism in BddDistLat
back into a BoundedLatticeHom
.
Equations
Instances For
Typecheck a BoundedLatticeHom
as a morphism in BddDistLat
.
Equations
Instances For
Use the ConcreteCategory.hom
projection for @[simps]
lemmas.
Equations
- BddDistLat.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
.
Equations
- BddDistLat.instInhabited = { default := BddDistLat.of PUnit.{?u.3 + 1} }
Turn a BddDistLat
into a BddLat
by forgetting it is distributive.
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.
Constructs an equivalence between bounded distributive lattices 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
The equivalence between BddDistLat
and itself induced by OrderDual
both ways.
Equations
- One or more equations did not get rendered due to their size.