module Cat.Diagram.Biproduct wheremodule _ {o β} (C : Precategory o β) where
  open Cat.Reasoning CBiproductsπ
Recall that, in an -enriched category, products and coproducts automatically coincide: these are called biproducts and are written
Following (Karvonen 2020), we can define what it means to be a biproduct in a general category: we say that a diagram like the one above is a biproduct diagram if is a product diagram, is a coproduct diagram, and the following equations relating the product projections and coproduct injections hold:
  record is-biproduct {A B P}
    (Οβ : Hom P A) (Οβ : Hom P B)
    (ΞΉβ : Hom A P) (ΞΉβ : Hom B P)
    : Type (o β β) where
    field
      has-is-product   : is-product C Οβ Οβ
      has-is-coproduct : is-coproduct C ΞΉβ ΞΉβ
      ΟΞΉβ : Οβ β ΞΉβ β‘ id
      ΟΞΉβ : Οβ β ΞΉβ β‘ id
      ΞΉΟ-comm : ΞΉβ β Οβ β ΞΉβ β Οβ β‘ ΞΉβ β Οβ β ΞΉβ β Οβ
  record Biproduct (A B : Ob) : Type (o β β) where
    field
      biapex : Ob
      Οβ : Hom biapex A
      Οβ : Hom biapex B
      ΞΉβ : Hom A biapex
      ΞΉβ : Hom B biapex
      has-is-biproduct : is-biproduct Οβ Οβ ΞΉβ ΞΉβ    open is-biproduct has-is-biproduct public
    product : Product C A B
    product = record
      { apex = biapex ; Οβ = Οβ ; Οβ = Οβ ; has-is-product = has-is-product }
    coproduct : Coproduct C A B
    coproduct = record
      { coapex = biapex ; ΞΉβ = ΞΉβ ; ΞΉβ = ΞΉβ ; has-is-coproduct = has-is-coproduct }
    open Product product public
      hiding (Οβ; Οβ)
      renaming (uniqueβ to β¨β©-uniqueβ)
    open Coproduct coproduct public
      hiding (ΞΉβ; ΞΉβ)
      renaming (uniqueβ to []-uniqueβ)Semiadditive categoriesπ
