add inverse form of ni_commutes
[coq-categories.git] / src / NaturalIsomorphisms_ch7_5.v
index f82dba8..ce97ce6 100644 (file)
@@ -1,5 +1,5 @@
 Generalizable All Variables.
-Require Import Preamble.
+Require Import Notations.
 Require Import Categories_ch1_3.
 Require Import Functors_ch1_4.
 Require Import Isomorphisms_ch1_5.
@@ -19,6 +19,17 @@ Implicit Arguments ni_commutes [Ob Hom Ob0 Hom0 C1 C2 Fobj1 Fobj2 F1 F2 A B].
 Coercion ni_iso : NaturalIsomorphism >-> Funclass.
 Notation "F <~~~> G" := (@NaturalIsomorphism _ _ _ _ _ _ _ _ F G) : category_scope.
 
+(* same as ni_commutes, but phrased in terms of inverses *)
+Lemma ni_commutes' `(ni:NaturalIsomorphism) : forall `(f:A~>B), F2 \ f >>> #(ni_iso ni B)⁻¹ ~~ #(ni_iso ni A)⁻¹ >>> F1 \ f.
+  intros.
+  apply iso_shift_right'.
+  setoid_rewrite <- associativity.
+  symmetry.
+  apply iso_shift_left'.
+  symmetry.
+  apply ni_commutes.
+  Qed.
+
 (* FIXME: Lemma 7.11: natural isos are natural transformations in which every morphism is an iso *)
 
 (* every natural iso is invertible, and that inverse is also a natural iso *)
@@ -78,6 +89,35 @@ Definition ni_comp `{C:Category}`{D:Category}
              reflexivity).
    Defined.
 
+Definition ni_respects1
+  `{A:Category}`{B:Category}
+  {Fobj}(F:Functor A B Fobj)
+  `{C:Category}
+  {G0obj}(G0:Functor B C G0obj)
+  {G1obj}(G1:Functor B C G1obj)
+  : (G0 <~~~> G1) -> ((F >>>> G0) <~~~> (F >>>> G1)).
+  intro phi.
+  destruct phi as [ phi_niso phi_comm ].
+  refine {| ni_iso := (fun a => phi_niso (Fobj a)) |}.
+  intros; simpl; apply phi_comm.
+  Defined.
+
+Definition ni_respects2
+  `{A:Category}`{B:Category}
+  {F0obj}(F0:Functor A B F0obj)
+  {F1obj}(F1:Functor A B F1obj)
+  `{C:Category}
+  {Gobj}(G:Functor B C Gobj)
+  : (F0 <~~~> F1) -> ((F0 >>>> G) <~~~> (F1 >>>> G)).
+  intro phi.
+  destruct phi as [ phi_niso phi_comm ].
+  refine {| ni_iso := fun a => (@functors_preserve_isos _ _ _ _ _ _ _ G) _ _ (phi_niso a) |}.
+  intros; simpl.
+  setoid_rewrite fmor_preserves_comp.
+  apply fmor_respects.
+  apply phi_comm.
+  Defined.
+
 Definition ni_respects
   `{A:Category}`{B:Category}
   {F0obj}(F0:Functor A B F0obj)