]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma
O-Basic Topologies do form a category.
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-concrete_spaces.ma
index 633f0b89e01591db6d5c73c75800fa608ed50385..5b08b59170acbd478bc720c8f646e102ba09b256 100644 (file)
@@ -72,14 +72,28 @@ interpretation "o-concrete space downarrow" 'downarrow x =
 definition bp': concrete_space → basic_pair ≝ λc.bp c.
 coercion bp'.
 
+lemma setoid_OF_OA : OA → setoid.
+intros; apply (oa_P o);
+qed.
+
+coercion setoid_OF_OA.
+
+definition binary_downarrow : 
+  ∀C:concrete_space.binary_morphism1 (form C) (form C) (form C).
+intros; constructor 1;
+[ intros; apply (↓ c ∧ ↓ c1);
+| intros; apply ((†H)‡(†H1));]
+qed.
+
+interpretation "concrete_space binary ↓" 'fintersects a b = (fun1 _ _ _ (binary_downarrow _) a b).
+
 record convergent_relation_pair (CS1,CS2: concrete_space) : Type ≝
  { rp:> arrows1 ? CS1 CS2;
    respects_converges:
-    ∀b,c.
-     extS ?? rp \sub\c (BPextS CS2 (b ↓ c)) =
-     BPextS CS1 ((extS ?? rp \sub\f b) ↓ (extS ?? rp \sub\f c));
+    ∀b,c. eq1 ? (rp\sub\c⎻ (Ext⎽CS2 (b ↓ c))) (Ext⎽CS1 (rp\sub\f⎻ b ↓ rp\sub\f⎻ c));
    respects_all_covered:
-    extS ?? rp\sub\c (BPextS CS2 (form CS2)) = BPextS CS1 (form CS1)
+     eq1 ? (rp\sub\c⎻ (Ext⎽CS2 (oa_one (form CS2))))
+           (Ext⎽CS1 (oa_one (form CS1)))
  }.
 
 definition rp' : ∀CS1,CS2. convergent_relation_pair CS1 CS2 → relation_pair CS1 CS2 ≝
@@ -108,34 +122,27 @@ definition convergent_relation_space_composition:
  ∀o1,o2,o3: concrete_space.
   binary_morphism1
    (convergent_relation_space_setoid o1 o2)
-   (convergent_relation_space_setoid o2 o3)
+   (convergentin ⊢ (? (? ? ? (? ? ? (? ? ? ? ? (? ? ? (? ? ? (% ? ?))) ?)) ?) ? ? ?)_relation_space_setoid o2 o3)
    (convergent_relation_space_setoid o1 o3).
  intros; constructor 1;
      [ intros; whd in c c1 ⊢ %;
        constructor 1;
-        [ apply (fun1 ??? (comp1 BP ???)); [apply (bp o2) |*: apply rp; assumption]
-        | intros;
-          change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
-          change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? (? ? ? (? ? ? %) ?) ?)))
-            with (c1 \sub \f ∘ c \sub \f);
-          change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? ? (? ? ? (? ? ? %) ?))))
-            with (c1 \sub \f ∘ c \sub \f);
-          apply (.= (extS_com ??????));
-          apply (.= (†(respects_converges ?????)));
-          apply (.= (respects_converges ?????));
-          apply (.= (†(((extS_com ??????) \sup -1)‡(extS_com ??????)\sup -1)));
+        [ apply (c1 ∘ c);
+        | intros;  
+          change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (b↓c2))));
+          alias symbol "trans" = "trans1".
+          apply (.= († (respects_converges : ?)));
+          apply (.= (respects_converges : ?));
           apply refl1;
-        | change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
-          apply (.= (extS_com ??????));
-          apply (.= (†(respects_all_covered ???)));
-          apply (.= respects_all_covered ???);
+        | change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (oa_one (form o3)))));
+          apply (.= (†(respects_all_covered :?)));
+          apply (.= (respects_all_covered :?));
           apply refl1]
      | intros;
-       change with (b ∘ a = b' ∘ a');
+       change with (b ∘ a = b' ∘ a'); 
        change in H with (rp'' ?? a = rp'' ?? a');
        change in H1 with (rp'' ?? b = rp ?? b');
-       apply (.= (H‡H1));
-       apply refl1]
+       apply ( (H‡H1));]
 qed.
 
 definition CSPA: category1.
@@ -144,25 +151,16 @@ definition CSPA: category1.
   | apply convergent_relation_space_setoid
   | intro; constructor 1;
      [ apply id1
-     | intros;
-       unfold id; simplify;
-       apply (.= (equalset_extS_id_X_X ??));
-       apply (.= (†((equalset_extS_id_X_X ??)\sup -1‡
-                    (equalset_extS_id_X_X ??)\sup -1)));
-       apply refl1;
-     | apply (.= (equalset_extS_id_X_X ??));
-       apply refl1]
+     | intros; apply refl1;
+     | apply refl1]
   | apply convergent_relation_space_composition
   | intros; simplify;
     change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12);
-    apply (.= ASSOC1);
-    apply refl1
+    apply ASSOC1;
   | intros; simplify;
     change with (a ∘ id1 ? o1 = a);
-    apply (.= id_neutral_right1 ????);
-    apply refl1
+    apply (id_neutral_right1 : ?);
   | intros; simplify;
     change with (id1 ? o2 ∘ a = a);
-    apply (.= id_neutral_left1 ????);
-    apply refl1]
+    apply (id_neutral_left1 : ?);]
 qed.