Documentation

Mathlib.CategoryTheory.GuitartExact

Guitart exact squares #

Given four functors T, L, R and B, a 2-square TwoSquare T L R B consists of a natural transformation w : T ⋙ R ⟶ L ⋙ B:

     T
  C₁ ⥤ C₂
L |     | R
  v     v
  C₃ ⥤ C₄
     B

In this file, we define a typeclass w.GuitartExact which expresses that this square is exact in the sense of Guitart. This means that for any X₃ : C₃, the induced functor CostructuredArrow L X₃ ⥤ CostructuredArrow R (B.obj X₃) is final. It is also equivalent to the fact that for any X₂ : C₂, the induced functor StructuredArrow X₂ T ⥤ StructuredArrow (R.obj X₂) B is initial.

Various categorical notions (fully faithful functors, adjunctions, etc.) can be characterized in terms of Guitart exact squares. Their particular role in pointwise Kan extensions shall also be used in the construction of derived functors.

TODO #

References #

A 2-square consists of a natural transformation T ⋙ R ⟶ L ⋙ B involving fours functors T, L, R, B that are on the top/left/right/bottom sides of a square of categories.

Equations
Instances For
    theorem CategoryTheory.TwoSquare.ext {C₁ : Type u₁} {C₂ : Type u₂} {C₃ : Type u₃} {C₄ : Type u₄} [CategoryTheory.Category.{v₁, u₁} C₁] [CategoryTheory.Category.{v₂, u₂} C₂] [CategoryTheory.Category.{v₃, u₃} C₃] [CategoryTheory.Category.{v₄, u₄} C₄] {T : CategoryTheory.Functor C₁ C₂} {L : CategoryTheory.Functor C₁ C₃} {R : CategoryTheory.Functor C₂ C₄} {B : CategoryTheory.Functor C₃ C₄} (w : CategoryTheory.TwoSquare T L R B) (w' : CategoryTheory.TwoSquare T L R B) (h : ∀ (X : C₁), w.app X = w'.app X) :
    w = w'

    Given w : TwoSquare T L R B and X₃ : C₃, this is the obvious functor CostructuredArrow L X₃ ⥤ CostructuredArrow R (B.obj X₃).

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

      Given w : TwoSquare T L R B and X₂ : C₂, this is the obvious functor StructuredArrow X₂ T ⥤ StructuredArrow (R.obj X₂) B.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[inline, reducible]
        abbrev CategoryTheory.TwoSquare.StructuredArrowRightwards {C₁ : Type u₁} {C₂ : Type u₂} {C₃ : Type u₃} {C₄ : Type u₄} [CategoryTheory.Category.{v₁, u₁} C₁] [CategoryTheory.Category.{v₂, u₂} C₂] [CategoryTheory.Category.{v₃, u₃} C₃] [CategoryTheory.Category.{v₄, u₄} C₄] {T : CategoryTheory.Functor C₁ C₂} {L : CategoryTheory.Functor C₁ C₃} {R : CategoryTheory.Functor C₂ C₄} {B : CategoryTheory.Functor C₃ C₄} (w : CategoryTheory.TwoSquare T L R B) {X₂ : C₂} {X₃ : C₃} (g : R.obj X₂ B.obj X₃) :
        Type (max (max u₁ v₃) v₂)

        Given w : TwoSquare T L R B and a morphism g : R.obj X₂ ⟶ B.obj X₃, this is the category StructuredArrow (CostructuredArrow.mk g) (w.costructuredArrowRightwards X₃), see the constructor StructuredArrowRightwards.mk for the data that is involved.

        Equations
        Instances For
          @[inline, reducible]
          abbrev CategoryTheory.TwoSquare.CostructuredArrowDownwards {C₁ : Type u₁} {C₂ : Type u₂} {C₃ : Type u₃} {C₄ : Type u₄} [CategoryTheory.Category.{v₁, u₁} C₁] [CategoryTheory.Category.{v₂, u₂} C₂] [CategoryTheory.Category.{v₃, u₃} C₃] [CategoryTheory.Category.{v₄, u₄} C₄] {T : CategoryTheory.Functor C₁ C₂} {L : CategoryTheory.Functor C₁ C₃} {R : CategoryTheory.Functor C₂ C₄} {B : CategoryTheory.Functor C₃ C₄} (w : CategoryTheory.TwoSquare T L R B) {X₂ : C₂} {X₃ : C₃} (g : R.obj X₂ B.obj X₃) :
          Type (max (max u₁ v₂) v₃)

          Given w : TwoSquare T L R B and a morphism g : R.obj X₂ ⟶ B.obj X₃, this is the category CostructuredArrow (w.structuredArrowDownwards X₂) (StructuredArrow.mk g), see the constructor CostructuredArrowDownwards.mk for the data that is involved.

          Equations
          Instances For
            @[inline, reducible]
            abbrev CategoryTheory.TwoSquare.StructuredArrowRightwards.mk {C₁ : Type u₁} {C₂ : Type u₂} {C₃ : Type u₃} {C₄ : Type u₄} [CategoryTheory.Category.{v₁, u₁} C₁] [CategoryTheory.Category.{v₂, u₂} C₂] [CategoryTheory.Category.{v₃, u₃} C₃] [CategoryTheory.Category.{v₄, u₄} C₄] {T : CategoryTheory.Functor C₁ C₂} {L : CategoryTheory.Functor C₁ C₃} {R : CategoryTheory.Functor C₂ C₄} {B : CategoryTheory.Functor C₃ C₄} (w : CategoryTheory.TwoSquare T L R B) {X₂ : C₂} {X₃ : C₃} (g : R.obj X₂ B.obj X₃) (X₁ : C₁) (a : X₂ T.obj X₁) (b : L.obj X₁ X₃) (comm : CategoryTheory.CategoryStruct.comp (R.map a) (CategoryTheory.CategoryStruct.comp (w.app X₁) (B.map b)) = g) :

            Constructor for objects in w.StructuredArrowRightwards g.

            Equations
            Instances For
              @[inline, reducible]
              abbrev CategoryTheory.TwoSquare.CoStructuredArrowDownwards.mk {C₁ : Type u₁} {C₂ : Type u₂} {C₃ : Type u₃} {C₄ : Type u₄} [CategoryTheory.Category.{v₁, u₁} C₁] [CategoryTheory.Category.{v₂, u₂} C₂] [CategoryTheory.Category.{v₃, u₃} C₃] [CategoryTheory.Category.{v₄, u₄} C₄] {T : CategoryTheory.Functor C₁ C₂} {L : CategoryTheory.Functor C₁ C₃} {R : CategoryTheory.Functor C₂ C₄} {B : CategoryTheory.Functor C₃ C₄} (w : CategoryTheory.TwoSquare T L R B) {X₂ : C₂} {X₃ : C₃} (g : R.obj X₂ B.obj X₃) (X₁ : C₁) (a : X₂ T.obj X₁) (b : L.obj X₁ X₃) (comm : CategoryTheory.CategoryStruct.comp (R.map a) (CategoryTheory.CategoryStruct.comp (w.app X₁) (B.map b)) = g) :

              Constructor for objects in w.CostructuredArrowDownwards g.

              Equations
              Instances For

                Given w : TwoSquare T L R B and a morphism g : R.obj X₂ ⟶ B.obj X₃, this is the obvious functor w.StructuredArrowRightwards g ⥤ w.CostructuredArrowDownwards g.

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

                  Given w : TwoSquare T L R B and a morphism g : R.obj X₂ ⟶ B.obj X₃, this is the obvious functor w.CostructuredArrowDownwards g ⥤ w.StructuredArrowRightwards g.

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

                    Given w : TwoSquare T L R B and a morphism g : R.obj X₂ ⟶ B.obj X₃, this is the obvious equivalence of categories w.StructuredArrowRightwards g ≌ w.CostructuredArrowDownwards g.

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

                      Condition on w : TwoSquare T L R B expressing that it is a Guitart exact square. It is equivalent to saying that for any X₃ : C₃, the induced functor CostructuredArrow L X₃ ⥤ CostructuredArrow R (B.obj X₃) is final (see guitartExact_iff_final) or equivalently that for any X₂ : C₂, the induced functor StructuredArrow X₂ T ⥤ StructuredArrow (R.obj X₂) B is initial (see guitartExact_iff_initial). See also guitartExact_iff_isConnected_rightwards, guitartExact_iff_isConnected_downwards for characterizations in terms of the connectedness of auxiliary categories.

                      Instances