Documentation

Mathlib.CategoryTheory.Category.Basic

Categories #

Defines a category, as a type class parametrised by the type of objects.

Notations #

Introduces notations in the CategoryTheory scope

Users may like to add g ⊚ f for composition in the standard convention, using

local notation:80 g " ⊚ " f:80 => CategoryTheory.CategoryStruct.comp f g    -- type as \oo
class CategoryTheory.CategoryStruct (obj : Type u) extends Quiver obj :
Type (max u (v + 1))

A preliminary structure on the way to defining a category, containing the data, but none of the axioms.

  • Hom : objobjType v
  • id (X : obj) : X X

    The identity morphism on an object.

  • comp {X Y Z : obj} : (X Y) → (Y Z) → (X Z)

    Composition of morphisms in a category, written f ≫ g.

Instances

    Notation for the identity morphism in a category.

    Equations
    Instances For

      Notation for composition of morphisms in a category.

      Equations
      Instances For

        Close the main goal with sorry if its type contains sorry, and fail otherwise.

        Equations
        Instances For

          Close the main goal with sorry if its type contains sorry, and fail otherwise.

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

            rfl_cat is a macro for intros; rfl which is attempted in aesop_cat before doing the more expensive aesop tactic.

            This gives a speedup because simp (called by aesop) is too slow. There is a fix for this slowness in https://github.com/leanprover/lean4/pull/7428. So, when that is resolved, the performance impact of rfl_cat should be measured again.

            Note on refine id ?_: In some cases it is important that the type of the proof matches the expected type exactly. e.g. if the goal is 2 = 1 + 1, the rfl tactic will give a proof of type 2 = 2. Starting a proof with refine id ?_ is a trick to make sure that the proof has exactly the expected type, in this case 2 = 1 + 1.

            Equations
            Instances For

              A thin wrapper for aesop which adds the CategoryTheory rule set and allows aesop to look through semireducible definitions when calling intros. This tactic fails when it is unable to solve the goal, making it suitable for use in auto-params.

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

                We also use aesop_cat? to pass along a Try this suggestion when using aesop_cat

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

                  A variant of aesop_cat which does not fail when it is unable to solve the goal. Use this only for exploration! Nonterminal aesop is even worse than nonterminal simp.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    class CategoryTheory.Category (obj : Type u) extends CategoryTheory.CategoryStruct.{v, u} obj :
                    Type (max u (v + 1))

                    The typeclass Category C describes morphisms associated to objects of type C. The universe levels of the objects and morphisms are unconstrained, and will often need to be specified explicitly, as Category.{v} C. (See also LargeCategory and SmallCategory.)

                    Instances
                      @[reducible, inline]
                      abbrev CategoryTheory.LargeCategory (C : Type (u + 1)) :
                      Type (u + 1)

                      A LargeCategory has objects in one universe level higher than the universe level of the morphisms. It is useful for examples such as the category of types, or the category of groups, etc.

                      Equations
                      Instances For
                        @[reducible, inline]
                        abbrev CategoryTheory.SmallCategory (C : Type u) :
                        Type (u + 1)

                        A SmallCategory has objects and morphisms in the same universe level.

                        Equations
                        Instances For
                          theorem CategoryTheory.eq_whisker {C : Type u} [Category.{v, u} C] {X Y Z : C} {f g : X Y} (w : f = g) (h : Y Z) :

                          postcompose an equation between morphisms by another morphism

                          theorem CategoryTheory.whisker_eq {C : Type u} [Category.{v, u} C] {X Y Z : C} (f : X Y) {g h : Y Z} (w : g = h) :

                          precompose an equation between morphisms by another morphism

                          Notation for whiskering an equation by a morphism (on the right). If f g : X ⟶ Y and w : f = g and h : Y ⟶ Z, then w =≫ h : f ≫ h = g ≫ h.

                          Equations
                          Instances For

                            Notation for whiskering an equation by a morphism (on the left). If g h : Y ⟶ Z and w : g = h and h : X ⟶ Y, then f ≫= w : f ≫ g = f ≫ h.

                            Equations
                            Instances For
                              theorem CategoryTheory.eq_of_comp_left_eq {C : Type u} [Category.{v, u} C] {X Y : C} {f g : X Y} (w : ∀ {Z : C} (h : Y Z), CategoryStruct.comp f h = CategoryStruct.comp g h) :
                              f = g
                              theorem CategoryTheory.eq_of_comp_right_eq {C : Type u} [Category.{v, u} C] {Y Z : C} {f g : Y Z} (w : ∀ {X : C} (h : X Y), CategoryStruct.comp h f = CategoryStruct.comp h g) :
                              f = g
                              theorem CategoryTheory.eq_of_comp_left_eq' {C : Type u} [Category.{v, u} C] {X Y : C} (f g : X Y) (w : (fun {Z : C} (h : Y Z) => CategoryStruct.comp f h) = fun {Z : C} (h : Y Z) => CategoryStruct.comp g h) :
                              f = g
                              theorem CategoryTheory.eq_of_comp_right_eq' {C : Type u} [Category.{v, u} C] {Y Z : C} (f g : Y Z) (w : (fun {X : C} (h : X Y) => CategoryStruct.comp h f) = fun {X : C} (h : X Y) => CategoryStruct.comp h g) :
                              f = g
                              theorem CategoryTheory.id_of_comp_left_id {C : Type u} [Category.{v, u} C] {X : C} (f : X X) (w : ∀ {Y : C} (g : X Y), CategoryStruct.comp f g = g) :
                              theorem CategoryTheory.id_of_comp_right_id {C : Type u} [Category.{v, u} C] {X : C} (f : X X) (w : ∀ {Y : C} (g : Y X), CategoryStruct.comp g f = g) :
                              theorem CategoryTheory.comp_ite {C : Type u} [Category.{v, u} C] {P : Prop} [Decidable P] {X Y Z : C} (f : X Y) (g g' : Y Z) :
                              theorem CategoryTheory.ite_comp {C : Type u} [Category.{v, u} C] {P : Prop} [Decidable P] {X Y Z : C} (f f' : X Y) (g : Y Z) :
                              theorem CategoryTheory.comp_dite {C : Type u} [Category.{v, u} C] {P : Prop} [Decidable P] {X Y Z : C} (f : X Y) (g : P → (Y Z)) (g' : ¬P → (Y Z)) :
                              CategoryStruct.comp f (if h : P then g h else g' h) = if h : P then CategoryStruct.comp f (g h) else CategoryStruct.comp f (g' h)
                              theorem CategoryTheory.dite_comp {C : Type u} [Category.{v, u} C] {P : Prop} [Decidable P] {X Y Z : C} (f : P → (X Y)) (f' : ¬P → (X Y)) (g : Y Z) :
                              CategoryStruct.comp (if h : P then f h else f' h) g = if h : P then CategoryStruct.comp (f h) g else CategoryStruct.comp (f' h) g
                              class CategoryTheory.Epi {C : Type u} [Category.{v, u} C] {X Y : C} (f : X Y) :

                              A morphism f is an epimorphism if it can be cancelled when precomposed: f ≫ g = f ≫ h implies g = h.

                              Instances
                                class CategoryTheory.Mono {C : Type u} [Category.{v, u} C] {X Y : C} (f : X Y) :

                                A morphism f is a monomorphism if it can be cancelled when postcomposed: g ≫ f = h ≫ f implies g = h.

                                Instances
                                  theorem CategoryTheory.cancel_epi {C : Type u} [Category.{v, u} C] {X Y Z : C} (f : X Y) [Epi f] {g h : Y Z} :
                                  theorem CategoryTheory.cancel_mono {C : Type u} [Category.{v, u} C] {X Y Z : C} (f : X Y) [Mono f] {g h : Z X} :
                                  theorem CategoryTheory.cancel_epi_id {C : Type u} [Category.{v, u} C] {X Y : C} (f : X Y) [Epi f] {h : Y Y} :
                                  theorem CategoryTheory.cancel_mono_id {C : Type u} [Category.{v, u} C] {X Y : C} (f : X Y) [Mono f] {g : X X} :
                                  instance CategoryTheory.epi_comp {C : Type u} [Category.{v, u} C] {X Y Z : C} (f : X Y) [Epi f] (g : Y Z) [Epi g] :

                                  The composition of epimorphisms is again an epimorphism. This version takes Epi f and Epi g as typeclass arguments. For a version taking them as explicit arguments, see epi_comp'.

                                  theorem CategoryTheory.epi_comp' {C : Type u} [Category.{v, u} C] {X Y Z : C} {f : X Y} {g : Y Z} (hf : Epi f) (hg : Epi g) :

                                  The composition of epimorphisms is again an epimorphism. This version takes Epi f and Epi g as explicit arguments. For a version taking them as typeclass arguments, see epi_comp.

                                  instance CategoryTheory.mono_comp {C : Type u} [Category.{v, u} C] {X Y Z : C} (f : X Y) [Mono f] (g : Y Z) [Mono g] :

                                  The composition of monomorphisms is again a monomorphism. This version takes Mono f and Mono g as typeclass arguments. For a version taking them as explicit arguments, see mono_comp'.

                                  theorem CategoryTheory.mono_comp' {C : Type u} [Category.{v, u} C] {X Y Z : C} {f : X Y} {g : Y Z} (hf : Mono f) (hg : Mono g) :

                                  The composition of monomorphisms is again a monomorphism. This version takes Mono f and Mono g as explicit arguments. For a version taking them as typeclass arguments, see mono_comp.

                                  theorem CategoryTheory.mono_of_mono {C : Type u} [Category.{v, u} C] {X Y Z : C} (f : X Y) (g : Y Z) [Mono (CategoryStruct.comp f g)] :
                                  theorem CategoryTheory.mono_of_mono_fac {C : Type u} [Category.{v, u} C] {X Y Z : C} {f : X Y} {g : Y Z} {h : X Z} [Mono h] (w : CategoryStruct.comp f g = h) :
                                  theorem CategoryTheory.epi_of_epi {C : Type u} [Category.{v, u} C] {X Y Z : C} (f : X Y) (g : Y Z) [Epi (CategoryStruct.comp f g)] :
                                  Epi g
                                  theorem CategoryTheory.epi_of_epi_fac {C : Type u} [Category.{v, u} C] {X Y Z : C} {f : X Y} {g : Y Z} {h : X Z} [Epi h] (w : CategoryStruct.comp f g = h) :
                                  Epi g
                                  instance CategoryTheory.instMono {C : Type u} [Category.{v, u} C] {X Y : C} [Quiver.IsThin C] (f : X Y) :
                                  instance CategoryTheory.instEpi {C : Type u} [Category.{v, u} C] {X Y : C} [Quiver.IsThin C] (f : X Y) :
                                  Epi f
                                  Equations
                                  • One or more equations did not get rendered due to their size.