Documentation

Mathlib.Order.Category.FinBddDistLat

The category of finite bounded distributive lattices #

This file defines FinBddDistLat, the category of finite distributive lattices with bounded lattice homomorphisms.

structure FinBddDistLatextends BddDistLat :
Type (u_1 + 1)

The category of finite distributive lattices with bounded lattice morphisms.

Instances For
    @[reducible, inline]

    Construct a bundled FinBddDistLat from a Fintype BoundedOrder DistribLattice.

    Equations
    Instances For
      @[reducible, inline]

      Construct a bundled FinBddDistLat from a Nonempty Fintype DistribLattice.

      Equations
      Instances For
        structure FinBddDistLat.Hom (X Y : FinBddDistLat) :

        The type of morphisms in FinBddDistLat R.

        Instances For
          theorem FinBddDistLat.Hom.ext {X Y : FinBddDistLat} {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 FinBddDistLat back into a BoundedLatticeHom.

          Equations
          Instances For
            @[reducible, inline]

            Typecheck a BoundedLatticeHom as a morphism in FinBddDistLat.

            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.

                theorem FinBddDistLat.hom_ext {X Y : FinBddDistLat} {f g : X Y} (hf : Hom.hom f = Hom.hom g) :
                f = g
                @[simp]
                theorem FinBddDistLat.ofHom_hom {X Y : FinBddDistLat} (f : X Y) :
                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 FinBddDistLat.Iso.mk {α β : FinBddDistLat} (e : α.toDistLat ≃o β.toDistLat) :
                α β

                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
                  @[simp]
                  theorem FinBddDistLat.Iso.mk_hom {α β : FinBddDistLat} (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' := })
                  @[simp]
                  theorem FinBddDistLat.Iso.mk_inv {α β : FinBddDistLat} (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' := })

                  OrderDual as a functor.

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

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

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