]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/ordered_uniform.ma
0bc8a32255cde1804c708a588c4f25f772215ed6
[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: ∀U.us_unifbase ous_stuff U → convex ous_stuff U
38 }.   
39
40 definition Type_of_ordered_uniform_space : ordered_uniform_space → Type.
41 intro; compose ordered_set_OF_ordered_uniform_space with os_l.
42 apply (hos_carr (f o));
43 qed.
44
45 definition Type_of_ordered_uniform_space_dual : ordered_uniform_space → Type.
46 intro; compose ordered_set_OF_ordered_uniform_space with os_r.
47 apply (hos_carr (f o));
48 qed.
49
50 coercion Type_of_ordered_uniform_space.
51 coercion Type_of_ordered_uniform_space_dual.
52
53 definition half_ordered_set_OF_ordered_uniform_space : ordered_uniform_space → half_ordered_set.
54 intro; compose ordered_set_OF_ordered_uniform_space with os_l. apply (f o);
55 qed.
56
57 definition invert_os_relation ≝
58   λC:ordered_set.λU:C squareO → Prop.
59     λx:C squareO. U 〈\snd x,\fst x〉.
60
61 interpretation "relation invertion" 'invert a = (invert_os_relation _ a).
62 interpretation "relation invertion" 'invert_symbol = (invert_os_relation _).
63 interpretation "relation invertion" 'invert_appl a x = (invert_os_relation _ a x).
64
65 lemma segment_square_of_ordered_set_square: 
66   ∀O:ordered_set.∀u,v:O.∀x:O squareO.
67    \fst x ∈ [u,v] → \snd x ∈ [u,v] → {[u,v]} squareO.
68 intros; split; exists; [1: apply (\fst x) |3: apply (\snd x)] assumption;
69 qed.
70
71 coercion segment_square_of_ordered_set_square with 0 2 nocomposites.
72
73 alias symbol "pi1" (instance 4) = "exT \fst".
74 alias symbol "pi1" (instance 2) = "exT \fst".
75 lemma ordered_set_square_of_segment_square : 
76  ∀O:ordered_set.∀u,v:O.{[u,v]} squareO → O squareO ≝ 
77   λO:ordered_set.λu,v:O.λb:{[u,v]} squareO.〈\fst(\fst b),\fst(\snd b)〉.
78
79 coercion ordered_set_square_of_segment_square nocomposites.
80
81 lemma restriction_agreement : 
82   ∀O:ordered_uniform_space.∀l,r:O.∀P:{[l,r]} squareO → Prop.∀OP:O squareO → Prop.Prop.
83 apply(λO:ordered_uniform_space.λl,r:O.
84        λP:{[l,r]} squareO → Prop. λOP:O squareO → Prop.
85           ∀b:O squareO.∀H1,H2.(P b → OP b) ∧ (OP b → P b));
86 [5,7: apply H1|6,8:apply H2]skip;
87 qed.
88
89 lemma unrestrict: ∀O:ordered_uniform_space.∀l,r:O.∀U,u.∀x:{[l,r]} squareO.
90   restriction_agreement ? l r U u → U x → u x.
91 intros 7; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b b1 x; 
92 cases (H 〈w1,w2〉 H1 H2) (L _); intro Uw; apply L; apply Uw;
93 qed.
94
95 lemma restrict: ∀O:ordered_uniform_space.∀l,r:O.∀U,u.∀x:{[l,r]} squareO.
96   restriction_agreement ? l r U u → u x → U x.
97 intros 6; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b1 b x;
98 intros (Ra uw); cases (Ra 〈w1,w2〉 H1 H2) (_ R); apply R; apply uw;
99 qed.
100
101 lemma invert_restriction_agreement: 
102   ∀O:ordered_uniform_space.∀l,r:O.
103    ∀U:{[l,r]} squareO → Prop.∀u:O squareO → Prop.
104     restriction_agreement ? l r U u →
105     restriction_agreement ? l r (\inv U) (\inv u).
106 intros 9; split; intro;
107 [1: apply (unrestrict ????? (segment_square_of_ordered_set_square ??? 〈\snd b,\fst b〉 H2 H1) H H3);
108 |2: apply (restrict ????? (segment_square_of_ordered_set_square ??? 〈\snd b,\fst b〉 H2 H1) H H3);]
109 qed. 
110
111 lemma bs2_of_bss2: 
112  ∀O:ordered_set.∀u,v:O.(bishop_set_of_ordered_set {[u,v]}) squareB → (bishop_set_of_ordered_set O) squareB ≝ 
113   λO:ordered_set.λu,v:O.λb:{[u,v]} squareO.〈\fst(\fst b),\fst(\snd b)〉.
114
115 coercion bs2_of_bss2 nocomposites.
116
117 (*
118 notation < "x \sub \neq" with precedence 91 for @{'bsss $x}.
119 interpretation "bs_of_ss" 'bsss x = (bs_of_ss _ _ _ x).
120 *)
121
122 (*
123 lemma ss_of_bs: 
124  ∀O:ordered_set.∀u,v:O.
125   ∀b:O squareO.\fst b ∈ [u,v] → \snd b ∈ [u,v] → {[u,v]} squareO ≝ 
126  λO:ordered_set.λu,v:O.
127   λb:O squareB.λH1,H2.〈≪\fst b,H1≫,≪\snd b,H2≫〉.
128 *)
129
130 (*
131 notation < "x \sub \nleq" with precedence 91 for @{'ssbs $x}.
132 interpretation "ss_of_bs" 'ssbs x = (ss_of_bs _ _ _ x _ _).
133 *)
134
135 lemma segment_ordered_uniform_space: 
136   ∀O:ordered_uniform_space.∀u,v:O.ordered_uniform_space.
137 intros (O l r); apply mk_ordered_uniform_space;
138 [1: apply (mk_ordered_uniform_space_ {[l,r]});
139     [1: alias symbol "and" = "constructive and".
140         letin f ≝ (λP:{[l,r]} squareO → Prop. ∃OP:O squareO → Prop.
141                     (us_unifbase O OP) ∧ restriction_agreement ??? P OP);
142         apply (mk_uniform_space (bishop_set_of_ordered_set {[l,r]}) f);
143         [1: intros (U H); intro x; simplify; 
144             cases H (w Hw); cases Hw (Gw Hwp); clear H Hw; intro Hm;
145             lapply (us_phi1 O w Gw x Hm) as IH;
146             apply (restrict ? l r ??? Hwp IH); STOP
147         |2: intros (U V HU HV); cases HU (u Hu); cases HV (v Hv); clear HU HV;
148             cases Hu (Gu HuU); cases Hv (Gv HvV); clear Hu Hv;
149             cases (us_phi2 ??? Gu Gv) (w HW); cases HW (Gw Hw); clear HW;
150             exists; [apply (λb:{[l,r]} square.w b)] split;
151             [1: unfold f; simplify; clearbody f;
152                 exists; [apply w]; split; [assumption] intro b; simplify;
153                 unfold segment_square_of_ordered_set_square;
154                 cases b; intros; split; intros; assumption;
155             |2: intros 2 (x Hx); cases (Hw ? Hx); split;
156                 [apply (restrict ?????? HuU H)|apply (restrict ?????? HvV H1);]]
157         |3: intros (U Hu); cases Hu (u HU); cases HU (Gu HuU); clear Hu HU;
158             cases (us_phi3 ?? Gu) (w HW); cases HW (Gw Hwu); clear HW;
159             exists; [apply (λx:{[l,r]} square.w x)] split;
160             [1: exists;[apply w];split;[assumption] intros; simplify; intro;
161                 unfold segment_square_of_ordered_set_square;
162                 cases b; intros; split; intro; assumption;
163             |2: intros 2 (x Hx); apply (restrict ?????? HuU); apply Hwu; 
164                 cases Hx (m Hm); exists[apply (\fst m)] apply Hm;]
165         |4: intros (U HU x); cases HU (u Hu); cases Hu (Gu HuU); clear HU Hu;
166             cases (us_phi4 ?? Gu x) (Hul Hur);
167             split; intros; 
168             [1: lapply (invert_restriction_agreement ????? HuU) as Ra;
169                 apply (restrict ????? x Ra);
170                 apply Hul; apply (unrestrict ?????? HuU H);
171             |2: apply (restrict ?????? HuU); apply Hur; 
172                 apply (unrestrict ?????? (invert_restriction_agreement ????? HuU) H);]]
173     |2: simplify; reflexivity;]
174 |2: simplify; unfold convex; intros;
175     cases H (u HU); cases HU (Gu HuU); clear HU H; 
176     lapply (ous_convex ?? Gu (bs_of_ss ? l r p) ? H2 (bs_of_ss ? l r y) H3) as Cu;
177     [1: apply (unrestrict ?????? HuU); apply H1;
178     |2: apply (restrict ?????? HuU Cu);]]
179 qed.
180
181 interpretation "Ordered uniform space segment" 'segment_set a b = 
182  (segment_ordered_uniform_space _ a b).
183
184 (* Lemma 3.2 *)
185 alias symbol "pi1" = "exT \fst".
186 lemma restric_uniform_convergence:
187  ∀O:ordered_uniform_space.∀l,u:O.
188   ∀x:{[l,u]}.
189    ∀a:sequence {[l,u]}.
190      ⌊n,\fst (a n)⌋ uniform_converges (\fst x) → 
191       a uniform_converges x.
192 intros 8; cases H1; cases H2; clear H2 H1;
193 cases (H ? H3) (m Hm); exists [apply m]; intros; 
194 apply (restrict ? l u ??? H4); apply (Hm ? H1);
195 qed.
196
197 definition order_continuity ≝
198   λC:ordered_uniform_space.∀a:sequence C.∀x:C.
199     (a ↑ x → a uniform_converges x) ∧ (a ↓ x → a uniform_converges x).