]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/library/formal_topology/relations_to_o-algebra.ma
Preparing for 0.5.9 release.
[helm.git] / helm / software / matita / library / formal_topology / relations_to_o-algebra.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 "formal_topology/relations.ma".
16 include "formal_topology/o-algebra.ma".
17
18 definition POW': objs1 SET → OAlgebra.
19  intro A; constructor 1;
20   [ apply (Ω^A);
21   | apply subseteq;
22   | apply overlaps;
23   | apply big_intersects;
24   | apply big_union;
25   | apply ({x | True});
26     simplify; intros; apply (refl1 ? (eq1 CPROP));
27   | apply ({x | False});
28     simplify; intros; apply (refl1 ? (eq1 CPROP));
29   | intros; whd; intros; assumption
30   | intros; whd; split; assumption
31   | intros; apply transitive_subseteq_operator; [2: apply f; | skip | assumption]
32   | intros; cases f; exists [apply w] assumption
33   | intros; split; [ intros 4; apply (f ? f1 i); | intros 3; intro; apply (f i ? f1); ]
34   | intros; split;
35      [ intros 4; apply f; exists; [apply i] assumption;
36      | intros 3; intros; cases f1; apply (f w a x); ]
37   | intros 3; cases f;
38   | intros 3; constructor 1;
39   | intros; cases f; exists; [apply w]
40      [ assumption
41      | whd; intros; cases i; simplify; assumption]
42   | intros; split; intro;
43      [ (** screenshot "screen-pow". *) cases f; cases x1; exists [apply w1] exists [apply w] assumption;
44      | cases e; cases x; exists; [apply w1] [ assumption | exists; [apply w] assumption]]
45   | intros; intros 2; cases (f {(a)} ?); 
46      [ exists; [apply a] [assumption | change with (a = a); apply refl1;]
47      | change in x1 with (a = w); change with (mem A a q); apply (. (x1‡#));
48        assumption]]
49 qed.
50
51 definition powerset_of_POW': ∀A.oa_P (POW' A) → Ω^A ≝ λA,x.x.
52 coercion powerset_of_POW'.
53
54 definition orelation_of_relation: ∀o1,o2:REL. o1 ⇒_\r1 o2 → (POW' o1) ⇒_\o2 (POW' o2).
55  intros;
56  constructor 1;
57   [ apply rule c;
58   | apply rule (c⎻* ); 
59   | apply rule (c* );
60   | apply rule (c⎻);
61   | intros; split; intro;
62      [ intros 2; intros 2; apply (f y); exists[apply a] split; assumption;
63      | intros 2; change with (a ∈ q); cases f1; cases x; clear f1 x; 
64        apply (f w f3); assumption; ]
65   | unfold foo; intros; split; intro;
66      [ intros 2; intros 2; apply (f x); exists [apply a] split; assumption;
67      | intros 2; change with (a ∈ q); cases f1; cases x; apply (f w f3); assumption;]
68   | intros; split; unfold foo; unfold image_coercion; simplify; intro; cases f; clear f;
69      [ cases x; cases x2; clear x x2; exists; [apply w1]
70         [ assumption | exists; [apply w] split; assumption]
71      | cases x1; cases x2; clear x1 x2; exists; [apply w1]
72         [ exists; [apply w] split; assumption;
73         | assumption; ]]]
74 qed.
75
76 lemma orelation_of_relation_preserves_equality:
77  ∀o1,o2:REL.∀t,t': o1 ⇒_\r1 o2. 
78    t = t' → orelation_of_relation ?? t =_2 orelation_of_relation ?? t'.
79  intros; split; unfold orelation_of_relation; unfold foo; simplify;
80  change in e with (t =_2 t'); unfold image_coercion; apply (†e);
81 qed.
82
83 lemma minus_image_id : ∀o:REL.((id1 REL o))⎻ =_1 (id2 SET1 Ω^o).
84 unfold foo; intro o; intro; unfold minus_image; simplify; split; simplify; intros;
85 [ cases e; cases x; change with (a1 ∈ a); change in f with (a1 =_1 w); 
86   apply (. f‡#); assumption;
87 | change in f with (a1 ∈ a); exists [ apply a1] split; try assumption; 
88   change with (a1 =_1 a1); apply refl1;]
89 qed.
90
91 lemma star_image_id : ∀o:REL.  ((id1 REL o))* =_1 (id2 SET1 Ω^o).
92 unfold foo; intro o; intro; unfold star_image; simplify; split; simplify; intros;
93 [ change with (a1 ∈ a); apply f; change with (a1 =_1 a1); apply rule refl1;
94 | change in f1 with (a1 =_1 y); apply (. f1^-1‡#); apply f;]
95 qed.
96     
97 lemma orelation_of_relation_preserves_identity:
98  ∀o1:REL. orelation_of_relation ?? (id1 ? o1) =_2 id2 OA (POW' o1).
99  intros; split; 
100  (unfold orelation_of_relation; unfold OA; unfold foo; simplify);
101  [ apply (minus_star_image_id o1);
102  | apply (minus_image_id o1); 
103  | apply (image_id o1);
104  | apply (star_image_id o1) ]
105 qed.
106  
107 (*
108   split; whd; intro; 
109   [ change with ((∀x. x ♮(id1 REL o1) a1→x∈a) → a1 ∈ a); intros;
110     apply (f a1); change with (a1 = a1); apply refl1;
111   | change with (a1 ∈ a → ∀x. x ♮(id1 REL o1) a1→x∈a); intros;
112     change in f1 with (x = a1); apply (. f1‡#); apply f;
113   | alias symbol "and" = "and_morphism".
114     change with ((∃y:o1.a1 ♮(id1 REL o1) y ∧ y∈a) → a1 ∈ a);
115     intro; cases e; clear e; cases x; clear x; change in f with (a1=w);
116     apply (. f‡#); apply f1;
117   | change with (a1 ∈ a → ∃y:o1.a1 ♮(id1 REL o1) y ∧ y∈a);
118     intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f]
119   | change with ((∃x:o1.x ♮(id1 REL o1) a1∧x∈a) → a1 ∈ a);
120     intro; cases e; clear e; cases x; clear x; change in f with (w=a1);
121     apply (. f^-1‡#); apply f1;
122   | change with (a1 ∈ a → ∃x:o1.x ♮(id1 REL o1) a1∧x∈a);
123     intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f]
124   | change with ((∀y.a1 ♮(id1 REL o1) y→y∈a) → a1 ∈ a); intros;
125     apply (f a1); change with (a1 = a1); apply refl1;
126   | change with (a1 ∈ a → ∀y.a1 ♮(id1 REL o1) y→y∈a); intros;
127     change in f1 with (a1 = y); apply (. f1^-1‡#); apply f;]
128 qed.
129 *)
130
131 (* CSC: ???? forse un uncertain mancato *)
132 alias symbol "eq" = "setoid2 eq".
133 alias symbol "compose" = "category1 composition".
134 lemma orelation_of_relation_preserves_composition:
135  ∀o1,o2,o3:REL.∀F: o1 ⇒_\r1 o2.∀G: o2 ⇒_\r1 o3.
136   orelation_of_relation ?? (G ∘ F) = 
137   comp2 OA ??? (orelation_of_relation ?? F) (orelation_of_relation ?? G).
138  intros; split; intro; split; whd; intro; whd in ⊢ (% → %); intros;
139   [ whd; intros; apply f; exists; [ apply x] split; assumption; 
140   | cases f1; clear f1; cases x1; clear x1; apply (f w); assumption;
141   | cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ]
142     split; [ assumption | exists; [apply w] split; assumption ]
143   | cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ]
144     split; [ exists; [apply w] split; assumption | assumption ]
145   | unfold arrows1_of_ORelation_setoid; 
146     cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ]
147     split; [ assumption | exists; [apply w] split; assumption ]
148   | unfold arrows1_of_ORelation_setoid in e; 
149     cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ]
150     split; [ exists; [apply w] split; assumption | assumption ]
151   | whd; intros; apply f; exists; [ apply y] split; assumption;
152   | cases f1; clear f1; cases x; clear x; apply (f w); assumption;]
153 qed.
154
155 definition POW: carr3 (arrows3 CAT2 (category2_of_category1 REL) OA).
156  constructor 1;
157   [ apply POW';
158   | intros; constructor 1;
159      [ apply (orelation_of_relation S T);
160      | intros; apply (orelation_of_relation_preserves_equality S T a a' e); ]
161   | apply orelation_of_relation_preserves_identity;
162   | apply orelation_of_relation_preserves_composition; ]
163 qed.
164
165 theorem POW_faithful: faithful2 ?? POW.
166  intros 5; unfold POW in e; simplify in e; cases e;
167  unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4;
168  intros 2; simplify; unfold image_coercion in e3; cases (e3 {(x)});
169  split; intro; [ lapply (s y); | lapply (s1 y); ]
170   [2,4: exists; [1,3:apply x] split; [1,3: assumption |*: change with (x=x); apply rule #]
171   |*: cases Hletin; cases x1; change in f3 with (x =_1 w); apply (. f3‡#); assumption;]
172 qed.
173
174
175 (*
176 lemma currify: ∀A,B,C. (A × B ⇒_1 C) → A → (B ⇒_1 C).
177 intros; constructor 1; [ apply (b c); | intros; apply (#‡e);]
178 qed.
179 *)
180
181 include "formal_topology/notation.ma".
182
183 theorem POW_full: full2 ?? POW. 
184  intros 3 (S T); exists;
185   [ constructor 1; constructor 1;
186      [ apply (λx:carr S.λy:carr T. y ∈ f {(x)});
187      | apply hide; intros; unfold FunClass_1_OF_carr2; lapply (.= e1‡#);
188         [4: apply mem; |6: apply Hletin;|1,2,3,5: skip]
189        lapply (#‡prop11 ?? f ?? (†e)); [6: apply Hletin; |*:skip ]]  
190      | (split; intro; split; simplify);
191            [ change with (∀a1.(∀x. a1 ∈ (f {(x):S}) → x ∈ a) → a1 ∈ f⎻* a);
192            | change with (∀a1.a1 ∈ f⎻* a → (∀x.a1 ∈ f {(x):S} → x ∈ a)); 
193            | alias symbol "and" (instance 4) = "and_morphism".
194 change with (∀a1.(∃y:carr T. y ∈ f {(a1):S} ∧ y ∈ a) → a1 ∈ f⎻ a);
195            | alias symbol "and" (instance 2) = "and_morphism".
196 change with (∀a1.a1 ∈ f⎻ a → (∃y:carr T.y ∈ f {(a1):S} ∧ y ∈ a)); 
197            | alias symbol "and" (instance 3) = "and_morphism".
198 change with (∀a1.(∃x:carr S. a1 ∈ f {(x):S} ∧ x ∈ a) → a1 ∈ f a);
199            | change with (∀a1.a1 ∈. f a → (∃x:carr S. a1 ∈ f {(x):S} ∧ x ∈ a));
200            | change with (∀a1.(∀y. y ∈ f {(a1):S} → y ∈ a) → a1 ∈ f* a);
201            | change with (∀a1.a1 ∈ f* a → (∀y. y ∈ f {(a1):S} → y ∈ a)); ]
202         [ intros; apply ((. (or_prop2 ?? f (singleton ? a1) a)^-1) ? a1);
203            [ intros 2; apply (f1 a2); change in f2 with (a2 ∈ f⎻ (singleton ? a1));
204              lapply (. (or_prop3 ?? f (singleton ? a2) (singleton ? a1)));
205               [ cases Hletin; change in x1 with (eq1 ? a1 w);
206                 apply (. x1‡#); assumption;
207               | exists; [apply a2] [change with (a2=a2); apply rule #; | assumption]]
208            | change with (a1 = a1); apply rule #; ]
209         | intros; apply ((. (or_prop2 ?? f (singleton ? a1) a)) ? x);
210            [ intros 2; change in f3 with (eq1 ? a1 a2); change with (a2 ∈ f⎻* a); apply (. f3^-1‡#);
211              assumption;
212            | lapply (. (or_prop3 ?? f (singleton ? x) (singleton ? a1))^-1);
213               [ cases Hletin; change in x1 with (eq1 ? x w);
214                 change with (x ∈ f⎻ (singleton ? a1)); apply (. x1‡#); assumption;
215               | exists; [apply a1] [assumption | change with (a1=a1); apply rule #; ]]]
216         | intros; cases e; cases x; clear e x;
217           lapply (. (or_prop3 ?? f (singleton ? a1) a)^-1);
218            [ cases Hletin; change in x with (eq1 ? a1 w1); apply (. x‡#); assumption;
219            | exists; [apply w] assumption ]
220         | intros; lapply (. (or_prop3 ?? f (singleton ? a1) a));
221            [ cases Hletin; exists; [apply w] split; assumption;
222            | exists; [apply a1] [change with (a1=a1); apply rule #; | assumption ]] 
223         | intros; cases e; cases x; clear e x;
224           apply (f_image_monotone ?? f (singleton ? w) a ? a1);
225            [ intros 2; change in f3 with (eq1 ? w a2); change with (a2 ∈ a);
226              apply (. f3^-1‡#); assumption;
227            | assumption; ]
228         | intros; lapply (. (or_prop3 ?? f a (singleton ? a1))^-1);
229            [ cases Hletin; exists; [apply w] split;
230               [ lapply (. (or_prop3 ?? f (singleton ? w) (singleton ? a1)));
231                  [ cases Hletin1; change in x3 with (eq1 ? a1 w1); apply (. x3‡#); assumption;
232                  | exists; [apply w] [change with (w=w); apply rule #; | assumption ]]
233               | assumption ]
234            | exists; [apply a1] [ assumption; | change with (a1=a1); apply rule #;]]
235         | intros; apply ((. (or_prop1 ?? f (singleton ? a1) a)^-1) ? a1);
236            [ apply f1; | change with (a1=a1); apply rule #; ]
237         | intros; apply ((. (or_prop1 ?? f (singleton ? a1) a)) ? y);
238            [ intros 2; change in f3 with (eq1 ? a1 a2); change with (a2 ∈ f* a);
239              apply (. f3^-1‡#); assumption;
240            | assumption ]]]
241 qed.