Documentation

Mathlib.CategoryTheory.Preadditive.Basic

Preadditive categories #

A preadditive category is a category in which X ⟶ Y is an abelian group in such a way that composition of morphisms is linear in both variables.

This file contains a definition of preadditive category that directly encodes the definition given above. The definition could also be phrased as follows: A preadditive category is a category enriched over the category of Abelian groups. Once the general framework to state this in Lean is available, the contents of this file should become obsolete.

Main results #

Implementation notes #

The simp normal form for negation and composition is to push negations as far as possible to the outside. For example, f ≫ (-g) and (-f) ≫ g both become -(f ≫ g), and (-f) ≫ (-g) is simplified to f ≫ g.

References #

Tags #

additive, preadditive, Hom group, Ab-category, Ab-enriched

A category is called preadditive if P ⟶ Q is an abelian group such that composition is linear in both variables.

Instances
    theorem CategoryTheory.Preadditive.add_comp_assoc {C : Type u} {inst✝ : Category.{v, u} C} [self : Preadditive C] (P Q R : C) (f f' : P Q) (g : Q R) {Z : C} (h : R Z) :
    theorem CategoryTheory.Preadditive.comp_add_assoc {C : Type u} {inst✝ : Category.{v, u} C} [self : Preadditive C] (P Q R : C) (f : P Q) (g g' : Q R) {Z : C} (h : R Z) :
    Equations
    • One or more equations did not get rendered due to their size.
    def CategoryTheory.InducedCategory.homAddEquiv {C : Type u} [Category.{v, u} C] [Preadditive C] {D : Type u'} {F : DC} {X Y : InducedCategory C F} :
    (X Y) ≃+ (F X F Y)

    The additive equivalence (X ⟶ Y) ≃+ (F X ⟶ F Y) when F : D → C and C is a preadditive category.

    Equations
    Instances For
      @[simp]
      theorem CategoryTheory.InducedCategory.homAddEquiv_apply {C : Type u} [Category.{v, u} C] [Preadditive C] {D : Type u'} {F : DC} {X Y : InducedCategory C F} (f : X Y) :
      @[simp]
      theorem CategoryTheory.InducedCategory.homAddEquiv_symm_apply_hom {C : Type u} [Category.{v, u} C] [Preadditive C] {D : Type u'} {F : DC} {X Y : InducedCategory C F} (f : F X F Y) :
      def CategoryTheory.Preadditive.leftComp {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q : C} (R : C) (f : P Q) :
      (Q R) →+ (P R)

      Composition by a fixed left argument as a group homomorphism

      Equations
      Instances For
        def CategoryTheory.Preadditive.rightComp {C : Type u} [Category.{v, u} C] [Preadditive C] (P : C) {Q R : C} (g : Q R) :
        (P Q) →+ (P R)

        Composition by a fixed right argument as a group homomorphism

        Equations
        Instances For
          def CategoryTheory.Preadditive.compHom {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} :
          (P Q) →+ (Q R) →+ (P R)

          Composition as a bilinear group homomorphism

          Equations
          Instances For
            @[simp]
            theorem CategoryTheory.Preadditive.sub_comp {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} (f f' : P Q) (g : Q R) :
            @[simp]
            theorem CategoryTheory.Preadditive.comp_sub {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} (f : P Q) (g g' : Q R) :
            @[simp]
            theorem CategoryTheory.Preadditive.neg_comp {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} (f : P Q) (g : Q R) :
            @[simp]
            theorem CategoryTheory.Preadditive.comp_neg {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} (f : P Q) (g : Q R) :
            theorem CategoryTheory.Preadditive.nsmul_comp {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} (f : P Q) (g : Q R) (n : ) :
            theorem CategoryTheory.Preadditive.comp_nsmul {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} (f : P Q) (g : Q R) (n : ) :
            theorem CategoryTheory.Preadditive.zsmul_comp {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} (f : P Q) (g : Q R) (n : ) :
            theorem CategoryTheory.Preadditive.comp_zsmul {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} (f : P Q) (g : Q R) (n : ) :
            theorem CategoryTheory.Preadditive.comp_sum {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} {J : Type u_1} (s : Finset J) (f : P Q) (g : J → (Q R)) :
            CategoryStruct.comp f (∑ js, g j) = js, CategoryStruct.comp f (g j)
            theorem CategoryTheory.Preadditive.comp_sum_assoc {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} {J : Type u_1} (s : Finset J) (f : P Q) (g : J → (Q R)) {Z : C} (h : R Z) :
            CategoryStruct.comp f (CategoryStruct.comp (∑ js, g j) h) = CategoryStruct.comp (∑ js, CategoryStruct.comp f (g j)) h
            theorem CategoryTheory.Preadditive.sum_comp {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} {J : Type u_1} (s : Finset J) (f : J → (P Q)) (g : Q R) :
            CategoryStruct.comp (∑ js, f j) g = js, CategoryStruct.comp (f j) g
            theorem CategoryTheory.Preadditive.sum_comp_assoc {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} {J : Type u_1} (s : Finset J) (f : J → (P Q)) (g : Q R) {Z : C} (h : R Z) :
            CategoryStruct.comp (∑ js, f j) (CategoryStruct.comp g h) = CategoryStruct.comp (∑ js, CategoryStruct.comp (f j) g) h
            theorem CategoryTheory.Preadditive.sum_comp' {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R S : C} {J : Type u_1} (s : Finset J) (f : J → (P Q)) (g : J → (Q R)) (h : R S) :
            CategoryStruct.comp (∑ js, CategoryStruct.comp (f j) (g j)) h = js, CategoryStruct.comp (f j) (CategoryStruct.comp (g j) h)
            theorem CategoryTheory.Preadditive.sum_comp'_assoc {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R S : C} {J : Type u_1} (s : Finset J) (f : J → (P Q)) (g : J → (Q R)) (h : R S) {Z : C} (h✝ : S Z) :
            CategoryStruct.comp (∑ js, CategoryStruct.comp (f j) (g j)) (CategoryStruct.comp h h✝) = CategoryStruct.comp (∑ js, CategoryStruct.comp (f j) (CategoryStruct.comp (g j) h)) h✝
            instance CategoryTheory.Preadditive.instEpiNegHom {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q : C} {f : P Q} [Epi f] :
            Epi (-f)
            instance CategoryTheory.Preadditive.instMonoNegHom {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q : C} {f : P Q} [Mono f] :
            Mono (-f)

            This instance is split off from the Ring (End X) instance to speed up instance search.

            Equations
            • One or more equations did not get rendered due to their size.
            Equations
            • One or more equations did not get rendered due to their size.
            Equations
            theorem CategoryTheory.Preadditive.mono_of_cancel_zero {C : Type u} [Category.{v, u} C] [Preadditive C] {Q R : C} (f : Q R) (h : ∀ {P : C} (g : P Q), CategoryStruct.comp g f = 0g = 0) :
            theorem CategoryTheory.Preadditive.mono_iff_cancel_zero {C : Type u} [Category.{v, u} C] [Preadditive C] {Q R : C} (f : Q R) :
            Mono f ∀ (P : C) (g : P Q), CategoryStruct.comp g f = 0g = 0
            theorem CategoryTheory.Preadditive.epi_of_cancel_zero {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q : C} (f : P Q) (h : ∀ {R : C} (g : Q R), CategoryStruct.comp f g = 0g = 0) :
            Epi f
            theorem CategoryTheory.Preadditive.epi_iff_cancel_zero {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q : C} (f : P Q) :
            Epi f ∀ (R : C) (g : Q R), CategoryStruct.comp f g = 0g = 0
            @[simp]
            theorem CategoryTheory.Preadditive.IsIso.comp_left_eq_zero {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} (f : P Q) (g : Q R) [IsIso f] :
            @[simp]
            theorem CategoryTheory.Preadditive.IsIso.comp_right_eq_zero {C : Type u} [Category.{v, u} C] [Preadditive C] {P Q R : C} (f : P Q) (g : Q R) [IsIso g] :

            Map a kernel cone on the difference of two morphisms to the equalizer fork.

            Equations
            Instances For
              @[simp]

              Map any equalizer fork to a cone on the difference of the two morphisms.

              Equations
              Instances For

                A kernel of f - g is an equalizer of f and g.

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

                  An equalizer of f and g is a kernel of f - g.

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

                    A preadditive category has an equalizer for f and g if it has a kernel for f - g.

                    A preadditive category has a kernel for f - g if it has an equalizer for f and g.

                    Map a cokernel cocone on the difference of two morphisms to the coequalizer cofork.

                    Equations
                    Instances For

                      Map any coequalizer cofork to a cocone on the difference of the two morphisms.

                      Equations
                      Instances For

                        A cokernel of f - g is a coequalizer of f and g.

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

                          A coequalizer of f and g is a cokernel of f - g.

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

                            A preadditive category has a coequalizer for f and g if it has a cokernel for f - g.

                            A preadditive category has a cokernel for f - g if it has a coequalizer for f and g.

                            If a preadditive category has all kernels, then it also has all equalizers.

                            If a preadditive category has all cokernels, then it also has all coequalizers.

                            Equations
                            @[simp]
                            theorem CategoryTheory.Preadditive.smul_iso_hom {C : Type u_1} [Category.{v_1, u_1} C] [Preadditive C] {X Y : C} (a : ˣ) (e : X Y) :
                            (a e).hom = a e.hom
                            @[simp]
                            theorem CategoryTheory.Preadditive.smul_iso_inv {C : Type u_1} [Category.{v_1, u_1} C] [Preadditive C] {X Y : C} (a : ˣ) (e : X Y) :
                            (a e).inv = a⁻¹ e.inv
                            Equations
                            @[simp]
                            theorem CategoryTheory.Preadditive.neg_iso_hom {C : Type u_1} [Category.{v_1, u_1} C] [Preadditive C] {X Y : C} (e : X Y) :
                            (-e).hom = -e.hom
                            @[simp]
                            theorem CategoryTheory.Preadditive.neg_iso_inv {C : Type u_1} [Category.{v_1, u_1} C] [Preadditive C] {X Y : C} (e : X Y) :
                            (-e).inv = -e.inv