]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma
o_continous_relations are really o_relation_pair... up to a bug of Matita
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_pairs_to_o-basic_topologies.ma
index b1655dc814342451ee91647435b9808b82acd72e..f2f6af0208c5c914fa25ab41db74f806e7da9804 100644 (file)
@@ -97,11 +97,11 @@ lemma lemma_10_2_d: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* p) ≤ p.
  apply oa_leq_refl.
 qed.
 
-lemma lemma_10_3_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* p)) = R⎻* p.
+lemma lemma_10_3_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* (R⎻ p)) = R⎻ p.
  intros; apply oa_leq_antisym;
-  [ apply f_minus_star_image_monotone;
-    apply (lemma_10_2_b ?? R p);
-  | apply lemma_10_2_a; ]
+  [ apply lemma_10_2_b;
+  | apply f_minus_image_monotone;
+    apply lemma_10_2_a; ]
 qed.
 
 lemma lemma_10_3_b: ∀S,T.∀R:arrows2 OA S T.∀p. R* (R (R* p)) = R* p.
@@ -111,10 +111,24 @@ lemma lemma_10_3_b: ∀S,T.∀R:arrows2 OA S T.∀p. R* (R (R* p)) = R* p.
   | apply lemma_10_2_c; ]
 qed.
 
-lemma lemma_10_4_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* (R⎻ (R⎻* p))) = R⎻ (R⎻* p).
+lemma lemma_10_3_c: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R p)) = R p.
+ intros; apply oa_leq_antisym;
+  [ apply lemma_10_2_d;
+  | apply f_image_monotone;
+    apply (lemma_10_2_c ?? R p); ]
+qed.
+
+lemma lemma_10_3_d: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* p)) = R⎻* p.
+ intros; apply oa_leq_antisym;
+  [ apply f_minus_star_image_monotone;
+    apply (lemma_10_2_b ?? R p);
+  | apply lemma_10_2_a; ]
+qed.
+
+lemma lemma_10_4_a: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻* (R⎻ (R⎻* (R⎻ p))) = R⎻* (R⎻ p).
  intros;
  (* BAD *)
- lapply (†(lemma_10_3_a ?? R p)); [2: apply (R⎻); | skip | apply Hletin ]
+ lapply (†(lemma_10_3_a ?? R p)); [2: apply (R⎻); | skip | apply Hletin ]
 qed.
 
 (* VEERY BAD! *)
@@ -125,8 +139,12 @@ axiom lemma_10_4_b: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R (R* p))) = R (R* p
  lapply (†(lemma_10_3_b ?? R p)); [2: apply rule R; | skip | apply Hletin ]
 qed. *)
 
+lemma oa_overlap_sym': ∀o:OA.∀U,V:o. (U >< V) = (V >< U).
+ intros; split; intro; apply oa_overlap_sym; assumption.
+qed.
+
 (* Qui, per fare le cose per bene, ci serve la nozione di funtore categorico *)
