]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/apply_functor.ma
New version.
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / apply_functor.ma
index e10bd351c85fc45c8598da8ebfd2c6ccc228d3e1..e52a420745a16ee725726d3a7afad5aed761af78 100644 (file)
@@ -21,9 +21,17 @@ record Fo (C1,C2:CAT2) (F:arrows3 CAT2 C1 C2) : Type2 ≝ {
   FP: map_objs2 ?? F F1 =_\ID F2
 }.
 
+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). 
+
+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). 
+
 lemma REW : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.∀X,Y:Fo ?? F.
-  arrows2 C2 (F (F1 ??? X)) (F (F1 ??? Y)) → 
-  arrows2 C2 (F2 ??? X) (F2 ??? Y).           
+  arrows2 C2 (F (ℱ_1 X)) (F (ℱ_1 Y)) → 
+  arrows2 C2 (ℱ_2 X) (ℱ_2 Y).           
 intros 5; cases X; cases Y; clear X Y; 
 cases H; cases H1; intros; assumption;
 qed.           
@@ -34,6 +42,14 @@ record Fm_c (C1,C2:CAT2) (F:arrows3 CAT2 C1 C2) (X,Y:Fo ?? F) : Type2 ≝ {
   FmP: REW ?? F X Y (map_arrows2 ?? F ?? Fm1) = Fm2
 }.
 
+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). 
+
+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). 
+
 definition Fm : 
  ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.
    Fo ?? F → Fo ?? F → setoid2. 
@@ -57,16 +73,14 @@ definition F_comp :
     binary_morphism2 (Fm ?? F o1 o2) (Fm ?? F o2 o3) (Fm ?? F o1 o3).
 intros; constructor 1;
 [ intros (f g); constructor 1;
-    [ apply (comp2 C2 ??? (Fm2 ????? f) (Fm2 ????? g));
-    | apply (comp2 C1 ??? (Fm1 ????? f) (Fm1 ????? g));
+    [ apply (comp2 C2 ??? (ℳ_2 f) (ℳ_2 g));
+    | apply (comp2 C1 ??? (ℳ_1 f) (ℳ_1 g));
     | apply hide; cases o1 in f; cases o2 in g; cases o3; clear o1 o2 o3;
       cases H; cases H1; cases H2; intros 2; cases c; cases c1; clear c c1;
       simplify; apply (.= (respects_comp2:?)); apply (e1‡e);]
-| intros 6; change with 
-    ((Fm2 C1 C2 F o2 o3 b∘Fm2 C1 C2 F o1 o2 a) = 
-    (Fm2 C1 C2 F o2 o3 b'∘Fm2 C1 C2 F o1 o2 a'));
-  change in e1 with (Fm2 ?? F ?? b = Fm2 ?? F ?? b');
-  change in e with (Fm2 ?? F ?? a = Fm2 ?? F ?? a');
+| intros 6; change with ((ℳ_2 b ∘ ℳ_2 a) = (ℳ_2 b' ∘ ℳ_2 a'));
+  change in e1 with (ℳ_2 b = ℳ_2 b');
+  change in e with (ℳ_2 a = ℳ_2 a');
   apply (e‡e1);]
 qed.
 
@@ -78,7 +92,31 @@ constructor 1;
 | apply (Fm ?? F); 
 | apply F_id; 
 | apply F_comp;
-| intros; apply (comp_assoc2 C2 ???? (Fm2 ????? a12) (Fm2 ????? a23) (Fm2 ????? a34));
-| intros; apply (id_neutral_right2 C2 ?? (Fm2 ????? a));
-| intros; apply (id_neutral_left2 C2 ?? (Fm2 ????? a));]
+| intros; apply (comp_assoc2 C2 ???? (ℳ_2 a12) (ℳ_2 a23) (ℳ_2 a34));
+| intros; apply (id_neutral_right2 C2 ?? (ℳ_2 a));
+| intros; apply (id_neutral_left2 C2 ?? (ℳ_2 a));]
 qed.
+
+definition faithful ≝  
+   λC1,C2.λF:arrows3 CAT2 C1 C2.∀S,T.∀f,g:arrows2 C1 S T.
+     map_arrows2 ?? F ?? f = map_arrows2 ?? F ?? g → f=g.
+
+definition Ylppa : ∀C1,C2: CAT2.∀F:arrows3 CAT2 C1 C2.
+  faithful ?? F →  let rC2 ≝ Apply ?? F in arrows3 CAT2 rC2 C1.
+intros; constructor 1;
+[ intro; apply (ℱ_1 o);
+| intros; constructor 1; 
+  [ intros; apply (ℳ_1 c);
+  | apply hide; intros; apply f;  change in e with (ℳ_2 a = ℳ_2 a');
+    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;]
+    clear H1 H2 e; cases S in a a' Hcut; cases T;
+    cases H; cases H1; simplify; intros; assumption;]
+| intro; apply rule #;
+| intros; simplify; apply rule #;]
+qed.
+
+
+