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.
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 : Ato an objectF.obj Xwith an epimorphismπ.app X : ι.obj (F.obj X) ⟶ X the natural epimorphism
Instances For
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
- Λ.chainComplexXZeroIso X = CategoryTheory.Iso.refl ((Λ.chainComplex X).X 0)
Instances For
Given Λ : LeftResolution ι, the chain complex Λ.chainComplex X
identifies in degree 1 to Λ.F.obj (kernel (Λ.π.app X)).
Equations
- Λ.chainComplexXOneIso X = CategoryTheory.Iso.refl ((Λ.chainComplex X).X 1)
Instances For
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
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
Given ι : C ⥤ A, Λ : LeftResolution ι, this is a
functor A ⥤ ChainComplex C ℕ which sends X : A to a resolution consisting
of objects in C.
Equations
- Λ.chainComplexFunctor = { obj := fun (X : A) => Λ.chainComplex X, map := fun {X Y : A} (f : X ⟶ Y) => Λ.chainComplexMap f, map_id := ⋯, map_comp := ⋯ }