|
1 (* Title: HOL/Integ/Int_lemmas.ML |
|
2 ID: $Id$ |
|
3 Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 Copyright 1998 University of Cambridge |
|
5 |
|
6 Type "int" is a linear order |
|
7 |
|
8 And many further lemmas |
|
9 *) |
|
10 |
|
11 (* legacy ML bindings *) |
|
12 |
|
13 structure Int = |
|
14 struct |
|
15 val thy = the_context (); |
|
16 val zabs_def = thm "zabs_def"; |
|
17 val nat_def = thm "nat_def"; |
|
18 end; |
|
19 |
|
20 open Int; |
|
21 |
|
22 Goal "int 0 = (0::int)"; |
|
23 by (simp_tac (simpset() addsimps [Zero_int_def]) 1); |
|
24 qed "int_0"; |
|
25 |
|
26 Goal "int 1 = 1"; |
|
27 by (simp_tac (simpset() addsimps [One_int_def]) 1); |
|
28 qed "int_1"; |
|
29 |
|
30 Goal "int (Suc 0) = 1"; |
|
31 by (simp_tac (simpset() addsimps [One_int_def, One_nat_def]) 1); |
|
32 qed "int_Suc0_eq_1"; |
|
33 |
|
34 Goalw [zdiff_def,zless_def] "neg x = (x < 0)"; |
|
35 by Auto_tac; |
|
36 qed "neg_eq_less_0"; |
|
37 |
|
38 Goalw [zle_def] "(~neg x) = (0 <= x)"; |
|
39 by (simp_tac (simpset() addsimps [neg_eq_less_0]) 1); |
|
40 qed "not_neg_eq_ge_0"; |
|
41 |
|
42 (** Needed to simplify inequalities when Numeral1 can get simplified to 1 **) |
|
43 |
|
44 Goal "~ neg 0"; |
|
45 by (simp_tac (simpset() addsimps [One_int_def, neg_eq_less_0]) 1); |
|
46 qed "not_neg_0"; |
|
47 |
|
48 Goal "~ neg 1"; |
|
49 by (simp_tac (simpset() addsimps [One_int_def, neg_eq_less_0]) 1); |
|
50 qed "not_neg_1"; |
|
51 |
|
52 Goal "iszero 0"; |
|
53 by (simp_tac (simpset() addsimps [iszero_def]) 1); |
|
54 qed "iszero_0"; |
|
55 |
|
56 Goal "~ iszero 1"; |
|
57 by (simp_tac (simpset() addsimps [Zero_int_def, One_int_def, One_nat_def, |
|
58 iszero_def]) 1); |
|
59 qed "not_iszero_1"; |
|
60 |
|
61 Goal "0 < (1::int)"; |
|
62 by (simp_tac (simpset() addsimps [Zero_int_def, One_int_def, One_nat_def]) 1); |
|
63 qed "int_0_less_1"; |
|
64 |
|
65 Goal "0 \\<noteq> (1::int)"; |
|
66 by (simp_tac (simpset() addsimps [Zero_int_def, One_int_def, One_nat_def]) 1); |
|
67 qed "int_0_neq_1"; |
|
68 |
|
69 Addsimps [int_0, int_1, int_0_neq_1]; |
|
70 |
|
71 |
|
72 (*** Abel_Cancel simproc on the integers ***) |
|
73 |
|
74 (* Lemmas needed for the simprocs *) |
|
75 |
|
76 (*Deletion of other terms in the formula, seeking the -x at the front of z*) |
|
77 Goal "((x::int) + (y + z) = y + u) = ((x + z) = u)"; |
|
78 by (stac zadd_left_commute 1); |
|
79 by (rtac zadd_left_cancel 1); |
|
80 qed "zadd_cancel_21"; |
|
81 |
|
82 (*A further rule to deal with the case that |
|
83 everything gets cancelled on the right.*) |
|
84 Goal "((x::int) + (y + z) = y) = (x = -z)"; |
|
85 by (stac zadd_left_commute 1); |
|
86 by (res_inst_tac [("t", "y")] (zadd_0_right RS subst) 1 |
|
87 THEN stac zadd_left_cancel 1); |
|
88 by (simp_tac (simpset() addsimps [eq_zdiff_eq RS sym]) 1); |
|
89 qed "zadd_cancel_end"; |
|
90 |
|
91 |
|
92 structure Int_Cancel_Data = |
|
93 struct |
|
94 val ss = HOL_ss |
|
95 val eq_reflection = eq_reflection |
|
96 |
|
97 val sg_ref = Sign.self_ref (Theory.sign_of (the_context ())) |
|
98 val T = HOLogic.intT |
|
99 val zero = Const ("0", HOLogic.intT) |
|
100 val restrict_to_left = restrict_to_left |
|
101 val add_cancel_21 = zadd_cancel_21 |
|
102 val add_cancel_end = zadd_cancel_end |
|
103 val add_left_cancel = zadd_left_cancel |
|
104 val add_assoc = zadd_assoc |
|
105 val add_commute = zadd_commute |
|
106 val add_left_commute = zadd_left_commute |
|
107 val add_0 = zadd_0 |
|
108 val add_0_right = zadd_0_right |
|
109 |
|
110 val eq_diff_eq = eq_zdiff_eq |
|
111 val eqI_rules = [zless_eqI, zeq_eqI, zle_eqI] |
|
112 fun dest_eqI th = |
|
113 #1 (HOLogic.dest_bin "op =" HOLogic.boolT |
|
114 (HOLogic.dest_Trueprop (concl_of th))) |
|
115 |
|
116 val diff_def = zdiff_def |
|
117 val minus_add_distrib = zminus_zadd_distrib |
|
118 val minus_minus = zminus_zminus |
|
119 val minus_0 = zminus_0 |
|
120 val add_inverses = [zadd_zminus_inverse, zadd_zminus_inverse2] |
|
121 val cancel_simps = [zadd_zminus_cancel, zminus_zadd_cancel] |
|
122 end; |
|
123 |
|
124 structure Int_Cancel = Abel_Cancel (Int_Cancel_Data); |
|
125 |
|
126 Addsimprocs [Int_Cancel.sum_conv, Int_Cancel.rel_conv]; |
|
127 |
|
128 |
|
129 |
|
130 (*** misc ***) |
|
131 |
|
132 Goal "- (z - y) = y - (z::int)"; |
|
133 by (Simp_tac 1); |
|
134 qed "zminus_zdiff_eq"; |
|
135 Addsimps [zminus_zdiff_eq]; |
|
136 |
|
137 Goal "(w<z) = neg(w-z)"; |
|
138 by (simp_tac (simpset() addsimps [zless_def]) 1); |
|
139 qed "zless_eq_neg"; |
|
140 |
|
141 Goal "(w=z) = iszero(w-z)"; |
|
142 by (simp_tac (simpset() addsimps [iszero_def, zdiff_eq_eq]) 1); |
|
143 qed "eq_eq_iszero"; |
|
144 |
|
145 Goal "(w<=z) = (~ neg(z-w))"; |
|
146 by (simp_tac (simpset() addsimps [zle_def, zless_def]) 1); |
|
147 qed "zle_eq_not_neg"; |
|
148 |
|
149 (** Inequality reasoning **) |
|
150 |
|
151 Goal "(w < z + (1::int)) = (w<z | w=z)"; |
|
152 by (auto_tac (claset(), |
|
153 simpset() addsimps [zless_iff_Suc_zadd, int_Suc, |
|
154 gr0_conv_Suc, zero_reorient])); |
|
155 by (res_inst_tac [("x","Suc n")] exI 1); |
|
156 by (simp_tac (simpset() addsimps [int_Suc]) 1); |
|
157 qed "zless_add1_eq"; |
|
158 |
|
159 Goal "(w + (1::int) <= z) = (w<z)"; |
|
160 by (asm_full_simp_tac (simpset() addsimps [zle_def, zless_add1_eq]) 1); |
|
161 by (auto_tac (claset() addIs [zle_anti_sym], |
|
162 simpset() addsimps [order_less_imp_le, symmetric zle_def])); |
|
163 qed "add1_zle_eq"; |
|
164 |
|
165 Goal "((1::int) + w <= z) = (w<z)"; |
|
166 by (stac zadd_commute 1); |
|
167 by (rtac add1_zle_eq 1); |
|
168 qed "add1_left_zle_eq"; |
|
169 |
|
170 |
|
171 (*** Monotonicity results ***) |
|
172 |
|
173 Goal "(v+z < w+z) = (v < (w::int))"; |
|
174 by (Simp_tac 1); |
|
175 qed "zadd_right_cancel_zless"; |
|
176 |
|
177 Goal "(z+v < z+w) = (v < (w::int))"; |
|
178 by (Simp_tac 1); |
|
179 qed "zadd_left_cancel_zless"; |
|
180 |
|
181 Addsimps [zadd_right_cancel_zless, zadd_left_cancel_zless]; |
|
182 |
|
183 Goal "(v+z <= w+z) = (v <= (w::int))"; |
|
184 by (Simp_tac 1); |
|
185 qed "zadd_right_cancel_zle"; |
|
186 |
|
187 Goal "(z+v <= z+w) = (v <= (w::int))"; |
|
188 by (Simp_tac 1); |
|
189 qed "zadd_left_cancel_zle"; |
|
190 |
|
191 Addsimps [zadd_right_cancel_zle, zadd_left_cancel_zle]; |
|
192 |
|
193 (*"v<=w ==> v+z <= w+z"*) |
|
194 bind_thm ("zadd_zless_mono1", zadd_right_cancel_zless RS iffD2); |
|
195 |
|
196 (*"v<=w ==> z+v <= z+w"*) |
|
197 bind_thm ("zadd_zless_mono2", zadd_left_cancel_zless RS iffD2); |
|
198 |
|
199 (*"v<=w ==> v+z <= w+z"*) |
|
200 bind_thm ("zadd_zle_mono1", zadd_right_cancel_zle RS iffD2); |
|
201 |
|
202 (*"v<=w ==> z+v <= z+w"*) |
|
203 bind_thm ("zadd_zle_mono2", zadd_left_cancel_zle RS iffD2); |
|
204 |
|
205 Goal "[| w'<=w; z'<=z |] ==> w' + z' <= w + (z::int)"; |
|
206 by (etac (zadd_zle_mono1 RS zle_trans) 1); |
|
207 by (Simp_tac 1); |
|
208 qed "zadd_zle_mono"; |
|
209 |
|
210 Goal "[| w'<w; z'<=z |] ==> w' + z' < w + (z::int)"; |
|
211 by (etac (zadd_zless_mono1 RS order_less_le_trans) 1); |
|
212 by (Simp_tac 1); |
|
213 qed "zadd_zless_mono"; |
|
214 |
|
215 |
|
216 (*** Comparison laws ***) |
|
217 |
|
218 Goal "(- x < - y) = (y < (x::int))"; |
|
219 by (simp_tac (simpset() addsimps [zless_def, zdiff_def] @ zadd_ac) 1); |
|
220 qed "zminus_zless_zminus"; |
|
221 Addsimps [zminus_zless_zminus]; |
|
222 |
|
223 Goal "(- x <= - y) = (y <= (x::int))"; |
|
224 by (simp_tac (simpset() addsimps [zle_def]) 1); |
|
225 qed "zminus_zle_zminus"; |
|
226 Addsimps [zminus_zle_zminus]; |
|
227 |
|
228 (** The next several equations can make the simplifier loop! **) |
|
229 |
|
230 Goal "(x < - y) = (y < - (x::int))"; |
|
231 by (simp_tac (simpset() addsimps [zless_def, zdiff_def] @ zadd_ac) 1); |
|
232 qed "zless_zminus"; |
|
233 |
|
234 Goal "(- x < y) = (- y < (x::int))"; |
|
235 by (simp_tac (simpset() addsimps [zless_def, zdiff_def] @ zadd_ac) 1); |
|
236 qed "zminus_zless"; |
|
237 |
|
238 Goal "(x <= - y) = (y <= - (x::int))"; |
|
239 by (simp_tac (simpset() addsimps [zle_def, zminus_zless]) 1); |
|
240 qed "zle_zminus"; |
|
241 |
|
242 Goal "(- x <= y) = (- y <= (x::int))"; |
|
243 by (simp_tac (simpset() addsimps [zle_def, zless_zminus]) 1); |
|
244 qed "zminus_zle"; |
|
245 |
|
246 Goal "(x = - y) = (y = - (x::int))"; |
|
247 by Auto_tac; |
|
248 qed "equation_zminus"; |
|
249 |
|
250 Goal "(- x = y) = (- (y::int) = x)"; |
|
251 by Auto_tac; |
|
252 qed "zminus_equation"; |
|
253 |
|
254 |
|
255 (** Instances of the equations above, for zero **) |
|
256 |
|
257 (*instantiate a variable to zero and simplify*) |
|
258 fun zero_instance v th = simplify (simpset()) (inst v "0" th); |
|
259 |
|
260 Addsimps [zero_instance "x" zless_zminus, |
|
261 zero_instance "y" zminus_zless, |
|
262 zero_instance "x" zle_zminus, |
|
263 zero_instance "y" zminus_zle, |
|
264 zero_instance "x" equation_zminus, |
|
265 zero_instance "y" zminus_equation]; |
|
266 |
|
267 |
|
268 Goal "- (int (Suc n)) < 0"; |
|
269 by (simp_tac (simpset() addsimps [zless_def]) 1); |
|
270 qed "negative_zless_0"; |
|
271 |
|
272 Goal "- (int (Suc n)) < int m"; |
|
273 by (rtac (negative_zless_0 RS order_less_le_trans) 1); |
|
274 by (Simp_tac 1); |
|
275 qed "negative_zless"; |
|
276 AddIffs [negative_zless]; |
|
277 |
|
278 Goal "- int n <= 0"; |
|
279 by (simp_tac (simpset() addsimps [zminus_zle]) 1); |
|
280 qed "negative_zle_0"; |
|
281 |
|
282 Goal "- int n <= int m"; |
|
283 by (simp_tac (simpset() addsimps [zless_def, zle_def, zdiff_def, zadd_int]) 1); |
|
284 qed "negative_zle"; |
|
285 AddIffs [negative_zle]; |
|
286 |
|
287 Goal "~(0 <= - (int (Suc n)))"; |
|
288 by (stac zle_zminus 1); |
|
289 by (Simp_tac 1); |
|
290 qed "not_zle_0_negative"; |
|
291 Addsimps [not_zle_0_negative]; |
|
292 |
|
293 Goal "(int n <= - int m) = (n = 0 & m = 0)"; |
|
294 by Safe_tac; |
|
295 by (Simp_tac 3); |
|
296 by (dtac (zle_zminus RS iffD1) 2); |
|
297 by (ALLGOALS (dtac (negative_zle_0 RSN(2,zle_trans)))); |
|
298 by (ALLGOALS Asm_full_simp_tac); |
|
299 qed "int_zle_neg"; |
|
300 |
|
301 Goal "~(int n < - int m)"; |
|
302 by (simp_tac (simpset() addsimps [symmetric zle_def]) 1); |
|
303 qed "not_int_zless_negative"; |
|
304 |
|
305 Goal "(- int n = int m) = (n = 0 & m = 0)"; |
|
306 by (rtac iffI 1); |
|
307 by (rtac (int_zle_neg RS iffD1) 1); |
|
308 by (dtac sym 1); |
|
309 by (ALLGOALS Asm_simp_tac); |
|
310 qed "negative_eq_positive"; |
|
311 |
|
312 Addsimps [negative_eq_positive, not_int_zless_negative]; |
|
313 |
|
314 |
|
315 Goal "(w <= z) = (EX n. z = w + int n)"; |
|
316 by (auto_tac (claset() addIs [inst "x" "0::nat" exI] |
|
317 addSIs [not_sym RS not0_implies_Suc], |
|
318 simpset() addsimps [zless_iff_Suc_zadd, int_le_less])); |
|
319 qed "zle_iff_zadd"; |
|
320 |
|
321 Goal "abs (int m) = int m"; |
|
322 by (simp_tac (simpset() addsimps [zabs_def]) 1); |
|
323 qed "abs_int_eq"; |
|
324 Addsimps [abs_int_eq]; |
|
325 |
|
326 |
|
327 (**** nat: magnitide of an integer, as a natural number ****) |
|
328 |
|
329 Goalw [nat_def] "nat(int n) = n"; |
|
330 by Auto_tac; |
|
331 qed "nat_int"; |
|
332 Addsimps [nat_int]; |
|
333 |
|
334 Goalw [nat_def] "nat(- (int n)) = 0"; |
|
335 by (auto_tac (claset(), |
|
336 simpset() addsimps [neg_eq_less_0, zero_reorient, zminus_zless])); |
|
337 qed "nat_zminus_int"; |
|
338 Addsimps [nat_zminus_int]; |
|
339 |
|
340 Goalw [Zero_int_def] "nat 0 = 0"; |
|
341 by (rtac nat_int 1); |
|
342 qed "nat_zero"; |
|
343 Addsimps [nat_zero]; |
|
344 |
|
345 Goal "~ neg z ==> int (nat z) = z"; |
|
346 by (dtac (not_neg_eq_ge_0 RS iffD1) 1); |
|
347 by (dtac zle_imp_zless_or_eq 1); |
|
348 by (auto_tac (claset(), simpset() addsimps [zless_iff_Suc_zadd])); |
|
349 qed "not_neg_nat"; |
|
350 |
|
351 Goal "neg x ==> EX n. x = - (int (Suc n))"; |
|
352 by (auto_tac (claset(), |
|
353 simpset() addsimps [neg_eq_less_0, zless_iff_Suc_zadd, |
|
354 zdiff_eq_eq RS sym, zdiff_def])); |
|
355 qed "negD"; |
|
356 |
|
357 Goalw [nat_def] "neg z ==> nat z = 0"; |
|
358 by Auto_tac; |
|
359 qed "neg_nat"; |
|
360 |
|
361 Goal "(m < nat z) = (int m < z)"; |
|
362 by (case_tac "neg z" 1); |
|
363 by (etac (not_neg_nat RS subst) 2); |
|
364 by (auto_tac (claset(), simpset() addsimps [neg_nat])); |
|
365 by (auto_tac (claset() addDs [order_less_trans], |
|
366 simpset() addsimps [neg_eq_less_0])); |
|
367 qed "zless_nat_eq_int_zless"; |
|
368 |
|
369 Goal "0 <= z ==> int (nat z) = z"; |
|
370 by (asm_full_simp_tac |
|
371 (simpset() addsimps [neg_eq_less_0, zle_def, not_neg_nat]) 1); |
|
372 qed "nat_0_le"; |
|
373 |
|
374 Goal "z <= 0 ==> nat z = 0"; |
|
375 by (auto_tac (claset(), |
|
376 simpset() addsimps [order_le_less, neg_eq_less_0, |
|
377 zle_def, neg_nat])); |
|
378 qed "nat_le_0"; |
|
379 Addsimps [nat_0_le, nat_le_0]; |
|
380 |
|
381 (*An alternative condition is 0 <= w *) |
|
382 Goal "0 < z ==> (nat w < nat z) = (w < z)"; |
|
383 by (stac (zless_int RS sym) 1); |
|
384 by (asm_simp_tac (simpset() addsimps [not_neg_nat, not_neg_eq_ge_0, |
|
385 order_le_less]) 1); |
|
386 by (case_tac "neg w" 1); |
|
387 by (asm_simp_tac (simpset() addsimps [not_neg_nat]) 2); |
|
388 by (asm_full_simp_tac (simpset() addsimps [neg_eq_less_0, neg_nat]) 1); |
|
389 by (blast_tac (claset() addIs [order_less_trans]) 1); |
|
390 val lemma = result(); |
|
391 |
|
392 Goal "(nat w < nat z) = (0 < z & w < z)"; |
|
393 by (case_tac "0 < z" 1); |
|
394 by (auto_tac (claset(), simpset() addsimps [lemma, linorder_not_less])); |
|
395 qed "zless_nat_conj"; |
|
396 |
|
397 |
|
398 (* a case theorem distinguishing non-negative and negative int *) |
|
399 |
|
400 val prems = Goal |
|
401 "[|!! n. z = int n ==> P; !! n. z = - (int (Suc n)) ==> P |] ==> P"; |
|
402 by (case_tac "neg z" 1); |
|
403 by (fast_tac (claset() addSDs [negD] addSEs prems) 1); |
|
404 by (dtac (not_neg_nat RS sym) 1); |
|
405 by (eresolve_tac prems 1); |
|
406 qed "int_cases"; |
|
407 |
|
408 fun int_case_tac x = res_inst_tac [("z",x)] int_cases; |
|
409 |
|
410 |
|
411 (*** Monotonicity of Multiplication ***) |
|
412 |
|
413 Goal "i <= (j::int) ==> i * int k <= j * int k"; |
|
414 by (induct_tac "k" 1); |
|
415 by (stac int_Suc 2); |
|
416 by (ALLGOALS |
|
417 (asm_simp_tac (simpset() addsimps [zadd_zmult_distrib2, zadd_zle_mono, |
|
418 int_Suc0_eq_1]))); |
|
419 val lemma = result(); |
|
420 |
|
421 Goal "[| i <= j; (0::int) <= k |] ==> i*k <= j*k"; |
|
422 by (res_inst_tac [("t", "k")] (not_neg_nat RS subst) 1); |
|
423 by (etac lemma 2); |
|
424 by (full_simp_tac (simpset() addsimps [not_neg_eq_ge_0]) 1); |
|
425 qed "zmult_zle_mono1"; |
|
426 |
|
427 Goal "[| i <= j; k <= (0::int) |] ==> j*k <= i*k"; |
|
428 by (rtac (zminus_zle_zminus RS iffD1) 1); |
|
429 by (asm_simp_tac (simpset() addsimps [zmult_zminus_right RS sym, |
|
430 zmult_zle_mono1, zle_zminus]) 1); |
|
431 qed "zmult_zle_mono1_neg"; |
|
432 |
|
433 Goal "[| i <= j; (0::int) <= k |] ==> k*i <= k*j"; |
|
434 by (dtac zmult_zle_mono1 1); |
|
435 by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); |
|
436 qed "zmult_zle_mono2"; |
|
437 |
|
438 Goal "[| i <= j; k <= (0::int) |] ==> k*j <= k*i"; |
|
439 by (dtac zmult_zle_mono1_neg 1); |
|
440 by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); |
|
441 qed "zmult_zle_mono2_neg"; |
|
442 |
|
443 (* <= monotonicity, BOTH arguments*) |
|
444 Goal "[| i <= j; k <= l; (0::int) <= j; (0::int) <= k |] ==> i*k <= j*l"; |
|
445 by (etac (zmult_zle_mono1 RS order_trans) 1); |
|
446 by (assume_tac 1); |
|
447 by (etac zmult_zle_mono2 1); |
|
448 by (assume_tac 1); |
|
449 qed "zmult_zle_mono"; |
|
450 |
|
451 |
|
452 (** strict, in 1st argument; proof is by induction on k>0 **) |
|
453 |
|
454 Goal "i<j ==> 0<k --> int k * i < int k * j"; |
|
455 by (induct_tac "k" 1); |
|
456 by (stac int_Suc 2); |
|
457 by (case_tac "n=0" 2); |
|
458 by (ALLGOALS (asm_full_simp_tac |
|
459 (simpset() addsimps [zadd_zmult_distrib, zadd_zless_mono, |
|
460 int_Suc0_eq_1, order_le_less]))); |
|
461 val lemma = result(); |
|
462 |
|
463 Goal "[| i<j; (0::int) < k |] ==> k*i < k*j"; |
|
464 by (res_inst_tac [("t", "k")] (not_neg_nat RS subst) 1); |
|
465 by (etac (lemma RS mp) 2); |
|
466 by (asm_simp_tac (simpset() addsimps [not_neg_eq_ge_0, |
|
467 order_le_less]) 1); |
|
468 by (forward_tac [conjI RS (zless_nat_conj RS iffD2)] 1); |
|
469 by Auto_tac; |
|
470 qed "zmult_zless_mono2"; |
|
471 |
|
472 Goal "[| i<j; (0::int) < k |] ==> i*k < j*k"; |
|
473 by (dtac zmult_zless_mono2 1); |
|
474 by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [zmult_commute]))); |
|
475 qed "zmult_zless_mono1"; |
|
476 |
|
477 (* < monotonicity, BOTH arguments*) |
|
478 Goal "[| i < j; k < l; (0::int) < j; (0::int) < k |] ==> i*k < j*l"; |
|
479 by (etac (zmult_zless_mono1 RS order_less_trans) 1); |
|
480 by (assume_tac 1); |
|
481 by (etac zmult_zless_mono2 1); |
|
482 by (assume_tac 1); |
|
483 qed "zmult_zless_mono"; |
|
484 |
|
485 Goal "[| i<j; k < (0::int) |] ==> j*k < i*k"; |
|
486 by (rtac (zminus_zless_zminus RS iffD1) 1); |
|
487 by (asm_simp_tac (simpset() addsimps [zmult_zminus_right RS sym, |
|
488 zmult_zless_mono1, zless_zminus]) 1); |
|
489 qed "zmult_zless_mono1_neg"; |
|
490 |
|
491 Goal "[| i<j; k < (0::int) |] ==> k*j < k*i"; |
|
492 by (rtac (zminus_zless_zminus RS iffD1) 1); |
|
493 by (asm_simp_tac (simpset() addsimps [zmult_zminus RS sym, |
|
494 zmult_zless_mono2, zless_zminus]) 1); |
|
495 qed "zmult_zless_mono2_neg"; |
|
496 |
|
497 |
|
498 Goal "(m*n = (0::int)) = (m = 0 | n = 0)"; |
|
499 by (case_tac "m < (0::int)" 1); |
|
500 by (auto_tac (claset(), |
|
501 simpset() addsimps [linorder_not_less, order_le_less, |
|
502 linorder_neq_iff])); |
|
503 by (REPEAT |
|
504 (force_tac (claset() addDs [zmult_zless_mono1_neg, zmult_zless_mono1], |
|
505 simpset()) 1)); |
|
506 qed "zmult_eq_0_iff"; |
|
507 AddIffs [zmult_eq_0_iff]; |
|
508 |
|
509 |
|
510 (** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =, |
|
511 but not (yet?) for k*m < n*k. **) |
|
512 |
|
513 Goal "(m*k < n*k) = (((0::int) < k & m<n) | (k < 0 & n<m))"; |
|
514 by (case_tac "k = (0::int)" 1); |
|
515 by (auto_tac (claset(), simpset() addsimps [linorder_neq_iff, |
|
516 zmult_zless_mono1, zmult_zless_mono1_neg])); |
|
517 by (auto_tac (claset(), |
|
518 simpset() addsimps [linorder_not_less, |
|
519 inst "y1" "m*k" (linorder_not_le RS sym), |
|
520 inst "y1" "m" (linorder_not_le RS sym)])); |
|
521 by (ALLGOALS (etac notE)); |
|
522 by (auto_tac (claset(), simpset() addsimps [order_less_imp_le, zmult_zle_mono1, |
|
523 zmult_zle_mono1_neg])); |
|
524 qed "zmult_zless_cancel2"; |
|
525 |
|
526 |
|
527 Goal "(k*m < k*n) = (((0::int) < k & m<n) | (k < 0 & n<m))"; |
|
528 by (simp_tac (simpset() addsimps [inst "z" "k" zmult_commute, |
|
529 zmult_zless_cancel2]) 1); |
|
530 qed "zmult_zless_cancel1"; |
|
531 |
|
532 Goal "(m*k <= n*k) = (((0::int) < k --> m<=n) & (k < 0 --> n<=m))"; |
|
533 by (simp_tac (simpset() addsimps [linorder_not_less RS sym, |
|
534 zmult_zless_cancel2]) 1); |
|
535 qed "zmult_zle_cancel2"; |
|
536 |
|
537 Goal "(k*m <= k*n) = (((0::int) < k --> m<=n) & (k < 0 --> n<=m))"; |
|
538 by (simp_tac (simpset() addsimps [linorder_not_less RS sym, |
|
539 zmult_zless_cancel1]) 1); |
|
540 qed "zmult_zle_cancel1"; |
|
541 |
|
542 Goal "(m*k = n*k) = (k = (0::int) | m=n)"; |
|
543 by (cut_facts_tac [linorder_less_linear] 1); |
|
544 by Safe_tac; |
|
545 by Auto_tac; |
|
546 by (REPEAT |
|
547 (force_tac (claset() addD2 ("mono_neg", zmult_zless_mono1_neg) |
|
548 addD2 ("mono_pos", zmult_zless_mono1), |
|
549 simpset() addsimps [linorder_neq_iff]) 1)); |
|
550 |
|
551 qed "zmult_cancel2"; |
|
552 |
|
553 Goal "(k*m = k*n) = (k = (0::int) | m=n)"; |
|
554 by (simp_tac (simpset() addsimps [inst "z" "k" zmult_commute, |
|
555 zmult_cancel2]) 1); |
|
556 qed "zmult_cancel1"; |
|
557 Addsimps [zmult_cancel1, zmult_cancel2]; |
|
558 |
|
559 |
|
560 (*Analogous to zadd_int*) |
|
561 Goal "n<=m --> int m - int n = int (m-n)"; |
|
562 by (induct_thm_tac diff_induct "m n" 1); |
|
563 by (auto_tac (claset(), simpset() addsimps [int_Suc, symmetric zdiff_def])); |
|
564 qed_spec_mp "zdiff_int"; |