]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/formal_topology/overlap/categories.ma
17c5f498f20df3af32034a893a73493076c89e01
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / categories.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 "cprop_connectives.ma".
16
17 record equivalence_relation (A:Type0) : Type1 ≝
18  { eq_rel:2> A → A → CProp0;
19    refl: reflexive ? eq_rel;
20    sym: symmetric ? eq_rel;
21    trans: transitive ? eq_rel
22  }.
23
24 record setoid : Type1 ≝
25  { carr:> Type0;
26    eq: equivalence_relation carr
27  }.
28
29 record equivalence_relation1 (A:Type1) : Type2 ≝
30  { eq_rel1:2> A → A → CProp1;
31    refl1: reflexive1 ? eq_rel1;
32    sym1: symmetric1 ? eq_rel1;
33    trans1: transitive1 ? eq_rel1
34  }.
35
36 record setoid1: Type2 ≝
37  { carr1:> Type1;
38    eq1: equivalence_relation1 carr1
39  }.
40
41 definition setoid1_of_setoid: setoid → setoid1.
42  intro;
43  constructor 1;
44   [ apply (carr s)
45   | constructor 1;
46     [ apply (eq_rel s);
47       apply (eq s)
48     | apply (refl s)
49     | apply (sym s)
50     | apply (trans s)]]
51 qed.
52
53 coercion setoid1_of_setoid.
54 prefer coercion Type_OF_setoid.
55
56 record equivalence_relation2 (A:Type2) : Type3 ≝
57  { eq_rel2:2> A → A → CProp2;
58    refl2: reflexive2 ? eq_rel2;
59    sym2: symmetric2 ? eq_rel2;
60    trans2: transitive2 ? eq_rel2
61  }.
62
63 record setoid2: Type3 ≝
64  { carr2:> Type2;
65    eq2: equivalence_relation2 carr2
66  }.
67
68 definition setoid2_of_setoid1: setoid1 → setoid2.
69  intro;
70  constructor 1;
71   [ apply (carr1 s)
72   | constructor 1;
73     [ apply (eq_rel1 s);
74       apply (eq1 s)
75     | apply (refl1 s)
76     | apply (sym1 s)
77     | apply (trans1 s)]]
78 qed.
79
80 coercion setoid2_of_setoid1.
81 prefer coercion Type_OF_setoid2. 
82 prefer coercion Type_OF_setoid. 
83 prefer coercion Type_OF_setoid1. 
84 (* we prefer 0 < 1 < 2 *)
85
86 record equivalence_relation3 (A:Type3) : Type4 ≝
87  { eq_rel3:2> A → A → CProp3;
88    refl3: reflexive3 ? eq_rel3;
89    sym3: symmetric3 ? eq_rel3;
90    trans3: transitive3 ? eq_rel3
91  }.
92
93 record setoid3: Type4 ≝
94  { carr3:> Type3;
95    eq3: equivalence_relation3 carr3
96  }.
97
98
99 interpretation "setoid3 eq" 'eq x y = (eq_rel3 _ (eq3 _) x y).
100 interpretation "setoid2 eq" 'eq x y = (eq_rel2 _ (eq2 _) x y).
101 interpretation "setoid1 eq" 'eq x y = (eq_rel1 _ (eq1 _) x y).
102 interpretation "setoid eq" 'eq x y = (eq_rel _ (eq _) x y).
103
104 notation < "hvbox(a break = \sub 1 b)" non associative with precedence 45
105 for @{ 'eq1 $a $b }.
106
107 notation < "hvbox(a break = \sub 2 b)" non associative with precedence 45
108 for @{ 'eq2 $a $b }.
109
110 notation < "hvbox(a break = \sub 3 b)" non associative with precedence 45
111 for @{ 'eq3 $a $b }.
112
113 notation > "hvbox(a break =_12 b)" non associative with precedence 45
114 for @{ eq_rel2 (carr2 (setoid2_of_setoid1 ?)) (eq2 (setoid2_of_setoid1 ?)) $a $b }.
115 notation > "hvbox(a break =_0 b)" non associative with precedence 45
116 for @{ eq_rel ? (eq ?) $a $b }.
117 notation > "hvbox(a break =_1 b)" non associative with precedence 45
118 for @{ eq_rel1 ? (eq1 ?) $a $b }.
119 notation > "hvbox(a break =_2 b)" non associative with precedence 45
120 for @{ eq_rel2 ? (eq2 ?) $a $b }.
121 notation > "hvbox(a break =_3 b)" non associative with precedence 45
122 for @{ eq_rel3 ? (eq3 ?) $a $b }.
123
124 interpretation "setoid3 eq explicit" 'eq3 x y = (eq_rel3 _ (eq3 _) x y).
125 interpretation "setoid2 eq explicit" 'eq2 x y = (eq_rel2 _ (eq2 _) x y).
126 interpretation "setoid1 eq explicit" 'eq1 x y = (eq_rel1 _ (eq1 _) x y).
127
128 interpretation "setoid3 symmetry" 'invert r = (sym3 ____ r).
129 interpretation "setoid2 symmetry" 'invert r = (sym2 ____ r).
130 interpretation "setoid1 symmetry" 'invert r = (sym1 ____ r).
131 interpretation "setoid symmetry" 'invert r = (sym ____ r).
132 notation ".= r" with precedence 50 for @{'trans $r}.
133 interpretation "trans3" 'trans r = (trans3 _____ r).
134 interpretation "trans2" 'trans r = (trans2 _____ r).
135 interpretation "trans1" 'trans r = (trans1 _____ r).
136 interpretation "trans" 'trans r = (trans _____ r).
137
138 record unary_morphism (A,B: setoid) : Type0 ≝
139  { fun1:1> A → B;
140    prop1: ∀a,a'. eq ? a a' → eq ? (fun1 a) (fun1 a')
141  }.
142
143 record unary_morphism1 (A,B: setoid1) : Type1 ≝
144  { fun11:1> A → B;
145    prop11: ∀a,a'. eq1 ? a a' → eq1 ? (fun11 a) (fun11 a')
146  }.
147
148 record unary_morphism2 (A,B: setoid2) : Type2 ≝
149  { fun12:1> A → B;
150    prop12: ∀a,a'. eq2 ? a a' → eq2 ? (fun12 a) (fun12 a')
151  }.
152
153 record unary_morphism3 (A,B: setoid3) : Type3 ≝
154  { fun13:1> A → B;
155    prop13: ∀a,a'. eq3 ? a a' → eq3 ? (fun13 a) (fun13 a')
156  }.
157
158 record binary_morphism (A,B,C:setoid) : Type0 ≝
159  { fun2:2> A → B → C;
160    prop2: ∀a,a',b,b'. eq ? a a' → eq ? b b' → eq ? (fun2 a b) (fun2 a' b')
161  }.
162
163 record binary_morphism1 (A,B,C:setoid1) : Type1 ≝
164  { fun21:2> A → B → C;
165    prop21: ∀a,a',b,b'. eq1 ? a a' → eq1 ? b b' → eq1 ? (fun21 a b) (fun21 a' b')
166  }.
167
168 record binary_morphism2 (A,B,C:setoid2) : Type2 ≝
169  { fun22:2> A → B → C;
170    prop22: ∀a,a',b,b'. eq2 ? a a' → eq2 ? b b' → eq2 ? (fun22 a b) (fun22 a' b')
171  }.
172
173 record binary_morphism3 (A,B,C:setoid3) : Type3 ≝
174  { fun23:2> A → B → C;
175    prop23: ∀a,a',b,b'. eq3 ? a a' → eq3 ? b b' → eq3 ? (fun23 a b) (fun23 a' b')
176  }.
177
178 notation "† c" with precedence 90 for @{'prop1 $c }.
179 notation "l ‡ r" with precedence 90 for @{'prop2 $l $r }.
180 notation "#" with precedence 90 for @{'refl}.
181 interpretation "prop1" 'prop1 c  = (prop1 _____ c).
182 interpretation "prop11" 'prop1 c = (prop11 _____ c).
183 interpretation "prop12" 'prop1 c = (prop12 _____ c).
184 interpretation "prop13" 'prop1 c = (prop13 _____ c).
185 interpretation "prop2" 'prop2 l r = (prop2 ________ l r).
186 interpretation "prop21" 'prop2 l r = (prop21 ________ l r).
187 interpretation "prop22" 'prop2 l r = (prop22 ________ l r).
188 interpretation "prop23" 'prop2 l r = (prop23 ________ l r).
189 interpretation "refl" 'refl = (refl ___).
190 interpretation "refl1" 'refl = (refl1 ___).
191 interpretation "refl2" 'refl = (refl2 ___).
192 interpretation "refl3" 'refl = (refl3 ___).
193
194 definition unary_morphism2_of_unary_morphism1: ∀S,T.unary_morphism1 S T → unary_morphism2 S T.
195  intros;
196  constructor 1;
197   [ apply (fun11 ?? u);
198   | apply (prop11 ?? u); ]
199 qed.
200
201 definition CPROP: setoid1.
202  constructor 1;
203   [ apply CProp0
204   | constructor 1;
205      [ apply Iff
206      | intros 1; split; intro; assumption
207      | intros 3; cases i; split; assumption
208      | intros 5; cases i; cases i1; split; intro;
209         [ apply (f2 (f x1)) | apply (f1 (f3 z1))]]]
210 qed.
211
212 definition CProp0_of_CPROP: carr1 CPROP → CProp0 ≝ λx.x.
213 coercion CProp0_of_CPROP.
214
215 alias symbol "eq" = "setoid1 eq".
216 definition fi': ∀A,B:CPROP. A = B → B → A.
217  intros; apply (fi ?? e); assumption.
218 qed.
219
220 notation ". r" with precedence 50 for @{'fi $r}.
221 interpretation "fi" 'fi r = (fi' __ r).
222
223 definition and_morphism: binary_morphism1 CPROP CPROP CPROP.
224  constructor 1;
225   [ apply And
226   | intros; split; intro; cases a1; split;
227      [ apply (if ?? e a2)
228      | apply (if ?? e1 b1)
229      | apply (fi ?? e a2)
230      | apply (fi ?? e1 b1)]]
231 qed.
232
233 interpretation "and_morphism" 'and a b = (fun21 ___ and_morphism a b).
234
235 definition or_morphism: binary_morphism1 CPROP CPROP CPROP.
236  constructor 1;
237   [ apply Or
238   | intros; split; intro; cases o; [1,3:left |2,4: right]
239      [ apply (if ?? e a1)
240      | apply (fi ?? e a1)
241      | apply (if ?? e1 b1)
242      | apply (fi ?? e1 b1)]]
243 qed.
244
245 interpretation "or_morphism" 'or a b = (fun21 ___ or_morphism a b).
246
247 definition if_morphism: binary_morphism1 CPROP CPROP CPROP.
248  constructor 1;
249   [ apply (λA,B. A → B)
250   | intros; split; intros;
251      [ apply (if ?? e1); apply f; apply (fi ?? e); assumption
252      | apply (fi ?? e1); apply f; apply (if ?? e); assumption]]
253 qed.
254
255
256 record category : Type1 ≝
257  { objs:> Type0;
258    arrows: objs → objs → setoid;
259    id: ∀o:objs. arrows o o;
260    comp: ∀o1,o2,o3. binary_morphism (arrows o1 o2) (arrows o2 o3) (arrows o1 o3);
261    comp_assoc: ∀o1,o2,o3,o4. ∀a12,a23,a34.
262     comp o1 o3 o4 (comp o1 o2 o3 a12 a23) a34 = comp o1 o2 o4 a12 (comp o2 o3 o4 a23 a34);
263    id_neutral_left: ∀o1,o2. ∀a: arrows o1 o2. comp ??? (id o1) a = a;
264    id_neutral_right: ∀o1,o2. ∀a: arrows o1 o2. comp ??? a (id o2) = a
265  }.
266
267 record category1 : Type2 ≝
268  { objs1:> Type1;
269    arrows1: objs1 → objs1 → setoid1;
270    id1: ∀o:objs1. arrows1 o o;
271    comp1: ∀o1,o2,o3. binary_morphism1 (arrows1 o1 o2) (arrows1 o2 o3) (arrows1 o1 o3);
272    comp_assoc1: ∀o1,o2,o3,o4. ∀a12,a23,a34.
273     comp1 o1 o3 o4 (comp1 o1 o2 o3 a12 a23) a34 = comp1 o1 o2 o4 a12 (comp1 o2 o3 o4 a23 a34);
274    id_neutral_right1: ∀o1,o2. ∀a: arrows1 o1 o2. comp1 ??? (id1 o1) a = a;
275    id_neutral_left1: ∀o1,o2. ∀a: arrows1 o1 o2. comp1 ??? a (id1 o2) = a
276  }.
277
278 record category2 : Type3 ≝
279  { objs2:> Type2;
280    arrows2: objs2 → objs2 → setoid2;
281    id2: ∀o:objs2. arrows2 o o;
282    comp2: ∀o1,o2,o3. binary_morphism2 (arrows2 o1 o2) (arrows2 o2 o3) (arrows2 o1 o3);
283    comp_assoc2: ∀o1,o2,o3,o4. ∀a12,a23,a34.
284     comp2 o1 o3 o4 (comp2 o1 o2 o3 a12 a23) a34 = comp2 o1 o2 o4 a12 (comp2 o2 o3 o4 a23 a34);
285    id_neutral_right2: ∀o1,o2. ∀a: arrows2 o1 o2. comp2 ??? (id2 o1) a = a;
286    id_neutral_left2: ∀o1,o2. ∀a: arrows2 o1 o2. comp2 ??? a (id2 o2) = a
287  }.
288
289 record category3 : Type4 ≝
290  { objs3:> Type3;
291    arrows3: objs3 → objs3 → setoid3;
292    id3: ∀o:objs3. arrows3 o o;
293    comp3: ∀o1,o2,o3. binary_morphism3 (arrows3 o1 o2) (arrows3 o2 o3) (arrows3 o1 o3);
294    comp_assoc3: ∀o1,o2,o3,o4. ∀a12,a23,a34.
295     comp3 o1 o3 o4 (comp3 o1 o2 o3 a12 a23) a34 = comp3 o1 o2 o4 a12 (comp3 o2 o3 o4 a23 a34);
296    id_neutral_right3: ∀o1,o2. ∀a: arrows3 o1 o2. comp3 ??? (id3 o1) a = a;
297    id_neutral_left3: ∀o1,o2. ∀a: arrows3 o1 o2. comp3 ??? a (id3 o2) = a
298  }.
299
300 notation "'ASSOC'" with precedence 90 for @{'assoc}.
301
302 interpretation "category2 composition" 'compose x y = (fun22 ___ (comp2 ____) y x).
303 interpretation "category2 assoc" 'assoc = (comp_assoc2 ________).
304 interpretation "category1 composition" 'compose x y = (fun21 ___ (comp1 ____) y x).
305 interpretation "category1 assoc" 'assoc = (comp_assoc1 ________).
306 interpretation "category composition" 'compose x y = (fun2 ___ (comp ____) y x).
307 interpretation "category assoc" 'assoc = (comp_assoc ________).
308
309 definition category2_of_category1: category1 → category2.
310  intro;
311  constructor 1;
312   [ apply (objs1 c);
313   | intros; apply (setoid2_of_setoid1 (arrows1 c o o1));
314   | apply (id1 c);
315   | intros;
316     constructor 1;
317      [ intros; apply (comp1 c o1 o2 o3 c1 c2);
318      | intros; whd in e e1 a a' b b'; change with (eq1 ? (b∘a) (b'∘a')); apply (e‡e1); ]
319   | intros; simplify; whd in a12 a23 a34; whd; apply rule (ASSOC);
320   | intros; simplify; whd in a; whd; apply id_neutral_right1;
321   | intros; simplify; whd in a; whd; apply id_neutral_left1; ]
322 qed.
323 (*coercion category2_of_category1.*)
324
325 record functor2 (C1: category2) (C2: category2) : Type3 ≝
326  { map_objs2:1> C1 → C2;
327    map_arrows2: ∀S,T. unary_morphism2 (arrows2 ? S T) (arrows2 ? (map_objs2 S) (map_objs2 T));
328    respects_id2: ∀o:C1. map_arrows2 ?? (id2 ? o) = id2 ? (map_objs2 o);
329    respects_comp2:
330      ∀o1,o2,o3.∀f1:arrows2 ? o1 o2.∀f2:arrows2 ? o2 o3.
331      map_arrows2 ?? (f2 ∘ f1) = map_arrows2 ?? f2 ∘ map_arrows2 ?? f1}.
332
333 definition functor2_setoid: category2 → category2 → setoid3.
334  intros (C1 C2);
335  constructor 1;
336   [ apply (functor2 C1 C2);
337   | constructor 1;
338      [ intros (f g);
339        apply (∀c:C1. cic:/matita/logic/equality/eq.ind#xpointer(1/1) ? (f c) (g c));
340      | simplify; intros; apply cic:/matita/logic/equality/eq.ind#xpointer(1/1/1);
341      | simplify; intros; apply cic:/matita/logic/equality/sym_eq.con; apply H;
342      | simplify; intros; apply cic:/matita/logic/equality/trans_eq.con;
343         [2: apply H; | skip | apply H1;]]]
344 qed.
345
346 definition functor2_of_functor2_setoid: ∀S,T. functor2_setoid S T → functor2 S T ≝ λS,T,x.x.
347 coercion functor2_of_functor2_setoid.
348
349 definition CAT2: category3.
350  constructor 1;
351   [ apply category2;
352   | apply functor2_setoid;
353   | intros; constructor 1;
354      [ apply (λx.x);
355      | intros; constructor 1;
356         [ apply (λx.x);
357         | intros; assumption;]
358      | intros; apply rule #;
359      | intros; apply rule #; ]
360   | intros; constructor 1;
361      [ intros; constructor 1;
362         [ intros; apply (c1 (c o));
363         | intros; constructor 1;
364            [ intro; apply (map_arrows2 ?? c1 ?? (map_arrows2 ?? c ?? c2));
365            | intros; apply (††e); ]
366         | intros; simplify;
367           apply (.= †(respects_id2 : ?));
368           apply (respects_id2 : ?);
369         | intros; simplify;
370           apply (.= †(respects_comp2 : ?));
371           apply (respects_comp2 : ?); ]
372         | intros; intro; simplify;
373           apply (cic:/matita/logic/equality/eq_ind.con ????? (e ?));
374           apply (cic:/matita/logic/equality/eq_ind.con ????? (e1 ?));
375           constructor 1; ]
376         | intros; intro; simplify; constructor 1;
377         | intros; intro; simplify; constructor 1;
378         | intros; intro; simplify; constructor 1; ]
379 qed.
380
381 definition category2_of_objs3_CAT2: objs3 CAT2 → category2 ≝ λx.x.
382 coercion category2_of_objs3_CAT2.
383
384 definition functor2_setoid_of_arrows3_CAT2: ∀S,T. arrows3 CAT2 S T → functor2_setoid S T ≝ λS,T,x.x.
385 coercion functor2_setoid_of_arrows3_CAT2.
386
387 definition unary_morphism_setoid: setoid → setoid → setoid.
388  intros;
389  constructor 1;
390   [ apply (unary_morphism s s1);
391   | constructor 1;
392      [ intros (f g); apply (∀a:s. eq ? (f a) (g a));
393      | intros 1; simplify; intros; apply refl;
394      | simplify; intros; apply sym; apply f;
395      | simplify; intros; apply trans; [2: apply f; | skip | apply f1]]]
396 qed.
397
398 definition SET: category1.
399  constructor 1;
400   [ apply setoid;
401   | apply rule (λS,T:setoid.setoid1_of_setoid (unary_morphism_setoid S T));
402   | intros; constructor 1; [ apply (λx:carr o.x); | intros; assumption ]
403   | intros; constructor 1; [ intros; constructor 1; [ apply (λx. c1 (c x)); | intros;
404      apply († (†e));]
405   | intros; whd; intros; simplify; whd in H1; whd in H;
406     apply trans; [ apply (b (a' a1)); | lapply (prop1 ?? b (a a1) (a' a1));
407      [ apply Hletin | apply (e a1); ]  | apply e1; ]]
408   | intros; whd; intros; simplify; apply refl;
409   | intros; simplify; whd; intros; simplify; apply refl;
410   | intros; simplify; whd; intros; simplify; apply refl;
411   ]
412 qed.
413
414 definition setoid_of_SET: objs1 SET → setoid ≝ λx.x.
415 coercion setoid_of_SET.
416
417 definition unary_morphism_setoid_of_arrows1_SET: 
418   ∀P,Q.arrows1 SET P Q → unary_morphism_setoid P Q  ≝ λP,Q,x.x.
419 coercion unary_morphism_setoid_of_arrows1_SET.
420
421 notation "hbox(a break ⇒ b)" right associative with precedence 20 for @{ 'Imply $a $b }.
422 interpretation "unary morphism" 'Imply a b = (arrows1 SET a b).
423
424 definition unary_morphism1_setoid1: setoid1 → setoid1 → setoid1.
425  intros;
426  constructor 1;
427   [ apply (unary_morphism1 s s1);
428   | constructor 1;
429      [ intros (f g);
430        alias symbol "eq" = "setoid1 eq".
431        apply (∀a: carr1 s. f a = g a);
432      | intros 1; simplify; intros; apply refl1;
433      | simplify; intros; apply sym1; apply f;
434      | simplify; intros; apply trans1; [2: apply f; | skip | apply f1]]]
435 qed.
436
437 definition unary_morphism1_of_unary_morphism1_setoid1 : 
438   ∀S,T. unary_morphism1_setoid1 S T → unary_morphism1 S T ≝ λP,Q,x.x.
439 coercion unary_morphism1_of_unary_morphism1_setoid1.
440
441 definition SET1: objs3 CAT2.
442  constructor 1;
443   [ apply setoid1;
444   | apply rule (λS,T.setoid2_of_setoid1 (unary_morphism1_setoid1 S T));
445   | intros; constructor 1; [ apply (λx.x); | intros; assumption ]
446   | intros; constructor 1; [ intros; constructor 1; [ apply (λx. c1 (c x)); | intros;
447      apply († (†e));]
448   | intros; whd; intros; simplify; whd in H1; whd in H;
449     apply trans1; [ apply (b (a' a1)); | lapply (prop11 ?? b (a a1) (a' a1));
450      [ apply Hletin | apply (e a1); ]  | apply e1; ]]
451   | intros; whd; intros; simplify; apply refl1;
452   | intros; simplify; whd; intros; simplify; apply refl1;
453   | intros; simplify; whd; intros; simplify; apply refl1;
454   ]
455 qed.
456
457 definition setoid1_of_SET1: objs2 SET1 → setoid1 ≝ λx.x.
458 coercion setoid1_of_SET1.
459
460 definition unary_morphism1_setoid1_of_arrows2_SET1: 
461   ∀P,Q.arrows2 SET1 P Q → unary_morphism1_setoid1 P Q ≝ λP,Q,x.x.
462 coercion unary_morphism1_setoid1_of_arrows2_SET1.
463  
464 variant objs2_of_category1: objs1 SET → objs2 SET1 ≝ setoid1_of_setoid.
465 coercion objs2_of_category1.
466
467 prefer coercion Type_OF_setoid. (* we prefer the lower carrier projection *)
468 prefer coercion Type_OF_objs1.
469
470 interpretation "unary morphism1" 'Imply a b = (arrows2 SET1 a b).