Documentation

Mathlib.Order.Category.CompleteLat

The category of complete lattices #

This file defines CompleteLat, the category of complete lattices.

structure CompleteLat :
Type (u_1 + 1)

The category of complete lattices.

Instances For
    theorem CompleteLat.coe_of (α : Type u_1) [CompleteLattice α] :
    { carrier := α, str := inst✝ } = α
    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.
    Equations
    • One or more equations did not get rendered due to their size.
    def CompleteLat.Iso.mk {α β : CompleteLat} (e : α ≃o β) :
    α β

    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_hom {α β : CompleteLat} (e : α ≃o β) :
      (mk e).hom = CategoryTheory.ConcreteCategory.ofHom { toFun := e, map_sInf' := , map_sSup' := }
      @[simp]
      theorem CompleteLat.Iso.mk_inv {α β : CompleteLat} (e : α ≃o β) :
      (mk e).inv = CategoryTheory.ConcreteCategory.ofHom { toFun := e.symm, map_sInf' := , map_sSup' := }

      OrderDual as a functor.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        theorem CompleteLat.dual_map {x✝ x✝¹ : CompleteLat} (a : CompleteLatticeHom x✝ x✝¹) :

        The equivalence between CompleteLat and itself induced by OrderDual both ways.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For