]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/apply_functor.ma
...
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / apply_functor.ma
index 7b10b4b9fc98f05bd0736344545809630c9f89c4..e52a420745a16ee725726d3a7afad5aed761af78 100644 (file)
@@ -96,3 +96,27 @@ constructor 1;
 | 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.
+
+
+