Documentation

Mathlib.CategoryTheory.Abelian.ProjectiveResolution

Abelian categories with enough projectives have projective resolutions #

Main results #

When the underlying category is abelian:

noncomputable def CategoryTheory.ProjectiveResolution.liftFZero {C : Type u} [Category.{v, u} C] [Limits.HasZeroObject C] [Limits.HasZeroMorphisms C] {Y Z : C} (f : Y Z) (P : ProjectiveResolution Y) (Q : ProjectiveResolution Z) :
P.complex.X 0 Q.complex.X 0

Auxiliary construction for lift.

Equations
Instances For
    theorem CategoryTheory.ProjectiveResolution.exact₀ {C : Type u} [Category.{v, u} C] [Abelian C] {Z : C} (P : ProjectiveResolution Z) :
    (ShortComplex.mk (P.complex.d 1 0) (P.f 0) ).Exact
    noncomputable def CategoryTheory.ProjectiveResolution.liftFOne {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} (f : Y Z) (P : ProjectiveResolution Y) (Q : ProjectiveResolution Z) :
    P.complex.X 1 Q.complex.X 1

    Auxiliary construction for lift.

    Equations
    Instances For
      @[simp]
      theorem CategoryTheory.ProjectiveResolution.liftFOne_zero_comm {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} (f : Y Z) (P : ProjectiveResolution Y) (Q : ProjectiveResolution Z) :
      CategoryStruct.comp (liftFOne f P Q) (Q.complex.d 1 0) = CategoryStruct.comp (P.complex.d 1 0) (liftFZero f P Q)
      noncomputable def CategoryTheory.ProjectiveResolution.liftFSucc {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} (P : ProjectiveResolution Y) (Q : ProjectiveResolution Z) (n : ) (g : P.complex.X n Q.complex.X n) (g' : P.complex.X (n + 1) Q.complex.X (n + 1)) (w : CategoryStruct.comp g' (Q.complex.d (n + 1) n) = CategoryStruct.comp (P.complex.d (n + 1) n) g) :
      (g'' : P.complex.X (n + 2) Q.complex.X (n + 2)) ×' CategoryStruct.comp g'' (Q.complex.d (n + 2) (n + 1)) = CategoryStruct.comp (P.complex.d (n + 2) (n + 1)) g'

      Auxiliary construction for lift.

      Equations
      Instances For
        noncomputable def CategoryTheory.ProjectiveResolution.lift {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} (f : Y Z) (P : ProjectiveResolution Y) (Q : ProjectiveResolution Z) :
        P.complex Q.complex

        A morphism in C lift to a chain map between projective resolutions.

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

          The resolution maps intertwine the lift of a morphism and that morphism.

          @[simp]
          @[simp]
          theorem CategoryTheory.ProjectiveResolution.lift_commutes_zero_assoc {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} (f : Y Z) (P : ProjectiveResolution Y) (Q : ProjectiveResolution Z) {Z✝ : C} (h : ((ChainComplex.single₀ C).obj Z).X 0 Z✝) :
          noncomputable def CategoryTheory.ProjectiveResolution.liftHomotopyZeroZero {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (f : P.complex Q.complex) (comm : CategoryStruct.comp f Q = 0) :
          P.complex.X 0 Q.complex.X 1

          An auxiliary definition for liftHomotopyZero.

          Equations
          Instances For
            @[simp]
            theorem CategoryTheory.ProjectiveResolution.liftHomotopyZeroZero_comp {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (f : P.complex Q.complex) (comm : CategoryStruct.comp f Q = 0) :
            CategoryStruct.comp (liftHomotopyZeroZero f comm) (Q.complex.d 1 0) = f.f 0
            @[simp]
            theorem CategoryTheory.ProjectiveResolution.liftHomotopyZeroZero_comp_assoc {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (f : P.complex Q.complex) (comm : CategoryStruct.comp f Q = 0) {Z✝ : C} (h : Q.complex.X 0 Z✝) :
            noncomputable def CategoryTheory.ProjectiveResolution.liftHomotopyZeroOne {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (f : P.complex Q.complex) (comm : CategoryStruct.comp f Q = 0) :
            P.complex.X 1 Q.complex.X 2

            An auxiliary definition for liftHomotopyZero.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem CategoryTheory.ProjectiveResolution.liftHomotopyZeroOne_comp {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (f : P.complex Q.complex) (comm : CategoryStruct.comp f Q = 0) :
              CategoryStruct.comp (liftHomotopyZeroOne f comm) (Q.complex.d 2 1) = f.f 1 - CategoryStruct.comp (P.complex.d 1 0) (liftHomotopyZeroZero f comm)
              @[simp]
              theorem CategoryTheory.ProjectiveResolution.liftHomotopyZeroOne_comp_assoc {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (f : P.complex Q.complex) (comm : CategoryStruct.comp f Q = 0) {Z✝ : C} (h : Q.complex.X 1 Z✝) :
              noncomputable def CategoryTheory.ProjectiveResolution.liftHomotopyZeroSucc {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (f : P.complex Q.complex) (n : ) (g : P.complex.X n Q.complex.X (n + 1)) (g' : P.complex.X (n + 1) Q.complex.X (n + 2)) (w : f.f (n + 1) = CategoryStruct.comp (P.complex.d (n + 1) n) g + CategoryStruct.comp g' (Q.complex.d (n + 2) (n + 1))) :
              P.complex.X (n + 2) Q.complex.X (n + 3)

              An auxiliary definition for liftHomotopyZero.

              Equations
              Instances For
                @[simp]
                theorem CategoryTheory.ProjectiveResolution.liftHomotopyZeroSucc_comp {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (f : P.complex Q.complex) (n : ) (g : P.complex.X n Q.complex.X (n + 1)) (g' : P.complex.X (n + 1) Q.complex.X (n + 2)) (w : f.f (n + 1) = CategoryStruct.comp (P.complex.d (n + 1) n) g + CategoryStruct.comp g' (Q.complex.d (n + 2) (n + 1))) :
                CategoryStruct.comp (liftHomotopyZeroSucc f n g g' w) (Q.complex.d (n + 3) (n + 2)) = f.f (n + 2) - CategoryStruct.comp (P.complex.d (n + 2) (n + 1)) g'
                @[simp]
                theorem CategoryTheory.ProjectiveResolution.liftHomotopyZeroSucc_comp_assoc {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (f : P.complex Q.complex) (n : ) (g : P.complex.X n Q.complex.X (n + 1)) (g' : P.complex.X (n + 1) Q.complex.X (n + 2)) (w : f.f (n + 1) = CategoryStruct.comp (P.complex.d (n + 1) n) g + CategoryStruct.comp g' (Q.complex.d (n + 2) (n + 1))) {Z✝ : C} (h : Q.complex.X (n + 2) Z✝) :
                CategoryStruct.comp (liftHomotopyZeroSucc f n g g' w) (CategoryStruct.comp (Q.complex.d (n + 3) (n + 2)) h) = CategoryStruct.comp (f.f (n + 2) - CategoryStruct.comp (P.complex.d (n + 2) (n + 1)) g') h
                noncomputable def CategoryTheory.ProjectiveResolution.liftHomotopyZero {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (f : P.complex Q.complex) (comm : CategoryStruct.comp f Q = 0) :

                Any lift of the zero morphism is homotopic to zero.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  noncomputable def CategoryTheory.ProjectiveResolution.liftHomotopy {C : Type u} [Category.{v, u} C] [Abelian C] {Y Z : C} (f : Y Z) {P : ProjectiveResolution Y} {Q : ProjectiveResolution Z} (g h : P.complex Q.complex) (g_comm : CategoryStruct.comp g Q = CategoryStruct.comp P ((ChainComplex.single₀ C).map f)) (h_comm : CategoryStruct.comp h Q = CategoryStruct.comp P ((ChainComplex.single₀ C).map f)) :

                  Two lifts of the same morphism are homotopic.

                  Equations
                  Instances For

                    The lift of the identity morphism is homotopic to the identity chain map.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      noncomputable def CategoryTheory.ProjectiveResolution.liftCompHomotopy {C : Type u} [Category.{v, u} C] [Abelian C] {X Y Z : C} (f : X Y) (g : Y Z) (P : ProjectiveResolution X) (Q : ProjectiveResolution Y) (R : ProjectiveResolution Z) :

                      The lift of a composition is homotopic to the composition of the lifts.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        noncomputable def CategoryTheory.ProjectiveResolution.homotopyEquiv {C : Type u} [Category.{v, u} C] [Abelian C] {X : C} (P Q : ProjectiveResolution X) :
                        HomotopyEquiv P.complex Q.complex

                        Any two projective resolutions are homotopy equivalent.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          @[simp]
                          theorem CategoryTheory.ProjectiveResolution.homotopyEquiv_hom_π {C : Type u} [Category.{v, u} C] [Abelian C] {X : C} (P Q : ProjectiveResolution X) :
                          CategoryStruct.comp (P.homotopyEquiv Q).hom Q = P
                          @[simp]
                          theorem CategoryTheory.ProjectiveResolution.homotopyEquiv_inv_π {C : Type u} [Category.{v, u} C] [Abelian C] {X : C} (P Q : ProjectiveResolution X) :
                          CategoryStruct.comp (P.homotopyEquiv Q).inv P = Q
                          @[reducible, inline]

                          An arbitrarily chosen projective resolution of an object.

                          Equations
                          Instances For

                            Taking projective resolutions is functorial, if considered with target the homotopy category (-indexed chain complexes and chain maps up to homotopy).

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

                              If P : ProjectiveResolution X, then the chosen (projectiveResolutions C).obj X is isomorphic (in the homotopy category) to P.complex.

                              Equations
                              Instances For
                                theorem CategoryTheory.exact_d_f {C : Type u} [Category.{v, u} C] [Abelian C] [EnoughProjectives C] {X Y : C} (f : X Y) :
                                (ShortComplex.mk (Projective.d f) f ).Exact

                                Our goal is to define ProjectiveResolution.of Z : ProjectiveResolution Z. The 0-th object in this resolution will just be Projective.over Z, i.e. an arbitrarily chosen projective object with a map to Z. After that, we build the n+1-st object as Projective.syzygies applied to the previously constructed morphism, and the map from the n-th object as Projective.d.

                                Auxiliary definition for ProjectiveResolution.of.

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

                                  In any abelian category with enough projectives, ProjectiveResolution.of Z constructs an projective resolution of the object Z.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    theorem CategoryTheory.ProjectiveResolution.of_def {C : Type u_1} [Category.{u_2, u_1} C] [Abelian C] [EnoughProjectives C] (Z : C) :
                                    of Z = mk (ofComplex Z) (((ofComplex Z).toSingle₀Equiv Z).symm Projective.π Z, )