From 0c6a1a312ce70ba673671291577f066e2a0a1bcb Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 25 Apr 2024 01:59:39 +0100 Subject: [PATCH 01/31] biproducts, semiadditive and additive categories Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Biproducts.v | 1017 +++++++++++++++++++++ theories/WildCat/Bifunctor.v | 34 +- theories/WildCat/Monoidal.v | 54 ++ theories/WildCat/Products.v | 2 +- 4 files changed, 1103 insertions(+), 4 deletions(-) create mode 100644 theories/Algebra/Homological/Biproducts.v diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v new file mode 100644 index 00000000000..e54b118cbdd --- /dev/null +++ b/theories/Algebra/Homological/Biproducts.v @@ -0,0 +1,1017 @@ +Require Import Basics.Overture Basics.Decidable Basics.Tactics Basics.Trunc. +Require Import WildCat.Core WildCat.Products WildCat.Coproducts WildCat.Equiv. +Require Import WildCat.PointedCat WildCat.Bifunctor WildCat.NatTrans WildCat.Prod. +Require Import WildCat.Monoidal. +Require Import Types.Bool Types.Forall. +Require Import Algebra.AbGroups.AbelianGroup. +Require Import WildCat.Square. + +(** * Categories with biproducts *) + +(** ** Biproducts *) + +Class Biproduct (I : Type) `{DecidablePaths I} {A : Type} + `{HasEquivs A, !IsPointedCat A} {x : I -> A} + := Build_Biproduct' { + biproduct_product :: Product I x; + biproduct_coproduct :: Coproduct I x; + catie_cat_coprod_prod_incl :: CatIsEquiv (cat_coprod_prod_incl x); +}. + +Arguments Biproduct I {_ A _ _ _ _ _ _} x. + +Section Biproducts. + Context {I : Type} {A : Type} (x : I -> A) `{Biproduct I A x}. + + Definition cate_coprod_prod : cat_coprod I x $<~> cat_prod I x + := Build_CatEquiv (cat_coprod_prod_incl x). + + Definition cat_biprod : A + := cat_prod I x. + + Definition cat_biprod_pr : forall (i : I), cat_biprod $-> x i + := cat_pr. + + Definition cat_biprod_corec {z : A} + : (forall i, z $-> x i) -> z $-> cat_biprod + := cat_prod_corec I. + + Definition cat_biprod_corec_beta {z : A} (f : forall i, z $-> x i) + : forall i, cat_biprod_pr i $o cat_biprod_corec f $== f i + := cat_prod_beta I f. + + Definition cat_biprod_corec_eta {z : A} (f : z $-> cat_biprod) + : cat_biprod_corec (fun i => cat_biprod_pr i $o f) $== f + := cat_prod_eta I f. + + Definition cat_biprod_corec_eta' {z : A} {f f' : forall i, z $-> x i} + : (forall i, f i $== f' i) -> cat_biprod_corec f $== cat_biprod_corec f' + := cat_prod_corec_eta I. + + Definition cat_biprod_pr_eta {z : A} {f f' : z $-> cat_biprod} + : (forall i, cat_biprod_pr i $o f $== cat_biprod_pr i $o f') -> f $== f' + := cat_prod_pr_eta I. + + Definition cat_biprod_in : forall (i : I), x i $-> cat_biprod + := fun i => cate_coprod_prod $o cat_in i. + + Definition cat_biprod_rec {z : A} + : (forall i, x i $-> z) -> cat_biprod $-> z + := fun f => cat_coprod_rec I f $o cate_coprod_prod^-1$. + + Definition cat_biprod_rec_beta {z : A} (f : forall i, x i $-> z) + : forall i, cat_biprod_rec f $o cat_biprod_in i $== f i. + Proof. + intros i. + unfold cat_biprod_rec, cat_biprod_in. + refine (_ $@ cat_coprod_beta I f i). + nrefine (cat_assoc _ _ _ $@ (_ $@L _)). + apply compose_V_hh. + Defined. + + Definition cat_biprod_rec_eta {z : A} (f : cat_biprod $-> z) + : cat_biprod_rec (fun i => f $o cat_biprod_in i) $== f. + Proof. + unfold cat_biprod_rec. + apply cate_moveR_eV. + apply cat_coprod_in_eta. + intros i. + refine (cat_coprod_beta I _ i $@ _^$). + apply cat_assoc. + Defined. + + Definition cat_biprod_rec_eta' {z : A} {f f' : forall i, x i $-> z} + : (forall i, f i $== f' i) -> cat_biprod_rec f $== cat_biprod_rec f'. + Proof. + intros p. + unfold cat_biprod_rec. + exact (cat_coprod_rec_eta I p $@R _). + Defined. + + Definition cat_biprod_in_eta {z : A} {f f' : cat_biprod $-> z} + : (forall i, f $o cat_biprod_in i $== f' $o cat_biprod_in i) -> f $== f'. + Proof. + intros p. + apply (cate_epic_equiv cate_coprod_prod). + apply cat_coprod_in_eta. + intros i. + exact (cat_assoc _ _ _ $@ p i $@ (cat_assoc _ _ _)^$). + Defined. + +End Biproducts. + +Arguments cat_biprod I {A} x {_ _ _ _ _ _ _ _}. +Arguments cat_biprod_pr {I A x _ _ _ _ _ _ _ _} i. +Arguments cat_biprod_in {I A x _ _ _ _ _ _ _ _} i. +Arguments cate_coprod_prod {I A} x {_ _ _ _ _ _ _ _}. + +(** An inclusion followed by a projection of the same index is the identity. *) +Definition cat_biprod_pr_in (I : Type) {A : Type} (x : I -> A) + `{Biproduct I A x} (i : I) + : cat_biprod_pr i $o cat_biprod_in i $== Id _. +Proof. + unfold cat_biprod_pr, cat_biprod_in. + refine ((_ $@L _) $@ _). + { refine ((cate_buildequiv_fun _ $@R _) $@ _). + apply cat_coprod_beta. } + refine (cat_prod_beta _ _ _ $@ _). + destruct (dec_paths i i); simpl. + - (** We cannot eliminate [p : i = i] with path induction, but we do have decidable equality, hence by Hedberg's theorem we must be in a hset and we can replace this with the identity. *) + assert (r : (idpath = p)) by apply path_ishprop. + by destruct r. + - contradiction n. + reflexivity. +Defined. + +(** An inclusion followed by a projection of a different index is zero. *) +Definition cat_biprod_pr_in_ne (I : Type) {A : Type} (x : I -> A) + `{Biproduct I A x} {i j : I} (p : i <> j) + : cat_biprod_pr j $o cat_biprod_in i $== zero_morphism. +Proof. + unfold cat_biprod_pr, cat_biprod_in. + refine ((_ $@L _) $@ _). + { refine ((cate_buildequiv_fun _ $@R _) $@ _). + apply cat_coprod_beta. } + refine (cat_prod_beta _ _ _ $@ _). + destruct (dec_paths i j); simpl. + - contradiction p. + - reflexivity. +Defined. + +Definition cat_biprod_diag I {A} (x : A) `{Biproduct I A (fun _ => x)} + : x $-> cat_biprod I (fun _ => x) + := cat_prod_diag x. + +Definition cat_biprod_codiag I {A} (x : A) `{Biproduct I A (fun _ => x)} + : cat_biprod I (fun _ => x) $-> x + := cat_coprod_fold x $o (cate_coprod_prod (fun _ => x))^-1$. + +(** Compatability of [cat_biprod_rec] and [cat_biprod_corec]. *) +Definition cat_biprod_corec_rec I `{DecidablePaths I} {A : Type} + `{HasEquivs A, !IsPointedCat A} {x y : I -> A} + `{!Biproduct I x, !Biproduct I y} + (f : forall i, x i $-> y i) + : cat_biprod_corec y (fun i => f i $o cat_biprod_pr i) + $== cat_biprod_rec x (fun i => cat_biprod_in i $o f i). +Proof. + apply cat_biprod_pr_eta. + intros i. + refine (cat_biprod_corec_beta _ _ i $@ _). + apply cat_biprod_in_eta. + intros j. + refine (_ $@ (_ $@L (cat_biprod_rec_beta _ _ _)^$) $@ (cat_assoc _ _ _)^$). + refine (cat_assoc _ _ _ $@ _ $@ cat_assoc _ _ _). + destruct (dec_paths j i) as [p | np]. + - destruct p. + refine ((_ $@L _) $@ cat_idr _ $@ (cat_idl _)^$ $@ (_^$ $@R _)). + 1,2: apply cat_biprod_pr_in. + - refine ((_ $@L _) $@ cat_zero_r _ $@ (cat_zero_l _)^$ $@ (_^$ $@R _)). + 1,2: apply cat_biprod_pr_in_ne, np. +Defined. + +(** *** Existence of biproducts *) + +Class HasBiproducts (I A : Type) + `{DecidablePaths I, HasEquivs A, !IsPointedCat A} + := has_biproducts :: forall (x : I -> A), Biproduct I x. + +(** *** Binary biproducts *) + +Class BinaryBiproduct {A : Type} `{HasEquivs A, !IsPointedCat A} (x y : A) + := binary_biproduct :: Biproduct Bool (fun b => if b then x else y). + +Class HasBinaryBiproducts (A : Type) `{HasEquivs A, !IsPointedCat A} + := has_binary_biproducts :: forall (x y : A), BinaryBiproduct x y. + +Global Instance hasbinarybiproducts_hasbiproductsbool {A : Type} + `{HasEquivs A, !IsPointedCat A, !HasBiproducts Bool A} + : HasBinaryBiproducts A + := fun x y => has_biproducts (fun b : Bool => if b then x else y). + +Global Instance hasbinaryproducts_hasbinarybiproducts {A : Type} + `{H : HasBinaryBiproducts A} + : HasBinaryProducts A. +Proof. + intros x y. + apply biproduct_product. +Defined. + +Global Instance hasbinarycoproducts_hasbinarybiproducts {A : Type} + `{H : HasBinaryBiproducts A} + : HasBinaryCoproducts A. +Proof. + intros x y. + nrapply biproduct_coproduct. + apply H. +Defined. + +Section BinaryBiproducts. + + Context {A : Type} `{HasEquivs A, !IsPointedCat A} {x y : A} + `{!BinaryBiproduct x y}. + + Definition cat_binbiprod : A + := cat_biprod Bool (fun b => if b then x else y). + + Definition cat_binbiprod_pr1 : cat_binbiprod $-> x := cat_biprod_pr true. + Definition cat_binbiprod_pr2 : cat_binbiprod $-> y := cat_biprod_pr false. + + Definition cat_binbiprod_corec {z : A} (f : z $-> x) (g : z $-> y) + : z $-> cat_binbiprod. + Proof. + apply cat_biprod_corec. + intros [|]. + - exact f. + - exact g. + Defined. + + Definition cat_binbiprod_corec_beta_pr1 {z : A} (f : z $-> x) (g : z $-> y) + : cat_binbiprod_pr1 $o cat_binbiprod_corec f g $== f + := cat_biprod_corec_beta _ _ true. + + Definition cat_binbiprod_corec_beta_pr2 {z : A} (f : z $-> x) (g : z $-> y) + : cat_binbiprod_pr2 $o cat_binbiprod_corec f g $== g + := cat_biprod_corec_beta _ _ false. + + Definition cat_binbiprod_corec_eta {z : A} (f : z $-> cat_binbiprod) + : cat_binbiprod_corec (cat_binbiprod_pr1 $o f) (cat_binbiprod_pr2 $o f) + $== f. + Proof. + apply cat_biprod_pr_eta. + intros [|]. + - exact (cat_binbiprod_corec_beta_pr1 _ _). + - exact (cat_binbiprod_corec_beta_pr2 _ _). + Defined. + + Definition cat_binbiprod_corec_eta' {z : A} {f f' : z $-> x} {g g' : z $-> y} + : f $== f' -> g $== g' + -> cat_binbiprod_corec f g $== cat_binbiprod_corec f' g'. + Proof. + intros p q. + apply cat_biprod_corec_eta'. + intros [|]. + - exact p. + - exact q. + Defined. + + Definition cat_binbiprod_pr_eta {z : A} {f f' : z $-> cat_binbiprod} + : cat_binbiprod_pr1 $o f $== cat_binbiprod_pr1 $o f' + -> cat_binbiprod_pr2 $o f $== cat_binbiprod_pr2 $o f' + -> f $== f'. + Proof. + intros p q. + apply cat_biprod_pr_eta. + intros [|]. + - exact p. + - exact q. + Defined. + + Definition cat_binbiprod_inl : x $-> cat_binbiprod := cat_biprod_in true. + Definition cat_binbiprod_inr : y $-> cat_binbiprod := cat_biprod_in false. + + Definition cat_binbiprod_rec {z : A} (f : x $-> z) (g : y $-> z) + : cat_binbiprod $-> z. + Proof. + apply cat_biprod_rec. + intros [|]. + - exact f. + - exact g. + Defined. + + Definition cat_binbiprod_rec_beta_inl {z : A} (f : x $-> z) (g : y $-> z) + : cat_binbiprod_rec f g $o cat_binbiprod_inl $== f + := cat_biprod_rec_beta _ _ true. + + Definition cat_binbiprod_rec_beta_inr {z : A} (f : x $-> z) (g : y $-> z) + : cat_binbiprod_rec f g $o cat_binbiprod_inr $== g + := cat_biprod_rec_beta _ _ false. + + Definition cat_binbiprod_rec_eta {z : A} (f : cat_binbiprod $-> z) + : cat_binbiprod_rec (f $o cat_binbiprod_inl) (f $o cat_binbiprod_inr) $== f. + Proof. + apply cat_biprod_in_eta. + intros [|]. + - exact (cat_binbiprod_rec_beta_inl _ _). + - exact (cat_binbiprod_rec_beta_inr _ _). + Defined. + + Definition cat_binbiprod_rec_eta' {z : A} {f f' : x $-> z} {g g' : y $-> z} + : f $== f' -> g $== g' + -> cat_binbiprod_rec f g $== cat_binbiprod_rec f' g'. + Proof. + intros p q. + apply cat_biprod_rec_eta'. + intros [|]. + - exact p. + - exact q. + Defined. + + Definition cat_binbiprod_in_eta {z : A} {f f' : cat_binbiprod $-> z} + : f $o cat_binbiprod_inl $== f' $o cat_binbiprod_inl + -> f $o cat_binbiprod_inr $== f' $o cat_binbiprod_inr + -> f $== f'. + Proof. + intros p q. + apply cat_biprod_in_eta. + intros [|]. + - exact p. + - exact q. + Defined. + + Definition cat_binbiprod_pr1_inl + : cat_binbiprod_pr1 $o cat_binbiprod_inl $== Id _ + := cat_biprod_pr_in Bool _ true. + + Definition cat_binbiprod_pr2_inr + : cat_binbiprod_pr2 $o cat_binbiprod_inr $== Id _ + := cat_biprod_pr_in Bool _ false. + + Definition cat_binbiprod_pr2_inl + : cat_binbiprod_pr2 $o cat_binbiprod_inl $== zero_morphism. + Proof. + snrapply (cat_biprod_pr_in_ne Bool _ (i := true) (j := false)). + apply (not_fixed_negb false). + Defined. + + Definition cat_binbiprod_pr1_inr + : cat_binbiprod_pr1 $o cat_binbiprod_inr $== zero_morphism. + Proof. + snrapply (cat_biprod_pr_in_ne Bool _ (i := false) (j := true)). + apply (not_fixed_negb true). + Defined. + +End BinaryBiproducts. + +Arguments cat_binbiprod {A _ _ _ _ _ _} x y {_}. + +(** Annoyingly this doesn't follow directly from the general diagonal since [fun b => if b then x else x] is not definitionally equal to [fun _ => x]. *) +Definition cat_binbiprod_diag {A : Type} + `{HasEquivs A, !IsPointedCat A} (x : A) `{!BinaryBiproduct x x} + : x $-> cat_binbiprod x x. +Proof. + snrapply cat_binbiprod_corec; exact (Id _). +Defined. + +Definition cat_binbiprod_codiag {A : Type} + `{HasEquivs A, !IsPointedCat A} (x : A) `{!BinaryBiproduct x x} + : cat_binbiprod x x $-> x. +Proof. + snrapply cat_binbiprod_rec; exact (Id _). +Defined. + +(** Compatability of [cat_binprod_rec] and [cat_binprod_corec]. *) +Definition cat_binbiprod_corec_rec {A : Type} + `{HasEquivs A, !IsPointedCat A} {w x y z : A} + `{!BinaryBiproduct w x, !BinaryBiproduct y z} + (f : w $-> y) (g : x $-> z) + : cat_binbiprod_corec (f $o cat_binbiprod_pr1) (g $o cat_binbiprod_pr2) + $== cat_binbiprod_rec (cat_binbiprod_inl $o f) (cat_binbiprod_inr $o g). +Proof. + unfold cat_binbiprod_corec, cat_binbiprod_rec. + nrefine (_ $@ _ $@ _). + 2: snrapply cat_biprod_corec_rec; by intros [|]. + 1: snrapply cat_biprod_corec_eta'; by intros [|]. + snrapply cat_biprod_rec_eta'; by intros [|]. +Defined. + +(** *** Binary biproduct bifunctor *) + +Global Instance is0bifunctor_cat_binbiprod {A : Type} + `{HasBinaryBiproducts A} + : Is0Bifunctor (fun x y : A => cat_binbiprod x y) + := is0bifunctor_cat_binprod. + +Global Instance is1bifunctor_cat_binbiprod {A : Type} + `{HasBinaryBiproducts A} + : Is1Bifunctor (fun x y : A => cat_binbiprod x y) + := is1bifunctor_cat_binprod. + +(** Binary biproducts are functorial in each argument. *) +Global Instance is0functor_cat_binbiprod_l {A : Type} + `{HasBinaryBiproducts A} (y : A) + : Is0Functor (fun x => cat_binbiprod x y) + := bifunctor_is0functor10 y. + +Global Instance is1functor_cat_binbiprod_l {A : Type} + `{HasBinaryBiproducts A} (y : A) + : Is1Functor (fun x => cat_binbiprod x y) + := bifunctor_is1functor10 y. + +Global Instance is0functor_cat_binbiprod_r {A : Type} + `{HasBinaryBiproducts A} (x : A) + : Is0Functor (fun y => cat_binbiprod x y) + := bifunctor_is0functor01 x. + +Global Instance is1functor_cat_binbiprod_r {A : Type} + `{HasBinaryBiproducts A} (x : A) + : Is1Functor (fun y => cat_binbiprod x y) + := bifunctor_is1functor01 x. + +(** *** Symmetry *) + +Section Symmetry. + Context {A : Type} `{HasBinaryBiproducts A}. + + Definition cat_binbiprod_swap (x y : A) + : cat_binbiprod x y $-> cat_binbiprod y x + := cat_binbiprod_corec cat_binbiprod_pr2 cat_binbiprod_pr1. + + Lemma cat_binbiprod_swap_swap (x y : A) + : cat_binbiprod_swap x y $o cat_binbiprod_swap y x $== Id _. + Proof. + rapply cat_binprod_swap_cat_binprod_swap. + Defined. + + Local Instance symmetricbraiding_cat_binbiprod + : SymmetricBraiding (fun x y => cat_binbiprod x y). + Proof. + rapply symmetricbraiding_binprod. + Defined. + + Lemma cat_binbiprod_swap_inl (x y : A) + : cat_binbiprod_swap x y $o cat_binbiprod_inl + $== cat_binbiprod_inr. + Proof. + apply cat_binbiprod_pr_eta. + - refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + refine (_ $@ _^$). + 1: apply cat_binbiprod_pr2_inl. + apply cat_binbiprod_pr1_inr. + - refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + refine (_ $@ _^$). + 1: apply cat_binbiprod_pr1_inl. + apply cat_binbiprod_pr2_inr. + Defined. + + Lemma cat_binbiprod_swap_inr (x y : A) + : cat_binbiprod_swap x y $o cat_binbiprod_inr + $== cat_binbiprod_inl. + Proof. + apply cat_binbiprod_pr_eta. + - refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + refine (_ $@ _^$). + 1: apply cat_binbiprod_pr2_inr. + apply cat_binbiprod_pr1_inl. + - refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + refine (_ $@ _^$). + 1: apply cat_binbiprod_pr1_inr. + apply cat_binbiprod_pr2_inl. + Defined. + + (** The swap map preserves the diagonal. *) + Lemma cat_binbiprod_swap_diag (x : A) + : cat_binbiprod_swap x x $o cat_binbiprod_diag x $== cat_binbiprod_diag x. + Proof. + apply cat_binbiprod_pr_eta. + - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + refine (cat_binbiprod_corec_beta_pr2 _ _ $@ _^$). + apply cat_binbiprod_corec_beta_pr1. + - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + refine (cat_binbiprod_corec_beta_pr1 _ _ $@ _^$). + apply cat_binbiprod_corec_beta_pr2. + Defined. + + (** The swap map preserves the codiagonal. TODO: proof can probably be shorter. *) + Lemma cat_binbiprod_swap_codiag (x : A) + : cat_binbiprod_codiag x $o cat_binbiprod_swap x x $== cat_binbiprod_codiag x. + Proof. + apply cat_binbiprod_in_eta. + - refine (_ $@ (cat_binbiprod_rec_beta_inl _ _)^$). + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + + apply cat_binbiprod_pr_eta. + * refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + refine (_ $@ _^$). + 1: apply cat_binbiprod_pr2_inl. + apply cat_binbiprod_pr1_inr. + * refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + refine (_ $@ _^$). + 1: apply cat_binbiprod_pr1_inl. + apply cat_binbiprod_pr2_inr. + + apply cat_binbiprod_rec_beta_inr. + - refine (_ $@ (cat_binbiprod_rec_beta_inr _ _)^$). + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + + apply cat_binbiprod_pr_eta. + * refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + refine (_ $@ _^$). + 1: apply cat_binbiprod_pr2_inr. + apply cat_binbiprod_pr1_inl. + * refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + refine (_ $@ _^$). + 1: apply cat_binbiprod_pr1_inr. + apply cat_binbiprod_pr2_inl. + + apply cat_binbiprod_rec_beta_inl. + Defined. + +End Symmetry. + +(** *** Associativity of binary products *) + +Section Associativity. + + Context {A : Type} `{HasBinaryBiproducts A}. + + Local Instance associator_binbiprod + : Associator (fun x y => cat_binbiprod x y) + := associator_binprod. + + Lemma cate_binbiprod_assoc (x y z : A) + : cat_binbiprod x (cat_binbiprod y z) + $<~> cat_binbiprod (cat_binbiprod x y) z. + Proof. + rapply associator_binbiprod. + Defined. + + Definition cat_binbiprod_twist (x y z : A) + : cat_binbiprod x (cat_binbiprod y z) + $-> cat_binbiprod y (cat_binbiprod x z) + := cat_binprod_twist x y z. + + Lemma cat_binbiprod_twist_inl (x y z : A) + : cat_binbiprod_twist x y z $o cat_binbiprod_inl + $== cat_binbiprod_inr $o cat_binbiprod_inl. + Proof. + apply cat_binbiprod_pr_eta. + - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + refine (cat_assoc _ _ _ $@ _ $@ cat_assoc _ _ _). + refine ((_ $@L _) $@ _ $@ (_^$ $@R _)). + 1: apply cat_binbiprod_pr2_inl. + 2: apply cat_binbiprod_pr1_inr. + refine (_ $@ _^$). + 1: apply cat_zero_r. + apply cat_zero_l. + - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + refine (cat_binbiprod_rec_beta_inl _ _ $@ cat_idr _ $@ _). + refine ((cat_idl _)^$ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + apply cat_binbiprod_pr2_inr. + Defined. + + Lemma cat_binbiprod_twist_inr_inl (x y z : A) + : cat_binbiprod_twist x y z $o cat_binbiprod_inr $o cat_binbiprod_inl + $== cat_binbiprod_inl. + Proof. + apply cat_binbiprod_pr_eta. + - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + 1: apply cat_binbiprod_corec_beta_pr1. } + refine (cat_assoc _ _ _ $@ _). + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_pr2_inr. + apply cat_idl. } + refine (_ $@ _^$). + 1,2: apply cat_binbiprod_pr1_inl. + - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + 1: apply cat_binbiprod_corec_beta_pr2. } + refine (cat_assoc _ _ _ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_rec. + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + 1: apply cat_binbiprod_pr2_inl. + refine (_ $@ _^$). + 1: apply cat_zero_r. + apply cat_binbiprod_pr2_inl. + Defined. + + Lemma cat_binbiprod_twist_inr_inr (x y z : A) + : cat_binbiprod_twist x y z $o cat_binbiprod_inr $o cat_binbiprod_inr + $== cat_binbiprod_inr $o cat_binbiprod_inr. + Proof. + apply cat_binbiprod_pr_eta. + - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + 1: apply cat_binbiprod_corec_beta_pr1. } + refine (cat_assoc _ _ _ $@ _). + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_pr2_inr. + apply cat_idl. } + refine (_ $@ _^$). + 1: apply cat_binbiprod_pr1_inr. + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_pr1_inr. + apply cat_zero_l. + - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + 1: apply cat_binbiprod_corec_beta_pr2. } + refine (cat_assoc _ _ _ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_rec. + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + 1: apply cat_binbiprod_pr2_inr. + refine (cat_idr _ $@ _^$). + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_pr2_inr. + apply cat_idl. + Defined. + + Lemma cate_binbiprod_assoc_inl (x y z : A) + : cate_binbiprod_assoc x y z $o cat_binbiprod_inl + $== cat_binbiprod_inl $o cat_binbiprod_inl. + Proof. + refine ((_ $@R _) $@ _). + 1: nrapply Monoidal.associator_twist'_unfold. + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L (cat_assoc _ _ _ $@ (_ $@L _))) $@ _). + { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). + 1: apply cat_binbiprod_rec_beta_inl. + apply cat_idr. } + simpl. + refine ((_ $@L _) $@ _). + 1: apply cat_binbiprod_twist_inl. + refine ((cat_assoc _ _ _)^$ $@ _). + refine (_ $@R _). + apply cat_binbiprod_swap_inr. + Defined. + + Lemma cate_binbiprod_assoc_inr_inl (x y z : A) + : cate_binbiprod_assoc x y z $o cat_binbiprod_inr $o cat_binbiprod_inl + $== cat_binbiprod_inl $o cat_binbiprod_inr. + Proof. + refine (((_ $@R _) $@R _) $@ _). + 1: nrapply Monoidal.associator_twist'_unfold. + refine (cat_assoc _ _ _ $@ _). + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L (cat_assoc _ _ _ $@ (_ $@L _))) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inr. } + refine ((_ $@L _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ _). + refine (((cat_assoc _ _ _)^$ $@R _) $@ _). + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + 1: apply cat_binbiprod_swap_inl. + apply cat_binbiprod_twist_inr_inr. } + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + apply cat_binbiprod_swap_inr. + Defined. + + Lemma cate_binbiprod_assoc_inr_inr (x y z : A) + : cate_binbiprod_assoc x y z $o cat_binbiprod_inr $o cat_binbiprod_inr + $== cat_binbiprod_inr. + Proof. + refine (((_ $@R _) $@R _) $@ _). + 1: nrapply Monoidal.associator_twist'_unfold. + refine (cat_assoc _ _ _ $@ _). + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L (cat_assoc _ _ _ $@ (_ $@L _))) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inr. } + refine ((_ $@L _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ _). + refine (((cat_assoc _ _ _)^$ $@R _) $@ _). + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + 1: apply cat_binbiprod_swap_inr. + apply cat_binbiprod_twist_inr_inl. } + apply cat_binbiprod_swap_inl. + Defined. + +End Associativity. + +(** ** Semiadditive Categories *) + +Class IsSemiAdditive (A : Type) `{HasEquivs A} := { + semiadditive_pointed :: IsPointedCat A; + semiadditive_hasbiproducts :: HasBinaryBiproducts A; + semiadditive_hasmorext :: HasMorExt A; + semiadditive_ishset_hom :: forall (a b : A), IsHSet (a $-> b); +}. + +Section CMonHom. + + Context {A : Type} `{IsSemiAdditive A} (a b : A). + + (** We can show that the hom-types of a semiadditive category are commutative monoids. *) + + Local Instance zero_hom : Zero (a $-> b) := zero_morphism. + + Local Instance sgop_hom : SgOp (a $-> b) := fun f g => + cat_binbiprod_codiag b + $o fmap11 (fun x y => cat_binbiprod x y) f g + $o cat_binbiprod_diag a. + + Local Instance left_identity_hom : LeftIdentity sgop_hom zero_hom. + Proof. + intros g. + apply path_hom. + unfold sgop_hom, fmap11; cbn. + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L cat_assoc _ _ _) $@ _). + refine ((cat_assoc _ _ _)^$ $@ _). + snrefine ((_ $@L _) $@ _). + 1: exact cat_binbiprod_inr. + { apply cat_binbiprod_pr_eta. + - refine (_ $@ cat_binbiprod_pr1_inr^$). + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + refine (cat_assoc _ _ _ $@ _). + apply cat_zero_l. + - refine (_ $@ cat_binbiprod_pr2_inr^$). + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + refine (cat_assoc _ _ _ $@ _). + refine (cat_idl _ $@ _). + apply cat_binbiprod_corec_beta_pr2. } + nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + { refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_rec. + apply cat_binbiprod_rec_beta_inr. } + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idl. + Defined. + + Local Existing Instance symmetricbraiding_cat_binbiprod. + + (** Using the naturality of swap we can show that the operation is commutative. *) + Local Instance commutative_hom : Commutative sgop_hom. + Proof. + intros f g. + apply path_hom. + refine (cat_assoc _ _ _ $@ _). + refine ((_^$ $@R _) $@ _). + 1: apply cat_binbiprod_swap_codiag. + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). + refine (_ $@ (_ $@L _)). + 2: apply cat_binbiprod_swap_diag. + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ cat_assoc _ _ _). + rapply Monoidal.braid_nat. + Defined. + + Local Instance right_identity_ab_hom : RightIdentity sgop_hom zero_hom. + Proof. + intros f. + refine (commutativity _ _ @ _). + apply left_identity_hom. + Defined. + + Local Instance associative_ab_hom : Associative sgop_hom. + Proof. + (** There is a lot to unfold here. But drawing out the diagram we can split this into 5 separate diagrams that all commute. Much of the work here will be applying associativity and functor laws to get this diagram in the correct shape. *) + intros f g h. + apply path_hom. + refine (((_ $@L (_ $@R _)) $@R _) $@ _ $@ ((_ $@L (_ $@L _)^$) $@R _)). + 1,3: refine (fmap_comp _ _ _ $@ (_ $@R _)). + 1,2: refine (fmap_comp _ _ _ $@ (_ $@L _)). + 1,2: rapply fmap_comp. + do 2 change (fmap (flip (fun x y => cat_binbiprod x y) ?a) ?f) + with (fmap10 (fun x y => cat_binbiprod x y) f a); + do 2 change (fmap ((fun x y => cat_binbiprod x y) ?b) ?f) + with (fmap01 (fun x y => cat_binbiprod x y) b f). + Local Notation "a ⊕L g" := (fmap01 (fun x y => cat_binbiprod x y) a g) (at level 99). + Local Notation "f ⊕R b" := (fmap10 (fun x y => cat_binbiprod x y) f b) (at level 99). + Local Notation "'Δ' x" := (cat_binbiprod_diag x) (at level 10). + Local Notation "'∇' x" := (cat_binbiprod_codiag x) (at level 10). + change (?w $o ?x $== ?y $o ?z) with (Square z w x y). + apply move_right_top. + (** The following associativity rewrites are specially chosen so that the diagram can be decomposed easily. *) + rapply vconcatL. + 1: do 2 refine (cat_assoc _ _ _ $@ _). + 1: refine (((cat_assoc _ _ _)^$ $@R _)). + rapply vconcatR. + 2: do 2 refine ((cat_assoc _ _ _)^$ $@ (_ $@L (cat_assoc _ _ _)) $@ _). + 2: refine (cat_assoc _ _ _)^$. + srapply hconcat. + (** We now have to specify the morphism that we are composing along, i.e. the common edge of the two squares where they meet. This is chosen to be the associativity equivalence for biproducts. *) + 1: nrapply cate_binbiprod_assoc. + { nrefine ((_ $@R _) $@ _). + 1: rapply Monoidal.associator_twist'_unfold. + apply cat_binbiprod_pr_eta. + - srefine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: exact (fmap01 (fun x y => cat_binbiprod x y) b cat_binbiprod_pr1). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + refine (_^$ $@ _). + 1: rapply fmap01_comp. + refine (fmap02 _ _ _). + apply cat_binbiprod_corec_beta_pr2. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((((fmap02 _ _ _ $@ fmap01_id _ _ _)^$ $@ fmap01_comp _ _ _ _)^$ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + nrefine (cat_idl _ $@ _). + refine (_ $@ (_ $@L (cat_assoc _ _ _)^$)). + nrefine (_ $@ cat_assoc _ _ _). + refine (_ $@ (_^$ $@R _)). + 2: rapply cat_binbiprod_corec_beta_pr1. + refine ((_ $@L _) $@ (cat_assoc _ _ _)^$). + nrefine (_ $@ cat_assoc _ _ _). + refine (_ $@ (_^$ $@R _)). + 2: rapply cat_binbiprod_corec_beta_pr1. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_ $@ (_ $@L _^$)). + 2: rapply cat_binbiprod_corec_beta_pr1. + symmetry. + apply cat_idr. + - nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). + refine ((cat_assoc _ _ _)^$ $@ _). + nrefine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). + refine ((cat_assoc _ _ _)^$ $@ _). + nrefine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + apply cat_binbiprod_corec_beta_pr2. } + nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). + refine ((cat_assoc _ _ _)^$ $@ _). + nrefine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + nrefine ((_ $@L _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + apply cat_binbiprod_corec_beta_pr2. } + nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). + refine ((cat_assoc _ _ _)^$ $@ _). + nrefine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + nrefine (cat_idl _ $@ _). + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + nrefine (cat_assoc _ _ _ $@ _). + nrefine (cat_idl _ $@ _). + nrefine (cat_binbiprod_corec_beta_pr2 _ _ $@ _). + refine (_ $@ (_ $@L (cat_assoc _ _ _)^$)). + nrefine (_ $@ cat_assoc _ _ _). + refine (_ $@ (_^$ $@R _)). + 2: apply cat_binbiprod_corec_beta_pr2. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_ $@ (cat_idl _)^$). + refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: apply cat_binbiprod_corec_beta_pr2. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_ $@ (cat_idl _)^$). + symmetry. + apply cat_binbiprod_corec_beta_pr2. } + (** We split the square again, this time by picking yet another associativity map as the common edge. This leaves us with a naturality square for associativity. *) + srapply hconcat. + 1: nrapply cate_binbiprod_assoc. + 1: nrapply Monoidal.associator_nat_m. + (** Finally we are left with another polygon which we resolve by brute force. *) + refine ((cat_assoc _ _ _)^$ $@ _). + apply cat_binbiprod_in_eta. + - nrefine (cat_assoc _ _ _ $@ _). + nrefine ((_ $@L _) $@ _). + { nrefine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inl. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine (cat_idr _ $@ _). + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inl. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine (cat_idr _ $@ _). + refine (_ $@ _). + 1: apply cat_binbiprod_rec_beta_inl. + symmetry. + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + 1: apply cate_binbiprod_assoc_inl. + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + { refine (cat_assoc _ _ _ $@ (_ $@L _)). + refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inl. } + simpl. + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + { refine (cat_assoc _ _ _ $@ (_ $@L _)). + apply cat_binbiprod_rec_beta_inl. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + { refine (cat_assoc _ _ _ $@ (_ $@L _)). + refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). + 1: apply cat_binbiprod_rec_beta_inl. + apply cat_idr. } + refine (cat_idr _ $@ _). + apply cat_binbiprod_rec_beta_inl. + - nrefine (cat_assoc _ _ _ $@ _). + nrefine ((_ $@L _) $@ _). + { nrefine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inr. } + refine ((cat_assoc _ _ _)^$ $@ _). + simpl. + refine ((_ $@R _) $@ _). + { refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + 1: refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idl. } + simpl. + refine (_ $@ (cat_assoc _ _ _)^$). + apply cat_binbiprod_in_eta. + + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inl. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inl. + refine (cat_idr _ $@ _). + refine (_ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). + 2: apply cate_binbiprod_assoc_inr_inl. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_ $@ (_ $@L _^$)). + 2: { refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inl. + simpl. + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idr. } + simpl. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_^$ $@ (_ $@L _^$)). + 2: { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). + 1: apply cat_binbiprod_rec_beta_inl. + apply cat_idr. } + apply cat_binbiprod_rec_beta_inl. + + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inr. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + refine (cat_idl _ $@ _). + refine (_ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). + 2: apply cate_binbiprod_assoc_inr_inr. + simpl. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_^$ $@ (_ $@L _^$)). + 2: { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idr. } + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idl. + Defined. + + Local Instance issemigroup_hom : IsSemiGroup (a $-> b) := {}. + Local Instance ismonoid_hom : IsMonoid (a $-> b) := {}. + +End CMonHom. + +(** ** Additive Categories *) + +Class IsAdditive (A : Type) `{HasEquivs A} := { + additive_semiadditive :: IsSemiAdditive A; + additive_negate_hom : forall (a b : A), Negate (a $-> b); + additive_left_inverse_hom : forall (a b : A), + LeftInverse (sgop_hom a b) (additive_negate_hom a b) (zero_hom a b); + additive_right_inverse_hom : forall (a b : A), + RightInverse (sgop_hom a b) (additive_negate_hom a b) (zero_hom a b); +}. + +(** Homs in an aditive category form an abelian group. *) +Definition AbHom {A : Type} `{IsAdditive A} : A -> A -> AbGroup. +Proof. + intros a b. + snrapply Build_AbGroup. + 1: snrapply Build_Group. + 5: split. + - exact (a $-> b). + - exact (sgop_hom a b). + - exact (zero_hom a b). + - exact (additive_negate_hom a b). + - exact (ismonoid_hom a b). + - exact (additive_left_inverse_hom a b). + - exact (additive_right_inverse_hom a b). + - exact (commutative_hom a b). +Defined. + +(** TODO: Additive functors *) diff --git a/theories/WildCat/Bifunctor.v b/theories/WildCat/Bifunctor.v index dd6c4667ccc..e1f2e3e1982 100644 --- a/theories/WildCat/Bifunctor.v +++ b/theories/WildCat/Bifunctor.v @@ -367,7 +367,16 @@ Global Instance is0functor_uncurry_uncurry_left {A B C D E} !Is0Bifunctor F, !Is0Bifunctor G} : Is0Functor (uncurry (uncurry (fun x y z => G (F x y) z))). Proof. - rapply is0functor_uncurry_bifunctor. + nrapply is0functor_uncurry_bifunctor; exact _. +Defined. + +Global Instance is1functor_uncurry_uncurry_left {A B C D E} + (F : A -> B -> C) (G : C -> D -> E) + `{Is1Cat A, Is1Cat B, Is1Cat C, Is1Cat D, Is1Cat E, + !Is0Bifunctor F, !Is1Bifunctor F, !Is0Bifunctor G, !Is1Bifunctor G} + : Is1Functor (uncurry (uncurry (fun x y z => G (F x y) z))). +Proof. + nrapply is1functor_uncurry_bifunctor; exact _. Defined. Global Instance is0functor_uncurry_uncurry_right {A B C D E} @@ -381,10 +390,29 @@ Proof. 1: exact _. intros b. change (Is0Functor (uncurry (fun x y => G x (F y b)))). - apply is0functor_uncurry_bifunctor. - apply (is0bifunctor_precompose' (flip F b) G). + exact _. Defined. +Global Instance is1functor_uncurry_uncurry_right {A B C D E} + (F : A -> B -> D) (G : C -> D -> E) + `{Is1Cat A, Is1Cat B, Is1Cat C, Is1Cat D, Is1Cat E, + !Is0Bifunctor F, !Is1Bifunctor F, !Is0Bifunctor G, !Is1Bifunctor G} + : Is1Functor (uncurry (uncurry (fun x y z => G x (F y z)))). +Proof. + apply is1functor_uncurry_bifunctor. + nrapply Build_Is1Bifunctor. + 1: exact _. + - intros b. + change (Is1Functor (uncurry (fun x y => G x (F y b)))). + exact _. + - intros [c a] [c' a'] [h f] b b' g. + refine ((cat_assoc _ _ _)^$ $@ _ $@ (cat_assoc _ _ _)^$). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _)). + 2: exact (fmap11_coh G h (fmap01 F a g)). + rapply fmap_square; unfold Square. + exact (fmap11_coh F f g). +Defined. + (** ** Natural transformations between bifunctors *) (** We can show that an uncurried natural transformation between uncurried bifunctors by composing the naturality square in each variable. *) diff --git a/theories/WildCat/Monoidal.v b/theories/WildCat/Monoidal.v index e2e81366f07..9d6040a995b 100644 --- a/theories/WildCat/Monoidal.v +++ b/theories/WildCat/Monoidal.v @@ -147,6 +147,60 @@ Class IsSymmetricMonoidal (A : Type) `{HasEquivs A} cat_symm_tensor_hexagon :: HexagonIdentity cat_tensor; }. +(** *** Theory about [Associator] *) + +Section Associator. + Context {A : Type} {F : A -> A -> A} `{Associator A F, !HasEquivs A}. + + Local Definition associator_nat {x x' y y' z z'} + (f : x $-> x') (g : y $-> y') (h : z $-> z') + : associator x' y' z' $o fmap11 F f (fmap11 F g h) + $== fmap11 F (fmap11 F f g) h $o associator x y z. + Proof. + refine (_ $@ is1natural_associator_uncurried (x, y, z) (x', y', z') (f, g, h)). + refine (_ $@L ((_ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _^$))); rapply fmap_comp. + Defined. + + Local Definition associator_nat_l {x x' : A} (f : x $-> x') (y z : A) + : associator x' y z $o fmap10 F f (F y z) + $== fmap10 F (fmap10 F f y) z $o associator x y z. + Proof. + refine ((_ $@L _^$) $@ _ $@ (_ $@R _)). + 2: apply (is1natural_associator_uncurried (x, y, z) (x', y, z) (f, Id _, Id _)). + - refine ((_ $@@ _) $@ cat_idl _). + 2: exact (fmap02 _ _ (fmap01_id _ _ _) $@ fmap01_id _ _ _). + rapply fmap10_is_fmap11. + - refine ((fmap20 _ _ _ $@@ fmap01_id _ _ _) $@ cat_idl _). + rapply fmap10_is_fmap11. + Defined. + + Local Definition associator_nat_m (x : A) {y y' : A} (g : y $-> y') (z : A) + : associator x y' z $o fmap01 F x (fmap10 F g z) + $== fmap10 F (fmap01 F x g) z $o associator x y z. + Proof. + refine ((_ $@L _^$) $@ _ $@ (_ $@R _)). + 2: apply (is1natural_associator_uncurried (x, y, z) (x, y', z) (Id _, g, Id _)). + - refine ((fmap01_is_fmap11 _ _ _ $@@ _) $@ cat_idl _). + exact (fmap02 _ _ (fmap01_id _ _ _) $@ fmap01_id _ _ _). + - refine ((fmap20 _ _ _ $@@ fmap01_id _ _ _) $@ cat_idl _). + rapply fmap01_is_fmap11. + Defined. + + Local Definition associator_nat_r (x y : A) {z z' : A} (h : z $-> z') + : associator x y z' $o fmap01 F x (fmap01 F y h) + $== fmap01 F (F x y) h $o associator x y z. + Proof. + refine ((_ $@L _^$) $@ _ $@ (_ $@R _)). + 2: apply (is1natural_associator_uncurried (x, y, z) (x, y, z') (Id _, Id _, h)). + - refine ((_ $@L _) $@ cat_idr _). + rapply fmap11_id. + - refine ((_ $@L _) $@ cat_idr _). + refine (fmap20 _ _ _ $@ fmap10_id _ _ _). + rapply fmap11_id. + Defined. + +End Associator. + (** ** Theory about [SymmetricBraid] *) Section SymmetricBraid. diff --git a/theories/WildCat/Products.v b/theories/WildCat/Products.v index c50ab4ca525..3a9090e5411 100644 --- a/theories/WildCat/Products.v +++ b/theories/WildCat/Products.v @@ -437,7 +437,7 @@ Proof. Defined. (** As a special case of the product functor, restriction along [Bool_rec A] yields bifunctoriality of [cat_binprod]. *) -Global Instance isbifunctor_cat_binprod {A : Type} `{HasBinaryProducts A} +Global Instance is0bifunctor_cat_binprod {A : Type} `{HasBinaryProducts A} : Is0Bifunctor (fun x y => cat_binprod x y). Proof. pose (p:=@has_products _ _ _ _ _ _ hasproductsbool_hasbinaryproducts). From 970f3c337fde6315cfa941324187804f596eee1f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 25 Apr 2024 05:14:00 +0100 Subject: [PATCH 02/31] Endomorphism ring in an additive category Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Biproducts.v | 109 +++++++++++++++++++++- 1 file changed, 106 insertions(+), 3 deletions(-) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index e54b118cbdd..8ffd0a2ef1b 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -1,10 +1,10 @@ Require Import Basics.Overture Basics.Decidable Basics.Tactics Basics.Trunc. Require Import WildCat.Core WildCat.Products WildCat.Coproducts WildCat.Equiv. Require Import WildCat.PointedCat WildCat.Bifunctor WildCat.NatTrans WildCat.Prod. -Require Import WildCat.Monoidal. +Require Import WildCat.Monoidal WildCat.Square. Require Import Types.Bool Types.Forall. -Require Import Algebra.AbGroups.AbelianGroup. -Require Import WildCat.Square. +Require Import Algebra.AbGroups.AbelianGroup Algebra.Rings.Ring. +Require Import canonical_names. (** * Categories with biproducts *) @@ -407,6 +407,68 @@ Global Instance is1functor_cat_binbiprod_r {A : Type} : Is1Functor (fun y => cat_binbiprod x y) := bifunctor_is1functor01 x. +(** *** Diagonal lemmas *) + +Definition cat_binbiprod_diag_fmap11 {A : Type} + `{HasEquivs A, !IsPointedCat A, !HasBinaryBiproducts A} {x y : A} (f : x $-> y) + : cat_binbiprod_diag y $o f + $== fmap11 (fun x y => cat_binbiprod x y) f f $o cat_binbiprod_diag x. +Proof. + apply cat_binbiprod_pr_eta. + - refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ (_ $@ _)). + 1: apply cat_binbiprod_corec_beta_pr1. + 1: apply cat_idl. + refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: rapply cat_pr1_fmap11_binprod. + refine ((cat_idr _)^$ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). + apply cat_binbiprod_corec_beta_pr1. + - refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ (_ $@ _)). + 1: apply cat_binbiprod_corec_beta_pr2. + 1: apply cat_idl. + refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: rapply cat_pr2_fmap11_binprod. + refine ((cat_idr _)^$ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). + apply cat_binbiprod_corec_beta_pr2. +Defined. + +Definition cat_binbiprod_codiag_fmap11 {A : Type} + `{HasEquivs A, !IsPointedCat A, !HasBinaryBiproducts A} {x y : A} (f : x $-> y) + : f $o cat_binbiprod_codiag x + $== cat_binbiprod_codiag y $o fmap11 (fun x y => cat_binbiprod x y) f f. +Proof. + apply cat_binbiprod_in_eta. + - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ cat_idr _ $@ _). + 1: apply cat_binbiprod_rec_beta_inl. + refine (_ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). + 2: { refine ((_ $@L _) $@ (cat_assoc _ _ _)^$). + refine (_^$ $@ (cat_binbiprod_corec_rec _ _ $@R _)^$). + apply cat_binbiprod_rec_beta_inl. } + refine (_ $@ (_ $@L _)). + 2: { refine ((_ $@R _) $@ cat_assoc _ _ _). + refine ((_ $@ _)^$ $@ (cat_binbiprod_corec_rec _ _ $@R _)^$). + 1: apply cat_binbiprod_rec_beta_inl. + apply cat_idr. } + refine (_^$ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: apply cat_binbiprod_rec_beta_inl. + apply cat_idl. + - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ cat_idr _ $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + refine (_ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). + 2: { refine ((_ $@L _) $@ (cat_assoc _ _ _)^$). + refine ((_ $@ _)^$ $@ (cat_binbiprod_corec_rec _ _ $@R _)^$). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idr. } + refine (_ $@ (_ $@L _)). + 2: { refine ((_ $@R _) $@ _)^$. + 1: apply cat_binbiprod_corec_rec. + apply cat_binbiprod_rec_beta_inr. } + refine (_^$ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: apply cat_binbiprod_rec_beta_inr. + apply cat_idl. +Defined. + (** *** Symmetry *) Section Symmetry. @@ -988,6 +1050,8 @@ End CMonHom. (** ** Additive Categories *) +(** *** Definition *) + Class IsAdditive (A : Type) `{HasEquivs A} := { additive_semiadditive :: IsSemiAdditive A; additive_negate_hom : forall (a b : A), Negate (a $-> b); @@ -997,6 +1061,8 @@ Class IsAdditive (A : Type) `{HasEquivs A} := { RightInverse (sgop_hom a b) (additive_negate_hom a b) (zero_hom a b); }. +(** *** Abelian Group structure on Hom *) + (** Homs in an aditive category form an abelian group. *) Definition AbHom {A : Type} `{IsAdditive A} : A -> A -> AbGroup. Proof. @@ -1014,4 +1080,41 @@ Proof. - exact (commutative_hom a b). Defined. +(** *** Endomorphism ring *) + +(** The composition operation provides a multiplicative operation on the abelian group hom. Turning it into a ring of endomorphsims. *) +Definition End {A : Type} `{IsAdditive A} (X : A) : Ring. +Proof. + snrapply Build_Ring'; repeat split. + - exact (AbHom X X). + - exact (fun f g => f $o g). + - exact (Id _). + - intros f g h. + apply path_hom. + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _ $@ cat_assoc _ _ _). + 1: apply cat_binbiprod_codiag_fmap11. + refine (cat_assoc _ _ _ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). + rapply fmap11_comp. + - intros f g h. + apply path_hom. + refine (cat_assoc _ _ _ $@ _). + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). + refine ((_ $@L _) $@ _). + 1: apply cat_binbiprod_diag_fmap11. + refine ((cat_assoc _ _ _)^$ $@ (_^$ $@R _)). + rapply fmap11_comp. + - exact _. + - intros f g h. + apply path_hom. + symmetry. + apply cat_assoc. + - intros f. + apply path_hom. + apply cat_idl. + - intros g. + apply path_hom. + apply cat_idr. +Defined. + (** TODO: Additive functors *) From 7188dce961b9000a19848ba69652d1d36da7913f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 25 Apr 2024 14:38:35 +0100 Subject: [PATCH 03/31] reorganise additive categories Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Additive.v | 374 ++++++++++++++++++++++ theories/Algebra/Homological/Biproducts.v | 372 +-------------------- 2 files changed, 376 insertions(+), 370 deletions(-) create mode 100644 theories/Algebra/Homological/Additive.v diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v new file mode 100644 index 00000000000..3f9aad37d4f --- /dev/null +++ b/theories/Algebra/Homological/Additive.v @@ -0,0 +1,374 @@ +Require Import Basics.Overture. +Require Import WildCat.Core WildCat.Equiv WildCat.PointedCat WildCat.Bifunctor. +Require Import WildCat.Square. +Require Import Algebra.Homological.Biproducts. +Require Import Algebra.AbGroups.AbelianGroup Algebra.Rings.Ring. +Require Import canonical_names. + +(** * Semiadditive and Additive Categories *) + + +(** ** Semiadditive Categories *) + +Class IsSemiAdditive (A : Type) `{HasEquivs A} := { + semiadditive_pointed :: IsPointedCat A; + semiadditive_hasbiproducts :: HasBinaryBiproducts A; + semiadditive_hasmorext :: HasMorExt A; + semiadditive_ishset_hom :: forall (a b : A), IsHSet (a $-> b); +}. + +Section CMonHom. + + Context {A : Type} `{IsSemiAdditive A} (a b : A). + + (** We can show that the hom-types of a semiadditive category are commutative monoids. *) + + Local Instance zero_hom : Zero (a $-> b) := zero_morphism. + + Local Instance sgop_hom : SgOp (a $-> b) := fun f g => + cat_binbiprod_codiag b + $o fmap11 (fun x y => cat_binbiprod x y) f g + $o cat_binbiprod_diag a. + + Local Instance left_identity_hom : LeftIdentity sgop_hom zero_hom. + Proof. + intros g. + apply path_hom. + unfold sgop_hom, fmap11; cbn. + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L cat_assoc _ _ _) $@ _). + refine ((cat_assoc _ _ _)^$ $@ _). + snrefine ((_ $@L _) $@ _). + 1: exact cat_binbiprod_inr. + { apply cat_binbiprod_pr_eta. + - refine (_ $@ cat_binbiprod_pr1_inr^$). + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + refine (cat_assoc _ _ _ $@ _). + apply cat_zero_l. + - refine (_ $@ cat_binbiprod_pr2_inr^$). + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + refine (cat_assoc _ _ _ $@ _). + refine (cat_idl _ $@ _). + apply cat_binbiprod_corec_beta_pr2. } + nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + { refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_rec. + apply cat_binbiprod_rec_beta_inr. } + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idl. + Defined. + + Local Existing Instance symmetricbraiding_cat_binbiprod. + + (** Using the naturality of swap we can show that the operation is commutative. *) + Local Instance commutative_hom : Commutative sgop_hom. + Proof. + intros f g. + apply path_hom. + refine (cat_assoc _ _ _ $@ _). + refine ((_^$ $@R _) $@ _). + 1: apply cat_binbiprod_swap_codiag. + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). + refine (_ $@ (_ $@L _)). + 2: apply cat_binbiprod_swap_diag. + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ cat_assoc _ _ _). + rapply Monoidal.braid_nat. + Defined. + + Local Instance right_identity_ab_hom : RightIdentity sgop_hom zero_hom. + Proof. + intros f. + refine (commutativity _ _ @ _). + apply left_identity_hom. + Defined. + + Local Instance associative_ab_hom : Associative sgop_hom. + Proof. + (** There is a lot to unfold here. But drawing out the diagram we can split this into 5 separate diagrams that all commute. Much of the work here will be applying associativity and functor laws to get this diagram in the correct shape. *) + intros f g h. + apply path_hom. + refine (((_ $@L (_ $@R _)) $@R _) $@ _ $@ ((_ $@L (_ $@L _)^$) $@R _)). + 1,3: refine (fmap_comp _ _ _ $@ (_ $@R _)). + 1,2: refine (fmap_comp _ _ _ $@ (_ $@L _)). + 1,2: rapply fmap_comp. + do 2 change (fmap (flip (fun x y => cat_binbiprod x y) ?a) ?f) + with (fmap10 (fun x y => cat_binbiprod x y) f a); + do 2 change (fmap ((fun x y => cat_binbiprod x y) ?b) ?f) + with (fmap01 (fun x y => cat_binbiprod x y) b f). + Local Notation "a ⊕L g" := (fmap01 (fun x y => cat_binbiprod x y) a g) (at level 99). + Local Notation "f ⊕R b" := (fmap10 (fun x y => cat_binbiprod x y) f b) (at level 99). + Local Notation "'Δ' x" := (cat_binbiprod_diag x) (at level 10). + Local Notation "'∇' x" := (cat_binbiprod_codiag x) (at level 10). + change (?w $o ?x $== ?y $o ?z) with (Square z w x y). + apply move_right_top. + (** The following associativity rewrites are specially chosen so that the diagram can be decomposed easily. *) + rapply vconcatL. + 1: do 2 refine (cat_assoc _ _ _ $@ _). + 1: refine (((cat_assoc _ _ _)^$ $@R _)). + rapply vconcatR. + 2: do 2 refine ((cat_assoc _ _ _)^$ $@ (_ $@L (cat_assoc _ _ _)) $@ _). + 2: refine (cat_assoc _ _ _)^$. + srapply hconcat. + (** We now have to specify the morphism that we are composing along, i.e. the common edge of the two squares where they meet. This is chosen to be the associativity equivalence for biproducts. *) + 1: nrapply cate_binbiprod_assoc. + { nrefine ((_ $@R _) $@ _). + 1: rapply Monoidal.associator_twist'_unfold. + apply cat_binbiprod_pr_eta. + - srefine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: exact (fmap01 (fun x y => cat_binbiprod x y) b cat_binbiprod_pr1). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + refine (_^$ $@ _). + 1: rapply fmap01_comp. + refine (fmap02 _ _ _). + apply cat_binbiprod_corec_beta_pr2. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((((fmap02 _ _ _ $@ fmap01_id _ _ _)^$ $@ fmap01_comp _ _ _ _)^$ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + nrefine (cat_idl _ $@ _). + refine (_ $@ (_ $@L (cat_assoc _ _ _)^$)). + nrefine (_ $@ cat_assoc _ _ _). + refine (_ $@ (_^$ $@R _)). + 2: rapply cat_binbiprod_corec_beta_pr1. + refine ((_ $@L _) $@ (cat_assoc _ _ _)^$). + nrefine (_ $@ cat_assoc _ _ _). + refine (_ $@ (_^$ $@R _)). + 2: rapply cat_binbiprod_corec_beta_pr1. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_ $@ (_ $@L _^$)). + 2: rapply cat_binbiprod_corec_beta_pr1. + symmetry. + apply cat_idr. + - nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). + refine ((cat_assoc _ _ _)^$ $@ _). + nrefine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). + refine ((cat_assoc _ _ _)^$ $@ _). + nrefine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + apply cat_binbiprod_corec_beta_pr2. } + nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). + refine ((cat_assoc _ _ _)^$ $@ _). + nrefine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr1. + nrefine ((_ $@L _) $@ _). + { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + apply cat_binbiprod_corec_beta_pr2. } + nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). + refine ((cat_assoc _ _ _)^$ $@ _). + nrefine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + nrefine (cat_idl _ $@ _). + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). + 1: apply cat_binbiprod_corec_beta_pr2. + nrefine (cat_assoc _ _ _ $@ _). + nrefine (cat_idl _ $@ _). + nrefine (cat_binbiprod_corec_beta_pr2 _ _ $@ _). + refine (_ $@ (_ $@L (cat_assoc _ _ _)^$)). + nrefine (_ $@ cat_assoc _ _ _). + refine (_ $@ (_^$ $@R _)). + 2: apply cat_binbiprod_corec_beta_pr2. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_ $@ (cat_idl _)^$). + refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: apply cat_binbiprod_corec_beta_pr2. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_ $@ (cat_idl _)^$). + symmetry. + apply cat_binbiprod_corec_beta_pr2. } + (** We split the square again, this time by picking yet another associativity map as the common edge. This leaves us with a naturality square for associativity. *) + srapply hconcat. + 1: nrapply cate_binbiprod_assoc. + 1: nrapply Monoidal.associator_nat_m. + (** Finally we are left with another polygon which we resolve by brute force. *) + refine ((cat_assoc _ _ _)^$ $@ _). + apply cat_binbiprod_in_eta. + - nrefine (cat_assoc _ _ _ $@ _). + nrefine ((_ $@L _) $@ _). + { nrefine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inl. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine (cat_idr _ $@ _). + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inl. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine (cat_idr _ $@ _). + refine (_ $@ _). + 1: apply cat_binbiprod_rec_beta_inl. + symmetry. + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + 1: apply cate_binbiprod_assoc_inl. + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + { refine (cat_assoc _ _ _ $@ (_ $@L _)). + refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inl. } + simpl. + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + { refine (cat_assoc _ _ _ $@ (_ $@L _)). + apply cat_binbiprod_rec_beta_inl. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + { refine (cat_assoc _ _ _ $@ (_ $@L _)). + refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). + 1: apply cat_binbiprod_rec_beta_inl. + apply cat_idr. } + refine (cat_idr _ $@ _). + apply cat_binbiprod_rec_beta_inl. + - nrefine (cat_assoc _ _ _ $@ _). + nrefine ((_ $@L _) $@ _). + { nrefine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inr. } + refine ((cat_assoc _ _ _)^$ $@ _). + simpl. + refine ((_ $@R _) $@ _). + { refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + 1: refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idl. } + simpl. + refine (_ $@ (cat_assoc _ _ _)^$). + apply cat_binbiprod_in_eta. + + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inl. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inl. + refine (cat_idr _ $@ _). + refine (_ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). + 2: apply cate_binbiprod_assoc_inr_inl. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_ $@ (_ $@L _^$)). + 2: { refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inl. + simpl. + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idr. } + simpl. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_^$ $@ (_ $@L _^$)). + 2: { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). + 1: apply cat_binbiprod_rec_beta_inl. + apply cat_idr. } + apply cat_binbiprod_rec_beta_inl. + + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + apply cat_binbiprod_rec_beta_inr. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + refine (cat_idl _ $@ _). + refine (_ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). + 2: apply cate_binbiprod_assoc_inr_inr. + simpl. + refine (_ $@ (cat_assoc _ _ _)^$). + refine (_^$ $@ (_ $@L _^$)). + 2: { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idr. } + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. } + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_binbiprod_rec_beta_inr. + apply cat_idl. + Defined. + + Local Instance issemigroup_hom : IsSemiGroup (a $-> b) := {}. + Local Instance ismonoid_hom : IsMonoid (a $-> b) := {}. + +End CMonHom. + +(** ** Additive Categories *) + +(** *** Definition *) + +Class IsAdditive (A : Type) `{HasEquivs A} := { + additive_semiadditive :: IsSemiAdditive A; + additive_negate_hom : forall (a b : A), Negate (a $-> b); + additive_left_inverse_hom : forall (a b : A), + LeftInverse (sgop_hom a b) (additive_negate_hom a b) (zero_hom a b); + additive_right_inverse_hom : forall (a b : A), + RightInverse (sgop_hom a b) (additive_negate_hom a b) (zero_hom a b); +}. + +(** *** Abelian Group structure on Hom *) + +(** Homs in an aditive category form an abelian group. *) +Definition AbHom {A : Type} `{IsAdditive A} : A -> A -> AbGroup. +Proof. + intros a b. + snrapply Build_AbGroup. + 1: snrapply Build_Group. + 5: split. + - exact (a $-> b). + - exact (sgop_hom a b). + - exact (zero_hom a b). + - exact (additive_negate_hom a b). + - exact (ismonoid_hom a b). + - exact (additive_left_inverse_hom a b). + - exact (additive_right_inverse_hom a b). + - exact (commutative_hom a b). +Defined. + +(** *** Endomorphism ring *) + +(** The composition operation provides a multiplicative operation on the abelian group hom. Turning it into a ring of endomorphsims. *) +Definition End {A : Type} `{IsAdditive A} (X : A) : Ring. +Proof. + snrapply Build_Ring'; repeat split. + - exact (AbHom X X). + - exact (fun f g => f $o g). + - exact (Id _). + - intros f g h. + apply path_hom. + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _ $@ cat_assoc _ _ _). + 1: apply cat_binbiprod_codiag_fmap11. + refine (cat_assoc _ _ _ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). + rapply fmap11_comp. + - intros f g h. + apply path_hom. + refine (cat_assoc _ _ _ $@ _). + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). + refine ((_ $@L _) $@ _). + 1: apply cat_binbiprod_diag_fmap11. + refine ((cat_assoc _ _ _)^$ $@ (_^$ $@R _)). + rapply fmap11_comp. + - exact _. + - intros f g h. + apply path_hom. + symmetry. + apply cat_assoc. + - intros f. + apply path_hom. + apply cat_idl. + - intros g. + apply path_hom. + apply cat_idr. +Defined. + +(** TODO: Additive functors *) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 8ffd0a2ef1b..4cd212874e0 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -1,10 +1,7 @@ Require Import Basics.Overture Basics.Decidable Basics.Tactics Basics.Trunc. +Require Import Types.Forall Types.Bool. Require Import WildCat.Core WildCat.Products WildCat.Coproducts WildCat.Equiv. -Require Import WildCat.PointedCat WildCat.Bifunctor WildCat.NatTrans WildCat.Prod. -Require Import WildCat.Monoidal WildCat.Square. -Require Import Types.Bool Types.Forall. -Require Import Algebra.AbGroups.AbelianGroup Algebra.Rings.Ring. -Require Import canonical_names. +Require Import WildCat.PointedCat WildCat.Bifunctor WildCat.Square WildCat.Monoidal. (** * Categories with biproducts *) @@ -753,368 +750,3 @@ Section Associativity. Defined. End Associativity. - -(** ** Semiadditive Categories *) - -Class IsSemiAdditive (A : Type) `{HasEquivs A} := { - semiadditive_pointed :: IsPointedCat A; - semiadditive_hasbiproducts :: HasBinaryBiproducts A; - semiadditive_hasmorext :: HasMorExt A; - semiadditive_ishset_hom :: forall (a b : A), IsHSet (a $-> b); -}. - -Section CMonHom. - - Context {A : Type} `{IsSemiAdditive A} (a b : A). - - (** We can show that the hom-types of a semiadditive category are commutative monoids. *) - - Local Instance zero_hom : Zero (a $-> b) := zero_morphism. - - Local Instance sgop_hom : SgOp (a $-> b) := fun f g => - cat_binbiprod_codiag b - $o fmap11 (fun x y => cat_binbiprod x y) f g - $o cat_binbiprod_diag a. - - Local Instance left_identity_hom : LeftIdentity sgop_hom zero_hom. - Proof. - intros g. - apply path_hom. - unfold sgop_hom, fmap11; cbn. - refine (cat_assoc _ _ _ $@ _). - refine ((_ $@L cat_assoc _ _ _) $@ _). - refine ((cat_assoc _ _ _)^$ $@ _). - snrefine ((_ $@L _) $@ _). - 1: exact cat_binbiprod_inr. - { apply cat_binbiprod_pr_eta. - - refine (_ $@ cat_binbiprod_pr1_inr^$). - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. - refine (cat_assoc _ _ _ $@ _). - apply cat_zero_l. - - refine (_ $@ cat_binbiprod_pr2_inr^$). - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. - refine (cat_assoc _ _ _ $@ _). - refine (cat_idl _ $@ _). - apply cat_binbiprod_corec_beta_pr2. } - nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - { refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_rec. - apply cat_binbiprod_rec_beta_inr. } - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idl. - Defined. - - Local Existing Instance symmetricbraiding_cat_binbiprod. - - (** Using the naturality of swap we can show that the operation is commutative. *) - Local Instance commutative_hom : Commutative sgop_hom. - Proof. - intros f g. - apply path_hom. - refine (cat_assoc _ _ _ $@ _). - refine ((_^$ $@R _) $@ _). - 1: apply cat_binbiprod_swap_codiag. - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). - refine (_ $@ (_ $@L _)). - 2: apply cat_binbiprod_swap_diag. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ cat_assoc _ _ _). - rapply Monoidal.braid_nat. - Defined. - - Local Instance right_identity_ab_hom : RightIdentity sgop_hom zero_hom. - Proof. - intros f. - refine (commutativity _ _ @ _). - apply left_identity_hom. - Defined. - - Local Instance associative_ab_hom : Associative sgop_hom. - Proof. - (** There is a lot to unfold here. But drawing out the diagram we can split this into 5 separate diagrams that all commute. Much of the work here will be applying associativity and functor laws to get this diagram in the correct shape. *) - intros f g h. - apply path_hom. - refine (((_ $@L (_ $@R _)) $@R _) $@ _ $@ ((_ $@L (_ $@L _)^$) $@R _)). - 1,3: refine (fmap_comp _ _ _ $@ (_ $@R _)). - 1,2: refine (fmap_comp _ _ _ $@ (_ $@L _)). - 1,2: rapply fmap_comp. - do 2 change (fmap (flip (fun x y => cat_binbiprod x y) ?a) ?f) - with (fmap10 (fun x y => cat_binbiprod x y) f a); - do 2 change (fmap ((fun x y => cat_binbiprod x y) ?b) ?f) - with (fmap01 (fun x y => cat_binbiprod x y) b f). - Local Notation "a ⊕L g" := (fmap01 (fun x y => cat_binbiprod x y) a g) (at level 99). - Local Notation "f ⊕R b" := (fmap10 (fun x y => cat_binbiprod x y) f b) (at level 99). - Local Notation "'Δ' x" := (cat_binbiprod_diag x) (at level 10). - Local Notation "'∇' x" := (cat_binbiprod_codiag x) (at level 10). - change (?w $o ?x $== ?y $o ?z) with (Square z w x y). - apply move_right_top. - (** The following associativity rewrites are specially chosen so that the diagram can be decomposed easily. *) - rapply vconcatL. - 1: do 2 refine (cat_assoc _ _ _ $@ _). - 1: refine (((cat_assoc _ _ _)^$ $@R _)). - rapply vconcatR. - 2: do 2 refine ((cat_assoc _ _ _)^$ $@ (_ $@L (cat_assoc _ _ _)) $@ _). - 2: refine (cat_assoc _ _ _)^$. - srapply hconcat. - (** We now have to specify the morphism that we are composing along, i.e. the common edge of the two squares where they meet. This is chosen to be the associativity equivalence for biproducts. *) - 1: nrapply cate_binbiprod_assoc. - { nrefine ((_ $@R _) $@ _). - 1: rapply Monoidal.associator_twist'_unfold. - apply cat_binbiprod_pr_eta. - - srefine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: exact (fmap01 (fun x y => cat_binbiprod x y) b cat_binbiprod_pr1). - { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. - refine (_^$ $@ _). - 1: rapply fmap01_comp. - refine (fmap02 _ _ _). - apply cat_binbiprod_corec_beta_pr2. } - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((((fmap02 _ _ _ $@ fmap01_id _ _ _)^$ $@ fmap01_comp _ _ _ _)^$ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. - nrefine (cat_idl _ $@ _). - refine (_ $@ (_ $@L (cat_assoc _ _ _)^$)). - nrefine (_ $@ cat_assoc _ _ _). - refine (_ $@ (_^$ $@R _)). - 2: rapply cat_binbiprod_corec_beta_pr1. - refine ((_ $@L _) $@ (cat_assoc _ _ _)^$). - nrefine (_ $@ cat_assoc _ _ _). - refine (_ $@ (_^$ $@R _)). - 2: rapply cat_binbiprod_corec_beta_pr1. - refine (_ $@ (cat_assoc _ _ _)^$). - refine (_ $@ (_ $@L _^$)). - 2: rapply cat_binbiprod_corec_beta_pr1. - symmetry. - apply cat_idr. - - nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). - refine ((cat_assoc _ _ _)^$ $@ _). - nrefine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. - nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). - refine ((cat_assoc _ _ _)^$ $@ _). - nrefine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. - nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - apply cat_binbiprod_corec_beta_pr2. } - nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). - refine ((cat_assoc _ _ _)^$ $@ _). - nrefine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. - nrefine ((_ $@L _) $@ _). - { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - apply cat_binbiprod_corec_beta_pr2. } - nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). - refine ((cat_assoc _ _ _)^$ $@ _). - nrefine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. - nrefine (cat_idl _ $@ _). - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. - nrefine (cat_assoc _ _ _ $@ _). - nrefine (cat_idl _ $@ _). - nrefine (cat_binbiprod_corec_beta_pr2 _ _ $@ _). - refine (_ $@ (_ $@L (cat_assoc _ _ _)^$)). - nrefine (_ $@ cat_assoc _ _ _). - refine (_ $@ (_^$ $@R _)). - 2: apply cat_binbiprod_corec_beta_pr2. - refine (_ $@ (cat_assoc _ _ _)^$). - refine (_ $@ (cat_idl _)^$). - refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: apply cat_binbiprod_corec_beta_pr2. - refine (_ $@ (cat_assoc _ _ _)^$). - refine (_ $@ (cat_idl _)^$). - symmetry. - apply cat_binbiprod_corec_beta_pr2. } - (** We split the square again, this time by picking yet another associativity map as the common edge. This leaves us with a naturality square for associativity. *) - srapply hconcat. - 1: nrapply cate_binbiprod_assoc. - 1: nrapply Monoidal.associator_nat_m. - (** Finally we are left with another polygon which we resolve by brute force. *) - refine ((cat_assoc _ _ _)^$ $@ _). - apply cat_binbiprod_in_eta. - - nrefine (cat_assoc _ _ _ $@ _). - nrefine ((_ $@L _) $@ _). - { nrefine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inl. } - refine ((cat_assoc _ _ _)^$ $@ _). - refine (cat_idr _ $@ _). - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inl. } - refine ((cat_assoc _ _ _)^$ $@ _). - refine (cat_idr _ $@ _). - refine (_ $@ _). - 1: apply cat_binbiprod_rec_beta_inl. - symmetry. - refine (cat_assoc _ _ _ $@ _). - refine ((_ $@L _) $@ _). - 1: apply cate_binbiprod_assoc_inl. - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - { refine (cat_assoc _ _ _ $@ (_ $@L _)). - refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inl. } - simpl. - refine (cat_assoc _ _ _ $@ _). - refine ((_ $@L _) $@ _). - { refine (cat_assoc _ _ _ $@ (_ $@L _)). - apply cat_binbiprod_rec_beta_inl. } - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - { refine (cat_assoc _ _ _ $@ (_ $@L _)). - refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). - 1: apply cat_binbiprod_rec_beta_inl. - apply cat_idr. } - refine (cat_idr _ $@ _). - apply cat_binbiprod_rec_beta_inl. - - nrefine (cat_assoc _ _ _ $@ _). - nrefine ((_ $@L _) $@ _). - { nrefine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inr. } - refine ((cat_assoc _ _ _)^$ $@ _). - simpl. - refine ((_ $@R _) $@ _). - { refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - 1: refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idl. } - simpl. - refine (_ $@ (cat_assoc _ _ _)^$). - apply cat_binbiprod_in_eta. - + refine (cat_assoc _ _ _ $@ _). - refine ((_ $@L _) $@ _). - { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inl. } - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inl. - refine (cat_idr _ $@ _). - refine (_ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). - 2: apply cate_binbiprod_assoc_inr_inl. - refine (_ $@ (cat_assoc _ _ _)^$). - refine (_ $@ (_ $@L _^$)). - 2: { refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inl. - simpl. - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idr. } - simpl. - refine (_ $@ (cat_assoc _ _ _)^$). - refine (_^$ $@ (_ $@L _^$)). - 2: { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). - 1: apply cat_binbiprod_rec_beta_inl. - apply cat_idr. } - apply cat_binbiprod_rec_beta_inl. - + refine (cat_assoc _ _ _ $@ _). - refine ((_ $@L _) $@ _). - { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inr. } - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. - refine (cat_idl _ $@ _). - refine (_ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). - 2: apply cate_binbiprod_assoc_inr_inr. - simpl. - refine (_ $@ (cat_assoc _ _ _)^$). - refine (_^$ $@ (_ $@L _^$)). - 2: { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idr. } - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. } - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idl. - Defined. - - Local Instance issemigroup_hom : IsSemiGroup (a $-> b) := {}. - Local Instance ismonoid_hom : IsMonoid (a $-> b) := {}. - -End CMonHom. - -(** ** Additive Categories *) - -(** *** Definition *) - -Class IsAdditive (A : Type) `{HasEquivs A} := { - additive_semiadditive :: IsSemiAdditive A; - additive_negate_hom : forall (a b : A), Negate (a $-> b); - additive_left_inverse_hom : forall (a b : A), - LeftInverse (sgop_hom a b) (additive_negate_hom a b) (zero_hom a b); - additive_right_inverse_hom : forall (a b : A), - RightInverse (sgop_hom a b) (additive_negate_hom a b) (zero_hom a b); -}. - -(** *** Abelian Group structure on Hom *) - -(** Homs in an aditive category form an abelian group. *) -Definition AbHom {A : Type} `{IsAdditive A} : A -> A -> AbGroup. -Proof. - intros a b. - snrapply Build_AbGroup. - 1: snrapply Build_Group. - 5: split. - - exact (a $-> b). - - exact (sgop_hom a b). - - exact (zero_hom a b). - - exact (additive_negate_hom a b). - - exact (ismonoid_hom a b). - - exact (additive_left_inverse_hom a b). - - exact (additive_right_inverse_hom a b). - - exact (commutative_hom a b). -Defined. - -(** *** Endomorphism ring *) - -(** The composition operation provides a multiplicative operation on the abelian group hom. Turning it into a ring of endomorphsims. *) -Definition End {A : Type} `{IsAdditive A} (X : A) : Ring. -Proof. - snrapply Build_Ring'; repeat split. - - exact (AbHom X X). - - exact (fun f g => f $o g). - - exact (Id _). - - intros f g h. - apply path_hom. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _ $@ cat_assoc _ _ _). - 1: apply cat_binbiprod_codiag_fmap11. - refine (cat_assoc _ _ _ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). - rapply fmap11_comp. - - intros f g h. - apply path_hom. - refine (cat_assoc _ _ _ $@ _). - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). - refine ((_ $@L _) $@ _). - 1: apply cat_binbiprod_diag_fmap11. - refine ((cat_assoc _ _ _)^$ $@ (_^$ $@R _)). - rapply fmap11_comp. - - exact _. - - intros f g h. - apply path_hom. - symmetry. - apply cat_assoc. - - intros f. - apply path_hom. - apply cat_idl. - - intros g. - apply path_hom. - apply cat_idr. -Defined. - -(** TODO: Additive functors *) From 5e5023367d65298afc1c66af2277185c64fcb14a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 25 Apr 2024 15:06:09 +0100 Subject: [PATCH 04/31] shorten proof of swap and inl/inr Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Biproducts.v | 30 ++++------------------- 1 file changed, 5 insertions(+), 25 deletions(-) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 4cd212874e0..88e10f44b88 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -540,39 +540,19 @@ Section Symmetry. apply cat_binbiprod_corec_beta_pr2. Defined. - (** The swap map preserves the codiagonal. TODO: proof can probably be shorter. *) + (** The swap map preserves the codiagonal. *) Lemma cat_binbiprod_swap_codiag (x : A) : cat_binbiprod_codiag x $o cat_binbiprod_swap x x $== cat_binbiprod_codiag x. Proof. apply cat_binbiprod_in_eta. - refine (_ $@ (cat_binbiprod_rec_beta_inl _ _)^$). refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - + apply cat_binbiprod_pr_eta. - * refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. - refine (_ $@ _^$). - 1: apply cat_binbiprod_pr2_inl. - apply cat_binbiprod_pr1_inr. - * refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. - refine (_ $@ _^$). - 1: apply cat_binbiprod_pr1_inl. - apply cat_binbiprod_pr2_inr. - + apply cat_binbiprod_rec_beta_inr. + 1: apply cat_binbiprod_swap_inl. + apply cat_binbiprod_rec_beta_inr. - refine (_ $@ (cat_binbiprod_rec_beta_inr _ _)^$). refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - + apply cat_binbiprod_pr_eta. - * refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. - refine (_ $@ _^$). - 1: apply cat_binbiprod_pr2_inr. - apply cat_binbiprod_pr1_inl. - * refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. - refine (_ $@ _^$). - 1: apply cat_binbiprod_pr1_inr. - apply cat_binbiprod_pr2_inl. - + apply cat_binbiprod_rec_beta_inl. + 1: apply cat_binbiprod_swap_inr. + apply cat_binbiprod_rec_beta_inl. Defined. End Symmetry. From fb05a7719f7a3f0e67406863077b39885cdd8b09 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 25 Apr 2024 18:55:18 +0100 Subject: [PATCH 05/31] more comments + cleanup Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Additive.v | 21 +++++++++++++-------- 1 file changed, 13 insertions(+), 8 deletions(-) diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index 3f9aad37d4f..db27baa9d10 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -10,12 +10,19 @@ Require Import canonical_names. (** ** Semiadditive Categories *) +(** As semiadditive category is a a wild category with equivalences and the following data: *) Class IsSemiAdditive (A : Type) `{HasEquivs A} := { + (** It is a pointed category. *) semiadditive_pointed :: IsPointedCat A; + (** It has binary biproducts. *) semiadditive_hasbiproducts :: HasBinaryBiproducts A; + (** We have morphism extensionality. *) semiadditive_hasmorext :: HasMorExt A; + (** The homs are set. *) semiadditive_ishset_hom :: forall (a b : A), IsHSet (a $-> b); }. + +(** The final two conditions in the definition of a semiadditive category ensure that the hom types can become commutative monoids. This is an essential chracteristic of semiadditive categories making it equivalent to alternate defintiions where the category is semiadditive if it is enriched in commutative monoids. The machinary of encriched categories however is a bit heavy so we use this more lightweight definition where the commutative monoid structure appears naturally. *) Section CMonHom. @@ -23,8 +30,11 @@ Section CMonHom. (** We can show that the hom-types of a semiadditive category are commutative monoids. *) + (** The zero element is given by the zero morphism. This exists from our pointedness assumption. *) Local Instance zero_hom : Zero (a $-> b) := zero_morphism. + (** TODO: explain a bit more: diagonal duplicates and codiagonal sums. *) + (** The operation is given by the biproduct structure. *) Local Instance sgop_hom : SgOp (a $-> b) := fun f g => cat_binbiprod_codiag b $o fmap11 (fun x y => cat_binbiprod x y) f g @@ -34,7 +44,6 @@ Section CMonHom. Proof. intros g. apply path_hom. - unfold sgop_hom, fmap11; cbn. refine (cat_assoc _ _ _ $@ _). refine ((_ $@L cat_assoc _ _ _) $@ _). refine ((cat_assoc _ _ _)^$ $@ _). @@ -93,13 +102,9 @@ Section CMonHom. intros f g h. apply path_hom. refine (((_ $@L (_ $@R _)) $@R _) $@ _ $@ ((_ $@L (_ $@L _)^$) $@R _)). - 1,3: refine (fmap_comp _ _ _ $@ (_ $@R _)). - 1,2: refine (fmap_comp _ _ _ $@ (_ $@L _)). - 1,2: rapply fmap_comp. - do 2 change (fmap (flip (fun x y => cat_binbiprod x y) ?a) ?f) - with (fmap10 (fun x y => cat_binbiprod x y) f a); - do 2 change (fmap ((fun x y => cat_binbiprod x y) ?b) ?f) - with (fmap01 (fun x y => cat_binbiprod x y) b f). + 1,3: refine (_ $@ ((_ $@ (_ $@L _)) $@R _)). + 1-3: rapply fmap01_comp. + 1-3: rapply fmap10_comp. Local Notation "a ⊕L g" := (fmap01 (fun x y => cat_binbiprod x y) a g) (at level 99). Local Notation "f ⊕R b" := (fmap10 (fun x y => cat_binbiprod x y) f b) (at level 99). Local Notation "'Δ' x" := (cat_binbiprod_diag x) (at level 10). From b96d1299744e8a283cb330cd56243788ce2576a8 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sat, 27 Apr 2024 01:14:52 +0100 Subject: [PATCH 06/31] abelian categories Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Abelian.v | 144 ++++++++++++++++++++ theories/Algebra/Homological/Additive.v | 154 +++++++++++++++++++--- theories/Algebra/Homological/Biproducts.v | 83 +++++++++++- theories/WildCat/Coproducts.v | 19 +++ theories/WildCat/Opposite.v | 15 +++ 5 files changed, 391 insertions(+), 24 deletions(-) create mode 100644 theories/Algebra/Homological/Abelian.v diff --git a/theories/Algebra/Homological/Abelian.v b/theories/Algebra/Homological/Abelian.v new file mode 100644 index 00000000000..71a5ec69161 --- /dev/null +++ b/theories/Algebra/Homological/Abelian.v @@ -0,0 +1,144 @@ +Require Import Basics.Overture Basics.Tactics Basics.Equivalences. +Require Import WildCat.Core WildCat.Equiv WildCat.PointedCat WildCat.Opposite. +Require Import Algebra.AbGroups.AbelianGroup Algebra.Rings.Ring. +Require Import Algebra.Homological.Additive. +Require Import canonical_names. + +(** * Abelian Categories *) + +Local Open Scope mc_scope. +Local Open Scope mc_add_scope. + +(** ** Kernels and cokernels *) + +(** *** Kernels *) + +Class Kernel {A : Type} `{IsPointedCat A} {a b : A} (f : a $-> b) := { + cat_ker_obj : A; + cat_ker : cat_ker_obj $-> a; + + cat_ker_zero : f $o cat_ker $== zero_morphism; + + cat_ker_corec {x} (g : x $-> a) : f $o g $== zero_morphism -> x $-> cat_ker_obj; + + cat_ker_corec_beta {x} (g : x $-> a) (p : f $o g $== zero_morphism) + : cat_ker $o cat_ker_corec g p $== g; + + monic_cat_ker : Monic cat_ker; +}. + +Arguments cat_ker_obj {A _ _ _ _ _ a b} f {_}. +Arguments cat_ker_corec {A _ _ _ _ _ a b f _ x} g p. +Arguments cat_ker {A _ _ _ _ _ a b} f {_}. +Arguments cat_ker_zero {A _ _ _ _ _ a b} f {_}. +Arguments monic_cat_ker {A _ _ _ _ _ a b} f {_}. + +(** Kernels of monomorphisms are zero. *) +Definition ker_zero_monic {A : Type} `{IsPointedCat A} + {a b : A} (f : a $-> b) `{!Kernel f} + : Monic f -> cat_ker f $== zero_morphism. +Proof. + intros monic. + apply monic. + refine (cat_ker_zero f $@ _^$). + apply cat_zero_r. +Defined. + +(** Maps with zero kernel are monic. *) +Definition monic_ker_zero {A : Type} `{IsAdditive A} + {a b : A} (f : a $-> b) `{!Kernel f} + : cat_ker f $== zero_morphism -> Monic f. +Proof. + intros ker_zero c g h p. + apply GpdHom_path. + apply path_hom in p. + change (@paths (AbHom c a) g h). + change (@paths (AbHom c b) (f $o g) (f $o h)) in p. + apply grp_moveL_1M. + apply grp_moveL_1M in p. + apply GpdHom_path in p. + apply path_hom. + refine ((cat_ker_corec_beta (g + (-h)) _)^$ $@ (ker_zero $@R _) $@ cat_zero_l _). + refine (addcat_dist_l _ _ _ $@ _ $@ p). + apply GpdHom_path. + f_ap. + apply path_hom. + symmetry. + apply addcat_comp_negate_r. +Defined. + +(** *** Cokernels *) + +Class Cokernel {A : Type} `{IsPointedCat A} {a b : A} (f : a $-> b) + := cokernel_kernel_op :: Kernel (A := A^op) (f : Hom (A:= A^op) b a). + +Arguments Cokernel {A _ _ _ _ _ a b} f. + +Definition cat_coker_obj {A : Type} `{IsPointedCat A} {a b : A} (f : a $-> b) + `{!Cokernel f} : A + := cat_ker_obj (A:=A^op) (a:=b) f. + +Definition cat_coker {A : Type} `{IsPointedCat A} {a b : A} (f : a $-> b) + `{!Cokernel f} : b $-> cat_coker_obj f + := cat_ker (A:=A^op) (a:=b) (b:=a) f. + +Definition cat_coker_zero {A : Type} `{IsPointedCat A} {a b : A} (f : a $-> b) + `{!Cokernel f} : cat_coker f $o f $== zero_morphism + := cat_ker_zero (A:=A^op) (a:=b) f. + +Definition cat_coker_rec {A : Type} `{IsPointedCat A} {a b : A} {f : a $-> b} + `{!Cokernel f} {x} (g : b $-> x) + : g $o f $== zero_morphism -> cat_coker_obj f $-> x + := cat_ker_corec (A:=A^op) (a:=b) (b:=a) g. + +Definition cat_coker_rec_beta {A : Type} `{IsPointedCat A} {a b : A} (f : a $-> b) + `{!Cokernel f} {x} (g : b $-> x) (p : g $o f $== zero_morphism) + : cat_coker_rec g p $o cat_coker f $== g + := cat_ker_corec_beta (A:=A^op) (a:=b) (b:=a) g p. + +Definition epic_cat_coker {A : Type} `{IsPointedCat A} + {a b : A} (f : a $-> b) `{!Cokernel f} + : Epic (cat_coker f) + := monic_cat_ker (A:=A^op) (a:=b) (b:=a) f. + +Definition coker_zero_epic {A : Type} `{IsPointedCat A} + {a b : A} (f : a $-> b) `{!Cokernel f} + : Epic f -> cat_coker f $== zero_morphism + := ker_zero_monic (A:=A^op) (a:=b) f. + +(** Funext required all the way in Biproducts to show opposite of additive is additivef. *) +Definition epic_coker_zero `{Funext} {A : Type} `{IsAdditive A} + {a b : A} (f : a $-> b) {k : Cokernel f} + : cat_coker f $== zero_morphism -> Epic f + := monic_ker_zero (A:=A^op) (a:=b) (b:=a) f. + +(** ** Preabelian categories *) + +Class IsPreAbelian {A : Type} `{IsAdditive A} := { + preabelian_has_kernels :: forall {a b : A} (f : a $-> b), Kernel f; + preabelian_has_cokernels :: forall {a b : A} (f : a $-> b), Cokernel f; +}. + +Definition ispreabelian_canonical_map {A : Type} `{IsPreAbelian A} + {a b : A} (f : a $-> b) + : cat_coker_obj (cat_ker f) $-> cat_ker_obj (cat_coker f). +Proof. + snrapply cat_coker_rec. + - snrapply cat_ker_corec. + + exact f. + + apply cat_coker_zero. + - snrapply monic_cat_ker. + refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ _). + 1: apply cat_ker_corec_beta. + refine (_ $@ (cat_zero_r _)^$). + apply cat_ker_zero. +Defined. + +(** ** Abelian categories *) + +Class IsAbelian {A : Type} `{IsPreAbelian A} := { + catie_preabelian_canonical_map : forall a b (f : a $-> b), + CatIsEquiv (ispreabelian_canonical_map f); +}. + diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index db27baa9d10..9efd7e12cf9 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -1,7 +1,8 @@ Require Import Basics.Overture. Require Import WildCat.Core WildCat.Equiv WildCat.PointedCat WildCat.Bifunctor. -Require Import WildCat.Square. +Require Import WildCat.Square WildCat.Opposite. Require Import Algebra.Homological.Biproducts. +Require Import Algebra.Groups.Group. Require Import Algebra.AbGroups.AbelianGroup Algebra.Rings.Ring. Require Import canonical_names. @@ -35,7 +36,7 @@ Section CMonHom. (** TODO: explain a bit more: diagonal duplicates and codiagonal sums. *) (** The operation is given by the biproduct structure. *) - Local Instance sgop_hom : SgOp (a $-> b) := fun f g => + Global Instance sgop_hom : Plus (a $-> b) := fun f g => cat_binbiprod_codiag b $o fmap11 (fun x y => cat_binbiprod x y) f g $o cat_binbiprod_diag a. @@ -313,10 +314,10 @@ End CMonHom. Class IsAdditive (A : Type) `{HasEquivs A} := { additive_semiadditive :: IsSemiAdditive A; - additive_negate_hom : forall (a b : A), Negate (a $-> b); - additive_left_inverse_hom : forall (a b : A), + additive_negate_hom :: forall (a b : A), Negate (a $-> b); + additive_left_inverse_hom :: forall (a b : A), LeftInverse (sgop_hom a b) (additive_negate_hom a b) (zero_hom a b); - additive_right_inverse_hom : forall (a b : A), + additive_right_inverse_hom :: forall (a b : A), RightInverse (sgop_hom a b) (additive_negate_hom a b) (zero_hom a b); }. @@ -339,30 +340,108 @@ Proof. - exact (commutative_hom a b). Defined. +(** *** Distributivity of composition over addition *) + +Definition addcat_dist_l {A : Type} `{IsAdditive A} {a b c : A} + (f : b $-> c) (g h : a $-> b) + : f $o (g + h) $== (f $o g) + (f $o h). +Proof. + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _ $@ cat_assoc _ _ _). + 1: apply cat_binbiprod_codiag_fmap11. + refine (cat_assoc _ _ _ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). + rapply fmap11_comp. +Defined. + +(** Note that this is more general than the distributive law in the endomorphism ring. *) +Global Instance left_heterodistribute_hom {A : Type} `{IsAdditive A} {a b c : A} + : LeftHeteroDistribute cat_comp (sgop_hom a b) (sgop_hom a c). +Proof. + intros f g h. + apply path_hom. + apply addcat_dist_l. +Defined. + +Definition addcat_dist_r {A : Type} `{IsAdditive A} {a b c : A} + (f g : b $-> c) (h : a $-> b) + : (f + g) $o h $== (f $o h) + (g $o h). +Proof. + refine (cat_assoc _ _ _ $@ _). + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). + refine ((_ $@L _) $@ _). + 1: apply cat_binbiprod_diag_fmap11. + refine ((cat_assoc _ _ _)^$ $@ (_^$ $@R _)). + rapply fmap11_comp. +Defined. + +Global Instance right_heterodistribute_hom {A : Type} `{IsAdditive A} {a b c : A} + : RightHeteroDistribute cat_comp (sgop_hom a c) (sgop_hom b c). +Proof. + intros f g h. + apply path_hom. + apply addcat_dist_r. +Defined. + +Definition addcat_comp_negate_id_l {A : Type} `{IsAdditive A} {a b : A} + (f : a $-> b) + : (- Id _) $o f $== -f. +Proof. + apply GpdHom_path. + change (@paths (Hom a b) ?x ?y) with (@paths (AbHom a b : Type) x y). + rapply grp_moveL_1V. + transitivity ((-Id b $o f) + (Id b $o f)). + - f_ap; symmetry; apply path_hom, cat_idl. + - lhs_V rapply right_heterodistribute_hom. + transitivity (zero_morphism (b := b) $o f). + + f_ap; apply (grp_inv_l (Id b : AbHom b b)). + + apply path_hom. + apply cat_zero_l. +Defined. + +Definition addcat_comp_negate_id_r {A : Type} `{IsAdditive A} {a b : A} + (f : a $-> b) + : f $o (- Id _) $== -f. +Proof. + apply GpdHom_path. + change (@paths (Hom a b) ?x ?y) with (@paths (AbHom a b : Type) x y). + rapply grp_moveL_1V. + transitivity ((f $o -Id a) + (f $o Id a)). + - f_ap; symmetry; apply path_hom, cat_idr. + - lhs_V rapply left_heterodistribute_hom. + transitivity (f $o zero_morphism (a := a)). + + f_ap; apply (grp_inv_l (Id a : AbHom a a)). + + apply path_hom. + apply cat_zero_r. +Defined. + +Definition addcat_comp_negate_l {A : Type} `{IsAdditive A} {a b c : A} + (f : b $-> c) (g : a $-> b) + : - (f $o g) $== (- f) $o g. +Proof. + refine ((addcat_comp_negate_id_l _)^$ $@ _). + refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + apply addcat_comp_negate_id_l. +Defined. + +Definition addcat_comp_negate_r {A : Type} `{IsAdditive A} {a b c : A} + (f : b $-> c) (g : a $-> b) + : - (f $o g) $== f $o (- g). +Proof. + refine ((addcat_comp_negate_id_r _)^$ $@ _). + refine (cat_assoc _ _ _ $@ (_ $@L _)). + apply addcat_comp_negate_id_r. +Defined. + (** *** Endomorphism ring *) -(** The composition operation provides a multiplicative operation on the abelian group hom. Turning it into a ring of endomorphsims. *) Definition End {A : Type} `{IsAdditive A} (X : A) : Ring. Proof. snrapply Build_Ring'; repeat split. - exact (AbHom X X). - exact (fun f g => f $o g). - exact (Id _). - - intros f g h. - apply path_hom. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _ $@ cat_assoc _ _ _). - 1: apply cat_binbiprod_codiag_fmap11. - refine (cat_assoc _ _ _ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). - rapply fmap11_comp. - - intros f g h. - apply path_hom. - refine (cat_assoc _ _ _ $@ _). - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). - refine ((_ $@L _) $@ _). - 1: apply cat_binbiprod_diag_fmap11. - refine ((cat_assoc _ _ _)^$ $@ (_^$ $@R _)). - rapply fmap11_comp. + - exact left_heterodistribute_hom. + - exact right_heterodistribute_hom. - exact _. - intros f g h. apply path_hom. @@ -376,4 +455,37 @@ Proof. apply cat_idr. Defined. +(** ** Properties of additive categories *) + +(** The opposite of a semiadditive category is semiadditive. *) +Global Instance issemiadditive_op `{Funext} {A : Type} `{IsSemiAdditive A} + : IsSemiAdditive A^op. +Proof. + snrapply Build_IsSemiAdditive. + - exact _. + - exact _. + - exact _. + - intros a b. + change A in a, b. + change (IsHSet (b $-> a)). + exact _. +Defined. + +Axiom sorry : Empty. +Ltac sorry := apply Empty_rect; apply sorry. + +(** The opposite of an additive category is additive. *) +Global Instance isadditive_op `{F:Funext} {A : Type} `{H : IsAdditive A} + : IsAdditive A^op. +Proof. + snrapply Build_IsAdditive. + - exact _. + - intros a b. + change A in a, b. + change (Negate (b $-> a)). + exact _. + - sorry. + - sorry. +Defined. + (** TODO: Additive functors *) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 88e10f44b88..9c38a565632 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -1,7 +1,8 @@ -Require Import Basics.Overture Basics.Decidable Basics.Tactics Basics.Trunc. -Require Import Types.Forall Types.Bool. +Require Import Basics.Overture Basics.Decidable Basics.Tactics Basics.Trunc Basics.Equivalences. +Require Import Types.Forall Types.Bool Types.Paths Types.Empty Types.Equiv Types.Sigma. Require Import WildCat.Core WildCat.Products WildCat.Coproducts WildCat.Equiv. -Require Import WildCat.PointedCat WildCat.Bifunctor WildCat.Square WildCat.Monoidal. +Require Import WildCat.PointedCat WildCat.Bifunctor WildCat.Square. +Require Import WildCat.Opposite WildCat.Monoidal. (** * Categories with biproducts *) @@ -730,3 +731,79 @@ Section Associativity. Defined. End Associativity. + +(** *** Biproducts in the opposite category *) + +(** Biproducts exist in the opposite category if they exist in the original category. *) +Global Instance biproduct_op `{Funext} {I A : Type} {x : I -> A} `{Biproduct I A x} + : Biproduct (A:=A^op) I x. +Proof. + snrapply Build_Biproduct'. + (** Products in the opposite category are coproducts in the original category. *) + - exact _. + (** Coproducts in the opposite category are products in the original category. *) + - apply coproduct_op. + exact _. + - snrapply catie_homotopic. + + simpl; exact (cat_coprod_prod_incl (A:=A) x). + + simpl; exact _. + + (** Showing that these two maps are homotopic is a bit tricky. It boils down to showing an equality between [dec_paths (i = j)] and [dec_paths (j = i)] which is true with [Funext] but it is not certain if it holds without. *) + apply cat_coprod_in_eta. + intros i. + apply cat_prod_pr_eta. + intros j. + simpl. + refine (cat_assoc _ _ _ $@ _). + refine ((_ $@L _) $@ _). + 1: apply cat_coprod_beta. + refine (_ $@ _). + 1: rapply cat_prod_beta. + symmetry. + refine ((_ $@R _) $@ _). + 1: apply cat_prod_beta. + refine (_ $@ _). + 1: rapply cat_coprod_beta. + simpl. + generalize (dec_paths i j). + generalize (dec_paths j i). + intros p q. + assert (r : @decidable_equiv (i = j) (j = i) inverse _ q = p). + { destruct p as [p|np]. + - destruct q as [q|nq]. + + apply (ap inl). + apply path_ishprop. + + destruct (nq p^). + - destruct q as [|nq]. + + destruct (np p^). + + apply (ap inr). + (** Seems to be essential. *) + apply path_forall. + intros p. + destruct (np p). } + destruct r. + destruct q as [[]|q]; reflexivity. +Defined. + +Global Instance hasbiproducts_op `{Funext} {I A : Type} `{DecidablePaths I, HasEquivs A, !IsPointedCat A, !HasBiproducts I A} + : HasBiproducts I (A^op). +Proof. + intros x. + by rapply biproduct_op. +Defined. + +Global Instance binarybiproduct_op `{Funext} {A : Type} + `{HasEquivs A, !IsPointedCat A} {x y : A} {bb : BinaryBiproduct x y} + : BinaryBiproduct (A:=A^op) x y. +Proof. + nrapply biproduct_op. + exact bb. +Defined. + +Global Instance hasbinarybiproducts_op `{Funext} {A : Type} + `{HasEquivs A, !IsPointedCat A} {hbb : HasBinaryBiproducts A} + : HasBinaryBiproducts (A^op). +Proof. + intros x y. + rapply biproduct_op. + exact (hbb x y). +Defined. diff --git a/theories/WildCat/Coproducts.v b/theories/WildCat/Coproducts.v index 34ad81b9829..57e708290e9 100644 --- a/theories/WildCat/Coproducts.v +++ b/theories/WildCat/Coproducts.v @@ -303,3 +303,22 @@ Proof. exact (Id _). - apply zero_morphism. Defined. + +(** *** Coproducts in the opposite category *) + +Definition coproduct_op {I A : Type} (x : I -> A) + `{Is1Cat A} {H' : Product I x} + : Coproduct I (A:=A^op) x. +Proof. + snrapply Build_Product. + - exact (cat_prod I x). + - exact cat_pr. + - exact (fun z => cat_prod_corec I). + - intros z f i. + apply cat_prod_beta. + - intros z f g p. + apply cat_prod_pr_eta. + exact p. +Defined. + +Hint Immediate coproduct_op : typeclass_instances. diff --git a/theories/WildCat/Opposite.v b/theories/WildCat/Opposite.v index a862bdeca1c..710282d65b3 100644 --- a/theories/WildCat/Opposite.v +++ b/theories/WildCat/Opposite.v @@ -5,6 +5,7 @@ Require Import WildCat.Core. Require Import WildCat.Equiv. Require Import WildCat.NatTrans. Require Import WildCat.FunctorCat. +Require Import WildCat.PointedCat. (** ** Opposite categories *) @@ -200,3 +201,17 @@ Proof. rapply is1nat_op. Defined. +Global Instance isinitial_isterminal_op {A : Type} `{Is1Cat A} (x : A) + {t : IsTerminal x} : IsInitial (A := A^op) x + := t. + +Global Instance isterminal_isinitial_op {A : Type} `{Is1Cat A} (x : A) + {i : IsInitial x} : IsTerminal (A := A^op) x + := i. + +Global Instance ispointedcat_op {A : Type} `{IsPointedCat A} : IsPointedCat A^op. +Proof. + snrapply Build_IsPointedCat. + 1: unfold op; exact zero_object. + 1,2: exact _. +Defined. From 62280579c48a24316ab9f11e74fb772a0a7d4979 Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Sat, 27 Apr 2024 12:23:37 -0400 Subject: [PATCH 07/31] Abstract out some idioms for dealing with decidable types --- theories/Algebra/Homological/Biproducts.v | 15 +++---- theories/Basics/Decidable.v | 49 +++++++++++++++++++++++ 2 files changed, 55 insertions(+), 9 deletions(-) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 9c38a565632..df4108fee6d 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -113,12 +113,9 @@ Proof. { refine ((cate_buildequiv_fun _ $@R _) $@ _). apply cat_coprod_beta. } refine (cat_prod_beta _ _ _ $@ _). - destruct (dec_paths i i); simpl. - - (** We cannot eliminate [p : i = i] with path induction, but we do have decidable equality, hence by Hedberg's theorem we must be in a hset and we can replace this with the identity. *) - assert (r : (idpath = p)) by apply path_ishprop. - by destruct r. - - contradiction n. - reflexivity. + simpl. + generalize (dec_paths i i). + by nrapply decidable_paths_refl. Defined. (** An inclusion followed by a projection of a different index is zero. *) @@ -131,9 +128,9 @@ Proof. { refine ((cate_buildequiv_fun _ $@R _) $@ _). apply cat_coprod_beta. } refine (cat_prod_beta _ _ _ $@ _). - destruct (dec_paths i j); simpl. - - contradiction p. - - reflexivity. + decidable_false (dec_paths i j) p. + simpl. + reflexivity. Defined. Definition cat_biprod_diag I {A} (x : A) `{Biproduct I A (fun _ => x)} diff --git a/theories/Basics/Decidable.v b/theories/Basics/Decidable.v index fe4d7927216..8c3aef256fa 100644 --- a/theories/Basics/Decidable.v +++ b/theories/Basics/Decidable.v @@ -31,6 +31,40 @@ Ltac decide := | [|- ?A] => decide_type A end. +Definition decidable_true {A : Type} `{Decidable A} + (a : A) + (P : forall (p : Decidable A), Type) + (p : forall x, P (inl x)) + : forall p, P p. +Proof. + intros [x|n]. + - apply p. + - contradiction n. +Defined. + +(** Replace a term [p] of the form [Decidable A] with [inl x] if we know [A] is true. *) +Ltac decidable_true p a := + generalize p; + rapply (decidable_true a); + try intro. + +Definition decidable_false {A : Type} `{Decidable A} + (n : not A) + (P : forall (p : Decidable A), Type) + (p : forall n', P (inr n')) + : forall p, P p. +Proof. + intros [x|n']. + - contradiction n. + - apply p. +Defined. + +(** Replace a term [p] of the form [Decidable A] with [inr na] if we know [A] is false. *) +Ltac decidable_false p a := + generalize p; + rapply (decidable_false a); + try intro. + Class DecidablePaths (A : Type) := dec_paths : forall (x y : A), Decidable (x = y). Global Existing Instance dec_paths. @@ -181,12 +215,27 @@ Proof. intros x y; apply collapsible_hprop; exact _. Defined. +(** Hedberg's Theorem *) Corollary hset_decpaths (A : Type) `{DecidablePaths A} : IsHSet A. Proof. exact _. Defined. +(** We can use Hedberg's Theorem to simplify a goal of the form [forall (d : Decidable (x = x :> A)), P d] when [A] has decidable paths. *) +Definition decidable_paths_refl (A : Type) `{DecidablePaths A} + (x : A) + (P : forall (d : Decidable (x = x)), Type) + (Px : P (inl idpath)) + : forall d, P d. +Proof. + rapply (decidable_true idpath). + intro p. + (** We cannot eliminate [p : x = x] with path induction, but we can use Hedberg's theorem to replace this with [idpath]. *) + assert (r : (idpath = p)) by apply path_ishprop. + by destruct r. +Defined. + (** ** Truncation *) (** Having decidable equality (which implies being an hset, by Hedberg's theorem above) is itself an hprop. *) From 528e8322f5e7f4fd15998d7bc4bf434f3074e4b0 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Apr 2024 03:10:36 +0100 Subject: [PATCH 08/31] rename cat_coprod_prod_incl -> cat_coprod_prod Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Biproducts.v | 6 +++--- theories/Homotopy/Wedge.v | 2 +- theories/WildCat/Coproducts.v | 4 ++-- 3 files changed, 6 insertions(+), 6 deletions(-) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index df4108fee6d..b2e1a595724 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -13,7 +13,7 @@ Class Biproduct (I : Type) `{DecidablePaths I} {A : Type} := Build_Biproduct' { biproduct_product :: Product I x; biproduct_coproduct :: Coproduct I x; - catie_cat_coprod_prod_incl :: CatIsEquiv (cat_coprod_prod_incl x); + catie_cat_coprod_prod :: CatIsEquiv (cat_coprod_prod x); }. Arguments Biproduct I {_ A _ _ _ _ _ _} x. @@ -22,7 +22,7 @@ Section Biproducts. Context {I : Type} {A : Type} (x : I -> A) `{Biproduct I A x}. Definition cate_coprod_prod : cat_coprod I x $<~> cat_prod I x - := Build_CatEquiv (cat_coprod_prod_incl x). + := Build_CatEquiv (cat_coprod_prod x). Definition cat_biprod : A := cat_prod I x. @@ -742,7 +742,7 @@ Proof. - apply coproduct_op. exact _. - snrapply catie_homotopic. - + simpl; exact (cat_coprod_prod_incl (A:=A) x). + + simpl; exact (cat_coprod_prod (A:=A) x). + simpl; exact _. + (** Showing that these two maps are homotopic is a bit tricky. It boils down to showing an equality between [dec_paths (i = j)] and [dec_paths (j = i)] which is true with [Funext] but it is not certain if it holds without. *) apply cat_coprod_in_eta. diff --git a/theories/Homotopy/Wedge.v b/theories/Homotopy/Wedge.v index 21ffaa33d8e..69e1b10c1bb 100644 --- a/theories/Homotopy/Wedge.v +++ b/theories/Homotopy/Wedge.v @@ -254,7 +254,7 @@ Defined. (** Wedge inclusions into the product can be defined if the indexing type has decidable paths. This is because we need to choose which factor a given wedge summand should land in. *) Definition fwedge_incl `{Funext} (I : Type) `(DecidablePaths I) (X : I -> pType) : FamilyWedge I X $-> pproduct X - := cat_coprod_prod_incl X. + := cat_coprod_prod X. (** ** The pinch map on the suspension *) diff --git a/theories/WildCat/Coproducts.v b/theories/WildCat/Coproducts.v index 57e708290e9..c3bd206ef62 100644 --- a/theories/WildCat/Coproducts.v +++ b/theories/WildCat/Coproducts.v @@ -286,10 +286,10 @@ Defined. Global Instance hasbinarycoproducts_type : HasBinaryCoproducts Type := {}. -(** ** Coproduct-product inclusions *) +(** ** Canonical coproduct-product map *) (** There is a canonical map from a coproduct to a product when the indexing set has decidable equality and the category is pointed. *) -Definition cat_coprod_prod_incl {I : Type} `{DecidablePaths I} {A : Type} +Definition cat_coprod_prod {I : Type} `{DecidablePaths I} {A : Type} `{Is1Cat A, !IsPointedCat A} (x : I -> A) `{!Coproduct I x, !Product I x} : cat_coprod I x $-> cat_prod I x. From 265023f222f3d402b8aa15485efd52ece9b82c62 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Apr 2024 03:28:35 +0100 Subject: [PATCH 09/31] remove funext from biproduct_op Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Abelian.v | 3 +- theories/Algebra/Homological/Additive.v | 4 +-- theories/Algebra/Homological/Biproducts.v | 36 ++++++++++------------- 3 files changed, 18 insertions(+), 25 deletions(-) diff --git a/theories/Algebra/Homological/Abelian.v b/theories/Algebra/Homological/Abelian.v index 71a5ec69161..cdff06048da 100644 --- a/theories/Algebra/Homological/Abelian.v +++ b/theories/Algebra/Homological/Abelian.v @@ -106,8 +106,7 @@ Definition coker_zero_epic {A : Type} `{IsPointedCat A} : Epic f -> cat_coker f $== zero_morphism := ker_zero_monic (A:=A^op) (a:=b) f. -(** Funext required all the way in Biproducts to show opposite of additive is additivef. *) -Definition epic_coker_zero `{Funext} {A : Type} `{IsAdditive A} +Definition epic_coker_zero {A : Type} `{IsAdditive A} {a b : A} (f : a $-> b) {k : Cokernel f} : cat_coker f $== zero_morphism -> Epic f := monic_ker_zero (A:=A^op) (a:=b) (b:=a) f. diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index 9efd7e12cf9..c1a867c7665 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -458,7 +458,7 @@ Defined. (** ** Properties of additive categories *) (** The opposite of a semiadditive category is semiadditive. *) -Global Instance issemiadditive_op `{Funext} {A : Type} `{IsSemiAdditive A} +Global Instance issemiadditive_op {A : Type} `{IsSemiAdditive A} : IsSemiAdditive A^op. Proof. snrapply Build_IsSemiAdditive. @@ -475,7 +475,7 @@ Axiom sorry : Empty. Ltac sorry := apply Empty_rect; apply sorry. (** The opposite of an additive category is additive. *) -Global Instance isadditive_op `{F:Funext} {A : Type} `{H : IsAdditive A} +Global Instance isadditive_op {A : Type} `{H : IsAdditive A} : IsAdditive A^op. Proof. snrapply Build_IsAdditive. diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index b2e1a595724..caef9c20fed 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -732,7 +732,7 @@ End Associativity. (** *** Biproducts in the opposite category *) (** Biproducts exist in the opposite category if they exist in the original category. *) -Global Instance biproduct_op `{Funext} {I A : Type} {x : I -> A} `{Biproduct I A x} +Global Instance biproduct_op {I A : Type} {x : I -> A} `{Biproduct I A x} : Biproduct (A:=A^op) I x. Proof. snrapply Build_Biproduct'. @@ -744,7 +744,7 @@ Proof. - snrapply catie_homotopic. + simpl; exact (cat_coprod_prod (A:=A) x). + simpl; exact _. - + (** Showing that these two maps are homotopic is a bit tricky. It boils down to showing an equality between [dec_paths (i = j)] and [dec_paths (j = i)] which is true with [Funext] but it is not certain if it holds without. *) + + (** Showing that these two maps are homotopic is a bit tricky. *) apply cat_coprod_in_eta. intros i. apply cat_prod_pr_eta. @@ -761,34 +761,28 @@ Proof. refine (_ $@ _). 1: rapply cat_coprod_beta. simpl. - generalize (dec_paths i j). generalize (dec_paths j i). + generalize (dec_paths i j). intros p q. - assert (r : @decidable_equiv (i = j) (j = i) inverse _ q = p). - { destruct p as [p|np]. - - destruct q as [q|nq]. - + apply (ap inl). - apply path_ishprop. - + destruct (nq p^). - - destruct q as [|nq]. - + destruct (np p^). - + apply (ap inr). - (** Seems to be essential. *) - apply path_forall. - intros p. - destruct (np p). } - destruct r. - destruct q as [[]|q]; reflexivity. + destruct p as [p|np]. + * decidable_true q p^. + assert (r : p^ = x0) by apply path_ishprop. + destruct r. + destruct p. + reflexivity. + * destruct q. + 1: by destruct (np p^). + reflexivity. Defined. -Global Instance hasbiproducts_op `{Funext} {I A : Type} `{DecidablePaths I, HasEquivs A, !IsPointedCat A, !HasBiproducts I A} +Global Instance hasbiproducts_op {I A : Type} `{DecidablePaths I, HasEquivs A, !IsPointedCat A, !HasBiproducts I A} : HasBiproducts I (A^op). Proof. intros x. by rapply biproduct_op. Defined. -Global Instance binarybiproduct_op `{Funext} {A : Type} +Global Instance binarybiproduct_op {A : Type} `{HasEquivs A, !IsPointedCat A} {x y : A} {bb : BinaryBiproduct x y} : BinaryBiproduct (A:=A^op) x y. Proof. @@ -796,7 +790,7 @@ Proof. exact bb. Defined. -Global Instance hasbinarybiproducts_op `{Funext} {A : Type} +Global Instance hasbinarybiproducts_op {A : Type} `{HasEquivs A, !IsPointedCat A} {hbb : HasBinaryBiproducts A} : HasBinaryBiproducts (A^op). Proof. From b10c9e0a9cd0f3756a3e93ea808800159e958aef Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Apr 2024 03:53:32 +0100 Subject: [PATCH 10/31] fix 8.18 issues Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Biproducts.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index caef9c20fed..89dabb79cab 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -752,14 +752,14 @@ Proof. simpl. refine (cat_assoc _ _ _ $@ _). refine ((_ $@L _) $@ _). - 1: apply cat_coprod_beta. + 1: nrapply cat_coprod_beta. refine (_ $@ _). - 1: rapply cat_prod_beta. + 1: nrapply cat_prod_beta. symmetry. refine ((_ $@R _) $@ _). - 1: apply cat_prod_beta. + 1: nrapply cat_prod_beta. refine (_ $@ _). - 1: rapply cat_coprod_beta. + 1: nrapply cat_coprod_beta. simpl. generalize (dec_paths j i). generalize (dec_paths i j). From 1ea65010e41c751a9ae7257b9f61dbc66f5a4c22 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Apr 2024 04:00:53 +0100 Subject: [PATCH 11/31] rename canonical abelian map Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Abelian.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Algebra/Homological/Abelian.v b/theories/Algebra/Homological/Abelian.v index cdff06048da..c3750c0ffb7 100644 --- a/theories/Algebra/Homological/Abelian.v +++ b/theories/Algebra/Homological/Abelian.v @@ -118,7 +118,7 @@ Class IsPreAbelian {A : Type} `{IsAdditive A} := { preabelian_has_cokernels :: forall {a b : A} (f : a $-> b), Cokernel f; }. -Definition ispreabelian_canonical_map {A : Type} `{IsPreAbelian A} +Definition cat_coker_ker_ker_coker {A : Type} `{IsPreAbelian A} {a b : A} (f : a $-> b) : cat_coker_obj (cat_ker f) $-> cat_ker_obj (cat_coker f). Proof. @@ -138,6 +138,6 @@ Defined. Class IsAbelian {A : Type} `{IsPreAbelian A} := { catie_preabelian_canonical_map : forall a b (f : a $-> b), - CatIsEquiv (ispreabelian_canonical_map f); + CatIsEquiv (cat_coker_ker_ker_coker f); }. From 2b8177d3bd99a36e713507278f8762f47f59cded Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Apr 2024 02:59:08 +0100 Subject: [PATCH 12/31] optimize Additive.v Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Additive.v | 152 ++++++++++++------------ 1 file changed, 76 insertions(+), 76 deletions(-) diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index c1a867c7665..2311806e934 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -50,27 +50,27 @@ Section CMonHom. refine ((cat_assoc _ _ _)^$ $@ _). snrefine ((_ $@L _) $@ _). 1: exact cat_binbiprod_inr. - { apply cat_binbiprod_pr_eta. + { nrapply cat_binbiprod_pr_eta. - refine (_ $@ cat_binbiprod_pr1_inr^$). refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. + 1: nrapply cat_binbiprod_corec_beta_pr1. refine (cat_assoc _ _ _ $@ _). - apply cat_zero_l. + nrapply cat_zero_l. - refine (_ $@ cat_binbiprod_pr2_inr^$). refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. + 1: nrapply cat_binbiprod_corec_beta_pr2. refine (cat_assoc _ _ _ $@ _). refine (cat_idl _ $@ _). - apply cat_binbiprod_corec_beta_pr2. } + nrapply cat_binbiprod_corec_beta_pr2. } nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). { refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_rec. - apply cat_binbiprod_rec_beta_inr. } + 1: nrapply cat_binbiprod_corec_rec. + nrapply cat_binbiprod_rec_beta_inr. } refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idl. + 1: nrapply cat_binbiprod_rec_beta_inr. + nrapply cat_idl. Defined. Local Existing Instance symmetricbraiding_cat_binbiprod. @@ -82,10 +82,10 @@ Section CMonHom. apply path_hom. refine (cat_assoc _ _ _ $@ _). refine ((_^$ $@R _) $@ _). - 1: apply cat_binbiprod_swap_codiag. + 1: nrapply cat_binbiprod_swap_codiag. refine (cat_assoc _ _ _ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). refine (_ $@ (_ $@L _)). - 2: apply cat_binbiprod_swap_diag. + 2: nrapply cat_binbiprod_swap_diag. refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ cat_assoc _ _ _). rapply Monoidal.braid_nat. Defined. @@ -94,7 +94,7 @@ Section CMonHom. Proof. intros f. refine (commutativity _ _ @ _). - apply left_identity_hom. + nrapply left_identity_hom. Defined. Local Instance associative_ab_hom : Associative sgop_hom. @@ -111,7 +111,7 @@ Section CMonHom. Local Notation "'Δ' x" := (cat_binbiprod_diag x) (at level 10). Local Notation "'∇' x" := (cat_binbiprod_codiag x) (at level 10). change (?w $o ?x $== ?y $o ?z) with (Square z w x y). - apply move_right_top. + nrapply move_right_top. (** The following associativity rewrites are specially chosen so that the diagram can be decomposed easily. *) rapply vconcatL. 1: do 2 refine (cat_assoc _ _ _ $@ _). @@ -124,20 +124,20 @@ Section CMonHom. 1: nrapply cate_binbiprod_assoc. { nrefine ((_ $@R _) $@ _). 1: rapply Monoidal.associator_twist'_unfold. - apply cat_binbiprod_pr_eta. + nrapply cat_binbiprod_pr_eta. - srefine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). 1: exact (fmap01 (fun x y => cat_binbiprod x y) b cat_binbiprod_pr1). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. + 1: nrapply cat_binbiprod_corec_beta_pr1. refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. + 1: nrapply cat_binbiprod_corec_beta_pr2. refine (_^$ $@ _). 1: rapply fmap01_comp. refine (fmap02 _ _ _). - apply cat_binbiprod_corec_beta_pr2. } + nrapply cat_binbiprod_corec_beta_pr2. } refine ((cat_assoc _ _ _)^$ $@ _). refine ((((fmap02 _ _ _ $@ fmap01_id _ _ _)^$ $@ fmap01_comp _ _ _ _)^$ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. + 1: nrapply cat_binbiprod_corec_beta_pr1. nrefine (cat_idl _ $@ _). refine (_ $@ (_ $@L (cat_assoc _ _ _)^$)). nrefine (_ $@ cat_assoc _ _ _). @@ -151,156 +151,156 @@ Section CMonHom. refine (_ $@ (_ $@L _^$)). 2: rapply cat_binbiprod_corec_beta_pr1. symmetry. - apply cat_idr. + nrapply cat_idr. - nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). refine ((cat_assoc _ _ _)^$ $@ _). nrefine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. + 1: nrapply cat_binbiprod_corec_beta_pr2. nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). refine ((cat_assoc _ _ _)^$ $@ _). nrefine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. + 1: nrapply cat_binbiprod_corec_beta_pr1. nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - apply cat_binbiprod_corec_beta_pr2. } + nrapply cat_binbiprod_corec_beta_pr2. } nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). refine ((cat_assoc _ _ _)^$ $@ _). nrefine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. + 1: nrapply cat_binbiprod_corec_beta_pr1. nrefine ((_ $@L _) $@ _). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - apply cat_binbiprod_corec_beta_pr2. } + nrapply cat_binbiprod_corec_beta_pr2. } nrefine ((_ $@L (cat_assoc _ _ _)) $@ _). refine ((cat_assoc _ _ _)^$ $@ _). nrefine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. + 1: nrapply cat_binbiprod_corec_beta_pr2. nrefine (cat_idl _ $@ _). refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. + 1: nrapply cat_binbiprod_corec_beta_pr2. nrefine (cat_assoc _ _ _ $@ _). nrefine (cat_idl _ $@ _). nrefine (cat_binbiprod_corec_beta_pr2 _ _ $@ _). refine (_ $@ (_ $@L (cat_assoc _ _ _)^$)). nrefine (_ $@ cat_assoc _ _ _). refine (_ $@ (_^$ $@R _)). - 2: apply cat_binbiprod_corec_beta_pr2. + 2: nrapply cat_binbiprod_corec_beta_pr2. refine (_ $@ (cat_assoc _ _ _)^$). refine (_ $@ (cat_idl _)^$). refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: apply cat_binbiprod_corec_beta_pr2. + 2: nrapply cat_binbiprod_corec_beta_pr2. refine (_ $@ (cat_assoc _ _ _)^$). refine (_ $@ (cat_idl _)^$). symmetry. - apply cat_binbiprod_corec_beta_pr2. } + nrapply cat_binbiprod_corec_beta_pr2. } (** We split the square again, this time by picking yet another associativity map as the common edge. This leaves us with a naturality square for associativity. *) srapply hconcat. 1: nrapply cate_binbiprod_assoc. 1: nrapply Monoidal.associator_nat_m. (** Finally we are left with another polygon which we resolve by brute force. *) refine ((cat_assoc _ _ _)^$ $@ _). - apply cat_binbiprod_in_eta. + nrapply cat_binbiprod_in_eta. - nrefine (cat_assoc _ _ _ $@ _). nrefine ((_ $@L _) $@ _). { nrefine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inl. } + nrapply cat_binbiprod_rec_beta_inl. } refine ((cat_assoc _ _ _)^$ $@ _). refine (cat_idr _ $@ _). refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inl. } + nrapply cat_binbiprod_rec_beta_inl. } refine ((cat_assoc _ _ _)^$ $@ _). refine (cat_idr _ $@ _). refine (_ $@ _). - 1: apply cat_binbiprod_rec_beta_inl. + 1: nrapply cat_binbiprod_rec_beta_inl. symmetry. refine (cat_assoc _ _ _ $@ _). refine ((_ $@L _) $@ _). - 1: apply cate_binbiprod_assoc_inl. + 1: nrapply cate_binbiprod_assoc_inl. refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). { refine (cat_assoc _ _ _ $@ (_ $@L _)). refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inl. } + nrapply cat_binbiprod_rec_beta_inl. } simpl. refine (cat_assoc _ _ _ $@ _). refine ((_ $@L _) $@ _). { refine (cat_assoc _ _ _ $@ (_ $@L _)). - apply cat_binbiprod_rec_beta_inl. } + nrapply cat_binbiprod_rec_beta_inl. } refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). { refine (cat_assoc _ _ _ $@ (_ $@L _)). refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). - 1: apply cat_binbiprod_rec_beta_inl. - apply cat_idr. } + 1: nrapply cat_binbiprod_rec_beta_inl. + nrapply cat_idr. } refine (cat_idr _ $@ _). - apply cat_binbiprod_rec_beta_inl. + nrapply cat_binbiprod_rec_beta_inl. - nrefine (cat_assoc _ _ _ $@ _). nrefine ((_ $@L _) $@ _). { nrefine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inr. } + nrapply cat_binbiprod_rec_beta_inr. } refine ((cat_assoc _ _ _)^$ $@ _). simpl. refine ((_ $@R _) $@ _). { refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). 1: refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. + 1: nrapply cat_binbiprod_rec_beta_inr. refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idl. } + 1: nrapply cat_binbiprod_rec_beta_inr. + nrapply cat_idl. } simpl. refine (_ $@ (cat_assoc _ _ _)^$). - apply cat_binbiprod_in_eta. + nrapply cat_binbiprod_in_eta. + refine (cat_assoc _ _ _ $@ _). refine ((_ $@L _) $@ _). { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inl. } + nrapply cat_binbiprod_rec_beta_inl. } refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inl. + 1: nrapply cat_binbiprod_rec_beta_inl. refine (cat_idr _ $@ _). refine (_ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). - 2: apply cate_binbiprod_assoc_inr_inl. + 2: nrapply cate_binbiprod_assoc_inr_inl. refine (_ $@ (cat_assoc _ _ _)^$). refine (_ $@ (_ $@L _^$)). 2: { refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). 1: refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inl. + 1: nrapply cat_binbiprod_rec_beta_inl. simpl. refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idr. } + 1: nrapply cat_binbiprod_rec_beta_inr. + nrapply cat_idr. } simpl. refine (_ $@ (cat_assoc _ _ _)^$). refine (_^$ $@ (_ $@L _^$)). 2: { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). - 1: apply cat_binbiprod_rec_beta_inl. - apply cat_idr. } - apply cat_binbiprod_rec_beta_inl. + 1: nrapply cat_binbiprod_rec_beta_inl. + nrapply cat_idr. } + nrapply cat_binbiprod_rec_beta_inl. + refine (cat_assoc _ _ _ $@ _). refine ((_ $@L _) $@ _). { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inr. } + nrapply cat_binbiprod_rec_beta_inr. } refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. + 1: nrapply cat_binbiprod_rec_beta_inr. refine (cat_idl _ $@ _). refine (_ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). - 2: apply cate_binbiprod_assoc_inr_inr. + 2: nrapply cate_binbiprod_assoc_inr_inr. simpl. refine (_ $@ (cat_assoc _ _ _)^$). refine (_^$ $@ (_ $@L _^$)). 2: { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idr. } + 1: nrapply cat_binbiprod_rec_beta_inr. + nrapply cat_idr. } refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. } + 1: nrapply cat_binbiprod_rec_beta_inr. } refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idl. + 1: nrapply cat_binbiprod_rec_beta_inr. + nrapply cat_idl. Defined. Local Instance issemigroup_hom : IsSemiGroup (a $-> b) := {}. @@ -348,7 +348,7 @@ Definition addcat_dist_l {A : Type} `{IsAdditive A} {a b c : A} Proof. refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _ $@ cat_assoc _ _ _). - 1: apply cat_binbiprod_codiag_fmap11. + 1: nrapply cat_binbiprod_codiag_fmap11. refine (cat_assoc _ _ _ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). rapply fmap11_comp. Defined. @@ -359,7 +359,7 @@ Global Instance left_heterodistribute_hom {A : Type} `{IsAdditive A} {a b c : A} Proof. intros f g h. apply path_hom. - apply addcat_dist_l. + nrapply addcat_dist_l. Defined. Definition addcat_dist_r {A : Type} `{IsAdditive A} {a b c : A} @@ -369,7 +369,7 @@ Proof. refine (cat_assoc _ _ _ $@ _). refine (cat_assoc _ _ _ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). refine ((_ $@L _) $@ _). - 1: apply cat_binbiprod_diag_fmap11. + 1: nrapply cat_binbiprod_diag_fmap11. refine ((cat_assoc _ _ _)^$ $@ (_^$ $@R _)). rapply fmap11_comp. Defined. @@ -379,39 +379,39 @@ Global Instance right_heterodistribute_hom {A : Type} `{IsAdditive A} {a b c : A Proof. intros f g h. apply path_hom. - apply addcat_dist_r. + nrapply addcat_dist_r. Defined. Definition addcat_comp_negate_id_l {A : Type} `{IsAdditive A} {a b : A} (f : a $-> b) : (- Id _) $o f $== -f. Proof. - apply GpdHom_path. + nrapply GpdHom_path. change (@paths (Hom a b) ?x ?y) with (@paths (AbHom a b : Type) x y). rapply grp_moveL_1V. transitivity ((-Id b $o f) + (Id b $o f)). - f_ap; symmetry; apply path_hom, cat_idl. - lhs_V rapply right_heterodistribute_hom. transitivity (zero_morphism (b := b) $o f). - + f_ap; apply (grp_inv_l (Id b : AbHom b b)). + + f_ap; nrapply (grp_inv_l (Id b : AbHom b b)). + apply path_hom. - apply cat_zero_l. + nrapply cat_zero_l. Defined. Definition addcat_comp_negate_id_r {A : Type} `{IsAdditive A} {a b : A} (f : a $-> b) : f $o (- Id _) $== -f. Proof. - apply GpdHom_path. + nrapply GpdHom_path. change (@paths (Hom a b) ?x ?y) with (@paths (AbHom a b : Type) x y). rapply grp_moveL_1V. transitivity ((f $o -Id a) + (f $o Id a)). - f_ap; symmetry; apply path_hom, cat_idr. - lhs_V rapply left_heterodistribute_hom. transitivity (f $o zero_morphism (a := a)). - + f_ap; apply (grp_inv_l (Id a : AbHom a a)). + + f_ap; nrapply (grp_inv_l (Id a : AbHom a a)). + apply path_hom. - apply cat_zero_r. + nrapply cat_zero_r. Defined. Definition addcat_comp_negate_l {A : Type} `{IsAdditive A} {a b c : A} @@ -420,7 +420,7 @@ Definition addcat_comp_negate_l {A : Type} `{IsAdditive A} {a b c : A} Proof. refine ((addcat_comp_negate_id_l _)^$ $@ _). refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - apply addcat_comp_negate_id_l. + nrapply addcat_comp_negate_id_l. Defined. Definition addcat_comp_negate_r {A : Type} `{IsAdditive A} {a b c : A} @@ -429,7 +429,7 @@ Definition addcat_comp_negate_r {A : Type} `{IsAdditive A} {a b c : A} Proof. refine ((addcat_comp_negate_id_r _)^$ $@ _). refine (cat_assoc _ _ _ $@ (_ $@L _)). - apply addcat_comp_negate_id_r. + nrapply addcat_comp_negate_id_r. Defined. (** *** Endomorphism ring *) @@ -446,13 +446,13 @@ Proof. - intros f g h. apply path_hom. symmetry. - apply cat_assoc. + nrapply cat_assoc. - intros f. apply path_hom. - apply cat_idl. + nrapply cat_idl. - intros g. apply path_hom. - apply cat_idr. + nrapply cat_idr. Defined. (** ** Properties of additive categories *) From 7bdd03691570e580a46c5c7f51b1e2f10eab007c Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Apr 2024 03:05:08 +0100 Subject: [PATCH 13/31] optimize biproducts Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Biproducts.v | 230 +++++++++++----------- 1 file changed, 115 insertions(+), 115 deletions(-) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 89dabb79cab..623231c805d 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -64,18 +64,18 @@ Section Biproducts. unfold cat_biprod_rec, cat_biprod_in. refine (_ $@ cat_coprod_beta I f i). nrefine (cat_assoc _ _ _ $@ (_ $@L _)). - apply compose_V_hh. + nrapply compose_V_hh. Defined. Definition cat_biprod_rec_eta {z : A} (f : cat_biprod $-> z) : cat_biprod_rec (fun i => f $o cat_biprod_in i) $== f. Proof. unfold cat_biprod_rec. - apply cate_moveR_eV. - apply cat_coprod_in_eta. + nrapply cate_moveR_eV. + nrapply cat_coprod_in_eta. intros i. refine (cat_coprod_beta I _ i $@ _^$). - apply cat_assoc. + nrapply cat_assoc; exact _. Defined. Definition cat_biprod_rec_eta' {z : A} {f f' : forall i, x i $-> z} @@ -90,8 +90,8 @@ Section Biproducts. : (forall i, f $o cat_biprod_in i $== f' $o cat_biprod_in i) -> f $== f'. Proof. intros p. - apply (cate_epic_equiv cate_coprod_prod). - apply cat_coprod_in_eta. + nrapply (cate_epic_equiv cate_coprod_prod). + nrapply cat_coprod_in_eta. intros i. exact (cat_assoc _ _ _ $@ p i $@ (cat_assoc _ _ _)^$). Defined. @@ -111,7 +111,7 @@ Proof. unfold cat_biprod_pr, cat_biprod_in. refine ((_ $@L _) $@ _). { refine ((cate_buildequiv_fun _ $@R _) $@ _). - apply cat_coprod_beta. } + nrapply cat_coprod_beta. } refine (cat_prod_beta _ _ _ $@ _). simpl. generalize (dec_paths i i). @@ -126,7 +126,7 @@ Proof. unfold cat_biprod_pr, cat_biprod_in. refine ((_ $@L _) $@ _). { refine ((cate_buildequiv_fun _ $@R _) $@ _). - apply cat_coprod_beta. } + nrapply cat_coprod_beta. } refine (cat_prod_beta _ _ _ $@ _). decidable_false (dec_paths i j) p. simpl. @@ -149,19 +149,19 @@ Definition cat_biprod_corec_rec I `{DecidablePaths I} {A : Type} : cat_biprod_corec y (fun i => f i $o cat_biprod_pr i) $== cat_biprod_rec x (fun i => cat_biprod_in i $o f i). Proof. - apply cat_biprod_pr_eta. + nrapply cat_biprod_pr_eta. intros i. refine (cat_biprod_corec_beta _ _ i $@ _). - apply cat_biprod_in_eta. + nrapply cat_biprod_in_eta. intros j. refine (_ $@ (_ $@L (cat_biprod_rec_beta _ _ _)^$) $@ (cat_assoc _ _ _)^$). refine (cat_assoc _ _ _ $@ _ $@ cat_assoc _ _ _). destruct (dec_paths j i) as [p | np]. - destruct p. refine ((_ $@L _) $@ cat_idr _ $@ (cat_idl _)^$ $@ (_^$ $@R _)). - 1,2: apply cat_biprod_pr_in. + 1,2: nrapply cat_biprod_pr_in. - refine ((_ $@L _) $@ cat_zero_r _ $@ (cat_zero_l _)^$ $@ (_^$ $@R _)). - 1,2: apply cat_biprod_pr_in_ne, np. + 1,2: nrapply cat_biprod_pr_in_ne; assumption. Defined. (** *** Existence of biproducts *) @@ -197,7 +197,7 @@ Global Instance hasbinarycoproducts_hasbinarybiproducts {A : Type} Proof. intros x y. nrapply biproduct_coproduct. - apply H. + nrapply H. Defined. Section BinaryBiproducts. @@ -214,7 +214,7 @@ Section BinaryBiproducts. Definition cat_binbiprod_corec {z : A} (f : z $-> x) (g : z $-> y) : z $-> cat_binbiprod. Proof. - apply cat_biprod_corec. + nrapply cat_biprod_corec. intros [|]. - exact f. - exact g. @@ -232,7 +232,7 @@ Section BinaryBiproducts. : cat_binbiprod_corec (cat_binbiprod_pr1 $o f) (cat_binbiprod_pr2 $o f) $== f. Proof. - apply cat_biprod_pr_eta. + nrapply cat_biprod_pr_eta. intros [|]. - exact (cat_binbiprod_corec_beta_pr1 _ _). - exact (cat_binbiprod_corec_beta_pr2 _ _). @@ -243,7 +243,7 @@ Section BinaryBiproducts. -> cat_binbiprod_corec f g $== cat_binbiprod_corec f' g'. Proof. intros p q. - apply cat_biprod_corec_eta'. + nrapply cat_biprod_corec_eta'. intros [|]. - exact p. - exact q. @@ -255,7 +255,7 @@ Section BinaryBiproducts. -> f $== f'. Proof. intros p q. - apply cat_biprod_pr_eta. + nrapply cat_biprod_pr_eta. intros [|]. - exact p. - exact q. @@ -267,7 +267,7 @@ Section BinaryBiproducts. Definition cat_binbiprod_rec {z : A} (f : x $-> z) (g : y $-> z) : cat_binbiprod $-> z. Proof. - apply cat_biprod_rec. + nrapply cat_biprod_rec. intros [|]. - exact f. - exact g. @@ -284,7 +284,7 @@ Section BinaryBiproducts. Definition cat_binbiprod_rec_eta {z : A} (f : cat_binbiprod $-> z) : cat_binbiprod_rec (f $o cat_binbiprod_inl) (f $o cat_binbiprod_inr) $== f. Proof. - apply cat_biprod_in_eta. + nrapply cat_biprod_in_eta. intros [|]. - exact (cat_binbiprod_rec_beta_inl _ _). - exact (cat_binbiprod_rec_beta_inr _ _). @@ -295,7 +295,7 @@ Section BinaryBiproducts. -> cat_binbiprod_rec f g $== cat_binbiprod_rec f' g'. Proof. intros p q. - apply cat_biprod_rec_eta'. + nrapply cat_biprod_rec_eta'. intros [|]. - exact p. - exact q. @@ -307,7 +307,7 @@ Section BinaryBiproducts. -> f $== f'. Proof. intros p q. - apply cat_biprod_in_eta. + nrapply cat_biprod_in_eta. intros [|]. - exact p. - exact q. @@ -325,14 +325,14 @@ Section BinaryBiproducts. : cat_binbiprod_pr2 $o cat_binbiprod_inl $== zero_morphism. Proof. snrapply (cat_biprod_pr_in_ne Bool _ (i := true) (j := false)). - apply (not_fixed_negb false). + nrapply (not_fixed_negb false). Defined. Definition cat_binbiprod_pr1_inr : cat_binbiprod_pr1 $o cat_binbiprod_inr $== zero_morphism. Proof. snrapply (cat_biprod_pr_in_ne Bool _ (i := false) (j := true)). - apply (not_fixed_negb true). + nrapply (not_fixed_negb true). Defined. End BinaryBiproducts. @@ -409,23 +409,23 @@ Definition cat_binbiprod_diag_fmap11 {A : Type} : cat_binbiprod_diag y $o f $== fmap11 (fun x y => cat_binbiprod x y) f f $o cat_binbiprod_diag x. Proof. - apply cat_binbiprod_pr_eta. + nrapply cat_binbiprod_pr_eta. - refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ (_ $@ _)). - 1: apply cat_binbiprod_corec_beta_pr1. - 1: apply cat_idl. + 1: rapply cat_binbiprod_corec_beta_pr1. + 1: nrapply cat_idl. refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). 2: rapply cat_pr1_fmap11_binprod. refine ((cat_idr _)^$ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). - apply cat_binbiprod_corec_beta_pr1. + rapply cat_binbiprod_corec_beta_pr1. - refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ (_ $@ _)). - 1: apply cat_binbiprod_corec_beta_pr2. - 1: apply cat_idl. + 1: rapply cat_binbiprod_corec_beta_pr2. + 1: nrapply cat_idl. refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). 2: rapply cat_pr2_fmap11_binprod. refine ((cat_idr _)^$ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). - apply cat_binbiprod_corec_beta_pr2. + rapply cat_binbiprod_corec_beta_pr2. Defined. Definition cat_binbiprod_codiag_fmap11 {A : Type} @@ -433,35 +433,35 @@ Definition cat_binbiprod_codiag_fmap11 {A : Type} : f $o cat_binbiprod_codiag x $== cat_binbiprod_codiag y $o fmap11 (fun x y => cat_binbiprod x y) f f. Proof. - apply cat_binbiprod_in_eta. + nrapply cat_binbiprod_in_eta. - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ cat_idr _ $@ _). - 1: apply cat_binbiprod_rec_beta_inl. + 1: nrapply cat_binbiprod_rec_beta_inl. refine (_ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). 2: { refine ((_ $@L _) $@ (cat_assoc _ _ _)^$). refine (_^$ $@ (cat_binbiprod_corec_rec _ _ $@R _)^$). - apply cat_binbiprod_rec_beta_inl. } + nrapply cat_binbiprod_rec_beta_inl. } refine (_ $@ (_ $@L _)). 2: { refine ((_ $@R _) $@ cat_assoc _ _ _). refine ((_ $@ _)^$ $@ (cat_binbiprod_corec_rec _ _ $@R _)^$). - 1: apply cat_binbiprod_rec_beta_inl. - apply cat_idr. } + 1: nrapply cat_binbiprod_rec_beta_inl. + nrapply cat_idr. } refine (_^$ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: apply cat_binbiprod_rec_beta_inl. - apply cat_idl. + 2: nrapply cat_binbiprod_rec_beta_inl. + nrapply cat_idl; exact _. - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ cat_idr _ $@ _). - 1: apply cat_binbiprod_rec_beta_inr. + 1: nrapply cat_binbiprod_rec_beta_inr. refine (_ $@ (_ $@L _) $@ (cat_assoc _ _ _)^$). 2: { refine ((_ $@L _) $@ (cat_assoc _ _ _)^$). refine ((_ $@ _)^$ $@ (cat_binbiprod_corec_rec _ _ $@R _)^$). - 1: apply cat_binbiprod_rec_beta_inr. - apply cat_idr. } + 1: nrapply cat_binbiprod_rec_beta_inr. + nrapply cat_idr. } refine (_ $@ (_ $@L _)). 2: { refine ((_ $@R _) $@ _)^$. - 1: apply cat_binbiprod_corec_rec. - apply cat_binbiprod_rec_beta_inr. } + 1: rapply cat_binbiprod_corec_rec. + nrapply cat_binbiprod_rec_beta_inr. } refine (_^$ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: apply cat_binbiprod_rec_beta_inr. - apply cat_idl. + 2: nrapply cat_binbiprod_rec_beta_inr. + nrapply cat_idl; exact _. Defined. (** *** Symmetry *) @@ -489,68 +489,68 @@ Section Symmetry. : cat_binbiprod_swap x y $o cat_binbiprod_inl $== cat_binbiprod_inr. Proof. - apply cat_binbiprod_pr_eta. + nrapply cat_binbiprod_pr_eta. - refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. + 1: nrapply cat_binbiprod_corec_beta_pr1. refine (_ $@ _^$). - 1: apply cat_binbiprod_pr2_inl. - apply cat_binbiprod_pr1_inr. + 1: nrapply cat_binbiprod_pr2_inl. + nrapply cat_binbiprod_pr1_inr. - refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. + 1: nrapply cat_binbiprod_corec_beta_pr2. refine (_ $@ _^$). - 1: apply cat_binbiprod_pr1_inl. - apply cat_binbiprod_pr2_inr. + 1: nrapply cat_binbiprod_pr1_inl. + nrapply cat_binbiprod_pr2_inr. Defined. Lemma cat_binbiprod_swap_inr (x y : A) : cat_binbiprod_swap x y $o cat_binbiprod_inr $== cat_binbiprod_inl. Proof. - apply cat_binbiprod_pr_eta. + nrapply cat_binbiprod_pr_eta. - refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. + 1: nrapply cat_binbiprod_corec_beta_pr1. refine (_ $@ _^$). - 1: apply cat_binbiprod_pr2_inr. - apply cat_binbiprod_pr1_inl. + 1: nrapply cat_binbiprod_pr2_inr. + nrapply cat_binbiprod_pr1_inl. - refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. + 1: nrapply cat_binbiprod_corec_beta_pr2. refine (_ $@ _^$). - 1: apply cat_binbiprod_pr1_inr. - apply cat_binbiprod_pr2_inl. + 1: nrapply cat_binbiprod_pr1_inr. + nrapply cat_binbiprod_pr2_inl. Defined. (** The swap map preserves the diagonal. *) Lemma cat_binbiprod_swap_diag (x : A) : cat_binbiprod_swap x x $o cat_binbiprod_diag x $== cat_binbiprod_diag x. Proof. - apply cat_binbiprod_pr_eta. + nrapply cat_binbiprod_pr_eta. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. + 1: nrapply cat_binbiprod_corec_beta_pr1. refine (cat_binbiprod_corec_beta_pr2 _ _ $@ _^$). - apply cat_binbiprod_corec_beta_pr1. + nrapply cat_binbiprod_corec_beta_pr1. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. + 1: nrapply cat_binbiprod_corec_beta_pr2. refine (cat_binbiprod_corec_beta_pr1 _ _ $@ _^$). - apply cat_binbiprod_corec_beta_pr2. + nrapply cat_binbiprod_corec_beta_pr2. Defined. (** The swap map preserves the codiagonal. *) Lemma cat_binbiprod_swap_codiag (x : A) : cat_binbiprod_codiag x $o cat_binbiprod_swap x x $== cat_binbiprod_codiag x. Proof. - apply cat_binbiprod_in_eta. + nrapply cat_binbiprod_in_eta. - refine (_ $@ (cat_binbiprod_rec_beta_inl _ _)^$). refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - 1: apply cat_binbiprod_swap_inl. - apply cat_binbiprod_rec_beta_inr. + 1: nrapply cat_binbiprod_swap_inl. + nrapply cat_binbiprod_rec_beta_inr. - refine (_ $@ (cat_binbiprod_rec_beta_inr _ _)^$). refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - 1: apply cat_binbiprod_swap_inr. - apply cat_binbiprod_rec_beta_inl. + 1: nrapply cat_binbiprod_swap_inr. + nrapply cat_binbiprod_rec_beta_inl. Defined. End Symmetry. @@ -581,86 +581,86 @@ Section Associativity. : cat_binbiprod_twist x y z $o cat_binbiprod_inl $== cat_binbiprod_inr $o cat_binbiprod_inl. Proof. - apply cat_binbiprod_pr_eta. + nrapply cat_binbiprod_pr_eta. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr1. + 1: nrapply cat_binbiprod_corec_beta_pr1. refine (cat_assoc _ _ _ $@ _ $@ cat_assoc _ _ _). refine ((_ $@L _) $@ _ $@ (_^$ $@R _)). - 1: apply cat_binbiprod_pr2_inl. - 2: apply cat_binbiprod_pr1_inr. + 1: nrapply cat_binbiprod_pr2_inl. + 2: nrapply cat_binbiprod_pr1_inr. refine (_ $@ _^$). - 1: apply cat_zero_r. - apply cat_zero_l. + 1: nrapply cat_zero_r. + nrapply cat_zero_l. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_beta_pr2. + 1: nrapply cat_binbiprod_corec_beta_pr2. refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). refine (cat_binbiprod_rec_beta_inl _ _ $@ cat_idr _ $@ _). refine ((cat_idl _)^$ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - apply cat_binbiprod_pr2_inr. + nrapply cat_binbiprod_pr2_inr. Defined. Lemma cat_binbiprod_twist_inr_inl (x y z : A) : cat_binbiprod_twist x y z $o cat_binbiprod_inr $o cat_binbiprod_inl $== cat_binbiprod_inl. Proof. - apply cat_binbiprod_pr_eta. + nrapply cat_binbiprod_pr_eta. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - 1: apply cat_binbiprod_corec_beta_pr1. } + 1: nrapply cat_binbiprod_corec_beta_pr1. } refine (cat_assoc _ _ _ $@ _). refine (cat_assoc _ _ _ $@ _). refine ((_ $@L _) $@ _). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_pr2_inr. - apply cat_idl. } + 1: nrapply cat_binbiprod_pr2_inr. + nrapply cat_idl. } refine (_ $@ _^$). - 1,2: apply cat_binbiprod_pr1_inl. + 1,2: nrapply cat_binbiprod_pr1_inl. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - 1: apply cat_binbiprod_corec_beta_pr2. } + 1: nrapply cat_binbiprod_corec_beta_pr2. } refine (cat_assoc _ _ _ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_rec. + 1: nrapply cat_binbiprod_corec_rec. refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. + 1: nrapply cat_binbiprod_rec_beta_inr. refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - 1: apply cat_binbiprod_pr2_inl. + 1: nrapply cat_binbiprod_pr2_inl. refine (_ $@ _^$). - 1: apply cat_zero_r. - apply cat_binbiprod_pr2_inl. + 1: nrapply cat_zero_r. + nrapply cat_binbiprod_pr2_inl. Defined. Lemma cat_binbiprod_twist_inr_inr (x y z : A) : cat_binbiprod_twist x y z $o cat_binbiprod_inr $o cat_binbiprod_inr $== cat_binbiprod_inr $o cat_binbiprod_inr. Proof. - apply cat_binbiprod_pr_eta. + nrapply cat_binbiprod_pr_eta. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - 1: apply cat_binbiprod_corec_beta_pr1. } + 1: nrapply cat_binbiprod_corec_beta_pr1. } refine (cat_assoc _ _ _ $@ _). refine (cat_assoc _ _ _ $@ _). refine ((_ $@L _) $@ _). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_pr2_inr. - apply cat_idl. } + 1: nrapply cat_binbiprod_pr2_inr. + nrapply cat_idl. } refine (_ $@ _^$). - 1: apply cat_binbiprod_pr1_inr. + 1: nrapply cat_binbiprod_pr1_inr. refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_pr1_inr. - apply cat_zero_l. + 1: nrapply cat_binbiprod_pr1_inr. + nrapply cat_zero_l. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - 1: apply cat_binbiprod_corec_beta_pr2. } + 1: nrapply cat_binbiprod_corec_beta_pr2. } refine (cat_assoc _ _ _ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_corec_rec. + 1: nrapply cat_binbiprod_corec_rec. refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_rec_beta_inr. + 1: nrapply cat_binbiprod_rec_beta_inr. refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - 1: apply cat_binbiprod_pr2_inr. + 1: nrapply cat_binbiprod_pr2_inr. refine (cat_idr _ $@ _^$). refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: apply cat_binbiprod_pr2_inr. - apply cat_idl. + 1: nrapply cat_binbiprod_pr2_inr. + nrapply cat_idl. Defined. Lemma cate_binbiprod_assoc_inl (x y z : A) @@ -672,14 +672,14 @@ Section Associativity. refine (cat_assoc _ _ _ $@ _). refine ((_ $@L (cat_assoc _ _ _ $@ (_ $@L _))) $@ _). { refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ (_ $@ _)). - 1: apply cat_binbiprod_rec_beta_inl. - apply cat_idr. } + 1: nrapply cat_binbiprod_rec_beta_inl. + nrapply cat_idr. } simpl. refine ((_ $@L _) $@ _). - 1: apply cat_binbiprod_twist_inl. + 1: nrapply cat_binbiprod_twist_inl. refine ((cat_assoc _ _ _)^$ $@ _). refine (_ $@R _). - apply cat_binbiprod_swap_inr. + nrapply cat_binbiprod_swap_inr. Defined. Lemma cate_binbiprod_assoc_inr_inl (x y z : A) @@ -693,16 +693,16 @@ Section Associativity. refine ((_ $@L (cat_assoc _ _ _ $@ (_ $@L _))) $@ _). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inr. } + nrapply cat_binbiprod_rec_beta_inr. } refine ((_ $@L _) $@ _). { refine ((cat_assoc _ _ _)^$ $@ _). refine (((cat_assoc _ _ _)^$ $@R _) $@ _). refine (cat_assoc _ _ _ $@ _). refine ((_ $@L _) $@ _). - 1: apply cat_binbiprod_swap_inl. - apply cat_binbiprod_twist_inr_inr. } + 1: nrapply cat_binbiprod_swap_inl. + nrapply cat_binbiprod_twist_inr_inr. } refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - apply cat_binbiprod_swap_inr. + nrapply cat_binbiprod_swap_inr. Defined. Lemma cate_binbiprod_assoc_inr_inr (x y z : A) @@ -716,15 +716,15 @@ Section Associativity. refine ((_ $@L (cat_assoc _ _ _ $@ (_ $@L _))) $@ _). { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). refine ((cat_binbiprod_corec_rec _ _ $@R _) $@ _). - apply cat_binbiprod_rec_beta_inr. } + nrapply cat_binbiprod_rec_beta_inr. } refine ((_ $@L _) $@ _). { refine ((cat_assoc _ _ _)^$ $@ _). refine (((cat_assoc _ _ _)^$ $@R _) $@ _). refine (cat_assoc _ _ _ $@ _). refine ((_ $@L _) $@ _). - 1: apply cat_binbiprod_swap_inr. - apply cat_binbiprod_twist_inr_inl. } - apply cat_binbiprod_swap_inl. + 1: nrapply cat_binbiprod_swap_inr. + nrapply cat_binbiprod_twist_inr_inl. } + nrapply cat_binbiprod_swap_inl. Defined. End Associativity. @@ -739,15 +739,15 @@ Proof. (** Products in the opposite category are coproducts in the original category. *) - exact _. (** Coproducts in the opposite category are products in the original category. *) - - apply coproduct_op. + - nrapply coproduct_op. exact _. - snrapply catie_homotopic. + simpl; exact (cat_coprod_prod (A:=A) x). + simpl; exact _. + (** Showing that these two maps are homotopic is a bit tricky. *) - apply cat_coprod_in_eta. + nrapply cat_coprod_in_eta. intros i. - apply cat_prod_pr_eta. + nrapply cat_prod_pr_eta. intros j. simpl. refine (cat_assoc _ _ _ $@ _). From ef2848aa1d464468d7124704d1457e62c6990c9d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Apr 2024 03:24:17 +0100 Subject: [PATCH 14/31] simplify proof of opposite coproducts Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Abelian.v | 12 ++++++------ theories/WildCat/Coproducts.v | 14 ++------------ 2 files changed, 8 insertions(+), 18 deletions(-) diff --git a/theories/Algebra/Homological/Abelian.v b/theories/Algebra/Homological/Abelian.v index c3750c0ffb7..6ba9ae0fd67 100644 --- a/theories/Algebra/Homological/Abelian.v +++ b/theories/Algebra/Homological/Abelian.v @@ -39,7 +39,7 @@ Definition ker_zero_monic {A : Type} `{IsPointedCat A} : Monic f -> cat_ker f $== zero_morphism. Proof. intros monic. - apply monic. + nrapply monic. refine (cat_ker_zero f $@ _^$). apply cat_zero_r. Defined. @@ -50,7 +50,7 @@ Definition monic_ker_zero {A : Type} `{IsAdditive A} : cat_ker f $== zero_morphism -> Monic f. Proof. intros ker_zero c g h p. - apply GpdHom_path. + nrapply GpdHom_path. apply path_hom in p. change (@paths (AbHom c a) g h). change (@paths (AbHom c b) (f $o g) (f $o h)) in p. @@ -64,7 +64,7 @@ Proof. f_ap. apply path_hom. symmetry. - apply addcat_comp_negate_r. + nrapply addcat_comp_negate_r. Defined. (** *** Cokernels *) @@ -125,13 +125,13 @@ Proof. snrapply cat_coker_rec. - snrapply cat_ker_corec. + exact f. - + apply cat_coker_zero. + + nrapply cat_coker_zero. - snrapply monic_cat_ker. refine ((cat_assoc _ _ _)^$ $@ _). refine ((_ $@R _) $@ _). - 1: apply cat_ker_corec_beta. + 1: nrapply cat_ker_corec_beta. refine (_ $@ (cat_zero_r _)^$). - apply cat_ker_zero. + nrapply cat_ker_zero. Defined. (** ** Abelian categories *) diff --git a/theories/WildCat/Coproducts.v b/theories/WildCat/Coproducts.v index c3bd206ef62..c72bee7c6a2 100644 --- a/theories/WildCat/Coproducts.v +++ b/theories/WildCat/Coproducts.v @@ -308,17 +308,7 @@ Defined. Definition coproduct_op {I A : Type} (x : I -> A) `{Is1Cat A} {H' : Product I x} - : Coproduct I (A:=A^op) x. -Proof. - snrapply Build_Product. - - exact (cat_prod I x). - - exact cat_pr. - - exact (fun z => cat_prod_corec I). - - intros z f i. - apply cat_prod_beta. - - intros z f g p. - apply cat_prod_pr_eta. - exact p. -Defined. + : Coproduct I (A:=A^op) x + := H'. Hint Immediate coproduct_op : typeclass_instances. From dc29a67f9f0a7b6c1930ee31fddb6f552496f340 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Apr 2024 15:08:53 +0100 Subject: [PATCH 15/31] AbGroup has biproducts Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/AbPushout.v | 2 +- theories/Algebra/AbGroups/Biproduct.v | 74 ++++++++----- theories/Algebra/AbSES/Core.v | 2 +- theories/Algebra/Homological/Additive.v | 13 +-- theories/Algebra/Homological/Biproducts.v | 120 +++++++++++++++++++++- theories/Types/Bool.v | 10 ++ theories/WildCat/Coproducts.v | 9 +- 7 files changed, 195 insertions(+), 35 deletions(-) diff --git a/theories/Algebra/AbGroups/AbPushout.v b/theories/Algebra/AbGroups/AbPushout.v index 0f89c933c0d..26793306e21 100644 --- a/theories/Algebra/AbGroups/AbPushout.v +++ b/theories/Algebra/AbGroups/AbPushout.v @@ -71,7 +71,7 @@ Proof. srapply path_sigma_hprop. refine (grp_quotient_rec_beta _ Y _ _ @ _). apply equiv_path_grouphomomorphism; intro bc. - exact (ab_biprod_rec_beta' (phi $o grp_quotient_map) bc). + exact (ab_biprod_rec_eta (phi $o grp_quotient_map) bc). Defined. (** Restricting [ab_pushout_rec] along [ab_pushout_inl] recovers the left inducing map. *) diff --git a/theories/Algebra/AbGroups/Biproduct.v b/theories/Algebra/AbGroups/Biproduct.v index 57737bb8ebb..12c04f9ad08 100644 --- a/theories/Algebra/AbGroups/Biproduct.v +++ b/theories/Algebra/AbGroups/Biproduct.v @@ -3,6 +3,7 @@ Require Import WildCat. Require Import HSet. Require Import AbelianGroup. Require Import Modalities.ReflectiveSubuniverse. +Require Import Algebra.Homological.Biproducts. Local Open Scope mc_add_scope. @@ -47,7 +48,7 @@ Proof. intros [f g]. exact (ab_biprod_rec f g). Defined. -Proposition ab_biprod_rec_beta' {A B Y : AbGroup} +Proposition ab_biprod_rec_eta {A B Y : AbGroup} (u : ab_biprod A B $-> Y) : ab_biprod_rec (u $o ab_biprod_inl) (u $o ab_biprod_inr) == u. Proof. @@ -58,29 +59,19 @@ Proof. - exact (left_identity b). Defined. -Proposition ab_biprod_rec_beta `{Funext} {A B Y : AbGroup} - (u : ab_biprod A B $-> Y) - : ab_biprod_rec (u $o ab_biprod_inl) (u $o ab_biprod_inr) = u. -Proof. - apply equiv_path_grouphomomorphism. - exact (ab_biprod_rec_beta' u). -Defined. - -Proposition ab_biprod_rec_inl_beta `{Funext} {A B Y : AbGroup} +Proposition ab_biprod_rec_inl_beta {A B Y : AbGroup} (a : A $-> Y) (b : B $-> Y) - : (ab_biprod_rec a b) $o ab_biprod_inl = a. + : (ab_biprod_rec a b) $o ab_biprod_inl == a. Proof. - apply equiv_path_grouphomomorphism. intro x; simpl. rewrite (grp_homo_unit b). exact (right_identity (a x)). Defined. -Proposition ab_biprod_rec_inr_beta `{Funext} {A B Y : AbGroup} +Proposition ab_biprod_rec_inr_beta {A B Y : AbGroup} (a : A $-> Y) (b : B $-> Y) - : (ab_biprod_rec a b) $o ab_biprod_inr = b. + : (ab_biprod_rec a b) $o ab_biprod_inr == b. Proof. - apply equiv_path_grouphomomorphism. intro y; simpl. rewrite (grp_homo_unit a). exact (left_identity (b y)). @@ -93,11 +84,14 @@ Proof. - intro phi. exact (phi $o ab_biprod_inl, phi $o ab_biprod_inr). - intro phi. - exact (ab_biprod_rec_beta phi). + apply equiv_path_grouphomomorphism. + exact (ab_biprod_rec_eta phi). - intros [a b]. apply path_prod. - + apply ab_biprod_rec_inl_beta. - + apply ab_biprod_rec_inr_beta. + + apply equiv_path_grouphomomorphism. + apply ab_biprod_rec_inl_beta. + + apply equiv_path_grouphomomorphism. + apply ab_biprod_rec_inr_beta. Defined. (** Corecursion principle, inherited from Groups/Group.v. *) @@ -105,7 +99,7 @@ Definition ab_biprod_corec {A B X : AbGroup} (f : X $-> A) (g : X $-> B) : X $-> ab_biprod A B := grp_prod_corec f g. -Definition ab_corec_beta {X Y A B : AbGroup} (f : X $-> Y) (g0 : Y $-> A) (g1 : Y $-> B) +Definition ab_corec_eta {X Y A B : AbGroup} (f : X $-> Y) (g0 : Y $-> A) (g1 : Y $-> B) : ab_biprod_corec g0 g1 $o f $== ab_biprod_corec (g0 $o f) (g1 $o f) := fun _ => idpath. @@ -190,6 +184,18 @@ Proof. - exact (left_identity _)^. Defined. +Lemma ab_biprod_corec_eta' {A B X : AbGroup} (f g : ab_biprod A B $-> X) + : (f $o ab_biprod_inl $== g $o ab_biprod_inl) + -> (f $o ab_biprod_inr $== g $o ab_biprod_inr) + -> f $== g. +Proof. + intros h k. + intros [a b]. + refine (ap f (ab_biprod_decompose _ _) @ _ @ ap g (ab_biprod_decompose _ _)^). + refine (grp_homo_op _ _ _ @ _ @ (grp_homo_op _ _ _)^). + exact (ap011 (+) (h a) (k b)). +Defined. + (* Maps out of biproducts are determined on the two inclusions. *) Lemma equiv_path_biprod_corec `{Funext} {A B X : AbGroup} (phi psi : ab_biprod A B $-> X) : ((phi $o ab_biprod_inl == psi $o ab_biprod_inl) * (phi $o ab_biprod_inr == psi $o ab_biprod_inr)) @@ -197,10 +203,7 @@ Lemma equiv_path_biprod_corec `{Funext} {A B X : AbGroup} (phi psi : ab_biprod A Proof. apply equiv_iff_hprop. - intros [h k]. - intros [a b]. - refine (ap phi (ab_biprod_decompose _ _) @ _ @ ap psi (ab_biprod_decompose _ _)^). - refine (grp_homo_op _ _ _ @ _ @ (grp_homo_op _ _ _)^). - exact (ap011 (+) (h a) (k b)). + apply ab_biprod_corec_eta'; assumption. - intro h. exact (fun a => h _, fun b => h _). Defined. @@ -311,3 +314,28 @@ Proof. refine (abgroup_commutative _ _ _ @ _). exact (ap (fun a => a + snd x) (abgroup_commutative _ _ _)). Defined. + +(** ** AbGroup has binary biproducts *) + +Global Instance hasbinarybiproducts_ab : HasBinaryBiproducts AbGroup. +Proof. + intros A B. + snrapply Build_BinaryBiproduct. + - exact (ab_biprod A B). + - exact ab_biprod_pr1. + - exact ab_biprod_pr2. + - exact (fun _ => ab_biprod_corec). + - exact (fun _ f _ => Id f). + - exact (fun _ _ g => Id g). + - exact (fun _ _ _ p q a => path_prod' (p a) (q a)). + - exact ab_biprod_inl. + - exact ab_biprod_inr. + - exact (fun _ => ab_biprod_rec). + - exact (fun _ => ab_biprod_rec_inl_beta). + - exact (fun _ => ab_biprod_rec_inr_beta). + - exact (fun _ => ab_biprod_corec_eta'). + - cbn; reflexivity. + - cbn; reflexivity. + - cbn; reflexivity. + - cbn; reflexivity. +Defined. diff --git a/theories/Algebra/AbSES/Core.v b/theories/Algebra/AbSES/Core.v index 2601b0ac0b0..19e52086f48 100644 --- a/theories/Algebra/AbSES/Core.v +++ b/theories/Algebra/AbSES/Core.v @@ -625,7 +625,7 @@ Proposition projection_split_beta {B A : AbGroup} (E : AbSES B A) : projection_split_iso E h o (inclusion _) == ab_biprod_inl. Proof. intro a. - refine (ap _ (ab_corec_beta _ _ _ _) @ _). + refine (ap _ (ab_corec_eta _ _ _ _) @ _). refine (ab_biprod_functor_beta _ _ _ _ _ @ _). nrapply path_prod'. 2: rapply cx_isexact. diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index 2311806e934..e865911b92e 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -461,14 +461,11 @@ Defined. Global Instance issemiadditive_op {A : Type} `{IsSemiAdditive A} : IsSemiAdditive A^op. Proof. - snrapply Build_IsSemiAdditive. - - exact _. - - exact _. - - exact _. - - intros a b. - change A in a, b. - change (IsHSet (b $-> a)). - exact _. + snrapply Build_IsSemiAdditive; try exact _. + intros a b. + change A in a, b. + change (IsHSet (b $-> a)). + exact _. Defined. Axiom sorry : Empty. diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 623231c805d..24218a5fd1b 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -17,6 +17,7 @@ Class Biproduct (I : Type) `{DecidablePaths I} {A : Type} }. Arguments Biproduct I {_ A _ _ _ _ _ _} x. +Arguments Build_Biproduct' I {_ A _ _ _ _ _ _} x _ _ _. Section Biproducts. Context {I : Type} {A : Type} (x : I -> A) `{Biproduct I A x}. @@ -103,6 +104,62 @@ Arguments cat_biprod_pr {I A x _ _ _ _ _ _ _ _} i. Arguments cat_biprod_in {I A x _ _ _ _ _ _ _ _} i. Arguments cate_coprod_prod {I A} x {_ _ _ _ _ _ _ _}. +(** A smart constructor for biproducts. Note that compared to a typical definition of a biproduct, we have not asked for the sum of inclusions composed with projections to be the identity. There are a few reasons for this: In our general setting [I] may be infinite, so it doesn't make sense to ask for an infinite sum; Even if we somehow axiomtised something about the sum, we later show in [Additive.v] that the sum is infact an emergent property; Finally, it doesn't seem necessary to produce an isomorphism since we take that to be the identity in this lemma. We should also note that we require [A] to be a pointed category. In theory this requirement could be dropped. To do this we would weaken the equation showing [cat_pr j $o cat_in i] for different [i] and [j], to postulate that [cat_pr j $o cat_in i] is a split idempotent. This doesn't seem particularly useful however as the finite biproduct case implies the existance of zero objects. *) +Definition Build_Biproduct (I : Type) `{DecidablePaths I} + {A : Type} `{HasEquivs A, !IsPointedCat A} (x : I -> A) + (cat_biprod : A) + (** A biproduct is a product. *) + (cat_pr : forall i, cat_biprod $-> x i) + (corec : forall z, (forall i, z $-> x i) -> z $-> cat_biprod) + (corec_beta : forall z f i, cat_pr i $o corec z f $== f i) + (corec_eta : forall z (f g : z $-> cat_biprod), + (forall i, cat_pr i $o f $== cat_pr i $o g) -> f $== g) + (** A biproduct is a coproduct. *) + (cat_in : forall i, x i $-> cat_biprod) + (rec : forall z, (forall i, x i $-> z) -> cat_biprod $-> z) + (rec_beta : forall z f i, rec z f $o cat_in i $== f i) + (rec_eta : forall z (f g : cat_biprod $-> z), + (forall i, f $o cat_in i $== g $o cat_in i) -> f $== g) + (** The projections and inclusion maps satisfy some further properties. *) + (cat_pr_in : forall i, cat_pr i $o cat_in i $== Id _) + (cat_pr_in_ne : forall i j, i <> j -> cat_pr j $o cat_in i $== zero_morphism) + : Biproduct I x. +Proof. + (** We can show these directly when building the biproduct, however we would like to refer to them between goals so we prove them here first. *) + transparent assert (product : (Product I x)). + { snrapply Build_Product. + - exact cat_biprod. + - exact cat_pr. + - exact corec. + - exact corec_beta. + - exact corec_eta. } + transparent assert (coproduct : (Coproduct I x)). + { snrapply Build_Coproduct. + - exact cat_biprod. + - exact cat_in. + - exact rec. + - exact rec_beta. + - exact rec_eta. } + assert (r : (cat_coprod_prod x $== Id cat_biprod)). + { nrapply rec_eta. + intros i. + refine (rec_beta _ _ _ $@ _ $@ (cat_idl _)^$). + nrapply corec_eta. + intros j. + refine (corec_beta _ _ _ $@ _^$). + destruct (dec_paths i j) as [p|np]. + * destruct p. + exact (cat_pr_in i). + * exact (cat_pr_in_ne i j np). } + snrapply Build_Biproduct'. + - exact product. + - exact coproduct. + - snrapply catie_adjointify. + + exact (Id _). + + exact (cat_idr _ $@ r). + + exact (cat_idl _ $@ r). +Defined. + (** An inclusion followed by a projection of the same index is the identity. *) Definition cat_biprod_pr_in (I : Type) {A : Type} (x : I -> A) `{Biproduct I A x} (i : I) @@ -174,7 +231,68 @@ Class HasBiproducts (I A : Type) Class BinaryBiproduct {A : Type} `{HasEquivs A, !IsPointedCat A} (x y : A) := binary_biproduct :: Biproduct Bool (fun b => if b then x else y). - + +Definition Build_BinaryBiproduct' {A : Type} `{HasEquivs A, !IsPointedCat A} + {x y : A} {p : BinaryProduct x y} {c : BinaryCoproduct x y} + {ie : CatIsEquiv (cat_bincoprod_binprod x y)} + : BinaryBiproduct x y + := Build_Biproduct' _ _ p c ie. + +(** Smart constructor for binary biproducts. *) +Definition Build_BinaryBiproduct {A : Type} `{HasEquivs A, !IsPointedCat A} + {x y : A} (cat_biprod : A) + (** A binary biproduct is a product. *) + (cat_pr1 : cat_biprod $-> x) (cat_pr2 : cat_biprod $-> y) + (corec : forall z, (z $-> x) -> (z $-> y) -> z $-> cat_biprod) + (corec_beta_pr1 : forall z f g, cat_pr1 $o corec z f g $== f) + (corec_beta_pr2 : forall z f g, cat_pr2 $o corec z f g $== g) + (corec_eta : forall z (f g : z $-> cat_biprod), + cat_pr1 $o f $== cat_pr1 $o g -> cat_pr2 $o f $== cat_pr2 $o g -> f $== g) + (** A biproduct is a coproduct. *) + (cat_inl : x $-> cat_biprod) (cat_inr : y $-> cat_biprod) + (rec : forall z, (x $-> z) -> (y $-> z) -> cat_biprod $-> z) + (rec_beta_inl : forall z f g, rec z f g $o cat_inl $== f) + (rec_beta_inr : forall z f g, rec z f g $o cat_inr $== g) + (rec_eta : forall z (f g : cat_biprod $-> z), + f $o cat_inl $== g $o cat_inl -> f $o cat_inr $== g $o cat_inr -> f $== g) + (** The projections and inclusion maps satisfy some further properties. *) + (cat_pr1_inl : cat_pr1 $o cat_inl $== Id _) + (cat_pr2_inr : cat_pr2 $o cat_inr $== Id _) + (cat_pr1_inr : cat_pr1 $o cat_inr $== zero_morphism) + (cat_pr2_inl : cat_pr2 $o cat_inl $== zero_morphism) + : BinaryBiproduct x y. +Proof. + snrapply Build_Biproduct. + - exact cat_biprod. + - intros [ | ]. + + exact cat_pr1. + + exact cat_pr2. + - intros z f. + exact (corec z (f true) (f false)). + - intros z f [ | ]. + + exact (corec_beta_pr1 z (f true) (f false)). + + exact (corec_beta_pr2 z (f true) (f false)). + - intros z f g p. + exact (corec_eta z _ _ (p true) (p false)). + - intros [ | ]. + + exact cat_inl. + + exact cat_inr. + - intros z f. + exact (rec z (f true) (f false)). + - intros z f [ | ]. + + exact (rec_beta_inl z (f true) (f false)). + + exact (rec_beta_inr z (f true) (f false)). + - intros z f g p. + exact (rec_eta z _ _ (p true) (p false)). + - intros [ | ]. + + exact cat_pr1_inl. + + exact cat_pr2_inr. + - intros i j ne. + destruct (negb_ne' ne), i as [ | ]. + + exact cat_pr2_inl. + + exact cat_pr1_inr. +Defined. + Class HasBinaryBiproducts (A : Type) `{HasEquivs A, !IsPointedCat A} := has_binary_biproducts :: forall (x y : A), BinaryBiproduct x y. diff --git a/theories/Types/Bool.v b/theories/Types/Bool.v index 830a5b0987e..f13a2530159 100644 --- a/theories/Types/Bool.v +++ b/theories/Types/Bool.v @@ -92,6 +92,16 @@ Proof. - intros oops; case (oops idpath). Defined. +(** This version of [negb_ne] is more convenient to [destruct] against. *) +Definition negb_ne' {b1 b2 : Bool} + : (b1 <> b2) -> (negb b1 = b2). +Proof. + intros oops. + symmetry. + apply negb_ne. + intros p; symmetry in p; contradiction. +Defined. + (** ** Products as [forall] over [Bool] *) Section BoolForall. diff --git a/theories/WildCat/Coproducts.v b/theories/WildCat/Coproducts.v index c72bee7c6a2..046a01f77a1 100644 --- a/theories/WildCat/Coproducts.v +++ b/theories/WildCat/Coproducts.v @@ -51,7 +51,7 @@ Section Lemmata. : yon_0gpd z (cat_coprod I x) $<~> prod_0gpd I (fun i => yon_0gpd z (x i)) := cate_cat_prod_corec_inv I (A:=A^op) (x:=x). - Definition cate_cate_coprod_rec {z : A} + Definition cate_cat_coprod_rec {z : A} : prod_0gpd I (fun i => yon_0gpd z (x i)) $<~> yon_0gpd z (cat_coprod I x) := cate_cat_prod_corec I (A:=A^op) (x:=x). @@ -304,6 +304,13 @@ Proof. - apply zero_morphism. Defined. +Definition cat_bincoprod_binprod {A : Type} `{Is1Cat A, !IsPointedCat A} + (x y : A) `{!BinaryCoproduct x y, !BinaryProduct x y} + : cat_bincoprod x y $-> cat_binprod x y. +Proof. + nrapply cat_coprod_prod; exact _. +Defined. + (** *** Coproducts in the opposite category *) Definition coproduct_op {I A : Type} (x : I -> A) From 74c28c66c66866d1945fbc4faa13bde269b5e9d9 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Apr 2024 15:36:01 +0100 Subject: [PATCH 16/31] AbGroup is additive Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/AbHom.v | 29 +++++++++++++++++++++++++ theories/Algebra/AbGroups/Biproduct.v | 3 +++ theories/Algebra/Homological/Additive.v | 25 +-------------------- theories/Algebra/Homological/End.v | 29 +++++++++++++++++++++++++ 4 files changed, 62 insertions(+), 24 deletions(-) create mode 100644 theories/Algebra/Homological/End.v diff --git a/theories/Algebra/AbGroups/AbHom.v b/theories/Algebra/AbGroups/AbHom.v index 90880e2d074..aa60b1e129b 100644 --- a/theories/Algebra/AbGroups/AbHom.v +++ b/theories/Algebra/AbGroups/AbHom.v @@ -1,6 +1,8 @@ Require Import Basics Types. Require Import WildCat HSet Truncations.Core Modalities.ReflectiveSubuniverse. Require Import AbelianGroup Biproduct. +Require Import Algebra.Groups.Group. +Require Import Algebra.Homological.Additive. (** * Homomorphisms from a group to an abelian group form an abelian group. *) @@ -117,3 +119,30 @@ Proof. rapply (conn_map_elim (Tr (-1)) f). exact (equiv_path_grouphomomorphism^-1 p). Defined. + +(** ** Additivity of AbGroup *) + +(** Here is a sanity check that the abelian group structure that is induced from the semiadditive structure is homotopic to the handcrafted operation on hom. *) +Definition ab_homo_add_is_semiadditive_add `{Funext} + {A B : AbGroup} (f g : A $-> B) + : ab_homo_add f g = sgop_hom A B f g. +Proof. + apply equiv_path_grouphomomorphism. + reflexivity. +Defined. + +(** AbGroup is an additive category. *) +Global Instance additive_ab `{Funext} : IsAdditive AbGroup. +Proof. + snrapply Build_IsAdditive. + - exact _. + - exact _. + - intros A B f. + snrapply equiv_path_grouphomomorphism. + intros x. + exact (left_inverse (f x)). + - intros A B f. + snrapply equiv_path_grouphomomorphism. + intros x. + exact (right_inverse (f x)). +Defined. diff --git a/theories/Algebra/AbGroups/Biproduct.v b/theories/Algebra/AbGroups/Biproduct.v index 12c04f9ad08..a42bf877445 100644 --- a/theories/Algebra/AbGroups/Biproduct.v +++ b/theories/Algebra/AbGroups/Biproduct.v @@ -4,6 +4,7 @@ Require Import HSet. Require Import AbelianGroup. Require Import Modalities.ReflectiveSubuniverse. Require Import Algebra.Homological.Biproducts. +Require Import Algebra.Homological.Additive. Local Open Scope mc_add_scope. @@ -339,3 +340,5 @@ Proof. - cbn; reflexivity. - cbn; reflexivity. Defined. + +Global Instance issemiadditive_ab `{Funext} : IsSemiAdditive AbGroup := {}. diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index e865911b92e..0facbf5452e 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -3,7 +3,7 @@ Require Import WildCat.Core WildCat.Equiv WildCat.PointedCat WildCat.Bifunctor. Require Import WildCat.Square WildCat.Opposite. Require Import Algebra.Homological.Biproducts. Require Import Algebra.Groups.Group. -Require Import Algebra.AbGroups.AbelianGroup Algebra.Rings.Ring. +Require Import Algebra.AbGroups.AbelianGroup. Require Import canonical_names. (** * Semiadditive and Additive Categories *) @@ -432,29 +432,6 @@ Proof. nrapply addcat_comp_negate_id_r. Defined. -(** *** Endomorphism ring *) - -Definition End {A : Type} `{IsAdditive A} (X : A) : Ring. -Proof. - snrapply Build_Ring'; repeat split. - - exact (AbHom X X). - - exact (fun f g => f $o g). - - exact (Id _). - - exact left_heterodistribute_hom. - - exact right_heterodistribute_hom. - - exact _. - - intros f g h. - apply path_hom. - symmetry. - nrapply cat_assoc. - - intros f. - apply path_hom. - nrapply cat_idl. - - intros g. - apply path_hom. - nrapply cat_idr. -Defined. - (** ** Properties of additive categories *) (** The opposite of a semiadditive category is semiadditive. *) diff --git a/theories/Algebra/Homological/End.v b/theories/Algebra/Homological/End.v new file mode 100644 index 00000000000..9ca64185dc4 --- /dev/null +++ b/theories/Algebra/Homological/End.v @@ -0,0 +1,29 @@ +Require Import Basics.Overture. +Require Import WildCat.Core. +Require Import Algebra.AbGroups.AbelianGroup Algebra.Rings.Ring. +Require Import canonical_names. + +Require Import Algebra.Homological.Additive. + +(** * Endomorphism Rings *) + +Definition End {A : Type} `{IsAdditive A} (X : A) : Ring. +Proof. + snrapply Build_Ring'; repeat split. + - exact (AbHom X X). + - exact (fun f g => f $o g). + - exact (Id _). + - exact left_heterodistribute_hom. + - exact right_heterodistribute_hom. + - exact _. + - intros f g h. + apply path_hom. + symmetry. + nrapply cat_assoc. + - intros f. + apply path_hom. + nrapply cat_idl. + - intros g. + apply path_hom. + nrapply cat_idr. +Defined. From 74ea8ee48ee412474ab7d1ff0a238a907855aa86 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Apr 2024 15:54:10 +0100 Subject: [PATCH 17/31] remove try and use goal selectors Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Additive.v | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index 0facbf5452e..d53813bea42 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -438,7 +438,8 @@ Defined. Global Instance issemiadditive_op {A : Type} `{IsSemiAdditive A} : IsSemiAdditive A^op. Proof. - snrapply Build_IsSemiAdditive; try exact _. + snrapply Build_IsSemiAdditive. + 1-3: exact _. intros a b. change A in a, b. change (IsHSet (b $-> a)). From d7a027b281935818b680668e9b283d2bcaae482f Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Mon, 29 Apr 2024 11:43:04 -0400 Subject: [PATCH 18/31] Homological/Additive: fix a few typos --- theories/Algebra/Homological/Additive.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index d53813bea42..c40c5242c27 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -11,7 +11,7 @@ Require Import canonical_names. (** ** Semiadditive Categories *) -(** As semiadditive category is a a wild category with equivalences and the following data: *) +(** As semiadditive category is a wild category with equivalences and the following data: *) Class IsSemiAdditive (A : Type) `{HasEquivs A} := { (** It is a pointed category. *) semiadditive_pointed :: IsPointedCat A; @@ -19,11 +19,11 @@ Class IsSemiAdditive (A : Type) `{HasEquivs A} := { semiadditive_hasbiproducts :: HasBinaryBiproducts A; (** We have morphism extensionality. *) semiadditive_hasmorext :: HasMorExt A; - (** The homs are set. *) + (** The homs are sets. *) semiadditive_ishset_hom :: forall (a b : A), IsHSet (a $-> b); }. -(** The final two conditions in the definition of a semiadditive category ensure that the hom types can become commutative monoids. This is an essential chracteristic of semiadditive categories making it equivalent to alternate defintiions where the category is semiadditive if it is enriched in commutative monoids. The machinary of encriched categories however is a bit heavy so we use this more lightweight definition where the commutative monoid structure appears naturally. *) +(** The final two conditions in the definition of a semiadditive category ensure that the hom types can become commutative monoids. This is an essential chracteristic of semiadditive categories making it equivalent to alternate definitions where the category is semiadditive if it is enriched in commutative monoids. The machinary of encriched categories however is a bit heavy so we use this more lightweight definition where the commutative monoid structure appears naturally. *) Section CMonHom. @@ -323,7 +323,7 @@ Class IsAdditive (A : Type) `{HasEquivs A} := { (** *** Abelian Group structure on Hom *) -(** Homs in an aditive category form an abelian group. *) +(** Homs in an additive category form abelian groups. *) Definition AbHom {A : Type} `{IsAdditive A} : A -> A -> AbGroup. Proof. intros a b. From 62c2290b8e703e2a4eadc11e7611123a1a446d57 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Apr 2024 18:03:30 +0100 Subject: [PATCH 19/31] review suggestions Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Additive.v | 5 ++--- theories/Algebra/Homological/Biproducts.v | 5 +---- 2 files changed, 3 insertions(+), 7 deletions(-) diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index c40c5242c27..5e199ddf5ac 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -45,9 +45,8 @@ Section CMonHom. Proof. intros g. apply path_hom. - refine (cat_assoc _ _ _ $@ _). - refine ((_ $@L cat_assoc _ _ _) $@ _). - refine ((cat_assoc _ _ _)^$ $@ _). + unfold sgop_hom, fmap11. + nrefine ((cat_assoc_opp _ _ _ $@R _) $@ cat_assoc _ _ _ $@ _). snrefine ((_ $@L _) $@ _). 1: exact cat_binbiprod_inr. { nrapply cat_binbiprod_pr_eta. diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 24218a5fd1b..5e61799bfab 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -154,10 +154,7 @@ Proof. snrapply Build_Biproduct'. - exact product. - exact coproduct. - - snrapply catie_adjointify. - + exact (Id _). - + exact (cat_idr _ $@ r). - + exact (cat_idl _ $@ r). + - exact (catie_homotopic _ r^$). Defined. (** An inclusion followed by a projection of the same index is the identity. *) From 365399be5dc54bc9c88ec567122a7a66d69d692b Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Apr 2024 18:24:41 +0100 Subject: [PATCH 20/31] spelling Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Additive.v | 2 +- theories/Algebra/Homological/Biproducts.v | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index 5e199ddf5ac..1a3dcad5370 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -23,7 +23,7 @@ Class IsSemiAdditive (A : Type) `{HasEquivs A} := { semiadditive_ishset_hom :: forall (a b : A), IsHSet (a $-> b); }. -(** The final two conditions in the definition of a semiadditive category ensure that the hom types can become commutative monoids. This is an essential chracteristic of semiadditive categories making it equivalent to alternate definitions where the category is semiadditive if it is enriched in commutative monoids. The machinary of encriched categories however is a bit heavy so we use this more lightweight definition where the commutative monoid structure appears naturally. *) +(** The final two conditions in the definition of a semiadditive category ensure that the hom types can become commutative monoids. This is an essential characteristic of semiadditive categories making it equivalent to alternate definitions where the category is semiadditive if it is enriched in commutative monoids. The machinery of encriched categories however is a bit heavy so we use this more lightweight definition where the commutative monoid structure appears naturally. *) Section CMonHom. diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 5e61799bfab..2a594806597 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -104,7 +104,7 @@ Arguments cat_biprod_pr {I A x _ _ _ _ _ _ _ _} i. Arguments cat_biprod_in {I A x _ _ _ _ _ _ _ _} i. Arguments cate_coprod_prod {I A} x {_ _ _ _ _ _ _ _}. -(** A smart constructor for biproducts. Note that compared to a typical definition of a biproduct, we have not asked for the sum of inclusions composed with projections to be the identity. There are a few reasons for this: In our general setting [I] may be infinite, so it doesn't make sense to ask for an infinite sum; Even if we somehow axiomtised something about the sum, we later show in [Additive.v] that the sum is infact an emergent property; Finally, it doesn't seem necessary to produce an isomorphism since we take that to be the identity in this lemma. We should also note that we require [A] to be a pointed category. In theory this requirement could be dropped. To do this we would weaken the equation showing [cat_pr j $o cat_in i] for different [i] and [j], to postulate that [cat_pr j $o cat_in i] is a split idempotent. This doesn't seem particularly useful however as the finite biproduct case implies the existance of zero objects. *) +(** A smart constructor for biproducts. Note that compared to a typical definition of a biproduct, we have not asked for the sum of inclusions composed with projections to be the identity. There are a few reasons for this: In our general setting [I] may be infinite, so it doesn't make sense to ask for an infinite sum; Even if we somehow axiomatized something about the sum, we later show in [Additive.v] that the sum is in fact an emergent property; Finally, it doesn't seem necessary to produce an isomorphism since we take that to be the identity in this lemma. We should also note that we require [A] to be a pointed category. In theory this requirement could be dropped. To do this we would weaken the equation showing [cat_pr j $o cat_in i] for different [i] and [j], to postulate that [cat_pr j $o cat_in i] is a split idempotent. This doesn't seem particularly useful however as the finite biproduct case implies the existence of zero objects. *) Definition Build_Biproduct (I : Type) `{DecidablePaths I} {A : Type} `{HasEquivs A, !IsPointedCat A} (x : I -> A) (cat_biprod : A) From bb6077b381187d7dae01d8e9dfe3bb45fb874e6a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Apr 2024 21:09:20 +0100 Subject: [PATCH 21/31] inline definition Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/AbHom.v | 11 ++++------- 1 file changed, 4 insertions(+), 7 deletions(-) diff --git a/theories/Algebra/AbGroups/AbHom.v b/theories/Algebra/AbGroups/AbHom.v index aa60b1e129b..7c28473fc00 100644 --- a/theories/Algebra/AbGroups/AbHom.v +++ b/theories/Algebra/AbGroups/AbHom.v @@ -123,13 +123,10 @@ Defined. (** ** Additivity of AbGroup *) (** Here is a sanity check that the abelian group structure that is induced from the semiadditive structure is homotopic to the handcrafted operation on hom. *) -Definition ab_homo_add_is_semiadditive_add `{Funext} - {A B : AbGroup} (f g : A $-> B) - : ab_homo_add f g = sgop_hom A B f g. -Proof. - apply equiv_path_grouphomomorphism. - reflexivity. -Defined. +Definition ab_homo_add_is_semiadditive_add + {A B : AbGroup} (f g : A $-> B) `{Funext} + : ab_homo_add f g = sgop_hom A B f g :> (A -> B) + := 1. (** AbGroup is an additive category. *) Global Instance additive_ab `{Funext} : IsAdditive AbGroup. From f37c55c4fd5675de698ce7a72f0b1c47ca73ae43 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Apr 2024 22:53:57 +0100 Subject: [PATCH 22/31] simplify left_identity in Additive.v Signed-off-by: Ali Caglayan --- theories/Algebra/Homological/Additive.v | 34 +++------- theories/Algebra/Homological/Biproducts.v | 83 +++++++++++++++++++---- 2 files changed, 76 insertions(+), 41 deletions(-) diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index 1a3dcad5370..d745a9d6fd9 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -45,33 +45,15 @@ Section CMonHom. Proof. intros g. apply path_hom. - unfold sgop_hom, fmap11. - nrefine ((cat_assoc_opp _ _ _ $@R _) $@ cat_assoc _ _ _ $@ _). - snrefine ((_ $@L _) $@ _). - 1: exact cat_binbiprod_inr. - { nrapply cat_binbiprod_pr_eta. - - refine (_ $@ cat_binbiprod_pr1_inr^$). - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: nrapply cat_binbiprod_corec_beta_pr1. - refine (cat_assoc _ _ _ $@ _). - nrapply cat_zero_l. - - refine (_ $@ cat_binbiprod_pr2_inr^$). - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: nrapply cat_binbiprod_corec_beta_pr2. - refine (cat_assoc _ _ _ $@ _). - refine (cat_idl _ $@ _). - nrapply cat_binbiprod_corec_beta_pr2. } - nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - { refine ((_ $@R _) $@ _). - 1: nrapply cat_binbiprod_corec_rec. - nrapply cat_binbiprod_rec_beta_inr. } - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: nrapply cat_binbiprod_rec_beta_inr. - nrapply cat_idl. + unfold sgop_hom. + nrefine (cat_assoc _ _ _ $@ (_ $@L (_ $@ _)) $@ _). + 1: nrapply cat_binbiprod_fmap11_corec. + 1: nrapply (cat_binbiprod_corec_eta' (cat_idr _) (cat_idr _) $@ _). + 1: exact (cat_binbiprod_corec_zero_inr g). + nrefine (cat_assoc_opp _ _ _ $@ (_ $@R _) $@ cat_idl _). + nrapply cat_binbiprod_rec_beta_inr. Defined. - + Local Existing Instance symmetricbraiding_cat_binbiprod. (** Using the naturality of swap we can show that the operation is commutative. *) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 2a594806597..97aeab4a8af 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -450,6 +450,26 @@ Section BinaryBiproducts. nrapply (not_fixed_negb true). Defined. + Definition cat_binbiprod_corec_zero_inl {z} (f : z $-> x) + : cat_binbiprod_corec f zero_morphism $== cat_binbiprod_inl $o f. + Proof. + snrapply cat_binbiprod_pr_eta. + - refine (cat_binbiprod_corec_beta_pr1 _ _ $@ _^$). + exact ((cat_assoc _ _ _)^$ $@ (cat_binbiprod_pr1_inl $@R _) $@ cat_idl _). + - refine (cat_binbiprod_corec_beta_pr2 _ _ $@ _^$). + exact ((cat_assoc _ _ _)^$ $@ (cat_binbiprod_pr2_inl $@R _) $@ cat_zero_l _). + Defined. + + Definition cat_binbiprod_corec_zero_inr {z} (f : z $-> y) + : cat_binbiprod_corec zero_morphism f $== cat_binbiprod_inr $o f. + Proof. + snrapply cat_binbiprod_pr_eta. + - refine (cat_binbiprod_corec_beta_pr1 _ _ $@ _^$). + exact ((cat_assoc _ _ _)^$ $@ (cat_binbiprod_pr1_inr $@R _) $@ cat_zero_l _). + - refine (cat_binbiprod_corec_beta_pr2 _ _ $@ _^$). + exact ((cat_assoc _ _ _)^$ $@ (cat_binbiprod_pr2_inr $@R _) $@ cat_idl _). + Defined. + End BinaryBiproducts. Arguments cat_binbiprod {A _ _ _ _ _ _} x y {_}. @@ -517,30 +537,63 @@ Global Instance is1functor_cat_binbiprod_r {A : Type} : Is1Functor (fun y => cat_binbiprod x y) := bifunctor_is1functor01 x. -(** *** Diagonal lemmas *) +(** *** Functor lemmas *) + +Definition cat_binbiprod_fmap01_corec {A : Type} + `{HasEquivs A, !IsPointedCat A, !HasBinaryBiproducts A} {w x y z : A} + (f : w $-> z) (g : x $-> y) (h : w $-> x) + : fmap01 (fun x y => cat_binbiprod x y) z g $o cat_binbiprod_corec f h + $== cat_binbiprod_corec f (g $o h). +Proof. + snrapply cat_binbiprod_pr_eta. + - refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ cat_idl _ $@ _ $@ _^$). + 1-3: rapply cat_binbiprod_corec_beta_pr1. + - refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _) $@ _^$). + 1-3: rapply cat_binbiprod_corec_beta_pr2. +Defined. + +Definition cat_binbiprod_fmap10_corec {A : Type} + `{HasEquivs A, !IsPointedCat A, !HasBinaryBiproducts A} {w x y z : A} + (f : x $-> y) (g : w $-> x) (h : w $-> z) + : fmap10 (fun x y => cat_binbiprod x y) f z $o cat_binbiprod_corec g h + $== cat_binbiprod_corec (f $o g) h. +Proof. + snrapply cat_binbiprod_pr_eta. + - refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _) $@ _^$). + 1-3: nrapply cat_binbiprod_corec_beta_pr1. + - refine ((cat_assoc _ _ _)^$ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ cat_idl _ $@ _ $@ _^$). + 1-3: nrapply cat_binbiprod_corec_beta_pr2. +Defined. + +Definition cat_binbiprod_fmap11_corec {A : Type} + `{HasEquivs A, !IsPointedCat A, !HasBinaryBiproducts A} {v w x y z : A} + (f : w $-> y) (g : x $-> z) (h : v $-> w) (i : v $-> x) + : fmap11 (fun x y => cat_binbiprod x y) f g $o cat_binbiprod_corec h i + $== cat_binbiprod_corec (f $o h) (g $o i). +Proof. + nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + - nrapply cat_binbiprod_fmap10_corec. + - nrapply cat_binbiprod_fmap01_corec. +Defined. Definition cat_binbiprod_diag_fmap11 {A : Type} `{HasEquivs A, !IsPointedCat A, !HasBinaryBiproducts A} {x y : A} (f : x $-> y) : cat_binbiprod_diag y $o f $== fmap11 (fun x y => cat_binbiprod x y) f f $o cat_binbiprod_diag x. Proof. + refine (_ $@ _^$). + 2: nrapply cat_binbiprod_fmap11_corec. nrapply cat_binbiprod_pr_eta. - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ (_ $@ _)). - 1: rapply cat_binbiprod_corec_beta_pr1. - 1: nrapply cat_idl. - refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: rapply cat_pr1_fmap11_binprod. - refine ((cat_idr _)^$ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). - rapply cat_binbiprod_corec_beta_pr1. + refine ((_ $@R _) $@ cat_idl _ $@ (cat_idr _)^$ $@ _^$). + 1,2: rapply cat_binbiprod_corec_beta_pr1. - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ (_ $@ _)). - 1: rapply cat_binbiprod_corec_beta_pr2. - 1: nrapply cat_idl. - refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: rapply cat_pr2_fmap11_binprod. - refine ((cat_idr _)^$ $@ (_ $@L _^$) $@ (cat_assoc _ _ _)^$). - rapply cat_binbiprod_corec_beta_pr2. + refine ((_ $@R _) $@ cat_idl _ $@ (cat_idr _)^$ $@ _^$). + 1,2: rapply cat_binbiprod_corec_beta_pr2. Defined. Definition cat_binbiprod_codiag_fmap11 {A : Type} From 63c8cf7a2dc640bcdeccd72b191d3a1a1c1bec4f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Apr 2024 00:02:39 +0100 Subject: [PATCH 23/31] update comment and order of arguments Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/AbHom.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/theories/Algebra/AbGroups/AbHom.v b/theories/Algebra/AbGroups/AbHom.v index 7c28473fc00..553ae34e78d 100644 --- a/theories/Algebra/AbGroups/AbHom.v +++ b/theories/Algebra/AbGroups/AbHom.v @@ -122,9 +122,9 @@ Defined. (** ** Additivity of AbGroup *) -(** Here is a sanity check that the abelian group structure that is induced from the semiadditive structure is homotopic to the handcrafted operation on hom. *) -Definition ab_homo_add_is_semiadditive_add - {A B : AbGroup} (f g : A $-> B) `{Funext} +(** The group operation that is induced from the semiadditive structure agrees with the handcrafted operation on hom at the level of the underlying functions. *) +Definition ab_homo_add_is_semiadditive_add `{Funext} + {A B : AbGroup} (f g : A $-> B) : ab_homo_add f g = sgop_hom A B f g :> (A -> B) := 1. From acfe86040a614021f54114cd8a6626b46b774c2d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 1 May 2024 04:29:50 +0100 Subject: [PATCH 24/31] two proofs of pentagon and simplifications in biproducts Signed-off-by: Ali Caglayan --- theories/WildCat/Products.v | 481 ++++++++++++++++++++++-------------- 1 file changed, 292 insertions(+), 189 deletions(-) diff --git a/theories/WildCat/Products.v b/theories/WildCat/Products.v index a6adc9774ce..77e3c5ac6e3 100644 --- a/theories/WildCat/Products.v +++ b/theories/WildCat/Products.v @@ -331,6 +331,104 @@ Proof. + exact (p false). Defined. +Definition cat_binprod_eta_pr_x_xx {A : Type} `{HasBinaryProducts A} {w x y z : A} + (f g : w $-> cat_binprod x (cat_binprod y z)) + : cat_pr1 $o f $== cat_pr1 $o g + -> cat_pr1 $o cat_pr2 $o f $== cat_pr1 $o cat_pr2 $o g + -> cat_pr2 $o cat_pr2 $o f $== cat_pr2 $o cat_pr2 $o g + -> f $== g. +Proof. + intros p q r. + snrapply cat_binprod_eta_pr. + - exact p. + - snrapply cat_binprod_eta_pr. + + exact (cat_assoc_opp _ _ _ $@ q $@ cat_assoc _ _ _). + + exact (cat_assoc_opp _ _ _ $@ r $@ cat_assoc _ _ _). +Defined. + +Definition cat_binprod_eta_pr_xx_x {A : Type} `{HasBinaryProducts A} {w x y z : A} + (f g : w $-> cat_binprod (cat_binprod x y) z) + : cat_pr1 $o cat_pr1 $o f $== cat_pr1 $o cat_pr1 $o g + -> cat_pr2 $o cat_pr1 $o f $== cat_pr2 $o cat_pr1 $o g + -> cat_pr2 $o f $== cat_pr2 $o g + -> f $== g. +Proof. + intros p q r. + snrapply cat_binprod_eta_pr. + 2: exact r. + snrapply cat_binprod_eta_pr. + 1,2: refine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). + - exact p. + - exact q. +Defined. + +Definition cat_binprod_eta_pr_x_xx_id {A : Type} `{HasBinaryProducts A} {x y z : A} + (f : cat_binprod x (cat_binprod y z) $-> cat_binprod x (cat_binprod y z)) + : cat_pr1 $o f $== cat_pr1 + -> cat_pr1 $o cat_pr2 $o f $== cat_pr1 $o cat_pr2 + -> cat_pr2 $o cat_pr2 $o f $== cat_pr2 $o cat_pr2 + -> f $== Id _. +Proof. + intros p q r. + snrapply cat_binprod_eta_pr_x_xx. + - exact (p $@ (cat_idr _)^$). + - exact (q $@ (cat_idr _)^$). + - exact (r $@ (cat_idr _)^$). +Defined. + +Definition cat_binprod_eta_pr_x_xx_assoc {A : Type} `{HasBinaryProducts A} {u v w x y z : A} + (f : u $-> v) (g : u $-> w) (h : v $-> cat_binprod x (cat_binprod y z)) + (i : w $-> cat_binprod x (cat_binprod y z)) + : cat_pr1 $o h $o f $== cat_pr1 $o i $o g + -> cat_pr1 $o cat_pr2 $o h $o f $== cat_pr1 $o cat_pr2 $o i $o g + -> cat_pr2 $o cat_pr2 $o h $o f $== cat_pr2 $o cat_pr2 $o i $o g + -> h $o f $== i $o g. +Proof. + intros p q r. + snrapply cat_binprod_eta_pr_x_xx. + - exact (cat_assoc_opp _ _ _ $@ p $@ cat_assoc _ _ _). + - exact (cat_assoc_opp _ _ _ $@ q $@ cat_assoc _ _ _). + - exact (cat_assoc_opp _ _ _ $@ r $@ cat_assoc _ _ _). +Defined. + +Definition cat_binprod_eta_pr_x_x_xx {A : Type} `{HasBinaryProducts A} {v w x y z : A} + (f g : v $-> cat_binprod w (cat_binprod x (cat_binprod y z))) + : cat_pr1 $o f $== cat_pr1 $o g + -> cat_pr1 $o cat_pr2 $o f $== cat_pr1 $o cat_pr2 $o g + -> cat_pr1 $o cat_pr2 $o cat_pr2 $o f $== cat_pr1 $o cat_pr2 $o cat_pr2 $o g + -> cat_pr2 $o cat_pr2 $o cat_pr2 $o f $== cat_pr2 $o cat_pr2 $o cat_pr2 $o g + -> f $== g. +Proof. + intros p q r s. + snrapply cat_binprod_eta_pr_x_xx. + - exact p. + - exact q. + - snrapply cat_binprod_eta_pr. + 1,2: nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). + 1,2: nrefine ((cat_assoc_opp _ _ _ $@R _) $@ _ $@ (cat_assoc _ _ _ $@R _)). + + exact r. + + exact s. +Defined. + +Definition cat_binprod_eta_pr_xx_x_x {A : Type} `{HasBinaryProducts A} {v w x y z : A} + (f g : v $-> cat_binprod (cat_binprod (cat_binprod w x) y) z) + : cat_pr1 $o cat_pr1 $o cat_pr1 $o f $== cat_pr1 $o cat_pr1 $o cat_pr1 $o g + -> cat_pr2 $o cat_pr1 $o cat_pr1 $o f $== cat_pr2 $o cat_pr1 $o cat_pr1 $o g + -> cat_pr2 $o cat_pr1 $o f $== cat_pr2 $o cat_pr1 $o g + -> cat_pr2 $o f $== cat_pr2 $o g + -> f $== g. +Proof. + intros p q r s. + snrapply cat_binprod_eta_pr_xx_x. + 3: exact s. + 2: exact r. + nrefine (cat_assoc _ _ _ $@ _ $@ cat_assoc_opp _ _ _). + snrapply cat_binprod_eta_pr. + 1,2: do 2 nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). + - exact p. + - exact q. +Defined. + (** From binary products, all Bool-shaped products can be constructed. This should not be an instance to avoid a cycle with [hasbinaryproducts_hasproductsbool]. *) Definition hasproductsbool_hasbinaryproducts {A : Type} `{HasBinaryProducts A} : HasProducts Bool A. @@ -619,34 +717,37 @@ Section Associativity. - exact (cat_pr1 $o cat_pr2). - exact (fmap01 (fun x y => cat_binprod x y) x cat_pr2). Defined. + + Definition cat_binprod_pr1_twist (x y z : A) + : cat_pr1 $o cat_binprod_twist x y z $== cat_pr1 $o cat_pr2 + := cat_binprod_beta_pr1 _ _. + + Definition cat_binprod_pr1_pr2_twist (x y z : A) + : cat_pr1 $o cat_pr2 $o cat_binprod_twist x y z $== cat_pr1. + Proof. + nrefine (cat_assoc _ _ _ $@ _). + nrefine ((_ $@L cat_binprod_beta_pr2 _ _) $@ _). + nrapply cat_pr1_fmap01_binprod. + Defined. + + Definition cat_binprod_pr2_pr2_twist (x y z : A) + : cat_pr2 $o cat_pr2 $o cat_binprod_twist x y z $== cat_pr2 $o cat_pr2. + Proof. + nrefine (cat_assoc _ _ _ $@ _). + nrefine ((_ $@L cat_binprod_beta_pr2 _ _) $@ _). + nrapply cat_pr2_fmap01_binprod. + Defined. Lemma cat_binprod_twist_cat_binprod_twist (x y z : A) : cat_binprod_twist x y z $o cat_binprod_twist y x z $== Id _. Proof. - unfold cat_binprod_twist. - nrapply cat_binprod_eta_pr. - - refine ((cat_assoc _ _ _)^$ $@ _). - nrefine (cat_binprod_beta_pr1 _ _ $@R _ $@ _). - nrefine (cat_assoc _ _ _ $@ _). - nrefine (_ $@L cat_binprod_beta_pr2 _ _ $@ _). - nrefine (cat_binprod_beta_pr1 _ _ $@ _). - exact (cat_idl _ $@ (cat_idr _)^$). - - refine ((cat_assoc _ _ _)^$ $@ _). - nrefine (cat_binprod_beta_pr2 _ _ $@R _ $@ _). - nrapply cat_binprod_eta_pr. - + refine ((cat_assoc _ _ _)^$ $@ _). - nrefine (cat_binprod_beta_pr1 _ _ $@R _ $@ _). - nrefine (cat_assoc _ _ _ $@ cat_idl _ $@ _). - nrefine (cat_binprod_beta_pr1 _ _ $@ _). - nrefine (_ $@L _). - exact (cat_idr _)^$. - + refine ((cat_assoc _ _ _)^$ $@ _). - nrefine (cat_binprod_beta_pr2 _ _ $@R _ $@ _). - nrefine (cat_assoc _ _ _ $@ _). - nrefine (_ $@L cat_binprod_beta_pr2 _ _ $@ _). - nrefine (cat_binprod_beta_pr2 _ _ $@ _). - nrefine (_ $@L _). - exact (cat_idr _)^$. + nrapply cat_binprod_eta_pr_x_xx_id. + - nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_pr1_twist _ _ _ $@R _) $@ _). + nrapply cat_binprod_pr1_pr2_twist. + - nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_pr1_pr2_twist _ _ _ $@R _) $@ _). + nrapply cat_binprod_pr1_twist. + - nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_pr2_pr2_twist _ _ _ $@R _) $@ _). + nrapply cat_binprod_pr2_pr2_twist. Defined. Definition cate_binprod_twist (x y z : A) @@ -703,28 +804,37 @@ Section Associativity. - intros ? ? ? ? ? ?; exact cat_binprod_twist_nat. Defined. + Definition cat_pr1_pr1_associator_binprod x y z + : cat_pr1 $o cat_pr1 $o associator_binprod x y z $== cat_pr1. + Proof. + Admitted. + + Definition cat_pr2_pr1_associator_binprod x y z + : cat_pr2 $o cat_pr1 $o associator_binprod x y z $== cat_pr1 $o cat_pr2. + Proof. + Admitted. + + Definition cat_pr2_associator_binprod x y z + : cat_pr2 $o associator_binprod x y z $== cat_pr2 $o cat_pr2. + Proof. + Admitted. + Context (unit : A) `{!IsTerminal unit}. Local Instance right_unitor_binprod : RightUnitor (fun x y => cat_binprod x y) unit. Proof. snrapply Build_NatEquiv. - - intros a. - unfold flip. + - intros a; unfold flip. snrapply cate_adjointify. + exact cat_pr1. - + nrapply cat_binprod_corec. - * exact (Id _). - * exact (mor_terminal _ _). - + nrapply cat_binprod_beta_pr1. + + exact (cat_binprod_corec (Id _) (mor_terminal _ _)). + + exact (cat_binprod_beta_pr1 _ _). + nrapply cat_binprod_eta_pr. - * refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr1. + * nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr1 _ _ $@R _) $@ _). exact (cat_idl _ $@ (cat_idr _)^$). - * refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr2. - refine (_^$ $@ _). - 1,2: rapply mor_terminal_unique. + * nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr2 _ _ $@R _) $@ _). + exact ((mor_terminal_unique _ _ _)^$ $@ mor_terminal_unique _ _ _). - intros a b f. refine ((_ $@R _) $@ _ $@ (_ $@L _^$)). 1,3: nrapply cate_buildequiv_fun. @@ -738,28 +848,19 @@ Section Associativity. Proof. snrapply triangle_twist. intros a b. - refine (fmap02 _ _ _ $@ _). - 1: nrapply cate_buildequiv_fun. - refine (_ $@ ((_ $@L fmap02 _ _ _^$) $@R _)). - 2: nrapply cate_buildequiv_fun. + refine (fmap02 _ _ _ $@ _ $@ ((_ $@L fmap02 _ _ _^$) $@R _)). + 1,3: nrapply cate_buildequiv_fun. nrapply cat_binprod_eta_pr. - - refine (cat_pr1_fmap01_binprod _ _ $@ _). - nrefine (_ $@ cat_assoc _ _ _). + - nrefine (cat_pr1_fmap01_binprod _ _ $@ _ $@ cat_assoc _ _ _). refine (_ $@ (((_^$ $@R _) $@ cat_assoc _ _ _) $@R _)). 2: nrapply cat_binprod_beta_pr1. - refine (_ $@ (_^$ $@R _)). - 2: nrapply cat_pr2_fmap01_binprod. - refine (_ $@ (cat_assoc _ _ _)^$). - refine (_^$ $@ (_ $@L _^$)). - 2: nrapply cat_binprod_beta_pr2. - nrapply cat_pr1_fmap01_binprod. - - nrefine (_ $@ _). + refine ((_ $@R _) $@ _)^$. 1: nrapply cat_pr2_fmap01_binprod. - nrefine (_ $@ cat_assoc _ _ _). - refine (_ $@ ((_^$ $@R _) $@ cat_assoc _ _ _ $@R _)). - 2: nrapply cat_binprod_beta_pr2. - refine (_^$ $@ (_^$ $@R _)). - 2: nrapply cat_pr1_fmap01_binprod. + nrapply cat_binprod_pr1_pr2_twist. + - nrefine (cat_pr2_fmap01_binprod _ _ $@ _ $@ cat_assoc _ _ _). + refine (_ $@ (((cat_binprod_beta_pr2 _ _)^$ $@R _) $@ cat_assoc _ _ _ $@R _)). + refine ((_ $@R _) $@ _)^$. + 1: nrapply cat_pr1_fmap01_binprod. nrapply cat_binprod_beta_pr1. Defined. @@ -769,149 +870,151 @@ Section Associativity. snrapply pentagon_twist. intros a b c d. repeat nrefine (cat_assoc _ _ _ $@ _). - repeat refine (_ $@ (cat_assoc _ _ _)^$). + repeat refine (_ $@ cat_assoc_opp _ _ _). nrapply cat_binprod_eta_pr. - - refine ((cat_assoc _ _ _)^$ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1,3: nrapply cat_pr1_fmap01_binprod. - refine ((cat_assoc _ _ _)^$ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1: nrapply cat_binprod_beta_pr1. + { nrefine (cat_assoc_opp _ _ _ $@ _). + nrefine ((cat_pr1_fmap01_binprod _ _ $@R _) $@ _). + nrefine (cat_assoc_opp _ _ _ $@ _). + nrefine ((cat_binprod_pr1_twist _ _ _ $@R _) $@ _). + nrefine (cat_assoc_opp _ _ _ $@ _). + nrefine (((cat_assoc _ _ _ $@ (_ $@L _)) $@R _) $@ _). + 1: nrapply cat_binprod_beta_pr2. + nrefine (cat_assoc _ _ _ $@ (_ $@L (cat_assoc_opp _ _ _ $@ (_ $@R _)) $@ _) $@ _). + 1: nrapply cat_binprod_pr1_twist. + { nrefine (cat_assoc_opp _ _ _ $@ _). + nrefine ((cat_assoc_opp _ _ _ $@R _) $@ _). + nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + 1: nrapply cat_pr2_fmap01_binprod. + nrefine (cat_assoc_opp _ _ _ $@ ((cat_assoc _ _ _ $@ (_ $@L _)) $@R _)). + nrapply cat_binprod_beta_pr1. } + refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). 2: nrapply cat_pr1_fmap01_binprod. - refine ((cat_assoc _ _ _)^$ $@ _ $@ cat_assoc _ _ _). - refine (_ $@ _). - { do 2 refine (cat_assoc _ _ _ $@ _). - refine (_ $@L _). - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr2. - refine ((cat_assoc _ _ _)^$ $@ _). - refine ((_ $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr1. - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: nrapply cat_pr1_fmap01_binprod. + refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: nrapply cat_binprod_pr1_twist. + refine ((_^$ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _^$) $@ cat_assoc_opp _ _ _). + 2: nrapply cat_pr2_fmap01_binprod. + nrapply cat_binprod_pr1_twist. } + refine (cat_assoc_opp _ _ _ $@ (_ $@R _) $@ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 1,3: nrapply cat_pr2_fmap01_binprod. + refine (_ $@ (_ $@L ((_ $@R _)^$ $@ cat_assoc _ _ _)) $@ cat_assoc_opp _ _ _). + 2: nrapply cat_pr2_fmap01_binprod. + nrapply cat_binprod_eta_pr. + { nrefine (cat_assoc_opp _ _ _ $@ _). + nrefine (((cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr1 _ _ $@R _)) $@R _) $@ _). + nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_pr2_pr2_twist _ _ _ $@R _) $@ _). + nrefine (cat_assoc_opp _ _ _ $@ ((cat_assoc _ _ _ $@ (_ $@L _)) $@R _) $@ _). + 1: nrapply cat_binprod_beta_pr2. + nrefine (cat_assoc _ _ _ $@ (_ $@L cat_assoc_opp _ _ _ $@ (_ $@L ((_ $@R _) $@ _))) $@ _). + 1: nrapply cat_binprod_pr1_twist. + { nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). 1: nrapply cat_pr2_fmap01_binprod. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). + nrefine (cat_assoc_opp _ _ _ $@ (_ $@R _)). nrapply cat_binprod_beta_pr1. } - refine (_ $@ (_^$ $@R _)). - 2: nrapply cat_binprod_beta_pr1. - refine (_ $@ (cat_assoc _ _ _)^$). - refine (_ $@ (_ $@L _^$)). + nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). + refine (_ $@ cat_assoc_opp _ _ _ $@ ((cat_binprod_pr1_twist _ _ _)^$ $@R _)). + refine (_ $@ (_ $@L ((_^$ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L cat_assoc_opp _ _ _)))). 2: nrapply cat_pr2_fmap01_binprod. - refine (_ $@ cat_assoc _ _ _). - refine (_ $@ (_^$ $@R _)). + nrefine (_ $@ (_ $@L cat_assoc _ _ _)). + nrefine (_ $@ (cat_assoc _ _ _ $@R _) $@ cat_assoc _ _ _). + refine (_ $@ (((_^$ $@R _) $@ cat_assoc _ _ _ $@R _) $@R _)). 2: nrapply cat_binprod_beta_pr1. - symmetry; nrapply cat_assoc. - - refine ((cat_assoc _ _ _)^$ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1,3: nrapply cat_pr2_fmap01_binprod. - refine (cat_assoc _ _ _ $@ _ $@ (cat_assoc _ _ _)^$). - refine ((_ $@L _) $@ _). - { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - nrapply cat_binprod_beta_pr2. } - refine (_ $@ (_ $@L _)). - 2: { refine ((_^$ $@R _) $@ cat_assoc _ _ _). + nrefine (cat_assoc _ _ _ $@ _ $@ cat_assoc_opp _ _ _ $@ cat_assoc_opp _ _ _). + refine ((_ $@L _) $@ _)^$. + { do 2 nrefine (cat_assoc_opp _ _ _ $@ _). + nrefine ((cat_binprod_pr2_pr2_twist _ _ _ $@R _) $@ _). + nrefine (cat_assoc _ _ _ $@ (_ $@L _)). nrapply cat_pr2_fmap01_binprod. } + nrefine (cat_assoc_opp _ _ _ $@ cat_assoc_opp _ _ _ $@ (_ $@R _) $@ cat_assoc _ _ _). + nrapply cat_binprod_pr2_pr2_twist. } + nrefine (cat_assoc_opp _ _ _ $@ ((cat_assoc_opp _ _ _ $@ (_ $@R _)) $@R _) $@ _). + 1: nrapply cat_binprod_beta_pr2. + nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_pr1_pr2_twist _ _ _ $@R _) $@ _). + nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr1 _ _ $@R _) $@ _). + (* nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). *) + nrapply cat_binprod_eta_pr. + { do 2 nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). + refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). + 1,3: nrapply cat_binprod_pr1_pr2_twist. + refine (cat_pr1_fmap01_binprod _ _ $@ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: nrefine (cat_assoc_opp _ _ _ $@ (cat_pr1_fmap01_binprod _ _ $@R _)). + refine (_^$ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: nrapply cat_binprod_pr1_pr2_twist. + nrapply cat_pr1_fmap01_binprod. } + do 2 nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). + refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). + 1,3: nrapply cat_binprod_pr2_pr2_twist. + nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + 1: nrapply cat_pr2_fmap01_binprod. + nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr2 _ _ $@R _) $@ _). + nrefine (_ $@ cat_assoc_opp _ _ _). + refine (_ $@ (_ $@L ((((_^$ $@R _) $@ cat_assoc _ _ _) $@R _) $@ cat_assoc _ _ _))). + 2: nrapply cat_pr2_fmap01_binprod. + refine (_ $@ ((cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr2 _ _ $@R _))^$ $@R _) + $@ cat_assoc _ _ _ $@ (_ $@L cat_assoc_opp _ _ _)). + nrefine (_ $@ (_ $@L (_ $@ cat_assoc _ _ _)) $@ cat_assoc_opp _ _ _). + 2: refine ((_ $@L _^$) $@ cat_assoc_opp _ _ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 3: nrapply cat_binprod_pr2_pr2_twist. + 2: nrapply cat_pr2_fmap01_binprod. + do 2 nrefine (_ $@ cat_assoc _ _ _). + refine (_^$ $@R _). + nrapply cat_binprod_pr1_pr2_twist. + Defined. + + Definition pnetagon2 : PentagonIdentity (fun x y => cat_binprod x y). + Proof. + intros a b c d. + nrapply cat_binprod_eta_pr_xx_x. + - nrefine (cat_assoc_opp _ _ _ $@ (_ $@R _) $@ _). + 1: nrapply cat_pr1_pr1_associator_binprod. + refine (_ $@ (_ $@L ((((_^$ $@R _) $@ cat_assoc _ _ _) $@R _) + $@ cat_assoc _ _ _)) $@ cat_assoc_opp _ _ _). + 2: nrapply cat_pr1_fmap10_binprod. + do 2 nrefine (_ $@ (_ $@L cat_assoc_opp _ _ _)). nrapply cat_binprod_eta_pr. - + refine ((cat_assoc _ _ _)^$ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1,3: nrapply cat_binprod_beta_pr1. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _ $@ (cat_assoc _ _ _)^$). - 1: nrapply cat_pr2_fmap01_binprod. - refine (cat_assoc _ _ _ $@ _). - refine ((_ $@L ((cat_assoc _ _ _)^$ $@ (_ $@R _))) $@ _). - 1: nrapply cat_binprod_beta_pr2. - refine (_ $@ (_ $@L _)). - 2: { do 2 refine ((_ $@R _) $@ cat_assoc _ _ _). - symmetry. - nrapply cat_binprod_beta_pr2. } - refine (_ $@ cat_assoc _ _ _). - refine (_ $@ (_ $@R _)). - 2: { do 2 refine ((_ $@R _) $@ (cat_assoc _ _ _)). - symmetry. - nrapply cat_binprod_beta_pr1. } - do 2 refine (_ $@ (cat_assoc _ _ _)^$). - refine ((_ $@L _) $@ _). - { refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - nrapply cat_binprod_beta_pr1. } - refine (_ $@ (_ $@L (_ $@L _))). - 2: { refine ((_^$ $@R _) $@ cat_assoc _ _ _). - nrapply cat_binprod_beta_pr2. } - refine (_ $@ (_ $@L _)). - 2: { refine ((_^$ $@R _) $@ cat_assoc _ _ _). - nrapply cat_binprod_beta_pr2. } - refine (_ $@ (_ $@L _)). - 2: { refine ((_ $@L _^$) $@ (cat_assoc _ _ _)^$). - nrapply cat_binprod_beta_pr2. } - refine (_ $@ (_ $@L _)). - 2: { refine ((_^$ $@R _) $@ cat_assoc _ _ _). - nrapply cat_binprod_beta_pr2. } - refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: nrapply cat_binprod_beta_pr2. - refine ((_ $@L _) $@ _). - { refine (cat_assoc _ _ _ $@ (_ $@L _)). - nrapply cat_binprod_beta_pr2. } - refine ((_ $@L _) $@ _). - 1: refine ((cat_assoc _ _ _)^$ $@ (_ $@R _)). - 1: nrapply cat_binprod_beta_pr1. - symmetry; nrapply cat_assoc. - + refine ((cat_assoc _ _ _)^$ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1,3: nrapply cat_binprod_beta_pr2. - refine ((cat_assoc _ _ _)^$ $@ ((_ $@ _) $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr1. - 1: nrapply cat_idl. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr1. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr2. - refine (_ $@ (_ $@L _)). - 2: { refine ((_^$ $@R _) $@ cat_assoc _ _ _). - refine (cat_assoc _ _ _ $@ (_ $@L _)). - nrapply cat_binprod_beta_pr2. } - refine (_ $@ cat_assoc _ _ _). - refine (_ $@ (_ $@R _)). - 2: refine (cat_assoc _ _ _). - repeat refine (_ $@ (cat_assoc _ _ _)^$). - nrapply cat_binprod_eta_pr. - * refine ((cat_assoc _ _ _)^$ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1,3: nrapply cat_pr1_fmap01_binprod. - nrefine (cat_pr1_fmap01_binprod _ _ $@ _ $@ cat_assoc _ _ _). - refine (_ $@ (_^$ $@R _)). - 2: nrapply cat_pr1_fmap01_binprod. - refine (_^$ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: nrapply cat_pr1_fmap01_binprod. - nrapply cat_pr1_fmap01_binprod. - * refine ((cat_assoc _ _ _)^$ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1,3: nrapply cat_pr2_fmap01_binprod. - refine (cat_assoc _ _ _ $@ _ $@ (cat_assoc _ _ _)^$). - refine ((_ $@L _) $@ _). - 1: nrapply cat_pr2_fmap01_binprod. - refine ((cat_assoc _ _ _)^$ $@ (_ $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr2. - refine (_ $@ (_ $@L _)). - 2: { refine ((_^$ $@R _) $@ cat_assoc _ _ _). - nrapply cat_pr2_fmap01_binprod. } - refine (_ $@ cat_assoc _ _ _). - refine (_ $@ ((_^$ $@ _) $@R _)). - 3: nrapply cat_assoc. - 2: refine (_ $@R _). - 2: nrapply cat_binprod_beta_pr2. - refine (_ $@ cat_assoc _ _ _). - refine (_ $@ ((_^$ $@ _^$) $@R _)). - 3: nrapply cat_assoc; exact _. - 2: refine (_ $@L _). - 2: nrapply cat_pr2_fmap01_binprod; exact _. - refine (_ $@ (_ $@L ((_ $@L _^$) $@ (cat_assoc _ _ _)^$)) - $@ (cat_assoc _ _ _)^$). - 2: nrapply cat_pr2_fmap01_binprod. - refine (_ $@ (_ $@L (_ $@ _))). - 3: nrapply cat_assoc. - 2: refine (_^$ $@R _). - 2: nrapply cat_binprod_beta_pr2. - refine ((_^$ $@R _) $@ cat_assoc _ _ _). - nrapply cat_pr1_fmap01_binprod. + + nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). + refine (_ $@ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 1,3: nrapply cat_pr1_pr1_associator_binprod. + do 2 nrefine (_ $@ cat_assoc _ _ _). + refine (_^$ $@ (_^$ $@R _)). + 2: nrapply cat_pr1_pr1_associator_binprod. + nrapply cat_pr1_fmap01_binprod. + + nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). + refine (_ $@ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 1,3: nrapply cat_pr2_pr1_associator_binprod. + do 2 nrefine (_ $@ cat_assoc _ _ _). + refine (_ $@ ((cat_assoc _ _ _ $@ (_ $@L (_^$ $@ cat_assoc _ _ _)) + $@ cat_assoc_opp _ _ _ $@ cat_assoc_opp _ _ _) $@R _)). + 2: nrapply cat_pr2_pr1_associator_binprod. + refine (_^$ $@ (_ $@L _^$) $@ cat_assoc_opp _ _ _). + 2: nrapply cat_pr2_fmap01_binprod. + nrefine (cat_assoc_opp _ _ _ $@ (_ $@R _)). + nrapply cat_pr1_pr1_associator_binprod. + - nrefine (cat_assoc_opp _ _ _ $@ (_ $@R _) $@ _). + 1: nrapply cat_pr2_pr1_associator_binprod. + nrefine (cat_assoc _ _ _ $@ _ $@ cat_assoc_opp _ _ _). + nrefine ((_ $@L cat_pr2_associator_binprod _ _ _) $@ _). + refine (_ $@ (_ $@L ((((_^$ $@R _) $@ cat_assoc _ _ _) $@R _) $@ cat_assoc _ _ _))). + 2: nrapply cat_pr1_fmap10_binprod. + nrefine (_ $@ (_ $@L (cat_assoc_opp _ _ _ $@ cat_assoc_opp _ _ _))). + refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: nrapply cat_pr2_associator_binprod. + refine (_ $@ (_ $@L ((_^$ $@R _) $@ cat_assoc _ _ _ $@ cat_assoc _ _ _)) $@ cat_assoc_opp _ _ _). + 2: nrapply cat_pr2_pr1_associator_binprod. + refine (_ $@ (_ $@L ((_ $@L _^$) $@ cat_assoc_opp _ _ _))). + 2: nrapply cat_pr2_fmap01_binprod. + refine (cat_assoc_opp _ _ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _ $@ cat_assoc _ _ _). + nrapply cat_pr2_pr1_associator_binprod. + - nrefine (cat_assoc_opp _ _ _ $@ (cat_pr2_associator_binprod _ _ _ $@R _) $@ _). + nrefine (cat_assoc _ _ _ $@ (_ $@L (cat_pr2_associator_binprod _ _ _)) $@ _). + refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L (cat_assoc_opp _ _ _))). + 2: nrapply cat_pr2_fmap10_binprod. + refine (_ $@ cat_assoc_opp _ _ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _). + 2: nrapply cat_pr2_associator_binprod. + refine (cat_assoc_opp _ _ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _ + $@ (_ $@L (cat_pr2_fmap01_binprod _ _)^$)). + nrapply cat_pr2_associator_binprod. Defined. Local Instance hexagon_identity From 628ac55a66b07c0b365303a83c6ce157b7cd4eb8 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 1 May 2024 05:06:45 +0100 Subject: [PATCH 25/31] short proof of hexagon Signed-off-by: Ali Caglayan --- theories/WildCat/Products.v | 192 ++++++++---------------------------- 1 file changed, 42 insertions(+), 150 deletions(-) diff --git a/theories/WildCat/Products.v b/theories/WildCat/Products.v index 77e3c5ac6e3..2fde87999c9 100644 --- a/theories/WildCat/Products.v +++ b/theories/WildCat/Products.v @@ -376,59 +376,6 @@ Proof. - exact (r $@ (cat_idr _)^$). Defined. -Definition cat_binprod_eta_pr_x_xx_assoc {A : Type} `{HasBinaryProducts A} {u v w x y z : A} - (f : u $-> v) (g : u $-> w) (h : v $-> cat_binprod x (cat_binprod y z)) - (i : w $-> cat_binprod x (cat_binprod y z)) - : cat_pr1 $o h $o f $== cat_pr1 $o i $o g - -> cat_pr1 $o cat_pr2 $o h $o f $== cat_pr1 $o cat_pr2 $o i $o g - -> cat_pr2 $o cat_pr2 $o h $o f $== cat_pr2 $o cat_pr2 $o i $o g - -> h $o f $== i $o g. -Proof. - intros p q r. - snrapply cat_binprod_eta_pr_x_xx. - - exact (cat_assoc_opp _ _ _ $@ p $@ cat_assoc _ _ _). - - exact (cat_assoc_opp _ _ _ $@ q $@ cat_assoc _ _ _). - - exact (cat_assoc_opp _ _ _ $@ r $@ cat_assoc _ _ _). -Defined. - -Definition cat_binprod_eta_pr_x_x_xx {A : Type} `{HasBinaryProducts A} {v w x y z : A} - (f g : v $-> cat_binprod w (cat_binprod x (cat_binprod y z))) - : cat_pr1 $o f $== cat_pr1 $o g - -> cat_pr1 $o cat_pr2 $o f $== cat_pr1 $o cat_pr2 $o g - -> cat_pr1 $o cat_pr2 $o cat_pr2 $o f $== cat_pr1 $o cat_pr2 $o cat_pr2 $o g - -> cat_pr2 $o cat_pr2 $o cat_pr2 $o f $== cat_pr2 $o cat_pr2 $o cat_pr2 $o g - -> f $== g. -Proof. - intros p q r s. - snrapply cat_binprod_eta_pr_x_xx. - - exact p. - - exact q. - - snrapply cat_binprod_eta_pr. - 1,2: nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). - 1,2: nrefine ((cat_assoc_opp _ _ _ $@R _) $@ _ $@ (cat_assoc _ _ _ $@R _)). - + exact r. - + exact s. -Defined. - -Definition cat_binprod_eta_pr_xx_x_x {A : Type} `{HasBinaryProducts A} {v w x y z : A} - (f g : v $-> cat_binprod (cat_binprod (cat_binprod w x) y) z) - : cat_pr1 $o cat_pr1 $o cat_pr1 $o f $== cat_pr1 $o cat_pr1 $o cat_pr1 $o g - -> cat_pr2 $o cat_pr1 $o cat_pr1 $o f $== cat_pr2 $o cat_pr1 $o cat_pr1 $o g - -> cat_pr2 $o cat_pr1 $o f $== cat_pr2 $o cat_pr1 $o g - -> cat_pr2 $o f $== cat_pr2 $o g - -> f $== g. -Proof. - intros p q r s. - snrapply cat_binprod_eta_pr_xx_x. - 3: exact s. - 2: exact r. - nrefine (cat_assoc _ _ _ $@ _ $@ cat_assoc_opp _ _ _). - snrapply cat_binprod_eta_pr. - 1,2: do 2 nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). - - exact p. - - exact q. -Defined. - (** From binary products, all Bool-shaped products can be constructed. This should not be an instance to avoid a cycle with [hasbinaryproducts_hasproductsbool]. *) Definition hasproductsbool_hasbinaryproducts {A : Type} `{HasBinaryProducts A} : HasProducts Bool A. @@ -866,103 +813,6 @@ Section Associativity. Local Instance pentagon_binprod : PentagonIdentity (fun x y => cat_binprod x y). - Proof. - snrapply pentagon_twist. - intros a b c d. - repeat nrefine (cat_assoc _ _ _ $@ _). - repeat refine (_ $@ cat_assoc_opp _ _ _). - nrapply cat_binprod_eta_pr. - { nrefine (cat_assoc_opp _ _ _ $@ _). - nrefine ((cat_pr1_fmap01_binprod _ _ $@R _) $@ _). - nrefine (cat_assoc_opp _ _ _ $@ _). - nrefine ((cat_binprod_pr1_twist _ _ _ $@R _) $@ _). - nrefine (cat_assoc_opp _ _ _ $@ _). - nrefine (((cat_assoc _ _ _ $@ (_ $@L _)) $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr2. - nrefine (cat_assoc _ _ _ $@ (_ $@L (cat_assoc_opp _ _ _ $@ (_ $@R _)) $@ _) $@ _). - 1: nrapply cat_binprod_pr1_twist. - { nrefine (cat_assoc_opp _ _ _ $@ _). - nrefine ((cat_assoc_opp _ _ _ $@R _) $@ _). - nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - 1: nrapply cat_pr2_fmap01_binprod. - nrefine (cat_assoc_opp _ _ _ $@ ((cat_assoc _ _ _ $@ (_ $@L _)) $@R _)). - nrapply cat_binprod_beta_pr1. } - refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: nrapply cat_pr1_fmap01_binprod. - refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: nrapply cat_pr1_fmap01_binprod. - refine (_ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: nrapply cat_binprod_pr1_twist. - refine ((_^$ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _^$) $@ cat_assoc_opp _ _ _). - 2: nrapply cat_pr2_fmap01_binprod. - nrapply cat_binprod_pr1_twist. } - refine (cat_assoc_opp _ _ _ $@ (_ $@R _) $@ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 1,3: nrapply cat_pr2_fmap01_binprod. - refine (_ $@ (_ $@L ((_ $@R _)^$ $@ cat_assoc _ _ _)) $@ cat_assoc_opp _ _ _). - 2: nrapply cat_pr2_fmap01_binprod. - nrapply cat_binprod_eta_pr. - { nrefine (cat_assoc_opp _ _ _ $@ _). - nrefine (((cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr1 _ _ $@R _)) $@R _) $@ _). - nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_pr2_pr2_twist _ _ _ $@R _) $@ _). - nrefine (cat_assoc_opp _ _ _ $@ ((cat_assoc _ _ _ $@ (_ $@L _)) $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr2. - nrefine (cat_assoc _ _ _ $@ (_ $@L cat_assoc_opp _ _ _ $@ (_ $@L ((_ $@R _) $@ _))) $@ _). - 1: nrapply cat_binprod_pr1_twist. - { nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - 1: nrapply cat_pr2_fmap01_binprod. - nrefine (cat_assoc_opp _ _ _ $@ (_ $@R _)). - nrapply cat_binprod_beta_pr1. } - nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). - refine (_ $@ cat_assoc_opp _ _ _ $@ ((cat_binprod_pr1_twist _ _ _)^$ $@R _)). - refine (_ $@ (_ $@L ((_^$ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L cat_assoc_opp _ _ _)))). - 2: nrapply cat_pr2_fmap01_binprod. - nrefine (_ $@ (_ $@L cat_assoc _ _ _)). - nrefine (_ $@ (cat_assoc _ _ _ $@R _) $@ cat_assoc _ _ _). - refine (_ $@ (((_^$ $@R _) $@ cat_assoc _ _ _ $@R _) $@R _)). - 2: nrapply cat_binprod_beta_pr1. - nrefine (cat_assoc _ _ _ $@ _ $@ cat_assoc_opp _ _ _ $@ cat_assoc_opp _ _ _). - refine ((_ $@L _) $@ _)^$. - { do 2 nrefine (cat_assoc_opp _ _ _ $@ _). - nrefine ((cat_binprod_pr2_pr2_twist _ _ _ $@R _) $@ _). - nrefine (cat_assoc _ _ _ $@ (_ $@L _)). - nrapply cat_pr2_fmap01_binprod. } - nrefine (cat_assoc_opp _ _ _ $@ cat_assoc_opp _ _ _ $@ (_ $@R _) $@ cat_assoc _ _ _). - nrapply cat_binprod_pr2_pr2_twist. } - nrefine (cat_assoc_opp _ _ _ $@ ((cat_assoc_opp _ _ _ $@ (_ $@R _)) $@R _) $@ _). - 1: nrapply cat_binprod_beta_pr2. - nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_pr1_pr2_twist _ _ _ $@R _) $@ _). - nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr1 _ _ $@R _) $@ _). - (* nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). *) - nrapply cat_binprod_eta_pr. - { do 2 nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1,3: nrapply cat_binprod_pr1_pr2_twist. - refine (cat_pr1_fmap01_binprod _ _ $@ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: nrefine (cat_assoc_opp _ _ _ $@ (cat_pr1_fmap01_binprod _ _ $@R _)). - refine (_^$ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 2: nrapply cat_binprod_pr1_pr2_twist. - nrapply cat_pr1_fmap01_binprod. } - do 2 nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1,3: nrapply cat_binprod_pr2_pr2_twist. - nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). - 1: nrapply cat_pr2_fmap01_binprod. - nrefine (cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr2 _ _ $@R _) $@ _). - nrefine (_ $@ cat_assoc_opp _ _ _). - refine (_ $@ (_ $@L ((((_^$ $@R _) $@ cat_assoc _ _ _) $@R _) $@ cat_assoc _ _ _))). - 2: nrapply cat_pr2_fmap01_binprod. - refine (_ $@ ((cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr2 _ _ $@R _))^$ $@R _) - $@ cat_assoc _ _ _ $@ (_ $@L cat_assoc_opp _ _ _)). - nrefine (_ $@ (_ $@L (_ $@ cat_assoc _ _ _)) $@ cat_assoc_opp _ _ _). - 2: refine ((_ $@L _^$) $@ cat_assoc_opp _ _ _ $@ (_^$ $@R _) $@ cat_assoc _ _ _). - 3: nrapply cat_binprod_pr2_pr2_twist. - 2: nrapply cat_pr2_fmap01_binprod. - do 2 nrefine (_ $@ cat_assoc _ _ _). - refine (_^$ $@R _). - nrapply cat_binprod_pr1_pr2_twist. - Defined. - - Definition pnetagon2 : PentagonIdentity (fun x y => cat_binprod x y). Proof. intros a b c d. nrapply cat_binprod_eta_pr_xx_x. @@ -1078,6 +928,48 @@ Section Associativity. 1,2: nrapply cat_binprod_beta_pr2. nrapply cat_pr1_fmap01_binprod. Defined. + + Definition hexagon2 : HexagonIdentity (fun x y => cat_binprod x y). + Proof. + intros a b c. + nrefine (cat_assoc _ _ _ $@ _ $@ cat_assoc_opp _ _ _). + nrapply cat_binprod_eta_pr. + { nrefine (cat_assoc_opp _ _ _ $@ (cat_pr1_fmap10_binprod _ _ $@R _) $@ _). + nrefine (cat_assoc _ _ _ $@ _). + nrapply cat_binprod_eta_pr. + { nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _ $@ cat_assoc _ _ _). + refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). + 1: nrapply cat_binprod_beta_pr1. + 2: nrapply cat_pr1_pr1_associator_binprod. + nrefine (cat_assoc_opp _ _ _ $@ cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). + refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). + 1: nrapply cat_pr2_pr1_associator_binprod. + 2: nrapply cat_binprod_beta_pr1. + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ cat_assoc_opp _ _ _ $@ (_ $@R _) $@ _^$). + 1: nrapply cat_pr2_fmap01_binprod. + 2: nrapply cat_pr2_associator_binprod. + nrapply cat_binprod_beta_pr1. } + nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _ $@ cat_assoc _ _ _). + refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). + 1: nrapply cat_binprod_beta_pr2. + 2: nrapply cat_pr2_pr1_associator_binprod. + nrefine (cat_assoc_opp _ _ _ $@ cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). + refine ((_ $@R _) $@ _ $@ (((_ $@L _^$) $@ cat_assoc_opp _ _ _) $@R _)). + 1: nrapply cat_pr1_pr1_associator_binprod. + 2: nrapply cat_binprod_beta_pr2. + refine (cat_pr1_fmap01_binprod _ _ $@ _^$). + nrapply cat_pr1_pr1_associator_binprod. } + nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _ $@ cat_assoc _ _ _). + refine ((_ $@R _) $@ _ $@ ((_^$ $@R _) $@R _)). + 1: nrapply cat_pr2_fmap10_binprod. + 2: nrapply cat_pr2_associator_binprod. + nrefine (cat_assoc_opp _ _ _ $@ (cat_pr2_associator_binprod _ _ _ $@R _) $@ _). + nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _ $@ (cat_assoc_opp _ _ _ $@R _)). + 1: nrapply cat_pr2_fmap01_binprod. + refine (cat_assoc_opp _ _ _ $@ (_ $@R _) $@ _^$ $@ ((_ $@L _^$) $@R _)). + 1,3: nrapply cat_binprod_beta_pr2. + nrapply cat_pr2_pr1_associator_binprod. + Defined. Global Instance ismonoidal_binprod : IsMonoidal A (fun x y => cat_binprod x y) unit From f6a4978547844cbbad31b9989324d16105aeb033 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 2 May 2024 23:54:36 +0100 Subject: [PATCH 26/31] remove duplicated hexagon proof Signed-off-by: Ali Caglayan --- theories/WildCat/Products.v | 42 ------------------------------------- 1 file changed, 42 deletions(-) diff --git a/theories/WildCat/Products.v b/theories/WildCat/Products.v index 7cf57db2cf4..0866e8a5305 100644 --- a/theories/WildCat/Products.v +++ b/theories/WildCat/Products.v @@ -928,48 +928,6 @@ Section Associativity. nrapply cat_pr2_pr1_associator_binprod. Defined. - Definition hexagon2 : HexagonIdentity (fun x y => cat_binprod x y). - Proof. - intros a b c. - nrefine (cat_assoc _ _ _ $@ _ $@ cat_assoc_opp _ _ _). - nrapply cat_binprod_eta_pr. - { nrefine (cat_assoc_opp _ _ _ $@ (cat_pr1_fmap10_binprod _ _ $@R _) $@ _). - nrefine (cat_assoc _ _ _ $@ _). - nrapply cat_binprod_eta_pr. - { nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1: nrapply cat_binprod_beta_pr1. - 2: nrapply cat_pr1_pr1_associator_binprod. - nrefine (cat_assoc_opp _ _ _ $@ cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1: nrapply cat_pr2_pr1_associator_binprod. - 2: nrapply cat_binprod_beta_pr1. - refine (cat_assoc _ _ _ $@ (_ $@L _) $@ cat_assoc_opp _ _ _ $@ (_ $@R _) $@ _^$). - 1: nrapply cat_pr2_fmap01_binprod. - 2: nrapply cat_pr2_associator_binprod. - nrapply cat_binprod_beta_pr1. } - nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (_^$ $@R _)). - 1: nrapply cat_binprod_beta_pr2. - 2: nrapply cat_pr2_pr1_associator_binprod. - nrefine (cat_assoc_opp _ _ _ $@ cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ (((_ $@L _^$) $@ cat_assoc_opp _ _ _) $@R _)). - 1: nrapply cat_pr1_pr1_associator_binprod. - 2: nrapply cat_binprod_beta_pr2. - refine (cat_pr1_fmap01_binprod _ _ $@ _^$). - nrapply cat_pr1_pr1_associator_binprod. } - nrefine (cat_assoc_opp _ _ _ $@ _ $@ cat_assoc _ _ _ $@ cat_assoc _ _ _). - refine ((_ $@R _) $@ _ $@ ((_^$ $@R _) $@R _)). - 1: nrapply cat_pr2_fmap10_binprod. - 2: nrapply cat_pr2_associator_binprod. - nrefine (cat_assoc_opp _ _ _ $@ (cat_pr2_associator_binprod _ _ _ $@R _) $@ _). - nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _ $@ (cat_assoc_opp _ _ _ $@R _)). - 1: nrapply cat_pr2_fmap01_binprod. - refine (cat_assoc_opp _ _ _ $@ (_ $@R _) $@ _^$ $@ ((_ $@L _^$) $@R _)). - 1,3: nrapply cat_binprod_beta_pr2. - nrapply cat_pr2_pr1_associator_binprod. - Defined. - Global Instance ismonoidal_binprod : IsMonoidal A (fun x y => cat_binprod x y) unit := {}. From 3eb10fd9724849ba100b3f2dd114f2804ff23789 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 2 May 2024 23:56:08 +0100 Subject: [PATCH 27/31] undo whitespace changes Signed-off-by: Ali Caglayan --- theories/WildCat/Products.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/theories/WildCat/Products.v b/theories/WildCat/Products.v index 0866e8a5305..c051c5128c1 100644 --- a/theories/WildCat/Products.v +++ b/theories/WildCat/Products.v @@ -664,7 +664,7 @@ Section Associativity. - exact (cat_pr1 $o cat_pr2). - exact (fmap01 (fun x y => cat_binprod x y) x cat_pr2). Defined. - + Definition cat_binprod_pr1_twist (x y z : A) : cat_pr1 $o cat_binprod_twist x y z $== cat_pr1 $o cat_pr2 := cat_binprod_beta_pr1 _ _. @@ -676,7 +676,7 @@ Section Associativity. nrefine ((_ $@L cat_binprod_beta_pr2 _ _) $@ _). nrapply cat_pr1_fmap01_binprod. Defined. - + Definition cat_binprod_pr2_pr2_twist (x y z : A) : cat_pr2 $o cat_pr2 $o cat_binprod_twist x y z $== cat_pr2 $o cat_pr2. Proof. @@ -927,7 +927,7 @@ Section Associativity. 1,3: nrapply cat_binprod_beta_pr2. nrapply cat_pr2_pr1_associator_binprod. Defined. - + Global Instance ismonoidal_binprod : IsMonoidal A (fun x y => cat_binprod x y) unit := {}. From edd3c8d5fdf26320227e8668f53162a30c0d4723 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 6 May 2024 16:27:26 +0100 Subject: [PATCH 28/31] monoids and comonoids Signed-off-by: Ali Caglayan --- theories/Algebra/Categorical/MonoidObject.v | 197 ++++++++++++++++++++ theories/WildCat/Adjoint.v | 2 +- theories/WildCat/Bifunctor.v | 93 ++++++++- theories/WildCat/Equiv.v | 7 + theories/WildCat/Monoidal.v | 129 ++++++++++--- theories/WildCat/NatTrans.v | 13 +- theories/WildCat/Products.v | 132 ++++++++++--- 7 files changed, 509 insertions(+), 64 deletions(-) create mode 100644 theories/Algebra/Categorical/MonoidObject.v diff --git a/theories/Algebra/Categorical/MonoidObject.v b/theories/Algebra/Categorical/MonoidObject.v new file mode 100644 index 00000000000..817c225980e --- /dev/null +++ b/theories/Algebra/Categorical/MonoidObject.v @@ -0,0 +1,197 @@ +Require Import Basics.Overture Basics.Tactics. +Require Import WildCat.Core WildCat.Equiv WildCat.Monoidal WildCat.Bifunctor + WildCat.NatTrans WildCat.Opposite WildCat.Products. +Require Import abstract_algebra. + +(** * Monoids and Comonoids *) + +(** Here we define a monoid internal to a monoidal category. Various algebraic theories such as groups and rings may also be internalized, however these specifically require a cartesian monoidal structure. The theory of monoids however has no such requirement and can therefore be developed in much greater generality. This can be used to define a range of objects such as R-algebras, H-spaces, Hopf algebras and more. *) + +(** * Monoid objects *) + +Section MonoidObject. + Context {A : Type} {tensor : A -> A -> A} {unit : A} + `{HasEquivs A, !Is0Bifunctor tensor, !Is1Bifunctor tensor} + `{!Associator tensor, !LeftUnitor tensor unit, !RightUnitor tensor unit}. + + (** An object [x] of [A] is a monoid object if it comes with the following data: *) + Class IsMonoidObject (x : A) := { + (** A multiplication map from the tensor product of [x] with itself to [x]. *) + mo_mult : tensor x x $-> x; + (** A unit of the multplication. *) + mo_unit : unit $-> x; + (** The multiplication map is associative. *) + mo_assoc : mo_mult $o fmap10 tensor mo_mult x $o associator x x x + $== mo_mult $o fmap01 tensor x mo_mult; + (** The multiplication map is left unital. *) + mo_left_unit : mo_mult $o fmap10 tensor mo_unit x $== left_unitor x; + (** The multiplication map is right unital. *) + mo_right_unit : mo_mult $o fmap01 tensor x mo_unit $== right_unitor x; + }. + + Context `{!Braiding tensor}. + + (** An object [x] of [A] is a commutative monoid object if: *) + Class IsCommutativeMonoidObject (x : A) := { + (** It is a monoid object. *) + cmo_mo :: IsMonoidObject x; + (** The multiplication map is commutative. *) + cmo_comm : mo_mult $o braid x x $== mo_mult; + }. + +End MonoidObject. + +Arguments IsMonoidObject {A} tensor unit {_ _ _ _ _ _ _ _ _ _} x. +Arguments IsCommutativeMonoidObject {A} tensor unit {_ _ _ _ _ _ _ _ _ _ _} x. + +Section ComonoidObject. + Context {A : Type} (tensor : A -> A -> A) (unit : A) + `{HasEquivs A, !Is0Bifunctor tensor, !Is1Bifunctor tensor} + `{!Associator tensor, !LeftUnitor tensor unit, !RightUnitor tensor unit}. + + (** A comonoid object is a monoid object in the opposite category. *) + Class IsComonoidObject (x : A) + := ismonoid_comonoid_op :: IsMonoidObject (A:=A^op) tensor unit x. + + (** We can build comonoid objects from the following data: *) + Definition Build_IsComonoidObject (x : A) + (** A comultplication map. *) + (co_comult : x $-> tensor x x) + (** A counit. *) + (co_counit : x $-> unit) + (** The comultiplication is coassociative. *) + (co_coassoc : associator x x x $o fmap01 tensor x co_comult $o co_comult + $== fmap10 tensor co_comult x $o co_comult) + (** The comultiplication is left counital. *) + (co_left_counit : left_unitor x $o fmap10 tensor co_counit x $o co_comult $== Id x) + (** The comultiplication is right counital. *) + (co_right_counit : right_unitor x $o fmap01 tensor x co_counit $o co_comult $== Id x) + : IsComonoidObject x. + Proof. + snrapply Build_IsMonoidObject. + - exact co_comult. + - exact co_counit. + - nrapply cate_moveR_eV. + symmetry. + nrefine (cat_assoc _ _ _ $@ _). + rapply co_coassoc. + - simpl; nrefine (_ $@ cat_idr _). + nrapply cate_moveL_Ve. + nrefine (cat_assoc_opp _ _ _ $@ _). + exact co_left_counit. + - simpl; nrefine (_ $@ cat_idr _). + nrapply cate_moveL_Ve. + nrefine (cat_assoc_opp _ _ _ $@ _). + exact co_right_counit. + Defined. + + (** Comultiplication *) + Definition co_comult {x : A} `{!IsComonoidObject x} : x $-> tensor x x + := mo_mult (A:=A^op) (tensor:=tensor) (unit:=unit) (x:=x). + + (** Counit *) + Definition co_counit {x : A} `{!IsComonoidObject x} : x $-> unit + := mo_unit (A:=A^op) (tensor:=tensor) (unit:=unit) (x:=x). + + Context `{!Braiding tensor}. + + (** A cocommutative comonoid objects is a commutative monoid object in the opposite category. *) + Class IsCocommuativeComonoidObject (x : A) + := iscommuatativemonoid_cocomutativemonoid_op + :: IsCommutativeMonoidObject (A:=A^op) tensor unit x. + + (** We can build cocommutative comonoid objects from the following data: *) + Definition Build_IsCocommuatativeComonoidObject (x : A) + (** A comonoid. *) + `{!IsComonoidObject x} + (** Together with a proof of cocommutativity. *) + (cco_cocomm : braid x x $o co_comult $== co_comult) + : IsCocommuativeComonoidObject x. + Proof. + snrapply Build_IsCommutativeMonoidObject. + - exact _. + - exact cco_cocomm. + Defined. + +End ComonoidObject. + +(** ** Monoid enrichment *) + +(** A hom [x $-> y] in a cartesian category where [y] is a monoid object has the structure of a monoid. Equivalently, a hom [x $-> y] in a cartesian category where [x] is a comonoid object has the structure of a comonoid. *) + +Section MonoidEnriched. + Context {A : Type} `{HasEquivs A} `{!HasBinaryProducts A} + (unit : A) `{!IsTerminal unit} {x y : A} + `{!HasMorExt A} `{forall x y, IsHSet (x $-> y)}. + + Section Monoid. + Context `{!IsMonoidObject _ _ y}. + + Local Instance sgop_hom : SgOp (x $-> y) + := fun f g => mo_mult $o cat_binprod_corec f g. + + Local Instance monunit_hom : MonUnit (x $-> y) := mo_unit $o mor_terminal _ _. + + Local Instance associative_hom : Associative sgop_hom. + Proof. + intros f g h. + unfold sgop_hom. + rapply path_hom. + refine ((_ $@L cat_binprod_fmap01_corec _ _ _)^$ $@ _). + nrefine (cat_assoc_opp _ _ _ $@ _). + refine ((mo_assoc $@R _)^$ $@ _). + nrefine (_ $@ (_ $@L cat_binprod_fmap10_corec _ _ _)). + refine (cat_assoc _ _ _ $@ (_ $@L _) $@ cat_assoc _ _ _). + nrapply cat_binprod_associator_corec. + Defined. + + Local Instance leftidentity_hom : LeftIdentity sgop_hom mon_unit. + Proof. + intros f. + unfold sgop_hom, mon_unit. + rapply path_hom. + refine ((_ $@L (cat_binprod_fmap10_corec _ _ _)^$) $@ cat_assoc_opp _ _ _ $@ _). + nrefine (((mo_left_unit $@ _) $@R _) $@ _). + 1: nrapply cate_buildequiv_fun. + unfold trans_nattrans. + nrefine ((((_ $@R _) $@ _) $@R _) $@ _). + 1: nrapply cate_buildequiv_fun. + 1: nrapply cat_binprod_beta_pr1. + nrapply cat_binprod_beta_pr2. + Defined. + + Local Instance rightidentity_hom : RightIdentity sgop_hom mon_unit. + Proof. + intros f. + unfold sgop_hom, mon_unit. + rapply path_hom. + refine ((_ $@L (cat_binprod_fmap01_corec _ _ _)^$) $@ cat_assoc_opp _ _ _ $@ _). + nrefine (((mo_right_unit $@ _) $@R _) $@ _). + 1: nrapply cate_buildequiv_fun. + nrapply cat_binprod_beta_pr1. + Defined. + + Local Instance issemigroup_hom : IsSemiGroup (x $-> y) := {}. + Local Instance ismonoid_hom : IsMonoid (x $-> y) := {}. + + End Monoid. + + Context `{!IsCommutativeMonoidObject _ _ y}. + Local Existing Instances sgop_hom monunit_hom ismonoid_hom. + + Local Instance commutative_hom : Commutative sgop_hom. + Proof. + intros f g. + unfold sgop_hom. + rapply path_hom. + refine ((_ $@L _^$) $@ cat_assoc_opp _ _ _ $@ (cmo_comm $@R _)). + nrapply cat_binprod_swap_corec. + Defined. + + Local Instance iscommutativemonoid_hom : IsCommutativeMonoid (x $-> y) := {}. + +End MonoidEnriched. + + + + diff --git a/theories/WildCat/Adjoint.v b/theories/WildCat/Adjoint.v index 68d865e06e6..ef4ef9965fc 100644 --- a/theories/WildCat/Adjoint.v +++ b/theories/WildCat/Adjoint.v @@ -388,7 +388,7 @@ Proof. snrapply Build_Adjunction_natequiv_nat_right. { intros y. refine (natequiv_compose (natequiv_adjunction_l adj _) _). - rapply (natequiv_postwhisker _ (natequiv_op _ _ e)). } + rapply (natequiv_postwhisker _ (natequiv_op e)). } intros x. rapply is1natural_comp. Defined. diff --git a/theories/WildCat/Bifunctor.v b/theories/WildCat/Bifunctor.v index 998fe9a2597..ba3501f9e49 100644 --- a/theories/WildCat/Bifunctor.v +++ b/theories/WildCat/Bifunctor.v @@ -1,6 +1,7 @@ Require Import Basics.Overture Basics.Tactics. Require Import Types.Forall Types.Prod. -Require Import WildCat.Core WildCat.Prod WildCat.Equiv WildCat.NatTrans WildCat.Square. +Require Import WildCat.Core WildCat.Prod WildCat.Equiv WildCat.NatTrans + WildCat.Square WildCat.Opposite. (** * Bifunctors between WildCats *) @@ -14,6 +15,11 @@ Class Is0Bifunctor {A B C : Type} is0functor10_bifunctor :: forall b, Is0Functor (flip F b); }. +Arguments Is0Bifunctor {A B C _ _ _} F. +Arguments is0functor_bifunctor_uncurried {A B C _ _ _} F {_}. +Arguments is0functor01_bifunctor {A B C _ _ _} F {_} a : rename. +Arguments is0functor10_bifunctor {A B C _ _ _} F {_} b : rename. + (** We provide two alternate constructors, allowing the user to provide just the first field or the last two fields. *) Definition Build_Is0Bifunctor' {A B C : Type} `{Is01Cat A, Is01Cat B, IsGraph C} (F : A -> B -> C) @@ -76,6 +82,9 @@ Class Is1Bifunctor {A B C : Type} Arguments Is1Bifunctor {A B C _ _ _ _ _ _ _ _ _ _ _ _} F {Is0Bifunctor} : rename. Arguments Build_Is1Bifunctor {A B C _ _ _ _ _ _ _ _ _ _ _ _} F {_} _ _ _ _ _. +Arguments is1functor_bifunctor_uncurried {A B C _ _ _ _ _ _ _ _ _ _ _ _} F {_ _}. +Arguments is1functor01_bifunctor {A B C _ _ _ _ _ _ _ _ _ _ _ _} F {_ _} a : rename. +Arguments is1functor10_bifunctor {A B C _ _ _ _ _ _ _ _ _ _ _ _} F {_ _} b : rename. Arguments fmap11_is_fmap01_fmap10 {A B C _ _ _ _ _ _ _ _ _ _ _ _} F {Is0Bifunctor Is1Bifunctor} {a0 a1} f {b0 b1} g : rename. Arguments fmap11_is_fmap10_fmap01 {A B C _ _ _ _ _ _ _ _ _ _ _ _} F @@ -298,14 +307,14 @@ Global Instance is0bifunctor_postcompose {A B C D : Type} `{IsGraph A, IsGraph B, IsGraph C, IsGraph D} (F : A -> B -> C) {bf : Is0Bifunctor F} (G : C -> D) `{!Is0Functor G} - : Is0Bifunctor (fun a b => G (F a b)) + : Is0Bifunctor (fun a b => G (F a b)) | 10 := {}. Global Instance is1bifunctor_postcompose {A B C D : Type} `{Is1Cat A, Is1Cat B, Is1Cat C, Is1Cat D} (F : A -> B -> C) (G : C -> D) `{!Is0Functor G, !Is1Functor G} `{!Is0Bifunctor F} {bf : Is1Bifunctor F} - : Is1Bifunctor (fun a b => G (F a b)). + : Is1Bifunctor (fun a b => G (F a b)) | 10. Proof. snrapply Build_Is1Bifunctor. 1-3: exact _. @@ -319,7 +328,7 @@ Global Instance is0bifunctor_precompose {A B C D E : Type} `{IsGraph A, IsGraph B, IsGraph C, IsGraph D, IsGraph E} (G : A -> B) (K : E -> C) (F : B -> C -> D) `{!Is0Functor G, !Is0Bifunctor F, !Is0Functor K} - : Is0Bifunctor (fun a b => F (G a) (K b)). + : Is0Bifunctor (fun a b => F (G a) (K b)) | 10. Proof. snrapply Build_Is0Bifunctor. - change (Is0Functor (uncurry F o functor_prod G K)). @@ -335,7 +344,7 @@ Global Instance is1bifunctor_precompose {A B C D E : Type} (G : A -> B) (K : E -> C) (F : B -> C -> D) `{!Is0Functor G, !Is1Functor G, !Is0Bifunctor F, !Is1Bifunctor F, !Is0Functor K, !Is1Functor K} - : Is1Bifunctor (fun a b => F (G a) (K b)). + : Is1Bifunctor (fun a b => F (G a) (K b)) | 10. Proof. snrapply Build_Is1Bifunctor. - change (Is1Functor (uncurry F o functor_prod G K)). @@ -385,10 +394,8 @@ Definition fmap11_square {A B C : Type} `{Is1Cat A, Is1Cat B, Is1Cat C} (** We can show that an uncurried natural transformation between uncurried bifunctors by composing the naturality square in each variable. *) Global Instance is1natural_uncurry {A B C : Type} `{Is1Cat A, Is1Cat B, Is1Cat C} - (F : A -> B -> C) - `{!Is0Bifunctor F, !Is1Bifunctor F} - (G : A -> B -> C) - `{!Is0Bifunctor G, !Is1Bifunctor G} + (F : A -> B -> C) `{!Is0Bifunctor F, !Is1Bifunctor F} + (G : A -> B -> C) `{!Is0Bifunctor G, !Is1Bifunctor G} (alpha : uncurry F $=> uncurry G) (nat_l : forall b, Is1Natural (flip F b) (flip G b) (fun x : A => alpha (x, b))) (nat_r : forall a, Is1Natural (F a) (G a) (fun y : B => alpha (a, y))) @@ -402,3 +409,71 @@ Proof. 2: rapply (fmap11_is_fmap01_fmap10 G). exact (hconcat (nat_l _ _ _ f) (nat_r _ _ _ f')). Defined. + +(** Flipping a natural transformation between bifunctors. *) +Definition nattrans_flip {A B C : Type} + `{Is1Cat A, Is1Cat B, Is1Cat C} + {F : A -> B -> C} `{!Is0Bifunctor F, !Is1Bifunctor F} + {G : B -> A -> C} `{!Is0Bifunctor G, !Is1Bifunctor G} + : NatTrans (uncurry F) (uncurry (flip G)) + -> NatTrans (uncurry (flip F)) (uncurry G). +Proof. + intros [alpha nat]. + snrapply Build_NatTrans. + - exact (alpha o equiv_prod_symm _ _). + - intros [b a] [b' a'] [g f]. + exact (nat (a, b) (a', b') (f, g)). +Defined. + +Definition nattrans_flip' {A B C : Type} + `{Is1Cat A, Is1Cat B, Is1Cat C} + {F : A -> B -> C} `{!Is0Bifunctor F, !Is1Bifunctor F} + {G : B -> A -> C} `{!Is0Bifunctor G, !Is1Bifunctor G} + : NatTrans (uncurry (flip F)) (uncurry G) + -> NatTrans (uncurry F) (uncurry (flip G)) + := nattrans_flip (F:=flip F) (G:=flip G). + +(** ** Opposite Bifunctors *) + +(** There are a few more combinations we can do for this, such as profunctors, but we will leave those for later. *) + +Global Instance is0bifunctor_op A B C (F : A -> B -> C) `{Is0Bifunctor A B C F} + : Is0Bifunctor (F : A^op -> B^op -> C^op). +Proof. + snrapply Build_Is0Bifunctor. + - exact (is0functor_op _ _ (uncurry F)). + - intros a. + nrapply is0functor_op. + exact (is0functor01_bifunctor F a). + - intros b. + nrapply is0functor_op. + exact (is0functor10_bifunctor F b). +Defined. + +Global Instance is1bifunctor_op A B C (F : A -> B -> C) `{Is1Bifunctor A B C F} + : Is1Bifunctor (F : A^op -> B^op -> C^op). +Proof. + snrapply Build_Is1Bifunctor. + - exact (is1functor_op _ _ (uncurry F)). + - intros a. + nrapply is1functor_op. + exact (is1functor01_bifunctor F a). + - intros b. + nrapply is1functor_op. + exact (is1functor10_bifunctor F b). + - intros a0 a1 f b0 b1 g; cbn in f, g. + exact (fmap11_is_fmap10_fmap01 F f g). + - intros a0 a1 f b0 b1 g; cbn in f, g. + exact (fmap11_is_fmap01_fmap10 F f g). +Defined. + +Global Instance is0bifunctor_op' A B C (F : A^op -> B^op -> C^op) + `{IsGraph A, IsGraph B, IsGraph C, Fop : !Is0Bifunctor (F : A^op -> B^op -> C^op)} + : Is0Bifunctor (F : A -> B -> C) + := is0bifunctor_op A^op B^op C^op F. + +Global Instance is1bifunctor_op' A B C (F : A^op -> B^op -> C^op) + `{Is1Cat A, Is1Cat B, Is1Cat C, + !Is0Bifunctor (F : A^op -> B^op -> C^op), !Is1Bifunctor (F : A^op -> B^op -> C^op)} + : Is1Bifunctor (F : A -> B -> C) + := is1bifunctor_op A^op B^op C^op F. diff --git a/theories/WildCat/Equiv.v b/theories/WildCat/Equiv.v index 116b0b0e3ba..aafc5684307 100644 --- a/theories/WildCat/Equiv.v +++ b/theories/WildCat/Equiv.v @@ -413,6 +413,13 @@ Proof. apply cate_inv_adjointify. Defined. +Definition cate_inv_compose' {A} `{HasEquivs A} {a b c : A} (e : a $<~> b) (f : b $<~> c) + : cate_fun (f $oE e)^-1$ $== e^-1$ $o f^-1$. +Proof. + nrefine (_ $@ cate_buildequiv_fun _). + nrapply cate_inv_compose. +Defined. + Definition cate_inv_V {A} `{HasEquivs A} {a b : A} (e : a $<~> b) : cate_fun (e^-1$)^-1$ $== cate_fun e. Proof. diff --git a/theories/WildCat/Monoidal.v b/theories/WildCat/Monoidal.v index 38abf2de889..9eea53cfdf4 100644 --- a/theories/WildCat/Monoidal.v +++ b/theories/WildCat/Monoidal.v @@ -1,6 +1,6 @@ Require Import Basics.Overture Basics.Tactics Types.Forall. Require Import WildCat.Core WildCat.Bifunctor WildCat.Prod WildCat.Equiv. -Require Import WildCat.NatTrans WildCat.Square. +Require Import WildCat.NatTrans WildCat.Square WildCat.Opposite. (** * Monoidal Categories *) @@ -12,21 +12,33 @@ Require Import WildCat.NatTrans WildCat.Square. (** *** Associators *) -(** A natural equivalence witnessing the associativity of a bifunctor. *) Class Associator {A : Type} `{HasEquivs A} - (F : A -> A -> A) `{!Is0Bifunctor F, !Is1Bifunctor F} := { - (** An isomorphism [associator] witnessing associativity of [F]. *) - associator a b c : F a (F b c) $<~> F (F a b) c; - - (** The [associator] is a natural isomorphism. *) - is1natural_associator_uncurried - :: Is1Natural - (fun '(a, b, c) => F a (F b c)) - (fun '(a, b, c) => F (F a b) c) - (fun '(a, b, c) => associator a b c); -}. + (F : A -> A -> A) `{!Is0Bifunctor F, !Is1Bifunctor F} + := associator_uncurried + : NatEquiv (fun '(a, b, c) => F a (F b c)) (fun '(a, b, c) => F (F a b) c). + +Arguments associator_uncurried {A _ _ _ _ _ F _ _ _}. + +Definition associator {A : Type} `{HasEquivs A} {F : A -> A -> A} + `{!Is0Bifunctor F, !Is1Bifunctor F, !Associator F} + : forall a b c, F a (F b c) $<~> F (F a b) c + := fun a b c => associator_uncurried (a, b, c). Coercion associator : Associator >-> Funclass. -Arguments associator {A _ _ _ _ _ F _ _ _} a b c. + +Definition Build_Associator {A : Type} `{HasEquivs A} (F : A -> A -> A) + `{!Is0Bifunctor F, !Is1Bifunctor F} + (associator : forall a b c, F a (F b c) $<~> F (F a b) c) + (isnat_assoc : Is1Natural + (fun '(a, b, c) => F a (F b c)) + (fun '(a, b, c) => F (F a b) c) + (fun '(a, b, c) => associator a b c)) + : Associator F. +Proof. + snrapply Build_NatEquiv. + - intros [[a b] c]. + exact (associator a b c). + - exact isnat_assoc. +Defined. (** *** Unitors *) @@ -69,18 +81,16 @@ Arguments pentagon_identity {A _ _ _ _ _} F {_ _ _}. (** *** Braiding *) Class Braiding {A : Type} `{Is1Cat A} - (F : A -> A -> A) `{!Is0Bifunctor F, !Is1Bifunctor F} := { - (** A morphism [braid] witnessing the symmetry of [F]. *) - braid a b : F a b $-> F b a; - (** The [braid] is a natural transformation. *) - is1natural_braiding_uncurried - : Is1Natural - (uncurry F) - (uncurry (flip F)) - (fun '(a, b) => braid a b); -}. + (F : A -> A -> A) `{!Is0Bifunctor F, !Is1Bifunctor F} + := braid_uncurried : NatTrans (uncurry F) (uncurry (flip F)). + +Arguments braid_uncurried {A _ _ _ _ F _ _ _}. + +Definition braid {A : Type} `{Is1Cat A} {F : A -> A -> A} + `{!Is0Bifunctor F, !Is1Bifunctor F, !Braiding F} + : forall a b, F a b $-> F b a + := fun a b => braid_uncurried (a, b). Coercion braid : Braiding >-> Funclass. -Arguments braid {A _ _ _ _ F _ _ _} a b. Class SymmetricBraiding {A : Type} `{Is1Cat A} (F : A -> A -> A) `{!Is0Bifunctor F, !Is1Bifunctor F} := { @@ -150,7 +160,8 @@ Class IsSymmetricMonoidal (A : Type) `{HasEquivs A} (** *** Theory about [Associator] *) Section Associator. - Context {A : Type} {F : A -> A -> A} `{assoc : Associator A F, !HasEquivs A}. + Context {A : Type} `{HasEquivs A} {F : A -> A -> A} + `{!Is0Bifunctor F, !Is1Bifunctor F, assoc : !Associator F}. Local Definition associator_nat {x x' y y' z z'} (f : x $-> x') (g : y $-> y') (h : z $-> z') @@ -191,8 +202,36 @@ Section Associator. - exact (fmap21 _ (fmap11_id _ _ _) _ $@ fmap01_is_fmap11 F _ _). Defined. + Global Instance associator_op : Associator (A:=A^op) F + := natequiv_inverse (natequiv_op assoc). + End Associator. +(** ** Theory about [LeftUnitor] and [RightUnitor] *) + +Section LeftUnitor. + Context {A : Type} `{HasEquivs A} {F : A -> A -> A} (unit : A) + `{!Is0Bifunctor F, !Is1Bifunctor F, !LeftUnitor F unit, !RightUnitor F unit}. + + Global Instance left_unitor_op : LeftUnitor (A:=A^op) F unit + := natequiv_inverse (natequiv_op left_unitor). + + Global Instance right_unitor_op : RightUnitor (A:=A^op) F unit + := natequiv_inverse (natequiv_op right_unitor). + +End LeftUnitor. + +(** ** Theory about [Braiding] *) + +Section Braiding. + Context {A : Type} `{HasEquivs A} {F : A -> A -> A} + `{!Is0Bifunctor F, !Is1Bifunctor F, braid : !Braiding F}. + + Global Instance braiding_op : Braiding (A:=A^op) F + := (nattrans_op (nattrans_flip braid)). + +End Braiding. + (** ** Theory about [SymmetricBraid] *) Section SymmetricBraid. @@ -358,14 +397,14 @@ Section SymmetricBraid. Local Definition braid_nat {a b c d} (f : a $-> c) (g : b $-> d) : braid c d $o fmap11 F f g $== fmap11 F g f $o braid a b. Proof. - exact (is1natural_braiding_uncurried (a, b) (c, d) (f, g)). + exact (isnat braid_uncurried (a := (a, b)) (a' := (c, d)) (f, g)). Defined. Local Definition braid_nat_l {a b c} (f : a $-> b) : braid b c $o fmap10 F f c $== fmap01 F c f $o braid a c. Proof. refine ((_ $@L (fmap10_is_fmap11 _ _ _)^$) $@ _ $@ (fmap01_is_fmap11 _ _ _ $@R _)). - exact (is1natural_braiding_uncurried (a, c) (b, c) (f, Id _)). + exact (isnat braid_uncurried (a := (a, c)) (a' := (b, c)) (f, Id _)). Defined. (** This is just the inverse of above. *) @@ -373,11 +412,43 @@ Section SymmetricBraid. : braid a c $o fmap01 F a g $== fmap10 F g a $o braid a b. Proof. refine ((_ $@L (fmap01_is_fmap11 _ _ _)^$) $@ _ $@ (fmap10_is_fmap11 _ _ _ $@R _)). - exact (is1natural_braiding_uncurried (a, b) (a, c) (Id _ , g)). + exact (isnat braid_uncurried (a := (a, b)) (a' := (a, c)) (Id _ , g)). + Defined. + + Global Instance symmetricbraiding_op : SymmetricBraiding (A:=A^op) F. + Proof. + snrapply Build_SymmetricBraiding. + - exact _. + - intros a b. + rapply braid_braid. Defined. End SymmetricBraid. +Global Instance ismonoidal_op {A : Type} (tensor : A -> A -> A) (unit : A) + `{IsMonoidal A tensor unit} + : IsMonoidal A^op tensor unit. +Proof. + snrapply Build_IsMonoidal. + 1-5: exact _. + - intros a b; unfold op in a, b; simpl. + refine (_^$ $@ _ $@ (_ $@L _)). + 1,3: exact (emap_inv _ _ $@ cate_buildequiv_fun _). + nrefine (cate_inv2 _ $@ cate_inv_compose' _ _). + refine (cate_buildequiv_fun _ $@ _ $@ ((cate_buildequiv_fun _)^$ $@R _) + $@ (cate_buildequiv_fun _)^$). + rapply cat_tensor_triangle_identity. + - intros a b c d; unfold op in a, b, c, d; simpl. + refine (_ $@ ((_ $@L _) $@@ _)). + 2,3: exact (emap_inv _ _ $@ cate_buildequiv_fun _). + refine ((cate_inv_compose' _ _)^$ $@ cate_inv2 _ $@ cate_inv_compose' _ _ + $@ (_ $@L cate_inv_compose' _ _)). + refine (cate_buildequiv_fun _ $@ _ $@ ((cate_buildequiv_fun _)^$ $@R _) + $@ (cate_buildequiv_fun _)^$). + refine (_ $@ (cate_buildequiv_fun _ $@@ (cate_buildequiv_fun _ $@R _))^$). + rapply cat_tensor_pentagon_identity. +Defined. + (** ** Building Symmetric Monoidal Categories *) (** The following construction is what we call the "twist construction". It is a way to build a symmetric monoidal category from simpler pieces than the axioms ask for. diff --git a/theories/WildCat/NatTrans.v b/theories/WildCat/NatTrans.v index 8ec964d6e81..07afbadea42 100644 --- a/theories/WildCat/NatTrans.v +++ b/theories/WildCat/NatTrans.v @@ -189,6 +189,17 @@ Proof. srapply isnat_tr. Defined. +Definition nattrans_op {A B : Type} `{Is01Cat A} `{Is1Cat B} + {F G : A -> B} `{!Is0Functor F, !Is0Functor G} + : NatTrans F G + -> NatTrans (A:=A^op) (B:=B^op) (G : A^op -> B^op) (F : A^op -> B^op). +Proof. + intros alpha. + snrapply Build_NatTrans. + - exact (transformation_op F G alpha). + - exact _. +Defined. + (** Natural equivalences *) Record NatEquiv {A B : Type} `{IsGraph A} `{HasEquivs B} @@ -291,7 +302,7 @@ Proof. Defined. Lemma natequiv_op {A B : Type} `{Is01Cat A} `{HasEquivs B} - (F G : A -> B) `{!Is0Functor F, !Is0Functor G} + {F G : A -> B} `{!Is0Functor F, !Is0Functor G} : NatEquiv F G -> NatEquiv (G : A^op -> B^op) F. Proof. intros [a n]. diff --git a/theories/WildCat/Products.v b/theories/WildCat/Products.v index 002ef175a61..c3010d0ab82 100644 --- a/theories/WildCat/Products.v +++ b/theories/WildCat/Products.v @@ -482,28 +482,28 @@ Global Instance is0functor_cat_binprod_l {A : Type} `{HasBinaryProducts A} (y : A) : Is0Functor (fun x => cat_binprod x y). Proof. - exact (is0functor10_bifunctor y). + exact (is0functor10_bifunctor _ y). Defined. Global Instance is1functor_cat_binprod_l {A : Type} `{HasBinaryProducts A} (y : A) : Is1Functor (fun x => cat_binprod x y). Proof. - exact (is1functor10_bifunctor y). + exact (is1functor10_bifunctor _ y). Defined. Global Instance is0functor_cat_binprod_r {A : Type} `{HasBinaryProducts A} (x : A) : Is0Functor (fun y => cat_binprod x y). Proof. - exact (is0functor01_bifunctor x). + exact (is0functor01_bifunctor _ x). Defined. Global Instance is1functor_cat_binprod_r {A : Type} `{HasBinaryProducts A} (x : A) : Is1Functor (fun y => cat_binprod x y). Proof. - exact (is1functor01_bifunctor x). + exact (is1functor01_bifunctor _ x). Defined. (** [cat_binprod_corec] is also functorial in each morphsism. *) @@ -556,6 +556,61 @@ Definition cat_pr2_fmap11_binprod {A : Type} `{HasBinaryProducts A} : cat_pr2 $o fmap11 (fun x y => cat_binprod x y) f g $== g $o cat_pr2 := cat_binprod_beta_pr2 _ _. +(** *** Diagonal *) + +(** Annoyingly this doesn't follow directly from the general diagonal since [fun b => if b then x else x] is not definitionally equal to [fun _ => x]. *) +Definition cat_binprod_diag {A : Type} + `{HasEquivs A} (x : A) `{!BinaryProduct x x} + : x $-> cat_binprod x x. +Proof. + snrapply cat_binprod_corec; exact (Id _). +Defined. + +Definition cat_binprod_fmap01_corec {A : Type} + `{Is1Cat A, !HasBinaryProducts A} {w x y z : A} + (f : w $-> z) (g : x $-> y) (h : w $-> x) + : fmap01 (fun x y => cat_binprod x y) z g $o cat_binprod_corec f h + $== cat_binprod_corec f (g $o h). +Proof. + snrapply cat_binprod_eta_pr. + - nrefine (cat_assoc_opp _ _ _ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ cat_idl _ $@ _ $@ _^$). + 1-3: rapply cat_binprod_beta_pr1. + - nrefine (cat_assoc_opp _ _ _ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _) $@ _^$). + 1-3: rapply cat_binprod_beta_pr2. +Defined. + +Definition cat_binprod_fmap10_corec {A : Type} + `{Is1Cat A, !HasBinaryProducts A} {w x y z : A} + (f : x $-> y) (g : w $-> x) (h : w $-> z) + : fmap10 (fun x y => cat_binprod x y) f z $o cat_binprod_corec g h + $== cat_binprod_corec (f $o g) h. +Proof. + snrapply cat_binprod_eta_pr. + - refine (cat_assoc_opp _ _ _ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _) $@ _^$). + 1-3: nrapply cat_binprod_beta_pr1. + - refine (cat_assoc_opp _ _ _ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ cat_idl _ $@ _ $@ _^$). + 1-3: nrapply cat_binprod_beta_pr2. +Defined. + +Definition cat_binprod_fmap11_corec {A : Type} + `{Is1Cat A, !HasBinaryProducts A} {v w x y z : A} + (f : w $-> y) (g : x $-> z) (h : v $-> w) (i : v $-> x) + : fmap11 (fun x y => cat_binprod x y) f g $o cat_binprod_corec h i + $== cat_binprod_corec (f $o h) (g $o i). +Proof. + snrapply cat_binprod_eta_pr. + - refine (cat_assoc_opp _ _ _ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _) $@ _^$). + 1-3: nrapply cat_binprod_beta_pr1. + - nrefine (cat_assoc_opp _ _ _ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _) $@ _^$). + 1-3: rapply cat_binprod_beta_pr2. +Defined. + (** *** Symmetry of binary products *) Section Symmetry. @@ -586,35 +641,30 @@ Section Symmetry. all: nrapply cat_binprod_swap_cat_binprod_swap. Defined. - Definition cat_binprod_swap_nat {a b c d : A} (f : a $-> c) (g : b $-> d) - : cat_binprod_swap c d $o fmap11 (fun x y : A => cat_binprod x y) f g - $== fmap11 (fun x y : A => cat_binprod x y) g f $o cat_binprod_swap a b. + Definition cat_binprod_swap_corec {a b c : A} (f : a $-> b) (g : a $-> c) + : cat_binprod_swap b c $o cat_binprod_corec f g $== cat_binprod_corec g f. Proof. nrapply cat_binprod_eta_pr. - - nrefine (cat_assoc_opp _ _ _ $@ _). - nrefine ((cat_binprod_beta_pr1 _ _ $@R _) $@ _). - nrefine (cat_pr2_fmap11_binprod _ _ $@ _). - nrefine (_ $@ cat_assoc _ _ _). - refine (_ $@ (_^$ $@R _)). - 2: nrapply cat_pr1_fmap11_binprod. - refine ((_ $@L _^$) $@ (cat_assoc _ _ _)^$). - nrapply cat_binprod_beta_pr1. - - refine ((cat_assoc _ _ _)^$ $@ _). - nrefine ((cat_binprod_beta_pr2 _ _ $@R _) $@ _). - nrefine (cat_pr1_fmap11_binprod _ _ $@ _). - nrefine (_ $@ cat_assoc _ _ _). - refine (_ $@ (_^$ $@R _)). - 2: nrapply cat_pr2_fmap11_binprod. - refine ((_ $@L _^$) $@ cat_assoc_opp _ _ _). + - refine (cat_assoc_opp _ _ _ $@ (_ $@R _) $@ (_ $@ _^$)). + 1,3: nrapply cat_binprod_beta_pr1. nrapply cat_binprod_beta_pr2. + - refine (cat_assoc_opp _ _ _ $@ (_ $@R _) $@ (_ $@ _^$)). + 1,3: nrapply cat_binprod_beta_pr2. + nrapply cat_binprod_beta_pr1. Defined. + Definition cat_binprod_swap_nat {a b c d : A} (f : a $-> c) (g : b $-> d) + : cat_binprod_swap c d $o fmap11 (fun x y : A => cat_binprod x y) f g + $== fmap11 (fun x y : A => cat_binprod x y) g f $o cat_binprod_swap a b + := cat_binprod_swap_corec _ _ $@ (cat_binprod_fmap11_corec _ _ _ _)^$. + Local Instance symmetricbraiding_binprod : SymmetricBraiding (fun x y => cat_binprod x y). Proof. snrapply Build_SymmetricBraiding. - - snrapply Build_Braiding. - + exact cat_binprod_swap. + - snrapply Build_NatTrans. + + intros [x y]. + exact (cat_binprod_swap x y). + intros [a b] [c d] [f g]; cbn in f, g. exact(cat_binprod_swap_nat f g). - exact cat_binprod_swap_cat_binprod_swap. @@ -655,6 +705,25 @@ Section Associativity. nrefine ((_ $@L cat_binprod_beta_pr2 _ _) $@ _). nrapply cat_pr2_fmap01_binprod. Defined. + + Definition cat_binprod_twist_corec {w x y z : A} + (f : w $-> x) (g : w $-> y) (h : w $-> z) + : cat_binprod_twist x y z $o cat_binprod_corec f (cat_binprod_corec g h) + $== cat_binprod_corec g (cat_binprod_corec f h). + Proof. + nrapply cat_binprod_eta_pr. + - nrefine (cat_assoc_opp _ _ _ $@ _). + refine ((_ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _) $@ (_ $@ _^$)). + 1: nrapply cat_binprod_pr1_twist. + 1: nrapply cat_binprod_beta_pr2. + 1,2: nrapply cat_binprod_beta_pr1. + - refine (cat_assoc_opp _ _ _ $@ (_ $@R _) $@ _ $@ (cat_binprod_beta_pr2 _ _)^$). + 1: nrapply cat_binprod_beta_pr2. + nrefine (cat_binprod_fmap01_corec _ _ _ $@ _). + nrapply cat_binprod_corec_eta. + 1: exact (Id _). + nrapply cat_binprod_beta_pr2. + Defined. Lemma cat_binprod_twist_cat_binprod_twist (x y z : A) : cat_binprod_twist x y z $o cat_binprod_twist y x z $== Id _. @@ -753,6 +822,21 @@ Section Associativity. nrefine (cat_assoc _ _ _ $@ (_ $@L cat_pr2_fmap01_binprod _ _) $@ _). exact (cat_assoc_opp _ _ _ $@ (cat_binprod_beta_pr1 _ _ $@R _)). Defined. + + Definition cat_binprod_associator_corec {w x y z} + (f : w $-> x) (g : w $-> y) (h : w $-> z) + : associator_binprod x y z $o cat_binprod_corec f (cat_binprod_corec g h) + $== cat_binprod_corec (cat_binprod_corec f g) h. + Proof. + nrefine ((Monoidal.associator_twist'_unfold _ _ _ _ _ _ _ _ $@R _) $@ _). + nrefine ((cat_assoc_opp _ _ _ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L (_ $@ _)) $@ _). + 1: nrapply cat_binprod_fmap01_corec. + 1: rapply (cat_binprod_corec_eta _ _ _ _ (Id _)). + 1: nrapply cat_binprod_swap_corec. + nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ _). + 1: nrapply cat_binprod_twist_corec. + nrapply cat_binprod_swap_corec. + Defined. Context (unit : A) `{!IsTerminal unit}. From 74166bf55dd251d8fb29db0072431248b2ff1053 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 8 May 2024 16:29:03 +0100 Subject: [PATCH 29/31] fixup Signed-off-by: Ali Caglayan --- theories/Algebra/Categorical/MonoidObject.v | 6 +-- theories/Algebra/Homological/Additive.v | 44 +++++++++++++++++++-- theories/Algebra/Homological/Biproducts.v | 7 ++++ 3 files changed, 51 insertions(+), 6 deletions(-) diff --git a/theories/Algebra/Categorical/MonoidObject.v b/theories/Algebra/Categorical/MonoidObject.v index 817c225980e..c793de21c60 100644 --- a/theories/Algebra/Categorical/MonoidObject.v +++ b/theories/Algebra/Categorical/MonoidObject.v @@ -96,17 +96,17 @@ Section ComonoidObject. Context `{!Braiding tensor}. (** A cocommutative comonoid objects is a commutative monoid object in the opposite category. *) - Class IsCocommuativeComonoidObject (x : A) + Class IsCocommutativeComonoidObject (x : A) := iscommuatativemonoid_cocomutativemonoid_op :: IsCommutativeMonoidObject (A:=A^op) tensor unit x. (** We can build cocommutative comonoid objects from the following data: *) - Definition Build_IsCocommuatativeComonoidObject (x : A) + Definition Build_IsCocommutativeComonoidObject (x : A) (** A comonoid. *) `{!IsComonoidObject x} (** Together with a proof of cocommutativity. *) (cco_cocomm : braid x x $o co_comult $== co_comult) - : IsCocommuativeComonoidObject x. + : IsCocommutativeComonoidObject x. Proof. snrapply Build_IsCommutativeMonoidObject. - exact _. diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index bbcb0bebe58..a8077e02198 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -1,10 +1,11 @@ Require Import Basics.Overture. Require Import WildCat.Core WildCat.Equiv WildCat.PointedCat WildCat.Bifunctor. -Require Import WildCat.Square WildCat.Opposite. +Require Import WildCat.Square WildCat.Opposite WildCat.Monoidal NatTrans WildCat.Products. Require Import Algebra.Homological.Biproducts. Require Import Algebra.Groups.Group. Require Import Algebra.AbGroups.AbelianGroup. Require Import canonical_names. +Require Import MonoidObject. (** * Semiadditive and Additive Categories *) @@ -25,14 +26,51 @@ Class IsSemiAdditive (A : Type) `{HasEquivs A} := { (** The final two conditions in the definition of a semiadditive category ensure that the hom types can become commutative monoids. This is an essential characteristic of semiadditive categories making it equivalent to alternate definitions where the category is semiadditive if it is enriched in commutative monoids. The machinery of encriched categories however is a bit heavy so we use this more lightweight definition where the commutative monoid structure appears naturally. *) +Local Existing Instance issymmetricmonoidal_cat_binbiprod. + +Global Instance iscocommutativecomonoidobject_semiadditive + {A : Type} `{IsSemiAdditive A} (x : A) + : IsCocommutativeComonoidObject (fun x y => cat_binbiprod x y) zero_object x. +Proof. + snrapply Build_IsCocommutativeComonoidObject. + { snrapply Build_IsComonoidObject. + - exact (cat_binbiprod_diag x). + - exact zero_morphism. + - refine (_ $@ (cat_binbiprod_fmap10_corec _ _ _)^$). + nrefine (cat_assoc _ _ _ $@ (_ $@L cat_binbiprod_fmap01_corec _ _ _) $@ _). + refine ((_ $@L (cat_binbiprod_corec_eta' (Id _) (cat_idr _))) $@ _ + $@ cat_binbiprod_corec_eta' (cat_idr _)^$ (Id _)). + nrapply Products.cat_binprod_associator_corec. + - refine (cat_assoc _ _ _ $@ (_ $@L cat_binbiprod_fmap10_corec _ _ _) $@ _). + nrefine ((cate_buildequiv_fun _ $@R _) $@ _). + unfold trans_nattrans. + nrefine (cat_assoc _ _ _ $@ (_ $@L cat_binprod_swap_corec _ _) $@ _). + nrefine ((cate_buildequiv_fun _ $@R _) $@ _). + nrapply cat_binbiprod_corec_beta_pr1. + - refine (cat_assoc _ _ _ $@ (_ $@L cat_binbiprod_fmap01_corec _ _ _) $@ _). + nrefine ((cate_buildequiv_fun _ $@R _) $@ _). + nrapply cat_binbiprod_corec_beta_pr1. } + rapply cat_binprod_swap_corec. +Defined. + Section CMonHom. Context {A : Type} `{IsSemiAdditive A} (a b : A). (** We can show that the hom-types of a semiadditive category are commutative monoids. *) - (** The zero element is given by the zero morphism. This exists from our pointedness assumption. *) - Local Instance zero_hom : Zero (a $-> b) := zero_morphism. + Local Instance zero_hom : Zero (a $-> b). + Proof. + change (MonUnit (Hom (A:=A^op) b a)). + snrapply (MonoidObject.monunit_hom (A:=A^op) (zero_object (A:=A^op)) (x := b) (y := a)). + 1-11: exact _. + (* change (IsMonoidObject (A:=A^op) ?F ?z ?x) with (IsComonoidObject (A:=A) F z x). + (IsComonoidObject (fun x y => cat_binbiprod x y) _ b). *) + pose proof (iscocommutativecomonoidobject_semiadditive a). + unfold IsCocommutativeComonoidObject in X. + rapply cmo_mo. + + := MonoidObject.monunit_hom (A:=A^op) _. (** TODO: explain a bit more: diagonal duplicates and codiagonal sums. *) (** The operation is given by the biproduct structure. *) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 18ff8edab7d..0dad774ffcf 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -883,6 +883,13 @@ Section Associativity. nrapply cat_binbiprod_swap_inl. Defined. + Definition issymmetricmonoidal_cat_binbiprod + : IsSymmetricMonoidal A (fun x y => cat_binbiprod x y) zero_object. + Proof. + change (IsSymmetricMonoidal A (fun x y => cat_binprod x y) zero_object). + exact _. + Defined. + End Associativity. (** *** Biproducts in the opposite category *) From 93182b240363dcbb07b7771adec63df942b29d2e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 26 May 2024 15:00:13 +0100 Subject: [PATCH 30/31] wip cocommutative comonoid structure gives monoid structure Signed-off-by: Ali Caglayan --- theories/Algebra/Categorical/MonoidObject.v | 151 +++++++++++++----- theories/Algebra/Homological/Additive.v | 42 ++++- theories/Algebra/Homological/Biproducts.v | 2 +- theories/WildCat/Bifunctor.v | 22 +-- theories/WildCat/Coproducts.v | 165 +++++++++++++++----- theories/WildCat/Monoidal.v | 82 +++++----- theories/WildCat/Opposite.v | 4 +- theories/WildCat/Products.v | 2 + 8 files changed, 333 insertions(+), 137 deletions(-) diff --git a/theories/Algebra/Categorical/MonoidObject.v b/theories/Algebra/Categorical/MonoidObject.v index 730d4639e66..018f8656a74 100644 --- a/theories/Algebra/Categorical/MonoidObject.v +++ b/theories/Algebra/Categorical/MonoidObject.v @@ -47,7 +47,7 @@ Arguments IsMonoidObject {A} tensor unit {_ _ _ _ _ _ _ _ _ _} x. Arguments IsCommutativeMonoidObject {A} tensor unit {_ _ _ _ _ _ _ _ _ _ _} x. Section ComonoidObject. - Context {A : Type} (tensor : A -> A -> A) (unit : A) + Context {A : Type} {tensor : A -> A -> A} {unit : A} `{HasEquivs A, !Is0Bifunctor tensor, !Is1Bifunctor tensor} `{!Associator tensor, !LeftUnitor tensor unit, !RightUnitor tensor unit}. @@ -95,6 +95,37 @@ Section ComonoidObject. Definition co_counit {x : A} `{!IsComonoidObject x} : x $-> unit := mo_unit (A:=A^op) (tensor:=tensor) (unit:=unit) (x:=x). + (** Coassociativity *) + Definition co_coassoc {x : A} `{co : !IsComonoidObject x} + : associator x x x $o fmap01 tensor x co_comult $o co_comult + $== fmap10 tensor co_comult x $o co_comult. + Proof. + nrefine (cat_assoc _ _ _ $@ _). + symmetry. + nrapply cate_moveL_Me. + exact (mo_assoc (A:=A^op) (tensor:=tensor) (unit:=unit) (x:=x)). + Defined. + + (** Left counitality *) + Definition co_left_counit {x : A} `{co : !IsComonoidObject x} + : left_unitor x $o fmap10 tensor co_counit x $o co_comult $== Id x. + Proof. + nrefine (cat_assoc _ _ _ $@ _). + nrapply cate_moveR_Me. + refine (_ $@ (cat_idr _)^$). + exact (mo_left_unit (A:=A^op) (tensor:=tensor) (unit:=unit) (x:=x)). + Defined. + + (** Right counitality *) + Definition co_right_counit {x : A} `{co : !IsComonoidObject x} + : right_unitor x $o fmap01 tensor x co_counit $o co_comult $== Id x. + Proof. + nrefine (cat_assoc _ _ _ $@ _). + nrapply cate_moveR_Me. + refine (_ $@ (cat_idr _)^$). + exact (mo_right_unit (A:=A^op) (tensor:=tensor) (unit:=unit) (x:=x)). + Defined. + Context `{!Braiding tensor}. (** A cocommutative comonoid objects is a commutative monoid object in the opposite category. *) @@ -115,8 +146,20 @@ Section ComonoidObject. - exact cco_cocomm. Defined. + Global Instance cco_co {x : A} + : IsCocommutativeComonoidObject x -> IsComonoidObject x + := fun _ => cmo_mo (A:=A^op) (tensor:=tensor) (unit:=unit) (x:=x). + + (** Cocommutativity *) + Definition cco_cocomm {x : A} `{!IsCocommutativeComonoidObject x} + : braid x x $o co_comult $== co_comult + := cmo_comm (A:=A^op) (tensor:=tensor) (unit:=unit) (x:=x). + End ComonoidObject. +Arguments IsComonoidObject {A} tensor unit {_ _ _ _ _ _ _ _ _ _} x. +Arguments IsCocommutativeComonoidObject {A} tensor unit {_ _ _ _ _ _ _ _ _ _ _} x. + (** ** Monoid enrichment *) (** A hom [x $-> y] in a cartesian category where [y] is a monoid object has the structure of a monoid. *) @@ -129,28 +172,29 @@ Section MonoidEnriched. Section Monoid. Context `{!IsMonoidObject _ _ y}. - Local Instance sgop_hom : SgOp (x $-> y) + Local Instance sgop_hom_mo : SgOp (x $-> y) := fun f g => mo_mult $o cat_binprod_corec f g. - Local Instance monunit_hom : MonUnit (x $-> y) := mo_unit $o mor_terminal _ _. + Local Instance monunit_hom_mo : MonUnit (x $-> y) + := mo_unit $o mor_terminal _ _. - Local Instance associative_hom : Associative sgop_hom. + Local Instance associative_hom_mo : Associative sgop_hom_mo. Proof. intros f g h. - unfold sgop_hom. + unfold sgop_hom_mo. rapply path_hom. refine ((_ $@L cat_binprod_fmap01_corec _ _ _)^$ $@ _). nrefine (cat_assoc_opp _ _ _ $@ _). - refine ((mo_assoc $@R _)^$ $@ _). nrefine (_ $@ (_ $@L cat_binprod_fmap10_corec _ _ _)). + refine ((mo_assoc $@R _)^$ $@ _). refine (cat_assoc _ _ _ $@ (_ $@L _) $@ cat_assoc _ _ _). nrapply cat_binprod_associator_corec. Defined. - Local Instance leftidentity_hom : LeftIdentity sgop_hom mon_unit. + Local Instance leftidentity_hom_mo : LeftIdentity sgop_hom_mo mon_unit. Proof. intros f. - unfold sgop_hom, mon_unit. + unfold sgop_hom_mo, mon_unit. rapply path_hom. refine ((_ $@L (cat_binprod_fmap10_corec _ _ _)^$) $@ cat_assoc_opp _ _ _ $@ _). nrefine (((mo_left_unit $@ _) $@R _) $@ _). @@ -162,29 +206,30 @@ Section MonoidEnriched. nrapply cat_binprod_beta_pr2. Defined. - Local Instance rightidentity_hom : RightIdentity sgop_hom mon_unit. + Local Instance rightidentity_hom_mo : RightIdentity sgop_hom_mo mon_unit. Proof. intros f. - unfold sgop_hom, mon_unit. + unfold sgop_hom_mo, mon_unit. rapply path_hom. - refine ((_ $@L (cat_binprod_fmap01_corec _ _ _)^$) $@ cat_assoc_opp _ _ _ $@ _). + refine ((_ $@L (cat_binprod_fmap01_corec _ _ _)^$) + $@ cat_assoc_opp _ _ _ $@ _). nrefine (((mo_right_unit $@ _) $@R _) $@ _). 1: nrapply cate_buildequiv_fun. nrapply cat_binprod_beta_pr1. Defined. - Local Instance issemigroup_hom : IsSemiGroup (x $-> y) := {}. - Local Instance ismonoid_hom : IsMonoid (x $-> y) := {}. + Local Instance issemigroup_hom_mo : IsSemiGroup (x $-> y) := {}. + Local Instance ismonoid_hom_mo : IsMonoid (x $-> y) := {}. End Monoid. Context `{!IsCommutativeMonoidObject _ _ y}. - Local Existing Instances sgop_hom monunit_hom ismonoid_hom. + Local Existing Instances sgop_hom_mo monunit_hom_mo ismonoid_hom_mo. - Local Instance commutative_hom : Commutative sgop_hom. + Local Instance commutative_hom : Commutative sgop_hom_mo. Proof. intros f g. - unfold sgop_hom. + unfold sgop_hom_mo. rapply path_hom. refine ((_ $@L _^$) $@ cat_assoc_opp _ _ _ $@ (cmo_comm $@R _)). nrapply cat_binprod_swap_corec. @@ -194,52 +239,78 @@ Section MonoidEnriched. End MonoidEnriched. -(** Equivalently, a hom [x $-> y] in a cartesian category where [x] is a comonoid object has the structure of a monoid. *) +(** A hom [x $-> y] in a cartesian category where [x] is a comonoid object has the structure of a monoid. It feels like this proof should be a formal dual of the one above. The proof here is a bit more non-trivial than the one above however. *) Section MonoidEnriched. Context {A : Type} `{HasEquivs A} `{!HasBinaryCoproducts A} (unit : A) `{!IsInitial unit} {x y : A} `{!HasMorExt A} `{forall x y, IsHSet (x $-> y)}. -Set Typeclasses Depth 1. Section Monoid. - Context `{!IsComonoidObject _ _ x}. + Context `{co : !IsComonoidObject _ _ x}. - Local Instance sgop_hom : SgOp (x $-> y) - := fun f g => cat_coproduct_corec f g $o co_comult. + Local Instance sgop_hom_co : SgOp (x $-> y) + := fun f g => cat_bincoprod_rec f g $o co_comult. - Local Instance monunit_hom : MonUnit (x $-> y) := co_counit $o mor_initial _ _. + Local Instance monunit_hom_co : MonUnit (x $-> y) + := mor_initial _ _ $o co_counit. - Local Instance associative_hom : Associative sgop_hom. + Local Instance associative_hom_co : Associative sgop_hom_co. Proof. intros f g h. - unfold sgop_hom. + unfold sgop_hom_co. rapply path_hom. - refine (_ $@ (cat_assoc _ _ _)^$). - refine (_ $@ (co_coassoc $@R _)^$). - refine (_ $@ cat_assoc _ _ _). - refine (_ $@ cat_coproduct_associator_corec). + refine ((cat_bincoprod_fmap01_rec _ _ _ $@R _)^$ $@ _). + nrefine (cat_assoc _ _ _ $@ _). + nrefine (_ $@ (cat_bincoprod_fmap10_rec _ _ _ $@R _)). + refine ((cat_bincoprod_rec_associator _ _ _ $@R _)^$ $@ _). + nrefine (cat_assoc _ _ _ $@ (_ $@L _) $@ cat_assoc_opp _ _ _). + nrefine (cat_assoc_opp _ _ _ $@ _). + nrapply co_coassoc. Defined. - Local Instance leftidentity_hom : LeftIdentity sgop_hom mon_unit. + Local Instance leftidentity_hom : LeftIdentity sgop_hom_co mon_unit. Proof. intros f. - unfold sgop_hom, mon_unit. + unfold sgop_hom_co, mon_unit, monunit_hom_co. rapply path_hom. - refine (_ $@ cat_assoc_opp _ _ _). - refine (_ $@ (co_left_counit $@R _)^$). - refine (_ $@ cat_assoc _ _ _). - refine (_ $@ cat_coproduct_beta_inl). + refine ((cat_bincoprod_fmap10_rec _ _ _ $@R _)^$ $@ cat_assoc _ _ _ $@ _). + snrefine ((_ $@R _) $@ _). + 1: exact (f $o left_unitor x). + { refine (_ $@ (_ $@L (_ $@ _))^$). + 2,3: nrapply cate_inv_adjointify. + nrefine (_ $@ cat_assoc _ _ _). + nrapply cate_moveL_eV. + nrefine ((_ $@L cate_buildequiv_fun _) $@ _). + nrapply cate_moveL_eV. + nrefine ((_ $@L cate_buildequiv_fun _) $@ _). + nrefine ((_ $@L cate_buildequiv_fun _) $@ _). + nrefine ((_ $@R _) $@ _). + 1: rapply cat_bincoprod_swap_rec. + apply cat_bincoprod_beta_inl. } + nrefine (cat_assoc _ _ _ $@ (_ $@L (cat_assoc_opp _ _ _ $@ _)) $@ cat_idr _). + nrapply co_left_counit. + (* TODO 5s! why so slow? *) Defined. - Local Instance rightidentity_hom : RightIdentity sgop_hom mon_unit. + Local Instance rightidentity_hom : RightIdentity sgop_hom_co mon_unit. Proof. intros f. - unfold sgop_hom, mon_unit. + unfold sgop_hom_co, mon_unit, monunit_hom_co. rapply path_hom. - refine (_ $@ cat_assoc_opp _ _ _). - refine (_ $@ (co_right_counit $@R _)^$). - refine (_ $@ cat_coproduct_beta_inr). + refine ((cat_bincoprod_fmap01_rec _ _ _ $@R _)^$ $@ cat_assoc _ _ _ $@ _). + snrefine ((_ $@R _) $@ _). + 1: exact (f $o right_unitor x). + { refine (_ $@ (_ $@L _)^$). + 2: nrapply cate_inv_adjointify. + rapply cat_bincoprod_eta_in. + - refine (cat_bincoprod_beta_inl _ _ $@ (cat_idr _)^$ + $@ (_ $@L _^$) $@ cat_assoc_opp _ _ _). + nrapply cat_bincoprod_beta_inl. + - nrefine (cat_bincoprod_beta_inr _ _ $@ _ $@ cat_assoc_opp _ _ _). + nrapply mor_initial_unique. } + nrefine (cat_assoc _ _ _ $@ (_ $@L (cat_assoc_opp _ _ _ $@ _)) $@ cat_idr _). + nrapply co_right_counit. Defined. Local Instance issemigroup_hom : IsSemiGroup (x $-> y) := {}. @@ -247,6 +318,4 @@ Set Typeclasses Depth 1. End Monoid. - End MonoidEnriched. - diff --git a/theories/Algebra/Homological/Additive.v b/theories/Algebra/Homological/Additive.v index a8077e02198..e55b43391dd 100644 --- a/theories/Algebra/Homological/Additive.v +++ b/theories/Algebra/Homological/Additive.v @@ -1,6 +1,6 @@ Require Import Basics.Overture. Require Import WildCat.Core WildCat.Equiv WildCat.PointedCat WildCat.Bifunctor. -Require Import WildCat.Square WildCat.Opposite WildCat.Monoidal NatTrans WildCat.Products. +Require Import WildCat.Square WildCat.Opposite WildCat.Monoidal NatTrans WildCat.Products WildCat.Coproducts. Require Import Algebra.Homological.Biproducts. Require Import Algebra.Groups.Group. Require Import Algebra.AbGroups.AbelianGroup. @@ -30,7 +30,7 @@ Local Existing Instance issymmetricmonoidal_cat_binbiprod. Global Instance iscocommutativecomonoidobject_semiadditive {A : Type} `{IsSemiAdditive A} (x : A) - : IsCocommutativeComonoidObject (fun x y => cat_binbiprod x y) zero_object x. + : IsCocommutativeComonoidObject (fun x y => cat_binprod x y) zero_object x. Proof. snrapply Build_IsCocommutativeComonoidObject. { snrapply Build_IsComonoidObject. @@ -61,16 +61,48 @@ Section CMonHom. Local Instance zero_hom : Zero (a $-> b). Proof. + snrapply MonoidObject.monunit_hom_co. + 1-5: exact _. + 1: exact zero_object. + 1: exact _. + (* Set Printing All. *) + (* Check ( cco_co (iscocommutativecomonoidobject_semiadditive a)). *) + unfold cat_bincoprod. + snrefine (transport (fun x => @IsComonoidObject _ _ _ _ _ _ _ _ _ x _ _ _ a) _ _). + 3:{ + (* TODO: we need to redefine biproducts so that the coproduct strucutre is emergent from the birpoduct structure rather than being part of the structure. Hopefully this is enough for the following term to be accepted. *) + pose (i := cco_co (iscocommutativecomonoidobject_semiadditive a)). + exact i. + (* Arguments IsComonoidObject : clear implicits. + Arguments IsCocommutativeComonoidObject : clear implicits. + pose proof (i := cco_co (iscocommutativecomonoidobject_semiadditive a)). + unfold cat_bincoprod. + unfold is1bifunctor_cat_bincoprod. + Arguments is1bifunctor_op' : clear implicits. + unfold is1bifunctor_op'. + unfold is1bifunctor_op. + + exact i. + + unfold Coproducts.is1bifunctor_cat_bincoprod. + unfold Coproducts.is1bifunctor_cat_bincoprod. + + exact i. + exact _. + + + + change (MonUnit (Hom (A:=A^op) b a)). - snrapply (MonoidObject.monunit_hom (A:=A^op) (zero_object (A:=A^op)) (x := b) (y := a)). + snrapply ( (A:=A^op) (zero_object (A:=A^op)) (x := b) (y := a)). 1-11: exact _. (* change (IsMonoidObject (A:=A^op) ?F ?z ?x) with (IsComonoidObject (A:=A) F z x). (IsComonoidObject (fun x y => cat_binbiprod x y) _ b). *) pose proof (iscocommutativecomonoidobject_semiadditive a). unfold IsCocommutativeComonoidObject in X. rapply cmo_mo. - - := MonoidObject.monunit_hom (A:=A^op) _. + *) + Admitted. (** TODO: explain a bit more: diagonal duplicates and codiagonal sums. *) (** The operation is given by the biproduct structure. *) diff --git a/theories/Algebra/Homological/Biproducts.v b/theories/Algebra/Homological/Biproducts.v index 0dad774ffcf..fb8ec756544 100644 --- a/theories/Algebra/Homological/Biproducts.v +++ b/theories/Algebra/Homological/Biproducts.v @@ -193,7 +193,7 @@ Definition cat_biprod_diag I {A} (x : A) `{Biproduct I A (fun _ => x)} Definition cat_biprod_codiag I {A} (x : A) `{Biproduct I A (fun _ => x)} : cat_biprod I (fun _ => x) $-> x - := cat_coprod_fold x $o (cate_coprod_prod (fun _ => x))^-1$. + := cat_coprod_codiag x $o (cate_coprod_prod (fun _ => x))^-1$. (** Compatability of [cat_biprod_rec] and [cat_biprod_corec]. *) Definition cat_biprod_corec_rec I `{DecidablePaths I} {A : Type} diff --git a/theories/WildCat/Bifunctor.v b/theories/WildCat/Bifunctor.v index 8359a7115c6..dbbcc183ea7 100644 --- a/theories/WildCat/Bifunctor.v +++ b/theories/WildCat/Bifunctor.v @@ -381,20 +381,14 @@ Global Instance is1functor_uncurry_uncurry_right {A B C D E} !Is0Bifunctor F, !Is1Bifunctor F, !Is0Bifunctor G, !Is1Bifunctor G} : Is1Functor (uncurry (uncurry (fun x y z => G x (F y z)))). Proof. - - (* apply is1functor_uncurry_bifunctor. - nrapply Build_Is1Bifunctor. - 1: exact _. - - intros b. - change (Is1Functor (uncurry (fun x y => G x (F y b)))). - exact _. - - intros [c a] [c' a'] [h f] b b' g. - refine ((cat_assoc _ _ _)^$ $@ _ $@ (cat_assoc _ _ _)^$). - refine ((_ $@R _) $@ cat_assoc _ _ _ $@ (_ $@L _)). - 2: exact (fmap11_coh G h (fmap01 F a g)). - rapply fmap_square; unfold Square. - exact (fmap11_coh F f g). *) -Admitted. + snrapply Build_Is1Functor. + - intros cab cab' [[h f] g] [[h' f'] g'] [[q p] r]. + exact (fmap22 G q (fmap22 F p r)). + - intros cab. + exact (fmap12 G _ (fmap11_id F _ _) $@ fmap11_id G _ _). + - intros cab cab' cab'' [[h f] g] [[h' f'] g']. + exact (fmap12 G _ (fmap11_comp F _ _ _ _) $@ fmap11_comp G _ _ _ _). +Defined. Definition fmap11_square {A B C : Type} `{Is1Cat A, Is1Cat B, Is1Cat C} (F : A -> B -> C) `{!Is0Bifunctor F, !Is1Bifunctor F} diff --git a/theories/WildCat/Coproducts.v b/theories/WildCat/Coproducts.v index 826d7d56800..2f7dc029063 100644 --- a/theories/WildCat/Coproducts.v +++ b/theories/WildCat/Coproducts.v @@ -83,7 +83,7 @@ End Lemmata. (** *** Codiagonal / fold map *) -Definition cat_coprod_fold {I : Type} {A : Type} (x : A) `{Coproduct I _ (fun _ => x)} +Definition cat_coprod_codiag {I : Type} {A : Type} (x : A) `{Coproduct I _ (fun _ => x)} : cat_coprod I (fun _ => x) $-> x := cat_prod_diag (A:=A^op) x. @@ -141,9 +141,9 @@ Defined. (** *** Binary coproducts *) Class BinaryCoproduct {A : Type} `{Is1Cat A} (x y : A) - := prod_co_bincoprod :: BinaryProduct (x : A^op) (y : A^op). + := prod_co_bincoprod :: BinaryProduct (A:=A^op) x y. -Definition cat_bincoprod {A : Type} `{Is1Cat A} (x y : A) `{!BinaryCoproduct x y} : A +Definition cat_bincoprod {A : Type} `{Is1Cat A} (x y : A) `{!BinaryCoproduct x y} : A := cat_binprod (x : A^op) y. Definition cat_inl {A : Type} `{Is1Cat A} {x y : A} `{!BinaryCoproduct x y} @@ -211,25 +211,6 @@ Section Lemmata. := @cat_binprod_corec_eta A^op _ _ _ _ x y _ _ f f' g g'. End Lemmata. -(** *** Symmetry of coproducts *) - -Definition cate_bincoprod_swap {A : Type} `{HasEquivs A} - {e : HasBinaryCoproducts A} (x y : A) - : cat_bincoprod x y $<~> cat_bincoprod y x. -Proof. - exact (@cate_binprod_swap A^op _ _ _ _ _ e _ _). -Defined. - -(** *** Associativity of coproducts *) - -Lemma cate_coprod_assoc {A : Type} `{HasEquivs A} - {e : HasBinaryCoproducts A} (x y z : A) - : cat_bincoprod x (cat_bincoprod y z) - $<~> cat_bincoprod (cat_bincoprod x y) z. -Proof. - exact (@associator_binprod A^op _ _ _ _ _ e x y z)^-1$. -Defined. - (** *** Binary coproduct functor *) (** Hint: Use [Set Printing Implicit] to see the implicit arguments in the following proofs. *) @@ -282,33 +263,139 @@ Proof. exact (is1bifunctor_cat_binprod (A:=A^op) (H0:=H0)). Defined. -(** *** Cocartesian Monoidal Structure *) +(** *** Products and coproducts in the opposite category *) -Global Instance ismonoidal_cat_bincoprod {A : Type} `{HasEquivs A} - {e : HasBinaryCoproducts A} (zero : A) `{!IsInitial zero} - : IsMonoidal A (fun x y => cat_bincoprod x y) zero. +Definition hasbinarycoproducts_op_hasbinaryproducts {A : Type} + `{H : HasBinaryProducts A} + : HasBinaryCoproducts A^op + := H. +Hint Immediate hasbinarycoproducts_op_hasbinaryproducts : typeclass_instances. + +Definition hasbinarycoproducts_hasbinaryproducts_op {A : Type} + `{HasEquivs A, H' : !HasBinaryProducts A^op} + : HasBinaryCoproducts A + := H'. +Hint Immediate hasbinarycoproducts_hasbinaryproducts_op : typeclass_instances. + +Definition hasbinaryproducts_op_hasbinarycoproducts {A : Type} + `{H : HasBinaryCoproducts A} + : HasBinaryProducts A^op + := H. +Hint Immediate hasbinarycoproducts_op_hasbinaryproducts : typeclass_instances. + +Definition hasbinaryproducts_hasbinarycoproducts_op {A : Type} + `{HasEquivs A, H' : !HasBinaryCoproducts A^op} + : HasBinaryProducts A + := H'. +Hint Immediate hasbinaryproducts_hasbinarycoproducts_op : typeclass_instances. + +(** *** Symmetry of coproducts *) + +Definition cat_bincoprod_swap {A : Type} `{HasEquivs A} + {e : HasBinaryCoproducts A} (x y : A) + : cat_bincoprod x y $-> cat_bincoprod y x. Proof. - + exact (@cat_binprod_swap A^op _ _ _ _ e _ _). +Defined. - unshelve epose proof (@ismonoidal_op A^op (fun x y => cat_bincoprod x y) zero _ _ _ _ _ - (ismonoidal_binprod (A:=A^op) zero) - ). - 4-7: exact _. - 1: exact _. - +Definition cate_bincoprod_swap {A : Type} `{HasEquivs A} + {e : HasBinaryCoproducts A} (x y : A) + : cat_bincoprod x y $<~> cat_bincoprod y x. +Proof. + exact (@cate_binprod_swap A^op _ _ _ _ _ e _ _). +Defined. +(** *** Associativity of coproducts *) - (* unshelve epose proof (ismonoidal_op (fun x y => cat_binprod x y) zero); try exact _. *) +Lemma cate_coprod_assoc {A : Type} `{HasEquivs A} + {e : HasBinaryCoproducts A} (x y z : A) + : cat_bincoprod x (cat_bincoprod y z) + $<~> cat_bincoprod (cat_bincoprod x y) z. +Proof. + exact (@associator_binprod A^op _ _ _ _ _ e x y z)^-1$. +Defined. +Definition associator_cat_bincoprod {A : Type} `{HasEquivs A} + {e : HasBinaryCoproducts A} + : Associator (fun x y => cat_bincoprod x y). +Proof. + unfold Associator. + snrapply associator_op'. + 1: exact _. + nrapply associator_binprod. +Defined. + +(** *** Codiagonal *) + +Definition cat_bincoprod_codiag {A : Type} + `{HasEquivs A} (x : A) `{!BinaryCoproduct x x} + : cat_bincoprod x x $-> x + := cat_binprod_diag (A:=A^op) x. + +(** *** Lemmas about [cat_bincoprod_rec] *) + +Definition cat_bincoprod_fmap01_rec {A : Type} + `{Is1Cat A, !HasBinaryCoproducts A} {w x y z : A} + (f : z $-> w) (g : y $-> x) (h : x $-> w) + : cat_bincoprod_rec f h $o fmap01 (fun x y => cat_bincoprod x y) z g + $== cat_bincoprod_rec f (h $o g) + := @cat_binprod_fmap01_corec A^op _ _ _ _ + hasbinaryproducts_op_hasbinarycoproducts _ _ _ _ f g h. + +Definition cat_bincoprod_fmap10_rec {A : Type} + `{Is1Cat A, !HasBinaryCoproducts A} {w x y z : A} + (f : y $-> x) (g : x $-> w) (h : z $-> w) + : cat_bincoprod_rec g h $o fmap10 (fun x y => cat_bincoprod x y) f z + $== cat_bincoprod_rec (g $o f) h + := @cat_binprod_fmap10_corec A^op _ _ _ _ + hasbinaryproducts_op_hasbinarycoproducts _ _ _ _ f g h. + +Definition cat_bincoprod_fmap11_rec {A : Type} + `{Is1Cat A, !HasBinaryCoproducts A} {v w x y z : A} + (f : y $-> w) (g : z $-> x) (h : w $-> v) (i : x $-> v) + : cat_bincoprod_rec h i $o fmap11 (fun x y => cat_binprod x y) f g + $== cat_bincoprod_rec (h $o f) (i $o g) + := @cat_binprod_fmap11_corec A^op _ _ _ _ + hasbinaryproducts_op_hasbinarycoproducts _ _ _ _ _ f g h i. + +Definition cat_bincoprod_rec_associator {A : Type} `{HasEquivs A} + `{!HasBinaryCoproducts A} {w x y z : A} (f : w $-> z) (g : x $-> z) (h : y $-> z) + : cat_bincoprod_rec (cat_bincoprod_rec f g) h $o associator_cat_bincoprod w x y + $== cat_bincoprod_rec f (cat_bincoprod_rec g h). +Proof. + nrapply cate_moveR_eV. + symmetry. + (** TODO not sure why this can't be inlined *) + pose (X := @cat_binprod_associator_corec A^op _ _ _ _ _ + hasbinaryproducts_op_hasbinarycoproducts _ _ _ _ f g h). + exact X. +Defined. + +Definition cat_bincoprod_swap_rec {A : Type} `{HasEquivs A} + `{!HasBinaryCoproducts A} {a b c : A} (f : a $-> c) (g : b $-> c) + : cat_bincoprod_rec f g $o cat_bincoprod_swap b a $== cat_bincoprod_rec g f + := @cat_binprod_swap_corec A^op _ _ _ _ + hasbinaryproducts_op_hasbinarycoproducts _ _ _ _ _. + +(** *** Cocartesian Monoidal Structure *) - snrapply Build_IsMonoidal. - 1-2: exact _. -Admitted. +Global Instance ismonoidal_cat_bincoprod {A : Type} `{HasEquivs A} + {e : HasBinaryCoproducts A} (zero : A) `{!IsInitial zero} + : IsMonoidal A (fun x y => cat_bincoprod x y) zero. +Proof. + nrapply ismonoidal_op'. + nrapply (ismonoidal_binprod (A:=A^op) zero). + by nrapply isterminal_op_isinitial. +Defined. Global Instance issymmetricmonoidal_cat_bincoprod {A : Type} `{HasEquivs A} {e : HasBinaryCoproducts A} (zero : A) `{!IsInitial zero} - : IsSymmetricMonoidal A (fun x y => cat_bincoprod x y) zero - := issymmetricmonoidal_binprod (A:=A^op) zero. + : IsSymmetricMonoidal A (fun x y => cat_bincoprod x y) zero. +Proof. + nrapply issymmetricmonoidal_op'. + nrapply (issymmetricmonoidal_binprod (A :=A^op) zero). + by nrapply isterminal_op_isinitial. +Defined. (** *** Coproducts in Type *) diff --git a/theories/WildCat/Monoidal.v b/theories/WildCat/Monoidal.v index 9f02fdc77b2..ea2adc38023 100644 --- a/theories/WildCat/Monoidal.v +++ b/theories/WildCat/Monoidal.v @@ -124,8 +124,8 @@ Class IsMonoidal (A : Type) `{HasEquivs A} (** These all satisfy the following properties: *) := { (** A [cat_tensor] is a 1-bifunctor. *) - is0bifunctor_cat_tensor :: Is0Bifunctor cat_tensor; - is1bifunctor_cat_tensor :: Is1Bifunctor cat_tensor; + is0bifunctor_cat_tensor : Is0Bifunctor cat_tensor; + is1bifunctor_cat_tensor : Is1Bifunctor cat_tensor; (** A natural isomorphism [associator] witnessing the associativity of the tensor product. *) cat_tensor_associator :: Associator cat_tensor; (** A natural isomorphism [left_unitor] witnessing the left unit law. *) @@ -138,6 +138,9 @@ Class IsMonoidal (A : Type) `{HasEquivs A} cat_tensor_pentagon_identity :: PentagonIdentity cat_tensor; }. +Existing Instance is0bifunctor_cat_tensor | 10. +Existing Instance is1bifunctor_cat_tensor | 10. + (** TODO: Braided monoidal categories *) (** ** Symmetric Monoidal Categories *) @@ -207,6 +210,11 @@ Section Associator. End Associator. +Definition associator_op' {A : Type} `{HasEquivs A} {F : A -> A -> A} + `{!Is0Bifunctor F, !Is1Bifunctor F, assoc : !Associator (A:=A^op) F} + : Associator F + := associator_op (A:=A^op) (assoc := assoc). + (** ** Theory about [LeftUnitor] and [RightUnitor] *) Section LeftUnitor. @@ -230,13 +238,8 @@ Global Instance braiding_op {A : Type} `{HasEquivs A} {F : A -> A -> A} Definition braiding_op' {A : Type} `{HasEquivs A} {F : A -> A -> A} `{!Is0Bifunctor F, !Is1Bifunctor F, braid : !Braiding (A:=A^op) F} - : Braiding F. -Proof. - - snrapply braiding_op. - - snrapply (nattrans_op (A:=A^op)). - := nattrans_op (A:=A^op) (nattrans_flip (A:=A^op * A^op) braid). + : Braiding F + := braiding_op (A:=A^op) (braid := braid). (** ** Theory about [SymmetricBraid] *) @@ -432,13 +435,10 @@ Section SymmetricBraid. End SymmetricBraid. Definition symmetricbraiding_op' {A : Type} {F : A -> A -> A} - `{HasEquivs A, !Is0Bifunctor F, !Is1Bifunctor F, !SymmetricBraiding (A:=A^op)F} - : SymmetricBraiding F. -Proof. - snrapply Build_SymmetricBraiding. - - -Admitted. - + `{HasEquivs A, !Is0Bifunctor F, !Is1Bifunctor F, + H' : !SymmetricBraiding (A:=A^op) F} + : SymmetricBraiding F + := symmetricbraiding_op (A:=A^op) (F := F). Global Instance ismonoidal_op {A : Type} (tensor : A -> A -> A) (unit : A) `{IsMonoidal A tensor unit} @@ -466,10 +466,8 @@ Defined. Definition ismonoidal_op' {A : Type} (tensor : A -> A -> A) (unit : A) `{HasEquivs A} `{!IsMonoidal A^op tensor unit} - : IsMonoidal A tensor unit. -Proof. - -Admitted. + : IsMonoidal A tensor unit + := ismonoidal_op (A:=A^op) tensor unit. Global Instance issymmetricmonoidal_op {A : Type} (tensor : A -> A -> A) (unit : A) `{IsSymmetricMonoidal A tensor unit} @@ -479,23 +477,37 @@ Proof. - rapply ismonoidal_op. - rapply symmetricbraiding_op. - intros a b c; unfold op in a, b, c; simpl. - refine (_ $@ (_ $@L _) $@ (_ $@R _)). - (** TODO should be easy *) -Admitted. + snrefine (_ $@ (_ $@L (_ $@R _))). + 2: exact ((braide _ _)^-1$). + 2: { nrapply cate_moveR_V1. + symmetry. + nrefine ((_ $@R _) $@ _). + 1: nrapply cate_buildequiv_fun. + rapply braid_braid. } + snrefine ((_ $@R _) $@ _). + { refine (emap _ _)^-1$. + rapply braide. } + { symmetry. + refine (cate_inv_adjointify _ _ _ _ $@ fmap2 _ _). + nrapply cate_inv_adjointify. } + snrefine ((_ $@L (_ $@L _)) $@ _). + { refine (emap (flip tensor c) _)^-1$. + rapply braide. } + { symmetry. + refine (cate_inv_adjointify _ _ _ _ $@ fmap2 _ _). + nrapply cate_inv_adjointify. } + refine ((_ $@L _)^$ $@ _^$ $@ cate_inv2 _ $@ _ $@ (_ $@L _)). + 1,2,4,5: rapply cate_inv_compose'. + refine (_ $@ (_ $@@ _) $@ _ $@ (_ $@R _)^$ $@ _^$). + 1-3,5-6: rapply cate_buildequiv_fun. + refine ((fmap02 _ _ _ $@@ ((_ $@ fmap20 _ _ _) $@R _)) $@ cat_symm_tensor_hexagon a b c $@ ((_ $@L _^$) $@R _)). + 1-4: nrapply cate_buildequiv_fun. +Defined. Definition issymmetricmonoidal_op' {A : Type} (tensor : A -> A -> A) (unit : A) - `{HasEquivs A} `{!IsSymmetricMonoidal A^op tensor unit} - : IsSymmetricMonoidal A tensor unit. -Proof. - snrapply Build_IsSymmetricMonoidal. - - nrapply ismonoidal_op'. - rapply issymmetricmonoidal_ismonoidal. - - nrapply symmetricbraiding_op. - - - - -Admitted. - + `{HasEquivs A} `{H' : !IsSymmetricMonoidal A^op tensor unit} + : IsSymmetricMonoidal A tensor unit + := issymmetricmonoidal_op (A:=A^op) tensor unit. (** ** Building Symmetric Monoidal Categories *) diff --git a/theories/WildCat/Opposite.v b/theories/WildCat/Opposite.v index 1c515862c38..8d26289d5bf 100644 --- a/theories/WildCat/Opposite.v +++ b/theories/WildCat/Opposite.v @@ -127,10 +127,10 @@ Proof. refine (@isequiv_Htpy_path _ _ _ _ _ H0 b a f g). Defined. -Global Instance isinitial_isterminal_op {A : Type} `{Is1Cat A} (x : A) +Global Instance isinitial_op_isterminal {A : Type} `{Is1Cat A} (x : A) {t : IsTerminal x} : IsInitial (A := A^op) x := t. -Global Instance isterminal_isinitial_op {A : Type} `{Is1Cat A} (x : A) +Global Instance isterminal_op_isinitial {A : Type} `{Is1Cat A} (x : A) {i : IsInitial x} : IsTerminal (A := A^op) x := i. diff --git a/theories/WildCat/Products.v b/theories/WildCat/Products.v index c3010d0ab82..c125fee610c 100644 --- a/theories/WildCat/Products.v +++ b/theories/WildCat/Products.v @@ -566,6 +566,8 @@ Proof. snrapply cat_binprod_corec; exact (Id _). Defined. +(** *** Lemmas about [cat_binprod_corec] *) + Definition cat_binprod_fmap01_corec {A : Type} `{Is1Cat A, !HasBinaryProducts A} {w x y z : A} (f : w $-> z) (g : x $-> y) (h : w $-> x) From 6fc082e342ad16aa9e93fac3af9f52f43ed57b16 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 2 Jul 2024 15:21:18 +0200 Subject: [PATCH 31/31] fix compilation issue in MonoidObject.v Signed-off-by: Ali Caglayan --- theories/Algebra/Categorical/MonoidObject.v | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/theories/Algebra/Categorical/MonoidObject.v b/theories/Algebra/Categorical/MonoidObject.v index 62dfd5013be..56ff2d163d2 100644 --- a/theories/Algebra/Categorical/MonoidObject.v +++ b/theories/Algebra/Categorical/MonoidObject.v @@ -1,6 +1,6 @@ Require Import Basics.Overture Basics.Tactics. Require Import WildCat.Core WildCat.Equiv WildCat.Monoidal WildCat.Bifunctor - WildCat.NatTrans WildCat.Opposite WildCat.Products WildCat.Coproducts. + WildCat.NatTrans WildCat.Opposite WildCat.Products. Require Import abstract_algebra. (** * Monoids and Comonoids *) @@ -47,7 +47,7 @@ Arguments IsMonoidObject {A} tensor unit {_ _ _ _ _ _ _ _ _ _} x. Arguments IsCommutativeMonoidObject {A} tensor unit {_ _ _ _ _ _ _ _ _ _ _} x. Section ComonoidObject. - Context {A : Type} {tensor : A -> A -> A} {unit : A} + Context {A : Type} (tensor : A -> A -> A) (unit : A) `{HasEquivs A, !Is0Bifunctor tensor, !Is1Bifunctor tensor} `{!Associator tensor, !LeftUnitor tensor unit, !RightUnitor tensor unit}. @@ -146,11 +146,12 @@ Section ComonoidObject. - exact cco_cocomm. Defined. - Global Instance co_cco {x : A} `{!IsCocommutativeComonoidObject x} + Definition co_cco {x : A} `{!IsCocommutativeComonoidObject x} : IsComonoidObject x. Proof. apply cmo_mo. Defined. + Hint Immediate co_cco : typeclass_instances. (** Cocommutativity *) Definition cco_cocomm {x : A} `{!IsCocommutativeComonoidObject x}