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;fp=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Fbasic_pairs_to_o-basic_pairs.ma;h=7dbd3a80ae6bd0c9bbe9b6c687e9404dda6420d8;hb=13114a0147a28f8c7359c9c19ee254716eb5f55a;hp=2dd4147ae592f6568a644d06c713eebef6a63e7c;hpb=071d7a246190074c97e78192839e4bb5d5a1eef4;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 2dd4147ae..7dbd3a80a 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,42 +17,85 @@ 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 ?? SUBSETS' (concr b)); + | apply (map_objs2 ?? SUBSETS' (form b)); + | apply (map_arrows2 ?? SUBSETS' (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 ?? SUBSETS' (concr BP1) (concr BP2) (r \sub \c)); + | apply (map_arrows2 ?? SUBSETS' (form BP1) (form BP2) (r \sub \f)); + | apply (.= (respects_comp2 ?? SUBSETS' (concr BP1) (concr BP2) (form BP2) r\sub\c (⊩\sub BP2) )^-1); + cut (⊩ \sub BP2∘r \sub \c = r\sub\f ∘ ⊩ \sub BP1) as H; + [ apply (.= †H); + apply (respects_comp2 ?? SUBSETS' (concr BP1) (form BP1) (form BP2) (⊩\sub BP1) r\sub\f); + | apply commute;]] qed. -(* -definition BP_to_OBP: carr3 (arrows3 CAT2 (category2_of_category1 cic:/matita/formal_topology/basic_pairs/BP.con) BP). +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; split; unfold o_relation_pair_of_relation_pair; simplify; - unfold o_basic_pair_of_basic_pair; simplify; ] - | simplify; intros; whd; split; unfold o_relation_pair_of_relation_pair; simplify; - unfold o_basic_pair_of_basic_pair; -simplify in ⊢ (? ? ? (? % ? ?) ?); -simplify in ⊢ (? ? ? (? ? % ?) ?); -simplify in ⊢ (? ? ? ? (? % ? ?)); -simplify in ⊢ (? ? ? ? (? ? % ?)); - | simplify; intros; whd; split;unfold o_relation_pair_of_relation_pair; simplify; - unfold o_basic_pair_of_basic_pair;simplify; + | 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 ?? SUBSETS' (concr S) (concr T) (form T) (a\sub\c) (⊩\sub T))^-1); + apply sym2; + apply (.= (respects_comp2 ?? SUBSETS' (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 ?? SUBSETS' (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 ?? SUBSETS' (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; unfold BP_to_OBP in e; simplify in e; cases e; + unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4; + intros 2; change in match or_f_ in e3 with (λq,w,x.fun12 ?? (or_f q w) x); + simplify in e3; STOP lapply (e3 (singleton ? x)); cases Hletin; + split; intro; [ lapply (s y); | lapply (s1 y); ] + [2,4: exists; [1,3:apply x] split; [1,3: assumption |*: change with (x=x); apply rule #] + |*: cases Hletin1; cases x1; change in f3 with (eq1 ? x w); apply (. f3‡#); assumption;] +qed. +*) + +(* +theorem SUBSETS_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? BP_to_OBP S T g = f). + intros; exists; + *) \ No newline at end of file