]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma
...
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_pairs_to_o-basic_topologies.ma
index f2f6af0208c5c914fa25ab41db74f806e7da9804..1bf31881c29d11fc3b0b1ddb90d96908efd3e552 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
+include "notation.ma".
 include "o-basic_pairs.ma".
 include "o-basic_topologies.ma".
 
-(* qui la notazione non va *)
-lemma leq_to_eq_join: ∀S:OA.∀p,q:S. p ≤ q → q = binary_join ? p q.
- intros;
- apply oa_leq_antisym;
-  [ apply oa_density; intros;
-    apply oa_overlap_sym;
-    unfold binary_join; simplify;
-    apply (. (oa_join_split : ?));
-    exists; [ apply false ]
-    apply oa_overlap_sym;
-    assumption
-  | unfold binary_join; simplify;
-    apply (. (oa_join_sup : ?)); intro;
-    cases i; whd in ⊢ (? ? ? ? ? % ?);
-     [ assumption | apply oa_leq_refl ]]
-qed.
-
-lemma overlap_monotone_left: ∀S:OA.∀p,q,r:S. p ≤ q → p >< r → q >< r.
- intros;
- apply (. (leq_to_eq_join : ?)‡#);
-  [ apply f;
-  | skip
-  | apply oa_overlap_sym;
-    unfold binary_join; simplify;
-    apply (. (oa_join_split : ?));
-    exists [ apply true ]
-    apply oa_overlap_sym;
-    assumption; ]
-qed.
-
-(* Part of proposition 9.9 *)
-lemma f_minus_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R⎻ p ≤ R⎻ q.
- intros;
- apply (. (or_prop2 : ?));
- apply oa_leq_trans; [2: apply f; | skip | apply (. (or_prop2 : ?)^ -1); apply oa_leq_refl;]
-qed.
-(* Part of proposition 9.9 *)
-lemma f_minus_star_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R⎻* p ≤ R⎻* q.
- intros;
- apply (. (or_prop2 : ?)^ -1);
- apply oa_leq_trans; [3: apply f; | skip | apply (. (or_prop2 : ?)); apply oa_leq_refl;]
-qed.
-
-(* Part of proposition 9.9 *)
-lemma f_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R p ≤ R q.
- intros;
- apply (. (or_prop1 : ?));
- apply oa_leq_trans; [2: apply f; | skip | apply (. (or_prop1 : ?)^ -1); apply oa_leq_refl;]
-qed.
-
-(* Part of proposition 9.9 *)
-lemma f_star_image_monotone: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R* p ≤ R* q.
- intros;
- apply (. (or_prop1 : ?)^ -1);
- apply oa_leq_trans; [3: apply f; | skip | apply (. (or_prop1 : ?)); apply oa_leq_refl;]
-qed.
-
-lemma lemma_10_2_a: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R⎻* (R⎻ p).
- intros;
- apply (. (or_prop2 : ?)^-1);
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_2_b: ∀S,T.∀R:arrows2 OA S T.∀p. R⎻ (R⎻* p) ≤ p.
- intros;
- apply (. (or_prop2 : ?));
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_2_c: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R* (R p).
- intros;
- apply (. (or_prop1 : ?)^-1);
- apply oa_leq_refl.
-qed.
-
-lemma lemma_10_2_d: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* p) ≤ p.
- intros;
- apply (. (or_prop1 : ?));
- apply oa_leq_refl.
-qed.
-
-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 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.
- intros; apply oa_leq_antisym;
-  [ apply f_star_image_monotone;
-    apply (lemma_10_2_d ?? R p);
-  | apply lemma_10_2_c; ]
-qed.
-
-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 ]
-qed.
-
-(* VEERY BAD! *)
-axiom lemma_10_4_b: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R (R* p))) = R (R* p).
-(*
- intros;
- (* BAD *)
- 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.
+alias symbol "eq" = "setoid1 eq".
 
 (* Qui, per fare le cose per bene, ci serve la nozione di funtore categorico *)
-definition o_basic_topology_of_o_basic_pair: BP → BTop.
- intro;
+definition o_basic_topology_of_o_basic_pair: OBP → OBTop.
+ intro t;
  constructor 1;
-  [ apply (form t);
+  [ apply (Oform t);
   | apply (□_t ∘ Ext⎽t);
   | apply (◊_t ∘ Rest⎽t);
-  | intros 2; split; intro;
+  | apply hide; intros 2; split; intro;
      [ change with ((⊩) \sup ⎻* ((⊩) \sup ⎻ U) ≤ (⊩) \sup ⎻* ((⊩) \sup ⎻ V));
        apply (. (#‡(lemma_10_4_a ?? (⊩) V)^-1));
        apply f_minus_star_image_monotone;
@@ -162,7 +37,7 @@ definition o_basic_topology_of_o_basic_pair: BP → BTop.
         | change with (U ≤ (⊩)⎻* ((⊩)⎻ U));
           apply (. (or_prop2 : ?) ^ -1);
           apply oa_leq_refl; ]]
-  | intros 2; split; intro;
+  | apply hide; intros 2; split; intro;
      [ change with (◊_t ((⊩) \sup * U) ≤ ◊_t ((⊩) \sup * V));
        apply (. ((lemma_10_4_b ?? (⊩) U)^-1)‡#);
        apply (f_image_monotone ?? (⊩) ? ((⊩)* V));
@@ -174,7 +49,7 @@ definition o_basic_topology_of_o_basic_pair: BP → BTop.
         | change with ((⊩) ((⊩)* V) ≤ V);
           apply (. (or_prop1 : ?));
           apply oa_leq_refl; ]]
-  | intros;
+  | apply hide; intros;
     apply (.= (oa_overlap_sym' : ?));
     change with ((◊_t ((⊩)* V) >< (⊩)⎻* ((⊩)⎻ U)) = (U >< (◊_t ((⊩)* V))));
     apply (.= (or_prop3 ?? (⊩) ((⊩)* V) ?));
@@ -184,44 +59,61 @@ definition o_basic_topology_of_o_basic_pair: BP → BTop.
 qed.
 
 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;
+ ∀BP1,BP2.arrows2 OBP BP1 BP2 →
+  arrows2 OBTop (o_basic_topology_of_o_basic_pair BP1) (o_basic_topology_of_o_basic_pair BP2).
+ intros (BP1 BP2 t);
  constructor 1;
   [ apply (t \sub \f);
-  | unfold o_basic_topology_of_o_basic_pair; simplify; intros;
+  | apply hide; 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;
+    apply (.= †(†e));
     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);
+    cut ((t \sub \f ∘ (⊩)) ((⊩)* U) = ((⊩) ∘ t \sub \c) ((⊩)* U)) as COM;[2:
+      cases (Ocommute ?? t); apply (e3 ^ -1 ((⊩)* U));]
+    apply (.= †COM);
     change in ⊢ (? ? ? % ?) with (((⊩) ∘ (⊩)* ) (((⊩) ∘ t \sub \c ∘ (⊩)* ) U));
     apply (.= (lemma_10_3_c ?? (⊩) (t \sub \c ((⊩)* U))));
-    apply (.= Hcut ^ -1);
+    apply (.= COM ^ -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;
+    change in e with (U=((⊩)∘(⊩ \sub BP1) \sup * ) U);
+    apply (†e^-1);
+  | apply hide; 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;
+    apply (.= †(†e));
     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);
+    cut ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩)⎻ U) = ((⊩)⎻* ∘ t \sub \c⎻* ) ((⊩)⎻ U)) as COM;[2:
+      cases (Ocommute ?? t); apply (e1 ^ -1 ((⊩)⎻ U));]
+    apply (.= †COM);
     change in ⊢ (? ? ? % ?) with (((⊩)⎻* ∘ (⊩)⎻ ) (((⊩)⎻* ∘ t \sub \c⎻* ∘ (⊩)⎻ ) U));
     apply (.= (lemma_10_3_d ?? (⊩) (t \sub \c⎻* ((⊩)⎻ U))));
-    apply (.= Hcut ^ -1);
+    apply (.= COM ^ -1);
     change in ⊢ (? ? ? % ?) with (t \sub \f⎻* (((⊩)⎻* ∘ (⊩)⎻ ) U));
-    apply (prop11 ?? t \sub \f⎻* );
-    apply (e ^ -1); ]
-qed.
\ No newline at end of file
+    change in e with (U=((⊩)⎻* ∘(⊩ \sub BP1)⎻ ) U);
+    apply (†e^-1);]
+qed.
+
+
+definition OR : carr3 (arrows3 CAT2 OBP OBTop).
+constructor 1;
+[ apply o_basic_topology_of_o_basic_pair;
+| intros; constructor 1;
+  [ apply o_continuous_relation_of_o_relation_pair;
+  | apply hide; 
+    intros; whd; unfold o_continuous_relation_of_o_relation_pair; simplify;;
+    change with ((a \sub \f ⎻* ∘ oA (o_basic_topology_of_o_basic_pair S)) =
+                 (a' \sub \f ⎻*∘ oA (o_basic_topology_of_o_basic_pair S)));
+    whd in e; cases e; clear e e2 e3 e4;
+    change in ⊢ (? ? ? (? ? ? ? ? % ?) ?) with ((⊩\sub S)⎻* ∘ (⊩\sub S)⎻);
+    apply (.= (comp_assoc2 ? ???? ?? a\sub\f⎻* ));
+    change in ⊢ (? ? ? (? ? ? ? ? ? %) ?) with (a\sub\f ∘ ⊩\sub S)⎻*;
+    apply (.= #‡†(Ocommute:?)^-1);
+    apply (.= #‡e1);
+    change in ⊢ (? ? ? (? ? ? ? ? ? %) ?) with (⊩\sub T ∘ a'\sub\c)⎻*;
+    apply (.= #‡†(Ocommute:?));    
+    change in ⊢ (? ? ? (? ? ? ? ? ? %) ?) with (a'\sub\f⎻* ∘ (⊩\sub S)⎻* );    
+    apply (.= (comp_assoc2 ? ???? ?? a'\sub\f⎻* )^-1);
+    apply refl2;]
+| intros 2 (o a); apply refl1;
+| intros 6; apply refl1;]
+qed.
+