Documentation

Mathlib.CategoryTheory.Limits.Fubini

A Fubini theorem for categorical (co)limits #

We prove that $lim_{J × K} G = lim_J (lim_K G(j, -))$ for a functor G : J × K ⥤ C, when all the appropriate limits exist.

We begin working with a functor F : J ⥤ K ⥤ C. We'll write G : J × K ⥤ C for the associated "uncurried" functor.

In the first part, given a coherent family D of limit cones over the functors F.obj j, and a cone c over G, we construct a cone over the cone points of D. We then show that if c is a limit cone, the constructed cone is also a limit cone.

In the second part, we state the Fubini theorem in the setting where limits are provided by suitable HasLimit classes.

We construct limitUncurryIsoLimitCompLim F : limit (uncurry.obj F) ≅ limit (F ⋙ lim) and give simp lemmas characterising it. For convenience, we also provide limitIsoLimitCurryCompLim G : limit G ≅ limit ((curry.obj G) ⋙ lim) in terms of the uncurried functor.

All statements have their counterpart for colimits.

A structure carrying a diagram of cones over the functors F.obj j.

Instances For
    theorem CategoryTheory.Limits.DiagramOfCones.comp {J : Type v} {K : Type v} [CategoryTheory.SmallCategory J] [CategoryTheory.SmallCategory K] {C : Type u} [CategoryTheory.Category.{v, u} C] {F : CategoryTheory.Functor J (CategoryTheory.Functor K C)} (self : CategoryTheory.Limits.DiagramOfCones F) {j₁ : J} {j₂ : J} {j₃ : J} (f : j₁ j₂) (g : j₂ j₃) :
    (self.map (CategoryTheory.CategoryStruct.comp f g)).hom = CategoryTheory.CategoryStruct.comp (self.map f).hom (self.map g).hom

    A structure carrying a diagram of cocones over the functors F.obj j.

    Instances For
      theorem CategoryTheory.Limits.DiagramOfCocones.comp {J : Type v} {K : Type v} [CategoryTheory.SmallCategory J] [CategoryTheory.SmallCategory K] {C : Type u} [CategoryTheory.Category.{v, u} C] {F : CategoryTheory.Functor J (CategoryTheory.Functor K C)} (self : CategoryTheory.Limits.DiagramOfCocones F) {j₁ : J} {j₂ : J} {j₃ : J} (f : j₁ j₂) (g : j₂ j₃) :
      (self.map (CategoryTheory.CategoryStruct.comp f g)).hom = CategoryTheory.CategoryStruct.comp (self.map f).hom (self.map g).hom

      Extract the functor J ⥤ C consisting of the cone points and the maps between them, from a DiagramOfCones.

      Equations
      • D.conePoints = { obj := fun (j : J) => (D.obj j).pt, map := fun {X Y : J} (f : X Y) => (D.map f).hom, map_id := , map_comp := }
      Instances For

        Extract the functor J ⥤ C consisting of the cocone points and the maps between them, from a DiagramOfCocones.

        Equations
        • D.coconePoints = { obj := fun (j : J) => (D.obj j).pt, map := fun {X Y : J} (f : X Y) => (D.map f).hom, map_id := , map_comp := }
        Instances For
          @[simp]
          theorem CategoryTheory.Limits.coneOfConeUncurry_π_app {J : Type v} {K : Type v} [CategoryTheory.SmallCategory J] [CategoryTheory.SmallCategory K] {C : Type u} [CategoryTheory.Category.{v, u} C] {F : CategoryTheory.Functor J (CategoryTheory.Functor K C)} {D : CategoryTheory.Limits.DiagramOfCones F} (Q : (j : J) → CategoryTheory.Limits.IsLimit (D.obj j)) (c : CategoryTheory.Limits.Cone (CategoryTheory.uncurry.obj F)) (j : J) :
          (CategoryTheory.Limits.coneOfConeUncurry Q c).app j = (Q j).lift { pt := c.pt, π := { app := fun (k : K) => c.app (j, k), naturality := } }

          Given a diagram D of limit cones over the F.obj j, and a cone over uncurry.obj F, we can construct a cone over the diagram consisting of the cone points from D.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            @[simp]
            theorem CategoryTheory.Limits.coconeOfCoconeUncurry_ι_app {J : Type v} {K : Type v} [CategoryTheory.SmallCategory J] [CategoryTheory.SmallCategory K] {C : Type u} [CategoryTheory.Category.{v, u} C] {F : CategoryTheory.Functor J (CategoryTheory.Functor K C)} {D : CategoryTheory.Limits.DiagramOfCocones F} (Q : (j : J) → CategoryTheory.Limits.IsColimit (D.obj j)) (c : CategoryTheory.Limits.Cocone (CategoryTheory.uncurry.obj F)) (j : J) :
            (CategoryTheory.Limits.coconeOfCoconeUncurry Q c).app j = (Q j).desc { pt := c.pt, ι := { app := fun (k : K) => c.app (j, k), naturality := } }

            Given a diagram D of colimit cocones over the F.obj j, and a cocone over uncurry.obj F, we can construct a cocone over the diagram consisting of the cocone points from D.

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

              coneOfConeUncurry Q c is a limit cone when c is a limit cone.

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

                coconeOfCoconeUncurry Q c is a colimit cocone when c is a colimit cocone.

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

                  Given a functor F : J ⥤ K ⥤ C, with all needed limits, we can construct a diagram consisting of the limit cone over each functor F.obj j, and the universal cone morphisms between these.

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

                    The Fubini theorem for a functor F : J ⥤ K ⥤ C, showing that the limit of uncurry.obj F can be computed as the limit of the limits of the functors F.obj j.

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

                      Given a functor F : J ⥤ K ⥤ C, with all needed colimits, we can construct a diagram consisting of the colimit cocone over each functor F.obj j, and the universal cocone morphisms between these.

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

                        The Fubini theorem for a functor F : J ⥤ K ⥤ C, showing that the colimit of uncurry.obj F can be computed as the colimit of the colimits of the functors F.obj j.

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

                          The limit of F.flip ⋙ lim is isomorphic to the limit of F ⋙ lim.

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

                            The colimit of F.flip ⋙ colim is isomorphic to the colimit of F ⋙ colim.

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For
                              noncomputable def CategoryTheory.Limits.limitIsoLimitCurryCompLim {J : Type v} {K : Type v} [CategoryTheory.SmallCategory J] [CategoryTheory.SmallCategory K] {C : Type u} [CategoryTheory.Category.{v, u} C] (G : CategoryTheory.Functor (J × K) C) [CategoryTheory.Limits.HasLimitsOfShape K C] [CategoryTheory.Limits.HasLimit G] [CategoryTheory.Limits.HasLimit ((CategoryTheory.curry.obj G).comp CategoryTheory.Limits.lim)] :
                              CategoryTheory.Limits.limit G CategoryTheory.Limits.limit ((CategoryTheory.curry.obj G).comp CategoryTheory.Limits.lim)

                              The Fubini theorem for a functor G : J × K ⥤ C, showing that the limit of G can be computed as the limit of the limits of the functors G.obj (j, _).

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

                                The Fubini theorem for a functor G : J × K ⥤ C, showing that the colimit of G can be computed as the colimit of the colimits of the functors G.obj (j, _).

                                Equations
                                • One or more equations did not get rendered due to their size.
                                Instances For
                                  noncomputable def CategoryTheory.Limits.limitCurrySwapCompLimIsoLimitCurryCompLim {J : Type v} {K : Type v} [CategoryTheory.SmallCategory J] [CategoryTheory.SmallCategory K] {C : Type u} [CategoryTheory.Category.{v, u} C] (G : CategoryTheory.Functor (J × K) C) [CategoryTheory.Limits.HasLimits C] :
                                  CategoryTheory.Limits.limit ((CategoryTheory.curry.obj ((CategoryTheory.Prod.swap K J).comp G)).comp CategoryTheory.Limits.lim) CategoryTheory.Limits.limit ((CategoryTheory.curry.obj G).comp CategoryTheory.Limits.lim)

                                  A variant of the Fubini theorem for a functor G : J × K ⥤ C, showing that $\lim_k \lim_j G(j,k) ≅ \lim_j \lim_k G(j,k)$.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    noncomputable def CategoryTheory.Limits.colimitCurrySwapCompColimIsoColimitCurryCompColim {J : Type v} {K : Type v} [CategoryTheory.SmallCategory J] [CategoryTheory.SmallCategory K] {C : Type u} [CategoryTheory.Category.{v, u} C] (G : CategoryTheory.Functor (J × K) C) [CategoryTheory.Limits.HasColimits C] :
                                    CategoryTheory.Limits.colimit ((CategoryTheory.curry.obj ((CategoryTheory.Prod.swap K J).comp G)).comp CategoryTheory.Limits.colim) CategoryTheory.Limits.colimit ((CategoryTheory.curry.obj G).comp CategoryTheory.Limits.colim)

                                    A variant of the Fubini theorem for a functor G : J × K ⥤ C, showing that $\colim_k \colim_j G(j,k) ≅ \colim_j \colim_k G(j,k)$.

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