]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/property_exhaustivity.ma
exhaustivity completed
[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 (* Fact 2.18 *)
44 lemma segment_cauchy:
45   ∀C:ordered_uniform_space.∀l,u:C.∀a:sequence {[l,u]}.
46     a is_cauchy → (λn:nat.fst (a n)) is_cauchy.
47 intros 7; 
48 alias symbol "pi1" (instance 3) = "pair pi1".
49 apply (H (λx:{[l,u]} square.U 〈fst (fst x),fst (snd x)〉));
50 (unfold segment_ordered_uniform_space; simplify);
51 exists [apply U] split; [assumption;]
52 intro; cases b; intros; simplify; split; intros; assumption;
53 qed.       
54
55 (* Lemma 3.8 *)
56 lemma restrict_uniform_convergence:
57   ∀C:ordered_uniform_space.property_sigma C →
58     ∀l,u:C.exhaustive {[l,u]} →
59      ∀a:sequence {[l,u]}.∀x:C. (λn.fst (a n)) ↑ x → 
60       x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges (sig_in ?? x h).
61 intros; cases H2 (Ha Hx); clear H2; cases Hx; split;
62 [1: split;
63     [1: apply (supremum_is_upper_bound C ?? Hx u); 
64         apply (segment_upperbound ? l);
65     |2: apply (le_transitive ?? (fst (a 0))); [2: apply H2;]
66         apply (segment_lowerbound ?l u);]
67 |2: intros;
68     lapply (uparrow_upperlocated ? a (sig_in ?? x h)) as Ha1;
69       [2: apply segment_preserves_uparrow;split; assumption;] 
70     lapply (segment_preserves_supremum ?l u a (sig_in ??? h)) as Ha2; 
71       [2:split; assumption]; cases Ha2; clear Ha2;
72     cases (H1 a a); lapply (H6 H4 Ha1) as HaC;
73     lapply (segment_cauchy ? l u ? HaC) as Ha;
74     lapply (sigma_cauchy ? H  ? x ? Ha); [split; assumption]
75     apply restric_uniform_convergence; assumption;]
76 qed.
77