1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| A.Asperti, C.Sacerdoti Coen, *)
8 (* ||A|| E.Tassi, S.Zacchiroli *)
10 (* \ / Matita is distributed under the terms of the *)
11 (* v GNU Lesser General Public License Version 2.1 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/nat/minimization".
17 include "nat/minus.ma".
25 | (S j) \Rightarrow max j f ]].
27 theorem max_O_f : \forall f: nat \to bool. max O f = O.
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.
42 theorem le_max_n : \forall f: nat \to bool. \forall n:nat.
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.
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.
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)).
58 apply le_S.apply le_max_n.
59 elim H3.rewrite > H5.apply le_n.
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).
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.
76 definition max_spec \def \lambda f:nat \to bool.\lambda n: nat.
77 \exists i. (le i n) \land (f i = true) \to
78 (f n) = true \land (\forall i. i < n \to (f i = false)).
80 theorem f_max_true : \forall f:nat \to bool. \forall n:nat.
81 (\exists i:nat. le i n \land f i = true) \to f (max n f) = true.
83 elim n.elim H.elim H1.generalize in match H3.
84 apply (le_n_O_elim a H2).intro.simplify.rewrite > H4.
87 apply (bool_ind (\lambda b:bool.
88 (f (S n1) = b) \to (f (match b in bool with
89 [ true \Rightarrow (S n1)
90 | false \Rightarrow (max n1 f)])) = true)).
91 simplify.intro.assumption.
92 simplify.intro.apply H.
93 elim H1.elim H3.generalize in match H5.
94 apply (le_n_Sm_elim a n1 H4).
96 apply (ex_intro nat ? a).
97 split.apply le_S_S_to_le.assumption.assumption.
98 intros.apply False_ind.apply not_eq_true_false.
99 rewrite < H2.rewrite < H7.rewrite > H6. reflexivity.
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.
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 elim (max_S_max f n1).
112 absurd (m \le S n1).assumption.
113 apply lt_to_not_le.rewrite < H6.assumption.
115 apply (le_n_Sm_elim m n1 H2).
117 apply H.rewrite < H6.assumption.
118 apply le_S_S_to_le.assumption.
119 intro.rewrite > H7.assumption.
122 let rec min_aux off n f \def
124 [ true \Rightarrow (n-off)
128 | (S p) \Rightarrow min_aux p n f]].
130 definition min : nat \to (nat \to bool) \to nat \def
131 \lambda n.\lambda f. min_aux n n f.
133 theorem min_aux_O_f: \forall f:nat \to bool. \forall i :nat.
135 intros.simplify.rewrite < minus_n_O.
136 elim (f i).reflexivity.
137 simplify.reflexivity.
140 theorem min_O_f : \forall f:nat \to bool.
142 intro.apply (min_aux_O_f f O).
145 theorem min_aux_S : \forall f: nat \to bool. \forall i,n:nat.
146 (f (n -(S i)) = true \land min_aux (S i) n f = (n - (S i))) \lor
147 (f (n -(S i)) = false \land min_aux (S i) n f = min_aux i n f).
148 intros.simplify.elim (f (n - (S i))).
149 simplify.left.split.reflexivity.reflexivity.
150 simplify.right.split.reflexivity.reflexivity.
153 theorem f_min_aux_true: \forall f:nat \to bool. \forall off,m:nat.
154 (\exists i. le (m-off) i \land le i m \land f i = true) \to
155 f (min_aux off m f) = true.
157 elim off.elim H.elim H1.elim H2.
159 rewrite > (min_aux_O_f f).rewrite < Hcut.assumption.
160 apply (antisym_le a m).assumption.rewrite > (minus_n_O m).assumption.
162 apply (bool_ind (\lambda b:bool.
163 (f (m-(S n)) = b) \to (f (match b in bool with
164 [ true \Rightarrow m-(S n)
165 | false \Rightarrow (min_aux n m f)])) = true)).
166 simplify.intro.assumption.
167 simplify.intro.apply H.
168 elim H1.elim H3.elim H4.
169 elim (le_to_or_lt_eq (m-(S n)) a H6).
170 apply (ex_intro nat ? a).
172 apply lt_minus_S_n_to_le_minus_n.assumption.
173 assumption.assumption.
174 absurd (f a = false).rewrite < H8.assumption.
176 apply not_eq_true_false.
180 theorem lt_min_aux_to_false : \forall f:nat \to bool.
181 \forall n,off,m:nat. (n-off) \leq m \to m < (min_aux off n f) \to f m = false.
183 elim off.absurd (le n m).rewrite > minus_n_O.assumption.
184 apply lt_to_not_le.rewrite < (min_aux_O_f f n).assumption.
185 generalize in match H1.
186 elim (min_aux_S f n1 n).
188 absurd (n - S n1 \le m).assumption.
189 apply lt_to_not_le.rewrite < H6.assumption.
191 elim (le_to_or_lt_eq (n -(S n1)) m).
192 apply H.apply lt_minus_S_n_to_le_minus_n.assumption.
193 rewrite < H6.assumption.
194 rewrite < H7.assumption.
198 theorem le_min_aux : \forall f:nat \to bool.
199 \forall n,off:nat. (n-off) \leq (min_aux off n f).
201 elim off.rewrite < minus_n_O.
202 rewrite > (min_aux_O_f f n).apply le_n.
203 elim (min_aux_S f n1 n).
204 elim H1.rewrite > H3.apply le_n.
205 elim H1.rewrite > H3.
206 apply (trans_le (n-(S n1)) (n-n1)).
207 apply monotonic_le_minus_r.
212 theorem le_min_aux_r : \forall f:nat \to bool.
213 \forall n,off:nat. (min_aux off n f) \le n.
215 elim off.simplify.rewrite < minus_n_O.
216 elim (f n).simplify.apply le_n.
218 simplify.elim (f (n -(S n1))).
219 simplify.apply le_plus_to_minus.
220 rewrite < sym_plus.apply le_plus_n.