Just like terminal objects (resp. initial objects) are 0-ary products (resp. coproducts), zero objects are 0-ary biproducts. A category with a zero object and binary biproducts (hence all finite biproducts) is called semiadditive.
  record is-semiadditive : Type (o β β) where
    field
      has-zero : Zero C
      has-biproducts : β {A B} β Biproduct A BAs the name hints, every additive category is semiadditive. However, quite surprisingly, it turns out that the structure1 of a semiadditive category is already enough to define an enrichment of in commutative monoids!
    open Zero has-zero public
    module Biprod {A B} = Biproduct (has-biproducts {A} {B})
    open Biprod using (ΟΞΉβ; ΟΞΉβ; ΞΉΟ-comm)
    open Binary-products C (Ξ» _ _ β Biprod.product)
    open Binary-coproducts C (Ξ» _ _ β Biprod.coproduct)
    cartesian : Cartesian-category C
    cartesian = record { products = Ξ» _ _ β Biprod.product ; terminal = terminal }
    open Monoidal-category (Cartesian-monoidal cartesian) using (associator; module β)We define the addition of morphisms by the following diagram, where (resp. is the diagonal map (resp. codiagonal map).
    _+β_ : β {x y} β Hom x y β Hom x y β Hom x y
    f +β g = β β (f ββ g) β Ξ΄We start by noticing a few properties of finite biproducts. First, while we are of course justified in writing without ambiguity for objects, we must prove that so that we can write without ambiguity for morphisms.
To that end, we start by showing that
and
are zero morphisms.
    Οβ-ΞΉβ : β {a b} β Οβ {a} {b} β ΞΉβ β‘ zeroβ
    Οβ-ΞΉβ : β {a b} β Οβ {a} {b} β ΞΉβ β‘ zeroβ
The proofs follow (Karvonen
2020) and are unenlightening.
    Οβ-ΞΉβ : β {a b} β Οβ {a} {b} β ΞΉβ β‘ zeroβ
    Οβ-ΞΉβ : β {a b} β Οβ {a} {b} β ΞΉβ β‘ zeroβ    Οβ-ΞΉβ = zero-unique Ξ» f g β
      let h = β¨ Οβ β ΞΉβ β g , f β© in
      (Οβ β ΞΉβ) β f                   β‘β¨ insertl ΟΞΉβ β©
      Οβ β ΞΉβ β (Οβ β ΞΉβ) β f         β‘Λβ¨ reflβ©ββ¨ reflβ©ββ¨ assoc _ _ _ β (reflβ©ββ¨ Οβββ¨β©) β©
      Οβ β ΞΉβ β Οβ β ΞΉβ β Οβ β h      β‘β¨ reflβ©ββ¨ pulll4 ΞΉΟ-comm β©
      Οβ β (ΞΉβ β Οβ β ΞΉβ β Οβ) β h    β‘β¨ reflβ©ββ¨ pullr (pullr (pullr Οβββ¨β©)) β©
      Οβ β ΞΉβ β Οβ β ΞΉβ β Οβ β ΞΉβ β g β‘Λβ¨ reflβ©ββ¨ extendl4 ΞΉΟ-comm β©
      Οβ β ΞΉβ β Οβ β ΞΉβ β Οβ β ΞΉβ β g β‘β¨ cancell ΟΞΉβ β©
      Οβ β ΞΉβ β Οβ β ΞΉβ β g           β‘β¨ pushr (reflβ©ββ¨ cancell ΟΞΉβ) β©
      (Οβ β ΞΉβ) β g                   β
    Οβ-ΞΉβ = zero-unique Ξ» f g β
      let h = β¨ f , Οβ β ΞΉβ β g β© in
      (Οβ β ΞΉβ) β f                   β‘β¨ insertl ΟΞΉβ β©
      Οβ β ΞΉβ β (Οβ β ΞΉβ) β f         β‘Λβ¨ reflβ©ββ¨ reflβ©ββ¨ assoc _ _ _ β (reflβ©ββ¨ Οβββ¨β©) β©
      Οβ β ΞΉβ β Οβ β ΞΉβ β Οβ β h      β‘Λβ¨ reflβ©ββ¨ pushl4 ΞΉΟ-comm β©
      Οβ β (ΞΉβ β Οβ β ΞΉβ β Οβ) β h    β‘β¨ reflβ©ββ¨ pullr (pullr (pullr Οβββ¨β©)) β©
      Οβ β ΞΉβ β Οβ β ΞΉβ β Οβ β ΞΉβ β g β‘β¨ reflβ©ββ¨ extendl4 ΞΉΟ-comm β©
      Οβ β ΞΉβ β Οβ β ΞΉβ β Οβ β ΞΉβ β g β‘β¨ cancell ΟΞΉβ β©
      Οβ β ΞΉβ β Οβ β ΞΉβ β g           β‘β¨ pushr (reflβ©ββ¨ cancell ΟΞΉβ) β©
      (Οβ β ΞΉβ) β g                   βNext, we note that maps are uniquely determined, thanks to the universal properties of the product and coproduct, by the four components for In other words, admits a unique matrix representation
    unique-matrix
      : β {a b c d} {f g : Hom (a ββ b) (c ββ d)}
      β (Οβ β f β ΞΉβ β‘ Οβ β g β ΞΉβ)
      β (Οβ β f β ΞΉβ β‘ Οβ β g β ΞΉβ)
      β (Οβ β f β ΞΉβ β‘ Οβ β g β ΞΉβ)
      β (Οβ β f β ΞΉβ β‘ Οβ β g β ΞΉβ)
      β f β‘ g
    unique-matrix pββ pββ pββ pββ = Biprod.[]-uniqueβ
      (Biprod.β¨β©-uniqueβ pββ pββ refl refl)
      (Biprod.β¨β©-uniqueβ pββ pββ refl refl)
      refl reflThis implies that and are equal, as they both have the same matrix representation:
    βββ‘ββ : β {a b c d} {f : Hom a b} {g : Hom c d} β f ββ g β‘ f ββ g
    βββ‘ββ = unique-matrix
      ((reflβ©ββ¨ []βΞΉβ) β cancell ΟΞΉβ β sym (pulll Οβββ¨β© β cancelr ΟΞΉβ))
      ((reflβ©ββ¨ []βΞΉβ) β pulll Οβ-ΞΉβ β zero-βr _ β sym (pulll Οβββ¨β© β pullr Οβ-ΞΉβ β zero-βl _))
      ((reflβ©ββ¨ []βΞΉβ) β pulll Οβ-ΞΉβ β zero-βr _ β sym (pulll Οβββ¨β© β pullr Οβ-ΞΉβ β zero-βl _))
      ((reflβ©ββ¨ []βΞΉβ) β cancell ΟΞΉβ β sym (pulll Οβββ¨β© β cancelr ΟΞΉβ))
Similarly, we show that the associators and braidings for products
and coproducts coincide.
    coassocβ‘assoc : β {a b c} β β-assoc {a} {b} {c} β‘ Γ-assoc
    coswapβ‘swap : β {a b} β coswap {a} {b} β‘ swap
    coassocβ‘assoc : β {a b c} β β-assoc {a} {b} {c} β‘ Γ-assoc
    coswapβ‘swap : β {a b} β coswap {a} {b} β‘ swap    coassocβ‘assoc = unique-matrix
      ((reflβ©ββ¨ []βΞΉβ) β cancell ΟΞΉβ β sym (pulll Οβββ¨β© β Biprod.β¨β©-uniqueβ
        (pulll Οβββ¨β© β ΟΞΉβ)
        (pulll Οβββ¨β© β pullr Οβ-ΞΉβ β zero-βl _)
        ΟΞΉβ Οβ-ΞΉβ))
      ((reflβ©ββ¨ []βΞΉβ) β pulll Οβ-ΞΉβ β zero-βr _ β sym (pulll Οβββ¨β© β pullr Οβ-ΞΉβ β zero-βl _))
      ((reflβ©ββ¨ []βΞΉβ) β unique-matrix
        ((reflβ©ββ¨ pullr []βΞΉβ β cancell ΟΞΉβ) β Οβ-ΞΉβ β sym (pulll (pulll Οβββ¨β©) β (Οβ-ΞΉβ β©ββ¨refl) β zero-βr _))
        ((reflβ©ββ¨ pullr []βΞΉβ β cancell ΟΞΉβ) β ΟΞΉβ β sym (pulll (pulll Οβββ¨β©) β (cancelr ΟΞΉβ β©ββ¨refl) β ΟΞΉβ))
        ((reflβ©ββ¨ pullr []βΞΉβ β Οβ-ΞΉβ) β zero-βl _ β sym (pulll (pulll Οβββ¨β©) β (Οβ-ΞΉβ β©ββ¨refl) β zero-βr _))
        ((reflβ©ββ¨ pullr []βΞΉβ β Οβ-ΞΉβ) β zero-βl _ β sym (pulll (pulll Οβββ¨β©) β (cancelr ΟΞΉβ β©ββ¨refl) β Οβ-ΞΉβ))
      β sym (pulll Οβββ¨β©))
      ((reflβ©ββ¨ []βΞΉβ) β Biprod.[]-uniqueβ
        (pullr []βΞΉβ β pulll Οβ-ΞΉβ β zero-βr _)
        (pullr []βΞΉβ β ΟΞΉβ)
        ((cancelr ΟΞΉβ β©ββ¨refl) β Οβ-ΞΉβ)
        ((cancelr ΟΞΉβ β©ββ¨refl) β ΟΞΉβ)
      β sym (pulll Οβββ¨β©))
    coswapβ‘swap = β¨β©-unique
      (Biprod.[]-uniqueβ
        (pullr []βΞΉβ β Οβ-ΞΉβ) (pullr []βΞΉβ β ΟΞΉβ)
        Οβ-ΞΉβ ΟΞΉβ)
      (Biprod.[]-uniqueβ
        (pullr []βΞΉβ β ΟΞΉβ) (pullr []βΞΉβ β Οβ-ΞΉβ)
        ΟΞΉβ Οβ-ΞΉβ)We are finally ready to show that addition of morphisms is associative, commutative and unital. These properties essentially follow from the corresponding properties of biproducts. For associativity, we use the following diagram:
    +-assoc : β {x y} {f g h : Hom x y} β f +β (g +β h) β‘ (f +β g) +β h
    +-assoc {f = f} {g} {h} =
      β β (f ββ (β β (g ββ h) β Ξ΄)) β Ξ΄                           β‘Λβ¨ reflβ©ββ¨ β.pulll3 (idl _ β idr _ ,β refl) β©
      β β (id ββ β) β (f ββ (g ββ h)) β (id ββ Ξ΄) β Ξ΄             β‘Λβ¨ reflβ©ββ¨ βββ‘ββ β©ββ¨refl β©
      β β (id ββ β) β (f ββ (g ββ h)) β (id ββ Ξ΄) β Ξ΄             β‘Λβ¨ pushl β-assoc β©
      (β β (β ββ id) β β-assoc) β (f ββ (g ββ h)) β (id ββ Ξ΄) β Ξ΄ β‘β¨ (reflβ©ββ¨ βββ‘ββ β©ββ¨refl) β©ββ¨refl β©
      (β β (β ββ id) β β-assoc) β (f ββ (g ββ h)) β (id ββ Ξ΄) β Ξ΄ β‘β¨ pullr (pullr (coassocβ‘assoc β©ββ¨refl)) β©
      β β (β ββ id) β Γ-assoc β (f ββ (g ββ h)) β (id ββ Ξ΄) β Ξ΄   β‘β¨ reflβ©ββ¨ reflβ©ββ¨ extendl (associator .IsoβΏ.from .is-natural _ _ _) β©
      β β (β ββ id) β ((f ββ g) ββ h) β Γ-assoc β (id ββ Ξ΄) β Ξ΄   β‘β¨ reflβ©ββ¨ reflβ©ββ¨ reflβ©ββ¨ assoc-Ξ΄ β©
      β β (β ββ id) β ((f ββ g) ββ h) β (Ξ΄ ββ id) β Ξ΄             β‘β¨ reflβ©ββ¨ β.pulll3 (refl ,β idl _ β idr _) β©
      β β ((β β (f ββ g) β Ξ΄) ββ h) β Ξ΄                           βCommutativity follows from the following diagram, where is the braiding:
    +-comm : β {x y} {f g : Hom x y} β f +β g β‘ g +β f
    +-comm {f = f} {g} =
      β β (f ββ g) β Ξ΄          β‘Λβ¨ pulll β-coswap β©
      β β coswap β (f ββ g) β Ξ΄ β‘β¨ reflβ©ββ¨ coswapβ‘swap β©ββ¨refl β©
      β β swap β (f ββ g) β Ξ΄   β‘Λβ¨ reflβ©ββ¨ extendl (swap-natural _) β©
      β β (g ββ f) β swap β Ξ΄   β‘β¨ reflβ©ββ¨ reflβ©ββ¨ swap-Ξ΄ β©
      β β (g ββ f) β Ξ΄          βSince addition is commutative, it suffices to show that the zero morphism is a left unit. We again use the analogous property of biproducts, which is that the zero object is a left unit for
    β-Β‘l : β {a} β β {a} β (Β‘ ββ id) β‘ Οβ
    β-Β‘l = Biprod.[]-uniqueβ
      (Β‘-uniqueβ _ _) (pullr []βΞΉβ β cancell []βΞΉβ)
      refl ΟΞΉβ
    +-idl : β {x y} {f : Hom x y} β zeroβ +β f β‘ f
    +-idl {f = f} =
      β β (zeroβ ββ f) β Ξ΄         β‘β¨ reflβ©ββ¨ β.pushl (refl ,β sym (idl _)) β©
      β β (Β‘ ββ id) β (! ββ f) β Ξ΄ β‘Λβ¨ reflβ©ββ¨ βββ‘ββ β©ββ¨refl β©
      β β (Β‘ ββ id) β (! ββ f) β Ξ΄ β‘β¨ pulll β-Β‘l β©
      Οβ β (! ββ f) β Ξ΄            β‘β¨ pulll Οβββ¨β© β©
      (f β Οβ) β Ξ΄                 β‘β¨ cancelr Οβββ¨β© β©
      f                            β
    +-idr : β {x y} {f : Hom x y} β f +β zeroβ β‘ f
    +-idr = +-comm β +-idlNaturality of the (co)diagonals implies that composition is bilinear with respect to our defined addition.
    β-linear-l
      : β {a b c} {f g : Hom b c} {h : Hom a b}
      β f β h +β g β h β‘ (f +β g) β h
    β-linear-l {f = f} {g} {h} =
      β β ((f β h) ββ (g β h)) β Ξ΄ β‘β¨ reflβ©ββ¨ β.pushl refl β©
      β β (f ββ g) β (h ββ h) β Ξ΄  β‘Λβ¨ pullr (pullr (Ξ΄-natural _ _ _)) β©
      (β β (f ββ g) β Ξ΄) β h       β
    β-linear-r
      : β {a b c} {f g : Hom a b} {h : Hom b c}
      β h β f +β h β g β‘ h β (f +β g)
    β-linear-r {f = f} {g} {h} =
      β β ((h β f) ββ (h β g)) β Ξ΄ β‘β¨ reflβ©ββ¨ β.pushl refl β©
      β β (h ββ h) β (f ββ g) β Ξ΄  β‘Λβ¨ reflβ©ββ¨ βββ‘ββ β©ββ¨refl β©
      β β (h ββ h) β (f ββ g) β Ξ΄  β‘β¨ extendl (β-natural _ _ _) β©
      h β β β (f ββ g) β Ξ΄         βThis provides an alternative route to defining additive categories: instead of starting with an -enriched category and requiring finite (co)products, we can start with a semiadditive category and require that every be a group.
In order to construct a semiadditive category from a category with a zero object, binary products and coproducts, it suffices to require that the map defined by the matrix representation
is invertible.
  record make-semiadditive : Type (o β β) where
    field
      has-zero : Zero C
      has-coprods : β a b β Coproduct C a b
      has-prods : β a b β Product C a b
    open Zero has-zero
    open Binary-products C has-prods
    open Binary-coproducts C has-coprods
    coproductβproduct : β {a b} β Hom (a ββ b) (a ββ b)
    coproductβproduct = β¨ [ id , zeroβ ]
                        , [ zeroβ , id ] β©
    field
      coproductβ
product : β {a b} β is-invertible (coproductβproduct {a} {b})If that is the case, then the coproduct is a biproduct when equipped with the projections and
    has-biproducts : β {a b} β Biproduct a b
    has-biproducts {a} {b} .Biproduct.biapex = a ββ b
    has-biproducts .Biproduct.Οβ = [ id , zeroβ ]
    has-biproducts .Biproduct.Οβ = [ zeroβ , id ]
    has-biproducts .Biproduct.ΞΉβ = ΞΉβ
    has-biproducts .Biproduct.ΞΉβ = ΞΉβ
    has-biproducts .Biproduct.has-is-biproduct .is-biproduct.has-is-product =
        is-product-iso-apex coproductβ
product Οβββ¨β© Οβββ¨β© has-is-product
    has-biproducts .Biproduct.has-is-biproduct .is-biproduct.has-is-coproduct = has-is-coproduct
    has-biproducts .Biproduct.has-is-biproduct .is-biproduct.ΟΞΉβ = []βΞΉβ
    has-biproducts .Biproduct.has-is-biproduct .is-biproduct.ΟΞΉβ = []βΞΉβ
    has-biproducts .Biproduct.has-is-biproduct .is-biproduct.ΞΉΟ-comm =
      ΞΉβ β [ id , zeroβ ] β ΞΉβ β [ zeroβ , id ] β‘β¨ reflβ©ββ¨ pulll []βΞΉβ β©
      ΞΉβ β zeroβ β [ zeroβ , id ]               β‘β¨ pulll (zero-βl _) β zero-βr _ β©
      zeroβ                                     β‘Λβ¨ pulll (zero-βl _) β zero-βr _ β©
      ΞΉβ β zeroβ β [ id , zeroβ ]               β‘Λβ¨ reflβ©ββ¨ pulll []βΞΉβ β©
      ΞΉβ β [ zeroβ , id ] β ΞΉβ β [ id , zeroβ ] β
  open make-semiadditive
  to-semiadditive : make-semiadditive β is-semiadditive
  to-semiadditive mk .is-semiadditive.has-zero = has-zero mk
  to-semiadditive mk .is-semiadditive.has-biproducts = has-biproducts mkReally property, in a univalent category.β©οΈ
References
- Karvonen, Martti. 2020. βBiproducts Without Pointedness.β https://arxiv.org/abs/1801.06488.