]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma
faithful
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_topologies.ma
index a958da4258b31de8376571f0eafbbe52ffa6b7fe..8847eef3157fbdf9c8ec2969894997ea49b4b493 100644 (file)
@@ -1,4 +1,4 @@
-(**************************************************************************)
+ (**************************************************************************)
 (*       ___                                                              *)
 (*      ||M||                                                             *)
 (*      ||A||       A project by Andrea Asperti                           *)
 include "o-algebra.ma".
 include "o-saturations.ma".
 
-record basic_topology: Type ≝
+record basic_topology: Type2 ≝
  { carrbt:> OA;
-   A: arrows1 SET (oa_P carrbt) (oa_P carrbt);
-   J: arrows1 SET (oa_P carrbt) (oa_P carrbt);
-   A_is_saturation: is_saturation ? A;
-   J_is_reduction: is_reduction ? J;
+   A: carrbt ⇒ carrbt;
+   J: carrbt ⇒ carrbt;
+   A_is_saturation: is_o_saturation ? A;
+   J_is_reduction: is_o_reduction ? J;
    compatibility: ∀U,V. (A U >< J V) = (U >< J V)
  }.
 
-record continuous_relation (S,T: basic_topology) : Type ≝
- { cont_rel:> arrows1 ? S T;
+record continuous_relation (S,T: basic_topology) : Type2 ≝
+ { cont_rel:> arrows2 OA S T;
+   (* reduces uses eq1, saturated uses eq!!! *)
    reduced: ∀U. U = J ? U → cont_rel U = J ? (cont_rel U);
    saturated: ∀U. U = A ? U → cont_rel⎻* U = A ? (cont_rel⎻* U)
  }. 
 
-definition continuous_relation_setoid: basic_topology → basic_topology → setoid1.
+definition continuous_relation_setoid: basic_topology → basic_topology → setoid2.
  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)));