-definition o_basic_topology_of_basic_pair: BP → BTop.
+definition o_basic_topology_of_o_basic_pair: BP → BTop.
  intro;
  constructor 1;
   [ apply (form t);
@@ -134,34 +152,76 @@ definition o_basic_topology_of_basic_pair: BP → BTop.
   | apply (◊_t ∘ Rest⎽t);
   | intros 2; split; intro;
      [ change with ((⊩) \sup ⎻* ((⊩) \sup ⎻ U) ≤ (⊩) \sup ⎻* ((⊩) \sup ⎻ V));
-     (*  apply (.= #‡
-    (* BAD *)
-    whd in t;
-    apply oa_leq_antisym;
-    lapply depth=0 (or_prop1 ?? (rel t));
-    lapply depth=0 (or_prop2 ?? (rel t));
-    *)
-     |
-     ]
+       apply (. (#‡(lemma_10_4_a ?? (⊩) V)^-1));
+       apply f_minus_star_image_monotone;
+       apply f_minus_image_monotone;
+       assumption
+     | apply oa_leq_trans;
+        [3: apply f;
+        | skip
+        | change with (U ≤ (⊩)⎻* ((⊩)⎻ U));
+          apply (. (or_prop2 : ?) ^ -1);
+          apply oa_leq_refl; ]]
   | intros 2; split; intro;
      [ change with (◊_t ((⊩) \sup * U) ≤ ◊_t ((⊩) \sup * V));
-       (*apply (.= ((lemma_10_4_b (concr t) (form t) (⊩) U)^-1)‡#);*)
-     |
-     ]
-  |
-  ]
+       apply (. ((lemma_10_4_b ?? (⊩) U)^-1)‡#);
+       apply (f_image_monotone ?? (⊩) ? ((⊩)* V));
+       apply f_star_image_monotone;
+       assumption;
+     | apply oa_leq_trans;
+        [2: apply f;
+        | skip
+        | change with ((⊩) ((⊩)* V) ≤ V);
+          apply (. (or_prop1 : ?));
+          apply oa_leq_refl; ]]
+  | intros;
+    apply (.= (oa_overlap_sym' : ?));
+    change with ((◊_t ((⊩)* V) >< (⊩)⎻* ((⊩)⎻ U)) = (U >< (◊_t ((⊩)* V))));
+    apply (.= (or_prop3 ?? (⊩) ((⊩)* V) ?));
+    apply (.= #‡(lemma_10_3_a : ?));
+    apply (.= (or_prop3 : ?)^-1);
+    apply (oa_overlap_sym' ? ((⊩) ((⊩)* V)) U); ]
 qed.
 
-definition o_convergent_relation_pair_of_convergent_relation_pair:
- ∀BP1,BP2.cic:/matita/formal_topology/concrete_spaces/convergent_relation_pair.ind#xpointer(1/1) BP1 BP2 →
-  convergent_relation_pair (o_concrete_space_of_concrete_space BP1) (o_concrete_space_of_concrete_space BP2).
+definition o_continuous_relation_of_o_relation_pair:
+ ∀BP1,BP2.arrows2 BP BP1 BP2 →
+  arrows2 BTop (o_basic_topology_of_o_basic_pair BP1) (o_basic_topology_of_o_basic_pair BP2).
  intros;
  constructor 1;
-  [ apply (orelation_of_relation ?? (r \sub \c));
-  | apply (orelation_of_relation ?? (r \sub \f));
-  | lapply (commute ?? r);
-    lapply (orelation_of_relation_preserves_equality ???? Hletin);
-    apply (.= (orelation_of_relation_preserves_composition (concr BP1) ??? (rel BP2)) ^ -1);
-    apply (.= (orelation_of_relation_preserves_equality ???? (commute ?? r)));
-    apply (orelation_of_relation_preserves_composition ?? (form BP2)  (rel BP1) ?); ]
-qed.
+  [ apply (t \sub \f);
+  | unfold o_basic_topology_of_o_basic_pair; simplify; intros;
+    apply sym1;
+    alias symbol "refl" = "refl1".
+    apply (.= †?); [1: apply (t \sub \f (((◊_BP1∘(⊩)* ) U))); |
+    lapply (†e); [2: apply rule t \sub \f; | skip | apply Hletin]]
+    change in ⊢ (? ? ? % ?) with ((◊_BP2 ∘(⊩)* ) ((t \sub \f ∘ (◊_BP1∘(⊩)* )) U));
+    lapply (comp_assoc2 ????? (⊩)* (⊩) t \sub \f);
+    apply (.= †(Hletin ?)); clear Hletin;
+    change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f ∘ (⊩)) ((⊩)* U));
+    cut ?;
+     [3: apply CProp1; |5: cases (commute ?? t); [2: apply (e3 ^ -1 ((⊩)* U));] | 2,4: skip]
+    apply (.= †Hcut);
+    change in ⊢ (? ? ? % ?) with (((⊩) ∘ (⊩)* ) (((⊩) ∘ t \sub \c ∘ (⊩)* ) U));
+    apply (.= (lemma_10_3_c ?? (⊩) (t \sub \c ((⊩)* U))));
+    apply (.= Hcut ^ -1);
+    change in ⊢ (? ? ? % ?) with (t \sub \f (((⊩) ∘ (⊩)* ) U));
+    apply (prop11 ?? t \sub \f);
+    apply (e ^ -1);
+  | unfold o_basic_topology_of_o_basic_pair; simplify; intros;
+    apply sym1;
+    apply (.= †?); [1: apply (t \sub \f⎻* ((((⊩)⎻* ∘ (⊩)⎻) U))); |
+    lapply (†e); [2: apply rule (t \sub \f⎻* ); | skip | apply Hletin]]
+    change in ⊢ (? ? ? % ?) with (((⊩)⎻* ∘(⊩)⎻ ) ((t \sub \f⎻* ∘ ((⊩)⎻*∘(⊩)⎻ )) U));
+    lapply (comp_assoc2 ????? (⊩)⎻ (⊩)⎻* t \sub \f⎻* );
+    apply (.= †(Hletin ?)); clear Hletin;
+    change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩)⎻ U));
+    cut ?;
+     [3: apply CProp1; |5: cases (commute ?? t); [2: apply (e1 ^ -1 ((⊩)⎻ U));] | 2,4: skip]
+    apply (.= †Hcut);
+    change in ⊢ (? ? ? % ?) with (((⊩)⎻* ∘ (⊩)⎻ ) (((⊩)⎻* ∘ t \sub \c⎻* ∘ (⊩)⎻ ) U));
+    apply (.= (lemma_10_3_d ?? (⊩) (t \sub \c⎻* ((⊩)⎻ U))));
+    apply (.= Hcut ^ -1);
+    change in ⊢ (? ? ? % ?) with (t \sub \f⎻* (((⊩)⎻* ∘ (⊩)⎻ ) U));
+    apply (prop11 ?? t \sub \f⎻* );
+    apply (e ^ -1); ]
+qed.
\ No newline at end of file