1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "dama/uniform.ma".
17 record ordered_uniform_space_ : Type ≝ {
19 ous_us_: uniform_space;
20 with_ : us_carr ous_us_ = bishop_set_of_ordered_set ous_os
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))]
33 coercion ous_unifspace.
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
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);
45 definition invert_os_relation ≝
46 λC:ordered_set.λU:C squareO → Prop.
47 λx:C squareO. U 〈\snd x,\fst x〉.
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).
53 lemma hint_segment: ∀O.
54 segment (Type_of_ordered_set O) →
55 segment (hos_carr (os_l O)).
59 coercion hint_segment nocomposites.
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;
67 coercion segment_square_of_ordered_set_square with 0 2 nocomposites.
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)〉.
75 coercion ordered_set_square_of_segment_square nocomposites.
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:{[s]} squareO.(P b → OP b) ∧ (OP b → P b));
84 lemma unrestrict: ∀O:ordered_uniform_space.∀s:‡O.∀U,u.∀x:{[s]} squareO.
85 restriction_agreement ? s U u → U x → u x.
86 intros 6; cases (H x); assumption;
89 lemma restrict: ∀O:ordered_uniform_space.∀s:‡O.∀U,u.∀x:{[s]} squareO.
90 restriction_agreement ? s U u → u x → U x.
91 intros 6; cases (H x); assumption;
94 lemma invert_restriction_agreement:
95 ∀O:ordered_uniform_space.∀s:‡O.
96 ∀U:{[s]} squareO → Prop.∀u:O squareO → Prop.
97 restriction_agreement ? s U u →
98 restriction_agreement ? s (\inv U) (\inv u).
100 generalize in match (H 〈h1,h〉); cases h; cases h1; simplify;
101 intro X; cases X; split; assumption;
105 ∀O:ordered_set.∀s:‡O.(bishop_set_of_ordered_set {[s]}) squareB → (bishop_set_of_ordered_set O) squareB ≝
106 λO:ordered_set.λs:‡O.λb:{[s]} squareO.〈\fst(\fst b),\fst(\snd b)〉.
108 coercion bs2_of_bss2 nocomposites.
112 ∀O:ordered_uniform_space.∀s:‡(ordered_set_OF_ordered_uniform_space O).
116 (bishop_set_of_ordered_set
118 (ordered_set_OF_ordered_uniform_space O) s))).
119 (\fst x) ≈ (\snd x) →
120 (\fst (bs2_of_bss2 (ordered_set_OF_ordered_uniform_space O) s x))
122 (\snd (bs2_of_bss2 (ordered_set_OF_ordered_uniform_space O) s x)).
123 intros 3; cases x (a b); clear x; simplify in match (\fst ?);
124 simplify in match (\snd ?); intros 2 (K H); apply K; clear K;
126 [ left; change in H1:(? ? % ?) with (\fst a);
127 change in H1:(? ? ? %) with (\fst b);
128 change in a with (hos_carr (half_segment_ordered_set ? s));
129 change in b with (hos_carr (half_segment_ordered_set ? s));
130 apply rule (x2sx_ ? s ?? H1);
131 | right; change in H1:(? ? % ?) with (\fst b);
132 change in H1:(? ? ? %) with (\fst a);
133 change in a with (hos_carr (half_segment_ordered_set ? s));
134 change in b with (hos_carr (half_segment_ordered_set ? s));
135 apply rule (x2sx_ ? s ?? H1);]
139 lemma segment_ordered_uniform_space:
140 ∀O:ordered_uniform_space.∀s:‡O.ordered_uniform_space.
141 intros (O s); apply mk_ordered_uniform_space;
142 [1: apply (mk_ordered_uniform_space_ {[s]});
143 [1: alias symbol "and" = "constructive and".
144 letin f ≝ (λP:{[s]} squareO → Prop. ∃OP:O squareO → Prop.
145 (us_unifbase O OP) ∧ restriction_agreement ?? P OP);
146 apply (mk_uniform_space (bishop_set_of_ordered_set {[s]}) f);
147 [1: intros (U H); intro x; simplify;
148 cases H (w Hw); cases Hw (Gw Hwp); clear H Hw; intro Hm;
149 lapply (us_phi1 O w Gw x (a2sa ??? Hm)) as IH;
150 apply (restrict ? s ??? Hwp IH);
151 |2: intros (U V HU HV); cases HU (u Hu); cases HV (v Hv); clear HU HV;
152 cases Hu (Gu HuU); cases Hv (Gv HvV); clear Hu Hv;
153 cases (us_phi2 O u v Gu Gv) (w HW); cases HW (Gw Hw); clear HW;
154 exists; [apply (λb:{[s]} squareB.w b)] split;
155 [1: unfold f; simplify; clearbody f;
156 exists; [apply w]; split; [assumption] intro b; simplify;
157 unfold segment_square_of_ordered_set_square;
158 cases b; intros; split; intros; assumption;
159 |2: intros 2 (x Hx); cases (Hw ? Hx); split;
160 [apply (restrict O s ??? HuU H)|apply (restrict O s ??? HvV H1);]]
161 |3: intros (U Hu); cases Hu (u HU); cases HU (Gu HuU); clear Hu HU;
162 cases (us_phi3 O u Gu) (w HW); cases HW (Gw Hwu); clear HW;
163 exists; [apply (λx:{[s]} squareB.w x)] split;
164 [1: exists;[apply w];split;[assumption] intros; simplify; intro;
165 unfold segment_square_of_ordered_set_square;
166 cases b; intros; split; intro; assumption;
167 |2: intros 2 (x Hx); apply (restrict O s ??? HuU); apply Hwu;
168 cases Hx (m Hm); exists[apply (\fst m)] apply Hm;]
169 |4: intros (U HU x); cases HU (u Hu); cases Hu (Gu HuU); clear HU Hu;
170 cases (us_phi4 O u Gu x) (Hul Hur);
172 [1: lapply (invert_restriction_agreement O s ?? HuU) as Ra;
173 apply (restrict O s ?? x Ra);
174 apply Hul; apply (unrestrict O s ??? HuU H);
175 |2: apply (restrict O s ??? HuU); apply Hur;
176 apply (unrestrict O s ??? (invert_restriction_agreement O s ?? HuU) H);]]
177 |2: simplify; reflexivity;]
178 |2: simplify; unfold convex; intros 3; cases s1; intros;
179 (* TODO: x2sx is for ≰, we need one for ≤ *)
180 cases H (u HU); cases HU (Gu HuU); clear HU H;
181 lapply depth=0 (ous_convex_l ?? Gu 〈\fst h,\fst h1〉 ???????) as K3;
182 [2: intro; apply H2; apply (x2sx_ (os_l O) s h h1 H);
183 |3: apply 〈\fst (\fst y),\fst (\snd y)〉;
184 |4: intro; change in H with (\fst (\fst y) ≰ \fst h1); apply H3;
185 apply (x2sx_ (os_l O) s (\fst y) h1 H);
186 |5: change with (\fst h ≤ \fst (\fst y)); intro; apply H4;
187 apply (x2sx_ (os_l O) s h (\fst y) H);
188 |6: change with (\fst (\snd y) ≤ \fst h1); intro; apply H5;
189 apply (x2sx_ (os_l O) s (\snd y) h1 H);
190 |7: change with (\fst (\fst y) ≤ \fst (\snd y)); intro; apply H6;
191 apply (x2sx_ (os_l O) s (\fst y) (\snd y) H);
192 |8: apply (restrict O s U u y HuU K3);
193 |1: apply (unrestrict O s ?? 〈h,h1〉 HuU H1);]
194 |3: simplify; unfold convex; intros 3; cases s1; intros;
195 cases H (u HU); cases HU (Gu HuU); clear HU H;
196 lapply depth=0 (ous_convex_r ?? Gu 〈\fst h,\fst h1〉 ???????) as K3;
197 [2: intro; apply H2; apply (x2sx_ (os_r O) s h h1 H);
198 |3: apply 〈\fst (\fst y),\fst (\snd y)〉;
200 apply (x2sx_ (os_r O) s (\fst y) h1 H);
202 apply (x2sx_ (os_r O) s h (\fst y) H);
204 apply (x2sx_ (os_r O) s (\snd y) h1 H);
206 apply (x2sx_ (os_r O) s (\fst y) (\snd y) H);
207 |8: apply (restrict O s U u y HuU K3);
208 |1: apply (unrestrict O s ?? 〈h,h1〉 HuU H1);]
212 interpretation "Ordered uniform space segment" 'segset a =
213 (segment_ordered_uniform_space _ a).
216 alias symbol "pi1" = "exT \fst".
217 lemma restric_uniform_convergence:
218 ∀O:ordered_uniform_space.∀s:‡O.
221 (⌊n, \fst (a n)⌋ : sequence O) uniform_converges (\fst x) →
222 a uniform_converges x.
223 intros 7; cases H1; cases H2; clear H2 H1;
224 cases (H ? H3) (m Hm); exists [apply m]; intros;
225 apply (restrict ? s ??? H4); apply (Hm ? H1);
228 definition order_continuity ≝
229 λC:ordered_uniform_space.∀a:sequence C.∀x:C.
230 (a ↑ x → a uniform_converges x) ∧ (a ↓ x → a uniform_converges x).
232 lemma hint_boh1: ∀O. Type_OF_ordered_uniform_space O → hos_carr (os_l O).
236 coercion hint_boh1 nocomposites.
238 lemma hint_boh2: ∀O:ordered_uniform_space. hos_carr (os_l O) → Type_OF_ordered_uniform_space O.
242 coercion hint_boh2 nocomposites.