]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma
Some more re-organization.
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_topologies.ma
index 64c867e2296fa35c75e9012e2dad1047923a10be..20923337ef6bdea87b4ecd23c8bc49497f9e0478 100644 (file)
@@ -47,19 +47,17 @@ definition continuous_relation_setoid: basic_topology → basic_topology → set
      | simplify; intros; apply trans2; [2: apply e |3: apply e1; |1: skip]]]
 qed.
 
-(*
-definition cont_rel': ∀S,T: basic_topology. continuous_relation_setoid S T → arrows1 ? S T ≝ cont_rel.
+definition cont_rel': ∀S,T: basic_topology. continuous_relation_setoid S T → arrows2 ? S T ≝ cont_rel.
 
 coercion cont_rel'.
 
 definition cont_rel'':
  ∀S,T: basic_topology.
-  continuous_relation_setoid S T → unary_morphism (oa_P (carrbt S)) (oa_P (carrbt T)).
+  carr2 (continuous_relation_setoid S T) → ORelation_setoid (carrbt S) (carrbt T).
  intros; apply rule cont_rel; apply c;
 qed.
 
 coercion cont_rel''.
-*)
 
 (*
 theorem continuous_relation_eq':
@@ -109,66 +107,78 @@ theorem continuous_relation_eq_inv':
  split; apply Hcut; [2: assumption | intro; apply sym1; apply H]
 qed.
 *)
+
+axiom daemon: False.
 definition continuous_relation_comp:
  ∀o1,o2,o3.
   continuous_relation_setoid o1 o2 →
    continuous_relation_setoid o2 o3 →
     continuous_relation_setoid o1 o3.
  intros (o1 o2 o3 r s); constructor 1;
-  [ apply (s ∘ r)
+  [ apply (s ∘ r);
   | intros;
     apply sym1;
     change in match ((s ∘ r) U) with (s (r U));
-    (*BAD*) unfold FunClass_1_OF_carr1;
-    apply (.= ((reduced : ?)\sup -1));
+    (*<BAD>*) unfold FunClass_1_OF_Type_OF_setoid2;
+    unfold objs2_OF_basic_topology1; unfold hint;
+    letin reduced := reduced; clearbody reduced;
+    unfold uncurry_arrows in reduced ⊢ %; (*</BAD>*)
+    apply (.= (reduced : ?)\sup -1);
      [ (*BAD*) change with (eq1 ? (r U) (J ? (r U)));
        (* BAD U *) apply (.= (reduced ??? U ?)); [ assumption | apply refl1 ]
      | apply refl1]
   | intros;
-    apply sym;
+    apply sym1;
     change in match ((s ∘ r)⎻* U) with (s⎻* (r⎻* U));
     apply (.= (saturated : ?)\sup -1);
-     [ apply (.= (saturated : ?)); [ assumption | apply refl ]
-     | apply refl]]
+     [ apply (.= (saturated : ?)); [ assumption | apply refl1 ]
+     | apply refl1]]
+qed.
+
+lemma hintx: ∀S,T. (S ⇒ T) → unary_morphism1 S T.
+ intros; apply t;
 qed.
+coercion hintx.
 
-definition BTop: category1.
+definition BTop: category2.
  constructor 1;
   [ apply basic_topology
   | apply continuous_relation_setoid
   | intro; constructor 1;
-     [ apply id1
-     | intros; apply H;
-     | intros; apply H;]
+     [ apply id2
+     | intros; apply e;
+     | intros; apply e;]
   | intros; constructor 1;
      [ apply continuous_relation_comp;
-     | intros; simplify; (*intro x; simplify;*)
+     | intros; simplify;
+       change with ((b⎻* ∘ a⎻* ) ∘ A o1 = ((b'⎻* ∘ a'⎻* ) ∘ A o1)); 
        change with (b⎻* ∘ (a⎻* ∘ A o1) = b'⎻* ∘ (a'⎻* ∘ A o1));
-       change in H with (a⎻* ∘ A o1 = a'⎻* ∘ A o1);
-       change in H1 with (b⎻* ∘ A o2 = b'⎻* ∘ A o2);
-       apply (.= H‡#);
+       change in e with (a⎻* ∘ A o1 = a'⎻* ∘ A o1);
+       change in e1 with (b⎻* ∘ A o2 = b'⎻* ∘ A o2);
+       apply (.= e‡#);
        intro x;
-        
-       change with (eq1 (oa_P (carrbt o3)) (b⎻* (a'⎻* (A o1 x))) (b'⎻*(a'⎻* (A o1 x))));
+       change with (eq1 o3 (b⎻* (a'⎻* (A o1 x))) (b'⎻*(a'⎻* (A o1 x))));
        lapply (saturated o1 o2 a' (A o1 x):?) as X;
          [ apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1) ]
-       change in X with (eq1 (oa_P (carrbt o2)) (a'⎻* (A o1 x)) (A o2 (a'⎻* (A o1 x)))); 
-       unfold uncurry_arrows;
-       apply (.= †X); whd in H1;
-       lapply (H1 (a'⎻* (A o1 x))) as X1;
+       change in X with (eq1 ? (a'⎻* (A o1 x)) (A o2 (a'⎻* (A o1 x)))); 
+       alias symbol "trans" = "trans1".
+       alias symbol "prop1" = "prop11".
+       apply (.= †X);
+       whd in e1;
+       lapply (e1 (a'⎻* (A o1 x))) as X1;
        change in X1 with (eq1 (oa_P (carrbt o3)) (b⎻* (A o2 (a'⎻* (A o1 x)))) (b'⎻* (A o2 (a' \sup ⎻* (A o1 x)))));
        apply (.= X1);
-       unfold uncurry_arrows;
+       alias symbol "invert" = "setoid1 symmetry".
        apply (†(X\sup -1));]
   | intros; simplify;
     change with (((a34⎻* ∘ a23⎻* ) ∘ a12⎻* ) ∘ A o1 = ((a34⎻* ∘ (a23⎻* ∘ a12⎻* )) ∘ A o1));
     apply rule (#‡ASSOC1\sup -1);
   | intros; simplify;
-    change with ((a⎻* ∘ (id1 ? o1)⎻* ) ∘ A o1 = a⎻* ∘ A o1);
-    apply (#‡(id_neutral_right1 : ?));
+    change with ((a⎻* ∘ (id2 ? o1)⎻* ) ∘ A o1 = a⎻* ∘ A o1);
+    apply (#‡(id_neutral_right2 : ?));
   | intros; simplify;
-    change with (((id1 ? o2)⎻* ∘ a⎻* ) ∘ A o1 = a⎻* ∘ A o1);
-    apply (#‡(id_neutral_left1 : ?));]
+    change with (((id2 ? o2)⎻* ∘ a⎻* ) ∘ A o1 = a⎻* ∘ A o1);
+    apply (#‡(id_neutral_left2 : ?));]
 qed.
 
 (*