Abelian categories
This file contains the definition and basic properties of abelian categories.
There are many definitions of abelian category. Our definition is as follows: A category is called abelian if it is preadditive, has a finite products, kernels and cokernels, and if every monomorphism and epimorphism is normal.
It should be noted that if we also assume coproducts, then preadditivity is
actually a consequence of the other properties, as we show in
non_preadditive_abelian.lean
. However, this fact is of little practical
relevance, since essentially all interesting abelian categories come with a
preadditive structure. In this way, by requiring preadditivity, we allow the
user to pass in the preadditive structure the specific category they are
working with has natively.
Main definitions
abelian
is the type class indicating that a category is abelian. It extendspreadditive
.abelian.image f
iskernel (cokernel.π f)
, andabelian.coimage f
iscokernel (kernel.ι f)
.
Main results
- In an abelian category, mono + epi = iso.
- If
f : X ⟶ Y
, then the mapfactor_thru_image f : X ⟶ image f
is an epimorphism, and the mapfactor_thru_coimage f : coimage f ⟶ Y
is a monomorphism. - Factoring through the image and coimage is a strong epi-mono factorisation. This means that
- every abelian category has images. We instantiated this in such a way that
abelian.image f
is definitionally equal tolimits.image f
, and - there is a canonical isomorphism
coimage_iso_image : coimage f ≅ image f
such thatcoimage.π f ≫ (coimage_iso_image f).hom ≫ image.ι f = f
. The lemma stating this is calledfull_image_factorisation
.
- every abelian category has images. We instantiated this in such a way that
- Every epimorphism is a cokernel of its kernel. Every monomorphism is a kernel of its cokernel.
- The pullback of an epimorphism is an epimorphism. The pushout of a monomorphism is a monomorphism. (This is not to be confused with the fact that the pullback of a monomorphism is a monomorphism, which is true in any category).
Implementation notes
The typeclass abelian
does not extend non_preadditive_abelian
,
to avoid having to deal with comparing the two has_zero_morphisms
instances
(one from preadditive
in abelian
, and the other a field of non_preadditive_abelian
).
As a consequence, at the beginning of this file we trivially build
a non_preadditive_abelian
instance from an abelian
instance,
and use this to restate a number of theorems,
in each case just reusing the proof from non_preadditive_abelian.lean
.
We don't show this yet, but abelian categories are finitely complete and finitely cocomplete. However, the limits we can construct at this level of generality will most likely be less nice than the ones that can be created in specific applications. For this reason, we adopt the following convention:
- If the statement of a theorem involves limits, the existence of these limits should be made an explicit typeclass parameter.
- If a limit only appears in a proof, but not in the statement of a theorem, the limit should not
be a typeclass parameter, but instead be created using
abelian.has_pullbacks
or a similar definition.
References
- [F. Borceux, Handbook of Categorical Algebra 2][borceux-vol2]
- [P. Aluffi, Algebra: Chaper 0][aluffi2016]
- to_preadditive : category_theory.preadditive C
- has_finite_products : category_theory.limits.has_finite_products C
- has_kernels : category_theory.limits.has_kernels C
- has_cokernels : category_theory.limits.has_cokernels C
- normal_mono : Π {X Y : C} (f : X ⟶ Y) [_inst_2 : category_theory.mono f], category_theory.normal_mono f
- normal_epi : Π {X Y : C} (f : X ⟶ Y) [_inst_2 : category_theory.epi f], category_theory.normal_epi f
A (preadditive) category C
is called abelian if it has all finite products,
all kernels and cokernels, and if every monomorphism is the kernel of some morphism
and every epimorphism is the cokernel of some morphism.
(This definition implies the existence of zero objects: finite products give a terminal object, and in a preadditive category any terminal object is a zero object.)
An abelian category has finite biproducts.
Every abelian category is, in particular, non_preadditive_abelian
.
Equations
- category_theory.abelian.non_preadditive_abelian = {has_zero_object := category_theory.limits.has_zero_object_of_has_terminal_object (category_theory.limits.has_limits_of_shape_discrete C pempty), has_zero_morphisms := category_theory.preadditive.preadditive_has_zero_morphisms category_theory.abelian.to_preadditive, has_kernels := category_theory.abelian.has_kernels _inst_2, has_cokernels := category_theory.abelian.has_cokernels _inst_2, has_finite_products := category_theory.abelian.has_finite_products _inst_2, has_finite_coproducts := category_theory.limits.has_finite_coproducts_of_has_finite_biproducts C category_theory.abelian.has_finite_biproducts, normal_mono := category_theory.abelian.normal_mono _inst_2, normal_epi := category_theory.abelian.normal_epi _inst_2}
In an abelian category, every epimorphism is strong.
In an abelian category, a monomorphism which is also an epimorphism is an isomorphism.
The kernel of the cokernel of f
is called the image of f
.
The inclusion of the image into the codomain.
There is a canonical epimorphism p : P ⟶ image f
for every f
.
f
factors through its image via the canonical morphism p
.
The map p : P ⟶ image f
is an epimorphism
Equations
- _ = _
Equations
- _ = _
Factoring through the image is a strong epi-mono factorisation.
Equations
- category_theory.abelian.images.image_strong_epi_mono_factorisation f = category_theory.limits.strong_epi_mono_factorisation.mk {I := category_theory.abelian.images.image f, m := category_theory.abelian.images.image.ι f, m_mono := _, e := category_theory.abelian.images.factor_thru_image f, fac' := _}
The cokernel of the kernel of f
is called the coimage of f
.
The projection onto the coimage.
There is a canonical monomorphism i : coimage f ⟶ Q
.
f
factors through its coimage via the canonical morphism p
.
The canonical morphism i : coimage f ⟶ Q
is a monomorphism
Equations
- _ = _
Equations
- _ = _
Factoring through the coimage is a strong epi-mono factorisation.
Equations
- category_theory.abelian.coimages.coimage_strong_epi_mono_factorisation f = category_theory.limits.strong_epi_mono_factorisation.mk {I := category_theory.abelian.coimages.coimage f, m := category_theory.abelian.coimages.factor_thru_coimage f, m_mono := _, e := category_theory.abelian.coimages.coimage.π f, fac' := _}
An abelian category has strong epi-mono factorisations.
Equations
There is a canonical isomorphism between the coimage and the image of a morphism.
In an abelian category, an epi is the cokernel of its kernel. More precisely:
If f
is an epimorphism and s
is some limit kernel cone on f
, then f
is a cokernel
of fork.ι s
.
In an abelian category, a mono is the kernel of its cokernel. More precisely:
If f
is a monomorphism and s
is some colimit cokernel cocone on f
, then f
is a kernel
of cofork.π s
.
Any abelian category has pullbacks
Any abelian category has pushouts
This section contains a slightly technical result about pullbacks and biproducts. We will need it in the proof that the pullback of an epimorphism is an epimorpism.
The canonical map pullback f g ⟶ X ⊞ Y
The canonical map pullback f g ⟶ X ⊞ Y
induces a kernel cone on the map
biproduct X Y ⟶ Z
induced by f
and g
. A slightly more intuitive way to think of
this may be that it induces an equalizer fork on the maps induced by (f, 0)
and
(0, g)
.
The canonical map pullback f g ⟶ X ⊞ Y
is a kernel of the map induced by
(f, -g)
.
Equations
- category_theory.abelian.pullback_to_biproduct_is_kernel.is_limit_pullback_to_biproduct f g = category_theory.limits.fork.is_limit.mk (category_theory.abelian.pullback_to_biproduct_is_kernel.pullback_to_biproduct_fork f g) (λ (s : category_theory.limits.fork (category_theory.limits.biprod.desc f (-g)) 0), category_theory.limits.pullback.lift (s.ι ≫ category_theory.limits.biprod.fst) (s.ι ≫ category_theory.limits.biprod.snd) _) _ _
The canonical map Y ⊞ Z ⟶ pushout f g
The canonical map Y ⊞ Z ⟶ pushout f g
induces a cokernel cofork on the map
X ⟶ Y ⊞ Z
induced by f
and -g
.
The cofork induced by the canonical map Y ⊞ Z ⟶ pushout f g
is in fact a colimit cokernel
cofork.
Equations
- category_theory.abelian.biproduct_to_pushout_is_cokernel.is_colimit_biproduct_to_pushout f g = category_theory.limits.cofork.is_colimit.mk (category_theory.abelian.biproduct_to_pushout_is_cokernel.biproduct_to_pushout_cofork f g) (λ (s : category_theory.limits.cofork (category_theory.limits.biprod.lift f (-g)) 0), category_theory.limits.pushout.desc (category_theory.limits.biprod.inl ≫ s.π) (category_theory.limits.biprod.inr ≫ s.π) _) _ _
In an abelian category, the pullback of an epimorphism is an epimorphism. Proof from [aluffi2016, IX.2.3], cf. [borceux-vol2, 1.7.6]
Equations
- _ = _
In an abelian category, the pullback of an epimorphism is an epimorphism.
Equations
- _ = _
Equations
- _ = _
Equations
- _ = _
Every non_preadditive_abelian category can be promoted to an abelian category.
Equations
- category_theory.non_preadditive_abelian.abelian C = {to_preadditive := {hom_group := category_theory.preadditive.hom_group category_theory.non_preadditive_abelian.preadditive, add_comp' := _, comp_add' := _}, has_finite_products := infer_instance category_theory.non_preadditive_abelian.has_finite_products, has_kernels := _.mpr infer_instance, has_cokernels := _.mpr infer_instance, normal_mono := λ (X Y : C) (f : X ⟶ Y) (_inst_2_1 : category_theory.mono f), _.mpr (category_theory.non_preadditive_abelian.normal_mono f), normal_epi := λ (X Y : C) (f : X ⟶ Y) (_inst_2_1 : category_theory.epi f), _.mpr (category_theory.non_preadditive_abelian.normal_epi f)}