Documentation

Mathlib.CategoryTheory.DiscreteCategory

Discrete categories #

We define Discrete α as a structure containing a term a : α for any type α, and use this type alias to provide a SmallCategory instance whose only morphisms are the identities.

There is an annoying technical difficulty that it has turned out to be inconvenient to allow categories with morphisms living in Prop, so instead of defining X ⟶ Y in Discrete α as X = Y, one might define it as PLift (X = Y). In fact, to allow Discrete α to be a SmallCategory (i.e. with morphisms in the same universe as the objects), we actually define the hom type X ⟶ Y as ULift (PLift (X = Y)).

Discrete.functor promotes a function f : I → C (for any category C) to a functor Discrete.functor f : Discrete I ⥤ C.

Similarly, Discrete.natTrans and Discrete.natIso promote I-indexed families of morphisms, or I-indexed families of isomorphisms to natural transformations or natural isomorphism.

We show equivalences of types are the same as (categorical) equivalences of the corresponding discrete categories.

theorem CategoryTheory.Discrete.ext {α : Type u₁} (x : CategoryTheory.Discrete α) (y : CategoryTheory.Discrete α) (as : x.as = y.as) :
x = y
structure CategoryTheory.Discrete (α : Type u₁) :
Type u₁

A wrapper for promoting any type to a category, with the only morphisms being equalities.

  • as : α

    A wrapper for promoting any type to a category, with the only morphisms being equalities.

