]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma
unary_morphism_N : seoidN -> setoidN -> setoidN (was ... -> setoidN+1)
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_topologies.ma
index 7f4270c79a6bb7fd100c91cfb4a5a88a7c8d8bd5..873a9df60988a632084b8c6f24e9c2b00b619590 100644 (file)
 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;
+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!!! *)
    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]]]
+     [ (*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 ?));
+     | 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 → 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':
  ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
@@ -100,73 +107,77 @@ theorem continuous_relation_eq_inv':
  split; apply Hcut; [2: assumption | intro; apply sym1; apply H]
 qed.
 *)
+
 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.
 
-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; (*
-       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 (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));
+       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);]]
+  | 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 btop_carr: BTop → Type1 ≝ λo:BTop. carr1 (oa_P (carrbt o)).
+coercion btop_carr.
+
 (*
 (*CSC: unused! *)
 (* this proof is more logic-oriented than set/lattice oriented *)
@@ -192,4 +203,4 @@ theorem continuous_relation_eqS:
   [2,4: intros; apply saturation_monotone; try (apply A_is_saturation); apply Hcut;]
  apply Hcut2; assumption.
 qed.
-*)
\ No newline at end of file
+*)