]> matita.cs.unibo.it Git - helm.git/blob - helm/matita/library/nat/minimization.ma
New version of the library. nth_prime, gcd, log.
[helm.git] / helm / matita / library / nat / minimization.ma
1 (**************************************************************************)
2 (*       ___                                                                *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||       A.Asperti, C.Sacerdoti Coen,                          *)
8 (*      ||A||       E.Tassi, S.Zacchiroli                                 *)
9 (*      \   /                                                             *)
10 (*       \ /        Matita is distributed under the terms of the          *)
11 (*        v         GNU Lesser General Public License Version 2.1         *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/nat/minimization".
16
17 include "nat/minus.ma".
18 include "datatypes/bool.ma".
19
20 let rec max i f \def
21   match (f i) with 
22   [ true \Rightarrow i
23   | false \Rightarrow 
24       match i with 
25       [ O \Rightarrow O
26       | (S j) \Rightarrow max j f ]].
27
28 theorem max_O_f : \forall f: nat \to bool. max O f = O.
29 intro. simplify.
30 elim f O.
31 simplify.reflexivity.
32 simplify.reflexivity.
33 qed. 
34
35 theorem max_S_max : \forall f: nat \to bool. \forall n:nat.
36 (f (S n) = true \land max (S n) f = (S n)) \lor 
37 (f (S n) = false \land max (S n) f = max n f).
38 intros.simplify.elim (f (S n)).
39 simplify.left.split.reflexivity.reflexivity.
40 simplify.right.split.reflexivity.reflexivity.
41 qed.
42
43 theorem le_max_n : \forall f: nat \to bool. \forall n:nat.
44 max n f \le n.
45 intros.elim n.rewrite > max_O_f.apply le_n.
46 simplify.elim (f (S n1)).simplify.apply le_n.
47 simplify.apply le_S.assumption.
48 qed.
49
50 theorem le_to_le_max : \forall f: nat \to bool. \forall n,m:nat.
51 n\le m  \to max n f \le max m f.
52 intros.elim H.
53 apply le_n.
54 apply trans_le ? (max n1 f).apply H2.
55 cut (f (S n1) = true \land max (S n1) f = (S n1)) \lor 
56 (f (S n1) = false \land max (S n1) f = max n1 f).
57 elim Hcut.elim H3.
58 rewrite > H5.
59 apply le_S.apply le_max_n.
60 elim H3.rewrite > H5.apply le_n.
61 apply max_S_max.
62 qed.
63
64 theorem f_m_to_le_max: \forall f: nat \to bool. \forall n,m:nat.
65 m\le n \to f m = true \to m \le max n f.
66 intros 3.elim n.apply le_n_O_elim m H.
67 apply le_O_n.
68 apply le_n_Sm_elim m n1 H1.
69 intro.apply trans_le ? (max n1 f).
70 apply H.apply le_S_S_to_le.assumption.assumption.
71 apply le_to_le_max.apply le_n_Sn.
72 intro.simplify.rewrite < H3. 
73 rewrite > H2.simplify.apply le_n.
74 qed.
75
76 definition max_spec \def \lambda f:nat \to bool.\lambda n: nat.
77 ex nat (\lambda i:nat. (le i n) \land (f i = true)) \to
78 (f n) = true \land (\forall i. i < n \to (f i = false)).
79
80 theorem f_max_true : \forall f:nat \to bool. \forall n:nat.
81 ex nat (\lambda i:nat. (le i n) \land (f i = true)) \to f (max n f) = true. 
82 intros 2.
83 elim n.elim H.elim H1.generalize in match H3.
84 apply le_n_O_elim a H2.intro.simplify.rewrite > H4.
85 simplify.assumption.
86 simplify.
87 apply bool_ind (\lambda b:bool.
88 (f (S n1) = b) \to (f ([\lambda b:bool.nat] match b in bool with
89 [ true \Rightarrow (S n1)
90 | false  \Rightarrow (max n1 f)])) = true) ? ? ?.
91 reflexivity.
92 simplify.intro.assumption.
93 simplify.intro.apply H.
94 elim H1.elim H3.generalize in match H5.
95 apply le_n_Sm_elim a n1 H4.
96 intros.
97 apply ex_intro nat ? a.
98 split.apply le_S_S_to_le.assumption.assumption.
99 intros.apply False_ind.apply not_eq_true_false ?.
100 rewrite < H2.rewrite < H7.rewrite > H6. reflexivity.
101 qed.
102
103 theorem lt_max_to_false : \forall f:nat \to bool. 
104 \forall n,m:nat. (max n f) < m \to m \leq n \to f m = false.
105 intros 2.
106 elim n.absurd le m O.assumption.
107 cut O < m.apply lt_O_n_elim m Hcut.exact not_le_Sn_O.
108 rewrite < max_O_f f.assumption.
109 generalize in match H1.
110 (* ?? non posso generalizzare su un goal implicativo ?? *)
111 elim max_S_max f n1.
112 elim H3.
113 absurd m \le S n1.assumption.
114 apply lt_to_not_le.rewrite < H6.assumption.
115 elim H3.
116 apply le_n_Sm_elim m n1 H2.
117 intro.
118 apply H.rewrite < H6.assumption.
119 apply le_S_S_to_le.assumption.
120 intro.rewrite > H7.assumption.
121 qed.
122
123 let rec min_aux off n f \def
124   match f (n-off) with 
125   [ true \Rightarrow (n-off)
126   | false \Rightarrow 
127       match off with
128       [ O \Rightarrow n
129       | (S p) \Rightarrow min_aux p n f]].
130
131 definition min : nat \to (nat \to bool) \to nat \def
132 \lambda n.\lambda f. min_aux n n f.
133
134 theorem min_aux_O_f: \forall f:nat \to bool. \forall i :nat.
135 min_aux O i f = i.
136 intros.simplify.rewrite < minus_n_O.
137 elim f i.
138 simplify.reflexivity.
139 simplify.reflexivity.
140 qed.
141
142 theorem min_O_f : \forall f:nat \to bool.
143 min O f = O.
144 intro.apply min_aux_O_f f O.
145 qed.
146
147 theorem min_aux_S : \forall f: nat \to bool. \forall i,n:nat.
148 (f (n -(S i)) = true \land min_aux (S i) n f = (n - (S i))) \lor 
149 (f (n -(S i)) = false \land min_aux (S i) n f = min_aux i n f).
150 intros.simplify.elim (f (n - (S i))).
151 simplify.left.split.reflexivity.reflexivity.
152 simplify.right.split.reflexivity.reflexivity.
153 qed.
154
155 theorem f_min_aux_true: \forall f:nat \to bool. \forall off,m:nat.
156 ex nat (\lambda i:nat. (le (m-off) i) \land (le i m) \land (f i = true)) \to
157 f (min_aux off m f) = true. 
158 intros 2.
159 elim off.elim H.elim H1.elim H2.
160 cut a = m.
161 rewrite > min_aux_O_f f.rewrite < Hcut.assumption.
162 apply antisym_le a m .assumption.rewrite > minus_n_O m.assumption.
163 simplify.
164 apply bool_ind (\lambda b:bool.
165 (f (m-(S n)) = b) \to (f ([\lambda b:bool.nat] match b in bool with
166 [ true \Rightarrow m-(S n)
167 | false  \Rightarrow (min_aux n m f)])) = true) ? ? ?.
168 reflexivity.
169 simplify.intro.assumption.
170 simplify.intro.apply H.
171 elim H1.elim H3.elim H4.
172 elim (le_to_or_lt_eq (m-(S n)) a H6). 
173 apply ex_intro nat ? a.
174 split.split.
175 apply lt_minus_S_n_to_le_minus_n.assumption.
176 assumption.assumption.
177 absurd f a = false.rewrite < H8.assumption.
178 rewrite > H5.
179 apply not_eq_true_false.
180 qed.
181
182 theorem lt_min_aux_to_false : \forall f:nat \to bool. 
183 \forall n,off,m:nat. (n-off) \leq m \to m < (min_aux off n f) \to f m = false.
184 intros 3.
185 elim off.absurd le n m.rewrite > minus_n_O.assumption.
186 apply lt_to_not_le.rewrite < min_aux_O_f f n.assumption.
187 generalize in match H1.
188 elim min_aux_S f n1 n.
189 elim H3.
190 absurd n - S n1 \le m.assumption.
191 apply lt_to_not_le.rewrite < H6.assumption.
192 elim H3.
193 elim le_to_or_lt_eq (n -(S n1)) m.
194 apply H.apply lt_minus_S_n_to_le_minus_n.assumption.
195 rewrite < H6.assumption.assumption.
196 rewrite < H7.assumption.
197 qed.
198
199 theorem le_min_aux : \forall f:nat \to bool. 
200 \forall n,off:nat. (n-off) \leq (min_aux off n f).
201 intros 3.
202 elim off.rewrite < minus_n_O.
203 rewrite > min_aux_O_f f n.apply le_n.
204 elim min_aux_S f n1 n.
205 elim H1.rewrite > H3.apply le_n.
206 elim H1.rewrite > H3.
207 apply trans_le (n-(S n1)) (n-n1) ?.
208 apply monotonic_le_minus_r.
209 apply le_n_Sn.
210 assumption.
211 qed.
212