feat(precategory): add composition of nat. trans. with functor

This commit is contained in:
Floris van Doorn 2015-04-19 16:03:52 -04:00 committed by Leonardo de Moura
parent a79a3043ed
commit b86ee9dfa6
2 changed files with 22 additions and 2 deletions

View file

@ -150,7 +150,7 @@ namespace iso
namespace iso
attribute to_hom [coercion]
definition MK (f : a ⟶ b) (g : b ⟶ a) (H1 : g ∘ f = id) (H2 : f ∘ g = id) :=
protected definition MK (f : a ⟶ b) (g : b ⟶ a) (H1 : g ∘ f = id) (H2 : f ∘ g = id) :=
@mk _ _ _ _ f (is_iso.mk H1 H2)
definition to_inv (f : a ≅ b) : b ⟶ a :=

View file

@ -15,7 +15,7 @@ structure nat_trans {C D : Precategory} (F G : C ⇒ D) :=
namespace nat_trans
infixl `⟹`:25 := nat_trans -- \==>
variables {C D E : Precategory} {F G H I : C ⇒ D} {F' G' : D ⇒ E}
variables {B C D E : Precategory} {F G H I : C ⇒ D} {F' G' : D ⇒ E}
attribute natural_map [coercion]
@ -108,6 +108,26 @@ namespace nat_trans
: (θ ∘nf G) ∘n (F' ∘fn η) = (G' ∘fn η) ∘n (θ ∘nf F) :=
nat_trans_eq_mk (λc, (naturality θ (η c))⁻¹)
definition fn_n_distrib (F' : D ⇒ E) (η : G ⟹ H) (θ : F ⟹ G)
: F' ∘fn (η ∘n θ) = (F' ∘fn η) ∘n (F' ∘fn θ) :=
nat_trans_eq_mk (λc, !respect_comp)
definition n_nf_distrib (η : G ⟹ H) (θ : F ⟹ G) (F' : B ⇒ C)
: (η ∘n θ) ∘nf F' = (η ∘nf F') ∘n (θ ∘nf F') :=
nat_trans_eq_mk (λc, idp)
definition fn_id (F' : D ⇒ E) : F' ∘fn nat_trans.ID F = nat_trans.id :=
nat_trans_eq_mk (λc, !respect_id)
definition id_nf (F' : B ⇒ C) : nat_trans.ID F ∘nf F' = nat_trans.id :=
nat_trans_eq_mk (λc, idp)
definition id_fn (η : G ⟹ H) (c : C) : (functor.id ∘fn η) c = η c :=
idp
definition nf_id (η : G ⟹ H) (c : C) : (η ∘nf functor.id) c = η c :=
idp
definition nat_trans_of_eq [reducible] (p : F = G) : F ⟹ G :=
nat_trans.mk (λc, hom_of_eq (ap010 to_fun_ob p c))
(λa b f, eq.rec_on p (!id_right ⬝ !id_left⁻¹))