13958
|
1 |
(* Title: IntFloor.ML
|
|
2 |
Author: Jacques D. Fleuriot
|
|
3 |
Copyright: 2001,2002 University of Edinburgh
|
|
4 |
Description: Floor and ceiling operations over reals
|
|
5 |
*)
|
|
6 |
|
|
7 |
Goal "((number_of n) < real (m::int)) = (number_of n < m)";
|
|
8 |
by Auto_tac;
|
|
9 |
by (rtac (real_of_int_less_iff RS iffD1) 1);
|
|
10 |
by (dtac (real_of_int_less_iff RS iffD2) 2);
|
|
11 |
by Auto_tac;
|
|
12 |
qed "number_of_less_real_of_int_iff";
|
|
13 |
Addsimps [number_of_less_real_of_int_iff];
|
|
14 |
|
|
15 |
Goal "(real (m::int) < (number_of n)) = (m < number_of n)";
|
|
16 |
by Auto_tac;
|
|
17 |
by (rtac (real_of_int_less_iff RS iffD1) 1);
|
|
18 |
by (dtac (real_of_int_less_iff RS iffD2) 2);
|
|
19 |
by Auto_tac;
|
|
20 |
qed "number_of_less_real_of_int_iff2";
|
|
21 |
Addsimps [number_of_less_real_of_int_iff2];
|
|
22 |
|
|
23 |
Goalw [real_le_def,zle_def]
|
|
24 |
"((number_of n) <= real (m::int)) = (number_of n <= m)";
|
|
25 |
by Auto_tac;
|
|
26 |
qed "number_of_le_real_of_int_iff";
|
|
27 |
Addsimps [number_of_le_real_of_int_iff];
|
|
28 |
|
|
29 |
Goalw [real_le_def,zle_def]
|
|
30 |
"(real (m::int) <= (number_of n)) = (m <= number_of n)";
|
|
31 |
by Auto_tac;
|
|
32 |
qed "number_of_le_real_of_int_iff2";
|
|
33 |
Addsimps [number_of_le_real_of_int_iff2];
|
|
34 |
|
|
35 |
Goalw [floor_def] "floor 0 = 0";
|
|
36 |
by (rtac Least_equality 1);
|
|
37 |
by Auto_tac;
|
|
38 |
qed "floor_zero";
|
|
39 |
Addsimps [floor_zero];
|
|
40 |
|
|
41 |
Goal "floor (real (0::nat)) = 0";
|
|
42 |
by Auto_tac;
|
|
43 |
qed "floor_real_of_nat_zero";
|
|
44 |
Addsimps [floor_real_of_nat_zero];
|
|
45 |
|
|
46 |
Goalw [floor_def] "floor (real (n::nat)) = int n";
|
|
47 |
by (rtac Least_equality 1);
|
|
48 |
by (dtac (real_of_int_real_of_nat RS ssubst) 2);
|
|
49 |
by (dtac (real_of_int_less_iff RS iffD1) 2);
|
|
50 |
by (auto_tac (claset(),simpset() addsimps [real_of_int_real_of_nat]));
|
|
51 |
qed "floor_real_of_nat";
|
|
52 |
Addsimps [floor_real_of_nat];
|
|
53 |
|
|
54 |
Goalw [floor_def] "floor (- real (n::nat)) = - int n";
|
|
55 |
by (rtac Least_equality 1);
|
|
56 |
by (dtac (real_of_int_real_of_nat RS ssubst) 2);
|
|
57 |
by (dtac (real_of_int_minus RS subst) 2);
|
|
58 |
by (dtac (real_of_int_less_iff RS iffD1) 2);
|
|
59 |
by (auto_tac (claset(),simpset() addsimps [real_of_int_real_of_nat]));
|
|
60 |
qed "floor_minus_real_of_nat";
|
|
61 |
Addsimps [floor_minus_real_of_nat];
|
|
62 |
|
|
63 |
Goalw [floor_def] "floor (real (n::int)) = n";
|
|
64 |
by (rtac Least_equality 1);
|
|
65 |
by (dtac (real_of_int_real_of_nat RS ssubst) 2);
|
|
66 |
by (dtac (real_of_int_less_iff RS iffD1) 2);
|
|
67 |
by Auto_tac;
|
|
68 |
qed "floor_real_of_int";
|
|
69 |
Addsimps [floor_real_of_int];
|
|
70 |
|
|
71 |
Goalw [floor_def] "floor (- real (n::int)) = - n";
|
|
72 |
by (rtac Least_equality 1);
|
|
73 |
by (dtac (real_of_int_minus RS subst) 2);
|
|
74 |
by (dtac (real_of_int_real_of_nat RS ssubst) 2);
|
|
75 |
by (dtac (real_of_int_less_iff RS iffD1) 2);
|
|
76 |
by Auto_tac;
|
|
77 |
qed "floor_minus_real_of_int";
|
|
78 |
Addsimps [floor_minus_real_of_int];
|
|
79 |
|
|
80 |
Goal "0 <= r ==> EX (n::nat). real (n - 1) <= r & r < real (n)";
|
|
81 |
by (cut_inst_tac [("x","r")] reals_Archimedean2 1);
|
|
82 |
by (Step_tac 1);
|
|
83 |
by (forw_inst_tac [("P","%k. r < real k"),("k","n"),("m","%x. x")]
|
|
84 |
(thm "ex_has_least_nat") 1);
|
|
85 |
by Auto_tac;
|
|
86 |
by (res_inst_tac [("x","x")] exI 1);
|
|
87 |
by (dres_inst_tac [("x","x - 1")] spec 1);
|
|
88 |
by (auto_tac (claset() addDs [ARITH_PROVE "x <= x - Suc 0 ==> x = (0::nat)"],
|
|
89 |
simpset()));
|
|
90 |
qed "reals_Archimedean6";
|
|
91 |
|
|
92 |
Goal "0 <= r ==> EX n. real (n) <= r & r < real (Suc n)";
|
|
93 |
by (dtac reals_Archimedean6 1);
|
|
94 |
by Auto_tac;
|
|
95 |
by (res_inst_tac [("x","n - 1")] exI 1);
|
|
96 |
by (subgoal_tac "0 < n" 1);
|
|
97 |
by (dtac order_le_less_trans 2);
|
|
98 |
by Auto_tac;
|
|
99 |
qed "reals_Archimedean6a";
|
|
100 |
|
|
101 |
Goal "0 <= r ==> EX n. real n <= r & r < real ((n::int) + 1)";
|
|
102 |
by (dtac reals_Archimedean6a 1);
|
|
103 |
by Auto_tac;
|
|
104 |
by (res_inst_tac [("x","int n")] exI 1);
|
|
105 |
by (auto_tac (claset(),simpset() addsimps [real_of_int_real_of_nat,
|
|
106 |
real_of_nat_Suc]));
|
|
107 |
qed "reals_Archimedean_6b_int";
|
|
108 |
|
|
109 |
Goal "r < 0 ==> EX n. real n <= r & r < real ((n::int) + 1)";
|
|
110 |
by (dtac (CLAIM "r < (0::real) ==> 0 <= -r") 1);
|
|
111 |
by (dtac reals_Archimedean_6b_int 1);
|
|
112 |
by Auto_tac;
|
|
113 |
by (dtac real_le_imp_less_or_eq 1 THEN Auto_tac);
|
|
114 |
by (res_inst_tac [("x","- n - 1")] exI 1);
|
|
115 |
by (res_inst_tac [("x","- n")] exI 2);
|
|
116 |
by Auto_tac;
|
|
117 |
qed "reals_Archimedean_6c_int";
|
|
118 |
|
|
119 |
Goal " EX (n::int). real n <= r & r < real ((n::int) + 1)";
|
|
120 |
by (cut_inst_tac [("r","r")] (CLAIM "0 <= r | r < (0::real)") 1);
|
|
121 |
by (blast_tac (claset() addIs [reals_Archimedean_6b_int,
|
|
122 |
reals_Archimedean_6c_int]) 1);
|
|
123 |
qed "real_lb_ub_int";
|
|
124 |
|
|
125 |
Goal "[| real n <= r; r < real y + 1 |] ==> n <= (y::int)";
|
|
126 |
by (dres_inst_tac [("x","real n"),("z","real y + 1")] order_le_less_trans 1);
|
|
127 |
by (rotate_tac 1 2);
|
|
128 |
by (dtac ((CLAIM "real (1::int) = 1") RS ssubst) 2);
|
|
129 |
by (rotate_tac 1 2);
|
|
130 |
by (dres_inst_tac [("x1","y")] (real_of_int_add RS subst) 2);
|
|
131 |
by (dtac (real_of_int_less_iff RS iffD1) 2);
|
|
132 |
by Auto_tac;
|
|
133 |
val lemma_floor = result();
|
|
134 |
|
|
135 |
Goalw [floor_def,Least_def] "real (floor r) <= r";
|
|
136 |
by (cut_inst_tac [("r","r")] real_lb_ub_int 1 THEN Step_tac 1);
|
|
137 |
by (rtac theI2 1);
|
|
138 |
by Auto_tac;
|
|
139 |
by (blast_tac (claset() addIs [lemma_floor]) 1);
|
|
140 |
by (ALLGOALS(dres_inst_tac [("x","n")] spec) THEN Auto_tac);
|
|
141 |
by (force_tac (claset() addDs [lemma_floor],simpset()) 1);
|
|
142 |
by (dtac (real_of_int_le_iff RS iffD2) 1);
|
|
143 |
by (blast_tac (claset() addIs [real_le_trans]) 1);
|
|
144 |
qed "real_of_int_floor_le";
|
|
145 |
Addsimps [real_of_int_floor_le];
|
|
146 |
|
|
147 |
Goalw [floor_def,Least_def]
|
|
148 |
"x < y ==> floor x <= floor y";
|
|
149 |
by (cut_inst_tac [("r","x")] real_lb_ub_int 1 THEN Step_tac 1);
|
|
150 |
by (cut_inst_tac [("r","y")] real_lb_ub_int 1 THEN Step_tac 1);
|
|
151 |
by (rtac theI2 1);
|
|
152 |
by (rtac theI2 3);
|
|
153 |
by Auto_tac;
|
|
154 |
by (auto_tac (claset() addIs [lemma_floor],simpset()));
|
|
155 |
by (ALLGOALS(force_tac (claset() addDs [lemma_floor],simpset())));
|
|
156 |
qed "floor_le";
|
|
157 |
|
|
158 |
Goal "x <= y ==> floor x <= floor y";
|
|
159 |
by (auto_tac (claset() addDs [real_le_imp_less_or_eq],simpset()
|
|
160 |
addsimps [floor_le]));
|
|
161 |
qed "floor_le2";
|
|
162 |
|
|
163 |
Goal "real na < real (x::int) + 1 ==> na <= x";
|
|
164 |
by (auto_tac (claset() addIs [lemma_floor],simpset()));
|
|
165 |
val lemma_floor2 = result();
|
|
166 |
|
|
167 |
Goalw [floor_def,Least_def]
|
|
168 |
"(real (floor x) = x) = (EX (n::int). x = real n)";
|
|
169 |
by (cut_inst_tac [("r","x")] real_lb_ub_int 1 THEN etac exE 1);
|
|
170 |
by (rtac theI2 1);
|
|
171 |
by (auto_tac (claset() addIs [lemma_floor],simpset()));
|
|
172 |
by (force_tac (claset() addDs [lemma_floor],simpset()) 1);
|
|
173 |
by (force_tac (claset() addDs [lemma_floor2],simpset()) 1);
|
|
174 |
qed "real_of_int_floor_cancel";
|
|
175 |
Addsimps [real_of_int_floor_cancel];
|
|
176 |
|
|
177 |
Goalw [floor_def]
|
|
178 |
"[| real n < x; x < real n + 1 |] ==> floor x = n";
|
|
179 |
by (rtac Least_equality 1);
|
|
180 |
by (auto_tac (claset() addIs [lemma_floor],simpset()));
|
|
181 |
qed "floor_eq";
|
|
182 |
|
|
183 |
Goalw [floor_def]
|
|
184 |
"[| real n <= x; x < real n + 1 |] ==> floor x = n";
|
|
185 |
by (rtac Least_equality 1);
|
|
186 |
by (auto_tac (claset() addIs [lemma_floor],simpset()));
|
|
187 |
qed "floor_eq2";
|
|
188 |
|
|
189 |
Goal "[| real n < x; x < real (Suc n) |] ==> nat(floor x) = n";
|
|
190 |
by (rtac (inj_int RS injD) 1);
|
|
191 |
by (rtac (CLAIM "0 <= x ==> int (nat x) = x" RS ssubst) 1);
|
|
192 |
by (rtac floor_eq 2);
|
|
193 |
by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc,
|
|
194 |
real_of_int_real_of_nat]));
|
|
195 |
by (rtac (floor_le RSN (2,zle_trans)) 1 THEN Auto_tac);
|
|
196 |
qed "floor_eq3";
|
|
197 |
|
|
198 |
Goal "[| real n <= x; x < real (Suc n) |] ==> nat(floor x) = n";
|
|
199 |
by (dtac order_le_imp_less_or_eq 1);
|
|
200 |
by (auto_tac (claset() addIs [floor_eq3],simpset()));
|
|
201 |
qed "floor_eq4";
|
|
202 |
|
|
203 |
Goalw [real_number_of_def]
|
|
204 |
"floor(number_of n :: real) = (number_of n :: int)";
|
|
205 |
by (rtac floor_eq2 1);
|
|
206 |
by (rtac (CLAIM "x < x + (1::real)") 2);
|
|
207 |
by (rtac real_le_refl 1);
|
|
208 |
qed "floor_number_of_eq";
|
|
209 |
Addsimps [floor_number_of_eq];
|
|
210 |
|
|
211 |
Goalw [floor_def,Least_def] "r - 1 <= real(floor r)";
|
|
212 |
by (cut_inst_tac [("r","r")] real_lb_ub_int 1 THEN Step_tac 1);
|
|
213 |
by (rtac theI2 1);
|
|
214 |
by (auto_tac (claset() addIs [lemma_floor],simpset()));
|
|
215 |
by (force_tac (claset() addDs [lemma_floor],simpset()) 1);
|
|
216 |
qed "real_of_int_floor_ge_diff_one";
|
|
217 |
Addsimps [real_of_int_floor_ge_diff_one];
|
|
218 |
|
|
219 |
Goal "r <= real(floor r) + 1";
|
|
220 |
by (cut_inst_tac [("r","r")] real_of_int_floor_ge_diff_one 1);
|
|
221 |
by (auto_tac (claset(),simpset() delsimps [real_of_int_floor_ge_diff_one]));
|
|
222 |
qed "real_of_int_floor_add_one_ge";
|
|
223 |
Addsimps [real_of_int_floor_add_one_ge];
|
|
224 |
|
|
225 |
|
|
226 |
(*--------------------------------------------------------------------------------*)
|
|
227 |
(* ceiling function for positive reals *)
|
|
228 |
(*--------------------------------------------------------------------------------*)
|
|
229 |
|
|
230 |
Goalw [ceiling_def] "ceiling 0 = 0";
|
|
231 |
by Auto_tac;
|
|
232 |
qed "ceiling_zero";
|
|
233 |
Addsimps [ceiling_zero];
|
|
234 |
|
|
235 |
Goalw [ceiling_def] "ceiling (real (n::nat)) = int n";
|
|
236 |
by Auto_tac;
|
|
237 |
qed "ceiling_real_of_nat";
|
|
238 |
Addsimps [ceiling_real_of_nat];
|
|
239 |
|
|
240 |
Goal "ceiling (real (0::nat)) = 0";
|
|
241 |
by Auto_tac;
|
|
242 |
qed "ceiling_real_of_nat_zero";
|
|
243 |
Addsimps [ceiling_real_of_nat_zero];
|
|
244 |
|
|
245 |
Goalw [ceiling_def] "ceiling (real (floor r)) = floor r";
|
|
246 |
by Auto_tac;
|
|
247 |
qed "ceiling_floor";
|
|
248 |
Addsimps [ceiling_floor];
|
|
249 |
|
|
250 |
Goalw [ceiling_def] "floor (real (ceiling r)) = ceiling r";
|
|
251 |
by Auto_tac;
|
|
252 |
qed "floor_ceiling";
|
|
253 |
Addsimps [floor_ceiling];
|
|
254 |
|
|
255 |
Goalw [ceiling_def] "r <= real (ceiling r)";
|
|
256 |
by Auto_tac;
|
|
257 |
by (rtac (CLAIM "x <= -y ==> (y::real) <= - x") 1);
|
|
258 |
by Auto_tac;
|
|
259 |
qed "real_of_int_ceiling_ge";
|
|
260 |
Addsimps [real_of_int_ceiling_ge];
|
|
261 |
|
|
262 |
Goalw [ceiling_def] "x < y ==> ceiling x <= ceiling y";
|
|
263 |
by (auto_tac (claset() addIs [floor_le],simpset()));
|
|
264 |
qed "ceiling_le";
|
|
265 |
|
|
266 |
Goalw [ceiling_def] "x <= y ==> ceiling x <= ceiling y";
|
|
267 |
by (auto_tac (claset() addIs [floor_le2],simpset()));
|
|
268 |
qed "ceiling_le2";
|
|
269 |
|
|
270 |
Goalw [ceiling_def] "(real (ceiling x) = x) = (EX (n::int). x = real n)";
|
|
271 |
by Auto_tac;
|
|
272 |
by (dtac (CLAIM "- x = y ==> (x::real) = -y") 1);
|
|
273 |
by Auto_tac;
|
|
274 |
by (res_inst_tac [("x","-n")] exI 1);
|
|
275 |
by Auto_tac;
|
|
276 |
qed "real_of_int_ceiling_cancel";
|
|
277 |
Addsimps [real_of_int_ceiling_cancel];
|
|
278 |
|
|
279 |
Goalw [ceiling_def]
|
|
280 |
"[| real n < x; x < real n + 1 |] ==> ceiling x = n + 1";
|
|
281 |
by (rtac (CLAIM "x = - y ==> - (x::int) = y") 1);
|
|
282 |
by (auto_tac (claset() addIs [floor_eq],simpset()));
|
|
283 |
qed "ceiling_eq";
|
|
284 |
|
|
285 |
Goalw [ceiling_def]
|
|
286 |
"[| real n < x; x <= real n + 1 |] ==> ceiling x = n + 1";
|
|
287 |
by (rtac (CLAIM "x = - y ==> - (x::int) = y") 1);
|
|
288 |
by (auto_tac (claset() addIs [floor_eq2],simpset()));
|
|
289 |
qed "ceiling_eq2";
|
|
290 |
|
|
291 |
Goalw [ceiling_def] "[| real n - 1 < x; x <= real n |] ==> ceiling x = n";
|
|
292 |
by (rtac (CLAIM "x = -(y::int) ==> - x = y") 1);
|
|
293 |
by (auto_tac (claset() addIs [floor_eq2],simpset()));
|
|
294 |
qed "ceiling_eq3";
|
|
295 |
|
|
296 |
Goalw [ceiling_def,real_number_of_def]
|
|
297 |
"ceiling (number_of n :: real) = (number_of n :: int)";
|
|
298 |
by (rtac (CLAIM "x = - y ==> - (x::int) = y") 1);
|
|
299 |
by (rtac (floor_minus_real_of_int RS ssubst) 1);
|
|
300 |
by Auto_tac;
|
|
301 |
qed "ceiling_number_of_eq";
|
|
302 |
Addsimps [ceiling_number_of_eq];
|
|
303 |
|
|
304 |
Goalw [ceiling_def] "real (ceiling r) - 1 <= r";
|
|
305 |
by (rtac (CLAIM "-x <= -y ==> (y::real) <= x") 1);
|
|
306 |
by (auto_tac (claset(),simpset() addsimps [real_diff_def]));
|
|
307 |
qed "real_of_int_ceiling_diff_one_le";
|
|
308 |
Addsimps [real_of_int_ceiling_diff_one_le];
|
|
309 |
|
|
310 |
Goal "real (ceiling r) <= r + 1";
|
|
311 |
by (cut_inst_tac [("r","r")] real_of_int_ceiling_diff_one_le 1);
|
|
312 |
by (auto_tac (claset(),simpset() delsimps [real_of_int_ceiling_diff_one_le]));
|
|
313 |
qed "real_of_int_ceiling_le_add_one";
|
|
314 |
Addsimps [real_of_int_ceiling_le_add_one];
|
|
315 |
|