Documentation

Mathlib.Order.Category.BddDistLat

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.

structure BddDistLatextends DistLat :
Type (u_1 + 1)

The category of bounded distributive lattices with bounded lattice morphisms.

Instances For
    @[reducible, inline]

    Construct a bundled BddDistLat from a BoundedOrder DistribLattice.

    Equations
    Instances For
      theorem BddDistLat.coe_of (α : Type u_1) [DistribLattice α] [BoundedOrder α] :
      (of α).toDistLat = α
      structure BddDistLat.Hom (X Y : BddDistLat) :

      The type of morphisms in BddDistLat R.

      Instances For
        theorem BddDistLat.Hom.ext {X Y : BddDistLat} {x y : X.Hom Y} (hom' : x.hom' = y.hom') :
        x = y
        Equations
        • One or more equations did not get rendered due to their size.
        @[reducible, inline]

        Turn a morphism in BddDistLat back into a BoundedLatticeHom.

        Equations
        Instances For
          @[reducible, inline]

          Typecheck a BoundedLatticeHom as a morphism in BddDistLat.

          Equations
          Instances For

            Use the ConcreteCategory.hom projection for @[simps] lemmas.

            Equations
            Instances For

              The results below duplicate the ConcreteCategory simp lemmas, but we can keep them for dsimp.

              @[simp]
              theorem BddDistLat.hom_ext {X Y : BddDistLat} {f g : X Y} (hf : Hom.hom f = Hom.hom g) :
              f = g
              @[simp]
              theorem BddDistLat.ofHom_hom {X Y : BddDistLat} (f : X Y) :

              Turn a BddDistLat into a BddLat by forgetting it is distributive.

              Equations
              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.
                def BddDistLat.Iso.mk {α β : BddDistLat} (e : α.toDistLat ≃o β.toDistLat) :
                α β

                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
                  @[simp]
                  theorem BddDistLat.Iso.mk_inv {α β : BddDistLat} (e : α.toDistLat ≃o β.toDistLat) :
                  (mk e).inv = ofHom (let __src := { toFun := e.symm, map_sup' := , map_inf' := }; { toFun := e.symm, map_sup' := , map_inf' := , map_top' := , map_bot' := })
                  @[simp]
                  theorem BddDistLat.Iso.mk_hom {α β : BddDistLat} (e : α.toDistLat ≃o β.toDistLat) :
                  (mk e).hom = ofHom (let __src := { toFun := e, map_sup' := , map_inf' := }; { toFun := e, map_sup' := , map_inf' := , map_top' := , map_bot' := })

                  OrderDual as a functor.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem BddDistLat.dual_map {X✝ Y✝ : BddDistLat} (f : X✝ Y✝) :

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

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