X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fbasic_pairs_to_o-basic_pairs.ma;h=d8813fdc1b1154c37458065c9f548d4ddddb3fea;hb=6ca487df4f361fd6b0a3b0734396bc6a62b520c3;hp=7fc6a07f46f0fc2f3c0bb43786a039dfc52b6d3e;hpb=2857d1c432f073379552e1572235a86509b665a4;p=helm.git diff --git a/helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_o-basic_pairs.ma b/helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_o-basic_pairs.ma index 7fc6a07f4..d8813fdc1 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_o-basic_pairs.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_o-basic_pairs.ma @@ -17,24 +17,96 @@ include "o-basic_pairs.ma". include "relations_to_o-algebra.ma". (* Qui, per fare le cose per bene, ci serve la nozione di funtore categorico *) -definition o_basic_pair_of_basic_pair: cic:/matita/formal_topology/basic_pairs/basic_pair.ind#xpointer(1/1) → basic_pair. - intro; +definition o_basic_pair_of_basic_pair: basic_pair → Obasic_pair. + intro b; constructor 1; - [ apply (SUBSETS (concr b)); - | apply (SUBSETS (form b)); - | apply (orelation_of_relation ?? (rel b)); ] + [ apply (map_objs2 ?? POW (concr b)); + | apply (map_objs2 ?? POW (form b)); + | apply (map_arrows2 ?? POW (concr b) (form b) (rel b)); ] qed. definition o_relation_pair_of_relation_pair: - ∀BP1,BP2.cic:/matita/formal_topology/basic_pairs/relation_pair.ind#xpointer(1/1) BP1 BP2 → - relation_pair (o_basic_pair_of_basic_pair BP1) (o_basic_pair_of_basic_pair BP2). + ∀BP1,BP2. relation_pair BP1 BP2 → + Orelation_pair (o_basic_pair_of_basic_pair BP1) (o_basic_pair_of_basic_pair BP2). intros; 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 (map_arrows2 ?? POW (concr BP1) (concr BP2) (r \sub \c)); + | apply (map_arrows2 ?? POW (form BP1) (form BP2) (r \sub \f)); + | apply (.= (respects_comp2 ?? POW (concr BP1) (concr BP2) (form BP2) r\sub\c (⊩\sub BP2) )^-1); + cut ( ⊩ \sub BP2 ∘ r \sub \c =_12 r\sub\f ∘ ⊩ \sub BP1) as H; + [ apply (.= †H); + apply (respects_comp2 ?? POW (concr BP1) (form BP1) (form BP2) (⊩\sub BP1) r\sub\f); + | apply commute;]] qed. + +definition BP_to_OBP: carr3 (arrows3 CAT2 (category2_of_category1 BP) OBP). + constructor 1; + [ apply o_basic_pair_of_basic_pair; + | intros; constructor 1; + [ apply (o_relation_pair_of_relation_pair S T); + | intros (a b Eab); split; unfold o_relation_pair_of_relation_pair; simplify; + unfold o_basic_pair_of_basic_pair; simplify; + [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x); + | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x); + | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x); + | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);] + simplify; + apply (prop12); + apply (.= (respects_comp2 ?? POW (concr S) (concr T) (form T) (a\sub\c) (⊩\sub T))^-1); + apply sym2; + apply (.= (respects_comp2 ?? POW (concr S) (concr T) (form T) (b\sub\c) (⊩\sub T))^-1); + apply sym2; + apply prop12; + apply Eab; + ] + | simplify; intros; whd; split; + [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x); + | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x); + | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x); + | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);] + simplify; + apply prop12; + apply prop22;[2,4,6,8: apply rule #;] + apply (respects_id2 ?? POW (concr o)); + | simplify; intros; whd; split; + [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x); + | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x); + | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x); + | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);] + simplify; + apply prop12; + apply prop22;[2,4,6,8: apply rule #;] + apply (respects_comp2 ?? POW (concr o1) (concr o2) (concr o3) f1\sub\c f2\sub\c);] +qed. + +theorem BP_to_OBP_faithful: + ∀S,T.∀f,g:arrows2 (category2_of_category1 BP) S T. + map_arrows2 ?? BP_to_OBP ?? f = map_arrows2 ?? BP_to_OBP ?? g → f=g. + intros; change with ( (⊩) ∘ f \sub \c = (⊩) ∘ g \sub \c); + apply (POW_faithful); + apply (.= respects_comp2 ?? POW (concr S) (concr T) (form T) f \sub \c (⊩ \sub T)); + apply sym2; + apply (.= respects_comp2 ?? POW (concr S) (concr T) (form T) g \sub \c (⊩ \sub T)); + apply sym2; + apply e; +qed. + +theorem BP_to_OBP_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? BP_to_OBP S T g = f). + intros; + cases (POW_full (concr S) (concr T) (Oconcr_rel ?? f)) (gc Hgc); + cases (POW_full (form S) (form T) (Oform_rel ?? f)) (gf Hgf); + exists[ + constructor 1; [apply gc|apply gf] + apply (POW_faithful); + apply (let xxxx ≝POW in .= respects_comp2 ?? POW (concr S) (concr T) (form T) gc (rel T)); + apply rule (.= Hgc‡#); + apply (.= Ocommute ?? f); + apply (.= #‡Hgf^-1); + apply (let xxxx ≝POW in (respects_comp2 ?? POW (concr S) (form S) (form T) (rel S) gf)^-1)] + split; + [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x); + | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x); + | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x); + | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);] + simplify; apply (†(Hgc‡#)); +qed.