Library UniMath.CategoryTheory.FunctorAlgebras
****************************************************************
Benedikt Ahrens
started March 2015
Extended by: Anders Mörtberg. October 2015
Rewritten using displayed categories by: Kobe Wullaert. October 2022
Extended by Ralph Matthes in February 2024 for the special case of functors that are binary or I-ary coproducts
***************************************************************
Contents :
- Category of algebras of an endofunctor
- This category is saturated if base precategory is
- Lambek's lemma: if (A,a) is an inital F-algebra then a is an iso
- The natural numbers are initial for X ↦ 1 + X
- the case of functors that are binary or I-ary coproducts
Require Import UniMath.Foundations.PartD.
Require Import UniMath.Foundations.Propositions.
Require Import UniMath.Foundations.Sets.
Require Import UniMath.MoreFoundations.Tactics.
Require Import UniMath.CategoryTheory.Core.Categories.
Require Import UniMath.CategoryTheory.Core.Isos.
Require Import UniMath.CategoryTheory.Core.Univalence.
Require Import UniMath.CategoryTheory.Core.Functors.
Require Import UniMath.CategoryTheory.whiskering.
Require Import UniMath.CategoryTheory.Limits.Initial.
Require Import UniMath.CategoryTheory.DisplayedCats.Core.
Require Import UniMath.CategoryTheory.DisplayedCats.Total.
Require Import UniMath.CategoryTheory.DisplayedCats.Constructions.
Require Import UniMath.CategoryTheory.DisplayedCats.Isos.
Require Import UniMath.CategoryTheory.DisplayedCats.Univalence.
Require Import UniMath.CategoryTheory.Limits.Terminal.
Require Import UniMath.CategoryTheory.Limits.BinCoproducts.
Require Import UniMath.CategoryTheory.Arithmetic.NNO.
Require Import UniMath.CategoryTheory.PrecategoryBinProduct.
Require Import UniMath.CategoryTheory.ProductCategory.
Require Import UniMath.CategoryTheory.Limits.Coproducts.
Local Open Scope cat.
Section Algebra_Definition.
Context {C : category} (F : functor C C).
Definition algebra_disp_cat_ob_mor : disp_cat_ob_mor C.
Proof.
use tpair.
- exact (λ x, F x --> x).
- exact (λ x y hx hy f, hx · f = #F f · hy).
Defined.
Definition algebra_disp_cat_id_comp
: disp_cat_id_comp C algebra_disp_cat_ob_mor.
Proof.
split.
- intros x hx ; cbn.
rewrite !functor_id.
rewrite id_left, id_right.
apply idpath.
- intros x y z f g hx hy hz hf hg ; cbn in *.
rewrite !functor_comp.
rewrite !assoc.
rewrite hf.
rewrite !assoc'.
rewrite hg.
apply idpath.
Qed.
Definition algebra_disp_cat_data : disp_cat_data C
:= algebra_disp_cat_ob_mor ,, algebra_disp_cat_id_comp.
Lemma is_locally_propositional_algebra_disp_cat :
locally_propositional algebra_disp_cat_data.
Proof.
red; intros ; apply homset_property.
Qed.
Definition algebra_disp_cat : disp_cat C
:= make_disp_cat_locally_prop is_locally_propositional_algebra_disp_cat.
Definition category_FunctorAlg : category
:= total_category algebra_disp_cat.
Definition FunctorAlg := category_FunctorAlg.
Definition algebra_ob : UU := ob FunctorAlg.
Definition alg_carrier (X : algebra_ob) : C := pr1 X.
Local Coercion alg_carrier : algebra_ob >-> ob.
Definition alg_map (X : algebra_ob) : F X --> X := pr2 X.
A morphism of F-algebras (F X, g : F X --> X) and (F Y, h : F Y --> Y)
is a morphism f : X --> Y such that the following diagram commutes:
>>>>>>> master F f F x ----> F y | | | g | h V V x ------> y f
Definition is_algebra_mor (X Y : algebra_ob) (f : alg_carrier X --> alg_carrier Y) : UU
:= alg_map X · f = #F f · alg_map Y.
Definition algebra_mor (X Y : algebra_ob) : UU := FunctorAlg⟦X,Y⟧.
Coercion mor_from_algebra_mor {X Y : algebra_ob} (f : algebra_mor X Y) : C⟦X, Y⟧ := pr1 f.
Lemma algebra_mor_commutes (X Y : algebra_ob) (f : algebra_mor X Y)
: alg_map X · f = #F f · alg_map Y.
Proof.
exact (pr2 f).
Qed.
End Algebra_Definition.
Definition algebra_mor_eq' {C : category} {F : functor C C} {X Y : algebra_ob F} (f g : algebra_mor F X Y)
: (f : alg_carrier F X --> alg_carrier F Y) = g ≃ f = g.
Proof.
apply invweq.
apply subtypeInjectivity.
intro a. apply C.
Defined.
Definition algebra_mor_eq {C : category} {F : functor C C} {X Y : FunctorAlg F} (f g : (FunctorAlg F)⟦X,Y⟧)
: ((pr1 f : alg_carrier F X --> alg_carrier F Y) = (pr1 g)) -> f = g.
Proof.
exact (algebra_mor_eq' f g).
Defined.
Section fixacategory.
Context {C : category}
(F : functor C C).
:= alg_map X · f = #F f · alg_map Y.
Definition algebra_mor (X Y : algebra_ob) : UU := FunctorAlg⟦X,Y⟧.
Coercion mor_from_algebra_mor {X Y : algebra_ob} (f : algebra_mor X Y) : C⟦X, Y⟧ := pr1 f.
Lemma algebra_mor_commutes (X Y : algebra_ob) (f : algebra_mor X Y)
: alg_map X · f = #F f · alg_map Y.
Proof.
exact (pr2 f).
Qed.
End Algebra_Definition.
Definition algebra_mor_eq' {C : category} {F : functor C C} {X Y : algebra_ob F} (f g : algebra_mor F X Y)
: (f : alg_carrier F X --> alg_carrier F Y) = g ≃ f = g.
Proof.
apply invweq.
apply subtypeInjectivity.
intro a. apply C.
Defined.
Definition algebra_mor_eq {C : category} {F : functor C C} {X Y : FunctorAlg F} (f g : (FunctorAlg F)⟦X,Y⟧)
: ((pr1 f : alg_carrier F X --> alg_carrier F Y) = (pr1 g)) -> f = g.
Proof.
exact (algebra_mor_eq' f g).
Defined.
Section fixacategory.
Context {C : category}
(F : functor C C).
forgetful functor from FunctorAlg to its underlying category
Definition forget_algebras : functor (category_FunctorAlg F) C := pr1_category (algebra_disp_cat F).
End fixacategory.
Section FunctorAlg_saturated.
Context {C : category}
(H : is_univalent C)
(F : functor C C).
Definition algebra_eq_type (X Y : FunctorAlg F) : UU
:= ∑ p : z_iso (pr1 X) (pr1 Y), is_algebra_mor F X Y p.
Definition algebra_ob_eq (X Y : FunctorAlg F) :
(X = Y) ≃ algebra_eq_type X Y.
Proof.
eapply weqcomp.
- apply total2_paths_equiv.
- set (H1 := make_weq _ (H (pr1 X) (pr1 Y))).
apply (weqbandf H1).
simpl.
intro p.
destruct X as [X α].
destruct Y as [Y β]; simpl in *.
destruct p.
rewrite idpath_transportf.
unfold is_algebra_mor; simpl.
rewrite functor_id.
rewrite id_left, id_right.
apply idweq.
Defined.
Definition is_z_iso_from_is_algebra_iso (X Y : FunctorAlg F) (f : X --> Y)
: is_z_isomorphism f → is_z_isomorphism (pr1 f).
Proof.
intro p.
set (H' := z_iso_inv_after_z_iso (make_z_iso' f p)).
set (H'':= z_iso_after_z_iso_inv (make_z_iso' f p)).
exists (pr1 (inv_from_z_iso (make_z_iso' f p))).
split; simpl.
- apply (maponpaths pr1 H').
- apply (maponpaths pr1 H'').
Defined.
Definition inv_algebra_mor_from_is_z_iso {X Y : FunctorAlg F} (f : X --> Y)
: is_z_isomorphism (pr1 f) → (Y --> X).
Proof.
intro T.
set (fiso:=make_z_iso' (pr1 f) T).
set (finv:=inv_from_z_iso fiso).
exists finv.
unfold finv.
apply pathsinv0.
apply z_iso_inv_on_left.
simpl.
rewrite functor_on_inv_from_z_iso.
rewrite <- assoc.
apply pathsinv0.
apply z_iso_inv_on_right.
simpl.
apply (pr2 f).
Defined.
Definition is_algebra_iso_from_is_z_iso {X Y : FunctorAlg F} (f : X --> Y)
: is_z_isomorphism (pr1 f) → is_z_isomorphism f.
Proof.
intro T.
exists (inv_algebra_mor_from_is_z_iso f T).
split; simpl.
- apply algebra_mor_eq.
apply (z_iso_inv_after_z_iso (make_z_iso' (pr1 f) T)).
- apply algebra_mor_eq.
apply (z_iso_after_z_iso_inv (make_z_iso' (pr1 f) T)).
Defined.
Definition algebra_iso_first_z_iso {X Y : FunctorAlg F}
: z_iso X Y ≃ ∑ f : X --> Y, is_z_isomorphism (pr1 f).
Proof.
apply (weqbandf (idweq _ )).
unfold idweq. simpl.
intro f.
apply weqimplimpl.
- apply is_z_iso_from_is_algebra_iso.
- apply is_algebra_iso_from_is_z_iso.
- apply (isaprop_is_z_isomorphism (C:=FunctorAlg F) f).
- apply (isaprop_is_z_isomorphism (pr1 f)).
Defined.
Definition swap (A B : UU) : A × B → B × A.
Proof.
intro ab.
exists (pr2 ab).
exact (pr1 ab).
Defined.
Definition swapweq (A B : UU) : (A × B) ≃ (B × A).
Proof.
exists (swap A B).
apply (isweq_iso _ (swap B A)).
- abstract ( intro ab; destruct ab; apply idpath ).
- abstract ( intro ba; destruct ba; apply idpath ).
Defined.
Definition algebra_z_iso_rearrange {X Y : FunctorAlg F}
: (∑ f : X --> Y, is_z_isomorphism (pr1 f)) ≃ algebra_eq_type X Y.
Proof.
eapply weqcomp.
- apply weqtotal2asstor.
- simpl. unfold algebra_eq_type.
apply invweq.
eapply weqcomp.
+ apply weqtotal2asstor.
+ simpl. apply (weqbandf (idweq _ )).
unfold idweq. simpl.
intro f; apply swapweq.
Defined.
Definition algebra_idtoiso (X Y : FunctorAlg F) :
(X = Y) ≃ z_iso X Y.
Proof.
eapply weqcomp.
- apply algebra_ob_eq.
- eapply weqcomp.
+ apply (invweq (algebra_z_iso_rearrange)).
+ apply (invweq algebra_iso_first_z_iso).
Defined.
Lemma isweq_idtoiso_FunctorAlg (X Y : FunctorAlg F)
: isweq (@idtoiso _ X Y).
Proof.
apply (isweqhomot (algebra_idtoiso X Y)).
- intro p. induction p.
simpl.
apply (z_iso_eq(C:=FunctorAlg F)). apply algebra_mor_eq.
apply idpath.
- apply (pr2 _ ).
Defined.
Lemma is_univalent_FunctorAlg : is_univalent (FunctorAlg F).
Proof.
intros X Y.
apply isweq_idtoiso_FunctorAlg.
Defined.
Lemma idtomor_FunctorAlg_commutes (X Y: FunctorAlg F) (e: X = Y)
: mor_from_algebra_mor F (idtomor _ _ e) = idtomor _ _ (maponpaths (alg_carrier F) e).
Proof.
induction e.
apply idpath.
Qed.
Corollary idtoiso_FunctorAlg_commutes (X Y: FunctorAlg F) (e: X = Y)
: mor_from_algebra_mor F (morphism_from_z_iso _ _ (idtoiso e))
= idtoiso (maponpaths (alg_carrier F) e).
Proof.
unfold morphism_from_z_iso.
rewrite eq_idtoiso_idtomor.
etrans.
2: { apply pathsinv0, eq_idtoiso_idtomor. }
apply idtomor_FunctorAlg_commutes.
Qed.
End FunctorAlg_saturated.
Section Lambeks_lemma.
Variables (C : category) (F : functor C C).
Variables (Aa : FunctorAlg F) (AaIsInitial : isInitial (FunctorAlg F) Aa).
Local Definition AaInitial : Initial (FunctorAlg F) :=
make_Initial _ AaIsInitial.
Local Notation A := (alg_carrier _ Aa).
Local Notation a := (alg_map _ Aa).
Local Definition FAa : FunctorAlg F := tpair (λ X, C ⟦F X,X⟧) (F A) (# F a).
Local Definition Fa' := InitialArrow AaInitial FAa.
Local Definition a' : C⟦A,F A⟧ := mor_from_algebra_mor F Fa'.
Local Definition Ha' := algebra_mor_commutes _ _ _ Fa'.
Lemma initialAlg_is_iso_subproof : is_inverse_in_precat a a'.
Proof.
assert (Ha'a : a' · a = identity A).
{ assert (algMor_a'a : is_algebra_mor _ _ _ (a' · a)).
{ unfold is_algebra_mor, a'; rewrite functor_comp.
eapply pathscomp0; [|eapply cancel_postcomposition; apply Ha'].
apply assoc. }
apply pathsinv0; set (X := tpair _ _ algMor_a'a).
apply (maponpaths pr1 (!@InitialEndo_is_identity _ AaInitial X)).
}
split; trivial.
eapply pathscomp0; [apply Ha'|]; cbn.
rewrite <- functor_comp.
eapply pathscomp0; [eapply maponpaths; apply Ha'a|].
apply functor_id.
Qed.
Lemma initialAlg_is_z_iso : is_z_isomorphism a.
Proof.
exists a'.
exact initialAlg_is_iso_subproof.
Defined.
End Lambeks_lemma.
The natural numbers are intial for X ↦ 1 + X
Section Nats.
Context (C : category).
Context (bc : BinCoproducts C).
Context (hsC : has_homsets C).
Context (T : Terminal C).
Local Notation "1" := T.
Local Notation "f + g" := (BinCoproductOfArrows _ _ _ f g).
Local Notation "[ f , g ]" := (BinCoproductArrow _ _ f g).
Let F : functor C C := BinCoproduct_of_functors _ _ bc
(constant_functor _ _ 1)
(functor_identity _).
F on objects: X ↦ 1 + X
Definition F_compute2 {x y : C} : ∏ f : x --> y, # F f = (identity 1) + f :=
fun c => (idpath _).
Definition nat_ob : UU := Initial (FunctorAlg F).
Definition nat_ob_carrier (N : nat_ob) : ob C :=
alg_carrier _ (InitialObject N).
Local Coercion nat_ob_carrier : nat_ob >-> ob.
fun c => (idpath _).
Definition nat_ob : UU := Initial (FunctorAlg F).
Definition nat_ob_carrier (N : nat_ob) : ob C :=
alg_carrier _ (InitialObject N).
Local Coercion nat_ob_carrier : nat_ob >-> ob.
We have an arrow alg_map : (F N = 1 + N) --> N,
so by the η-rule (UMP) for the coproduct, we can assume that it
arises from a pair of maps nat_ob_z,nat_ob_s by composing with
coproduct injections.
in1 in2 1 ----> 1 + N <---- N | | | nat_ob_z | | alg_map | nat_ob_s | V | +-------> N <-------+
Definition nat_ob_z (N : nat_ob) : (1 --> N) :=
BinCoproductIn1 (bc 1 (alg_carrier F (pr1 N))) · (alg_map _ (pr1 N)).
Definition nat_ob_s (N : nat_ob) : (N --> N) :=
BinCoproductIn2 (bc 1 (alg_carrier F (pr1 N))) · (alg_map _ (pr1 N)).
Local Notation "0" := (nat_ob_z _).
BinCoproductIn1 (bc 1 (alg_carrier F (pr1 N))) · (alg_map _ (pr1 N)).
Definition nat_ob_s (N : nat_ob) : (N --> N) :=
BinCoproductIn2 (bc 1 (alg_carrier F (pr1 N))) · (alg_map _ (pr1 N)).
Local Notation "0" := (nat_ob_z _).
Use the universal property of the coproduct to make any object with a
point and an endomorphism into an F-algebra
Definition make_F_alg {X : ob C} (f : 1 --> X) (g : X --> X) : ob (FunctorAlg F).
Proof.
refine (X,, _).
exact (BinCoproductArrow _ f g).
Defined.
Proof.
refine (X,, _).
exact (BinCoproductArrow _ f g).
Defined.
Using make_F_alg, X will be an F-algebra, and by initiality of N, there will
be a unique morphism of F-algebras N --> X, which can be projected to a
morphism in C.
Definition nat_ob_rec (N : nat_ob) {X : ob C} :
∏ (f : 1 --> X) (g : X --> X), (N --> X) :=
fun f g => mor_from_algebra_mor F (InitialArrow N (make_F_alg f g)).
∏ (f : 1 --> X) (g : X --> X), (N --> X) :=
fun f g => mor_from_algebra_mor F (InitialArrow N (make_F_alg f g)).
When calling the recursor on 0, you get the base case.
Specifically,
nat_ob_z · nat_ob_rec = f
Lemma nat_ob_rec_z (N : nat_ob) {X : ob C} :
∏ (f : 1 --> X) (g : X --> X), nat_ob_z N · nat_ob_rec N f g = f.
Proof.
intros f g.
pose (inlN := BinCoproductIn1 (bc 1 N)).
pose (succ := nat_ob_s N).
∏ (f : 1 --> X) (g : X --> X), nat_ob_z N · nat_ob_rec N f g = f.
Proof.
intros f g.
pose (inlN := BinCoproductIn1 (bc 1 N)).
pose (succ := nat_ob_s N).
By initiality of N, there is a unique morphism making the following
diagram commute:
This proof uses somewhat idiosyncratic "forward reasoning", transforming
the term "diagram" rather than the goal.
inlN identity 1 + nat_ob_rec 1 -----> 1 + N -------------------------> 1 + X | | alg_map N | | alg_map X V V N --------------------------> X nat_ob_rec
pose
(diagram :=
maponpaths
(fun x => inlN · x)
(algebra_mor_commutes F (pr1 N) _ (InitialArrow N (make_F_alg f g)))).
rewrite (F_compute2 _) in diagram.
(diagram :=
maponpaths
(fun x => inlN · x)
(algebra_mor_commutes F (pr1 N) _ (InitialArrow N (make_F_alg f g)))).
rewrite (F_compute2 _) in diagram.
Using the β-rules for coproducts, we can simplify some of the terms (identity 1 + _) · f, g --β--> identity 1 · f, _ · g
We can dispense with the identity
rewrite (id_left _) in diagram.
rewrite assoc in diagram.
rewrite (BinCoproductArrowEta C 1 N (bc _ _) _ _) in diagram.
refine (_ @ (BinCoproductIn1Commutes C _ _ (bc 1 _) _ f g)).
rewrite (!BinCoproductIn1Commutes C _ _ (bc 1 _) _ 0 succ).
unfold nat_ob_rec in *.
exact diagram.
Defined.
Opaque nat_ob_rec_z.
rewrite assoc in diagram.
rewrite (BinCoproductArrowEta C 1 N (bc _ _) _ _) in diagram.
refine (_ @ (BinCoproductIn1Commutes C _ _ (bc 1 _) _ f g)).
rewrite (!BinCoproductIn1Commutes C _ _ (bc 1 _) _ 0 succ).
unfold nat_ob_rec in *.
exact diagram.
Defined.
Opaque nat_ob_rec_z.
The succesor case:
nat_ob_s · nat_ob_rec = nat_ob_rec · g
The proof is very similar.
Lemma nat_ob_rec_s (N : nat_ob) {X : ob C} :
∏ (f : 1 --> X) (g : X --> X),
nat_ob_s N · nat_ob_rec N f g = nat_ob_rec N f g · g.
Proof.
intros f g.
pose (inrN := BinCoproductIn2 (bc 1 N)).
pose (succ := nat_ob_s N).
∏ (f : 1 --> X) (g : X --> X),
nat_ob_s N · nat_ob_rec N f g = nat_ob_rec N f g · g.
Proof.
intros f g.
pose (inrN := BinCoproductIn2 (bc 1 N)).
pose (succ := nat_ob_s N).
By initiality of N, there is a unique morphism making the same diagram
commute as above, but with "inrN" in place of "inlN".
pose
(diagram :=
maponpaths
(fun x => inrN · x)
(algebra_mor_commutes F (pr1 N) _ (InitialArrow N (make_F_alg f g)))).
rewrite (F_compute2 _) in diagram.
rewrite (BinCoproductArrowEta C 1 X (bc _ _) _ _) in diagram.
(diagram :=
maponpaths
(fun x => inrN · x)
(algebra_mor_commutes F (pr1 N) _ (InitialArrow N (make_F_alg f g)))).
rewrite (F_compute2 _) in diagram.
rewrite (BinCoproductArrowEta C 1 X (bc _ _) _ _) in diagram.
Using the β-rules for coproducts, we can simplify some of the terms (identity 1 + _) · f, g --β--> identity 1 · f, _ · g
rewrite (BinCoproductIn2Commutes C 1 N (bc 1 _) _ _ _) in diagram.
rewrite assoc in diagram.
rewrite (BinCoproductArrowEta C 1 N (bc _ _) _ _) in diagram.
refine
(_ @ maponpaths (fun x => nat_ob_rec N f g · x)
(BinCoproductIn2Commutes C _ _ (bc 1 _) _ f g)).
rewrite (!BinCoproductIn2Commutes C _ _ (bc 1 _) _ 0 (nat_ob_s N)).
unfold nat_ob_rec in *.
exact diagram.
Defined.
Opaque nat_ob_rec_s.
End Nats.
rewrite assoc in diagram.
rewrite (BinCoproductArrowEta C 1 N (bc _ _) _ _) in diagram.
refine
(_ @ maponpaths (fun x => nat_ob_rec N f g · x)
(BinCoproductIn2Commutes C _ _ (bc 1 _) _ f g)).
rewrite (!BinCoproductIn2Commutes C _ _ (bc 1 _) _ 0 (nat_ob_s N)).
unfold nat_ob_rec in *.
exact diagram.
Defined.
Opaque nat_ob_rec_s.
End Nats.
nat_ob implies NNO
Lemma nat_ob_NNO {C : category} (BC : BinCoproducts C) (hsC : has_homsets C) (TC : Terminal C) :
nat_ob _ BC TC → NNO TC.
Proof.
intros N.
use make_NNO.
- exact (nat_ob_carrier _ _ _ N).
- apply nat_ob_z.
- apply nat_ob_s.
- intros n z s.
use unique_exists.
+ apply (nat_ob_rec _ _ _ _ z s).
+ split; [ apply nat_ob_rec_z | apply nat_ob_rec_s ].
+ intros x; apply isapropdirprod; apply hsC.
+ intros x [H1 H2].
transparent assert (xalg : (FunctorAlg (BinCoproduct_of_functors C C BC
(constant_functor C C TC)
(functor_identity C))
⟦ InitialObject N, make_F_alg C BC TC z s ⟧)).
{ refine (x,,_).
abstract (apply pathsinv0; etrans; [apply precompWithBinCoproductArrow |];
rewrite id_left, <- H1;
etrans; [eapply maponpaths, pathsinv0, H2|];
now apply pathsinv0, BinCoproductArrowUnique; rewrite assoc;
apply maponpaths).
}
exact (maponpaths pr1 (InitialArrowUnique N (make_F_alg C BC TC z s) xalg)).
Defined.
Section AlgebrasForBinaryCoproducts.
Context {C : category} (BC : BinCoproducts C) (F1 F2 : functor C C).
Let F : functor C C := BinCoproduct_of_functors _ _ BC F1 F2.
Local Coercion alg_carrier : algebra_ob >-> ob.
Definition tau1_from_alg (X : algebra_ob F) : F1 X --> X.
Proof.
exact (BinCoproductIn1 (BC _ _) · alg_map _ X).
Defined.
Local Notation τ1 := tau1_from_alg.
Definition tau2_from_alg (X : algebra_ob F) : F2 X --> X.
Proof.
exact (BinCoproductIn2 (BC _ _) · alg_map _ X).
Defined.
Local Notation τ2 := tau2_from_alg.
Definition AlgebrasForBincoproductIntoProductcategory_data :
functor_data (FunctorAlg F) (category_binproduct (FunctorAlg F1) (FunctorAlg F2)).
Proof.
use make_functor_data.
- intro X.
split.
+ exists (pr1 X). exact (τ1 X).
+ exists (pr1 X). exact (τ2 X).
- intros X X' f.
split; cbn.
+ exists (pr1 f).
abstract (unfold τ1;
rewrite assoc';
etrans; [apply maponpaths; apply (pr2 f) |];
rewrite assoc;
etrans;
[apply cancel_postcomposition; apply BinCoproductIn1Commutes |];
apply assoc').
+ exists (pr1 f).
abstract (unfold τ2;
rewrite assoc';
etrans; [apply maponpaths; apply (pr2 f) |];
rewrite assoc;
etrans;
[apply cancel_postcomposition; apply BinCoproductIn2Commutes |];
apply assoc').
Defined.
Lemma AlgebrasForBincoproductIntoProductcategory_is_functor :
is_functor AlgebrasForBincoproductIntoProductcategory_data.
Proof.
split; red.
- intro X.
use dirprodeq;
(use mor_eq_total_category_when_locally_prop; [ apply is_locally_propositional_algebra_disp_cat | apply idpath]).
- intros X Y Z f g.
use dirprodeq;
(use mor_eq_total_category_when_locally_prop; [ apply is_locally_propositional_algebra_disp_cat | apply idpath]).
Qed.
Definition AlgebrasForBincoproductIntoProductcategory :
functor (FunctorAlg F) (category_binproduct (FunctorAlg F1 ) (FunctorAlg F2))
:= _ ,, AlgebrasForBincoproductIntoProductcategory_is_functor.
End AlgebrasForBinaryCoproducts.
Section AlgebrasForCoproducts.
Context (I : UU) {C : category} (HC : Coproducts I C) (F : I -> functor C C).
Let sF : functor C C := coproduct_of_functors _ _ _ HC F.
Local Coercion alg_carrier : algebra_ob >-> ob.
Definition taui_from_alg (X : algebra_ob sF) (i : I) : F i X --> X.
Proof.
exact (CoproductIn _ _ (HC _) i · alg_map _ X).
Defined.
Definition AlgebrasForCoproductIntoProductcategory_data :
functor_data (FunctorAlg sF) (product_category (fun i => FunctorAlg (F i))).
Proof.
use make_functor_data.
- intro X.
intro i.
exists (pr1 X). exact (taui_from_alg X i).
- intros X X' f.
intro i.
exists (pr1 f).
abstract (unfold taui_from_alg;
cbn;
rewrite assoc';
etrans; [apply maponpaths; apply (pr2 f) |];
rewrite assoc;
etrans;
[ apply cancel_postcomposition;
apply (CoproductInCommutes I C (λ i0 : I, F i0 (pr1 X)) (HC (λ i0 : I, F i0 (pr1 X)))) |];
apply assoc').
Defined.
Lemma AlgebrasForCoproductIntoProductcategory_is_functor :
is_functor AlgebrasForCoproductIntoProductcategory_data.
Proof.
split; red.
- intro X.
apply funextsec.
intro i.
transparent assert (xx : (ob (total_precategory (algebra_disp_cat (F i))))).
{ exists (pr1 X). exact (taui_from_alg X i). }
apply (mor_eq_total_category_when_locally_prop(D:=algebra_disp_cat (F i)) (is_locally_propositional_algebra_disp_cat (F i)) (xx:=xx)(yy:=xx)).
apply idpath.
- intros X Y Z f g.
apply funextsec.
intro i.
transparent assert (xx : (ob (total_precategory (algebra_disp_cat (F i))))).
{ exists (pr1 X). exact (taui_from_alg X i). }
transparent assert (zz : (ob (total_precategory (algebra_disp_cat (F i))))).
{ exists (pr1 Z). exact (taui_from_alg Z i). }
apply (mor_eq_total_category_when_locally_prop(D:=algebra_disp_cat (F i)) (is_locally_propositional_algebra_disp_cat (F i)) (xx:=xx)(yy:=zz)).
apply idpath.
Qed.
Definition AlgebrasForCoproductIntoProductcategory :
functor (FunctorAlg sF) (product_category (fun i => FunctorAlg (F i)))
:= _ ,, AlgebrasForCoproductIntoProductcategory_is_functor.
End AlgebrasForCoproducts.
nat_ob _ BC TC → NNO TC.
Proof.
intros N.
use make_NNO.
- exact (nat_ob_carrier _ _ _ N).
- apply nat_ob_z.
- apply nat_ob_s.
- intros n z s.
use unique_exists.
+ apply (nat_ob_rec _ _ _ _ z s).
+ split; [ apply nat_ob_rec_z | apply nat_ob_rec_s ].
+ intros x; apply isapropdirprod; apply hsC.
+ intros x [H1 H2].
transparent assert (xalg : (FunctorAlg (BinCoproduct_of_functors C C BC
(constant_functor C C TC)
(functor_identity C))
⟦ InitialObject N, make_F_alg C BC TC z s ⟧)).
{ refine (x,,_).
abstract (apply pathsinv0; etrans; [apply precompWithBinCoproductArrow |];
rewrite id_left, <- H1;
etrans; [eapply maponpaths, pathsinv0, H2|];
now apply pathsinv0, BinCoproductArrowUnique; rewrite assoc;
apply maponpaths).
}
exact (maponpaths pr1 (InitialArrowUnique N (make_F_alg C BC TC z s) xalg)).
Defined.
Section AlgebrasForBinaryCoproducts.
Context {C : category} (BC : BinCoproducts C) (F1 F2 : functor C C).
Let F : functor C C := BinCoproduct_of_functors _ _ BC F1 F2.
Local Coercion alg_carrier : algebra_ob >-> ob.
Definition tau1_from_alg (X : algebra_ob F) : F1 X --> X.
Proof.
exact (BinCoproductIn1 (BC _ _) · alg_map _ X).
Defined.
Local Notation τ1 := tau1_from_alg.
Definition tau2_from_alg (X : algebra_ob F) : F2 X --> X.
Proof.
exact (BinCoproductIn2 (BC _ _) · alg_map _ X).
Defined.
Local Notation τ2 := tau2_from_alg.
Definition AlgebrasForBincoproductIntoProductcategory_data :
functor_data (FunctorAlg F) (category_binproduct (FunctorAlg F1) (FunctorAlg F2)).
Proof.
use make_functor_data.
- intro X.
split.
+ exists (pr1 X). exact (τ1 X).
+ exists (pr1 X). exact (τ2 X).
- intros X X' f.
split; cbn.
+ exists (pr1 f).
abstract (unfold τ1;
rewrite assoc';
etrans; [apply maponpaths; apply (pr2 f) |];
rewrite assoc;
etrans;
[apply cancel_postcomposition; apply BinCoproductIn1Commutes |];
apply assoc').
+ exists (pr1 f).
abstract (unfold τ2;
rewrite assoc';
etrans; [apply maponpaths; apply (pr2 f) |];
rewrite assoc;
etrans;
[apply cancel_postcomposition; apply BinCoproductIn2Commutes |];
apply assoc').
Defined.
Lemma AlgebrasForBincoproductIntoProductcategory_is_functor :
is_functor AlgebrasForBincoproductIntoProductcategory_data.
Proof.
split; red.
- intro X.
use dirprodeq;
(use mor_eq_total_category_when_locally_prop; [ apply is_locally_propositional_algebra_disp_cat | apply idpath]).
- intros X Y Z f g.
use dirprodeq;
(use mor_eq_total_category_when_locally_prop; [ apply is_locally_propositional_algebra_disp_cat | apply idpath]).
Qed.
Definition AlgebrasForBincoproductIntoProductcategory :
functor (FunctorAlg F) (category_binproduct (FunctorAlg F1 ) (FunctorAlg F2))
:= _ ,, AlgebrasForBincoproductIntoProductcategory_is_functor.
End AlgebrasForBinaryCoproducts.
Section AlgebrasForCoproducts.
Context (I : UU) {C : category} (HC : Coproducts I C) (F : I -> functor C C).
Let sF : functor C C := coproduct_of_functors _ _ _ HC F.
Local Coercion alg_carrier : algebra_ob >-> ob.
Definition taui_from_alg (X : algebra_ob sF) (i : I) : F i X --> X.
Proof.
exact (CoproductIn _ _ (HC _) i · alg_map _ X).
Defined.
Definition AlgebrasForCoproductIntoProductcategory_data :
functor_data (FunctorAlg sF) (product_category (fun i => FunctorAlg (F i))).
Proof.
use make_functor_data.
- intro X.
intro i.
exists (pr1 X). exact (taui_from_alg X i).
- intros X X' f.
intro i.
exists (pr1 f).
abstract (unfold taui_from_alg;
cbn;
rewrite assoc';
etrans; [apply maponpaths; apply (pr2 f) |];
rewrite assoc;
etrans;
[ apply cancel_postcomposition;
apply (CoproductInCommutes I C (λ i0 : I, F i0 (pr1 X)) (HC (λ i0 : I, F i0 (pr1 X)))) |];
apply assoc').
Defined.
Lemma AlgebrasForCoproductIntoProductcategory_is_functor :
is_functor AlgebrasForCoproductIntoProductcategory_data.
Proof.
split; red.
- intro X.
apply funextsec.
intro i.
transparent assert (xx : (ob (total_precategory (algebra_disp_cat (F i))))).
{ exists (pr1 X). exact (taui_from_alg X i). }
apply (mor_eq_total_category_when_locally_prop(D:=algebra_disp_cat (F i)) (is_locally_propositional_algebra_disp_cat (F i)) (xx:=xx)(yy:=xx)).
apply idpath.
- intros X Y Z f g.
apply funextsec.
intro i.
transparent assert (xx : (ob (total_precategory (algebra_disp_cat (F i))))).
{ exists (pr1 X). exact (taui_from_alg X i). }
transparent assert (zz : (ob (total_precategory (algebra_disp_cat (F i))))).
{ exists (pr1 Z). exact (taui_from_alg Z i). }
apply (mor_eq_total_category_when_locally_prop(D:=algebra_disp_cat (F i)) (is_locally_propositional_algebra_disp_cat (F i)) (xx:=xx)(yy:=zz)).
apply idpath.
Qed.
Definition AlgebrasForCoproductIntoProductcategory :
functor (FunctorAlg sF) (product_category (fun i => FunctorAlg (F i)))
:= _ ,, AlgebrasForCoproductIntoProductcategory_is_functor.
End AlgebrasForCoproducts.