Documentation

Mathlib.Order.Category.Preord

Category of preorders #

This defines Preord, the category of preorders with monotone maps.

structure Preord :
Type (u_1 + 1)

The category of preorders.

  • carrier : Type u_1

    The underlying preordered type.

  • str : Preorder self
Instances For
    @[reducible, inline]
    abbrev Preord.of (X : Type u) [Preorder X] :

    Construct a bundled Preord from the underlying type and typeclass.

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

      The type of morphisms in Preord R.

      Instances For
        theorem Preord.Hom.ext {X Y : Preord} {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]
        abbrev Preord.Hom.hom {X Y : Preord} (f : X.Hom Y) :
        X →o Y

        Turn a morphism in Preord back into a OrderHom.

        Equations
        Instances For
          @[reducible, inline]
          abbrev Preord.ofHom {X Y : Type u} [Preorder X] [Preorder Y] (f : X →o Y) :
          of X of Y

          Typecheck a OrderHom as a morphism in Preord.

          Equations
          Instances For
            def Preord.Hom.Simps.hom (X Y : Preord) (f : X.Hom Y) :
            X →o Y

            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 Preord.ext {X Y : Preord} {f g : X Y} (w : ∀ (x : X), (CategoryTheory.ConcreteCategory.hom f) x = (CategoryTheory.ConcreteCategory.hom g) x) :
              f = g
              theorem Preord.coe_of (X : Type u) [Preorder X] :
              (of X) = X
              @[simp]
              theorem Preord.hom_comp {X Y Z : Preord} (f : X Y) (g : Y Z) :
              theorem Preord.hom_ext {X Y : Preord} {f g : X Y} (hf : Hom.hom f = Hom.hom g) :
              f = g
              @[simp]
              theorem Preord.hom_ofHom {X Y : Type u} [Preorder X] [Preorder Y] (f : X →o Y) :
              @[simp]
              theorem Preord.ofHom_hom {X Y : Preord} (f : X Y) :
              @[simp]
              theorem Preord.ofHom_comp {X Y Z : Type u} [Preorder X] [Preorder Y] [Preorder Z] (f : X →o Y) (g : Y →o Z) :
              theorem Preord.ofHom_apply {X Y : Type u} [Preorder X] [Preorder Y] (f : X →o Y) (x : X) :
              def Preord.Iso.mk {α β : Preord} (e : α ≃o β) :
              α β

              Constructs an equivalence between preorders from an order isomorphism between them.

              Equations
              Instances For
                @[simp]
                theorem Preord.Iso.mk_inv {α β : Preord} (e : α ≃o β) :
                (mk e).inv = ofHom e.symm
                @[simp]
                theorem Preord.Iso.mk_hom {α β : Preord} (e : α ≃o β) :
                (mk e).hom = ofHom e

                OrderDual as a functor.

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

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

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

                    The embedding of Preord into Cat.

                    Equations
                    Instances For
                      @[simp]
                      theorem preordToCat_map {X✝ Y✝ : Preord} (f : X✝ Y✝) :