Documentation

Mathlib.CategoryTheory.Abelian.Basic

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 NonPreadditiveAbelian.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 "native" preadditive structure for the specific category they are working with.

Main definitions #

Main results #

Implementation notes #

The typeclass Abelian does not extend NonPreadditiveAbelian, to avoid having to deal with comparing the two HasZeroMorphisms instances (one from Preadditive in Abelian, and the other a field of NonPreadditiveAbelian). As a consequence, at the beginning of this file we trivially build a NonPreadditiveAbelian instance from an Abelian instance, and use this to restate a number of theorems, in each case just reusing the proof from NonPreadditiveAbelian.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:

References #

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.)

Instances

    We begin by providing an alternative constructor: a preadditive category with kernels, cokernels, and finite products, in which the coimage-image comparison morphism is always an isomorphism, is an abelian category.

    A preadditive category with kernels, cokernels, and finite products, in which the coimage-image comparison morphism is always an isomorphism, is an abelian category.

    The Stacks project uses this characterisation at the definition of an abelian category. See https://stacks.math.columbia.edu/tag/0109.

    Instances For

      Every abelian category is, in particular, NonPreadditiveAbelian.

      Instances For

        We now promote some instances that were constructed using non_preadditive_abelian.

        Factoring through the image is a strong epi-mono factorisation.

        Instances For

          Factoring through the coimage is a strong epi-mono factorisation.

          Instances For
            @[inline, reducible]

            There is a canonical isomorphism between the abelian coimage and the abelian image of a morphism.

            Instances For
              @[inline, reducible]

              There is a canonical isomorphism between the abelian coimage and the categorical image of a morphism.

              Instances For
                @[inline, reducible]

                There is a canonical isomorphism between the abelian image and the categorical image of a morphism.

                Instances For

                  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.

                  Instances For

                    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.

                    Instances For

                      In an abelian category, any morphism that turns to zero when precomposed with the kernel of an epimorphism factors through that epimorphism.

                      Instances For

                        In an abelian category, any morphism that turns to zero when postcomposed with the cokernel of a monomorphism factors through that monomorphism.

                        Instances For

                          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 epimorphism.

                          @[inline, reducible]

                          The canonical map pullback f g ⟶ X ⊞ Y

                          Instances For
                            @[inline, reducible]

                            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).

                            Instances For
                              @[inline, reducible]

                              The canonical map Y ⊞ Z ⟶ pushout f g

                              Instances For
                                @[inline, reducible]

                                The canonical map Y ⊞ Z ⟶ pushout f g induces a cokernel cofork on the map X ⟶ Y ⊞ Z induced by f and -g.

                                Instances For

                                  The cofork induced by the canonical map Y ⊞ Z ⟶ pushout f g is in fact a colimit cokernel cofork.

                                  Instances For
                                    instance CategoryTheory.Abelian.epi_pullback_of_epi_f {C : Type u} [CategoryTheory.Category.{v, u} C] [CategoryTheory.Abelian C] [CategoryTheory.Limits.HasPullbacks C] {X : C} {Y : C} {Z : C} (f : X Z) (g : Y Z) [CategoryTheory.Epi f] :
                                    CategoryTheory.Epi CategoryTheory.Limits.pullback.snd

                                    In an abelian category, the pullback of an epimorphism is an epimorphism. Proof from [aluffi2016, IX.2.3], cf. [borceux-vol2, 1.7.6]

                                    instance CategoryTheory.Abelian.epi_pullback_of_epi_g {C : Type u} [CategoryTheory.Category.{v, u} C] [CategoryTheory.Abelian C] [CategoryTheory.Limits.HasPullbacks C] {X : C} {Y : C} {Z : C} (f : X Z) (g : Y Z) [CategoryTheory.Epi g] :
                                    CategoryTheory.Epi CategoryTheory.Limits.pullback.fst

                                    In an abelian category, the pullback of an epimorphism is an epimorphism.

                                    Suppose f and g are two morphisms with a common codomain and suppose we have written g as an epimorphism followed by a monomorphism. If f factors through the mono part of this factorization, then any pullback of g along f is an epimorphism.

                                    Suppose f and g are two morphisms with a common domain and suppose we have written g as an epimorphism followed by a monomorphism. If f factors through the epi part of this factorization, then any pushout of g along f is a monomorphism.

                                    Every NonPreadditiveAbelian category can be promoted to an abelian category.

                                    Instances For