Every non_preadditive_abelian category is preadditive #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
In mathlib, we define an abelian category as a preadditive category with a zero object, kernels and cokernels, products and coproducts and in which every monomorphism and epimorphis is normal.
While virtually every interesting abelian category has a natural preadditive structure (which is why it is included in the definition), preadditivity is not actually needed: Every category that has all of the other properties appearing in the definition of an abelian category admits a preadditive structure. This is the construction we carry out in this file.
The proof proceeds in roughly five steps:
- Prove some results (for example that all equalizers exist) that would be trivial if we already had the preadditive structure but are a bit of work without it.
- Develop images and coimages to show that every monomorphism is the kernel of its cokernel.
The results of the first two steps are also useful for the "normal" development of abelian categories, and will be used there.
- For every object
A
, define a "subtraction" morphismσ : A ⨯ A ⟶ A
and use it to define subtraction on morphisms asf - g := prod.lift f g ≫ σ
. - Prove a small number of identities about this subtraction from the definition of
σ
. - From these identities, prove a large number of other identities that imply that defining
f + g := f - (0 - g)
indeed gives an abelian group structure on morphisms such that composition is bilinear.
The construction is non-trivial and it is quite remarkable that this abelian group structure can
be constructed purely from the existence of a few limits and colimits. Even more remarkably,
since abelian categories admit exactly one preadditive structure (see
subsingleton_preadditive_of_has_binary_biproducts
), the construction manages to exactly
reconstruct any natural preadditive structure the category may have.
References #
- to_has_zero_morphisms : category_theory.limits.has_zero_morphisms C
- to_normal_mono_category : category_theory.normal_mono_category C
- to_normal_epi_category : category_theory.normal_epi_category C
- has_zero_object : category_theory.limits.has_zero_object C
- has_kernels : category_theory.limits.has_kernels C
- has_cokernels : category_theory.limits.has_cokernels C
- has_finite_products : category_theory.limits.has_finite_products C
- has_finite_coproducts : category_theory.limits.has_finite_coproducts C
We call a category non_preadditive_abelian
if it has a zero object, kernels, cokernels, binary
products and coproducts, and every monomorphism and every epimorphism is normal.
Instances for category_theory.non_preadditive_abelian
- category_theory.non_preadditive_abelian.has_sizeof_inst
The map p : P ⟶ image f
is an epimorphism
The canonical morphism i : coimage f ⟶ Q
is a monomorphism
In a non_preadditive_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
.
Equations
- category_theory.non_preadditive_abelian.epi_is_cokernel_of_kernel s h = category_theory.limits.is_cokernel.cokernel_iso s.ι f (category_theory.limits.cokernel.of_iso_comp ((category_theory.limits.limit.cone (category_theory.limits.parallel_pair f 0)).π.app category_theory.limits.walking_parallel_pair.zero) s.ι ((category_theory.limits.limit.is_limit (category_theory.limits.parallel_pair f 0)).cone_point_unique_up_to_iso h) _) (category_theory.as_iso (category_theory.abelian.factor_thru_coimage f)) _
In a non_preadditive_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
.
Equations
- category_theory.non_preadditive_abelian.mono_is_kernel_of_cokernel s h = category_theory.limits.is_kernel.iso_kernel s.π f (category_theory.limits.kernel.of_comp_iso ((category_theory.limits.colimit.cocone (category_theory.limits.parallel_pair f 0)).ι.app category_theory.limits.walking_parallel_pair.one) s.π (h.cocone_point_unique_up_to_iso (category_theory.limits.colimit.is_colimit (category_theory.limits.parallel_pair f 0))) _) (category_theory.as_iso (category_theory.abelian.factor_thru_image f)) _
The composite A ⟶ A ⨯ A ⟶ cokernel (Δ A)
, where the first map is (𝟙 A, 0)
and the second map
is the canonical projection into the cokernel.
The composite A ⨯ A ⟶ cokernel (diag A) ⟶ A
given by the natural projection into the cokernel
followed by the inverse of r
. In the category of modules, using the normal kernels and
cokernels, this map is equal to the map (a, b) ↦ a - b
, hence the name σ
for
"subtraction".
σ is a cokernel of Δ X.
Equations
- category_theory.non_preadditive_abelian.is_colimit_σ = category_theory.limits.cokernel.cokernel_iso (category_theory.limits.diag X) category_theory.non_preadditive_abelian.σ (category_theory.as_iso (category_theory.non_preadditive_abelian.r X)).symm category_theory.non_preadditive_abelian.is_colimit_σ._proof_1
This is the key identity satisfied by σ
.
Subtraction of morphisms in a non_preadditive_abelian
category.
Equations
Negation of morphisms in a non_preadditive_abelian
category.
Addition of morphisms in a non_preadditive_abelian
category.
Every non_preadditive_abelian
category is preadditive.
Equations
- category_theory.non_preadditive_abelian.preadditive = {hom_group := λ (X Y : C), {add := has_add.add category_theory.non_preadditive_abelian.has_add, add_assoc := _, zero := 0, zero_add := _, add_zero := _, nsmul := add_group.nsmul._default 0 has_add.add category_theory.non_preadditive_abelian.neg_neg category_theory.non_preadditive_abelian.add_zero, nsmul_zero' := _, nsmul_succ' := _, neg := λ (f : X ⟶ Y), -f, sub := add_group.sub._default has_add.add category_theory.non_preadditive_abelian.add_assoc 0 category_theory.non_preadditive_abelian.neg_neg category_theory.non_preadditive_abelian.add_zero (add_group.nsmul._default 0 has_add.add category_theory.non_preadditive_abelian.neg_neg category_theory.non_preadditive_abelian.add_zero) _ _ (λ (f : X ⟶ Y), -f), sub_eq_add_neg := _, zsmul := add_group.zsmul._default has_add.add category_theory.non_preadditive_abelian.add_assoc 0 category_theory.non_preadditive_abelian.neg_neg category_theory.non_preadditive_abelian.add_zero (add_group.nsmul._default 0 has_add.add category_theory.non_preadditive_abelian.neg_neg category_theory.non_preadditive_abelian.add_zero) _ _ (λ (f : X ⟶ Y), -f), zsmul_zero' := _, zsmul_succ' := _, zsmul_neg' := _, add_left_neg := _, add_comm := _}, add_comp' := _, comp_add' := _}