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=1791079573076a20be6ff31c9b872efda139f2a4;hpb=0080faad4e730c227b6bbb2549407b23703b477a;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 179107957..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 @@ -17,133 +17,11 @@ 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. +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; split; intro; @@ -180,7 +58,7 @@ definition o_basic_topology_of_o_basic_pair: BP → BTop. qed. definition o_continuous_relation_of_o_relation_pair: - ∀BP1,BP2.arrows2 BP BP1 BP2 → + ∀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; @@ -190,7 +68,7 @@ definition o_continuous_relation_of_o_relation_pair: apply (.= †(†e)); change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f ∘ (⊩)) ((⊩)* U)); cut ((t \sub \f ∘ (⊩)) ((⊩)* U) = ((⊩) ∘ t \sub \c) ((⊩)* U)) as COM;[2: - cases (commute ?? t); apply (e3 ^ -1 ((⊩)* U));] + 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)))); @@ -203,7 +81,7 @@ definition o_continuous_relation_of_o_relation_pair: apply (.= †(†e)); change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩)⎻ U)); cut ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩)⎻ U) = ((⊩)⎻* ∘ t \sub \c⎻* ) ((⊩)⎻ U)) as COM;[2: - cases (commute ?? t); apply (e1 ^ -1 ((⊩)⎻ U));] + 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))));