X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fo-basic_pairs_to_o-basic_topologies.ma;h=b78e7b037046bededec4973267647e8e115270bc;hb=b6e187ff7580c3dbec8bf467915d0ccd0dfd65a8;hp=2ee78912afa42af75a8874e3244e5b43090a435b;hpb=b000797a7e07f511926a19d947feae90406f6c89;p=helm.git 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..b78e7b037 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,69 +15,78 @@ include "o-basic_pairs.ma". include "o-basic_topologies.ma". -lemma pippo: ∀S:OA.∀p,q,r:S. p ≤ q → p >< r → q >< r. - intros; - cut (r = binary_meet ? r r); (* la notazione non va ??? *) - [ apply (. (#‡Hcut)); - apply oa_overlap_preservers_meet; - | - ] - -(* Part of proposition 9.9 *) -lemma lemmax: ∀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 - -(* 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 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 oa_leq_refl. -qed. - -lemma lemma_10_3: ∀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;] -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_basic_pair: BP → BTop. - intro; +definition o_basic_topology_of_o_basic_pair: OBP → BTop. + intro t; constructor 1; - [ apply (form t); + [ apply (Oform t); | apply (□_t ∘ Ext⎽t); | apply (◊_t ∘ Rest⎽t); - | intros 2; - lapply depth=0 (or_prop1 ?? (rel t)); - lapply depth=0 (or_prop2 ?? (rel t)); - - | - | - ] + | 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; + 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 ?? (⊩) 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). - intros; +definition o_continuous_relation_of_o_relation_pair: + ∀BP1,BP2.arrows2 OBP BP1 BP2 → + arrows2 BTop (o_basic_topology_of_o_basic_pair BP1) (o_basic_topology_of_o_basic_pair BP2). + intros (BP1 BP2 t); 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) ?); ] + [ apply (t \sub \f); + | unfold o_basic_topology_of_o_basic_pair; simplify; intros; + apply sym1; + apply (.= †(†e)); + change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f ∘ (⊩)) ((⊩)* U)); + 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 (.= COM ^ -1); + change in ⊢ (? ? ? % ?) with (t \sub \f (((⊩) ∘ (⊩)* ) U)); + change in e with (U=((⊩)∘(⊩ \sub BP1) \sup * ) U); + apply (†e^-1); + | unfold o_basic_topology_of_o_basic_pair; simplify; intros; + apply sym1; + apply (.= †(†e)); + change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩)⎻ U)); + 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 (.= COM ^ -1); + change in ⊢ (? ? ? % ?) with (t \sub \f⎻* (((⊩)⎻* ∘ (⊩)⎻ ) U)); + change in e with (U=((⊩)⎻* ∘(⊩ \sub BP1)⎻ ) U); + apply (†e^-1);] qed.