]> matita.cs.unibo.it Git - helm.git/blob - matita/library/nat/minimization.ma
8b4b83c2fbc7d910fc9c3d91dbd4dc130e713c30
[helm.git] / 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
19 let rec max i f \def
20   match (f i) with 
21   [ true \Rightarrow i
22   | false \Rightarrow 
23       match i with 
24       [ O \Rightarrow O
25       | (S j) \Rightarrow max j f ]].
26
27 theorem max_O_f : \forall f: nat \to bool. max O f = O.
28 intro. simplify.
29 elim (f O).
30 simplify.reflexivity.
31 simplify.reflexivity.
32 qed. 
33
34 theorem max_S_max : \forall f: nat \to bool. \forall n:nat.
35 (f (S n) = true \land max (S n) f = (S n)) \lor 
36 (f (S n) = false \land max (S n) f = max n f).
37 intros.simplify.elim (f (S n)).
38 simplify.left.split.reflexivity.reflexivity.
39 simplify.right.split.reflexivity.reflexivity.
40 qed.
41
42 theorem le_max_n : \forall f: nat \to bool. \forall n:nat.
43 max n f \le n.
44 intros.elim n.rewrite > max_O_f.apply le_n.
45 simplify.elim (f (S n1)).simplify.apply le_n.
46 simplify.apply le_S.assumption.
47 qed.
48
49 theorem le_to_le_max : \forall f: nat \to bool. \forall n,m:nat.
50 n\le m  \to max n f \le max m f.
51 intros.elim H.
52 apply le_n.
53 apply (trans_le ? (max n1 f)).apply H2.
54 cut ((f (S n1) = true \land max (S n1) f = (S n1)) \lor 
55 (f (S n1) = false \land max (S n1) f = max n1 f)).
56 elim Hcut.elim H3.
57 rewrite > H5.
58 apply le_S.apply le_max_n.
59 elim H3.rewrite > H5.apply le_n.
60 apply max_S_max.
61 qed.
62
63 theorem f_m_to_le_max: \forall f: nat \to bool. \forall n,m:nat.
64 m\le n \to f m = true \to m \le max n f.
65 intros 3.elim n.apply (le_n_O_elim m H).
66 apply le_O_n.
67 apply (le_n_Sm_elim m n1 H1).
68 intro.apply (trans_le ? (max n1 f)).
69 apply H.apply le_S_S_to_le.assumption.assumption.
70 apply le_to_le_max.apply le_n_Sn.
71 intro.simplify.rewrite < H3. 
72 rewrite > H2.simplify.apply le_n.
73 qed.
74
75 theorem f_max_true : \forall f:nat \to bool. \forall n:nat.
76 (\exists i:nat. le i n \land f i = true) \to f (max n f) = true. 
77 intros 2.
78 elim n.elim H.elim H1.generalize in match H3.
79 apply (le_n_O_elim a H2).intro.simplify.rewrite > H4.
80 simplify.assumption.
81 simplify.
82 apply (bool_ind (\lambda b:bool.
83 (f (S n1) = b) \to (f (match b in bool with
84 [ true \Rightarrow (S n1)
85 | false  \Rightarrow (max n1 f)])) = true)).
86 simplify.intro.assumption.
87 simplify.intro.apply H.
88 elim H1.elim H3.generalize in match H5.
89 apply (le_n_Sm_elim a n1 H4).
90 intros.
91 apply (ex_intro nat ? a).
92 split.apply le_S_S_to_le.assumption.assumption.
93 intros.apply False_ind.apply not_eq_true_false.
94 rewrite < H2.rewrite < H7.rewrite > H6. reflexivity.
95 reflexivity.
96 qed.
97
98 theorem lt_max_to_false : \forall f:nat \to bool. 
99 \forall n,m:nat. (max n f) < m \to m \leq n \to f m = false.
100 intros 2.
101 elim n.absurd (le m O).assumption.
102 cut (O < m).apply (lt_O_n_elim m Hcut).exact not_le_Sn_O.
103 rewrite < (max_O_f f).assumption.
104 generalize in match H1.
105 elim (max_S_max f n1).
106 elim H3.
107 absurd (m \le S n1).assumption.
108 apply lt_to_not_le.rewrite < H6.assumption.
109 elim H3.
110 apply (le_n_Sm_elim m n1 H2).
111 intro.
112 apply H.rewrite < H6.assumption.
113 apply le_S_S_to_le.assumption.
114 intro.rewrite > H7.assumption.
115 qed.
116
117 definition max_spec \def \lambda f:nat \to bool.\lambda n,m: nat.
118 m \le n \land (f m)=true \land (\forall i. m < i \to i \le n \to (f i = false)).
119
120 theorem max_spec_to_max: \forall f:nat \to bool.\forall n,m:nat.
121 max_spec f n m \to max n f = m.
122 intros 2.
123 elim n
124   [elim H.elim H1.apply (le_n_O_elim ? H3).
125    apply max_O_f
126   |elim H1.
127    elim (max_S_max f n1)
128     [elim H4.
129      rewrite > H6.
130      apply le_to_le_to_eq
131       [apply not_lt_to_le.
132        unfold Not.intro.
133        apply not_eq_true_false.
134        rewrite < H5.
135        apply H3
136         [assumption|apply le_n]
137       |elim H2.assumption
138       ]
139     |elim H4.
140      rewrite > H6.
141      apply H.
142      elim H2.
143      split
144       [split
145         [elim (le_to_or_lt_eq ? ? H7)
146           [apply le_S_S_to_le.assumption
147           |apply False_ind.
148            apply not_eq_true_false.
149            rewrite < H8.
150            rewrite > H9.
151            assumption
152           ]
153         |assumption
154         ]
155       |intros.
156        apply H3
157         [assumption|apply le_S.assumption]
158       ]
159     ]
160   ]
161 qed.
162  
163 let rec min_aux off n f \def
164   match f n with 
165   [ true \Rightarrow n
166   | false \Rightarrow 
167       match off with
168       [ O \Rightarrow n
169       | (S p) \Rightarrow min_aux p (S n) f]].
170
171 definition min : nat \to (nat \to bool) \to nat \def
172 \lambda n.\lambda f. min_aux n O f.
173
174 theorem min_aux_O_f: \forall f:nat \to bool. \forall i :nat.
175 min_aux O i f = i.
176 intros.simplify.
177 elim (f i).reflexivity.
178 simplify.reflexivity.
179 qed.
180
181 theorem min_O_f : \forall f:nat \to bool.
182 min O f = O.
183 intro.apply (min_aux_O_f f O).
184 qed.
185
186 theorem min_aux_S : \forall f: nat \to bool. \forall i,n:nat.
187 ((f n) = true \land min_aux (S i) n f = n) \lor 
188 ((f n) = false \land min_aux (S i) n f = min_aux i (S n) f).
189 intros.simplify.elim (f n).
190 simplify.left.split.reflexivity.reflexivity.
191 simplify.right.split.reflexivity.reflexivity.
192 qed.
193
194 theorem f_min_aux_true: \forall f:nat \to bool. \forall off,m:nat.
195 (\exists i. le m i \land le i (off + m) \land f i = true) \to
196 f (min_aux off m f) = true. 
197 intros 2.
198 elim off.elim H.elim H1.elim H2.
199 cut (a = m).
200 rewrite > (min_aux_O_f f).rewrite < Hcut.assumption.
201 apply (antisym_le a m).assumption.assumption.
202 simplify.
203 apply (bool_ind (\lambda b:bool.
204 (f m = b) \to (f (match b in bool with
205 [ true \Rightarrow m
206 | false  \Rightarrow (min_aux n (S m) f)])) = true)).
207 intro; assumption.
208 intro. apply H.
209 elim H1.elim H3.elim H4.
210 elim (le_to_or_lt_eq ? a H6). 
211 apply (ex_intro nat ? a).
212 split.split.
213 assumption.
214 rewrite < plus_n_Sm; assumption.
215 assumption.
216 absurd (f a = false).rewrite < H8.assumption.
217 rewrite > H5.
218 apply not_eq_true_false.
219 reflexivity.
220 qed.
221
222 theorem lt_min_aux_to_false : \forall f:nat \to bool. 
223 \forall n,off,m:nat. n \leq m \to m < (min_aux off n f) \to f m = false.
224 intros 3.
225 generalize in match n; clear n.
226 elim off.absurd (le n1 m).assumption.
227 apply lt_to_not_le.rewrite < (min_aux_O_f f n1).assumption.
228 elim (le_to_or_lt_eq ? ? H1);
229  [ unfold lt in H3;
230    apply (H (S n1));
231    [ assumption
232    | elim (min_aux_S f n n1);
233      [ elim H4;
234        elim (not_le_Sn_n n1);
235        rewrite > H6 in H2;
236        apply (lt_to_le (S n1) n1 ?).
237        apply (le_to_lt_to_lt (S n1) m n1 ? ?);
238         [apply (H3).
239         |apply (H2).
240         ]
241      | elim H4;
242        rewrite < H6;
243        assumption
244      ]
245    ]
246  | rewrite < H3 in H2 ⊢ %.
247    elim (min_aux_S f n n1);
248     [ elim H4;
249       rewrite > H6 in H2;
250       unfold lt in H2;
251       elim (not_le_Sn_n ? H2)
252     | elim H4;
253       assumption
254     ]
255  ]
256 qed.
257
258
259 lemma le_min_aux : \forall f:nat \to bool. 
260 \forall n,off:nat. n \leq (min_aux off n f).
261 intros 3.
262 generalize in match n. clear n.
263 elim off.
264 rewrite > (min_aux_O_f f n1).apply le_n.
265 elim (min_aux_S f n n1).
266 elim H1.rewrite > H3.apply le_n.
267 elim H1.rewrite > H3.
268 apply (transitive_le ? (S n1));
269  [ apply le_n_Sn
270  | apply (H (S n1))
271  ]
272 qed.
273
274 theorem le_min_aux_r : \forall f:nat \to bool. 
275 \forall n,off:nat. (min_aux off n f) \le n+off.
276 intros.
277 generalize in match n. clear n.
278 elim off.simplify.
279 elim (f n1).simplify.rewrite < plus_n_O.apply le_n.
280 simplify.rewrite < plus_n_O.apply le_n.
281 simplify.elim (f n1).
282 simplify.
283 apply (le_plus_n_r (S n) n1).
284 simplify.rewrite < plus_n_Sm.
285 apply (H (S n1)).
286 qed.