]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/apply_functor.ma
some notation for map_arrows2
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / apply_functor.ma
index e52a420745a16ee725726d3a7afad5aed761af78..81cf6d8e216b25aadc308877e97eeddfb83861bf 100644 (file)
@@ -23,11 +23,11 @@ record Fo (C1,C2:CAT2) (F:arrows3 CAT2 C1 C2) : Type2 ≝ {
 
 notation "ℱ\sub 1 x" non associative with precedence 60 for @{'F1 $x}.
 notation > "ℱ_1" non associative with precedence 90 for @{F1 ???}.
-interpretation "F1" 'F1 x = (F1 ___ x). 
+interpretation "F1" 'F1 x = (F1 ??? x). 
 
 notation "ℱ\sub 2 x" non associative with precedence 60 for @{'F2 $x}.
 notation > "ℱ_2" non associative with precedence 90 for @{F2 ???}.
-interpretation "F2" 'F2 x = (F2 ___ x). 
+interpretation "F2" 'F2 x = (F2 ??? x). 
 
 lemma REW : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀X,Y:Fo ?? F.
   arrows2 C2 (F (ℱ_1 X)) (F (ℱ_1 Y)) → 
@@ -44,11 +44,11 @@ record Fm_c (C1,C2:CAT2) (F:arrows3 CAT2 C1 C2) (X,Y:Fo ?? F) : Type2 ≝ {
 
 notation "ℳ\sub 1 x" non associative with precedence 60 for @{'Fm1 $x}.
 notation > "ℳ_1" non associative with precedence 90 for @{Fm1 ?????}.
-interpretation "Fm1" 'Fm1 x = (Fm1 _____ x). 
+interpretation "Fm1" 'Fm1 x = (Fm1 ????? x). 
 
 notation "ℳ\sub 2 x" non associative with precedence 60 for @{'Fm2 $x}.
 notation > "ℳ_2" non associative with precedence 90 for @{Fm2 ?????}.
-interpretation "Fm2" 'Fm2 x = (Fm2 _____ x). 
+interpretation "Fm2" 'Fm2 x = (Fm2 ????? x). 
 
 definition Fm : 
  ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.
@@ -70,7 +70,7 @@ qed.
 
 definition F_comp : 
   ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀o1,o2,o3.
-    binary_morphism2 (Fm ?? F o1 o2) (Fm ?? F o2 o3) (Fm ?? F o1 o3).
+    (Fm ?? F o1 o2) × (Fm ?? F o2 o3) ⇒_2 (Fm ?? F o1 o3).
 intros; constructor 1;
 [ intros (f g); constructor 1;
     [ apply (comp2 C2 ??? (ℳ_2 f) (ℳ_2 g));
@@ -111,7 +111,7 @@ intros; constructor 1;
     lapply (FmP ????? a) as H1; lapply (FmP ????? a') as H2;
     cut (REW ????? (map_arrows2 ?? F ?? (ℳ_1 a)) = 
          REW ????? (map_arrows2 ?? F ?? (ℳ_1 a')));[2:
-      apply (.= H1); apply (.= e); apply H2^-1;]
+      apply (.= H1); apply (.= e); apply (H2^-1);]
     clear H1 H2 e; cases S in a a' Hcut; cases T;
     cases H; cases H1; simplify; intros; assumption;]
 | intro; apply rule #;