F.cones
is the functor assigning to an object X
the type of
natural transformations from the constant functor with value X
to F
.
An object representing this functor is a limit of F
.
Equations
F.cocones
is the functor assigning to an object X
the type of
natural transformations from F
to the constant functor with value X
.
An object corepresenting this functor is a colimit of F
.
Equations
Functorially associated to each functor J ⥤ C
, we have the C
-presheaf consisting of
cones with a given cone point.
Equations
- category_theory.cones J C = {obj := category_theory.functor.cones _inst_2, map := λ (F G : J ⥤ C) (f : F ⟶ G), category_theory.whisker_left (category_theory.functor.const J).op (category_theory.yoneda.map f), map_id' := _, map_comp' := _}
Contravariantly associated to each functor J ⥤ C
, we have the C
-copresheaf consisting of
cocones with a given cocone point.
Equations
- category_theory.cocones J C = {obj := λ (F : (J ⥤ C)ᵒᵖ), (opposite.unop F).cocones, map := λ (F G : (J ⥤ C)ᵒᵖ) (f : F ⟶ G), category_theory.whisker_left (category_theory.functor.const J) (category_theory.coyoneda.map f), map_id' := _, map_comp' := _}
- X : C
- π : (category_theory.functor.const J).obj c.X ⟶ F
A c : cone F
is:
- an object
c.X
and - a natural transformation
c.π : c.X ⟶ F
from the constantc.X
functor toF
.
cone F
is equivalent, via cone.equiv
below, to Σ X, F.cones.obj X
.
- X : C
- ι : F ⟶ (category_theory.functor.const J).obj c.X
A c : cocone F
is
- an object
c.X
and - a natural transformation
c.ι : F ⟶ c.X
fromF
to the constantc.X
functor.
cocone F
is equivalent, via cone.equiv
below, to Σ X, F.cocones.obj X
.
The isomorphism between a cone on F
and an element of the functor F.cones
.
Equations
- category_theory.limits.cone.equiv F = {hom := λ (c : category_theory.limits.cone F), ⟨opposite.op c.X, c.π⟩, inv := λ (c : Σ (X : Cᵒᵖ), F.cones.obj X), {X := opposite.unop c.fst, π := c.snd}, hom_inv_id' := _, inv_hom_id' := _}
Equations
- c.extensions = {app := λ (X : Cᵒᵖ) (f : (category_theory.yoneda.obj c.X).obj X), (category_theory.functor.const J).map f ≫ c.π, naturality' := _}
A map to the vertex of a cone induces a cone by composition.
Equations
- c.extend f = {X := X, π := c.extensions.app (opposite.op X) f}
Whisker a cone by precomposition of a functor.
Equations
- category_theory.limits.cone.whisker E c = {X := c.X, π := category_theory.whisker_left E c.π}
The isomorphism between a cocone on F
and an element of the functor F.cocones
.
Equations
- category_theory.limits.cocone.equiv F = {hom := λ (c : category_theory.limits.cocone F), ⟨c.X, c.ι⟩, inv := λ (c : Σ (X : C), F.cocones.obj X), {X := c.fst, ι := c.snd}, hom_inv_id' := _, inv_hom_id' := _}
Equations
- c.extensions = {app := λ (X : C) (f : (category_theory.coyoneda.obj (opposite.op c.X)).obj X), c.ι ≫ (category_theory.functor.const J).map f, naturality' := _}
A map from the vertex of a cocone induces a cocone by composition.
Whisker a cocone by precomposition of a functor. See whiskering
for a functorial
version.
Equations
- category_theory.limits.cocone.whisker E c = {X := c.X, ι := category_theory.whisker_left E c.ι}
A cone morphism between two cones for the same diagram is a morphism of the cone points which commutes with the cone legs.
The category of cones on a given diagram.
Equations
- category_theory.limits.cone.category = {to_category_struct := {to_has_hom := {hom := λ (A B : category_theory.limits.cone F), category_theory.limits.cone_morphism A B}, id := λ (B : category_theory.limits.cone F), {hom := 𝟙 B.X, w' := _}, comp := λ (X Y Z : category_theory.limits.cone F) (f : X ⟶ Y) (g : Y ⟶ Z), {hom := f.hom ≫ g.hom, w' := _}}, id_comp' := _, comp_id' := _, assoc' := _}
To give an isomorphism between cones, it suffices to give an isomorphism between their vertices which commutes with the cone maps.
Equations
- category_theory.limits.cones.ext φ w = {hom := {hom := φ.hom, w' := _}, inv := {hom := φ.inv, w' := _}, hom_inv_id' := _, inv_hom_id' := _}
Given a cone morphism whose object part is an isomorphism, produce an isomorphism of cones.
Equations
- category_theory.limits.cones.cone_iso_of_hom_iso f = {inv := {hom := category_theory.is_iso.inv f.hom i, w' := _}, hom_inv_id' := _, inv_hom_id' := _}
Functorially postcompose a cone for F
by a natural transformation F ⟶ G
to give a cone for G
.
Postcomposing a cone by the composite natural transformation α ≫ β
is the same as
postcomposing by α
and then by β
.
Equations
Postcomposing by the identity does not change the cone up to isomorphism.
Equations
- category_theory.limits.cones.postcompose_id = category_theory.nat_iso.of_components (λ (s : category_theory.limits.cone F), category_theory.limits.cones.ext (category_theory.iso.refl ((category_theory.limits.cones.postcompose (𝟙 F)).obj s).X) _) category_theory.limits.cones.postcompose_id._proof_2
If F
and G
are naturally isomorphic functors, then they have equivalent categories of
cones.
Equations
- category_theory.limits.cones.postcompose_equivalence α = {functor := category_theory.limits.cones.postcompose α.hom, inverse := category_theory.limits.cones.postcompose α.inv, unit_iso := category_theory.nat_iso.of_components (λ (s : category_theory.limits.cone F), category_theory.limits.cones.ext (category_theory.iso.refl ((𝟭 (category_theory.limits.cone F)).obj s).X) _) _, counit_iso := category_theory.nat_iso.of_components (λ (s : category_theory.limits.cone G), category_theory.limits.cones.ext (category_theory.iso.refl ((category_theory.limits.cones.postcompose α.inv ⋙ category_theory.limits.cones.postcompose α.hom).obj s).X) _) _, functor_unit_iso_comp' := _}
Whiskering on the left by E : K ⥤ J
gives a functor from cone F
to cone (E ⋙ F)
.
Equations
- category_theory.limits.cones.whiskering E = {obj := λ (c : category_theory.limits.cone F), category_theory.limits.cone.whisker E c, map := λ (c c' : category_theory.limits.cone F) (f : c ⟶ c'), {hom := f.hom, w' := _}, map_id' := _, map_comp' := _}
Whiskering by an equivalence gives an equivalence between categories of cones.
Equations
- category_theory.limits.cones.whiskering_equivalence e = {functor := category_theory.limits.cones.whiskering e.functor, inverse := category_theory.limits.cones.whiskering e.inverse ⋙ category_theory.limits.cones.postcompose ((e.inverse.associator e.functor F).inv ≫ category_theory.whisker_right e.counit_iso.hom F ≫ F.left_unitor.hom), unit_iso := category_theory.nat_iso.of_components (λ (s : category_theory.limits.cone F), category_theory.limits.cones.ext (category_theory.iso.refl ((𝟭 (category_theory.limits.cone F)).obj s).X) _) _, counit_iso := category_theory.nat_iso.of_components (λ (s : category_theory.limits.cone (e.functor ⋙ F)), category_theory.limits.cones.ext (category_theory.iso.refl (((category_theory.limits.cones.whiskering e.inverse ⋙ category_theory.limits.cones.postcompose ((e.inverse.associator e.functor F).inv ≫ category_theory.whisker_right e.counit_iso.hom F ≫ F.left_unitor.hom)) ⋙ category_theory.limits.cones.whiskering e.functor).obj s).X) _) _, functor_unit_iso_comp' := _}
The categories of cones over F
and G
are equivalent if F
and G
are naturally isomorphic
(possibly after changing the indexing category by an equivalence).
Forget the cone structure and obtain just the cone point.
Equations
- category_theory.limits.cones.forget F = {obj := λ (t : category_theory.limits.cone F), t.X, map := λ (s t : category_theory.limits.cone F) (f : s ⟶ t), f.hom, map_id' := _, map_comp' := _}
Equations
Equations
- category_theory.limits.cones.functoriality_full F G = {preimage := λ (X Y : category_theory.limits.cone F) (t : (category_theory.limits.cones.functoriality F G).obj X ⟶ (category_theory.limits.cones.functoriality F G).obj Y), {hom := G.preimage t.hom, w' := _}, witness' := _}
Equations
- _ = _
If F
reflects isomorphisms, then cones.functoriality F
reflects isomorphisms
as well.
Equations
- category_theory.limits.cones.reflects_cone_isomorphism F K = {reflects := λ (A B : category_theory.limits.cone K) (f : A ⟶ B) (_inst_3_1 : category_theory.is_iso ((category_theory.limits.cones.functoriality K F).map f)), category_theory.limits.cones.cone_iso_of_hom_iso f}
A cocone morphism between two cocones for the same diagram is a morphism of the cocone points which commutes with the cocone legs.
Equations
- category_theory.limits.cocone.category = {to_category_struct := {to_has_hom := {hom := λ (A B : category_theory.limits.cocone F), category_theory.limits.cocone_morphism A B}, id := λ (B : category_theory.limits.cocone F), {hom := 𝟙 B.X, w' := _}, comp := λ (_x _x_1 _x_2 : category_theory.limits.cocone F) (f : _x ⟶ _x_1) (g : _x_1 ⟶ _x_2), {hom := f.hom ≫ g.hom, w' := _}}, id_comp' := _, comp_id' := _, assoc' := _}
To give an isomorphism between cocones, it suffices to give an isomorphism between their vertices which commutes with the cocone maps.
Equations
- category_theory.limits.cocones.ext φ w = {hom := {hom := φ.hom, w' := w}, inv := {hom := φ.inv, w' := _}, hom_inv_id' := _, inv_hom_id' := _}
Given a cocone morphism whose object part is an isomorphism, produce an isomorphism of cocones.
Equations
- category_theory.limits.cocones.cocone_iso_of_hom_iso f = {inv := {hom := category_theory.is_iso.inv f.hom i, w' := _}, hom_inv_id' := _, inv_hom_id' := _}
Functorially precompose a cocone for F
by a natural transformation G ⟶ F
to give a cocone for G
.
Precomposing a cocone by the composite natural transformation α ≫ β
is the same as
precomposing by β
and then by α
.
Precomposing by the identity does not change the cocone up to isomorphism.
Equations
- category_theory.limits.cocones.precompose_id = category_theory.nat_iso.of_components (λ (s : category_theory.limits.cocone F), category_theory.limits.cocones.ext (category_theory.iso.refl ((category_theory.limits.cocones.precompose (𝟙 F)).obj s).X) _) category_theory.limits.cocones.precompose_id._proof_2
If F
and G
are naturally isomorphic functors, then they have equivalent categories of
cocones.
Equations
- category_theory.limits.cocones.precompose_equivalence α = {functor := category_theory.limits.cocones.precompose α.hom, inverse := category_theory.limits.cocones.precompose α.inv, unit_iso := category_theory.nat_iso.of_components (λ (s : category_theory.limits.cocone F), category_theory.limits.cocones.ext (category_theory.iso.refl ((𝟭 (category_theory.limits.cocone F)).obj s).X) _) _, counit_iso := category_theory.nat_iso.of_components (λ (s : category_theory.limits.cocone G), category_theory.limits.cocones.ext (category_theory.iso.refl ((category_theory.limits.cocones.precompose α.inv ⋙ category_theory.limits.cocones.precompose α.hom).obj s).X) _) _, functor_unit_iso_comp' := _}
Whiskering on the left by E : K ⥤ J
gives a functor from cocone F
to cocone (E ⋙ F)
.
Equations
- category_theory.limits.cocones.whiskering E = {obj := λ (c : category_theory.limits.cocone F), category_theory.limits.cocone.whisker E c, map := λ (c c' : category_theory.limits.cocone F) (f : c ⟶ c'), {hom := f.hom, w' := _}, map_id' := _, map_comp' := _}
Whiskering by an equivalence gives an equivalence between categories of cones.
Equations
- category_theory.limits.cocones.whiskering_equivalence e = {functor := category_theory.limits.cocones.whiskering e.functor, inverse := category_theory.limits.cocones.whiskering e.inverse ⋙ category_theory.limits.cocones.precompose (F.left_unitor.inv ≫ category_theory.whisker_right e.counit_iso.inv F ≫ (e.inverse.associator e.functor F).inv), unit_iso := category_theory.nat_iso.of_components (λ (s : category_theory.limits.cocone F), category_theory.limits.cocones.ext (category_theory.iso.refl ((𝟭 (category_theory.limits.cocone F)).obj s).X) _) _, counit_iso := category_theory.nat_iso.of_components (λ (s : category_theory.limits.cocone (e.functor ⋙ F)), category_theory.limits.cocones.ext (category_theory.iso.refl (((category_theory.limits.cocones.whiskering e.inverse ⋙ category_theory.limits.cocones.precompose (F.left_unitor.inv ≫ category_theory.whisker_right e.counit_iso.inv F ≫ (e.inverse.associator e.functor F).inv)) ⋙ category_theory.limits.cocones.whiskering e.functor).obj s).X) _) _, functor_unit_iso_comp' := _}
The categories of cocones over F
and G
are equivalent if F
and G
are naturally isomorphic
(possibly after changing the indexing category by an equivalence).
Forget the cocone structure and obtain just the cocone point.
Equations
- category_theory.limits.cocones.forget F = {obj := λ (t : category_theory.limits.cocone F), t.X, map := λ (s t : category_theory.limits.cocone F) (f : s ⟶ t), f.hom, map_id' := _, map_comp' := _}
Equations
- category_theory.limits.cocones.functoriality F G = {obj := λ (A : category_theory.limits.cocone F), {X := G.obj A.X, ι := {app := λ (j : J), G.map (A.ι.app j), naturality' := _}}, map := λ (_x _x_1 : category_theory.limits.cocone F) (f : _x ⟶ _x_1), {hom := G.map f.hom, w' := _}, map_id' := _, map_comp' := _}
Equations
- category_theory.limits.cocones.functoriality_full F G = {preimage := λ (X Y : category_theory.limits.cocone F) (t : (category_theory.limits.cocones.functoriality F G).obj X ⟶ (category_theory.limits.cocones.functoriality F G).obj Y), {hom := G.preimage t.hom, w' := _}, witness' := _}
Equations
- _ = _
If F
reflects isomorphisms, then cocones.functoriality F
reflects isomorphisms
as well.
Equations
- category_theory.limits.cocones.reflects_cocone_isomorphism F K = {reflects := λ (A B : category_theory.limits.cocone K) (f : A ⟶ B) (_inst_3_1 : category_theory.is_iso ((category_theory.limits.cocones.functoriality K F).map f)), category_theory.limits.cocones.cocone_iso_of_hom_iso f}
The image of a cone in C under a functor G : C ⥤ D is a cone in D.
Equations
- H.map_cone c = (category_theory.limits.cones.functoriality F H).obj c
The image of a cocone in C under a functor G : C ⥤ D is a cocone in D.
Equations
- H.map_cocone c = (category_theory.limits.cocones.functoriality F H).obj c
If H
is an equivalence, we invert H.map_cone
and get an original cone for F
from a cone
for F ⋙ H
.
Equations
- H.map_cone_inv c = let t : category_theory.limits.cone ((F ⋙ H) ⋙ H.inv) := H.inv.map_cone c, α : (F ⋙ H) ⋙ H.inv ⟶ F := category_theory.whisker_left F H.fun_inv_id.hom ≫ F.right_unitor.hom in {X := t.X, π := ((category_theory.cones J C).map α).app (opposite.op t.X) t.π}
Given a cone morphism c ⟶ c'
, construct a cone morphism on the mapped cones functorially.
Equations
Given a cocone morphism c ⟶ c'
, construct a cocone morphism on the mapped cocones functorially.
Equations
map_cone
is the left inverse to map_cone_inv
.
Equations
- H.map_cone_map_cone_inv c = category_theory.limits.cones.ext (H.inv_fun_id.app c.X) _
map_cone
is the right inverse to map_cone_inv
.
Equations
- H.map_cone_inv_map_cone c = category_theory.limits.cones.ext (H.fun_inv_id.app c.X) _
Change a cocone on F.left_op : Jᵒᵖ ⥤ C
to a cocone on F : J ⥤ Cᵒᵖ
.
Equations
Change a cone on F : J ⥤ Cᵒᵖ
to a cocone on F.left_op : Jᵒᵖ ⥤ C
.
Equations
Change a cone on F.left_op : Jᵒᵖ ⥤ C
to a cocone on F : J ⥤ Cᵒᵖ
.
Equations
Change a cocone on F : J ⥤ Cᵒᵖ
to a cone on F.left_op : Jᵒᵖ ⥤ C
.