]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/property_exhaustivity.ma
cf3a5b0574cfe099fd04f79cb29f62ef087cb416
[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 (* Definition 3.7 *)
19 definition exhaustive ≝
20   λC:ordered_uniform_space.
21    ∀a,b:sequence C.
22      (a is_increasing → a is_upper_located → a is_cauchy) ∧
23      (b is_decreasing → b is_lower_located → b is_cauchy).
24
25 lemma segment_upperbound:
26   ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.u is_upper_bound (λn.fst (a n)).
27 intros 5; change with (fst (a n) ≤ u); cases (a n); cases H; assumption;
28 qed.
29
30 lemma segment_lowerbound:
31   ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.l is_lower_bound (λn.fst (a n)).
32 intros 5; change with (l ≤ fst (a n)); cases (a n); cases H; assumption;
33 qed.
34
35 lemma segment_preserves_uparrow:
36   ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.∀x,h. 
37     (λn.fst (a n)) ↑ x → a ↑ (sig_in ?? x h).
38 intros; cases H (Ha Hx); split [apply Ha] cases Hx; 
39 split; [apply H1] intros;
40 cases (H2 (fst y) H3); exists [apply w] assumption;
41 qed.
42
43 lemma segment_preserves_downarrow:
44   ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.∀x,h. 
45     (λn.fst (a n)) ↓ x → a ↓ (sig_in ?? x h).
46 intros; cases H (Ha Hx); split [apply Ha] cases Hx; 
47 split; [apply H1] intros;
48 cases (H2 (fst y) H3); exists [apply w] assumption;
49 qed.
50     
51 (* Fact 2.18 *)
52 lemma segment_cauchy:
53   ∀C:ordered_uniform_space.∀l,u:C.∀a:sequence {[l,u]}.
54     a is_cauchy → (λn:nat.fst (a n)) is_cauchy.
55 intros 7; 
56 alias symbol "pi1" (instance 3) = "pair pi1".
57 apply (H (λx:{[l,u]} square.U 〈fst (fst x),fst (snd x)〉));
58 (unfold segment_ordered_uniform_space; simplify);
59 exists [apply U] split; [assumption;]
60 intro; cases b; intros; simplify; split; intros; assumption;
61 qed.       
62
63 (* Lemma 3.8 *)
64 lemma restrict_uniform_convergence_uparrow:
65   ∀C:ordered_uniform_space.property_sigma C →
66     ∀l,u:C.exhaustive {[l,u]} →
67      ∀a:sequence {[l,u]}.∀x:C. (λn.fst (a n)) ↑ x → 
68       x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges (sig_in ?? x h).
69 intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
70 [1: split;
71     [1: apply (supremum_is_upper_bound C ?? Hx u); 
72         apply (segment_upperbound ? l);
73     |2: apply (le_transitive ?? (fst (a 0))); [2: apply H2;]
74         apply (segment_lowerbound ?l u);]
75 |2: intros;
76     lapply (uparrow_upperlocated ? a (sig_in ?? x h)) as Ha1;
77       [2: apply segment_preserves_uparrow;split; assumption;] 
78     lapply (segment_preserves_supremum ?l u a (sig_in ??? h)) as Ha2; 
79       [2:split; assumption]; cases Ha2; clear Ha2;
80     cases (H1 a a); lapply (H6 H4 Ha1) as HaC;
81     lapply (segment_cauchy ? l u ? HaC) as Ha;
82     lapply (sigma_cauchy ? H  ? x ? Ha); [left; split; assumption]
83     apply restric_uniform_convergence; assumption;]
84 qed.
85       
86 lemma restrict_uniform_convergence_downarrow:
87   ∀C:ordered_uniform_space.property_sigma C →
88     ∀l,u:C.exhaustive {[l,u]} →
89      ∀a:sequence {[l,u]}.∀x:C. (λn.fst (a n)) ↓ x → 
90       x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges (sig_in ?? x h).
91 intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
92 [1: split;
93     [2: apply (infimum_is_lower_bound C ?? Hx l); 
94         apply (segment_lowerbound ? l u);
95     |1: apply (le_transitive ?? (fst (a 0))); [apply H2;]
96         apply (segment_upperbound ? l u);]
97 |2: intros;
98     lapply (downarrow_lowerlocated ? a (sig_in ?? x h)) as Ha1;
99       [2: apply segment_preserves_downarrow;split; assumption;] 
100     lapply (segment_preserves_infimum ?l u a (sig_in ??? h)) as Ha2; 
101       [2:split; assumption]; cases Ha2; clear Ha2;
102     cases (H1 a a); lapply (H7 H4 Ha1) as HaC;
103     lapply (segment_cauchy ? l u ? HaC) as Ha;
104     lapply (sigma_cauchy ? H  ? x ? Ha); [right; split; assumption]
105     apply restric_uniform_convergence; assumption;]
106 qed.