Limits and colimits
We set up the general theory of limits and colimits in a category. In this introduction we only describe the setup for limits; it is repeated, with slightly different names, for colimits.
The three main structures involved are
is_limit c
, forc : cone F
,F : J ⥤ C
, expressing thatc
is a limit cone,limit_cone F
, which consists of a choice of cone forF
and the fact it is a limit cone, andhas_limit F
, asserting the mere existence of some limit cone forF
.
has_limit
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).
Typically there are two different ways one can use the limits library:
- working with particular cones, and terms of type
is_limit
- working solely with
has_limit
.
While has_limit
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 has_limit F
:
limit F : C
, producing some limit object (of course all such are isomorphic)limit.π F j : limit F ⟶ F.obj j
, the morphisms out of the limit,limit.lift F c : c.X ⟶ limit F
, the universal morphism from any otherc : cone F
, etc.
Key to using the has_limit
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 has_limits_of_shape J C
and has_limits 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 is_limit
,
and then a result in terms of has_limit
derived from this.
At this point, however, this is far from uniformly achieved in mathlib ---
often statements are only written in terms of has_limit
.
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
- lift : Π (s : category_theory.limits.cone F), s.X ⟶ t.X
- fac' : (∀ (s : category_theory.limits.cone F) (j : J), c.lift s ≫ t.π.app j = s.π.app j) . "obviously"
- uniq' : (∀ (s : category_theory.limits.cone F) (m : s.X ⟶ t.X), (∀ (j : J), m ≫ t.π.app j = s.π.app j) → m = c.lift s) . "obviously"
A cone t
on F
is a limit cone if each cone on F
admits a unique
cone morphism to t
.
See https://stacks.math.columbia.edu/tag/002E.
Given a natural transformation α : F ⟶ G
, we give a morphism from the cone point
of any cone over F
to the cone point of a limit cone over G
.
Equations
The universal morphism from any other cone to a limit cone.
Alternative constructor for is_limit
,
providing a morphism of cones rather than a morphism between the cone points
and separately the factorisation condition.
Equations
- category_theory.limits.is_limit.mk_cone_morphism lift uniq' = {lift := λ (s : category_theory.limits.cone F), (lift s).hom, fac' := _, uniq' := _}
Limit cones on F
are unique up to isomorphism.
Equations
- P.unique_up_to_iso Q = {hom := Q.lift_cone_morphism s, inv := P.lift_cone_morphism t, hom_inv_id' := _, inv_hom_id' := _}
Any cone morphism between limit cones is an isomorphism.
Equations
- P.hom_is_iso Q f = {inv := P.lift_cone_morphism t, hom_inv_id' := _, inv_hom_id' := _}
Limits of F
are unique up to isomorphism.
Equations
Transport evidence that a cone is a limit cone across an isomorphism of cones.
Equations
- P.of_iso_limit i = category_theory.limits.is_limit.mk_cone_morphism (λ (s : category_theory.limits.cone F), P.lift_cone_morphism s ≫ i.hom) _
Isomorphism of cones preserves whether or not they are limiting cones.
Equations
- category_theory.limits.is_limit.equiv_iso_limit i = {to_fun := λ (h : category_theory.limits.is_limit r), h.of_iso_limit i, inv_fun := λ (h : category_theory.limits.is_limit t), h.of_iso_limit i.symm, left_inv := _, right_inv := _}
If the canonical morphism from a cone point to a limiting cone point is an iso, then the first cone was limiting also.
Equations
Two morphisms into a limit are equal if their compositions with each cone morphism are equal.
Given a right adjoint functor between categories of cones, the image of a limit cone is a limit cone.
Equations
Given two functors which have equivalent categories of cones, we can transport a limiting cone across the equivalence.
Equations
- category_theory.limits.is_limit.of_cone_equiv h = {to_fun := λ (P : category_theory.limits.is_limit (h.functor.obj c)), (category_theory.limits.is_limit.of_right_adjoint h.inverse P).of_iso_limit (h.unit_iso.symm.app c), inv_fun := category_theory.limits.is_limit.of_right_adjoint h.functor c, left_inv := _, right_inv := _}
A cone postcomposed with a natural isomorphism is a limit cone if and only if the original cone is.
A cone postcomposed with the inverse of a natural isomorphism is a limit cone if and only if the original cone is.
The cone points of two limit cones for naturally isomorphic functors are themselves isomorphic.
Equations
- P.cone_points_iso_of_nat_iso Q w = {hom := category_theory.limits.is_limit.map s Q w.hom, inv := category_theory.limits.is_limit.map t P w.inv, hom_inv_id' := _, inv_hom_id' := _}
If s : cone F
is a limit cone, so is s
whiskered by an equivalence e
.
We can prove two cone points (s : cone F).X
and (t.cone F).X
are isomorphic if
- both cones are limit cones
- their indexing categories are equivalent via some
e : J ≌ K
, - the triangle of functors commutes up to a natural isomorphism:
e.functor ⋙ G ≅ F
.
This is the most general form of uniqueness of cone points, allowing relabelling of both the indexing category (up to equivalence) and the functor (up to natural isomorphism).
Equations
- P.cone_points_iso_of_equivalence Q e w = let w' : e.inverse ⋙ F ≅ G := (category_theory.iso_whisker_left e.inverse w).symm ≪≫ e.inv_fun_id_assoc G in {hom := Q.lift ((category_theory.limits.cones.equivalence_of_reindexing e.symm w').functor.obj s), inv := P.lift ((category_theory.limits.cones.equivalence_of_reindexing e w).functor.obj t), hom_inv_id' := _, inv_hom_id' := _}
The universal property of a limit cone: a map W ⟶ X
is the same as
a cone on F
with vertex W
.
The limit of F
represents the functor taking W
to
the set of cones on F
with vertex W
.
Equations
- h.nat_iso = category_theory.nat_iso.of_components (λ (W : Cᵒᵖ), h.hom_iso (opposite.unop W)) _
Another, more explicit, formulation of the universal property of a limit cone.
See also hom_iso
.
If G : C → D is a faithful functor which sends t to a limit cone, then it suffices to check that the induced maps for the image of t can be lifted to maps of C.
If F
and G
are naturally isomorphic, then F.map_cone c
being a limit implies
G.map_cone c
is also a limit.
Equations
- category_theory.limits.is_limit.map_cone_equiv h t = {lift := λ (s : category_theory.limits.cone (K ⋙ G)), category_theory.limits.is_limit.map s t (category_theory.iso_whisker_left K h).inv ≫ h.hom.app c.X, fac' := _, uniq' := _}
A cone is a limit cone exactly if there is a unique cone morphism from any other cone.
Equations
- category_theory.limits.is_limit.iso_unique_cone_morphism = {hom := λ (h : category_theory.limits.is_limit t) (s : category_theory.limits.cone F), {to_inhabited := {default := h.lift_cone_morphism s}, uniq := _}, inv := λ (h : Π (s : category_theory.limits.cone F), unique (s ⟶ t)), {lift := λ (s : category_theory.limits.cone F), (default (s ⟶ t)).hom, fac' := _, uniq' := _}, hom_inv_id' := _, inv_hom_id' := _}
If F.cones
is represented by X
, each morphism f : Y ⟶ X
gives a cone with cone point Y
.
Equations
- category_theory.limits.is_limit.of_nat_iso.cone_of_hom h f = {X := Y, π := h.hom.app (opposite.op Y) f}
If F.cones
is represented by X
, each cone s
gives a morphism s.X ⟶ X
.
Equations
If F.cones
is represented by X
, the cone corresponding to the identity morphism on X
will be a limit cone.
If F.cones
is represented by X
, the cone corresponding to a morphism f : Y ⟶ X
is
the limit cone extended by f
.
If F.cones
is represented by X
, any cone is the extension of the limit cone by the
corresponding morphism.
If F.cones
is representable, then the cone corresponding to the identity morphism on
the representing object is a limit cone.
Equations
- category_theory.limits.is_limit.of_nat_iso h = {lift := λ (s : category_theory.limits.cone F), category_theory.limits.is_limit.of_nat_iso.hom_of_cone h s, fac' := _, uniq' := _}
- desc : Π (s : category_theory.limits.cocone F), t.X ⟶ s.X
- fac' : (∀ (s : category_theory.limits.cocone F) (j : J), t.ι.app j ≫ c.desc s = s.ι.app j) . "obviously"
- uniq' : (∀ (s : category_theory.limits.cocone F) (m : t.X ⟶ s.X), (∀ (j : J), t.ι.app j ≫ m = s.ι.app j) → m = c.desc s) . "obviously"
A cocone t
on F
is a colimit cocone if each cocone on F
admits a unique
cocone morphism from t
.
See https://stacks.math.columbia.edu/tag/002F.
Given a natural transformation α : F ⟶ G
, we give a morphism from the cocone point
of a colimit cocone over F
to the cocone point of any cocone over G
.
Equations
- P.map t α = P.desc ((category_theory.limits.cocones.precompose α).obj t)
The universal morphism from a colimit cocone to any other cocone.
Alternative constructor for is_colimit
,
providing a morphism of cocones rather than a morphism between the cocone points
and separately the factorisation condition.
Equations
- category_theory.limits.is_colimit.mk_cocone_morphism desc uniq' = {desc := λ (s : category_theory.limits.cocone F), (desc s).hom, fac' := _, uniq' := _}
Colimit cocones on F
are unique up to isomorphism.
Equations
- P.unique_up_to_iso Q = {hom := P.desc_cocone_morphism t, inv := Q.desc_cocone_morphism s, hom_inv_id' := _, inv_hom_id' := _}
Any cocone morphism between colimit cocones is an isomorphism.
Equations
- P.hom_is_iso Q f = {inv := Q.desc_cocone_morphism s, hom_inv_id' := _, inv_hom_id' := _}
Colimits of F
are unique up to isomorphism.
Equations
Transport evidence that a cocone is a colimit cocone across an isomorphism of cocones.
Equations
- P.of_iso_colimit i = category_theory.limits.is_colimit.mk_cocone_morphism (λ (s : category_theory.limits.cocone F), i.inv ≫ P.desc_cocone_morphism s) _
Isomorphism of cocones preserves whether or not they are colimiting cocones.
Equations
- category_theory.limits.is_colimit.equiv_iso_colimit i = {to_fun := λ (h : category_theory.limits.is_colimit r), h.of_iso_colimit i, inv_fun := λ (h : category_theory.limits.is_colimit t), h.of_iso_colimit i.symm, left_inv := _, right_inv := _}
If the canonical morphism to a cocone point from a colimiting cocone point is an iso, then the first cocone was colimiting also.
Equations
Two morphisms out of a colimit are equal if their compositions with each cocone morphism are equal.
Given a left adjoint functor between categories of cocones, the image of a colimit cocone is a colimit cocone.
Equations
Given two functors which have equivalent categories of cocones, we can transport a colimiting cocone across the equivalence.
Equations
- category_theory.limits.is_colimit.of_cocone_equiv h = {to_fun := λ (P : category_theory.limits.is_colimit (h.functor.obj c)), (category_theory.limits.is_colimit.of_left_adjoint h.inverse P).of_iso_colimit (h.unit_iso.symm.app c), inv_fun := category_theory.limits.is_colimit.of_left_adjoint h.functor c, left_inv := _, right_inv := _}
A cocone precomposed with a natural isomorphism is a colimit cocone if and only if the original cocone is.
A cocone precomposed with the inverse of a natural isomorphism is a colimit cocone if and only if the original cocone is.
The cocone points of two colimit cocones for naturally isomorphic functors are themselves isomorphic.
Equations
- P.cocone_points_iso_of_nat_iso Q w = {hom := P.map t w.hom, inv := Q.map s w.inv, hom_inv_id' := _, inv_hom_id' := _}
If s : cone F
is a limit cone, so is s
whiskered by an equivalence e
.
We can prove two cocone points (s : cocone F).X
and (t.cocone F).X
are isomorphic if
- both cocones are colimit ccoones
- their indexing categories are equivalent via some
e : J ≌ K
, - the triangle of functors commutes up to a natural isomorphism:
e.functor ⋙ G ≅ F
.
This is the most general form of uniqueness of cocone points, allowing relabelling of both the indexing category (up to equivalence) and the functor (up to natural isomorphism).
Equations
- P.cocone_points_iso_of_equivalence Q e w = let w' : e.inverse ⋙ F ≅ G := (category_theory.iso_whisker_left e.inverse w).symm ≪≫ e.inv_fun_id_assoc G in {hom := P.desc ((category_theory.limits.cocones.equivalence_of_reindexing e w).functor.obj t), inv := Q.desc ((category_theory.limits.cocones.equivalence_of_reindexing e.symm w').functor.obj s), hom_inv_id' := _, inv_hom_id' := _}
The universal property of a colimit cocone: a map X ⟶ W
is the same as
a cocone on F
with vertex W
.
The colimit of F
represents the functor taking W
to
the set of cocones on F
with vertex W
.
Equations
Another, more explicit, formulation of the universal property of a colimit cocone.
See also hom_iso
.
If G : C → D is a faithful functor which sends t to a colimit cocone, then it suffices to check that the induced maps for the image of t can be lifted to maps of C.
A cocone is a colimit cocone exactly if there is a unique cocone morphism from any other cocone.
Equations
- category_theory.limits.is_colimit.iso_unique_cocone_morphism = {hom := λ (h : category_theory.limits.is_colimit t) (s : category_theory.limits.cocone F), {to_inhabited := {default := h.desc_cocone_morphism s}, uniq := _}, inv := λ (h : Π (s : category_theory.limits.cocone F), unique (t ⟶ s)), {desc := λ (s : category_theory.limits.cocone F), (default (t ⟶ s)).hom, fac' := _, uniq' := _}, hom_inv_id' := _, inv_hom_id' := _}
If F.cocones
is corepresented by X
, each morphism f : X ⟶ Y
gives a cocone with cone point Y
.
If F.cocones
is corepresented by X
, each cocone s
gives a morphism X ⟶ s.X
.
If F.cocones
is corepresented by X
, the cocone corresponding to the identity morphism on X
will be a colimit cocone.
If F.cocones
is corepresented by X
, the cocone corresponding to a morphism f : Y ⟶ X
is
the colimit cocone extended by f
.
If F.cocones
is corepresented by X
, any cocone is the extension of the colimit cocone by the
corresponding morphism.
If F.cocones
is corepresentable, then the cocone corresponding to the identity morphism on
the representing object is a colimit cocone.
Equations
- category_theory.limits.is_colimit.of_nat_iso h = {desc := λ (s : category_theory.limits.cocone F), category_theory.limits.is_colimit.of_nat_iso.hom_of_cocone h s, fac' := _, uniq' := _}
- cone : category_theory.limits.cone F
- is_limit : category_theory.limits.is_limit c.cone
limit_cone F
contains a cone over F
together with the information that it is a limit.
- exists_limit : nonempty (category_theory.limits.limit_cone F)
has_limit F
represents the mere existence of a limit for F
.
Instances
- category_theory.limits.has_limit_of_has_limits_of_shape
- category_theory.limits.has_product_of_has_biproduct
- category_theory.limits.has_limit_equivalence_comp
- category_theory.limits.has_binary_biproduct.has_limit_pair
- category_theory.adjunction.has_limit_comp_equivalence
- category_theory.under.has_limit
- category_theory.comp_comparison_forget_has_limit
- category_theory.comp_comparison_has_limit
Use the axiom of choice to extract explicit limit_cone F
from has_limit F
.
- has_limit : ∀ (F : J ⥤ C), category_theory.limits.has_limit F
C
has limits of shape J
if there exists a limit for every functor F : J ⥤ C
.
Instances
- category_theory.limits.has_limits_of_shape_of_has_limits
- category_theory.limits.has_limits_of_shape_of_has_finite_limits
- category_theory.limits.has_limits_of_shape_wide_pullback_shape
- category_theory.limits.has_limits_of_shape_discrete
- category_theory.limits.functor_category_has_limits_of_shape
- category_theory.under.has_limits_of_shape
- category_theory.over.has_connected_limits
- has_limits_of_shape : ∀ (J : Type ?) [𝒥 : category_theory.small_category J], category_theory.limits.has_limits_of_shape J C
C
has all (small) limits if it has limits of every shape.
Instances
- category_theory.limits.complete_lattice.has_limits_of_complete_lattice
- category_theory.limits.types.category_theory.limits.has_limits
- Mon.has_limits
- AddMon.has_limits
- CommMon.has_limits
- AddCommMon.has_limits
- Group.has_limits
- AddGroup.has_limits
- CommGroup.has_limits
- AddCommGroup.has_limits
- Module.has_limits
- SemiRing.has_limits
- CommSemiRing.has_limits
- Ring.has_limits
- CommRing.has_limits
- Algebra.has_limits
- category_theory.limits.category_theory.limits.has_limits
- Top.Top_has_limits
- category_theory.limits.functor_category_has_limits
- Top.category_theory.limits.has_limits
- category_theory.under.has_limits
- category_theory.over.has_limits
- Mon_.has_limits
An arbitrary choice of limit cone for a functor.
Equations
An arbitrary choice of limit object of a functor.
Equations
The projection from the limit object to a value of the functor.
Equations
Evidence that the arbitrary choice of cone provied by limit.cone F
is a limit cone.
The morphism from the cone point of any other cone to the limit object.
Equations
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
.
The cone morphism from any cone to the arbitrary choice of limit cone.
Given any other limit cone for F
, the chosen limit F
is isomorphic to the cone point.
The isomorphism (in Type
) between
morphisms from a specified object W
to the limit object,
and cones with cone point W
.
Equations
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
If a functor F
has a limit, so does any naturally isomorphic functor.
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.
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.
The canonical morphism from the limit of F
to the limit of E ⋙ F
.
-
If we have particular limit cones available for E ⋙ F
and for F
,
we obtain a formula for limit.pre F E
.
The canonical morphism from G
applied to the limit of F
to the limit of F ⋙ G
.
Equations
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
- category_theory.limits.lim = {obj := λ (F : J ⥤ C), category_theory.limits.limit F, map := λ (F G : J ⥤ C) (α : F ⟶ G), category_theory.limits.lim_map α, map_id' := _, map_comp' := _}
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
- category_theory.limits.lim_yoneda = category_theory.nat_iso.of_components (λ (F : J ⥤ C), category_theory.nat_iso.of_components (λ (W : Cᵒᵖ), category_theory.limits.limit.hom_iso F (opposite.unop W)) _) category_theory.limits.lim_yoneda._proof_2
We can transport limits of shape J
along an equivalence J ≌ J'
.
- cocone : category_theory.limits.cocone F
- is_colimit : category_theory.limits.is_colimit c.cocone
colimit_cocone F
contains a cocone over F
together with the information that it is a
colimit.
- exists_colimit : nonempty (category_theory.limits.colimit_cocone F)
has_colimit F
represents the mere existence of a colimit for F
.
Instances
- category_theory.limits.has_colimit_of_has_colimits_of_shape
- category_theory.limits.has_coproduct_of_has_biproduct
- category_theory.cofinal.comp_has_colimit
- category_theory.limits.has_colimit_equivalence_comp
- category_theory.limits.has_binary_biproduct.has_colimit_pair
- category_theory.adjunction.has_colimit_comp_equivalence
- category_theory.over.has_colimit
Use the axiom of choice to extract explicit colimit_cocone F
from has_colimit F
.
- has_colimit : ∀ (F : J ⥤ C), category_theory.limits.has_colimit F
C
has colimits of shape J
if there exists a colimit for every functor F : J ⥤ C
.
Instances
- category_theory.limits.has_colimits_of_shape_of_has_colimits
- category_theory.limits.has_colimits_of_shape_of_has_finite_colimits
- category_theory.limits.has_colimits_of_shape_wide_pushout_shape
- category_theory.limits.has_colimits_of_shape_discrete
- category_theory.limits.functor_category_has_colimits_of_shape
- category_theory.over.has_colimits_of_shape
- has_colimits_of_shape : ∀ (J : Type ?) [𝒥 : category_theory.small_category J], category_theory.limits.has_colimits_of_shape J C
C
has all (small) colimits if it has colimits of every shape.
Instances
- category_theory.limits.complete_lattice.has_colimits_of_complete_lattice
- category_theory.limits.types.category_theory.limits.has_colimits
- CommRing.colimits.has_colimits_CommRing
- AddCommGroup.colimits.has_colimits_AddCommGroup
- Mon.colimits.has_colimits_Mon
- category_theory.limits.category_theory.limits.has_colimits
- Top.Top_has_colimits
- category_theory.limits.functor_category_has_colimits
- Top.category_theory.limits.has_colimits
- algebraic_geometry.PresheafedSpace.category_theory.limits.has_colimits
- category_theory.over.has_colimits
An arbitrary choice of colimit cocone of a functor.
An arbitrary choice of colimit object of a functor.
Equations
The coprojection from a value of the functor to the colimit object.
Equations
Evidence that the arbitrary choice of cocone is a colimit cocone.
The morphism from the colimit object to the cone point of any other cocone.
Equations
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
)
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
.
The cocone morphism from the arbitrary choice of colimit cocone to any cocone.
Given any other colimit cocone for F
, the chosen colimit F
is isomorphic to the cocone point.
The isomorphism (in Type
) between
morphisms from the colimit object to a specified object W
,
and cocones with cone point W
.
Equations
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
If F
has a colimit, so does any naturally isomorphic functor.
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.
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.
The canonical morphism from the colimit of E ⋙ F
to the colimit of F
.
-
If we have particular colimit cocones available for E ⋙ F
and for F
,
we obtain a formula for colimit.pre F E
.
The canonical morphism from G
applied to the colimit of F ⋙ G
to G
applied to the colimit of F
.
Equations
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
- category_theory.limits.colim = {obj := λ (F : J ⥤ C), category_theory.limits.colimit F, map := λ (F G : J ⥤ C) (α : F ⟶ G), category_theory.limits.colim_map α, map_id' := _, map_comp' := _}
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
- category_theory.limits.colim_coyoneda = category_theory.nat_iso.of_components (λ (F : (J ⥤ C)ᵒᵖ), category_theory.nat_iso.of_components (category_theory.limits.colimit.hom_iso (opposite.unop F)) _) category_theory.limits.colim_coyoneda._proof_3
We can transport colimits of shape J
along an equivalence J ≌ J'
.