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