]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma
the new coercion behaviour (variants + composition with ID) and the new
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_topologies.ma
index 68b9befb82dd93e3e5acb5e4d4fd0feefc9b0b3a..8847eef3157fbdf9c8ec2969894997ea49b4b493 100644 (file)
@@ -1,4 +1,4 @@
-(**************************************************************************)
+ (**************************************************************************)
 (*       ___                                                              *)
 (*      ||M||                                                             *)
 (*      ||A||       A project by Andrea Asperti                           *)
@@ -24,11 +24,6 @@ record basic_topology: Type2 ≝
    compatibility: ∀U,V. (A U >< J V) = (U >< J V)
  }.
 
-lemma hint: OA → objs2 OA.
- intro; apply t;
-qed.
-coercion hint.
-
 record continuous_relation (S,T: basic_topology) : Type2 ≝
  { cont_rel:> arrows2 OA S T;
    (* reduces uses eq1, saturated uses eq!!! *)
@@ -40,24 +35,17 @@ definition continuous_relation_setoid: basic_topology → basic_topology → set
  intros (S T); constructor 1;
   [ apply (continuous_relation S T)
   | constructor 1;
-     [ (*apply (λr,s:continuous_relation S T.∀b. eq1 (oa_P (carrbt S)) (A ? (r⎻ b)) (A ? (s⎻ b)));*)
-       apply (λr,s:continuous_relation S T.r⎻* ∘ (A S) = s⎻* ∘ (A ?));
+     [ alias symbol "eq" = "setoid2 eq".
+       alias symbol "compose" = "category2 composition".
+       apply (λr,s:continuous_relation S T. (r⎻* ) ∘ (A S) = (s⎻* ∘ (A ?)));
      | simplify; intros; apply refl2;
      | simplify; intros; apply sym2; apply e
      | 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 → arrows2 ? S T ≝ cont_rel.
-
-coercion cont_rel'.
-
-definition cont_rel'':
- ∀S,T: basic_topology.
-  carr2 (continuous_relation_setoid S T) → ORelation_setoid (carrbt S) (carrbt T).
- intros; apply rule cont_rel; apply c;
-qed.
-
-coercion cont_rel''.
+definition continuous_relation_of_continuous_relation_setoid: 
+  ∀P,Q. continuous_relation_setoid P Q → continuous_relation P Q ≝ λP,Q,c.c.
+coercion continuous_relation_of_continuous_relation_setoid.
 
 (*
 theorem continuous_relation_eq':
@@ -118,11 +106,8 @@ definition continuous_relation_comp:
   | intros;
     apply sym1; 
     change in match ((s ∘ r) U) with (s (r U));
-    (*<BAD>*) unfold FunClass_1_OF_Type_OF_setoid21;
-    unfold objs2_OF_basic_topology1; unfold hint;
-    letin reduced := reduced; clearbody reduced;
-    unfold uncurry_arrows in reduced ⊢ %; (*</BAD>*)
-    apply (.= (reduced : ?)\sup -1);
+    (*<BAD>*) unfold FunClass_1_OF_carr2;
+              apply (.= (reduced : ?)\sup -1);
      [ (*BAD*) change with (eq1 ? (r U) (J ? (r U)));
        (* BAD U *) apply (.= (reduced ??? U ?)); [ assumption | apply refl1 ]
      | apply refl1]
@@ -150,20 +135,15 @@ definition BTop: category2.
        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 (b⎻* (a'⎻* (A o1 x)) = b'⎻*(a'⎻* (A o1 x)));
-       alias symbol "trans" = "trans1".
-       alias symbol "prop1" = "prop11".
-       alias symbol "invert" = "setoid1 symmetry".
-       lapply (.= †(saturated o1 o2 a' (A o1 x) : ?));
-        [3: apply (b⎻* ); | 5: apply Hletin; |1,2: skip;
-        |apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1); ]
-       change in e1 with (∀x.b⎻* (A o2 x) = b'⎻* (A o2 x));
+       intro x;          
+       change with (eq1 ? (b⎻* (a'⎻* (A o1 x))) (b'⎻*(a'⎻* (A o1 x)))); 
+       apply (.= †(saturated o1 o2 a' (A o1 x) ?)); [
+         apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1);]
        apply (.= (e1 (a'⎻* (A o1 x))));
-       alias symbol "invert" = "setoid1 symmetry".
-       lapply (†((saturated ?? a' (A o1 x) : ?) ^ -1));
-        [2: apply (b'⎻* ); |4: apply Hletin; | skip;
-        |apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1);]]
+       change with (eq1 ? (b'⎻* (A o2 (a'⎻* (A o1 x)))) (b'⎻*(a'⎻* (A o1 x))));   
+       apply (.= †(saturated o1 o2 a' (A o1 x):?)^-1); [
+         apply ((o_saturation_idempotent ?? (A_is_saturation o1) x)^-1);]
+       apply rule #;]
   | intros; simplify;
     change with (((a34⎻* ∘ a23⎻* ) ∘ a12⎻* ) ∘ A o1 = ((a34⎻* ∘ (a23⎻* ∘ a12⎻* )) ∘ A o1));
     apply rule (#‡ASSOC ^ -1);
@@ -175,8 +155,12 @@ definition BTop: category2.
     apply (#‡(id_neutral_left2 : ?));]
 qed.
 
-definition btop_carr: BTop → Type1 ≝ λo:BTop. carr1 (oa_P (carrbt o)).
-coercion btop_carr.
+definition basic_topology_of_BTop: objs2 BTop → basic_topology ≝ λx.x.
+coercion basic_topology_of_BTop.
+
+definition continuous_relation_setoid_of_arrows2_BTop : 
+  ∀P,Q. arrows2 BTop P Q → continuous_relation_setoid P Q ≝ λP,Q,x.x.
+coercion continuous_relation_setoid_of_arrows2_BTop.
 
 (*
 (*CSC: unused! *)