The category of complete lattices #
This file defines CompleteLat
, the category of complete lattices.
The category of complete lattices.
- carrier : Type u_1
The underlying lattice.
- str : CompleteLattice ↑self
Instances For
Equations
- CompleteLat.instCoeSortType = { coe := CompleteLat.carrier }
@[reducible, inline]
Construct a bundled CompleteLat
from the underlying type and typeclass.
Equations
Instances For
Equations
- CompleteLat.instInhabited = { default := CompleteLat.of PUnit.{?u.3 + 1} }
instance
CompleteLat.instConcreteCategoryCompleteLatticeHomCarrier :
CategoryTheory.ConcreteCategory CompleteLat fun (x1 x2 : CompleteLat) => CompleteLatticeHom ↑x1 ↑x2
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 isomorphism of complete lattices from an order isomorphism between them.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
theorem
CompleteLat.Iso.mk_inv
{α β : CompleteLat}
(e : ↑α ≃o ↑β)
:
(mk e).inv = CategoryTheory.ConcreteCategory.ofHom { toFun := ⇑e.symm, map_sInf' := ⋯, map_sSup' := ⋯ }
@[simp]
OrderDual
as a functor.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
The equivalence between CompleteLat
and itself induced by OrderDual
both ways.
Equations
- One or more equations did not get rendered due to their size.