]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/property_exhaustivity.ma
WIP
[helm.git] / helm / software / matita / contribs / dama / dama / property_exhaustivity.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 "ordered_uniform.ma".
16 include "property_sigma.ma".
17
18 lemma h_segment_upperbound:
19   ∀C:half_ordered_set.
20    ∀s:segment C.
21     ∀a:sequence (half_segment_ordered_set C s).
22      (seg_u C s) (upper_bound ? ⌊n,\fst (a n)⌋). 
23 intros; cases (wloss_prop C); unfold; rewrite < H; simplify; intro n;
24 cases (a n); simplify; unfold in H1; rewrite < H in H1; cases H1; 
25 simplify in H2 H3; rewrite < H in H2 H3; assumption;
26 qed.
27
28 notation "'segment_upperbound'" non associative with precedence 90 for @{'segment_upperbound}.
29 notation "'segment_lowerbound'" non associative with precedence 90 for @{'segment_lowerbound}.
30
31 interpretation "segment_upperbound" 'segment_upperbound = (h_segment_upperbound (os_l _)).
32 interpretation "segment_lowerbound" 'segment_lowerbound = (h_segment_upperbound (os_r _)).
33
34 lemma h_segment_preserves_uparrow:
35   ∀C:half_ordered_set.∀s:segment C.∀a:sequence (half_segment_ordered_set C s).
36    ∀x,h. uparrow C ⌊n,\fst (a n)⌋ x → uparrow (half_segment_ordered_set C s) a ≪x,h≫.
37 intros; cases H (Ha Hx); split;
38 [ intro n; intro H; apply (Ha n); apply (sx2x ???? H);
39 | cases Hx; split; 
40   [ intro n; intro H; apply (H1 n);apply (sx2x ???? H); 
41   | intros; cases (H2 (\fst y)); [2: apply (sx2x ???? H3);] 
42     exists [apply w] apply (x2sx ?? (a w) y H4);]]
43 qed.
44
45 notation "'segment_preserves_uparrow'" non associative with precedence 90 for @{'segment_preserves_uparrow}.
46 notation "'segment_preserves_downarrow'" non associative with precedence 90 for @{'segment_preserves_downarrow}.
47
48 interpretation "segment_preserves_uparrow" 'segment_preserves_uparrow = (h_segment_preserves_uparrow (os_l _)).
49 interpretation "segment_preserves_downarrow" 'segment_preserves_downarrow = (h_segment_preserves_uparrow (os_r _)).
50  
51 (* Fact 2.18 *)
52 lemma segment_cauchy:
53   ∀C:ordered_uniform_space.∀s:‡C.∀a:sequence {[s]}.
54     a is_cauchy → ⌊n,\fst (a n)⌋ is_cauchy.
55 intros 6; 
56 alias symbol "pi1" (instance 3) = "pair pi1".
57 alias symbol "pi2" = "pair pi2".
58 apply (H (λx:{[s]} squareB.U 〈\fst (\fst x),\fst (\snd x)〉));
59 (unfold segment_ordered_uniform_space; simplify);
60 exists [apply U] split; [assumption;]
61 intro; cases b; intros; simplify; split; intros; assumption;
62 qed.       
63
64 (* Definition 3.7 *)
65 definition exhaustive ≝
66   λC:ordered_uniform_space.
67    ∀a,b:sequence C.
68      (a is_increasing → a is_upper_located → a is_cauchy) ∧
69      (b is_decreasing → b is_lower_located → b is_cauchy).
70
71 lemma prove_in_segment: 
72  ∀O:ordered_set.∀s:segment (os_l O).∀x:O.
73    𝕝_s (λl.l ≤ x) → 𝕦_s (λu.x ≤ u) → x ∈ s.
74 intros; unfold; cases (wloss_prop (os_l O)); rewrite < H2;
75 split; assumption;
76 qed.
77
78 lemma under_wloss_upperbound: 
79  ∀C:half_ordered_set.∀s:segment C.∀a:sequence C.
80   seg_u C s (upper_bound C a) → 
81     ∀i.seg_u C s (λu.a i ≤≤ u).
82 intros; unfold in H; unfold;
83 cases (wloss_prop C); rewrite <H1 in H ⊢ %;
84 apply (H i);
85 qed. 
86
87
88 (* Lemma 3.8 NON DUALIZZATO *)
89 lemma restrict_uniform_convergence_uparrow:
90   ∀C:ordered_uniform_space.property_sigma C →
91     ∀s:segment (os_l C).exhaustive (segment_ordered_uniform_space C s) →
92      ∀a:sequence (segment_ordered_uniform_space C s).
93       ∀x:C. ⌊n,\fst (a n)⌋ ↑ x → 
94        in_segment (os_l C) s x ∧ ∀h:x ∈ s.a uniform_converges ≪x,h≫.
95 intros; split;
96 [1: unfold in H2; cases H2; clear H2;unfold in H3 H4; cases H4; clear H4; unfold in H2;
97     cases (wloss_prop (os_l C)) (W W); apply prove_in_segment; unfold; rewrite <W;
98     simplify;
99     [ apply (le_transitive ?? x ? (H2 O));  
100       lapply (under_wloss_upperbound (os_r C) ?? (h_segment_upperbound (os_r C) s a) O) as K;
101       unfold in K; whd in K:(?%????); simplify in K; rewrite <W in K; apply K;
102     | intro; cases (H5 ? H4); clear H5 H4;    
103       lapply (under_wloss_upperbound (os_l C) ?? (segment_upperbound s a) w) as K;
104       unfold in K; whd in K:(?%????); simplify in K; rewrite <W in K;
105       apply K; apply H6;
106     | intro; unfold in H4; rewrite <W in H4;
107       lapply depth=0 (H5 (seg_u_ (os_l C) s)) as k; unfold in k:(%???→?);
108       simplify in k; rewrite <W in k; lapply (k
109      simplify;intro; cases (H5 ? H4); clear H5 H4;    
110       lapply (under_wloss_upperbound (os_l C) ?? (segment_upperbound s a) w) as K;
111       unfold in K; whd in K:(?%????); simplify in K; rewrite <W in K;
112       apply K; apply H6;    
113       
114       
115       
116     cases H2 (Ha Hx); clear H2; cases Hx; split; 
117     lapply depth=0 (under_wloss_upperbound (os_l C) ?? (segment_upperbound s a) O) as W1;
118     lapply depth=0 (under_wloss_upperbound (os_r C) ?? (h_segment_upperbound (os_r C) s a) O) as W2;
119     lapply (H2 O); simplify in Hletin; simplify in W2 W1;
120     cases a in Hletin W2 W1; simplify; cases (f O); simplify; intros;
121     whd in H6:(? % ? ? ? ?);
122     simplify in H6:(%);
123     cases (wloss_prop (os_l C)); rewrite <H8 in H5 H6 ⊢ %;
124     [ change in H6 with (le (os_l C) (seg_l_ (os_l C) s) w);
125       apply (le_transitive ??? H6 H7);
126     | apply (le_transitive (seg_u_ (os_l C) s) w x H6 H7);
127     |  
128       lapply depth=0 (supremum_is_upper_bound ? x Hx (seg_u_ (os_l C) s)) as K;    
129       lapply depth=0 (under_wloss_upperbound (os_l C) ?? (segment_upperbound s a));
130       apply K; intro; lapply (Hletin n); unfold; unfold in Hletin1;
131       rewrite < H8 in Hletin1; intro; apply Hletin1; clear Hletin1; apply H9;
132     | lapply depth=0 (h_supremum_is_upper_bound (os_r C) ⌊n,\fst (a n)⌋ x Hx (seg_l_ (os_r C) s)) as K;    
133       lapply depth=0 (under_wloss_upperbound (os_r C) ?? (h_segment_upperbound (os_r C) s a));
134       apply K; intro; lapply (Hletin n); unfold; unfold in Hletin1;
135 whd in Hletin1:(? % ? ? ? ?);
136 simplify in Hletin1:(%);
137       rewrite < H8 in Hletin1; intro; apply Hletin1; clear Hletin1; apply H9;
138   
139       
140         apply (segment_upperbound ? l);
141     generalize in match (H2 O); generalize in match Hx; unfold supremum;
142     unfold upper_bound; whd in ⊢ (?→%→?); rewrite < H4;
143     split; unfold; rewrite < H4; simplify;
144       [1: lapply (infimum_is_lower_bound ? ? Hx u); 
145
146
147
148 split;
149     [1: apply (supremum_is_upper_bound ? x Hx u); 
150         apply (segment_upperbound ? l);
151     |2: apply (le_transitive l ? x ? (H2 O));
152         apply (segment_lowerbound ? l u a 0);]
153 |2: intros;
154     lapply (uparrow_upperlocated a ≪x,h≫) as Ha1;
155       [2: apply (segment_preserves_uparrow C l u);split; assumption;] 
156     lapply (segment_preserves_supremum C l u a ≪?,h≫) as Ha2; 
157       [2:split; assumption]; cases Ha2; clear Ha2;
158     cases (H1 a a); lapply (H6 H4 Ha1) as HaC;
159     lapply (segment_cauchy ? l u ? HaC) as Ha;
160     lapply (sigma_cauchy ? H  ? x ? Ha); [left; split; assumption]
161     apply restric_uniform_convergence; assumption;]
162 qed.
163
164 lemma hint_mah1:
165   ∀C. Type_OF_ordered_uniform_space1 C → hos_carr (os_r C).
166   intros; assumption; qed.
167   
168 coercion hint_mah1 nocomposites.
169
170 lemma hint_mah2:
171   ∀C. sequence (hos_carr (os_l C)) → sequence (hos_carr (os_r C)).
172   intros; assumption; qed.
173   
174 coercion hint_mah2 nocomposites.
175
176 lemma hint_mah3:
177   ∀C. Type_OF_ordered_uniform_space C → hos_carr (os_r C).
178   intros; assumption; qed.
179   
180 coercion hint_mah3 nocomposites.
181     
182 lemma hint_mah4:
183   ∀C. sequence (hos_carr (os_r C)) → sequence (hos_carr (os_l C)).
184   intros; assumption; qed.
185   
186 coercion hint_mah4 nocomposites.
187
188 lemma restrict_uniform_convergence_downarrow:
189   ∀C:ordered_uniform_space.property_sigma C →
190     ∀l,u:C.exhaustive {[l,u]} →
191      ∀a:sequence {[l,u]}.∀x: C. ⌊n,\fst (a n)⌋ ↓ x → 
192       x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges ≪x,h≫.
193 intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
194 [1: split;
195     [2: apply (infimum_is_lower_bound ? x Hx l); 
196         apply (segment_lowerbound ? l u);
197     |1: lapply (ge_transitive ? ? x ? (H2 O)); [apply u||assumption]
198         apply (segment_upperbound ? l u a 0);]
199 |2: intros;
200     lapply (downarrow_lowerlocated a ≪x,h≫) as Ha1;
201       [2: apply (segment_preserves_downarrow ? l u);split; assumption;]
202           lapply (segment_preserves_infimum C l u a ≪x,h≫) as Ha2; 
203       [2:split; assumption]; cases Ha2; clear Ha2;
204     cases (H1 a a); lapply (H7 H4 Ha1) as HaC;
205     lapply (segment_cauchy ? l u ? HaC) as Ha;
206     lapply (sigma_cauchy ? H  ? x ? Ha); [right; split; assumption]
207     apply restric_uniform_convergence; assumption;]
208 qed.