mathlib documentation

category_theory.limits.limits

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

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:

  1. working with particular cones, and terms of type is_limit
  2. 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:

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

@[nolint]
structure category_theory.limits.is_limit {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] {F : J C} (t : category_theory.limits.cone F) :
Type (max u v)

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
@[simp]
theorem category_theory.limits.is_limit.map_π_assoc {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] {F G : J C} (c : category_theory.limits.cone F) {d : category_theory.limits.cone G} (hd : category_theory.limits.is_limit d) (α : F G) (j : J) {X' : C} (f' : G.obj j X') :

The universal morphism from any other cone to a limit cone.

Equations

Alternative constructor for is_limit, providing a morphism of cones rather than a morphism between the cone points and separately the factorisation condition.

Equations

Any cone morphism between limit cones is an isomorphism.

Equations

Transport evidence that a cone is a limit cone across an isomorphism of cones.

Equations

If the canonical morphism from a cone point to a limiting cone point is an iso, then the first cone was limiting also.

Equations
theorem category_theory.limits.is_limit.hom_lift {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] {F : J C} {t : category_theory.limits.cone F} (h : category_theory.limits.is_limit t) {W : C} (m : W t.X) :
m = h.lift {X := W, π := {app := λ (b : J), m t.π.app b, naturality' := _}}

theorem category_theory.limits.is_limit.hom_ext {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] {F : J C} {t : category_theory.limits.cone F} (h : category_theory.limits.is_limit t) {W : C} {f f' : W t.X} (w : ∀ (j : J), f t.π.app j = f' t.π.app j) :
f = f'

Two morphisms into a limit are equal if their compositions with each cone morphism are equal.

The cone points of two limit cones for naturally isomorphic functors are themselves isomorphic.

Equations

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

The universal property of a limit cone: a map W ⟶ X is the same as a cone on F with vertex W.

Equations
@[simp]

The limit of F represents the functor taking W to the set of cones on F with vertex W.

Equations
def category_theory.limits.is_limit.hom_iso' {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] {F : J C} {t : category_theory.limits.cone F} (h : category_theory.limits.is_limit t) (W : C) :
(W t.X) {p // ∀ {j j' : J} (f : j j'), p j F.map f = p j'}

Another, more explicit, formulation of the universal property of a limit cone. See also hom_iso.

Equations

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.

Equations

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

A cone is a limit cone exactly if there is a unique cone morphism from any other cone.

Equations

If F.cones is represented by X, each morphism f : Y ⟶ X gives a cone with cone point Y.

Equations

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.

Equations

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
@[nolint]
structure category_theory.limits.is_colimit {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] {F : J C} (t : category_theory.limits.cocone F) :
Type (max u v)

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
@[simp]
theorem category_theory.limits.is_colimit.ι_map_assoc {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] {F G : J C} {c : category_theory.limits.cocone F} (hc : category_theory.limits.is_colimit c) (d : category_theory.limits.cocone G) (α : F G) (j : J) {X' : C} (f' : d.X X') :
c.ι.app j hc.map d α f' = α.app j d.ι.app j f'

@[simp]

The universal morphism from a colimit cocone to any other cocone.

Equations

Alternative constructor for is_colimit, providing a morphism of cocones rather than a morphism between the cocone points and separately the factorisation condition.

Equations

Any cocone morphism between colimit cocones is an isomorphism.

Equations

Transport evidence that a cocone is a colimit cocone across an isomorphism of cocones.

Equations

If the canonical morphism to a cocone point from a colimiting cocone point is an iso, then the first cocone was colimiting also.

Equations
theorem category_theory.limits.is_colimit.hom_desc {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] {F : J C} {t : category_theory.limits.cocone F} (h : category_theory.limits.is_colimit t) {W : C} (m : t.X W) :
m = h.desc {X := W, ι := {app := λ (b : J), t.ι.app b m, naturality' := _}}

theorem category_theory.limits.is_colimit.hom_ext {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] {F : J C} {t : category_theory.limits.cocone F} (h : category_theory.limits.is_colimit t) {W : C} {f f' : t.X W} (w : ∀ (j : J), t.ι.app j f = t.ι.app j f') :
f = f'

Two morphisms out of a colimit are equal if their compositions with each cocone morphism are equal.

The cocone points of two colimit cocones for naturally isomorphic functors are themselves isomorphic.

Equations

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

The universal property of a colimit cocone: a map X ⟶ W is the same as a cocone on F with vertex W.

Equations

The colimit of F represents the functor taking W to the set of cocones on F with vertex W.

Equations
def category_theory.limits.is_colimit.hom_iso' {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] {F : J C} {t : category_theory.limits.cocone F} (h : category_theory.limits.is_colimit t) (W : C) :
(t.X W) {p // ∀ {j j' : J} (f : j j'), F.map f p j' = p j}

Another, more explicit, formulation of the universal property of a colimit cocone. See also hom_iso.

Equations

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.

Equations

If F.cocones is corepresented by X, each morphism f : X ⟶ Y gives a cocone with cone point Y.

Equations

If F.cocones is corepresented by X, each cocone s gives a morphism X ⟶ s.X.

Equations

If F.cocones is corepresented by X, the cocone corresponding to the identity morphism on X will be a colimit cocone.

Equations

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
@[nolint]
structure category_theory.limits.limit_cone {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] (F : J C) :
Type (max u v)

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

An arbitrary choice of limit object of a functor.

Equations

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

Equations
@[simp]
theorem category_theory.limits.limit.w_assoc {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] (F : J C) [category_theory.limits.has_limit F] {j j' : J} (f : j j') {X' : C} (f' : F.obj j' X') :

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

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

Equations
def category_theory.limits.limit.hom_iso' {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] (F : J C) [category_theory.limits.has_limit F] (W : C) :
(W category_theory.limits.limit F) {p // ∀ {j j' : J} (f : j j'), 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

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 : K ⥤ C are isomorphic, if there is an equivalence e : J ≌ K making the triangle commute up to natural isomorphism.

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

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

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

@[nolint]
structure category_theory.limits.colimit_cocone {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] (F : J C) :
Type (max u v)

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

An arbitrary choice of colimit object of a functor.

Equations

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

Equations
@[simp]

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.

Equations

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

Equations
def category_theory.limits.colimit.hom_iso' {J : Type v} [category_theory.small_category J] {C : Type u} [category_theory.category C] (F : J C) [category_theory.limits.has_colimit F] (W : C) :
(category_theory.limits.colimit F W) {p // ∀ {j j' : J} (f : j j'), 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

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.

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

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

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