The category of finite bounded distributive lattices #
This file defines FinBddDistLat
, the category of finite distributive lattices with
bounded lattice homomorphisms.
The category of finite distributive lattices with bounded lattice morphisms.
- str : DistribLattice ↑self.toDistLat
- isBoundedOrder : BoundedOrder ↑self.toDistLat
Instances For
Equations
- FinBddDistLat.instCoeSortType = { coe := fun (X : FinBddDistLat) => ↑X.toDistLat }
Equations
Equations
Construct a bundled FinBddDistLat
from a Fintype
BoundedOrder
DistribLattice
.
Equations
Instances For
Construct a bundled FinBddDistLat
from a Nonempty
Fintype
DistribLattice
.
Equations
Instances For
The type of morphisms in FinBddDistLat 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 FinBddDistLat
back into a BoundedLatticeHom
.
Equations
Instances For
Typecheck a BoundedLatticeHom
as a morphism in FinBddDistLat
.
Equations
Instances For
Use the ConcreteCategory.hom
projection for @[simps]
lemmas.
Equations
- FinBddDistLat.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
- FinBddDistLat.instInhabited = { default := FinBddDistLat.of PUnit.{?u.3 + 1} }
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 finite 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 FinBddDistLat
and itself induced by OrderDual
both ways.
Equations
- One or more equations did not get rendered due to their size.