192 |
187 |
193 text{*Strict Monotonicity of Multiplication*} |
188 text{*Strict Monotonicity of Multiplication*} |
194 |
189 |
195 text{*strict, in 1st argument; proof is by induction on k>0*} |
190 text{*strict, in 1st argument; proof is by induction on k>0*} |
196 lemma zmult_zless_mono2_lemma: |
191 lemma zmult_zless_mono2_lemma: |
197 "(i::int)<j ==> 0<k ==> int k * i < int k * j" |
192 "(i::int)<j ==> 0<k ==> of_nat k * i < of_nat k * j" |
198 apply (induct "k", simp) |
193 apply (induct "k", simp) |
199 apply (simp add: left_distrib) |
194 apply (simp add: left_distrib) |
200 apply (case_tac "k=0") |
195 apply (case_tac "k=0") |
201 apply (simp_all add: add_strict_mono) |
196 apply (simp_all add: add_strict_mono) |
202 done |
197 done |
203 |
198 |
204 lemma zero_le_imp_eq_int: "(0::int) \<le> k ==> \<exists>n. k = int n" |
199 lemma zero_le_imp_eq_int: "(0::int) \<le> k ==> \<exists>n. k = of_nat n" |
205 apply (cases k) |
200 apply (cases k) |
206 apply (auto simp add: le add int_def Zero_int_def) |
201 apply (auto simp add: le add int_def Zero_int_def) |
207 apply (rule_tac x="x-y" in exI, simp) |
202 apply (rule_tac x="x-y" in exI, simp) |
208 done |
203 done |
209 |
204 |
210 lemma zero_less_imp_eq_int: "(0::int) < k ==> \<exists>n>0. k = int n" |
205 lemma zero_less_imp_eq_int: "(0::int) < k ==> \<exists>n>0. k = of_nat n" |
211 apply (cases k) |
206 apply (cases k) |
212 apply (simp add: less int_def Zero_int_def) |
207 apply (simp add: less int_def Zero_int_def) |
213 apply (rule_tac x="x-y" in exI, simp) |
208 apply (rule_tac x="x-y" in exI, simp) |
214 done |
209 done |
215 |
210 |
256 by (simp add: congruent_def) arith |
251 by (simp add: congruent_def) arith |
257 thus ?thesis |
252 thus ?thesis |
258 by (simp add: nat_def UN_equiv_class [OF equiv_intrel]) |
253 by (simp add: nat_def UN_equiv_class [OF equiv_intrel]) |
259 qed |
254 qed |
260 |
255 |
261 lemma nat_int [simp]: "nat (int n) = n" |
256 lemma nat_int [simp]: "nat (of_nat n) = n" |
262 by (simp add: nat int_def) |
257 by (simp add: nat int_def) |
263 |
258 |
264 lemma nat_zero [simp]: "nat 0 = 0" |
259 lemma nat_zero [simp]: "nat 0 = 0" |
265 by (simp add: Zero_int_def nat) |
260 by (simp add: Zero_int_def nat) |
266 |
261 |
267 lemma int_nat_eq [simp]: "int (nat z) = (if 0 \<le> z then z else 0)" |
262 lemma int_nat_eq [simp]: "of_nat (nat z) = (if 0 \<le> z then z else 0)" |
268 by (cases z, simp add: nat le int_def Zero_int_def) |
263 by (cases z, simp add: nat le int_def Zero_int_def) |
269 |
264 |
270 corollary nat_0_le: "0 \<le> z ==> int (nat z) = z" |
265 corollary nat_0_le: "0 \<le> z ==> of_nat (nat z) = z" |
271 by simp |
266 by simp |
272 |
267 |
273 lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0" |
268 lemma nat_le_0 [simp]: "z \<le> 0 ==> nat z = 0" |
274 by (cases z, simp add: nat le Zero_int_def) |
269 by (cases z, simp add: nat le Zero_int_def) |
275 |
270 |
288 lemma zless_nat_conj [simp]: "(nat w < nat z) = (0 < z & w < z)" |
283 lemma zless_nat_conj [simp]: "(nat w < nat z) = (0 < z & w < z)" |
289 apply (cases w, cases z) |
284 apply (cases w, cases z) |
290 apply (simp add: nat le Zero_int_def linorder_not_le [symmetric], arith) |
285 apply (simp add: nat le Zero_int_def linorder_not_le [symmetric], arith) |
291 done |
286 done |
292 |
287 |
293 lemma nonneg_eq_int: "[| 0 \<le> z; !!m. z = int m ==> P |] ==> P" |
288 lemma nonneg_eq_int: |
294 by (blast dest: nat_0_le sym) |
289 fixes z :: int |
295 |
290 assumes "0 \<le> z" and "\<And>m. z = of_nat m \<Longrightarrow> P" |
296 lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = int m else m=0)" |
291 shows P |
|
292 using assms by (blast dest: nat_0_le sym) |
|
293 |
|
294 lemma nat_eq_iff: "(nat w = m) = (if 0 \<le> w then w = of_nat m else m=0)" |
297 by (cases w, simp add: nat le int_def Zero_int_def, arith) |
295 by (cases w, simp add: nat le int_def Zero_int_def, arith) |
298 |
296 |
299 corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = int m else m=0)" |
297 corollary nat_eq_iff2: "(m = nat w) = (if 0 \<le> w then w = of_nat m else m=0)" |
300 by (simp only: eq_commute [of m] nat_eq_iff) |
298 by (simp only: eq_commute [of m] nat_eq_iff) |
301 |
299 |
302 lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < int m)" |
300 lemma nat_less_iff: "0 \<le> w ==> (nat w < m) = (w < of_nat m)" |
303 apply (cases w) |
301 apply (cases w) |
304 apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith) |
302 apply (simp add: nat le int_def Zero_int_def linorder_not_le [symmetric], arith) |
305 done |
303 done |
306 |
304 |
307 lemma int_eq_iff: "(int m = z) = (m = nat z & 0 \<le> z)" |
305 lemma int_eq_iff: "(of_nat m = z) = (m = nat z & 0 \<le> z)" |
308 by (auto simp add: nat_eq_iff2) |
306 by (auto simp add: nat_eq_iff2) |
309 |
307 |
310 lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)" |
308 lemma zero_less_nat_eq [simp]: "(0 < nat z) = (0 < z)" |
311 by (insert zless_nat_conj [of 0], auto) |
309 by (insert zless_nat_conj [of 0], auto) |
312 |
310 |
317 lemma nat_diff_distrib: |
315 lemma nat_diff_distrib: |
318 "[| (0::int) \<le> z'; z' \<le> z |] ==> nat (z-z') = nat z - nat z'" |
316 "[| (0::int) \<le> z'; z' \<le> z |] ==> nat (z-z') = nat z - nat z'" |
319 by (cases z, cases z', |
317 by (cases z, cases z', |
320 simp add: nat add minus diff_minus le Zero_int_def) |
318 simp add: nat add minus diff_minus le Zero_int_def) |
321 |
319 |
322 lemma nat_zminus_int [simp]: "nat (- (int n)) = 0" |
320 lemma nat_zminus_int [simp]: "nat (- (of_nat n)) = 0" |
323 by (simp add: int_def minus nat Zero_int_def) |
321 by (simp add: int_def minus nat Zero_int_def) |
324 |
322 |
325 lemma zless_nat_eq_int_zless: "(m < nat z) = (int m < z)" |
323 lemma zless_nat_eq_int_zless: "(m < nat z) = (of_nat m < z)" |
326 by (cases z, simp add: nat less int_def, arith) |
324 by (cases z, simp add: nat less int_def, arith) |
327 |
325 |
328 |
326 |
329 subsection{*Lemmas about the Function @{term int} and Orderings*} |
327 subsection{*Lemmas about the Function @{term of_nat} and Orderings*} |
330 |
328 |
331 lemma negative_zless_0: "- (int (Suc n)) < 0" |
329 lemma negative_zless_0: "- (of_nat (Suc n)) < (0 \<Colon> int)" |
332 by (simp add: order_less_le del: of_nat_Suc) |
330 by (simp add: order_less_le del: of_nat_Suc) |
333 |
331 |
334 lemma negative_zless [iff]: "- (int (Suc n)) < int m" |
332 lemma negative_zless [iff]: "- (of_nat (Suc n)) < (of_nat m \<Colon> int)" |
335 by (rule negative_zless_0 [THEN order_less_le_trans], simp) |
333 by (rule negative_zless_0 [THEN order_less_le_trans], simp) |
336 |
334 |
337 lemma negative_zle_0: "- int n \<le> 0" |
335 lemma negative_zle_0: "- of_nat n \<le> (0 \<Colon> int)" |
338 by (simp add: minus_le_iff) |
336 by (simp add: minus_le_iff) |
339 |
337 |
340 lemma negative_zle [iff]: "- int n \<le> int m" |
338 lemma negative_zle [iff]: "- of_nat n \<le> (of_nat m \<Colon> int)" |
341 by (rule order_trans [OF negative_zle_0 of_nat_0_le_iff]) |
339 by (rule order_trans [OF negative_zle_0 of_nat_0_le_iff]) |
342 |
340 |
343 lemma not_zle_0_negative [simp]: "~ (0 \<le> - (int (Suc n)))" |
341 lemma not_zle_0_negative [simp]: "~ (0 \<le> - (of_nat (Suc n) \<Colon> int))" |
344 by (subst le_minus_iff, simp del: of_nat_Suc) |
342 by (subst le_minus_iff, simp del: of_nat_Suc) |
345 |
343 |
346 lemma int_zle_neg: "(int n \<le> - int m) = (n = 0 & m = 0)" |
344 lemma int_zle_neg: "((of_nat n \<Colon> int) \<le> - of_nat m) = (n = 0 & m = 0)" |
347 by (simp add: int_def le minus Zero_int_def) |
345 by (simp add: int_def le minus Zero_int_def) |
348 |
346 |
349 lemma not_int_zless_negative [simp]: "~ (int n < - int m)" |
347 lemma not_int_zless_negative [simp]: "~ ((of_nat n \<Colon> int) < - of_nat m)" |
350 by (simp add: linorder_not_less) |
348 by (simp add: linorder_not_less) |
351 |
349 |
352 lemma negative_eq_positive [simp]: "(- int n = int m) = (n = 0 & m = 0)" |
350 lemma negative_eq_positive [simp]: "((- of_nat n \<Colon> int) = of_nat m) = (n = 0 & m = 0)" |
353 by (force simp add: order_eq_iff [of "- int n"] int_zle_neg) |
351 by (force simp add: order_eq_iff [of "- of_nat n"] int_zle_neg) |
354 |
352 |
355 lemma zle_iff_zadd: "(w \<le> z) = (\<exists>n. z = w + int n)" |
353 lemma zle_iff_zadd: "(w\<Colon>int) \<le> z \<longleftrightarrow> (\<exists>n. z = w + of_nat n)" |
356 proof - |
354 proof - |
357 have "(w \<le> z) = (0 \<le> z - w)" |
355 have "(w \<le> z) = (0 \<le> z - w)" |
358 by (simp only: le_diff_eq add_0_left) |
356 by (simp only: le_diff_eq add_0_left) |
359 also have "\<dots> = (\<exists>n. z - w = int n)" |
357 also have "\<dots> = (\<exists>n. z - w = of_nat n)" |
360 by (auto elim: zero_le_imp_eq_int) |
358 by (auto elim: zero_le_imp_eq_int) |
361 also have "\<dots> = (\<exists>n. z = w + int n)" |
359 also have "\<dots> = (\<exists>n. z = w + of_nat n)" |
362 by (simp only: group_simps) |
360 by (simp only: group_simps) |
363 finally show ?thesis . |
361 finally show ?thesis . |
364 qed |
362 qed |
365 |
363 |
366 lemma zadd_int_left: "(int m) + (int n + z) = int (m + n) + z" |
364 lemma zadd_int_left: "of_nat m + (of_nat n + z) = of_nat (m + n) + (z\<Colon>int)" |
367 by simp |
365 by simp |
368 |
366 |
369 lemma int_Suc0_eq_1: "int (Suc 0) = 1" |
367 lemma int_Suc0_eq_1: "of_nat (Suc 0) = (1\<Colon>int)" |
370 by simp |
368 by simp |
371 |
|
372 lemma abs_of_nat [simp]: "\<bar>of_nat n::'a::ordered_idom\<bar> = of_nat n" |
|
373 by (rule of_nat_0_le_iff [THEN abs_of_nonneg]) (* belongs in Nat.thy *) |
|
374 |
369 |
375 text{*This version is proved for all ordered rings, not just integers! |
370 text{*This version is proved for all ordered rings, not just integers! |
376 It is proved here because attribute @{text arith_split} is not available |
371 It is proved here because attribute @{text arith_split} is not available |
377 in theory @{text Ring_and_Field}. |
372 in theory @{text Ring_and_Field}. |
378 But is it really better than just rewriting with @{text abs_if}?*} |
373 But is it really better than just rewriting with @{text abs_if}?*} |
488 text{*Class for unital rings with characteristic zero. |
483 text{*Class for unital rings with characteristic zero. |
489 Includes non-ordered rings like the complex numbers.*} |
484 Includes non-ordered rings like the complex numbers.*} |
490 class ring_char_0 = ring_1 + semiring_char_0 |
485 class ring_char_0 = ring_1 + semiring_char_0 |
491 |
486 |
492 lemma of_int_eq_iff [simp]: |
487 lemma of_int_eq_iff [simp]: |
493 "(of_int w = (of_int z::'a::ring_char_0)) = (w = z)" |
488 "of_int w = (of_int z \<Colon> 'a\<Colon>ring_char_0) \<longleftrightarrow> w = z" |
494 apply (cases w, cases z, simp add: of_int) |
489 apply (cases w, cases z, simp add: of_int) |
495 apply (simp only: diff_eq_eq diff_add_eq eq_diff_eq) |
490 apply (simp only: diff_eq_eq diff_add_eq eq_diff_eq) |
496 apply (simp only: of_nat_add [symmetric] of_nat_eq_iff) |
491 apply (simp only: of_nat_add [symmetric] of_nat_eq_iff) |
497 done |
492 done |
498 |
493 |
584 |
579 |
585 text{*Now we replace the case analysis rule by a more conventional one: |
580 text{*Now we replace the case analysis rule by a more conventional one: |
586 whether an integer is negative or not.*} |
581 whether an integer is negative or not.*} |
587 |
582 |
588 lemma zless_iff_Suc_zadd: |
583 lemma zless_iff_Suc_zadd: |
589 "(w < z) = (\<exists>n. z = w + int (Suc n))" |
584 "(w \<Colon> int) < z \<longleftrightarrow> (\<exists>n. z = w + of_nat (Suc n))" |
590 apply (cases z, cases w) |
585 apply (cases z, cases w) |
591 apply (auto simp add: less add int_def) |
586 apply (auto simp add: less add int_def) |
592 apply (rename_tac a b c d) |
587 apply (rename_tac a b c d) |
593 apply (rule_tac x="a+d - Suc(c+b)" in exI) |
588 apply (rule_tac x="a+d - Suc(c+b)" in exI) |
594 apply arith |
589 apply arith |
595 done |
590 done |
596 |
591 |
597 lemma negD: "x<0 ==> \<exists>n. x = - (int (Suc n))" |
592 lemma negD: "(x \<Colon> int) < 0 \<Longrightarrow> \<exists>n. x = - (of_nat (Suc n))" |
598 apply (cases x) |
593 apply (cases x) |
599 apply (auto simp add: le minus Zero_int_def int_def order_less_le) |
594 apply (auto simp add: le minus Zero_int_def int_def order_less_le) |
600 apply (rule_tac x="y - Suc x" in exI, arith) |
595 apply (rule_tac x="y - Suc x" in exI, arith) |
601 done |
596 done |
602 |
597 |
603 theorem int_cases [cases type: int, case_names nonneg neg]: |
598 theorem int_cases [cases type: int, case_names nonneg neg]: |
604 "[|!! n. z = int n ==> P; !! n. z = - (int (Suc n)) ==> P |] ==> P" |
599 "[|!! n. (z \<Colon> int) = of_nat n ==> P; !! n. z = - (of_nat (Suc n)) ==> P |] ==> P" |
605 apply (cases "z < 0", blast dest!: negD) |
600 apply (cases "z < 0", blast dest!: negD) |
606 apply (simp add: linorder_not_less del: of_nat_Suc) |
601 apply (simp add: linorder_not_less del: of_nat_Suc) |
607 apply (blast dest: nat_0_le [THEN sym]) |
602 apply (blast dest: nat_0_le [THEN sym]) |
608 done |
603 done |
609 |
604 |
610 theorem int_induct [induct type: int, case_names nonneg neg]: |
605 theorem int_induct [induct type: int, case_names nonneg neg]: |
611 "[|!! n. P (int n); !!n. P (- (int (Suc n))) |] ==> P z" |
606 "[|!! n. P (of_nat n \<Colon> int); !!n. P (- (of_nat (Suc n))) |] ==> P z" |
612 by (cases z rule: int_cases) auto |
607 by (cases z rule: int_cases) auto |
613 |
608 |
614 text{*Contributed by Brian Huffman*} |
609 text{*Contributed by Brian Huffman*} |
615 theorem int_diff_cases [case_names diff]: |
610 theorem int_diff_cases [case_names diff]: |
616 assumes prem: "!!m n. z = int m - int n ==> P" shows "P" |
611 assumes prem: "!!m n. (z\<Colon>int) = of_nat m - of_nat n ==> P" shows "P" |
617 apply (cases z rule: eq_Abs_Integ) |
612 apply (cases z rule: eq_Abs_Integ) |
618 apply (rule_tac m=x and n=y in prem) |
613 apply (rule_tac m=x and n=y in prem) |
619 apply (simp add: int_def diff_def minus add) |
614 apply (simp add: int_def diff_def minus add) |
620 done |
615 done |
621 |
616 |
671 lemmas int_less_0_conv = of_nat_less_0_iff [where 'a=int and m="?k"] |
666 lemmas int_less_0_conv = of_nat_less_0_iff [where 'a=int and m="?k"] |
672 lemmas zero_less_int_conv = of_nat_0_less_iff [where 'a=int] |
667 lemmas zero_less_int_conv = of_nat_0_less_iff [where 'a=int] |
673 lemmas zle_int = of_nat_le_iff [where 'a=int] |
668 lemmas zle_int = of_nat_le_iff [where 'a=int] |
674 lemmas zero_zle_int = of_nat_0_le_iff [where 'a=int] |
669 lemmas zero_zle_int = of_nat_0_le_iff [where 'a=int] |
675 lemmas int_le_0_conv = of_nat_le_0_iff [where 'a=int and m="?n"] |
670 lemmas int_le_0_conv = of_nat_le_0_iff [where 'a=int and m="?n"] |
676 lemmas int_0 = of_nat_0 [where ?'a_1.0=int] |
671 lemmas int_0 = of_nat_0 [where 'a=int] |
677 lemmas int_1 = of_nat_1 [where 'a=int] |
672 lemmas int_1 = of_nat_1 [where 'a=int] |
678 lemmas int_Suc = of_nat_Suc [where ?'a_1.0=int] |
673 lemmas int_Suc = of_nat_Suc [where 'a=int] |
679 lemmas abs_int_eq = abs_of_nat [where 'a=int and n="?m"] |
674 lemmas abs_int_eq = abs_of_nat [where 'a=int and n="?m"] |
680 lemmas of_int_int_eq = of_int_of_nat_eq [where 'a=int] |
675 lemmas of_int_int_eq = of_int_of_nat_eq [where 'a=int] |
681 lemmas zdiff_int = of_nat_diff [where 'a=int, symmetric] |
676 lemmas zdiff_int = of_nat_diff [where 'a=int, symmetric] |
682 lemmas zless_le = less_int_def [THEN meta_eq_to_obj_eq] |
677 lemmas zless_le = less_int_def [THEN meta_eq_to_obj_eq] |
683 lemmas int_eq_of_nat = TrueI |
678 lemmas int_eq_of_nat = TrueI |
684 |
679 |
685 abbreviation |
680 abbreviation |
|
681 int :: "nat \<Rightarrow> int" |
|
682 where |
|
683 "int \<equiv> of_nat" |
|
684 |
|
685 abbreviation |
686 int_of_nat :: "nat \<Rightarrow> int" |
686 int_of_nat :: "nat \<Rightarrow> int" |
687 where |
687 where |
688 "int_of_nat \<equiv> of_nat" |
688 "int_of_nat \<equiv> of_nat" |
689 |
689 |
690 end |
690 end |