]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma
7292b4ecb336c0a86580a2e5f6559bf8bc5a771f
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-concrete_spaces.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 "o-basic_pairs.ma".
16 include "o-saturations.ma".
17
18 lemma xxx : ∀x.carr x → carr1 (setoid1_of_setoid x). intros; assumption; qed.
19 coercion xxx.
20
21 record concrete_space : Type ≝
22  { bp:> BP;
23    downarrow: unary_morphism (oa_P (form bp)) (oa_P (form bp));
24    downarrow_is_sat: is_saturation ? downarrow;
25    converges: ∀q1,q2.
26      or_f_minus ?? (⊩) q1 ∧ or_f_minus ?? (⊩) q2 = 
27      or_f_minus ?? (⊩) ((downarrow q1) ∧ (downarrow q2));
28    all_covered: (*⨍^-_bp*) or_f_minus ?? (⊩) (oa_one (form bp)) = oa_one (concr bp);
29    il2: ∀I:setoid.∀p:ums I (oa_P (form bp)).
30      downarrow (oa_join ? I (mk_unary_morphism ?? (λi:I. downarrow (p i)) ?)) =
31      oa_join ? I (mk_unary_morphism ?? (λi:I. downarrow (p i)) ?)
32  }.
33
34 definition bp': concrete_space → basic_pair ≝ λc.bp c.
35
36 coercion bp'.
37
38 record convergent_relation_pair (CS1,CS2: concrete_space) : Type ≝
39  { rp:> arrows1 ? CS1 CS2;
40    respects_converges:
41     ∀b,c.
42      extS ?? rp \sub\c (BPextS CS2 (b ↓ c)) =
43      BPextS CS1 ((extS ?? rp \sub\f b) ↓ (extS ?? rp \sub\f c));
44    respects_all_covered:
45     extS ?? rp\sub\c (BPextS CS2 (form CS2)) = BPextS CS1 (form CS1)
46  }.
47
48 definition rp' : ∀CS1,CS2. convergent_relation_pair CS1 CS2 → relation_pair CS1 CS2 ≝
49  λCS1,CS2,c. rp CS1 CS2 c.
50  
51 coercion rp'.
52
53 definition convergent_relation_space_setoid: concrete_space → concrete_space → setoid1.
54  intros;
55  constructor 1;
56   [ apply (convergent_relation_pair c c1)
57   | constructor 1;
58      [ intros;
59        apply (relation_pair_equality c c1 c2 c3);
60      | intros 1; apply refl1;
61      | intros 2; apply sym1; 
62      | intros 3; apply trans1]]
63 qed.
64
65 definition rp'': ∀CS1,CS2.convergent_relation_space_setoid CS1 CS2 → arrows1 BP CS1 CS2 ≝
66  λCS1,CS2,c.rp ?? c.
67
68 coercion rp''.
69
70 definition convergent_relation_space_composition:
71  ∀o1,o2,o3: concrete_space.
72   binary_morphism1
73    (convergent_relation_space_setoid o1 o2)
74    (convergent_relation_space_setoid o2 o3)
75    (convergent_relation_space_setoid o1 o3).
76  intros; constructor 1;
77      [ intros; whd in c c1 ⊢ %;
78        constructor 1;
79         [ apply (fun1 ??? (comp1 BP ???)); [apply (bp o2) |*: apply rp; assumption]
80         | intros;
81           change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
82           change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? (? ? ? (? ? ? %) ?) ?)))
83             with (c1 \sub \f ∘ c \sub \f);
84           change in ⊢ (? ? ? ? (? ? ? ? (? ? ? ? ? ? (? ? ? (? ? ? %) ?))))
85             with (c1 \sub \f ∘ c \sub \f);
86           apply (.= (extS_com ??????));
87           apply (.= (†(respects_converges ?????)));
88           apply (.= (respects_converges ?????));
89           apply (.= (†(((extS_com ??????) \sup -1)‡(extS_com ??????)\sup -1)));
90           apply refl1;
91         | change in ⊢ (? ? ? (? ? ? (? ? ? %) ?) ?) with (c1 \sub \c ∘ c \sub \c);
92           apply (.= (extS_com ??????));
93           apply (.= (†(respects_all_covered ???)));
94           apply (.= respects_all_covered ???);
95           apply refl1]
96      | intros;
97        change with (b ∘ a = b' ∘ a');
98        change in H with (rp'' ?? a = rp'' ?? a');
99        change in H1 with (rp'' ?? b = rp ?? b');
100        apply (.= (H‡H1));
101        apply refl1]
102 qed.
103
104 definition CSPA: category1.
105  constructor 1;
106   [ apply concrete_space
107   | apply convergent_relation_space_setoid
108   | intro; constructor 1;
109      [ apply id1
110      | intros;
111        unfold id; simplify;
112        apply (.= (equalset_extS_id_X_X ??));
113        apply (.= (†((equalset_extS_id_X_X ??)\sup -1‡
114                     (equalset_extS_id_X_X ??)\sup -1)));
115        apply refl1;
116      | apply (.= (equalset_extS_id_X_X ??));
117        apply refl1]
118   | apply convergent_relation_space_composition
119   | intros; simplify;
120     change with (a34 ∘ (a23 ∘ a12) = (a34 ∘ a23) ∘ a12);
121     apply (.= ASSOC1);
122     apply refl1
123   | intros; simplify;
124     change with (a ∘ id1 ? o1 = a);
125     apply (.= id_neutral_right1 ????);
126     apply refl1
127   | intros; simplify;
128     change with (id1 ? o2 ∘ a = a);
129     apply (.= id_neutral_left1 ????);
130     apply refl1]
131 qed.