Documentation

Mathlib.CategoryTheory.Limits.HasLimits

Existence of limits and colimits #

In CategoryTheory.Limits.IsLimit we defined IsLimit c, the data showing that a cone c is a limit cone.

The two main structures defined in this file are:

HasLimit is a propositional typeclass (it's important that it is a proposition merely asserting the existence of a limit, as otherwise we would have non-defeq problems from incompatible instances).

While HasLimit only asserts the existence of a limit cone, we happily use the axiom of choice in mathlib, so there are convenience functions all depending on HasLimit F:

Key to using the HasLimit interface is that there is an @[ext] lemma stating that to check f = g, for f g : Z ⟶ limit F, it suffices to check f ≫ limit.π F j = g ≫ limit.π F j for every j. This, combined with @[simp] lemmas, makes it possible to prove many easy facts about limits using automation (e.g. tidy).

There are abbreviations HasLimitsOfShape J C and HasLimits C asserting the existence of classes of limits. Later more are introduced, for finite limits, special shapes of limits, etc.

Ideally, many results about limits should be stated first in terms of IsLimit, and then a result in terms of HasLimit derived from this. At this point, however, this is far from uniformly achieved in mathlib --- often statements are only written in terms of HasLimit.

Implementation #

At present we simply say everything twice, in order to handle both limits and colimits. It would be highly desirable to have some automation support, e.g. a @[dualize] attribute that behaves similarly to @[to_additive].

References #

structure CategoryTheory.Limits.LimitCone {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) :
Type (max (max u u₁) v)

LimitCone F contains a cone over F together with the information that it is a limit.

  • cone : Cone F

    The cone itself

  • isLimit : IsLimit self.cone

    The proof that is the limit cone

Instances For

    HasLimit F represents the mere existence of a limit for F.

    Instances

      Use the axiom of choice to extract explicit LimitCone F from HasLimit F.

      Equations
      Instances For

        C has limits of shape J if there exists a limit for every functor F : J ⥤ C.

        • has_limit (F : Functor J C) : HasLimit F

          All functors F : J ⥤ C from J have limits

        Instances

          C has all limits of size v₁ u₁ (HasLimitsOfSize.{v₁ u₁} C) if it has limits of every shape J : Type u₁ with [Category.{v₁} J].

          Instances
            @[reducible, inline]

            C has all (small) limits if it has limits of every shape that is as big as its hom-sets.

            Equations
            Instances For

              An arbitrary choice of limit cone for a functor.

              Equations
              Instances For

                An arbitrary choice of limit object of a functor.

                Equations
                Instances For
                  def CategoryTheory.Limits.limit.π {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] (j : J) :
                  limit F F.obj j

                  The projection from the limit object to a value of the functor.

                  Equations
                  Instances For
                    theorem CategoryTheory.Limits.limit.π_comp_eqToHom {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] {j j' : J} (hj : j = j') :
                    CategoryStruct.comp (π F j) (eqToHom ) = π F j'
                    theorem CategoryTheory.Limits.limit.π_comp_eqToHom_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] {j j' : J} (hj : j = j') {Z : C} (h : F.obj j' Z) :
                    @[simp]
                    @[simp]
                    theorem CategoryTheory.Limits.limit.cone_π {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] :
                    (cone F).app = π F
                    @[simp]
                    theorem CategoryTheory.Limits.limit.w {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] {j j' : J} (f : j j') :
                    CategoryStruct.comp (π F j) (F.map f) = π F j'
                    @[simp]
                    theorem CategoryTheory.Limits.limit.w_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] {j j' : J} (f : j j') {Z : C} (h : F.obj j' Z) :

                    Evidence that the arbitrary choice of cone provided by limit.cone F is a limit cone.

                    Equations
                    Instances For
                      def CategoryTheory.Limits.limit.lift {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] (c : Cone F) :
                      c.pt limit F

                      The morphism from the cone point of any other cone to the limit object.

                      Equations
                      Instances For
                        @[simp]
                        theorem CategoryTheory.Limits.limit.isLimit_lift {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] (c : Cone F) :
                        (isLimit F).lift c = lift F c
                        @[simp]
                        theorem CategoryTheory.Limits.limit.lift_π {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] (c : Cone F) (j : J) :
                        CategoryStruct.comp (lift F c) (π F j) = c.app j
                        @[simp]
                        theorem CategoryTheory.Limits.limit.lift_π_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] (c : Cone F) (j : J) {Z : C} (h : F.obj j Z) :
                        def CategoryTheory.Limits.limMap {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasLimit F] [HasLimit G] (α : F G) :

                        Functoriality of limits.

                        Usually this morphism should be accessed through lim.map, but may be needed separately when you have specified limits for the source and target functors, but not necessarily for all functors of shape J.

                        Equations
                        Instances For
                          @[simp]
                          theorem CategoryTheory.Limits.limMap_π {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasLimit F] [HasLimit G] (α : F G) (j : J) :
                          @[simp]
                          theorem CategoryTheory.Limits.limMap_π_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasLimit F] [HasLimit G] (α : F G) (j : J) {Z : C} (h : G.obj j Z) :

                          The cone morphism from any cone to the arbitrary choice of limit cone.

                          Equations
                          Instances For
                            @[simp]
                            theorem CategoryTheory.Limits.limit.coneMorphism_π {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] (c : Cone F) (j : J) :
                            CategoryStruct.comp (coneMorphism c).hom (π F j) = c.app j
                            @[simp]
                            theorem CategoryTheory.Limits.limit.conePointUniqueUpToIso_hom_comp {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] {c : Cone F} (hc : IsLimit c) (j : J) :
                            CategoryStruct.comp (hc.conePointUniqueUpToIso (isLimit F)).hom (π F j) = c.app j
                            @[simp]
                            theorem CategoryTheory.Limits.limit.conePointUniqueUpToIso_hom_comp_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] {c : Cone F} (hc : IsLimit c) (j : J) {Z : C} (h : F.obj j Z) :
                            CategoryStruct.comp (hc.conePointUniqueUpToIso (isLimit F)).hom (CategoryStruct.comp (π F j) h) = CategoryStruct.comp (c.app j) h
                            @[simp]
                            theorem CategoryTheory.Limits.limit.conePointUniqueUpToIso_inv_comp {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] {c : Cone F} (hc : IsLimit c) (j : J) :
                            CategoryStruct.comp ((isLimit F).conePointUniqueUpToIso hc).inv (π F j) = c.app j
                            @[simp]
                            theorem CategoryTheory.Limits.limit.conePointUniqueUpToIso_inv_comp_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] {c : Cone F} (hc : IsLimit c) (j : J) {Z : C} (h : F.obj j Z) :
                            CategoryStruct.comp ((isLimit F).conePointUniqueUpToIso hc).inv (CategoryStruct.comp (π F j) h) = CategoryStruct.comp (c.app j) h
                            theorem CategoryTheory.Limits.limit.existsUnique {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] (t : Cone F) :
                            ∃! l : t.pt limit F, ∀ (j : J), CategoryStruct.comp l (π F j) = t.app j

                            Given any other limit cone for F, the chosen limit F is isomorphic to the cone point.

                            Equations
                            Instances For
                              @[simp]
                              theorem CategoryTheory.Limits.limit.isoLimitCone_hom_π {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] (t : LimitCone F) (j : J) :
                              CategoryStruct.comp (isoLimitCone t).hom (t.cone.app j) = π F j
                              @[simp]
                              theorem CategoryTheory.Limits.limit.isoLimitCone_hom_π_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] (t : LimitCone F) (j : J) {Z : C} (h : F.obj j Z) :
                              @[simp]
                              theorem CategoryTheory.Limits.limit.isoLimitCone_inv_π {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] (t : LimitCone F) (j : J) :
                              CategoryStruct.comp (isoLimitCone t).inv (π F j) = t.cone.app j
                              @[simp]
                              theorem CategoryTheory.Limits.limit.isoLimitCone_inv_π_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] (t : LimitCone F) (j : J) {Z : C} (h : F.obj j Z) :
                              theorem CategoryTheory.Limits.limit.hom_ext {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] {X : C} {f f' : X limit F} (w : ∀ (j : J), CategoryStruct.comp f (π F j) = CategoryStruct.comp f' (π F j)) :
                              f = f'
                              @[simp]
                              theorem CategoryTheory.Limits.limit.lift_map {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasLimit F] [HasLimit G] (c : Cone F) (α : F G) :
                              @[simp]
                              theorem CategoryTheory.Limits.limit.lift_map_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasLimit F] [HasLimit G] (c : Cone F) (α : F G) {Z : C} (h : limit G Z) :
                              def CategoryTheory.Limits.limit.homIso {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] (W : C) :
                              ULift.{u₁, v} (W limit F) F.cones.obj (Opposite.op W)

                              The isomorphism (in Type) between morphisms from a specified object W to the limit object, and cones with cone point W.

                              Equations
                              Instances For
                                @[simp]
                                theorem CategoryTheory.Limits.limit.homIso_hom {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] {W : C} (f : ULift.{u₁, v} (W limit F)) :
                                (homIso F W).hom f = CategoryStruct.comp ((Functor.const J).map f.down) (cone F)
                                def CategoryTheory.Limits.limit.homIso' {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] (W : C) :
                                ULift.{u₁, v} (W limit F) { p : (j : J) → W F.obj j // ∀ {j j' : J} (f : j j'), CategoryStruct.comp (p j) (F.map f) = p j' }

                                The isomorphism (in Type) between morphisms from a specified object W to the limit object, and an explicit componentwise description of cones with cone point W.

                                Equations
                                Instances For
                                  theorem CategoryTheory.Limits.limit.lift_extend {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] (c : Cone F) {X : C} (f : X c.pt) :
                                  lift F (c.extend f) = CategoryStruct.comp f (lift F c)
                                  theorem CategoryTheory.Limits.hasLimitOfIso {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasLimit F] (α : F G) :

                                  If a functor F has a limit, so does any naturally isomorphic functor.

                                  theorem CategoryTheory.Limits.HasLimit.ofConesIso {C : Type u} [Category.{v, u} C] {J K : Type u₁} [Category.{v₁, u₁} J] [Category.{v₂, u₁} K] (F : Functor J C) (G : Functor K C) (h : F.cones G.cones) [HasLimit F] :

                                  If a functor G has the same collection of cones as a functor F which has a limit, then G also has a limit.

                                  The limits of F : J ⥤ C and G : J ⥤ C are isomorphic, if the functors are naturally isomorphic.

                                  Equations
                                  Instances For
                                    @[simp]
                                    @[simp]
                                    theorem CategoryTheory.Limits.HasLimit.isoOfNatIso_hom_π_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasLimit F] [HasLimit G] (w : F G) (j : J) {Z : C} (h : G.obj j Z) :
                                    @[simp]
                                    @[simp]
                                    theorem CategoryTheory.Limits.HasLimit.isoOfNatIso_inv_π_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasLimit F] [HasLimit G] (w : F G) (j : J) {Z : C} (h : F.obj j Z) :
                                    def CategoryTheory.Limits.HasLimit.isoOfEquivalence {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] {G : Functor K C} [HasLimit G] (e : J K) (w : e.functor.comp G F) :

                                    The limits of F : J ⥤ C and G : K ⥤ C are isomorphic, if there is an equivalence e : J ≌ K making the triangle commute up to natural isomorphism.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem CategoryTheory.Limits.HasLimit.isoOfEquivalence_hom_π {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] {G : Functor K C} [HasLimit G] (e : J K) (w : e.functor.comp G F) (k : K) :
                                      CategoryStruct.comp (isoOfEquivalence e w).hom (limit.π G k) = CategoryStruct.comp (limit.π F (e.inverse.obj k)) (CategoryStruct.comp (w.inv.app (e.inverse.obj k)) (G.map (e.counit.app k)))
                                      @[simp]
                                      theorem CategoryTheory.Limits.HasLimit.isoOfEquivalence_inv_π {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] {G : Functor K C} [HasLimit G] (e : J K) (w : e.functor.comp G F) (j : J) :
                                      CategoryStruct.comp (isoOfEquivalence e w).inv (limit.π F j) = CategoryStruct.comp (limit.π G (e.functor.obj j)) (w.hom.app j)
                                      def CategoryTheory.Limits.limit.pre {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] (E : Functor K J) [HasLimit (E.comp F)] :
                                      limit F limit (E.comp F)

                                      The canonical morphism from the limit of F to the limit of E ⋙ F.

                                      Equations
                                      Instances For
                                        @[simp]
                                        theorem CategoryTheory.Limits.limit.pre_π {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] (E : Functor K J) [HasLimit (E.comp F)] (k : K) :
                                        CategoryStruct.comp (pre F E) (π (E.comp F) k) = π F (E.obj k)
                                        @[simp]
                                        theorem CategoryTheory.Limits.limit.pre_π_assoc {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] (E : Functor K J) [HasLimit (E.comp F)] (k : K) {Z : C} (h : F.obj (E.obj k) Z) :
                                        CategoryStruct.comp (pre F E) (CategoryStruct.comp (π (E.comp F) k) h) = CategoryStruct.comp (π F (E.obj k)) h
                                        @[simp]
                                        theorem CategoryTheory.Limits.limit.lift_pre {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] (E : Functor K J) [HasLimit (E.comp F)] (c : Cone F) :
                                        CategoryStruct.comp (lift F c) (pre F E) = lift (E.comp F) (Cone.whisker E c)
                                        @[simp]
                                        theorem CategoryTheory.Limits.limit.pre_pre {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasLimit F] (E : Functor K J) [HasLimit (E.comp F)] {L : Type u₃} [Category.{v₃, u₃} L] (D : Functor L K) [h : HasLimit (D.comp (E.comp F))] :
                                        CategoryStruct.comp (pre F E) (pre (E.comp F) D) = pre F (D.comp E)
                                        theorem CategoryTheory.Limits.limit.pre_eq {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimit F] {E : Functor K J} [HasLimit (E.comp F)] (s : LimitCone (E.comp F)) (t : LimitCone F) :
                                        pre F E = CategoryStruct.comp (isoLimitCone t).hom (CategoryStruct.comp (s.isLimit.lift (Cone.whisker E t.cone)) (isoLimitCone s).inv)

                                        If we have particular limit cones available for E ⋙ F and for F, we obtain a formula for limit.pre F E.

                                        def CategoryTheory.Limits.limit.post {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (F : Functor J C) [HasLimit F] (G : Functor C D) [HasLimit (F.comp G)] :
                                        G.obj (limit F) limit (F.comp G)

                                        The canonical morphism from G applied to the limit of F to the limit of F ⋙ G.

                                        Equations
                                        Instances For
                                          @[simp]
                                          theorem CategoryTheory.Limits.limit.post_π {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (F : Functor J C) [HasLimit F] (G : Functor C D) [HasLimit (F.comp G)] (j : J) :
                                          CategoryStruct.comp (post F G) (π (F.comp G) j) = G.map (π F j)
                                          @[simp]
                                          theorem CategoryTheory.Limits.limit.post_π_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (F : Functor J C) [HasLimit F] (G : Functor C D) [HasLimit (F.comp G)] (j : J) {Z : D} (h : G.obj (F.obj j) Z) :
                                          CategoryStruct.comp (post F G) (CategoryStruct.comp (π (F.comp G) j) h) = CategoryStruct.comp (G.map (π F j)) h
                                          @[simp]
                                          theorem CategoryTheory.Limits.limit.lift_post {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (F : Functor J C) [HasLimit F] (G : Functor C D) [HasLimit (F.comp G)] (c : Cone F) :
                                          CategoryStruct.comp (G.map (lift F c)) (post F G) = lift (F.comp G) (G.mapCone c)
                                          @[simp]
                                          theorem CategoryTheory.Limits.limit.post_post {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (F : Functor J C) [HasLimit F] (G : Functor C D) [HasLimit (F.comp G)] {E : Type u''} [Category.{v'', u''} E] (H : Functor D E) [h : HasLimit ((F.comp G).comp H)] :
                                          CategoryStruct.comp (H.map (post F G)) (post (F.comp G) H) = post F (G.comp H)
                                          theorem CategoryTheory.Limits.limit.pre_post {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (E : Functor K J) (F : Functor J C) (G : Functor C D) [HasLimit F] [HasLimit (E.comp F)] [HasLimit (F.comp G)] [h : HasLimit ((E.comp F).comp G)] :
                                          CategoryStruct.comp (G.map (pre F E)) (post (E.comp F) G) = CategoryStruct.comp (post F G) (pre (F.comp G) E)
                                          instance CategoryTheory.Limits.hasLimitEquivalenceComp {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} (e : K J) [HasLimit F] :
                                          HasLimit (e.functor.comp F)
                                          theorem CategoryTheory.Limits.hasLimitOfEquivalenceComp {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} (e : K J) [HasLimit (e.functor.comp F)] :

                                          If a E ⋙ F has a limit, and E is an equivalence, we can construct a limit of F.

                                          limit F is functorial in F, when C has all limits of shape J.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            @[simp]
                                            theorem CategoryTheory.Limits.lim_map {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] [HasLimitsOfShape J C] {X✝ Y✝ : Functor J C} (α : X✝ Y✝) :
                                            lim.map α = limMap α
                                            @[simp]
                                            theorem CategoryTheory.Limits.limMap_eq {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimitsOfShape J C] {G : Functor J C} (α : F G) :
                                            limMap α = lim.map α
                                            theorem CategoryTheory.Limits.limit.map_pre {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimitsOfShape J C] {G : Functor J C} (α : F G) [HasLimitsOfShape K C] (E : Functor K J) :
                                            theorem CategoryTheory.Limits.limit.map_pre' {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] [HasLimitsOfShape J C] [HasLimitsOfShape K C] (F : Functor J C) {E₁ E₂ : Functor K J} (α : E₁ E₂) :
                                            pre F E₂ = CategoryStruct.comp (pre F E₁) (lim.map (whiskerRight α F))
                                            theorem CategoryTheory.Limits.limit.id_pre {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] [HasLimitsOfShape J C] (F : Functor J C) :
                                            pre F (Functor.id J) = lim.map F.leftUnitor.inv
                                            theorem CategoryTheory.Limits.limit.map_post {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasLimitsOfShape J C] {G : Functor J C} (α : F G) {D : Type u'} [Category.{v', u'} D] [HasLimitsOfShape J D] (H : Functor C D) :

                                            The isomorphism between morphisms from W to the cone point of the limit cone for F and cones over F with cone point W is natural in F.

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

                                              The constant functor and limit functor are adjoint to each other

                                              Equations
                                              • One or more equations did not get rendered due to their size.
                                              Instances For
                                                instance CategoryTheory.Limits.limMap_mono' {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasLimitsOfShape J C] (α : F G) [Mono α] :
                                                instance CategoryTheory.Limits.limMap_mono {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasLimit F] [HasLimit G] (α : F G) [∀ (j : J), Mono (α.app j)] :
                                                noncomputable def CategoryTheory.Limits.coneOfAdj {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {L : Functor (Functor J C) C} (adj : Functor.const J L) (F : Functor J C) :

                                                The limit cone obtained from a right adjoint of the constant functor.

                                                Equations
                                                Instances For
                                                  @[simp]
                                                  theorem CategoryTheory.Limits.coneOfAdj_π {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {L : Functor (Functor J C) C} (adj : Functor.const J L) (F : Functor J C) :
                                                  (coneOfAdj adj F) = adj.counit.app F
                                                  @[simp]
                                                  theorem CategoryTheory.Limits.coneOfAdj_pt {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {L : Functor (Functor J C) C} (adj : Functor.const J L) (F : Functor J C) :
                                                  (coneOfAdj adj F).pt = L.obj F

                                                  The cones defined by coneOfAdj are limit cones.

                                                  Equations
                                                  Instances For
                                                    @[simp]
                                                    theorem CategoryTheory.Limits.isLimitConeOfAdj_lift {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {L : Functor (Functor J C) C} (adj : Functor.const J L) (F : Functor J C) (s : Cone F) :
                                                    (isLimitConeOfAdj adj F).lift s = (adj.homEquiv s.pt F) s

                                                    We can transport limits of shape J along an equivalence J ≌ J'.

                                                    hasLimitsOfSizeShrink.{v u} C tries to obtain HasLimitsOfSize.{v u} C from some other HasLimitsOfSize C.

                                                    structure CategoryTheory.Limits.ColimitCocone {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) :
                                                    Type (max (max u u₁) v)

                                                    ColimitCocone F contains a cocone over F together with the information that it is a colimit.

                                                    • cocone : Cocone F

                                                      The cocone itself

                                                    • isColimit : IsColimit self.cocone

                                                      The proof that it is the colimit cocone

                                                    Instances For

                                                      HasColimit F represents the mere existence of a colimit for F.

                                                      Instances

                                                        Use the axiom of choice to extract explicit ColimitCocone F from HasColimit F.

                                                        Equations
                                                        Instances For

                                                          C has colimits of shape J if there exists a colimit for every functor F : J ⥤ C.

                                                          • has_colimit (F : Functor J C) : HasColimit F

                                                            All F : J ⥤ C have colimits for a fixed J

                                                          Instances

                                                            C has all colimits of size v₁ u₁ (HasColimitsOfSize.{v₁ u₁} C) if it has colimits of every shape J : Type u₁ with [Category.{v₁} J].

                                                            Instances
                                                              @[reducible, inline]

                                                              C has all (small) colimits if it has colimits of every shape that is as big as its hom-sets.

                                                              Equations
                                                              Instances For

                                                                An arbitrary choice of colimit cocone of a functor.

                                                                Equations
                                                                Instances For

                                                                  An arbitrary choice of colimit object of a functor.

                                                                  Equations
                                                                  Instances For
                                                                    def CategoryTheory.Limits.colimit.ι {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (j : J) :
                                                                    F.obj j colimit F

                                                                    The coprojection from a value of the functor to the colimit object.

                                                                    Equations
                                                                    Instances For
                                                                      theorem CategoryTheory.Limits.colimit.eqToHom_comp_ι {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] {j j' : J} (hj : j = j') :
                                                                      CategoryStruct.comp (eqToHom ) (ι F j) = ι F j'
                                                                      theorem CategoryTheory.Limits.colimit.eqToHom_comp_ι_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] {j j' : J} (hj : j = j') {Z : C} (h : colimit F Z) :
                                                                      @[simp]
                                                                      theorem CategoryTheory.Limits.colimit.cocone_ι {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] (j : J) :
                                                                      (cocone F).app j = ι F j
                                                                      @[simp]
                                                                      theorem CategoryTheory.Limits.colimit.w {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] {j j' : J} (f : j j') :
                                                                      CategoryStruct.comp (F.map f) (ι F j') = ι F j
                                                                      @[simp]
                                                                      theorem CategoryTheory.Limits.colimit.w_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] {j j' : J} (f : j j') {Z : C} (h : colimit F Z) :

                                                                      Evidence that the arbitrary choice of cocone is a colimit cocone.

                                                                      Equations
                                                                      Instances For

                                                                        The morphism from the colimit object to the cone point of any other cocone.

                                                                        Equations
                                                                        Instances For
                                                                          @[simp]
                                                                          @[simp]
                                                                          theorem CategoryTheory.Limits.colimit.ι_desc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] (c : Cocone F) (j : J) :
                                                                          CategoryStruct.comp (ι F j) (desc F c) = c.app j

                                                                          We have lots of lemmas describing how to simplify colimit.ι F j ≫ _, and combined with colimit.ext we rely on these lemmas for many calculations.

                                                                          However, since Category.assoc is a @[simp] lemma, often expressions are right associated, and it's hard to apply these lemmas about colimit.ι.

                                                                          We thus use reassoc to define additional @[simp] lemmas, with an arbitrary extra morphism. (see Tactic/reassoc_axiom.lean)

                                                                          @[simp]
                                                                          theorem CategoryTheory.Limits.colimit.ι_desc_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] (c : Cocone F) (j : J) {Z : C} (h : c.pt Z) :

                                                                          Functoriality of colimits.

                                                                          Usually this morphism should be accessed through colim.map, but may be needed separately when you have specified colimits for the source and target functors, but not necessarily for all functors of shape J.

                                                                          Equations
                                                                          Instances For
                                                                            @[simp]
                                                                            theorem CategoryTheory.Limits.ι_colimMap {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasColimit F] [HasColimit G] (α : F G) (j : J) :

                                                                            The cocone morphism from the arbitrary choice of colimit cocone to any cocone.

                                                                            Equations
                                                                            Instances For
                                                                              @[simp]
                                                                              theorem CategoryTheory.Limits.colimit.comp_coconePointUniqueUpToIso_hom {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] {c : Cocone F} (hc : IsColimit c) (j : J) :
                                                                              CategoryStruct.comp (ι F j) ((isColimit F).coconePointUniqueUpToIso hc).hom = c.app j
                                                                              @[simp]
                                                                              theorem CategoryTheory.Limits.colimit.comp_coconePointUniqueUpToIso_hom_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] {c : Cocone F} (hc : IsColimit c) (j : J) {Z : C} (h : c.pt Z) :
                                                                              CategoryStruct.comp (ι F j) (CategoryStruct.comp ((isColimit F).coconePointUniqueUpToIso hc).hom h) = CategoryStruct.comp (c.app j) h
                                                                              @[simp]
                                                                              theorem CategoryTheory.Limits.colimit.comp_coconePointUniqueUpToIso_inv {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] {c : Cocone F} (hc : IsColimit c) (j : J) :
                                                                              CategoryStruct.comp (ι F j) (hc.coconePointUniqueUpToIso (isColimit F)).inv = c.app j
                                                                              @[simp]
                                                                              theorem CategoryTheory.Limits.colimit.comp_coconePointUniqueUpToIso_inv_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] {c : Cocone F} (hc : IsColimit c) (j : J) {Z : C} (h : c.pt Z) :
                                                                              CategoryStruct.comp (ι F j) (CategoryStruct.comp (hc.coconePointUniqueUpToIso (isColimit F)).inv h) = CategoryStruct.comp (c.app j) h
                                                                              theorem CategoryTheory.Limits.colimit.existsUnique {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] (t : Cocone F) :
                                                                              ∃! d : colimit F t.pt, ∀ (j : J), CategoryStruct.comp (ι F j) d = t.app j

                                                                              Given any other colimit cocone for F, the chosen colimit F is isomorphic to the cocone point.

                                                                              Equations
                                                                              Instances For
                                                                                @[simp]
                                                                                @[simp]
                                                                                theorem CategoryTheory.Limits.colimit.isoColimitCocone_ι_hom_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] (t : ColimitCocone F) (j : J) {Z : C} (h : t.cocone.pt Z) :
                                                                                @[simp]
                                                                                theorem CategoryTheory.Limits.colimit.hom_ext {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] {X : C} {f f' : colimit F X} (w : ∀ (j : J), CategoryStruct.comp (ι F j) f = CategoryStruct.comp (ι F j) f') :
                                                                                f = f'
                                                                                def CategoryTheory.Limits.colimit.homIso {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (W : C) :
                                                                                ULift.{u₁, v} (colimit F W) F.cocones.obj W

                                                                                The isomorphism (in Type) between morphisms from the colimit object to a specified object W, and cocones with cone point W.

                                                                                Equations
                                                                                Instances For
                                                                                  @[simp]
                                                                                  theorem CategoryTheory.Limits.colimit.homIso_hom {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] {W : C} (f : ULift.{u₁, v} (colimit F W)) :
                                                                                  (homIso F W).hom f = CategoryStruct.comp (cocone F) ((Functor.const J).map f.down)
                                                                                  def CategoryTheory.Limits.colimit.homIso' {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (W : C) :
                                                                                  ULift.{u₁, v} (colimit F W) { p : (j : J) → F.obj j W // ∀ {j j' : J} (f : j j'), CategoryStruct.comp (F.map f) (p j') = p j }

                                                                                  The isomorphism (in Type) between morphisms from the colimit object to a specified object W, and an explicit componentwise description of cocones with cone point W.

                                                                                  Equations
                                                                                  Instances For
                                                                                    theorem CategoryTheory.Limits.colimit.desc_extend {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (c : Cocone F) {X : C} (f : c.pt X) :
                                                                                    desc F (c.extend f) = CategoryStruct.comp (desc F c) f

                                                                                    If F has a colimit, so does any naturally isomorphic functor.

                                                                                    theorem CategoryTheory.Limits.HasColimit.ofCoconesIso {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {K : Type u₁} [Category.{v₂, u₁} K] (F : Functor J C) (G : Functor K C) (h : F.cocones G.cocones) [HasColimit F] :

                                                                                    If a functor G has the same collection of cocones as a functor F which has a colimit, then G also has a colimit.

                                                                                    The colimits of F : J ⥤ C and G : J ⥤ C are isomorphic, if the functors are naturally isomorphic.

                                                                                    Equations
                                                                                    Instances For
                                                                                      def CategoryTheory.Limits.HasColimit.isoOfEquivalence {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] {G : Functor K C} [HasColimit G] (e : J K) (w : e.functor.comp G F) :

                                                                                      The colimits of F : J ⥤ C and G : K ⥤ C are isomorphic, if there is an equivalence e : J ≌ K making the triangle commute up to natural isomorphism.

                                                                                      Equations
                                                                                      Instances For
                                                                                        @[simp]
                                                                                        theorem CategoryTheory.Limits.HasColimit.isoOfEquivalence_hom_π {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] {G : Functor K C} [HasColimit G] (e : J K) (w : e.functor.comp G F) (j : J) :
                                                                                        CategoryStruct.comp (colimit.ι F j) (isoOfEquivalence e w).hom = CategoryStruct.comp (F.map (e.unit.app j)) (CategoryStruct.comp (w.inv.app ((e.functor.comp e.inverse).obj j)) (colimit.ι G (e.functor.obj ((e.functor.comp e.inverse).obj j))))
                                                                                        @[simp]
                                                                                        theorem CategoryTheory.Limits.HasColimit.isoOfEquivalence_inv_π {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] {G : Functor K C} [HasColimit G] (e : J K) (w : e.functor.comp G F) (k : K) :
                                                                                        CategoryStruct.comp (colimit.ι G k) (isoOfEquivalence e w).inv = CategoryStruct.comp (G.map (e.counitInv.app k)) (CategoryStruct.comp (w.hom.app (e.inverse.obj k)) (colimit.ι F (e.inverse.obj k)))
                                                                                        def CategoryTheory.Limits.colimit.pre {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (E : Functor K J) [HasColimit (E.comp F)] :
                                                                                        colimit (E.comp F) colimit F

                                                                                        The canonical morphism from the colimit of E ⋙ F to the colimit of F.

                                                                                        Equations
                                                                                        Instances For
                                                                                          @[simp]
                                                                                          theorem CategoryTheory.Limits.colimit.ι_pre {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (E : Functor K J) [HasColimit (E.comp F)] (k : K) :
                                                                                          CategoryStruct.comp (ι (E.comp F) k) (pre F E) = ι F (E.obj k)
                                                                                          @[simp]
                                                                                          theorem CategoryTheory.Limits.colimit.ι_pre_assoc {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (E : Functor K J) [HasColimit (E.comp F)] (k : K) {Z : C} (h : colimit F Z) :
                                                                                          CategoryStruct.comp (ι (E.comp F) k) (CategoryStruct.comp (pre F E) h) = CategoryStruct.comp (ι F (E.obj k)) h
                                                                                          @[simp]
                                                                                          theorem CategoryTheory.Limits.colimit.ι_inv_pre {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (E : Functor K J) [HasColimit (E.comp F)] [IsIso (pre F E)] (k : K) :
                                                                                          CategoryStruct.comp (ι F (E.obj k)) (inv (pre F E)) = ι (E.comp F) k
                                                                                          @[simp]
                                                                                          theorem CategoryTheory.Limits.colimit.ι_inv_pre_assoc {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (E : Functor K J) [HasColimit (E.comp F)] [IsIso (pre F E)] (k : K) {Z : C} (h : colimit (E.comp F) Z) :
                                                                                          CategoryStruct.comp (ι F (E.obj k)) (CategoryStruct.comp (inv (pre F E)) h) = CategoryStruct.comp (ι (E.comp F) k) h
                                                                                          @[simp]
                                                                                          theorem CategoryTheory.Limits.colimit.pre_desc {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (E : Functor K J) [HasColimit (E.comp F)] (c : Cocone F) :
                                                                                          CategoryStruct.comp (pre F E) (desc F c) = desc (E.comp F) (Cocone.whisker E c)
                                                                                          @[simp]
                                                                                          theorem CategoryTheory.Limits.colimit.pre_desc_assoc {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (E : Functor K J) [HasColimit (E.comp F)] (c : Cocone F) {Z : C} (h : c.pt Z) :
                                                                                          @[simp]
                                                                                          theorem CategoryTheory.Limits.colimit.pre_pre {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] (F : Functor J C) [HasColimit F] (E : Functor K J) [HasColimit (E.comp F)] {L : Type u₃} [Category.{v₃, u₃} L] (D : Functor L K) [h : HasColimit (D.comp (E.comp F))] :
                                                                                          CategoryStruct.comp (pre (E.comp F) D) (pre F E) = pre F (D.comp E)
                                                                                          theorem CategoryTheory.Limits.colimit.pre_eq {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimit F] {E : Functor K J} [HasColimit (E.comp F)] (s : ColimitCocone (E.comp F)) (t : ColimitCocone F) :
                                                                                          pre F E = CategoryStruct.comp (isoColimitCocone s).hom (CategoryStruct.comp (s.isColimit.desc (Cocone.whisker E t.cocone)) (isoColimitCocone t).inv)

                                                                                          If we have particular colimit cocones available for E ⋙ F and for F, we obtain a formula for colimit.pre F E.

                                                                                          def CategoryTheory.Limits.colimit.post {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) {D : Type u'} [Category.{v', u'} D] [HasColimit F] (G : Functor C D) [HasColimit (F.comp G)] :
                                                                                          colimit (F.comp G) G.obj (colimit F)

                                                                                          The canonical morphism from G applied to the colimit of F ⋙ G to G applied to the colimit of F.

                                                                                          Equations
                                                                                          Instances For
                                                                                            @[simp]
                                                                                            theorem CategoryTheory.Limits.colimit.ι_post {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) {D : Type u'} [Category.{v', u'} D] [HasColimit F] (G : Functor C D) [HasColimit (F.comp G)] (j : J) :
                                                                                            CategoryStruct.comp (ι (F.comp G) j) (post F G) = G.map (ι F j)
                                                                                            @[simp]
                                                                                            theorem CategoryTheory.Limits.colimit.ι_post_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) {D : Type u'} [Category.{v', u'} D] [HasColimit F] (G : Functor C D) [HasColimit (F.comp G)] (j : J) {Z : D} (h : G.obj (colimit F) Z) :
                                                                                            CategoryStruct.comp (ι (F.comp G) j) (CategoryStruct.comp (post F G) h) = CategoryStruct.comp (G.map (ι F j)) h
                                                                                            @[simp]
                                                                                            theorem CategoryTheory.Limits.colimit.post_desc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) {D : Type u'} [Category.{v', u'} D] [HasColimit F] (G : Functor C D) [HasColimit (F.comp G)] (c : Cocone F) :
                                                                                            CategoryStruct.comp (post F G) (G.map (desc F c)) = desc (F.comp G) (G.mapCocone c)
                                                                                            @[simp]
                                                                                            theorem CategoryTheory.Limits.colimit.post_post {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] (F : Functor J C) {D : Type u'} [Category.{v', u'} D] [HasColimit F] (G : Functor C D) [HasColimit (F.comp G)] {E : Type u''} [Category.{v'', u''} E] (H : Functor D E) [h : HasColimit ((F.comp G).comp H)] :
                                                                                            CategoryStruct.comp (post (F.comp G) H) (H.map (post F G)) = post F (G.comp H)
                                                                                            theorem CategoryTheory.Limits.colimit.pre_post {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {D : Type u'} [Category.{v', u'} D] (E : Functor K J) (F : Functor J C) (G : Functor C D) [HasColimit F] [HasColimit (E.comp F)] [HasColimit (F.comp G)] [h : HasColimit ((E.comp F).comp G)] :
                                                                                            CategoryStruct.comp (post (E.comp F) G) (G.map (pre F E)) = CategoryStruct.comp (pre (F.comp G) E) (post F G)
                                                                                            instance CategoryTheory.Limits.hasColimit_equivalence_comp {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} (e : K J) [HasColimit F] :
                                                                                            HasColimit (e.functor.comp F)
                                                                                            theorem CategoryTheory.Limits.hasColimit_of_equivalence_comp {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] {F : Functor J C} (e : K J) [HasColimit (e.functor.comp F)] :

                                                                                            If a E ⋙ F has a colimit, and E is an equivalence, we can construct a colimit of F.

                                                                                            colimit F is functorial in F, when C has all colimits of shape J.

                                                                                            Equations
                                                                                            • One or more equations did not get rendered due to their size.
                                                                                            Instances For
                                                                                              @[simp]
                                                                                              theorem CategoryTheory.Limits.colim_map {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] [HasColimitsOfShape J C] {X✝ Y✝ : Functor J C} (α : X✝ Y✝) :
                                                                                              colim.map α = colimMap α
                                                                                              theorem CategoryTheory.Limits.colimMap_eq {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimitsOfShape J C] {G : Functor J C} (α : F G) :
                                                                                              colimMap α = colim.map α
                                                                                              theorem CategoryTheory.Limits.colimit.ι_map {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimitsOfShape J C] {G : Functor J C} (α : F G) (j : J) :
                                                                                              CategoryStruct.comp (ι F j) (colim.map α) = CategoryStruct.comp (α.app j) (ι G j)
                                                                                              theorem CategoryTheory.Limits.colimit.ι_map_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimitsOfShape J C] {G : Functor J C} (α : F G) (j : J) {Z : C} (h : colim.obj G Z) :
                                                                                              @[simp]
                                                                                              theorem CategoryTheory.Limits.colimit.map_desc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimitsOfShape J C] {G : Functor J C} (α : F G) (c : Cocone G) :
                                                                                              @[simp]
                                                                                              theorem CategoryTheory.Limits.colimit.map_desc_assoc {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimitsOfShape J C] {G : Functor J C} (α : F G) (c : Cocone G) {Z : C} (h : c.pt Z) :
                                                                                              theorem CategoryTheory.Limits.colimit.pre_map' {J : Type u₁} [Category.{v₁, u₁} J] {K : Type u₂} [Category.{v₂, u₂} K] {C : Type u} [Category.{v, u} C] [HasColimitsOfShape J C] [HasColimitsOfShape K C] (F : Functor J C) {E₁ E₂ : Functor K J} (α : E₁ E₂) :
                                                                                              pre F E₁ = CategoryStruct.comp (colim.map (whiskerRight α F)) (pre F E₂)
                                                                                              theorem CategoryTheory.Limits.colimit.pre_id {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] [HasColimitsOfShape J C] (F : Functor J C) :
                                                                                              pre F (Functor.id J) = colim.map F.leftUnitor.hom
                                                                                              theorem CategoryTheory.Limits.colimit.map_post {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} [HasColimitsOfShape J C] {G : Functor J C} (α : F G) {D : Type u'} [Category.{v', u'} D] [HasColimitsOfShape J D] (H : Functor C D) :
                                                                                              CategoryStruct.comp (post F H) (H.map (colim.map α)) = CategoryStruct.comp (colim.map (whiskerRight α H)) (post G H)

                                                                                              The isomorphism between morphisms from the cone point of the colimit cocone for F to W and cocones over F with cone point W is natural in F.

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

                                                                                                The colimit functor and constant functor are adjoint to each other

                                                                                                Equations
                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                Instances For
                                                                                                  instance CategoryTheory.Limits.colimMap_epi' {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasColimitsOfShape J C] (α : F G) [Epi α] :
                                                                                                  instance CategoryTheory.Limits.colimMap_epi {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F G : Functor J C} [HasColimit F] [HasColimit G] (α : F G) [∀ (j : J), Epi (α.app j)] :

                                                                                                  We can transport colimits of shape J along an equivalence J ≌ J'.

                                                                                                  hasColimitsOfSizeShrink.{v u} C tries to obtain HasColimitsOfSize.{v u} C from some other HasColimitsOfSize C.

                                                                                                  def CategoryTheory.Limits.IsLimit.op {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} {t : Cone F} (P : IsLimit t) :

                                                                                                  If t : Cone F is a limit cone, then t.op : Cocone F.op is a colimit cocone.

                                                                                                  Equations
                                                                                                  Instances For

                                                                                                    If t : Cocone F is a colimit cocone, then t.op : Cone F.op is a limit cone.

                                                                                                    Equations
                                                                                                    Instances For
                                                                                                      def CategoryTheory.Limits.IsLimit.unop {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} {t : Cone F.op} (P : IsLimit t) :
                                                                                                      IsColimit t.unop

                                                                                                      If t : Cone F.op is a limit cone, then t.unop : Cocone F is a colimit cocone.

                                                                                                      Equations
                                                                                                      Instances For
                                                                                                        def CategoryTheory.Limits.IsColimit.unop {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} {t : Cocone F.op} (P : IsColimit t) :
                                                                                                        IsLimit t.unop

                                                                                                        If t : Cocone F.op is a colimit cocone, then t.unop : Cone F is a limit cone.

                                                                                                        Equations
                                                                                                        Instances For
                                                                                                          def CategoryTheory.Limits.isLimitOfOp {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} {t : Cone F} (P : IsColimit t.op) :

                                                                                                          If t.op : Cocone F.op is a colimit cocone, then t : Cone F is a limit cone.

                                                                                                          Equations
                                                                                                          Instances For

                                                                                                            If t.op : Cone F.op is a limit cone, then t : Cocone F is a colimit cocone.

                                                                                                            Equations
                                                                                                            Instances For
                                                                                                              def CategoryTheory.Limits.isLimitOfUnop {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} {t : Cone F.op} (P : IsColimit t.unop) :

                                                                                                              If t.unop : Cocone F is a colimit cocone, then t : Cone F.op is a limit cone.

                                                                                                              Equations
                                                                                                              Instances For
                                                                                                                def CategoryTheory.Limits.isColimitOfUnop {J : Type u₁} [Category.{v₁, u₁} J] {C : Type u} [Category.{v, u} C] {F : Functor J C} {t : Cocone F.op} (P : IsLimit t.unop) :

                                                                                                                If t.unop : Cone F is a limit cone, then t : Cocone F.op is a colimit cocone.

                                                                                                                Equations
                                                                                                                Instances For