]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/property_exhaustivity.ma
fc7121349d89c6917a7384710b211dc29a4a8e01
[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; cases H2 (Ha Hx); clear H2; cases Hx; split;
96 [1: apply prove_in_segment; 
97     lapply depth=0 (under_wloss_upperbound (os_l C) ?? (segment_upperbound s a) O) as W1;
98     lapply depth=0 (under_wloss_upperbound (os_r C) ?? (h_segment_upperbound (os_r C) s a) O) as W2;
99     lapply (H2 O); simplify in Hletin; simplify in W2 W1;
100     cases a in Hletin W2 W1; simplify; cases (f O); simplify; intros;
101     whd in H6:(? % ? ? ? ?);
102     simplify in H6:(%);
103     cases (wloss_prop (os_l C)); rewrite <H8 in H5 H6 ⊢ %;
104     [ change in H6 with (le (os_l C) (seg_l_ (os_l C) s) w);
105       apply (le_transitive ??? H6 H7);
106     | apply (le_transitive (seg_u_ (os_l C) s) w x H6 H7);
107     |  
108       lapply depth=0 (supremum_is_upper_bound ? x Hx (seg_u_ (os_l C) s)) as K;    
109       lapply depth=0 (under_wloss_upperbound (os_l C) ?? (segment_upperbound s a));
110       apply K; intro; lapply (Hletin n); unfold; unfold in Hletin1;
111       rewrite < H8 in Hletin1; intro; apply Hletin1; clear Hletin1; apply H9;
112     | 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;    
113       lapply depth=0 (under_wloss_upperbound (os_r C) ?? (h_segment_upperbound (os_r C) s a));
114       apply K; intro; lapply (Hletin n); unfold; unfold in Hletin1;
115 whd in Hletin1:(? % ? ? ? ?);
116 simplify in Hletin1:(%);
117       rewrite < H8 in Hletin1; intro; apply Hletin1; clear Hletin1; apply H9;
118   
119       
120         apply (segment_upperbound ? l);
121     generalize in match (H2 O); generalize in match Hx; unfold supremum;
122     unfold upper_bound; whd in ⊢ (?→%→?); rewrite < H4;
123     split; unfold; rewrite < H4; simplify;
124       [1: lapply (infimum_is_lower_bound ? ? Hx u); 
125
126
127
128 split;
129     [1: apply (supremum_is_upper_bound ? x Hx u); 
130         apply (segment_upperbound ? l);
131     |2: apply (le_transitive l ? x ? (H2 O));
132         apply (segment_lowerbound ? l u a 0);]
133 |2: intros;
134     lapply (uparrow_upperlocated a ≪x,h≫) as Ha1;
135       [2: apply (segment_preserves_uparrow C l u);split; assumption;] 
136     lapply (segment_preserves_supremum C l u a ≪?,h≫) as Ha2; 
137       [2:split; assumption]; cases Ha2; clear Ha2;
138     cases (H1 a a); lapply (H6 H4 Ha1) as HaC;
139     lapply (segment_cauchy ? l u ? HaC) as Ha;
140     lapply (sigma_cauchy ? H  ? x ? Ha); [left; split; assumption]
141     apply restric_uniform_convergence; assumption;]
142 qed.
143
144 lemma hint_mah1:
145   ∀C. Type_OF_ordered_uniform_space1 C → hos_carr (os_r C).
146   intros; assumption; qed.
147   
148 coercion hint_mah1 nocomposites.
149
150 lemma hint_mah2:
151   ∀C. sequence (hos_carr (os_l C)) → sequence (hos_carr (os_r C)).
152   intros; assumption; qed.
153   
154 coercion hint_mah2 nocomposites.
155
156 lemma hint_mah3:
157   ∀C. Type_OF_ordered_uniform_space C → hos_carr (os_r C).
158   intros; assumption; qed.
159   
160 coercion hint_mah3 nocomposites.
161     
162 lemma hint_mah4:
163   ∀C. sequence (hos_carr (os_r C)) → sequence (hos_carr (os_l C)).
164   intros; assumption; qed.
165   
166 coercion hint_mah4 nocomposites.
167
168 lemma restrict_uniform_convergence_downarrow:
169   ∀C:ordered_uniform_space.property_sigma C →
170     ∀l,u:C.exhaustive {[l,u]} →
171      ∀a:sequence {[l,u]}.∀x: C. ⌊n,\fst (a n)⌋ ↓ x → 
172       x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges ≪x,h≫.
173 intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
174 [1: split;
175     [2: apply (infimum_is_lower_bound ? x Hx l); 
176         apply (segment_lowerbound ? l u);
177     |1: lapply (ge_transitive ? ? x ? (H2 O)); [apply u||assumption]
178         apply (segment_upperbound ? l u a 0);]
179 |2: intros;
180     lapply (downarrow_lowerlocated a ≪x,h≫) as Ha1;
181       [2: apply (segment_preserves_downarrow ? l u);split; assumption;]
182           lapply (segment_preserves_infimum C l u a ≪x,h≫) as Ha2; 
183       [2:split; assumption]; cases Ha2; clear Ha2;
184     cases (H1 a a); lapply (H7 H4 Ha1) as HaC;
185     lapply (segment_cauchy ? l u ? HaC) as Ha;
186     lapply (sigma_cauchy ? H  ? x ? Ha); [right; split; assumption]
187     apply restric_uniform_convergence; assumption;]
188 qed.