]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma
Many universe inconsistency avoided here and there.
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-concrete_spaces.ma
index 80ee4649fee1f149370a456e5c81d28f9375d26e..0bd8715a67550692dbc7fbdc702e8ecbdff15f86 100644 (file)
@@ -47,7 +47,8 @@ lemma xxx : ∀x.carr x → carr1 (setoid1_of_setoid x). intros; assumption; qed
 coercion xxx nocomposites.
 
 lemma down_p : ∀S,I:SET.∀u:S⇒S.∀c:arrows1 SET I S.∀a:I.∀a':I.a=a'→u (c a)=u (c a').
-intros; unfold uncurry_arrows; apply (†(†H));
+intros; unfold uncurry_arrows; change in c with (I ⇒ S);
+apply (†(†H));
 qed.
 
 alias symbol "eq" = "setoid eq".
@@ -118,33 +119,26 @@ definition rp'': ∀CS1,CS2.convergent_relation_space_setoid CS1 CS2 → arrows1
 
 coercion rp''.
 
-definition prop_1_SET : 
- ∀A,B:SET.∀w:arrows1 SET A B.∀a,b:A.eq1 ? a b→eq1 ? (w a) (w b).
-intros; apply (prop_1 A B w a b H);
-qed.
-
-interpretation "SET dagger" 'prop1 h = (prop_1_SET _ _ _ _ _ h).
-
 definition convergent_relation_space_composition:
  ∀o1,o2,o3: concrete_space.
   binary_morphism1
    (convergent_relation_space_setoid o1 o2)
-   (convergentin ⊢ (? (? ? ? (? ? ? (? ? ? ? ? (? ? ? (? ? ? (% ? ?))) ?)) ?) ? ? ?)_relation_space_setoid o2 o3)
+   (convergent_relation_space_setoid o2 o3)
    (convergent_relation_space_setoid o1 o3).
  intros; constructor 1;
      [ intros; whd in c c1 ⊢ %;
        constructor 1;
         [ apply (c1 ∘ c);
-        | intros;  
+        | intros;
           change in ⊢ (? ? ? % ?) with (c\sub\c⎻ (c1\sub\c⎻ (Ext⎽o3 (b↓c2))));
+          unfold uncurry_arrows;
           alias symbol "trans" = "trans1".
           apply (.= († (respects_converges : ?)));
-          apply (.= (respects_converges : ?));
-          apply refl1;
+          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)))));
+          unfold uncurry_arrows;
           apply (.= (†(respects_all_covered :?)));
-          apply (.= (respects_all_covered :?));
-          apply refl1]
+          apply rule (respects_all_covered ?? c);]
      | intros;
        change with (b ∘ a = b' ∘ a'); 
        change in H with (rp'' ?? a = rp'' ?? a');
@@ -170,4 +164,4 @@ definition CSPA: category1.
   | intros; simplify;
     change with (id1 ? o2 ∘ a = a);
     apply (id_neutral_left1 : ?);]
-qed.
\ No newline at end of file
+qed.