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=1bf31881c29d11fc3b0b1ddb90d96908efd3e552;hb=ddc0a7b3f0acd57f879e540e696f69ca0c20bbf5;hp=b1655dc814342451ee91647435b9808b82acd72e;hpb=0d6f208ea4728aa106d9cf0965cec853551b0b02;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 b1655dc81..1bf31881c 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 @@ -12,156 +12,108 @@ (* *) (**************************************************************************) +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 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. *) +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 → 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 (.= #‡ - (* 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; + 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; ]] + | apply hide; intros 2; split; intro; [ change with (◊_t ((⊩) \sup * U) ≤ ◊_t ((⊩) \sup * V)); - (*apply (.= ((lemma_10_4_b (concr t) (form t) (⊩) U)^-1)‡#);*) - | - ] - | - ] + 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; ]] + | apply hide; 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 OBTop (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); + | apply hide; 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); + | apply hide; 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. + + +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. +