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=bfd76f99e31ed73bd6869051cceb62fd124d0741;hb=ccac4e720ff2a9bee8e1c9d5ba1ea6474db72572;hp=f2f6af0208c5c914fa25ab41db74f806e7da9804;hpb=8c0ccf03dbefd83818bc3b6849634f422f8ec727;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 f2f6af020..bfd76f99e 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,145 +12,20 @@ (* *) (**************************************************************************) +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 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; - (* 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. *) - -lemma oa_overlap_sym': ∀o:OA.∀U,V:o. (U >< V) = (V >< U). - intros; split; intro; apply oa_overlap_sym; assumption. -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_o_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; split; intro; + | apply hide; 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; @@ -162,7 +37,7 @@ definition o_basic_topology_of_o_basic_pair: BP → BTop. | change with (U ≤ (⊩)⎻* ((⊩)⎻ U)); apply (. (or_prop2 : ?) ^ -1); apply oa_leq_refl; ]] - | intros 2; split; intro; + | apply hide; intros 2; split; intro; [ change with (◊_t ((⊩) \sup * U) ≤ ◊_t ((⊩) \sup * V)); apply (. ((lemma_10_4_b ?? (⊩) U)^-1)‡#); apply (f_image_monotone ?? (⊩) ? ((⊩)* V)); @@ -174,7 +49,7 @@ definition o_basic_topology_of_o_basic_pair: BP → BTop. | change with ((⊩) ((⊩)* V) ≤ V); apply (. (or_prop1 : ?)); apply oa_leq_refl; ]] - | intros; + | apply hide; intros; apply (.= (oa_overlap_sym' : ?)); change with ((◊_t ((⊩)* V) >< (⊩)⎻* ((⊩)⎻ U)) = (U >< (◊_t ((⊩)* V)))); apply (.= (or_prop3 ?? (⊩) ((⊩)* V) ?)); @@ -184,44 +59,172 @@ 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; + intros (BP1 BP2 t); constructor 1; [ apply (t \sub \f); - | unfold o_basic_topology_of_o_basic_pair; simplify; intros; + | apply hide; unfold o_basic_topology_of_o_basic_pair; simplify; intros; apply sym1; - alias symbol "refl" = "refl1". - apply (.= †?); [1: apply (t \sub \f (((◊_BP1∘(⊩)* ) U))); | - lapply (†e); [2: apply rule t \sub \f; | skip | apply Hletin]] - change in ⊢ (? ? ? % ?) with ((◊_BP2 ∘(⊩)* ) ((t \sub \f ∘ (◊_BP1∘(⊩)* )) U)); - lapply (comp_assoc2 ????? (⊩)* (⊩) t \sub \f); - apply (.= †(Hletin ?)); clear Hletin; + apply (.= †(†e)); change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f ∘ (⊩)) ((⊩)* U)); - cut ?; - [3: apply CProp1; |5: cases (commute ?? t); [2: apply (e3 ^ -1 ((⊩)* U));] | 2,4: skip] - apply (.= †Hcut); + 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 (.= Hcut ^ -1); + apply (.= COM ^ -1); change in ⊢ (? ? ? % ?) with (t \sub \f (((⊩) ∘ (⊩)* ) U)); - apply (prop11 ?? t \sub \f); - apply (e ^ -1); - | unfold o_basic_topology_of_o_basic_pair; simplify; intros; + 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 (.= †?); [1: apply (t \sub \f⎻* ((((⊩)⎻* ∘ (⊩)⎻) U))); | - lapply (†e); [2: apply rule (t \sub \f⎻* ); | skip | apply Hletin]] - change in ⊢ (? ? ? % ?) with (((⊩)⎻* ∘(⊩)⎻ ) ((t \sub \f⎻* ∘ ((⊩)⎻*∘(⊩)⎻ )) U)); - lapply (comp_assoc2 ????? (⊩)⎻ (⊩)⎻* t \sub \f⎻* ); - apply (.= †(Hletin ?)); clear Hletin; + apply (.= †(†e)); change in ⊢ (? ? ? (? ? ? ? %) ?) with ((t \sub \f⎻* ∘ (⊩)⎻* ) ((⊩)⎻ U)); - cut ?; - [3: apply CProp1; |5: cases (commute ?? t); [2: apply (e1 ^ -1 ((⊩)⎻ U));] | 2,4: skip] - apply (.= †Hcut); + 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 (.= Hcut ^ -1); + apply (.= COM ^ -1); change in ⊢ (? ? ? % ?) with (t \sub \f⎻* (((⊩)⎻* ∘ (⊩)⎻ ) U)); - apply (prop11 ?? t \sub \f⎻* ); - apply (e ^ -1); ] -qed. \ No newline at end of file + change in e with (U=((⊩)⎻* ∘(⊩ \sub BP1)⎻ ) U); + apply (†e^-1);] +qed. + + +definition OR : carr3 (arrows3 CAT2 OBP BTop). +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 ⎻* ∘ A (o_basic_topology_of_o_basic_pair S)) = + (a' \sub \f ⎻*∘A (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 rule #; +| intros 6; apply refl1;] +qed. + +(* +axiom DDD : False. + +definition sigma_equivalence_relation2: + ∀C2:CAT2.∀Q.∀X,Y:exT22 ? (λy:C2.Q y).∀P. + equivalence_relation2 (exT22 ? (λf:arrows2 C2 (\fst X) (\fst Y).P f)). +intros; constructor 1; + [ intros(F G); apply (\fst F =_2 \fst G); + | intro; apply refl2; + | intros 3; apply sym2; assumption; + | intros 5; apply (trans2 ?? ??? x1 x2);] +qed. + +definition Apply : ∀C1,C2: CAT2.arrows3 CAT2 C1 C2 → CAT2. +intros (C1 C2 F); +constructor 1; +[ apply (exT22 ? (λx:C2.exT22 ? (λy:C1.map_objs2 ?? F y =_\ID x))); +| intros (X Y); constructor 1; + [ apply (exT22 ? (λf:arrows2 C2 (\fst X) (\fst Y). + exT22 ? (λg:arrows2 C1 (\fst (\snd X)) (\fst (\snd Y)). + ? (map_arrows2 ?? F ?? g) = f))); + intro; apply hide; clear g f; cases X in c; cases Y; cases x; cases x1; clear X Y x x1; + simplify; cases H; cases H1; intros; assumption; + | apply sigma_equivalence_relation2;] +| intro o; constructor 1; + [ apply (id2 C2 (\fst o)) + | exists[apply (id2 C1 (\fst (\snd o)))] + cases o; cases x; cases H; unfold hide; simplify; + apply (respects_id2 ?? F);] +| intros (o1 o2 o3); constructor 1; + [ intros (f g); whd in f g; constructor 1; + [ apply (comp2 C2 (\fst o1) (\fst o2) (\fst o3) (\fst f) (\fst g)); + | exists[apply (comp2 C1 (\fst (\snd o1)) (\fst (\snd o2)) (\fst (\snd o3)) (\fst (\snd f)) (\fst (\snd g)))] + cases o1; cases x; cases H; + +(* scrivo gli statement qua cosi' verra' un conflitto :-) + +1. definire il funtore OR +2. dimostrare che ORel e' faithful + +3. Definire la funzione + Apply: + \forall C1,C2: CAT2. F: arrows3 CAT2 C1 C2 -> CAT2 + := + constructor 1; + [ gli oggetti sono gli oggetti di C1 mappati da F + | i morfismi i morfismi di C1 mappati da F + | .... + ] + + E : objs CATS === Σx.∃y. F y = x + + Quindi (Apply C1 C2 F) (che usando da ora in avanti una coercion + scrivero' (F C1) ) e' l'immagine di C1 tramite F ed e' + una sottocategoria di C2 (qualcosa da dimostare qui??? vedi sotto + al punto 5) + +4. Definire rOBP (le OBP rappresentabili) come (BP_to_OBP BP) + [Si puo' fare lo stesso per le OA: rOA := Rel_to_OA REL ] + +5. Dimostrare che OR (il funtore faithful da OBP a OBTop) e' full + quando applicato a rOBP. + Nota: puo' darsi che faccia storie ad accettare lo statement. + Infatti rOBP e' (BP_to_OBP BP) ed e' "una sottocategoria di OBP" + e OR va da OBP a OBTop. Non so se tipa subito o se devi dare + una "proiezione" da rOBP a OBP. + +6. Definire rOBTop come (OBP_to_OBTop rOBP). + +7. Per composizione si ha un funtore full and faithful da BP a rOBTop: + basta prendere (OR \circ BP_to_OBP). + +8. Dimostrare (banale: quasi tutti i campi sono per conversione) che + esiste un funtore da rOBTop a BTop. Dimostrare che tale funtore e' + faithful e full (banale: tutta conversione). + +9. Per composizione si ha un funtore full and faithful da BP a BTop. + +10. Dimostrare che i seguenti funtori sono anche isomorphism-dense + (http://planetmath.org/encyclopedia/DenseFunctor.html): + + BP_to_OBP + OBP_to_OBTop quando applicato alle rOBP + OBTop_to_BTop quando applicato alle rOBTop + + Concludere per composizione che anche il funtore da BP a BTop e' + isomorphism-dense. + +====== Da qui in avanti non e' "necessario" nulla: + +== altre cose mancanti + +11. Dimostrare che le r* e le * orrizzontali + sono isomorfe dando il funtore da r* a * e dimostrando che componendo i + due funtori ottengo l'identita' + +12. La definizione di r* fa schifo: in pratica dici solo come ottieni + qualcosa, ma non come lo caratterizzeresti. Ora un teorema carino + e' che una a* (e.g. una aOBP) e' sempre una rOBP dove "a" sta per + atomic. Dimostrarlo per tutte le r*. + +== categorish/future works + +13. definire astrattamente la FG-completion e usare quella per + ottenere le BP da Rel e le OBP da OA. + +14. indebolire le OA, generalizzare le costruzioni, etc. come detto + con Giovanni + +*) +