]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_o-basic_pairs.ma
7dbd3a80ae6bd0c9bbe9b6c687e9404dda6420d8
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / basic_pairs_to_o-basic_pairs.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "basic_pairs.ma".
16 include "o-basic_pairs.ma".
17 include "relations_to_o-algebra.ma".
18
19 (* Qui, per fare le cose per bene, ci serve la nozione di funtore categorico *)
20 definition o_basic_pair_of_basic_pair: basic_pair → Obasic_pair.
21  intro b;
22  constructor 1;
23   [ apply (map_objs2 ?? SUBSETS' (concr b));
24   | apply (map_objs2 ?? SUBSETS' (form b));
25   | apply (map_arrows2 ?? SUBSETS' (concr b) (form b) (rel b)); ]
26 qed.
27
28 definition o_relation_pair_of_relation_pair:
29  ∀BP1,BP2. relation_pair BP1 BP2 →
30   Orelation_pair (o_basic_pair_of_basic_pair BP1) (o_basic_pair_of_basic_pair BP2).
31  intros;
32  constructor 1;
33   [ apply (map_arrows2 ?? SUBSETS' (concr BP1) (concr BP2) (r \sub \c));
34   | apply (map_arrows2 ?? SUBSETS' (form BP1) (form BP2) (r \sub \f));
35   | apply (.= (respects_comp2 ?? SUBSETS' (concr BP1) (concr BP2) (form BP2)  r\sub\c (⊩\sub BP2) )^-1);
36     cut (⊩ \sub BP2∘r \sub \c = r\sub\f ∘ ⊩ \sub BP1) as H;
37     [ apply (.= †H);
38       apply (respects_comp2 ?? SUBSETS' (concr BP1) (form BP1) (form BP2) (⊩\sub BP1) r\sub\f);
39     | apply commute;]]
40 qed.
41
42 definition BP_to_OBP: carr3 (arrows3 CAT2 (category2_of_category1 BP) OBP).
43  constructor 1;
44   [ apply o_basic_pair_of_basic_pair;
45   | intros; constructor 1;
46      [ apply (o_relation_pair_of_relation_pair S T);
47      | intros (a b Eab); split; unfold o_relation_pair_of_relation_pair; simplify;
48        unfold o_basic_pair_of_basic_pair; simplify;
49        [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x); 
50        | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
51        | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
52        | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
53        simplify;
54        apply (prop12);
55        apply (.= (respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) (a\sub\c) (⊩\sub T))^-1);
56        apply sym2;
57        apply (.= (respects_comp2 ?? SUBSETS' (concr S) (concr T) (form T) (b\sub\c) (⊩\sub T))^-1);
58        apply sym2;
59        apply prop12;
60        apply Eab;
61      ]
62   | simplify; intros; whd; split; 
63        [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x); 
64        | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
65        | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
66        | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
67     simplify;
68     apply prop12;
69     apply prop22;[2,4,6,8: apply rule #;]
70     apply (respects_id2 ?? SUBSETS' (concr o)); 
71   | simplify; intros; whd; split;
72        [ change in match or_f_minus_star_ with (λq,w,x.fun12 ?? (or_f_minus_star q w) x); 
73        | change in match or_f_minus_ with (λq,w,x.fun12 ?? (or_f_minus q w) x);
74        | change in match or_f_ with (λq,w,x.fun12 ?? (or_f q w) x);
75        | change in match or_f_star_ with (λq,w,x.fun12 ?? (or_f_star q w) x);]
76     simplify;
77     apply prop12;
78     apply prop22;[2,4,6,8: apply rule #;]
79     apply (respects_comp2 ?? SUBSETS' (concr o1) (concr o2) (concr o3) f1\sub\c f2\sub\c);]     
80 qed.
81
82
83 (*
84 theorem BP_to_OBP_faithful:
85  ∀S,T.∀f,g:arrows2 (category2_of_category1 BP) S T.
86   map_arrows2 ?? BP_to_OBP ?? f = map_arrows2 ?? BP_to_OBP ?? g → f=g.
87  intros; unfold BP_to_OBP in e; simplify in e; cases e;
88  unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4;
89  intros 2; change in match or_f_ in e3 with (λq,w,x.fun12 ?? (or_f q w) x);
90  simplify in e3; STOP lapply (e3 (singleton ? x)); cases Hletin;
91  split; intro; [ lapply (s y); | lapply (s1 y); ]
92   [2,4: exists; [1,3:apply x] split; [1,3: assumption |*: change with (x=x); apply rule #]
93   |*: cases Hletin1; cases x1; change in f3 with (eq1 ? x w); apply (. f3‡#); assumption;]
94 qed.
95 *)
96
97 (*
98 theorem SUBSETS_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? BP_to_OBP S T g = f).
99  intros; exists;
100  
101 *)