Instances For
    @[simp]
    theorem CategoryTheory.Discrete.mk_as {α : Type u₁} (X : CategoryTheory.Discrete α) :
    { as := X.as } = X
    @[simp]
    theorem CategoryTheory.discreteEquiv_apply {α : Type u₁} (self : CategoryTheory.Discrete α) :
    CategoryTheory.discreteEquiv self = self.as
    @[simp]
    theorem CategoryTheory.discreteEquiv_symm_apply_as {α : Type u₁} (as : α) :
    (CategoryTheory.discreteEquiv.symm as).as = as

    Discrete α is equivalent to the original type α.

    Equations
    • CategoryTheory.discreteEquiv = { toFun := CategoryTheory.Discrete.as, invFun := CategoryTheory.Discrete.mk, left_inv := , right_inv := }
    Instances For
      Equations

      The "Discrete" category on a type, whose morphisms are equalities.

      Because we do not allow morphisms in Prop (only in Type), somewhat annoyingly we have to define X ⟶ Y as ULift (PLift (X = Y)).

      See

      Equations
      Equations
      • CategoryTheory.Discrete.instInhabitedDiscrete = { default := { as := default } }

      A simple tactic to run cases on any Discrete α hypotheses.

      Equations
      Instances For

        Use:

        attribute [local aesop safe tactic (rule_sets := [CategoryTheory])]
          CategoryTheory.Discrete.discreteCases
        

        to locally gives aesop_cat the ability to call cases on Discrete and (_ : Discrete _) ⟶ (_ : Discrete _) hypotheses.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          Equations
          theorem CategoryTheory.Discrete.eq_of_hom {α : Type u₁} {X : CategoryTheory.Discrete α} {Y : CategoryTheory.Discrete α} (i : X Y) :
          X.as = Y.as

          Extract the equation from a morphism in a discrete category.

          @[inline, reducible]
          abbrev CategoryTheory.Discrete.eqToHom {α : Type u₁} {X : CategoryTheory.Discrete α} {Y : CategoryTheory.Discrete α} (h : X.as = Y.as) :
          X Y

          Promote an equation between the wrapped terms in X Y : Discrete α to a morphism X ⟶ Y in the discrete category.

          Equations
          Instances For
            @[inline, reducible]
            abbrev CategoryTheory.Discrete.eqToIso {α : Type u₁} {X : CategoryTheory.Discrete α} {Y : CategoryTheory.Discrete α} (h : X.as = Y.as) :
            X Y

            Promote an equation between the wrapped terms in X Y : Discrete α to an isomorphism X ≅ Y in the discrete category.

            Equations
            Instances For
              @[inline, reducible]
              abbrev CategoryTheory.Discrete.eqToHom' {α : Type u₁} {a : α} {b : α} (h : a = b) :
              { as := a } { as := b }

              A variant of eqToHom that lifts terms to the discrete category.

              Equations
              Instances For
                @[inline, reducible]
                abbrev CategoryTheory.Discrete.eqToIso' {α : Type u₁} {a : α} {b : α} (h : a = b) :
                { as := a } { as := b }

                A variant of eqToIso that lifts terms to the discrete category.

                Equations
                Instances For
                  @[simp]
                  theorem CategoryTheory.Discrete.id_def {α : Type u₁} (X : CategoryTheory.Discrete α) :
                  { down := { down := } } = CategoryTheory.CategoryStruct.id X

                  Any function I → C gives a functor Discrete I ⥤ C.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem CategoryTheory.Discrete.functor_obj {C : Type u₂} [CategoryTheory.Category.{v₂, u₂} C] {I : Type u₁} (F : IC) (i : I) :
                    (CategoryTheory.Discrete.functor F).obj { as := i } = F i

                    The discrete functor induced by a composition of maps can be written as a composition of two discrete functors.

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

                      For functors out of a discrete category, a natural transformation is just a collection of maps, as the naturality squares are trivial.

                      Equations
                      Instances For

                        For functors out of a discrete category, a natural isomorphism is just a collection of isomorphisms, as the naturality squares are trivial.

                        Equations
                        Instances For

                          Every functor F from a discrete category is naturally isomorphic (actually, equal) to Discrete.functor (F.obj).

                          Equations
                          Instances For

                            Composing Discrete.functor F with another functor G amounts to composing F with G.obj

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For
                              @[simp]
                              theorem CategoryTheory.Discrete.equivalence_inverse {I : Type u₁} {J : Type u₂} (e : I J) :
                              (CategoryTheory.Discrete.equivalence e).inverse = CategoryTheory.Discrete.functor (CategoryTheory.Discrete.mk e.symm)
                              @[simp]
                              theorem CategoryTheory.Discrete.equivalence_functor {I : Type u₁} {J : Type u₂} (e : I J) :
                              (CategoryTheory.Discrete.equivalence e).functor = CategoryTheory.Discrete.functor (CategoryTheory.Discrete.mk e)

                              We can promote a type-level Equiv to an equivalence between the corresponding discrete categories.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                @[simp]
                                theorem CategoryTheory.Discrete.equivOfEquivalence_symm_apply {α : Type u₁} {β : Type u₂} (h : CategoryTheory.Discrete α CategoryTheory.Discrete β) :
                                ∀ (a : β), (CategoryTheory.Discrete.equivOfEquivalence h).symm a = (CategoryTheory.Discrete.as h.inverse.obj CategoryTheory.Discrete.mk) a
                                @[simp]
                                theorem CategoryTheory.Discrete.equivOfEquivalence_apply {α : Type u₁} {β : Type u₂} (h : CategoryTheory.Discrete α CategoryTheory.Discrete β) :
                                ∀ (a : α), (CategoryTheory.Discrete.equivOfEquivalence h) a = (CategoryTheory.Discrete.as h.functor.obj CategoryTheory.Discrete.mk) a

                                We can convert an equivalence of discrete categories to a type-level Equiv.

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

                                  A discrete category is equivalent to its opposite category.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    @[simp]
                                    theorem CategoryTheory.piEquivalenceFunctorDiscrete_counitIso (J : Type u₂) (C : Type u₁) [CategoryTheory.Category.{v₁, u₁} C] :
                                    (CategoryTheory.piEquivalenceFunctorDiscrete J C).counitIso = CategoryTheory.NatIso.ofComponents (fun (F : CategoryTheory.Functor (CategoryTheory.Discrete J) C) => CategoryTheory.NatIso.ofComponents (fun (j : CategoryTheory.Discrete J) => CategoryTheory.Iso.refl (((CategoryTheory.Functor.comp { toPrefunctor := { obj := fun (F : CategoryTheory.Functor (CategoryTheory.Discrete J) C) (j : J) => F.obj { as := j }, map := fun {X Y : CategoryTheory.Functor (CategoryTheory.Discrete J) C} (f : X Y) (j : J) => f.app { as := j } }, map_id := , map_comp := } { toPrefunctor := { obj := fun (F : JC) => CategoryTheory.Discrete.functor F, map := fun {X Y : JC} (f : X Y) => CategoryTheory.Discrete.natTrans fun (j : CategoryTheory.Discrete J) => f j.as }, map_id := , map_comp := }).obj F).obj j)) )

                                    The equivalence of categories (J → C) ≌ (Discrete J ⥤ C).

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