X-Git-Url: http://git.megacz.com/?p=coq-categories.git;a=blobdiff_plain;f=src%2FIsomorphisms_ch1_5.v;h=422d4bbf75c4dfb0c8a138ad22f238a6f2f7d538;hp=65eaf4624265e65d2d567e8da5b31949cbce6bc7;hb=57dfee583c84915cb6b3aadfa1dba692681499b3;hpb=e928451c4c45cdbdd975bbfb229e8cc2616b8194 diff --git a/src/Isomorphisms_ch1_5.v b/src/Isomorphisms_ch1_5.v index 65eaf46..422d4bb 100644 --- a/src/Isomorphisms_ch1_5.v +++ b/src/Isomorphisms_ch1_5.v @@ -1,6 +1,5 @@ Generalizable All Variables. -Require Import Preamble. -Require Import General. +Require Import Notations. Require Import Categories_ch1_3. Require Import Functors_ch1_4. @@ -53,6 +52,7 @@ Definition iso_comp `{C:Category}{a b c:C}(i1:Isomorphic a b)(i2:Isomorphic b c) [ setoid_rewrite (iso_comp1 i1) | setoid_rewrite (iso_comp2 i2) ]; reflexivity. Defined. +Notation "a >>≅>> b" := (iso_comp a b). Definition functors_preserve_isos `{C1:Category}`{C2:Category}{Fo}(F:Functor C1 C2 Fo){a b:C1}(i:Isomorphic a b) : Isomorphic (F a) (F b). @@ -118,3 +118,32 @@ Lemma iso_inv_inv `{C:Category}{a}{b}(i:a ≅ b) : #(i⁻¹)⁻¹ ~~ #i. unfold iso_inv; simpl. reflexivity. Qed. + +(* the next four lemmas are handy for setoid_rewrite; they let you avoid having to get the associativities right *) +Lemma iso_comp2_right : forall `{C:Category}{a b}(i:a≅b) c (g:b~>c), iso_backward i >>> (iso_forward i >>> g) ~~ g. + intros. + setoid_rewrite <- associativity. + setoid_rewrite iso_comp2. + apply left_identity. + Qed. + +Lemma iso_comp2_left : forall `{C:Category}{a b}(i:a≅b) c (g:c~>b), (g >>> iso_backward i) >>> iso_forward i ~~ g. + intros. + setoid_rewrite associativity. + setoid_rewrite iso_comp2. + apply right_identity. + Qed. + +Lemma iso_comp1_right : forall `{C:Category}{a b}(i:a≅b) c (g:a~>c), iso_forward i >>> (iso_backward i >>> g) ~~ g. + intros. + setoid_rewrite <- associativity. + setoid_rewrite iso_comp1. + apply left_identity. + Qed. + +Lemma iso_comp1_left : forall `{C:Category}{a b}(i:a≅b) c (g:c~>a), (g >>> iso_forward i) >>> iso_backward i ~~ g. + intros. + setoid_rewrite associativity. + setoid_rewrite iso_comp1. + apply right_identity. + Qed.