Documentation

Mathlib.Algebra.Homology.LeftResolution.Basic

Left resolutions #

Given a fully faithful functor ι : C ⥤ A to an abelian category, we introduce a structure Abelian.LeftResolution ι which gives a functor F : A ⥤ C and a natural epimorphism π.app X : ι.obj (F.obj X) ⟶ X for all X : A. This is used in order to construct a resolution functor LeftResolution.chainComplexFunctor : A ⥤ ChainComplex C ℕ.

This shall be used in order to construct functorial flat resolutions.

structure CategoryTheory.Abelian.LeftResolution {A : Type u_1} {C : Type u_2} [Category.{u_3, u_2} C] [Category.{u_4, u_1} A] (ι : Functor C A) :
Type (max (max (max u_1 u_2) u_3) u_4)

Given a fully faithful functor ι : C ⥤ A, this structure contains the data of a functor F : A ⥤ C and a functorial epimorphism π.app X : ι.obj (F.obj X) ⟶ X for all X : A.

  • F : Functor A C

    a functor which sends X : A to an object F.obj X with an epimorphism π.app X : ι.obj (F.obj X) ⟶ X

  • π : self.F.comp ι Functor.id A

    the natural epimorphism

  • epi_π_app (X : A) : Epi (self.π.app X)
Instances For
    @[simp]
    theorem CategoryTheory.Abelian.LeftResolution.π_naturality {A : Type u_1} {C : Type u_2} [Category.{u_4, u_2} C] [Category.{u_3, u_1} A] {ι : Functor C A} (Λ : LeftResolution ι) (X Y : A) (f : X Y) :
    @[simp]
    theorem CategoryTheory.Abelian.LeftResolution.π_naturality_assoc {A : Type u_1} {C : Type u_2} [Category.{u_4, u_2} C] [Category.{u_3, u_1} A] {ι : Functor C A} (Λ : LeftResolution ι) (X Y : A) (f : X Y) {Z : A} (h : Y Z) :
    @[irreducible]

    Given ι : C ⥤ A, Λ : LeftResolution ι, X : A, this is a chain complex which is a (functorial) resolution of A that is obtained inductively by using the epimorphisms given by Λ.

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

      Given Λ : LeftResolution ι, the chain complex Λ.chainComplex X identifies in degree 0 to Λ.F.obj X.

      Equations
      Instances For

        Given Λ : LeftResolution ι, the chain complex Λ.chainComplex X identifies in degree 1 to Λ.F.obj (kernel (Λ.π.app X)).

        Equations
        Instances For
          noncomputable def CategoryTheory.Abelian.LeftResolution.chainComplexXIso {A : Type u_1} {C : Type u_2} [Category.{u_3, u_2} C] [Category.{u_4, u_1} A] {ι : Functor C A} (Λ : LeftResolution ι) (X : A) [ι.Full] [ι.Faithful] [Limits.HasZeroMorphisms C] [Abelian A] (n : ) :
          (Λ.chainComplex X).X (n + 2) Λ.F.obj (Limits.kernel (ι.map ((Λ.chainComplex X).d (n + 1) n)))

          The isomorphism which gives the inductive step of the construction of Λ.chainComplex X.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            theorem CategoryTheory.Abelian.LeftResolution.map_chainComplex_d {A : Type u_1} {C : Type u_2} [Category.{u_4, u_2} C] [Category.{u_3, u_1} A] {ι : Functor C A} (Λ : LeftResolution ι) (X : A) [ι.Full] [ι.Faithful] [Limits.HasZeroMorphisms C] [Abelian A] (n : ) :
            ι.map ((Λ.chainComplex X).d (n + 2) (n + 1)) = CategoryStruct.comp (ι.map (Λ.chainComplexXIso X n).hom) (CategoryStruct.comp (Λ.π.app (Limits.kernel (ι.map ((Λ.chainComplex X).d (n + 1) n)))) (Limits.kernel.ι (ι.map ((Λ.chainComplex X).d (n + 1) n))))
            noncomputable def CategoryTheory.Abelian.LeftResolution.chainComplexMap {A : Type u_1} {C : Type u_2} [Category.{u_3, u_2} C] [Category.{u_4, u_1} A] {ι : Functor C A} (Λ : LeftResolution ι) {X Y : A} (f : X Y) [ι.Full] [ι.Faithful] [Limits.HasZeroMorphisms C] [Abelian A] :

            The morphism Λ.chainComplex X ⟶ Λ.chainComplex Y of chain complexes induced by f : X ⟶ Y.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem CategoryTheory.Abelian.LeftResolution.chainComplexMap_f_succ_succ {A : Type u_1} {C : Type u_2} [Category.{u_3, u_2} C] [Category.{u_4, u_1} A] {ι : Functor C A} (Λ : LeftResolution ι) {X Y : A} (f : X Y) [ι.Full] [ι.Faithful] [Limits.HasZeroMorphisms C] [Abelian A] (n : ) :
              (Λ.chainComplexMap f).f (n + 2) = CategoryStruct.comp (Λ.chainComplexXIso X n).hom (CategoryStruct.comp (Λ.F.map (Limits.kernel.map (ι.map ((Λ.chainComplex X).d (n + 1) n)) (ι.map ((Λ.chainComplex Y).d (n + 1) n)) (ι.map ((Λ.chainComplexMap f).f (n + 1))) (ι.map ((Λ.chainComplexMap f).f n)) )) (Λ.chainComplexXIso Y n).inv)

              Given ι : C ⥤ A, Λ : LeftResolution ι, this is a functor A ⥤ ChainComplex C ℕ which sends X : A to a resolution consisting of objects in C.

              Equations
              Instances For