]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/models/nat_order_continuous.ma
general reorganization and first (unconditional) proof of lebesgue over naturals
[helm.git] / helm / software / matita / contribs / dama / dama / models / nat_order_continuous.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 "models/nat_dedekind_sigma_complete.ma".
16 include "models/nat_ordered_uniform.ma".
17
18 lemma nat_us_is_oc: ∀l,u:ℕ.order_continuity (segment_ordered_uniform_space ℕ l u).
19 intros 4; split; intros 3; cases H1; cases H2; clear H2 H1; cases H; clear H;
20 [1: cases (nat_dedekind_sigma_complete l u a H1 ? H2); 
21     exists [apply w1]; intros; 
22     apply (restrict nat_ordered_uniform_space l u ??? H4);     
23     generalize in match (H ? H5);
24     cases x; intro; 
25     generalize in match (refl_eq ? (a i):a i = a i);
26     generalize in match (a i) in ⊢ (?? % ? → %); intro X; cases X; clear X;  
27     intro; rewrite < H9 in H7; simplify; rewrite < H7; simplify;
28     apply (us_phi1 nat_uniform_space ? H3); simplify; apply eq_reflexive;
29 |2: cases (nat_dedekind_sigma_complete_r l u a H1 ? H2); 
30     exists [apply w1]; intros; 
31     apply (restrict nat_ordered_uniform_space l u ??? H4);     
32     generalize in match (H ? H5);
33     cases x; intro; 
34     generalize in match (refl_eq ? (a i):a i = a i);
35     generalize in match (a i) in ⊢ (?? % ? → %); intro X; cases X; clear X;  
36     intro; rewrite < H9 in H7; simplify; rewrite < H7; simplify;
37     apply (us_phi1 nat_uniform_space ? H3); simplify; apply eq_reflexive;]
38 qed.
39