-     | simplify; intros; apply refl1;
-     | simplify; intros; apply sym1; apply H
-     | simplify; intros; apply trans1; [2: apply H |3: apply H1; |1: skip]]]
-qed.
-
-definition cont_rel': ∀S,T: basic_topology. continuous_relation_setoid S T → arrows1 ? 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)).
- intros; apply rule cont_rel; apply c;
+     [ 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.
 
-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':
  ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
   a = a' → ∀X.a⎻* (A o1 X) = a'⎻* (A o1 X).
- intros;
- lapply (prop_1_SET ??? H);
-  split; intro; unfold minus_star_image; simplify; intros;
+ intros; apply oa_leq_antisym; intro; unfold minus_star_image; simplify; intros;
   [ cut (ext ?? a a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
     lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
     cut (A ? (ext ?? a' a1) ⊆ A ? X); [2: apply (. (H ?)‡#); assumption]
@@ -75,11 +68,11 @@ qed.
 
 theorem continuous_relation_eq_inv':
  ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
-  (∀X.minus_star_image ?? a (A o1 X) = minus_star_image ?? a' (A o1 X)) → a=a'.
+  (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) → a=a'.
  intros 6;
  cut (∀a,a': continuous_relation_setoid o1 o2.
-  (∀X.minus_star_image ?? a (A o1 X) = minus_star_image ?? a' (A o1 X)) → 
-   ∀V:o2. A ? (ext ?? a' V) ⊆ A ? (ext ?? a V));
+  (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) → 
+   ∀V:(oa_P (carrbt o2)). A o1 (a'⎻ V) ≤ A o1 (a⎻ V));
   [2: clear b H a' a; intros;
       lapply depth=0 (λV.saturation_expansive ??? (extS ?? a V)); [2: apply A_is_saturation;|skip]
        (* fundamental adjunction here! to be taken out *)
@@ -101,6 +94,7 @@ theorem continuous_relation_eq_inv':
       assumption;]
  split; apply Hcut; [2: assumption | intro; apply sym1; apply H]
 qed.
+*)
 
 definition continuous_relation_comp:
  ∀o1,o2,o3.
@@ -108,77 +102,67 @@ definition continuous_relation_comp:
    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_carr2;
+              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 sym1;
-    apply (.= †(image_comp ??????));
-    apply (.= (reduced ?????)\sup -1);
-     [ apply (.= (reduced ?????)); [ assumption | apply refl1 ]
-     | apply (.= (image_comp ??????)\sup -1);
-       apply refl1]
-     | intros;
-       apply sym1;
-       apply (.= †(minus_star_image_comp ??????));
-       apply (.= (saturated ?????)\sup -1);
-        [ apply (.= (saturated ?????)); [ assumption | apply refl1 ]
-        | apply (.= (minus_star_image_comp ??????)\sup -1);
-          apply refl1]]
+    change in match ((s ∘ r)⎻* U) with (s⎻* (r⎻* U));
+    apply (.= (saturated : ?)\sup -1);
+     [ apply (.= (saturated : ?)); [ assumption | apply refl1 ]
+     | apply refl1]]
 qed.
 
-definition BTop: category1.
+definition BTop: category2.
  constructor 1;
   [ apply basic_topology
   | apply continuous_relation_setoid
   | intro; constructor 1;
-     [ apply id1
-     | intros;
-       apply (.= (image_id ??));
-       apply sym1;
-       apply (.= †(image_id ??));
-       apply sym1;
-       assumption
-     | intros;
-       apply (.= (minus_star_image_id ??));
-       apply sym1;
-       apply (.= †(minus_star_image_id ??));
-       apply sym1;
-       assumption]
+     [ apply id2
+     | intros; apply e;
+     | intros; apply e;]
   | intros; constructor 1;
      [ apply continuous_relation_comp;
-     | intros; simplify; intro x; simplify;
-       lapply depth=0 (continuous_relation_eq' ???? H) as H';
-       lapply depth=0 (continuous_relation_eq' ???? H1) as H1';
-       letin K ≝ (λX.H1' (minus_star_image ?? a (A ? X))); clearbody K;
-       cut (∀X:Ω \sup o1.
-              minus_star_image o2 o3 b (A o2 (minus_star_image o1 o2 a (A o1 X)))
-            = minus_star_image o2 o3 b' (A o2 (minus_star_image o1 o2 a' (A o1 X))));
-        [2: intro; apply sym1; apply (.= #‡(†((H' ?)\sup -1))); apply sym1; apply (K X);]
-       clear K H' H1';
-       cut (∀X:Ω \sup o1.
-              minus_star_image o1 o3 (b ∘ a) (A o1 X) = minus_star_image o1 o3 (b'∘a') (A o1 X));
-        [2: intro;
-            apply (.= (minus_star_image_comp ??????));
-            apply (.= #‡(saturated ?????));
-             [ apply ((saturation_idempotent ????) \sup -1); apply A_is_saturation ]
-            apply sym1; 
-            apply (.= (minus_star_image_comp ??????));
-            apply (.= #‡(saturated ?????));
-             [ apply ((saturation_idempotent ????) \sup -1); apply A_is_saturation ]
-           apply ((Hcut X) \sup -1)]
-       clear Hcut; generalize in match x; clear x;
-       apply (continuous_relation_eq_inv');
-       apply Hcut1;]
-  | intros; simplify; intro; do 2 (unfold continuous_relation_comp); simplify;
-    apply (.= †(ASSOC1‡#));
-    apply refl1
-  | intros; simplify; intro; unfold continuous_relation_comp; simplify;
-    apply (.= †((id_neutral_right1 ????)‡#));
-    apply refl1
-  | intros; simplify; intro; simplify;
-    apply (.= †((id_neutral_left1 ????)‡#));
-    apply refl1]
+     | intros; simplify;
+       change with ((b⎻* ∘ a⎻* ) ∘ A o1 = ((b'⎻* ∘ a'⎻* ) ∘ A o1)); 
+       change with (b⎻* ∘ (a⎻* ∘ A o1) = b'⎻* ∘ (a'⎻* ∘ A o1));
+       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 ? (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))));
+       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);
+  | intros; simplify;
+    change with ((a⎻* ∘ (id2 ? o1)⎻* ) ∘ A o1 = a⎻* ∘ A o1);
+    apply (#‡(id_neutral_right2 : ?));
+  | intros; simplify;
+    change with (((id2 ? o2)⎻* ∘ a⎻* ) ∘ A o1 = a⎻* ∘ A o1);
+    apply (#‡(id_neutral_left2 : ?));]
 qed.
 
+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! *)
 (* this proof is more logic-oriented than set/lattice oriented *)
 theorem continuous_relation_eqS:
@@ -203,3 +187,4 @@ theorem continuous_relation_eqS:
   [2,4: intros; apply saturation_monotone; try (apply A_is_saturation); apply Hcut;]
  apply Hcut2; assumption.
 qed.
+*)