Documentation

Mathlib.CategoryTheory.Limits.Shapes.Grothendieck

(Co)limits on the (strict) Grothendieck Construction #

theorem CategoryTheory.Limits.hasColimit_ι_comp {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] (G : Functor (Grothendieck F) H) [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] (X : C) :
def CategoryTheory.Limits.fiberwiseColimit {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] (G : Functor (Grothendieck F) H) [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] :

A functor taking a colimit on each fiber of a functor G : Grothendieck F ⥤ H.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    @[simp]
    theorem CategoryTheory.Limits.fiberwiseColimit_map {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] (G : Functor (Grothendieck F) H) [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] {X Y : C} (f : X Y) :
    @[simp]
    theorem CategoryTheory.Limits.fiberwiseColimit_obj {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] (G : Functor (Grothendieck F) H) [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] (X : C) :
    (fiberwiseColimit G).obj X = colimit ((Grothendieck.ι F X).comp G)
    def CategoryTheory.Limits.fiberwiseColim {C : Type u₁} [Category.{v₁, u₁} C] (F : Functor C Cat) (H : Type u₂) [Category.{v₂, u₂} H] [∀ (c : C), HasColimitsOfShape (↑(F.obj c)) H] :

    Similar to colimit and colim, taking fiberwise colimits is a functor (Grothendieck F ⥤ H) ⥤ (C ⥤ H) between functor categories.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem CategoryTheory.Limits.fiberwiseColim_map_app {C : Type u₁} [Category.{v₁, u₁} C] (F : Functor C Cat) (H : Type u₂) [Category.{v₂, u₂} H] [∀ (c : C), HasColimitsOfShape (↑(F.obj c)) H] {X✝ Y✝ : Functor (Grothendieck F) H} (α : X✝ Y✝) (c : C) :
      ((fiberwiseColim F H).map α).app c = colim.map (whiskerLeft (Grothendieck.ι F c) α)
      @[simp]
      theorem CategoryTheory.Limits.fiberwiseColim_obj {C : Type u₁} [Category.{v₁, u₁} C] (F : Functor C Cat) (H : Type u₂) [Category.{v₂, u₂} H] [∀ (c : C), HasColimitsOfShape (↑(F.obj c)) H] (G : Functor (Grothendieck F) H) :

      Every functor G : Grothendieck F ⥤ H induces a natural transformation from G to the composition of the forgetful functor on Grothendieck F with the fiberwise colimit on G.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        theorem CategoryTheory.Limits.natTransIntoForgetCompFiberwiseColimit_app {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] (G : Functor (Grothendieck F) H) [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] (X : Grothendieck F) :
        def CategoryTheory.Limits.coconeFiberwiseColimitOfCocone {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] {G : Functor (Grothendieck F) H} [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] (c : Cocone G) :

        A cocone on a functor G : Grothendieck F ⥤ H induces a cocone on the fiberwise colimit on G.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          @[simp]
          theorem CategoryTheory.Limits.coconeFiberwiseColimitOfCocone_ι_app {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] {G : Functor (Grothendieck F) H} [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] (c : Cocone G) (X : C) :
          @[simp]
          theorem CategoryTheory.Limits.coconeFiberwiseColimitOfCocone_pt {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] {G : Functor (Grothendieck F) H} [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] (c : Cocone G) :

          If c is a colimit cocone on G : Grockendieck F ⥤ H, then the induced cocone on the fiberwise colimit on G is a colimit cocone, too.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            def CategoryTheory.Limits.coconeOfCoconeFiberwiseColimit {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] {G : Functor (Grothendieck F) H} [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] (c : Cocone (fiberwiseColimit G)) :

            For a functor G : Grothendieck F ⥤ H, every cocone over fiberwiseColimit G induces a cocone over G itself.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              @[simp]
              theorem CategoryTheory.Limits.coconeOfCoconeFiberwiseColimit_ι_app {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] {G : Functor (Grothendieck F) H} [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] (c : Cocone (fiberwiseColimit G)) (X : Grothendieck F) :
              (coconeOfCoconeFiberwiseColimit c).app X = CategoryStruct.comp (colimit.ι ((Grothendieck.ι F X.base).comp G) X.fiber) (c.app X.base)

              If a cocone c over a functor G : Grothendieck F ⥤ H is a colimit, than the induced cocone coconeOfFiberwiseCocone G c

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

                We can infer that a functor G : Grothendieck F ⥤ H, with F : C ⥤ Cat, has a colimit from the fact that each of its fibers has a colimit and that these fiberwise colimits, as a functor C ⥤ H have a colimit.

                For every functor G on the Grothendieck construction Grothendieck F, if G has a colimit and every fiber of G has a colimit, then taking this colimit is isomorphic to first taking the fiberwise colimit and then the colimit of the resulting functor.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[simp]
                  theorem CategoryTheory.Limits.ι_colimitFiberwiseColimitIso_hom {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] (G : Functor (Grothendieck F) H) [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] [HasColimit (fiberwiseColimit G)] (X : C) (d : (F.obj X)) :
                  @[simp]
                  theorem CategoryTheory.Limits.ι_colimitFiberwiseColimitIso_hom_assoc {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] (G : Functor (Grothendieck F) H) [∀ {X Y : C} (f : X Y), HasColimit (Functor.comp (F.map f) ((Grothendieck.ι F Y).comp G))] [HasColimit (fiberwiseColimit G)] (X : C) (d : (F.obj X)) {Z : H} (h : colimit G Z) :

                  The isomorphism colimitFiberwiseColimitIso induces an isomorphism of functors (J ⥤ C) ⥤ C between fiberwiseColim F H ⋙ colim and colim.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    def CategoryTheory.Limits.fiberwiseColimCompEvaluationIso {C : Type u₁} [Category.{v₁, u₁} C] {F : Functor C Cat} {H : Type u₂} [Category.{v₂, u₂} H] [∀ (c : C), HasColimitsOfShape (↑(F.obj c)) H] (c : C) :
                    (fiberwiseColim F H).comp ((evaluation C H).obj c) ((whiskeringLeft (↑(F.obj c)) (Grothendieck F) H).obj (Grothendieck.ι F c)).comp colim

                    Composing fiberwiseColim F H with the evaluation functor (evaluation C H).obj c is naturally isomorphic to precomposing the Grothendieck inclusion Grothendieck.ι to colim.

                    Equations
                    Instances For