]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma
the new coercion behaviour (variants + composition with ID) and the new
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-concrete_spaces.ma
index b005216336904fc0dd6c0939bf2646386841897e..d7e0bf649754b0e995b3e791585e729de854a1a4 100644 (file)
@@ -18,7 +18,7 @@ include "o-saturations.ma".
 definition A : ∀b:BP. unary_morphism1 (form b) (form b).
 intros; constructor 1;
  [ apply (λx.□_b (Ext⎽b x));
- | do 2 unfold FunClass_1_OF_Type_OF_setoid21; intros; apply  (†(†e));]
+ | intros; apply  (†(†e));]
 qed.
 
 lemma down_p : ∀S:SET1.∀I:SET.∀u:S⇒S.∀c:arrows2 SET1 I S.∀a:I.∀a':I.a=a'→u (c a)=u (c a').
@@ -42,18 +42,10 @@ record concrete_space : Type2 ≝
 interpretation "o-concrete space downarrow" 'downarrow x = 
   (fun11 __ (downarrow _) x).
 
-definition bp': concrete_space → basic_pair ≝ λc.bp c.
-coercion bp'.
-
-definition bp'': concrete_space → objs2 BP.
- intro; apply (bp' c);
-qed.
-coercion bp''.
-
 definition binary_downarrow : 
   ∀C:concrete_space.binary_morphism1 (form C) (form C) (form C).
 intros; constructor 1;
-[ intros; apply (↓ t ∧ ↓ t1);
+[ intros; apply (↓ c ∧ ↓ c1);
 | intros;
   alias symbol "prop2" = "prop21".
   alias symbol "prop1" = "prop11".
@@ -71,10 +63,6 @@ record convergent_relation_pair (CS1,CS2: concrete_space) : Type2 ≝
            (Ext⎽CS1 (oa_one (form CS1)))
  }.
 
-definition rp' : ∀CS1,CS2. convergent_relation_pair CS1 CS2 → relation_pair CS1 CS2 ≝
- λCS1,CS2,c. rp CS1 CS2 c.
-coercion rp'.
-
 definition convergent_relation_space_setoid: concrete_space → concrete_space → setoid2.
  intros;
  constructor 1;
@@ -87,19 +75,10 @@ definition convergent_relation_space_setoid: concrete_space → concrete_space 
      | intros 3; apply trans2]]
 qed.
 
-
-definition rp'': ∀CS1,CS2.carr2 (convergent_relation_space_setoid CS1 CS2) → arrows2 BP CS1 CS2 ≝
- λCS1,CS2,c.rp ?? c.
-coercion rp''.
-
-
-definition rp''': ∀CS1,CS2.Type_OF_setoid21 (convergent_relation_space_setoid CS1 CS2) → arrows2 BP CS1 CS2 ≝
- λCS1,CS2,c.rp ?? c.
-coercion rp'''.
-
-definition rp'''': ∀CS1,CS2.Type_OF_setoid21 (convergent_relation_space_setoid CS1 CS2) → carr2 (arrows2 BP CS1 CS2) ≝
- λCS1,CS2,c.rp ?? c.
-coercion rp''''.
+definition convergent_relation_space_of_convergent_relation_space_setoid: 
+  ∀CS1,CS2.carr2 (convergent_relation_space_setoid CS1 CS2) → 
+    convergent_relation_pair CS1 CS2  ≝ λP,Q,c.c.
+coercion convergent_relation_space_of_convergent_relation_space_setoid.
 
 definition convergent_relation_space_composition:
  ∀o1,o2,o3: concrete_space.
@@ -110,21 +89,19 @@ definition convergent_relation_space_composition:
  intros; constructor 1;
      [ intros; whd in t t1 ⊢ %;
        constructor 1;
-        [ apply (t1 ∘ t);
+        [ apply (c1 ∘ c);
         | intros;
-          change in ⊢ (? ? ? % ?) with (t\sub\c⎻ (t1\sub\c⎻ (Ext⎽o3 (b↓c))));
-          unfold FunClass_1_OF_Type_OF_setoid21;
+          change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (b↓c2))));
           alias symbol "trans" = "trans1".
           apply (.= († (respects_converges : ?)));
-          apply (respects_converges ?? t (t1\sub\f⎻ b) (t1\sub\f⎻ c));
-        | change in ⊢ (? ? ? % ?) with (t\sub\c⎻ (t1\sub\c⎻ (Ext⎽o3 (oa_one (form o3)))));
-          unfold FunClass_1_OF_Type_OF_setoid21;
+          apply (respects_converges ?? c (c1\sub\f⎻ b) (c1\sub\f⎻ c2));
+        | change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (oa_one (form o3)))));
           apply (.= (†(respects_all_covered :?)));
-          apply rule (respects_all_covered ?? t);]
+          apply rule (respects_all_covered ?? c);]
      | intros;
        change with (b ∘ a = b' ∘ a'); 
-       change in e with (rp'' ?? a = rp'' ?? a');
-       change in e1 with (rp'' ?? b = rp ?? b');
+       change in e with (rp ?? a = rp ?? a');
+       change in e1 with (rp ?? b = rp ?? b');
        apply (e‡e1);]
 qed.
 
@@ -147,3 +124,11 @@ definition CSPA: category2.
     change with (id2 ? o2 ∘ a = a);
     apply (id_neutral_left2 : ?);]
 qed.
+
+definition concrete_space_of_CSPA : objs2 CSPA → concrete_space ≝ λx.x.
+coercion concrete_space_of_CSPA.
+
+definition convergent_relation_space_setoid_of_arrows2_CSPA :
+ ∀P,Q. arrows2 CSPA P Q → convergent_relation_space_setoid P Q ≝ λP,Q,x.x.
+coercion convergent_relation_space_setoid_of_arrows2_CSPA.
+