From: Claudio Sacerdoti Coen Date: Mon, 12 Jan 2009 23:49:29 +0000 (+0000) Subject: Some work on o-algebras towards the proof that a and j are saturation/reduction X-Git-Tag: make_still_working~4264 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=0d6f208ea4728aa106d9cf0965cec853551b0b02;p=helm.git Some work on o-algebras towards the proof that a and j are saturation/reduction operators. --- diff --git a/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma b/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma index 838876f93..8ff55d0d6 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma @@ -101,6 +101,18 @@ interpretation "o-algebra meet" 'oa_meet f = interpretation "o-algebra meet with explicit function" 'oa_meet_mk f = (fun12 __ (oa_meet __) (mk_unary_morphism _ _ f _)). +notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (\emsp) \nbsp term 90 p)" +non associative with precedence 50 for @{ 'oa_join $p }. +notation < "hovbox(mstyle scriptlevel 1 scriptsizemultiplier 1.7 (∨) \below (ident i ∈  I) break term 90 p)" +non associative with precedence 50 for @{ 'oa_join_mk (λ${ident i}:$I.$p) }. + +notation > "hovbox(∨ f)" non associative with precedence 60 +for @{ 'oa_join $f }. +interpretation "o-algebra join" 'oa_join f = + (fun12 __ (oa_join __) f). +interpretation "o-algebra join with explicit function" 'oa_join_mk f = + (fun12 __ (oa_join __) (mk_unary_morphism _ _ f _)). + definition hint3: OAlgebra → setoid1. intro; apply (oa_P o); qed. @@ -125,6 +137,20 @@ qed. interpretation "o-algebra binary meet" 'and a b = (fun21 ___ (binary_meet _) a b). +definition binary_join : ∀O:OAlgebra. binary_morphism1 O O O. +intros; split; +[ intros (p q); + apply (∨ { x ∈ BOOL | match x with [ true ⇒ p | false ⇒ q ] | IF_THEN_ELSE_p ? p q }); +| intros; lapply (prop12 ? O (oa_join O BOOL)); + [2: apply ({ x ∈ BOOL | match x with [ true ⇒ a | false ⇒ b ] | IF_THEN_ELSE_p ? a b }); + |3: apply ({ x ∈ BOOL | match x with [ true ⇒ a' | false ⇒ b' ] | IF_THEN_ELSE_p ? a' b' }); + | apply Hletin;] + intro x; simplify; cases x; simplify; assumption;] +qed. + +interpretation "o-algebra binary join" 'or a b = + (fun21 ___ (binary_join _) a b). + coercion Type1_OF_OAlgebra nocomposites. lemma oa_overlap_preservers_meet: ∀O:OAlgebra.∀p,q:O.p >< q → p >< (p ∧ q). diff --git a/helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma b/helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma index 2ee78912a..b1655dc81 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/o-basic_pairs_to_o-basic_topologies.ma @@ -15,42 +15,115 @@ 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.