]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/formal_topology/overlap/relations_to_o-algebra.ma
Additional contribs.
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / 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 "relations.ma".
16 include "o-algebra.ma".
17
18 definition POW': objs1 SET → OAlgebra.
19  intro A; constructor 1;
20   [ apply (Ω \sup 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      [ 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) → Ω \sup A ≝ λA,x.x.
52 coercion powerset_of_POW'.
53
54 definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (POW' o1) (POW' o2).
55  intros;
56  constructor 1;
57   [ constructor 1; 
58      [ apply (λU.image ?? c U);
59      | intros; apply (#‡e); ]
60   | constructor 1;
61      [ apply (λU.minus_star_image ?? c U);
62      | intros; apply (#‡e); ]
63   | constructor 1;
64      [ apply (λU.star_image ?? c U);
65      | intros; apply (#‡e); ]
66   | constructor 1;
67      [ apply (λU.minus_image ?? c U);
68      | intros; apply (#‡e); ]
69   | intros; split; intro;
70      [ change in f with (∀a. a ∈ image ?? c p → a ∈ q);
71        change with (∀a:o1. a ∈ p → a ∈ star_image ?? c q);
72        intros 4; apply f; exists; [apply a] split; assumption;
73      | change in f with (∀a:o1. a ∈ p → a ∈ star_image ?? c q);
74        change with (∀a. a ∈ image ?? c p → a ∈ q);
75        intros; cases f1; cases x; clear f1 x; apply (f ? f3); assumption; ]
76   | intros; split; intro;
77      [ change in f with (∀a. a ∈ minus_image ?? c p → a ∈ q);
78        change with (∀a:o2. a ∈ p → a ∈ minus_star_image ?? c q);
79        intros 4; apply f; exists; [apply a] split; assumption;
80      | change in f with (∀a:o2. a ∈ p → a ∈ minus_star_image ?? c q);
81        change with (∀a. a ∈ minus_image ?? c p → a ∈ q);
82        intros; cases f1; cases x; clear f1 x; apply (f ? f3); assumption; ]
83   | intros; split; intro; cases f; clear f;
84      [ cases x; cases x2; clear x x2; exists; [apply w1]
85         [ assumption;
86         | exists; [apply w] split; assumption]
87      | cases x1; cases x2; clear x1 x2; exists; [apply w1]
88         [ exists; [apply w] split; assumption;
89         | assumption; ]]]
90 qed.
91
92 lemma orelation_of_relation_preserves_equality:
93  ∀o1,o2:REL.∀t,t': arrows1 ? o1 o2. t = t' → eq2 ? (orelation_of_relation ?? t) (orelation_of_relation ?? t').
94  intros; split; unfold orelation_of_relation; simplify; intro; split; intro;
95  simplify; whd in o1 o2;
96   [ change with (a1 ∈ minus_star_image ?? t a → a1 ∈ minus_star_image ?? t' a);
97     apply (. #‡(e^-1‡#));
98   | change with (a1 ∈ minus_star_image ?? t' a → a1 ∈ minus_star_image ?? t a);
99     apply (. #‡(e‡#));
100   | change with (a1 ∈ minus_image ?? t a → a1 ∈ minus_image ?? t' a);
101     apply (. #‡(e ^ -1‡#));
102   | change with (a1 ∈ minus_image ?? t' a → a1 ∈ minus_image ?? t a);
103     apply (. #‡(e‡#));
104   | change with (a1 ∈ image ?? t a → a1 ∈ image ?? t' a);
105     apply (. #‡(e ^ -1‡#));
106   | change with (a1 ∈ image ?? t' a → a1 ∈ image ?? t a);
107     apply (. #‡(e‡#));
108   | change with (a1 ∈ star_image ?? t a → a1 ∈ star_image ?? t' a);
109     apply (. #‡(e ^ -1‡#));
110   | change with (a1 ∈ star_image ?? t' a → a1 ∈ star_image ?? t a);
111     apply (. #‡(e‡#)); ]
112 qed.
113
114 lemma orelation_of_relation_preserves_identity:
115  ∀o1:REL. eq2 ? (orelation_of_relation ?? (id1 ? o1)) (id2 OA (POW' o1)).
116  intros; split; intro; split; whd; intro; 
117   [ change with ((∀x. x ♮(id1 REL o1) a1→x∈a) → a1 ∈ a); intros;
118     apply (f a1); change with (a1 = a1); apply refl1;
119   | change with (a1 ∈ a → ∀x. x ♮(id1 REL o1) a1→x∈a); intros;
120     change in f1 with (x = a1); apply (. f1‡#); apply f;
121   | alias symbol "and" = "and_morphism".
122     change with ((∃y:o1.a1 ♮(id1 REL o1) y ∧ y∈a) → a1 ∈ a);
123     intro; cases e; clear e; cases x; clear x; change in f with (a1=w);
124     apply (. f‡#); apply f1;
125   | change with (a1 ∈ a → ∃y:o1.a1 ♮(id1 REL o1) y ∧ y∈a);
126     intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f]
127   | change with ((∃x:o1.x ♮(id1 REL o1) a1∧x∈a) → a1 ∈ a);
128     intro; cases e; clear e; cases x; clear x; change in f with (w=a1);
129     apply (. f^-1‡#); apply f1;
130   | change with (a1 ∈ a → ∃x:o1.x ♮(id1 REL o1) a1∧x∈a);
131     intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f]
132   | change with ((∀y.a1 ♮(id1 REL o1) y→y∈a) → a1 ∈ a); intros;
133     apply (f a1); change with (a1 = a1); apply refl1;
134   | change with (a1 ∈ a → ∀y.a1 ♮(id1 REL o1) y→y∈a); intros;
135     change in f1 with (a1 = y); apply (. f1^-1‡#); apply f;]
136 qed.
137
138 (* CSC: ???? forse un uncertain mancato *)
139 alias symbol "eq" = "setoid2 eq".
140 alias symbol "compose" = "category1 composition".
141 lemma orelation_of_relation_preserves_composition:
142  ∀o1,o2,o3:REL.∀F: arrows1 ? o1 o2.∀G: arrows1 ? o2 o3.
143   orelation_of_relation ?? (G ∘ F) =
144    comp2 OA (POW' o1) (POW' o2) (POW' o3)
145     ?? (*(orelation_of_relation ?? F) (orelation_of_relation ?? G)*).
146  [ apply (orelation_of_relation ?? F); | apply (orelation_of_relation ?? G); ]
147  intros; split; intro; split; whd; intro; whd in ⊢ (% → %); intros;
148   [ whd; intros; apply f; exists; [ apply x] split; assumption; 
149   | cases f1; clear f1; cases x1; clear x1; apply (f w); assumption;
150   | cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ]
151     split; [ assumption | exists; [apply w] split; assumption ]
152   | cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ]
153     split; [ exists; [apply w] split; assumption | assumption ]
154   | unfold arrows1_of_ORelation_setoid; 
155     cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ]
156     split; [ assumption | exists; [apply w] split; assumption ]
157   | unfold arrows1_of_ORelation_setoid in e; 
158     cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ]
159     split; [ exists; [apply w] split; assumption | assumption ]
160   | whd; intros; apply f; exists; [ apply y] split; assumption;
161   | cases f1; clear f1; cases x; clear x; apply (f w); assumption;]
162 qed.
163
164 definition POW: carr3 (arrows3 CAT2 (category2_of_category1 REL) OA).
165  constructor 1;
166   [ apply POW';
167   | intros; constructor 1;
168      [ apply (orelation_of_relation S T);
169      | intros; apply (orelation_of_relation_preserves_equality S T a a' e); ]
170   | apply orelation_of_relation_preserves_identity;
171   | apply orelation_of_relation_preserves_composition; ]
172 qed.
173
174 theorem POW_faithful:
175  ∀S,T.∀f,g:arrows2 (category2_of_category1 REL) S T.
176   map_arrows2 ?? POW ?? f = map_arrows2 ?? POW ?? g → f=g.
177  intros; unfold POW in e; simplify in e; cases e;
178  unfold orelation_of_relation in e3; simplify in e3; clear e e1 e2 e4;
179  intros 2; cases (e3 {(x)}); 
180  split; intro; [ lapply (s y); | lapply (s1 y); ]
181   [2,4: exists; [1,3:apply x] split; [1,3: assumption |*: change with (x=x); apply rule #]
182   |*: cases Hletin; cases x1; change in f3 with (x =_1 w); apply (. f3‡#); assumption;]
183 qed.
184
185
186 lemma currify: ∀A,B,C. binary_morphism1 A B C → A → unary_morphism1 B C.
187 intros; constructor 1; [ apply (b c); | intros; apply (#‡e);]
188 qed.
189
190 (*
191 alias symbol "singl" = "singleton".
192 alias symbol "eq" = "setoid eq".
193 lemma in_singleton_to_eq : ∀A:setoid.∀y,x:A.y ∈ {(x)} → (eq1 A) y x.
194 intros; apply sym1; apply f;
195 qed.  
196
197 lemma eq_to_in_singleton : ∀A:setoid.∀y,x:A.eq1 A y x → y ∈ {(x)}.
198 intros; apply (e^-1);
199 qed.
200 *)
201
202 interpretation "lifting singl" 'singl x = 
203  (fun11 ? (objs2 (POW ?))  (singleton ?) x).
204
205 theorem POW_full: ∀S,T.∀f. exT22 ? (λg. map_arrows2 ?? POW S T g = f).
206  intros; exists;
207   [ constructor 1; constructor 1;
208      [ apply (λx:carr S.λy:carr T. y ∈ f {(x)});
209      | intros; unfold FunClass_1_OF_carr2; lapply (.= e1‡#);
210         [4: apply mem; |6: apply Hletin;|1,2,3,5: skip]
211        lapply (#‡prop11 ?? f ?? (†e)); [6: apply Hletin; |*:skip ]]  
212      | whd; split; whd; intro; simplify; unfold map_arrows2; simplify; 
213         [ split;
214            [ change with (∀a1.(∀x. a1 ∈ f (singleton S x) → x ∈ a) → a1 ∈ f⎻* a);
215            | change with (∀a1.a1 ∈ f⎻* a → (∀x.a1 ∈ f (singleton S x) → x ∈ a)); ]
216         | split;
217            [ change with (∀a1.(∃y:carr T. y ∈ f (singleton S a1) ∧ y ∈ a) → a1 ∈ f⎻ a);
218            | change with (∀a1.a1 ∈ f⎻ a → (∃y:carr T.y ∈ f (singleton S a1) ∧ y ∈ a)); ]
219         | split;
220            [ change with (∀a1.(∃x:carr S. a1 ∈ f (singleton S x) ∧ x ∈ a) → a1 ∈ f a);
221            | change with (∀a1.a1 ∈ f a → (∃x:carr S. a1 ∈ f (singleton S x) ∧ x ∈ a)); ]
222         | split;
223            [ change with (∀a1.(∀y. y ∈ f (singleton S a1) → y ∈ a) → a1 ∈ f* a);
224            | change with (∀a1.a1 ∈ f* a → (∀y. y ∈ f (singleton S a1) → y ∈ a)); ]]
225         [ intros; apply ((. (or_prop2 ?? f (singleton ? a1) a)^-1) ? a1);
226            [ intros 2; apply (f1 a2); change in f2 with (a2 ∈ f⎻ (singleton ? a1));
227              lapply (. (or_prop3 ?? f (singleton ? a2) (singleton ? a1)));
228               [ cases Hletin; change in x1 with (eq1 ? a1 w);
229                 apply (. x1‡#); assumption;
230               | exists; [apply a2] [change with (a2=a2); apply rule #; | assumption]]
231            | change with (a1 = a1); apply rule #; ]
232         | intros; apply ((. (or_prop2 ?? f (singleton ? a1) a)) ? x);
233            [ intros 2; change in f3 with (eq1 ? a1 a2); change with (a2 ∈ f⎻* a); apply (. f3^-1‡#);
234              assumption;
235            | lapply (. (or_prop3 ?? f (singleton ? x) (singleton ? a1))^-1);
236               [ cases Hletin; change in x1 with (eq1 ? x w);
237                 change with (x ∈ f⎻ (singleton ? a1)); apply (. x1‡#); assumption;
238               | exists; [apply a1] [assumption | change with (a1=a1); apply rule #; ]]]
239         | intros; cases e; cases x; clear e x;
240           lapply (. (or_prop3 ?? f (singleton ? a1) a)^-1);
241            [ cases Hletin; change in x with (eq1 ? a1 w1); apply (. x‡#); assumption;
242            | exists; [apply w] assumption ]
243         | intros; lapply (. (or_prop3 ?? f (singleton ? a1) a));
244            [ cases Hletin; exists; [apply w] split; assumption;
245            | exists; [apply a1] [change with (a1=a1); apply rule #; | assumption ]] 
246         | intros; cases e; cases x; clear e x;
247           apply (f_image_monotone ?? f (singleton ? w) a ? a1);
248            [ intros 2; change in f3 with (eq1 ? w a2); change with (a2 ∈ a);
249              apply (. f3^-1‡#); assumption;
250            | assumption; ]
251         | intros; lapply (. (or_prop3 ?? f a (singleton ? a1))^-1);
252            [ cases Hletin; exists; [apply w] split;
253               [ lapply (. (or_prop3 ?? f (singleton ? w) (singleton ? a1)));
254                  [ cases Hletin1; change in x3 with (eq1 ? a1 w1); apply (. x3‡#); assumption;
255                  | exists; [apply w] [change with (w=w); apply rule #; | assumption ]]
256               | assumption ]
257            | exists; [apply a1] [ assumption; | change with (a1=a1); apply rule #;]]
258         | intros; apply ((. (or_prop1 ?? f (singleton ? a1) a)^-1) ? a1);
259            [ apply f1; | change with (a1=a1); apply rule #; ]
260         | intros; apply ((. (or_prop1 ?? f (singleton ? a1) a)) ? y);
261            [ intros 2; change in f3 with (eq1 ? a1 a2); change with (a2 ∈ f* a);
262              apply (. f3^-1‡#); assumption;
263            | assumption ]]]
264 qed.