]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma
all pullbacks are attempted in sequence, removed many unfold
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_pairs_to_o-basic_topologies.ma
index bdadaf0fe376a197d162194faefb3dac21e69166..1791079573076a20be6ff31c9b872efda139f2a4 100644 (file)
@@ -17,6 +17,128 @@ include "o-basic_topologies.ma".
 
 alias symbol "eq" = "setoid1 eq".
 
+(* 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; apply (†(lemma_10_3_a ?? R p));
+qed.
+
+lemma lemma_10_4_b: ∀S,T.∀R:arrows2 OA S T.∀p. R (R* (R (R* p))) = R (R* p).
+intros; apply (†(lemma_10_3_b ?? R p));
+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_o_basic_pair: BP → BTop.
  intro t;
@@ -65,7 +187,6 @@ definition o_continuous_relation_of_o_relation_pair:
   [ apply (t \sub \f);
   | unfold o_basic_topology_of_o_basic_pair; simplify; intros;
     apply sym1;
-    unfold in ⊢ (? ? ? (? ? ? ? %) ?);
     apply (.= †(†e));
     change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f ∘ (⊩)) ((⊩)* U));
     cut ((t \sub \f ∘ (⊩)) ((⊩)* U) = ((⊩) ∘ t \sub \c) ((⊩)* U)) as COM;[2:
@@ -76,10 +197,9 @@ definition o_continuous_relation_of_o_relation_pair:
     apply (.= COM ^ -1);
     change in ⊢ (? ? ? % ?) with (t \sub \f (((⊩) ∘ (⊩)* ) U));
     change in e with (U=((⊩)∘(⊩ \sub BP1) \sup * ) U);
-    unfold in ⊢ (? ? ? % %); apply (†e^-1);
+    apply (†e^-1);
   | unfold o_basic_topology_of_o_basic_pair; simplify; intros;
     apply sym1;
-    unfold in ⊢ (? ? ? (? ? ? ? %) ?);
     apply (.= †(†e));
     change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩)⎻ U));
     cut ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩)⎻ U) = ((⊩)⎻* ∘ t \sub \c⎻* ) ((⊩)⎻ U)) as COM;[2:
@@ -90,5 +210,5 @@ definition o_continuous_relation_of_o_relation_pair:
     apply (.= COM ^ -1);
     change in ⊢ (? ? ? % ?) with (t \sub \f⎻* (((⊩)⎻* ∘ (⊩)⎻ ) U));
     change in e with (U=((⊩)⎻* ∘(⊩ \sub BP1)⎻ ) U);
-    unfold in ⊢ (? ? ? % %); apply (†e^-1);]
-qed.
\ No newline at end of file
+    apply (†e^-1);]
+qed.