]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/ordered_uniform.ma
lebesge works
[helm.git] / helm / software / matita / contribs / dama / dama / ordered_uniform.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 "uniform.ma".
16
17 record ordered_uniform_space_ : Type ≝ {
18   ous_os:> ordered_set;
19   ous_us_: uniform_space;
20   with_ : us_carr ous_us_ = bishop_set_of_ordered_set ous_os
21 }.
22
23 lemma ous_unifspace: ordered_uniform_space_ → uniform_space.
24 intro X; apply (mk_uniform_space (bishop_set_of_ordered_set X));
25 unfold bishop_set_OF_ordered_uniform_space_;
26 [1: rewrite < (with_ X); simplify; apply (us_unifbase (ous_us_ X));
27 |2: cases (with_ X); simplify; apply (us_phi1 (ous_us_ X));
28 |3: cases (with_ X); simplify; apply (us_phi2 (ous_us_ X));
29 |4: cases (with_ X); simplify; apply (us_phi3 (ous_us_ X));
30 |5: cases (with_ X); simplify; apply (us_phi4 (ous_us_ X))]
31 qed.
32
33 coercion ous_unifspace.
34
35 record ordered_uniform_space : Type ≝ {
36   ous_stuff :> ordered_uniform_space_;
37   ous_convex_l: ∀U.us_unifbase ous_stuff U → convex (os_l ous_stuff) U;
38   ous_convex_r: ∀U.us_unifbase ous_stuff U → convex (os_r ous_stuff) U
39 }.   
40
41 definition half_ordered_set_OF_ordered_uniform_space : ordered_uniform_space → half_ordered_set.
42 intro; compose ordered_set_OF_ordered_uniform_space with os_l. apply (f o);
43 qed.
44
45 definition invert_os_relation ≝
46   λC:ordered_set.λU:C squareO → Prop.
47     λx:C squareO. U 〈\snd x,\fst x〉.
48
49 interpretation "relation invertion" 'invert a = (invert_os_relation _ a).
50 interpretation "relation invertion" 'invert_symbol = (invert_os_relation _).
51 interpretation "relation invertion" 'invert_appl a x = (invert_os_relation _ a x).
52
53 lemma hint_segment: ∀O.
54  segment (Type_of_ordered_set O) →
55  segment (hos_carr (os_l O)).
56 intros; assumption;
57 qed.
58
59 coercion hint_segment nocomposites.
60
61 lemma segment_square_of_ordered_set_square: 
62   ∀O:ordered_set.∀s:‡O.∀x:O squareO.
63    \fst x ∈ s → \snd x ∈ s → {[s]} squareO.
64 intros; split; exists; [1: apply (\fst x) |3: apply (\snd x)] assumption;
65 qed.
66
67 coercion segment_square_of_ordered_set_square with 0 2 nocomposites.
68
69 alias symbol "pi1" (instance 4) = "exT \fst".
70 alias symbol "pi1" (instance 2) = "exT \fst".
71 lemma ordered_set_square_of_segment_square : 
72  ∀O:ordered_set.∀s:‡O.{[s]} squareO → O squareO ≝ 
73   λO:ordered_set.λs:‡O.λb:{[s]} squareO.〈\fst(\fst b),\fst(\snd b)〉.
74
75 coercion ordered_set_square_of_segment_square nocomposites.
76
77 lemma restriction_agreement : 
78   ∀O:ordered_uniform_space.∀s:‡O.∀P:{[s]} squareO → Prop.∀OP:O squareO → Prop.Prop.
79 apply(λO:ordered_uniform_space.λs:‡O.
80        λP:{[s]} squareO → Prop. λOP:O squareO → Prop.
81           ∀b:O squareO.∀H1,H2.(P b → OP b) ∧ (OP b → P b));
82 [5,7: apply H1|6,8:apply H2]skip;
83 qed.
84
85 lemma unrestrict: ∀O:ordered_uniform_space.∀s:‡O.∀U,u.∀x:{[s]} squareO.
86   restriction_agreement ? s U u → U x → u x.
87 intros 6; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b b1 x; 
88 cases (H 〈w1,w2〉 H1 H2) (L _); intro Uw; apply L; apply Uw;
89 qed.
90
91 lemma restrict: ∀O:ordered_uniform_space.∀s:‡O.∀U,u.∀x:{[s]} squareO.
92   restriction_agreement ? s U u → u x → U x.
93 intros 5; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b1 b x;
94 intros (Ra uw); cases (Ra 〈w1,w2〉 H1 H2) (_ R); apply R; apply uw;
95 qed.
96
97 lemma invert_restriction_agreement: 
98   ∀O:ordered_uniform_space.∀s:‡O.
99    ∀U:{[s]} squareO → Prop.∀u:O squareO → Prop.
100     restriction_agreement ? s U u →
101     restriction_agreement ? s (\inv U) (\inv u).
102 intros 8; split; intro;
103 [1: apply (unrestrict ???? (segment_square_of_ordered_set_square ?? 〈\snd b,\fst b〉 H2 H1) H H3);
104 |2: apply (restrict ???? (segment_square_of_ordered_set_square ?? 〈\snd b,\fst b〉 H2 H1) H H3);]
105 qed. 
106
107 lemma bs2_of_bss2: 
108  ∀O:ordered_set.∀s:‡O.(bishop_set_of_ordered_set {[s]}) squareB → (bishop_set_of_ordered_set O) squareB ≝ 
109   λO:ordered_set.λs:‡O.λb:{[s]} squareO.〈\fst(\fst b),\fst(\snd b)〉.
110
111 coercion bs2_of_bss2 nocomposites.
112
113
114 lemma a2sa :
115  ∀O:ordered_uniform_space.∀s:‡(ordered_set_OF_ordered_uniform_space O).
116  ∀x:
117   bs_carr
118   (square_bishop_set
119    (bishop_set_of_ordered_set
120      (segment_ordered_set 
121        (ordered_set_OF_ordered_uniform_space O) s))).
122  (\fst x) ≈ (\snd x) →
123  (\fst (bs2_of_bss2 (ordered_set_OF_ordered_uniform_space O) s x))
124  ≈
125  (\snd (bs2_of_bss2 (ordered_set_OF_ordered_uniform_space O) s x)).
126 intros 3; cases x (a b); clear x; cases a (x Hx); cases b (y Hy); clear a b;
127 simplify; intros 2 (K H); apply K; clear K; whd; whd in H; cases H;
128 [left|right] apply x2sx; assumption;
129 qed.
130
131
132 lemma segment_ordered_uniform_space: 
133   ∀O:ordered_uniform_space.∀s:‡O.ordered_uniform_space.
134 intros (O s); apply mk_ordered_uniform_space;
135 [1: apply (mk_ordered_uniform_space_ {[s]});
136     [1: alias symbol "and" = "constructive and".
137         letin f ≝ (λP:{[s]} squareO → Prop. ∃OP:O squareO → Prop.
138                     (us_unifbase O OP) ∧ restriction_agreement ?? P OP);
139         apply (mk_uniform_space (bishop_set_of_ordered_set {[s]}) f);
140         [1: intros (U H); intro x; simplify; 
141             cases H (w Hw); cases Hw (Gw Hwp); clear H Hw; intro Hm;
142             lapply (us_phi1 O w Gw x (a2sa ??? Hm)) as IH;
143             apply (restrict ? s ??? Hwp IH); 
144         |2: intros (U V HU HV); cases HU (u Hu); cases HV (v Hv); clear HU HV;
145             cases Hu (Gu HuU); cases Hv (Gv HvV); clear Hu Hv;
146             cases (us_phi2 O u v Gu Gv) (w HW); cases HW (Gw Hw); clear HW;
147             exists; [apply (λb:{[s]} squareB.w b)] split;
148             [1: unfold f; simplify; clearbody f;
149                 exists; [apply w]; split; [assumption] intro b; simplify;
150                 unfold segment_square_of_ordered_set_square;
151                 cases b; intros; split; intros; assumption;
152             |2: intros 2 (x Hx); cases (Hw ? Hx); split;
153                 [apply (restrict O s ??? HuU H)|apply (restrict O s ??? HvV H1);]]
154         |3: intros (U Hu); cases Hu (u HU); cases HU (Gu HuU); clear Hu HU;
155             cases (us_phi3 O u Gu) (w HW); cases HW (Gw Hwu); clear HW;
156             exists; [apply (λx:{[s]} squareB.w x)] split;
157             [1: exists;[apply w];split;[assumption] intros; simplify; intro;
158                 unfold segment_square_of_ordered_set_square;
159                 cases b; intros; split; intro; assumption;
160             |2: intros 2 (x Hx); apply (restrict O s ??? HuU); apply Hwu; 
161                 cases Hx (m Hm); exists[apply (\fst m)] apply Hm;]
162         |4: intros (U HU x); cases HU (u Hu); cases Hu (Gu HuU); clear HU Hu;
163             cases (us_phi4 O u Gu x) (Hul Hur);
164             split; intros; 
165             [1: lapply (invert_restriction_agreement O s ?? HuU) as Ra;
166                 apply (restrict O s ?? x Ra);
167                 apply Hul; apply (unrestrict O s ??? HuU H);
168             |2: apply (restrict O s ??? HuU); apply Hur; 
169                 apply (unrestrict O s ??? (invert_restriction_agreement O s ?? HuU) H);]]
170     |2: simplify; reflexivity;]
171 |2: simplify; unfold convex; intros 3; cases s1; intros;
172     (* TODO: x2sx is for ≰, we need one for ≤ *) 
173     cases H (u HU); cases HU (Gu HuU); clear HU H;
174     lapply depth=0 (ous_convex_l ?? Gu 〈\fst h,\fst h1〉 ???????) as K3;
175     [2: intro; apply H2; apply (x2sx (os_l O) s h h1 H);
176     |3: apply 〈\fst (\fst y),\fst (\snd y)〉;
177     |4: intro; change in H with (\fst (\fst y) ≰ \fst h1); apply H3;
178         apply (x2sx (os_l O) s (\fst y) h1 H);
179     |5: change with (\fst h ≤ \fst (\fst y)); intro; apply H4;
180         apply (x2sx (os_l O) s h (\fst y) H);
181     |6: change with (\fst (\snd y) ≤ \fst h1); intro; apply H5;
182         apply (x2sx (os_l O) s (\snd y) h1 H);
183     |7: change with (\fst (\fst y) ≤ \fst (\snd y)); intro; apply H6;
184         apply (x2sx (os_l O) s (\fst y) (\snd y) H);
185     |8: apply (restrict O s U u y HuU K3);
186     |1: apply (unrestrict O s ?? 〈h,h1〉 HuU H1);]
187 |3: simplify; unfold convex; intros 3; cases s1; intros; (* TODO *)
188     cases H (u HU); cases HU (Gu HuU); clear HU H;
189     lapply depth=0 (ous_convex_r ?? Gu 〈\fst h,\fst h1〉 ???????) as K3;
190     [2: intro; apply H2; apply (x2sx (os_r O) s h h1 H);
191     |3: apply 〈\fst (\fst y),\fst (\snd y)〉;
192     |4: intro; (*change in H with (\fst (\fst y) ≱ \fst h1);*) apply H3;
193         apply (x2sx (os_r O) s (\fst y) h1 H);
194     |5: (*change with (\fst h ≥ \fst (\fst y));*) intro; apply H4;
195         apply (x2sx (os_r O) s h (\fst y) H);
196     |6: (*change with (\fst (\snd y) ≤ \fst h1);*) intro; apply H5;
197         apply (x2sx (os_r O) s (\snd y) h1 H);
198     |7: (*change with (\fst (\fst y) ≤ \fst (\snd y));*) intro; apply H6;
199         apply (x2sx (os_r O) s (\fst y) (\snd y) H);
200     |8: apply (restrict O s U u y HuU K3);
201     |1: apply (unrestrict O s ?? 〈h,h1〉 HuU H1);]
202 ]  
203 qed.
204
205 interpretation "Ordered uniform space segment" 'segset a = 
206  (segment_ordered_uniform_space _ a).
207
208 (* Lemma 3.2 *)
209 alias symbol "pi1" = "exT \fst".
210 lemma restric_uniform_convergence:
211  ∀O:ordered_uniform_space.∀s:‡O.
212   ∀x:{[s]}.
213    ∀a:sequence {[s]}.
214     (⌊n, \fst (a n)⌋ : sequence O) uniform_converges (\fst x) → 
215       a uniform_converges x.
216 intros 7; cases H1; cases H2; clear H2 H1;
217 cases (H ? H3) (m Hm); exists [apply m]; intros; 
218 apply (restrict ? s ??? H4); apply (Hm ? H1);
219 qed.
220
221 definition order_continuity ≝
222   λC:ordered_uniform_space.∀a:sequence C.∀x:C.
223     (a ↑ x → a uniform_converges x) ∧ (a ↓ x → a uniform_converges x).
224
225 lemma hint_boh1: ∀O. Type_OF_ordered_uniform_space O → hos_carr (os_l O).
226 intros; assumption;
227 qed.
228
229 coercion hint_boh1 nocomposites. 
230
231 lemma hint_boh2: ∀O:ordered_uniform_space. hos_carr (os_l O) → Type_OF_ordered_uniform_space O.
232 intros; assumption;
233 qed.
234
235 coercion hint_boh2 nocomposites. 
236