]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/formal_topology/overlap/o-basic_topologies.ma
Ooops, I forgot to commit this in the previous 3-4 commits.
[helm.git] / helm / software / matita / contribs / formal_topology / overlap / o-basic_topologies.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-algebra.ma".
16 include "o-saturations.ma".
17
18 record basic_topology: Type2 ≝
19  { carrbt:> OA;
20    A: carrbt ⇒ carrbt;
21    J: carrbt ⇒ carrbt;
22    A_is_saturation: is_saturation ? A;
23    J_is_reduction: is_reduction ? J;
24    compatibility: ∀U,V. (A U >< J V) = (U >< J V)
25  }.
26
27 lemma hint: OA → objs2 OA.
28  intro; apply t;
29 qed.
30 coercion hint.
31
32 record continuous_relation (S,T: basic_topology) : Type2 ≝
33  { cont_rel:> arrows2 OA S T;
34    (* reduces uses eq1, saturated uses eq!!! *)
35    reduced: ∀U. U = J ? U → cont_rel U = J ? (cont_rel U);
36    saturated: ∀U. U = A ? U → cont_rel⎻* U = A ? (cont_rel⎻* U)
37  }. 
38
39 definition continuous_relation_setoid: basic_topology → basic_topology → setoid2.
40  intros (S T); constructor 1;
41   [ apply (continuous_relation S T)
42   | constructor 1;
43      [ (*apply (λr,s:continuous_relation S T.∀b. eq1 (oa_P (carrbt S)) (A ? (r⎻ b)) (A ? (s⎻ b)));*)
44        apply (λr,s:continuous_relation S T.r⎻* ∘ (A S) = s⎻* ∘ (A ?));
45      | simplify; intros; apply refl2;
46      | simplify; intros; apply sym2; apply e
47      | simplify; intros; apply trans2; [2: apply e |3: apply e1; |1: skip]]]
48 qed.
49
50 definition cont_rel': ∀S,T: basic_topology. continuous_relation_setoid S T → arrows2 ? S T ≝ cont_rel.
51
52 coercion cont_rel'.
53
54 definition cont_rel'':
55  ∀S,T: basic_topology.
56   carr2 (continuous_relation_setoid S T) → ORelation_setoid (carrbt S) (carrbt T).
57  intros; apply rule cont_rel; apply c;
58 qed.
59
60 coercion cont_rel''.
61
62 (*
63 theorem continuous_relation_eq':
64  ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
65   a = a' → ∀X.a⎻* (A o1 X) = a'⎻* (A o1 X).
66  intros; apply oa_leq_antisym; intro; unfold minus_star_image; simplify; intros;
67   [ cut (ext ?? a a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
68     lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
69     cut (A ? (ext ?? a' a1) ⊆ A ? X); [2: apply (. (H ?)‡#); assumption]
70     lapply (fi ?? (A_is_saturation ???) Hcut);
71     apply (Hletin1 x); change with (x ∈ ext ?? a' a1); split; simplify;
72      [ apply I | assumption ]
73   | cut (ext ?? a' a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;]
74     lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut;
75     cut (A ? (ext ?? a a1) ⊆ A ? X); [2: apply (. (H ?)\sup -1‡#); assumption]
76     lapply (fi ?? (A_is_saturation ???) Hcut);
77     apply (Hletin1 x); change with (x ∈ ext ?? a a1); split; simplify;
78      [ apply I | assumption ]]
79 qed.
80
81 theorem continuous_relation_eq_inv':
82  ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2.
83   (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) → a=a'.
84  intros 6;
85  cut (∀a,a': continuous_relation_setoid o1 o2.
86   (∀X.a⎻* (A o1 X) = a'⎻* (A o1 X)) → 
87    ∀V:(oa_P (carrbt o2)). A o1 (a'⎻ V) ≤ A o1 (a⎻ V));
88   [2: clear b H a' a; intros;
89       lapply depth=0 (λV.saturation_expansive ??? (extS ?? a V)); [2: apply A_is_saturation;|skip]
90        (* fundamental adjunction here! to be taken out *)
91        cut (∀V:Ω \sup o2.V ⊆ minus_star_image ?? a (A ? (extS ?? a V)));
92         [2: intro; intros 2; unfold minus_star_image; simplify; intros;
93             apply (Hletin V1 x); whd; split; [ exact I | exists; [apply a1] split; assumption]]
94        clear Hletin;
95        cut (∀V:Ω \sup o2.V ⊆ minus_star_image ?? a' (A ? (extS ?? a V)));
96         [2: intro; apply (. #‡(H ?)); apply Hcut] clear H Hcut;
97        (* second half of the fundamental adjunction here! to be taken out too *)
98       intro; lapply (Hcut1 (singleton ? V)); clear Hcut1;
99       unfold minus_star_image in Hletin; unfold singleton in Hletin; simplify in Hletin;
100       whd in Hletin; whd in Hletin:(?→?→%); simplify in Hletin;
101       apply (if ?? (A_is_saturation ???));
102       intros 2 (x H); lapply (Hletin V ? x ?);
103        [ apply refl | cases H; assumption; ]
104       change with (x ∈ A ? (ext ?? a V));
105       apply (. #‡(†(extS_singleton ????)));
106       assumption;]
107  split; apply Hcut; [2: assumption | intro; apply sym1; apply H]
108 qed.
109 *)
110
111 definition continuous_relation_comp:
112  ∀o1,o2,o3.
113   continuous_relation_setoid o1 o2 →
114    continuous_relation_setoid o2 o3 →
115     continuous_relation_setoid o1 o3.
116  intros (o1 o2 o3 r s); constructor 1;
117   [ apply (s ∘ r);
118   | intros;
119     apply sym1;
120     change in match ((s ∘ r) U) with (s (r U));
121     (*<BAD>*) unfold FunClass_1_OF_Type_OF_setoid2;
122     unfold objs2_OF_basic_topology1; unfold hint;
123     letin reduced := reduced; clearbody reduced;
124     unfold uncurry_arrows in reduced ⊢ %; (*</BAD>*)
125     apply (.= (reduced : ?)\sup -1);
126      [ (*BAD*) change with (eq1 ? (r U) (J ? (r U)));
127        (* BAD U *) apply (.= (reduced ??? U ?)); [ assumption | apply refl1 ]
128      | apply refl1]
129   | intros;
130     apply sym1;
131     change in match ((s ∘ r)⎻* U) with (s⎻* (r⎻* U));
132     apply (.= (saturated : ?)\sup -1);
133      [ apply (.= (saturated : ?)); [ assumption | apply refl1 ]
134      | apply refl1]]
135 qed.
136
137 definition BTop: category2.
138  constructor 1;
139   [ apply basic_topology
140   | apply continuous_relation_setoid
141   | intro; constructor 1;
142      [ apply id2
143      | intros; apply e;
144      | intros; apply e;]
145   | intros; constructor 1;
146      [ apply continuous_relation_comp;
147      | intros; simplify;
148        change with ((b⎻* ∘ a⎻* ) ∘ A o1 = ((b'⎻* ∘ a'⎻* ) ∘ A o1)); 
149        change with (b⎻* ∘ (a⎻* ∘ A o1) = b'⎻* ∘ (a'⎻* ∘ A o1));
150        change in e with (a⎻* ∘ A o1 = a'⎻* ∘ A o1);
151        change in e1 with (b⎻* ∘ A o2 = b'⎻* ∘ A o2);
152        apply (.= e‡#);
153        intro x;
154        change with (b⎻* (a'⎻* (A o1 x)) = b'⎻*(a'⎻* (A o1 x)));
155        alias symbol "trans" = "trans1".
156        alias symbol "prop1" = "prop11".
157        alias symbol "invert" = "setoid1 symmetry".
158        lapply (.= †(saturated o1 o2 a' (A o1 x) : ?));
159         [3: apply (b⎻* ); | 5: apply Hletin; |1,2: skip;
160         |apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1); ]
161        change in e1 with (∀x.b⎻* (A o2 x) = b'⎻* (A o2 x));
162        apply (.= (e1 (a'⎻* (A o1 x))));
163        alias symbol "invert" = "setoid1 symmetry".
164        lapply (†((saturated ?? a' (A o1 x) : ?) ^ -1));
165         [2: apply (b'⎻* ); |4: apply Hletin; | skip;
166         |apply ((saturation_idempotent ?? (A_is_saturation o1) x)^-1);]]
167   | intros; simplify;
168     change with (((a34⎻* ∘ a23⎻* ) ∘ a12⎻* ) ∘ A o1 = ((a34⎻* ∘ (a23⎻* ∘ a12⎻* )) ∘ A o1));
169     apply rule (#‡ASSOC ^ -1);
170   | intros; simplify;
171     change with ((a⎻* ∘ (id2 ? o1)⎻* ) ∘ A o1 = a⎻* ∘ A o1);
172     apply (#‡(id_neutral_right2 : ?));
173   | intros; simplify;
174     change with (((id2 ? o2)⎻* ∘ a⎻* ) ∘ A o1 = a⎻* ∘ A o1);
175     apply (#‡(id_neutral_left2 : ?));]
176 qed.
177
178 definition btop_carr: BTop → Type1 ≝ λo:BTop. carr1 (oa_P (carrbt o)).
179 coercion btop_carr.
180
181 (*
182 (*CSC: unused! *)
183 (* this proof is more logic-oriented than set/lattice oriented *)
184 theorem continuous_relation_eqS:
185  ∀o1,o2:basic_topology.∀a,a': continuous_relation_setoid o1 o2.
186   a = a' → ∀X. A ? (extS ?? a X) = A ? (extS ?? a' X).
187  intros;
188  cut (∀a: arrows1 ? o1 ?.∀x. x ∈ extS ?? a X → ∃y:o2.y ∈ X ∧ x ∈ ext ?? a y);
189   [2: intros; cases f; clear f; cases H1; exists [apply w] cases x1; split;
190       try assumption; split; assumption]
191  cut (∀a,a':continuous_relation_setoid o1 o2.eq1 ? a a' → ∀x. x ∈ extS ?? a X → ∃y:o2. y ∈ X ∧ x ∈ A ? (ext ?? a' y));
192   [2: intros; cases (Hcut ?? f); exists; [apply w] cases x1; split; try assumption;
193       apply (. #‡(H1 ?));
194       apply (saturation_expansive ?? (A_is_saturation o1) (ext ?? a1 w) x);
195       assumption;] clear Hcut;
196  split; apply (if ?? (A_is_saturation ???)); intros 2;
197   [lapply (Hcut1 a a' H a1 f) | lapply (Hcut1 a' a (H \sup -1) a1 f)]
198   cases Hletin; clear Hletin; cases x; clear x;
199  cut (∀a: arrows1 ? o1 ?. ext ?? a w ⊆ extS ?? a X);
200   [2,4: intros 3; cases f3; clear f3; simplify in f5; split; try assumption;
201       exists [1,3: apply w] split; assumption;]
202  cut (∀a. A ? (ext o1 o2 a w) ⊆ A ? (extS o1 o2 a X));
203   [2,4: intros; apply saturation_monotone; try (apply A_is_saturation); apply Hcut;]
204  apply Hcut2; assumption.
205 qed.
206 *)