]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/models/nat_dedekind_sigma_complete.ma
9681b4d93618b6fc4ef7a66dd21796b2fe5d0b3e
[helm.git] / helm / software / matita / contribs / dama / dama / models / nat_dedekind_sigma_complete.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_uniform.ma".
16 include "supremum.ma".
17 include "nat/le_arith.ma".
18 include "russell_support.ma".
19
20 lemma hint1:
21  ∀l,u.sequence (Type_of_ordered_set (segment_ordered_set nat_ordered_set l u))
22    → sequence (hos_carr (os_l (segment_ordered_set nat_ordered_set l u))).
23 intros; assumption;
24 qed.   
25    
26 coercion hint1 nocomposites.   
27    
28 alias symbol "pi1" = "exT \fst".
29 alias symbol "N" = "ordered set N".
30 lemma nat_dedekind_sigma_complete:
31   ∀l,u:ℕ.∀a:sequence {[l,u]}.a is_increasing → 
32     ∀x.x is_supremum a → ∃i.∀j.i ≤ j → \fst x = \fst (a j).
33 intros 5; cases x (s Hs); clear x; letin X ≝ ≪s,Hs≫; 
34 fold normalize X; intros; cases H1; 
35 alias symbol "N" = "Natural numbers".
36 letin spec ≝ (λi,j:ℕ.(u≤i ∧ s = \fst (a j)) ∨ (i < u ∧ s+i ≤ u + \fst (a j))); 
37 (* s - aj <= max 0 (u - i) *)  
38 letin m ≝ (hide ? (
39   let rec aux i ≝
40     match i with
41     [ O ⇒ O
42     | S m ⇒ 
43         let pred ≝ aux m in
44         let apred ≝ a pred in 
45         match cmp_nat s (\fst apred) with
46         [ cmp_le _ ⇒ pred
47         | cmp_gt nP ⇒ \fst (H3 apred ?)]]
48   in aux 
49    :
50    ∀i:nat.∃j:nat.spec i j));[whd; apply nP;] unfold spec in aux ⊢ %;
51 [3: unfold X in H2; clear H4 n aux spec H3 H1 H X;
52     cases u in H2 Hs a ⊢ %; intros (a Hs H);
53     [1: left; split; [apply le_n]
54         generalize in match H;
55         generalize in match Hs;
56         rewrite > (?:s = O); 
57         [2: cases Hs; lapply (os_le_to_nat_le ?? H1);
58             apply (symmetric_eq nat O s ?).apply (le_n_O_to_eq s ?).apply (Hletin).
59         |1: intros; lapply (os_le_to_nat_le (\fst (a O)) O (H2 O));
60             lapply (le_n_O_to_eq ? Hletin); assumption;]
61     |2: right; cases Hs; rewrite > (sym_plus s O); split; [apply le_S_S; apply le_O_n];
62         apply (trans_le ??? (os_le_to_nat_le ?? H1));
63         apply le_plus_n_r;] 
64 |2: clear H6; cut (s = \fst (a (aux n1))); [2:
65       cases (le_to_or_lt_eq ?? H5); [2: assumption]
66       cases (?:False); apply (H2 (aux n1) H6);] clear H5;
67       generalize in match Hcut; clear Hcut; intro H5;
68 |1: clear H6]
69 [2,1:
70     cases (aux n1) in H5 ⊢ %; intros;
71     change in match (a ≪w,H5≫) in H6 ⊢ % with (a w);
72     cases H5; clear H5; cases H7; clear H7;
73     [1: left; split; [ apply (le_S ?? H5); | assumption]
74     |3: cases (?:False); rewrite < H8 in H6; apply (not_le_Sn_n ? H6);
75     |*: cases (cmp_nat u (S n1));
76         [1,3: left; split; [1,3: assumption |2: assumption]
77             cut (u = S n1); [2: apply le_to_le_to_eq; assumption ]
78             clear H7 H5 H4;rewrite > Hcut in H8:(? ? (? % ?)); clear Hcut;
79             cut (s = S (\fst (a w)));
80             [2: apply le_to_le_to_eq; [2: assumption]
81                 change in H8 with (s + n1 ≤ S (n1 + \fst (a w)));
82                 rewrite > plus_n_Sm in H8; rewrite > sym_plus in H8;
83                 apply (le_plus_to_le ??? H8);]
84             cases (H3 (a w) H6);
85             change with (s = \fst (a w1));
86             change in H4 with (\fst (a w) < \fst (a w1));
87             apply le_to_le_to_eq; [ rewrite > Hcut; assumption ]
88             apply (os_le_to_nat_le (\fst (a w1)) s (H2 w1));
89         |*: right; split; try assumption;
90             [1: rewrite > sym_plus in ⊢ (? ? %);
91                 rewrite < H6; apply le_plus_r; assumption;
92             |2: cases (H3 (a w) H6);
93                 change with (s + S n1 ≤ u + \fst (a w1));rewrite < plus_n_Sm;
94                 apply (trans_le ??? (le_S_S ?? H8)); rewrite > plus_n_Sm;
95                 apply (le_plus ???? (le_n ?) H9);]]]]
96 clearbody m; unfold spec in m; clear spec;
97 letin find ≝ (
98  let rec find i u on u : nat ≝
99   match u with
100   [ O ⇒ (m i:nat)
101   | S w ⇒ match eqb (\fst (a (m i))) s with
102           [ true ⇒ (m i:nat)
103           | false ⇒ find (S i) w]]
104  in find
105  :
106   ∀i,bound.∃j.i + bound = u → s = \fst (a j));
107 [1: cases (find (S n) n2); intro; change with (s = \fst (a w));
108     apply H6; rewrite < H7; simplify; apply plus_n_Sm;
109 |2: intros; rewrite > (eqb_true_to_eq ?? H5); reflexivity
110 |3: intros; rewrite > sym_plus in H5; rewrite > H5; clear H5 H4 n n1;
111     cases (m u); cases H4; clear H4; cases H5; clear H5; [assumption]
112     cases (not_le_Sn_n ? H4)]
113 clearbody find; cases (find O u);
114 exists [apply w]; intros; change with (s = \fst (a j));
115 rewrite > (H4 ?); [2: reflexivity]
116 apply le_to_le_to_eq;
117 [1: apply os_le_to_nat_le;
118     apply (trans_increasing ? H ? ? (nat_le_to_os_le ?? H5));
119 |2: apply (trans_le ? s ?);[apply os_le_to_nat_le; apply (H2 j);]
120     rewrite < (H4 ?); [2: reflexivity] apply le_n;]
121 qed.
122
123 lemma hint2:
124  ∀l,u.sequence (Type_of_ordered_set (segment_ordered_set nat_ordered_set l u))
125    → sequence (hos_carr (os_r (segment_ordered_set nat_ordered_set l u))).
126 intros; assumption;
127 qed.   
128    
129 coercion hint2 nocomposites.   
130    
131 alias symbol "N" = "ordered set N".
132 axiom nat_dedekind_sigma_complete_r:
133   ∀l,u:ℕ.∀a:sequence {[l,u]}.a is_decreasing → 
134     ∀x.x is_infimum a → ∃i.∀j.i ≤ j → \fst x = \fst (a j).
135