]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/formal_topology/formal_topologies.ma
update in groud_2 and models
[helm.git] / helm / software / matita / library / formal_topology / formal_topologies.ma
index ba2d65ba0cb32dfe84fd3d47f248eb1289652930..177eb454e491611f5401e4cea853e4bdd9b00d8b 100644 (file)
 
 include "formal_topology/basic_topologies.ma".
 
-definition btop_carr: BTop → Type ≝ λo:BTop. carr (carrbt o).
-
-coercion btop_carr.
-
-definition btop_carr': BTop → setoid ≝ λo:BTop. carrbt o.
-
-coercion btop_carr'.
-
+(*
 definition downarrow: ∀S:BTop. unary_morphism (Ω \sup S) (Ω \sup S).
  intros; constructor 1;
   [ apply (λU:Ω \sup S.{a | ∃b:carrbt S. b ∈ U ∧ a ∈ A ? (singleton ? b)});
@@ -31,7 +24,7 @@ definition downarrow: ∀S:BTop. unary_morphism (Ω \sup S) (Ω \sup S).
     try assumption; [ apply (. #‡H) | apply (. #‡H \sup -1)] assumption]
 qed.
 
-interpretation "downarrow" 'downarrow a = (fun_1 __ (downarrow _) a).
+interpretation "downarrow" 'downarrow a = (fun_1 ?? (downarrow ?) a).
 
 definition ffintersects: ∀S:BTop. binary_morphism1 (Ω \sup S) (Ω \sup S) (Ω \sup S).
  intros; constructor 1;
@@ -39,16 +32,13 @@ definition ffintersects: ∀S:BTop. binary_morphism1 (Ω \sup S) (Ω \sup S) (Ω
   | intros; apply (.= (†H)‡(†H1)); apply refl1]
 qed.
 
-interpretation "ffintersects" 'fintersects U V = (fun1 ___ (ffintersects _) U V).
+interpretation "ffintersects" 'fintersects U V = (fun1 ??? (ffintersects ?) U V).
 
 record formal_topology: Type ≝
  { bt:> BTop;
    converges: ∀U,V: Ω \sup bt. A ? (U ↓ V) = A ? U ∩ A ? V
  }.
 
-definition bt': formal_topology → basic_topology ≝ λo:formal_topology.bt o.
-
-coercion bt'.
 
 definition ffintersects': ∀S:BTop. binary_morphism1 S S (Ω \sup S).
  intros; constructor 1;
@@ -56,7 +46,7 @@ definition ffintersects': ∀S:BTop. binary_morphism1 S S (Ω \sup S).
   | intros; apply (.= (†H)‡(†H1)); apply refl1]
 qed.
 
-interpretation "ffintersects'" 'fintersects U V = (fun1 ___ (ffintersects' _) U V).
+interpretation "ffintersects'" 'fintersects U V = (fun1 ??? (ffintersects' ?) U V).
 
 record formal_map (S,T: formal_topology) : Type ≝
  { cr:> continuous_relation_setoid S T;
@@ -64,11 +54,6 @@ record formal_map (S,T: formal_topology) : Type ≝
    C2: extS ?? cr T = S
  }.
 
-definition cr': ∀FT1,FT2.formal_map FT1 FT2 → continuous_relation FT1 FT2 ≝
- λFT1,FT2,c. cr FT1 FT2 c.
-
-coercion cr'.
-
 definition formal_map_setoid: formal_topology → formal_topology → setoid1.
  intros (S T); constructor 1;
   [ apply (formal_map S T);
@@ -79,17 +64,9 @@ definition formal_map_setoid: formal_topology → formal_topology → setoid1.
      | simplify; intros 3; apply trans1]]
 qed.
 
-definition cr'': ∀FT1,FT2.formal_map_setoid FT1 FT2 → arrows1 BTop FT1 FT2 ≝
- λFT1,FT2,c.cr ?? c.
-
-coercion cr''.
-
-(*
-theorem C1':
+axiom C1':
  ∀S,T: formal_topology.∀f:formal_map_setoid S T.∀U,V: Ω \sup T.
   extS ?? f (U ↓ V) = extS ?? f U ↓ extS ?? f V.
- intros;
-qed.
 
 definition formal_map_composition:
  ∀o1,o2,o3: formal_topology.
@@ -103,5 +80,18 @@ definition formal_map_composition:
      | intros;
        apply (.= (extS_com ??? c c1 ?));
        apply (.= †(C1 ?????));
-       apply (.= (C1 ?????));
+       apply (.= (C1' ?????));
+       apply (.= ((†((extS_singleton ????) \sup -1))‡(†((extS_singleton ????) \sup -1))));
+       apply (.= (extS_com ??????)\sup -1 ‡ (extS_com ??????) \sup -1);
+       apply (.= (extS_singleton ????)‡(extS_singleton ????));
+       apply refl1;
+     | apply (.= (extS_com ??? c c1 ?));
+       apply (.= (†(C2 ???)));
+       apply (.= (C2 ???));
+       apply refl1;]
+  | intros; simplify;
+    change with (comp1 BTop ??? a b = comp1 BTop ??? a' b');
+    apply prop1; assumption]
+qed.
+
 *)
\ No newline at end of file