]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/property_exhaustivity.ma
lebesge works
[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:half_ordered_set.∀s:segment O.∀x:O.
73    seg_l O s (λl.l ≤≤ x) → seg_u O s (λu.x ≤≤ u) → x ∈ s.
74 intros; unfold; cases (wloss_prop O); rewrite < H2;
75 split; assumption;
76 qed.
77
78 lemma h_uparrow_to_in_segment:
79   ∀C:half_ordered_set.
80    ∀s:segment C.
81      ∀a:sequence C.
82       (∀i.a i ∈ s) →
83        ∀x:C. uparrow C a x → 
84          in_segment C s x.
85 intros (C H a H1 x H2); unfold in H2; cases H2; clear H2;unfold in H3 H4; cases H4; clear H4; unfold in H2;
86 cases (wloss_prop C) (W W); apply prove_in_segment; unfold; rewrite <W;simplify;
87 [ apply (hle_transitive ??? x ? (H2 O)); lapply (H1 O) as K; unfold in K; rewrite <W in K;
88   cases K; unfold in H4 H6; rewrite <W in H6 H4; simplify in H4 H6; assumption;
89 | intro; cases (H5 ? H4); clear H5 H4;lapply(H1 w) as K; unfold in K; rewrite<W in K;
90   cases K; unfold in H5 H4; rewrite<W in H4 H5; simplify in H4 H5; apply (H5 H6);    
91 | apply (hle_transitive ??? x ?  (H2 O)); lapply (H1 0) as K; unfold in K; rewrite <W in K;
92   cases K; unfold in H4 H6; rewrite <W in H4 H6; simplify in H4 H6; assumption;
93 | intro; cases (H5 ? H4); clear H5 H4;lapply(H1 w) as K; unfold in K; rewrite<W in K;
94   cases K; unfold in H5 H4; rewrite<W in H4 H5; simplify in H4 H5; apply (H4 H6);]    
95 qed.
96
97 notation "'uparrow_to_in_segment'" non associative with precedence 90 for @{'uparrow_to_in_segment}.
98 notation "'downarrow_to_in_segment'" non associative with precedence 90 for @{'downarrow_to_in_segment}.
99
100 interpretation "uparrow_to_in_segment" 'uparrow_to_in_segment = (h_uparrow_to_in_segment (os_l _)).
101 interpretation "downarrow_to_in_segment" 'downarrow_to_in_segment = (h_uparrow_to_in_segment (os_r _)).
102  
103 (* Lemma 3.8 NON DUALIZZATO *)
104 lemma restrict_uniform_convergence_uparrow:
105   ∀C:ordered_uniform_space.property_sigma C →
106     ∀s:segment (os_l C).exhaustive (segment_ordered_uniform_space C s) →
107      ∀a:sequence (segment_ordered_uniform_space C s).
108       ∀x:C. ⌊n,\fst (a n)⌋ ↑ x → 
109        in_segment (os_l C) s x ∧ ∀h:x ∈ s.a uniform_converges ≪x,h≫.
110 intros; split;
111 [1: apply (uparrow_to_in_segment s ⌊n,\fst (a \sub n)⌋ ? x H2); 
112     simplify; intros; cases (a i); assumption;
113 |2: intros;
114     lapply (uparrow_upperlocated a ≪x,h≫) as Ha1;
115       [2: apply (segment_preserves_uparrow s); assumption;] 
116     lapply (segment_preserves_supremum s a ≪?,h≫ H2) as Ha2; 
117     cases Ha2; clear Ha2;
118     cases (H1 a a); lapply (H5 H3 Ha1) as HaC;
119     lapply (segment_cauchy C s ? HaC) as Ha;
120     lapply (sigma_cauchy ? H  ? x ? Ha); [left; assumption]
121     apply (restric_uniform_convergence C s ≪x,h≫ a Hletin)]
122 qed.
123
124 lemma hint_mah1:
125   ∀C. Type_OF_ordered_uniform_space1 C → hos_carr (os_r C).
126   intros; assumption; qed.
127   
128 coercion hint_mah1 nocomposites.
129
130 lemma hint_mah2:
131   ∀C. sequence (hos_carr (os_l C)) → sequence (hos_carr (os_r C)).
132   intros; assumption; qed.
133   
134 coercion hint_mah2 nocomposites.
135
136 lemma hint_mah3:
137   ∀C. Type_OF_ordered_uniform_space C → hos_carr (os_r C).
138   intros; assumption; qed.
139   
140 coercion hint_mah3 nocomposites.
141     
142 lemma hint_mah4:
143   ∀C. sequence (hos_carr (os_r C)) → sequence (hos_carr (os_l C)).
144   intros; assumption; qed.
145   
146 coercion hint_mah4 nocomposites.
147
148 lemma hint_mah5:
149   ∀C. segment (hos_carr (os_r C)) → segment (hos_carr (os_l C)).
150   intros; assumption; qed.
151   
152 coercion hint_mah5 nocomposites.
153
154 lemma hint_mah6:
155   ∀C. segment (hos_carr (os_l C)) → segment (hos_carr (os_r C)).
156   intros; assumption; qed.
157
158 coercion hint_mah6 nocomposites.
159
160 lemma restrict_uniform_convergence_downarrow:
161   ∀C:ordered_uniform_space.property_sigma C →
162     ∀s:segment (os_l C).exhaustive (segment_ordered_uniform_space C s) →
163      ∀a:sequence (segment_ordered_uniform_space C s).
164       ∀x:C. ⌊n,\fst (a n)⌋ ↓ x → 
165        in_segment (os_l C) s x ∧ ∀h:x ∈ s.a uniform_converges ≪x,h≫.
166 intros; split;       
167 [1: apply (downarrow_to_in_segment s ⌊n,\fst (a n)⌋ ? x); [2: apply H2]; 
168     simplify; intros; cases (a i); assumption;
169 |2: intros;
170     lapply (downarrow_lowerlocated a ≪x,h≫) as Ha1;
171       [2: apply (segment_preserves_downarrow s a x h H2);] 
172     lapply (segment_preserves_infimum s a ≪?,h≫ H2) as Ha2; 
173     cases Ha2; clear Ha2;
174     cases (H1 a a); lapply (H6 H3 Ha1) as HaC;
175     lapply (segment_cauchy C s ? HaC) as Ha;
176     lapply (sigma_cauchy ? H  ? x ? Ha); [right; assumption]
177     apply (restric_uniform_convergence C s ≪x,h≫ a Hletin)]
178 qed.