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.

structure CategoryTheory.Limits.DiagramOfCones {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] (F : Functor J (Functor K C)) :
Type (max (max (max (max u_1 u_2) u_3) u_4) u_6)

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

Instances For
    structure CategoryTheory.Limits.DiagramOfCocones {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] (F : Functor J (Functor K C)) :
    Type (max (max (max (max u_1 u_2) u_3) u_4) u_6)

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

    Instances For

      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
        @[simp]
        theorem CategoryTheory.Limits.DiagramOfCones.conePoints_map {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] {F : Functor J (Functor K C)} (D : DiagramOfCones F) {X✝ Y✝ : J} (f : X✝ Y✝) :
        D.conePoints.map f = (D.map f).hom

        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.DiagramOfCocones.coconePoints_map {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] {F : Functor J (Functor K C)} (D : DiagramOfCocones F) {X✝ Y✝ : J} (f : X✝ Y✝) :
          def CategoryTheory.Limits.coneOfConeUncurry {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] {F : Functor J (Functor K C)} {D : DiagramOfCones F} (Q : (j : J) → IsLimit (D.obj j)) (c : Cone (uncurry.obj F)) :

          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.coneOfConeUncurry_π_app {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] {F : Functor J (Functor K C)} {D : DiagramOfCones F} (Q : (j : J) → IsLimit (D.obj j)) (c : Cone (uncurry.obj F)) (j : J) :
            (coneOfConeUncurry Q c).π.app j = (Q j).lift { pt := c.pt, π := { app := fun (k : K) => c.π.app (j, k), naturality := } }
            @[simp]
            theorem CategoryTheory.Limits.coneOfConeUncurry_pt {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] {F : Functor J (Functor K C)} {D : DiagramOfCones F} (Q : (j : J) → IsLimit (D.obj j)) (c : Cone (uncurry.obj F)) :
            def CategoryTheory.Limits.coneOfConeCurry {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] (G : Functor (J × K) C) {D : DiagramOfCones (curry.obj G)} (Q : (j : J) → IsLimit (D.obj j)) (c : Cone G) :

            Given a diagram D of limit cones over the curry.obj G j, and a cone over G, 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.coneOfConeCurry_pt {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] (G : Functor (J × K) C) {D : DiagramOfCones (curry.obj G)} (Q : (j : J) → IsLimit (D.obj j)) (c : Cone G) :
              @[simp]
              theorem CategoryTheory.Limits.coneOfConeCurry_π_app {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] (G : Functor (J × K) C) {D : DiagramOfCones (curry.obj G)} (Q : (j : J) → IsLimit (D.obj j)) (c : Cone G) (j : J) :
              (coneOfConeCurry G Q c).π.app j = (Q j).lift { 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
                @[simp]
                theorem CategoryTheory.Limits.coconeOfCoconeUncurry_pt {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] {F : Functor J (Functor K C)} {D : DiagramOfCocones F} (Q : (j : J) → IsColimit (D.obj j)) (c : Cocone (uncurry.obj F)) :
                @[simp]
                theorem CategoryTheory.Limits.coconeOfCoconeUncurry_ι_app {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] {F : Functor J (Functor K C)} {D : DiagramOfCocones F} (Q : (j : J) → IsColimit (D.obj j)) (c : Cocone (uncurry.obj F)) (j : J) :
                (coconeOfCoconeUncurry Q c).ι.app j = (Q j).desc { pt := c.pt, ι := { app := fun (k : K) => c.ι.app (j, k), naturality := } }
                def CategoryTheory.Limits.coconeOfCoconeCurry {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] (G : Functor (J × K) C) {D : DiagramOfCocones (curry.obj G)} (Q : (j : J) → IsColimit (D.obj j)) (c : Cocone G) :

                Given a diagram D of colimit cocones under the curry.obj G j, and a cocone under G, we can construct a cocone under the diagram consisting of the cocone points from D.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[simp]
                  theorem CategoryTheory.Limits.coconeOfCoconeCurry_pt {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] (G : Functor (J × K) C) {D : DiagramOfCocones (curry.obj G)} (Q : (j : J) → IsColimit (D.obj j)) (c : Cocone G) :
                  @[simp]
                  theorem CategoryTheory.Limits.coconeOfCoconeCurry_ι_app {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] (G : Functor (J × K) C) {D : DiagramOfCocones (curry.obj G)} (Q : (j : J) → IsColimit (D.obj j)) (c : Cocone G) (j : J) :
                  (coconeOfCoconeCurry G Q c).ι.app j = (Q j).desc { pt := c.pt, ι := { app := fun (k : K) => c.ι.app (j, k), naturality := } }
                  def CategoryTheory.Limits.coneOfConeUncurryIsLimit {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] {F : Functor J (Functor K C)} {D : DiagramOfCones F} (Q : (j : J) → IsLimit (D.obj j)) {c : Cone (uncurry.obj F)} (P : IsLimit c) :

                  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
                        @[simp]
                        theorem CategoryTheory.Limits.DiagramOfCones.mkOfHasLimits_map_hom {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] (F : Functor J (Functor K C)) [HasLimitsOfShape K C] {j✝ j'✝ : J} (f : j✝ j'✝) :
                        ((mkOfHasLimits F).map f).hom = lim.map (F.map f)

                        Given a functor G : J × K ⥤ C such that (curry.obj G ⋙ lim) makes sense and has a limit, we can construct a cone over G with limit (curry.obj G ⋙ lim) as a cone point

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

                          The cone coneOfHasLimitCurryCompLim is in fact a limit cone.

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

                            The functor G has a limit if C has K-shaped limits and (curry.obj G ⋙ lim) has a limit.

                            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
                                @[simp]
                                theorem CategoryTheory.Limits.DiagramOfCocones.mkOfHasColimits_map_hom {J : Type u_1} {K : Type u_2} [Category.{u_4, u_1} J] [Category.{u_5, u_2} K] {C : Type u_3} [Category.{u_6, u_3} C] (F : Functor J (Functor K C)) [HasColimitsOfShape K C] {j✝ j'✝ : J} (f : j✝ j'✝) :

                                Given a functor G : J × K ⥤ C such that (curry.obj G ⋙ colim) makes sense and has a colimit, we can construct a cocone under G with colimit (curry.obj G ⋙ colim) as a cocone point

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

                                  The cocone coconeOfHasColimitCurryCompColim is in fact a limit cocone.

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

                                    The functor G has a colimit if C has K-shaped colimits and (curry.obj G ⋙ colim) has a colimit.

                                    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

                                          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

                                              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

                                                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