]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma
Some work on o-algebras towards the proof that a and j are saturation/reduction
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_pairs_to_o-basic_topologies.ma
index 2ee78912afa42af75a8874e3244e5b43090a435b..b1655dc814342451ee91647435b9808b82acd72e 100644 (file)
 include "o-basic_pairs.ma".
 include "o-basic_topologies.ma".
 
-lemma pippo: ∀S:OA.∀p,q,r:S. p ≤ q → p >< r → q >< r.
+(* qui la notazione non va *)
+lemma leq_to_eq_join: ∀S:OA.∀p,q:S. p ≤ q → q = binary_join ? p q.
  intros;
- cut (r = binary_meet ? r r); (* la notazione non va ??? *)
-  [ apply (. (#‡Hcut));
-    apply oa_overlap_preservers_meet;
-  | 
-  ]
+ 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 lemmax: ∀S,T.∀R:arrows2 OA S T.∀p,q. p ≤ q → R p ≤ R q.
+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 oa_density; intros;
- apply (. (or_prop3 : ?) ^ -1);
- apply 
+ 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 10.2, to be moved to OA *)
 lemma lemma_10_2_a: ∀S,T.∀R:arrows2 OA S T.∀p. p ≤ R⎻* (R⎻ p).
  intros;
- apply (. (or_prop2 : ?));
+ 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 : ?) ^ -1);
+ 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: ∀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;
-  [ lapply (lemma_10_2_b ?? R p);
-    
-  | apply lemma_10_2_a;]
+  [ apply f_minus_star_image_monotone;
+    apply (lemma_10_2_b ?? R p);
+  | 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_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. *)
 
 (* Qui, per fare le cose per bene, ci serve la nozione di funtore categorico *)
 definition o_basic_topology_of_basic_pair: BP → BTop.
@@ -59,11 +132,22 @@ definition o_basic_topology_of_basic_pair: BP → BTop.
   [ apply (form t);
   | apply (□_t ∘ Ext⎽t);
   | apply (◊_t ∘ Rest⎽t);
-  | intros 2;
+  | 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));
-    
-  |
+    *)
+     |
+     ]
+  | intros 2; split; intro;
+     [ change with (◊_t ((⊩) \sup * U) ≤ ◊_t ((⊩) \sup * V));
+       (*apply (.= ((lemma_10_4_b (concr t) (form t) (⊩) U)^-1)‡#);*)
+     |
+     ]
   |
   ]
 qed.