| author | wenzelm | 
| Mon, 15 Nov 2010 17:40:38 +0100 | |
| changeset 40547 | 05a82b4bccbc | 
| parent 39489 | 8bb7f32a3a08 | 
| child 41550 | efa734d9b221 | 
| permissions | -rw-r--r-- | 
| 3366 | 1 | (* Title: HOL/Divides.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 3 | Copyright 1999 University of Cambridge | 
| 18154 | 4 | *) | 
| 3366 | 5 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 6 | header {* The division operators div and mod *}
 | 
| 3366 | 7 | |
| 15131 | 8 | theory Divides | 
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33296diff
changeset | 9 | imports Nat_Numeral Nat_Transfer | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 10 | uses "~~/src/Provers/Arith/cancel_div_mod.ML" | 
| 15131 | 11 | begin | 
| 3366 | 12 | |
| 25942 | 13 | subsection {* Syntactic division operations *}
 | 
| 14 | ||
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 15 | class div = dvd + | 
| 27540 | 16 | fixes div :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "div" 70) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 17 | and mod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "mod" 70) | 
| 27540 | 18 | |
| 19 | ||
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 20 | subsection {* Abstract division in commutative semirings. *}
 | 
| 25942 | 21 | |
| 30930 | 22 | class semiring_div = comm_semiring_1_cancel + no_zero_divisors + div + | 
| 25942 | 23 | assumes mod_div_equality: "a div b * b + a mod b = a" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 24 | and div_by_0 [simp]: "a div 0 = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 25 | and div_0 [simp]: "0 div a = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 26 | and div_mult_self1 [simp]: "b \<noteq> 0 \<Longrightarrow> (a + c * b) div b = c + a div b" | 
| 30930 | 27 | and div_mult_mult1 [simp]: "c \<noteq> 0 \<Longrightarrow> (c * a) div (c * b) = a div b" | 
| 25942 | 28 | begin | 
| 29 | ||
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 30 | text {* @{const div} and @{const mod} *}
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 31 | |
| 26062 | 32 | lemma mod_div_equality2: "b * (a div b) + a mod b = a" | 
| 33 | unfolding mult_commute [of b] | |
| 34 | by (rule mod_div_equality) | |
| 35 | ||
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 36 | lemma mod_div_equality': "a mod b + a div b * b = a" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 37 | using mod_div_equality [of a b] | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 38 | by (simp only: add_ac) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 39 | |
| 26062 | 40 | lemma div_mod_equality: "((a div b) * b + a mod b) + c = a + c" | 
| 30934 | 41 | by (simp add: mod_div_equality) | 
| 26062 | 42 | |
| 43 | lemma div_mod_equality2: "(b * (a div b) + a mod b) + c = a + c" | |
| 30934 | 44 | by (simp add: mod_div_equality2) | 
| 26062 | 45 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 46 | lemma mod_by_0 [simp]: "a mod 0 = a" | 
| 30934 | 47 | using mod_div_equality [of a zero] by simp | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 48 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 49 | lemma mod_0 [simp]: "0 mod a = 0" | 
| 30934 | 50 | using mod_div_equality [of zero a] div_0 by simp | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 51 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 52 | lemma div_mult_self2 [simp]: | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 53 | assumes "b \<noteq> 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 54 | shows "(a + b * c) div b = c + a div b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 55 | using assms div_mult_self1 [of b a c] by (simp add: mult_commute) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 56 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 57 | lemma mod_mult_self1 [simp]: "(a + c * b) mod b = a mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 58 | proof (cases "b = 0") | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 59 | case True then show ?thesis by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 60 | next | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 61 | case False | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 62 | have "a + c * b = (a + c * b) div b * b + (a + c * b) mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 63 | by (simp add: mod_div_equality) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 64 | also from False div_mult_self1 [of b a c] have | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 65 | "\<dots> = (c + a div b) * b + (a + c * b) mod b" | 
| 29667 | 66 | by (simp add: algebra_simps) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 67 | finally have "a = a div b * b + (a + c * b) mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 68 | by (simp add: add_commute [of a] add_assoc left_distrib) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 69 | then have "a div b * b + (a + c * b) mod b = a div b * b + a mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 70 | by (simp add: mod_div_equality) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 71 | then show ?thesis by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 72 | qed | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 73 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 74 | lemma mod_mult_self2 [simp]: "(a + b * c) mod b = a mod b" | 
| 30934 | 75 | by (simp add: mult_commute [of b]) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 76 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 77 | lemma div_mult_self1_is_id [simp]: "b \<noteq> 0 \<Longrightarrow> b * a div b = a" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 78 | using div_mult_self2 [of b 0 a] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 79 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 80 | lemma div_mult_self2_is_id [simp]: "b \<noteq> 0 \<Longrightarrow> a * b div b = a" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 81 | using div_mult_self1 [of b 0 a] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 82 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 83 | lemma mod_mult_self1_is_0 [simp]: "b * a mod b = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 84 | using mod_mult_self2 [of 0 b a] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 85 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 86 | lemma mod_mult_self2_is_0 [simp]: "a * b mod b = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 87 | using mod_mult_self1 [of 0 a b] by simp | 
| 26062 | 88 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 89 | lemma div_by_1 [simp]: "a div 1 = a" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 90 | using div_mult_self2_is_id [of 1 a] zero_neq_one by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 91 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 92 | lemma mod_by_1 [simp]: "a mod 1 = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 93 | proof - | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 94 | from mod_div_equality [of a one] div_by_1 have "a + a mod 1 = a" by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 95 | then have "a + a mod 1 = a + 0" by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 96 | then show ?thesis by (rule add_left_imp_eq) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 97 | qed | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 98 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 99 | lemma mod_self [simp]: "a mod a = 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 100 | using mod_mult_self2_is_0 [of 1] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 101 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 102 | lemma div_self [simp]: "a \<noteq> 0 \<Longrightarrow> a div a = 1" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 103 | using div_mult_self2_is_id [of _ 1] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 104 | |
| 27676 | 105 | lemma div_add_self1 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 106 | assumes "b \<noteq> 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 107 | shows "(b + a) div b = a div b + 1" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 108 | using assms div_mult_self1 [of b a 1] by (simp add: add_commute) | 
| 26062 | 109 | |
| 27676 | 110 | lemma div_add_self2 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 111 | assumes "b \<noteq> 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 112 | shows "(a + b) div b = a div b + 1" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 113 | using assms div_add_self1 [of b a] by (simp add: add_commute) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 114 | |
| 27676 | 115 | lemma mod_add_self1 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 116 | "(b + a) mod b = a mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 117 | using mod_mult_self1 [of a 1 b] by (simp add: add_commute) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 118 | |
| 27676 | 119 | lemma mod_add_self2 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 120 | "(a + b) mod b = a mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 121 | using mod_mult_self1 [of a 1 b] by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 122 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 123 | lemma mod_div_decomp: | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 124 | fixes a b | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 125 | obtains q r where "q = a div b" and "r = a mod b" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 126 | and "a = q * b + r" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 127 | proof - | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 128 | from mod_div_equality have "a = a div b * b + a mod b" by simp | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 129 | moreover have "a div b = a div b" .. | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 130 | moreover have "a mod b = a mod b" .. | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 131 | note that ultimately show thesis by blast | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 132 | qed | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 133 | |
| 33364 | 134 | lemma dvd_eq_mod_eq_0 [code, code_unfold, code_inline del]: "a dvd b \<longleftrightarrow> b mod a = 0" | 
| 25942 | 135 | proof | 
| 136 | assume "b mod a = 0" | |
| 137 | with mod_div_equality [of b a] have "b div a * a = b" by simp | |
| 138 | then have "b = a * (b div a)" unfolding mult_commute .. | |
| 139 | then have "\<exists>c. b = a * c" .. | |
| 140 | then show "a dvd b" unfolding dvd_def . | |
| 141 | next | |
| 142 | assume "a dvd b" | |
| 143 | then have "\<exists>c. b = a * c" unfolding dvd_def . | |
| 144 | then obtain c where "b = a * c" .. | |
| 145 | then have "b mod a = a * c mod a" by simp | |
| 146 | then have "b mod a = c * a mod a" by (simp add: mult_commute) | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 147 | then show "b mod a = 0" by simp | 
| 25942 | 148 | qed | 
| 149 | ||
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 150 | lemma mod_div_trivial [simp]: "a mod b div b = 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 151 | proof (cases "b = 0") | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 152 | assume "b = 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 153 | thus ?thesis by simp | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 154 | next | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 155 | assume "b \<noteq> 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 156 | hence "a div b + a mod b div b = (a mod b + a div b * b) div b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 157 | by (rule div_mult_self1 [symmetric]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 158 | also have "\<dots> = a div b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 159 | by (simp only: mod_div_equality') | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 160 | also have "\<dots> = a div b + 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 161 | by simp | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 162 | finally show ?thesis | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 163 | by (rule add_left_imp_eq) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 164 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 165 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 166 | lemma mod_mod_trivial [simp]: "a mod b mod b = a mod b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 167 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 168 | have "a mod b mod b = (a mod b + a div b * b) mod b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 169 | by (simp only: mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 170 | also have "\<dots> = a mod b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 171 | by (simp only: mod_div_equality') | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 172 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 173 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 174 | |
| 29925 | 175 | lemma dvd_imp_mod_0: "a dvd b \<Longrightarrow> b mod a = 0" | 
| 29948 | 176 | by (rule dvd_eq_mod_eq_0[THEN iffD1]) | 
| 29925 | 177 | |
| 178 | lemma dvd_div_mult_self: "a dvd b \<Longrightarrow> (b div a) * a = b" | |
| 179 | by (subst (2) mod_div_equality [of b a, symmetric]) (simp add:dvd_imp_mod_0) | |
| 180 | ||
| 33274 
b6ff7db522b5
moved lemmas for dvd on nat to theories Nat and Power
 haftmann parents: 
32010diff
changeset | 181 | lemma dvd_mult_div_cancel: "a dvd b \<Longrightarrow> a * (b div a) = b" | 
| 
b6ff7db522b5
moved lemmas for dvd on nat to theories Nat and Power
 haftmann parents: 
32010diff
changeset | 182 | by (drule dvd_div_mult_self) (simp add: mult_commute) | 
| 
b6ff7db522b5
moved lemmas for dvd on nat to theories Nat and Power
 haftmann parents: 
32010diff
changeset | 183 | |
| 30052 | 184 | lemma dvd_div_mult: "a dvd b \<Longrightarrow> (b div a) * c = b * c div a" | 
| 185 | apply (cases "a = 0") | |
| 186 | apply simp | |
| 187 | apply (auto simp: dvd_def mult_assoc) | |
| 188 | done | |
| 189 | ||
| 29925 | 190 | lemma div_dvd_div[simp]: | 
| 191 | "a dvd b \<Longrightarrow> a dvd c \<Longrightarrow> (b div a dvd c div a) = (b dvd c)" | |
| 192 | apply (cases "a = 0") | |
| 193 | apply simp | |
| 194 | apply (unfold dvd_def) | |
| 195 | apply auto | |
| 196 | apply(blast intro:mult_assoc[symmetric]) | |
| 197 | apply(fastsimp simp add: mult_assoc) | |
| 198 | done | |
| 199 | ||
| 30078 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 200 | lemma dvd_mod_imp_dvd: "[| k dvd m mod n; k dvd n |] ==> k dvd m" | 
| 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 201 | apply (subgoal_tac "k dvd (m div n) *n + m mod n") | 
| 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 202 | apply (simp add: mod_div_equality) | 
| 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 203 | apply (simp only: dvd_add dvd_mult) | 
| 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 204 | done | 
| 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 205 | |
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 206 | text {* Addition respects modular equivalence. *}
 | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 207 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 208 | lemma mod_add_left_eq: "(a + b) mod c = (a mod c + b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 209 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 210 | have "(a + b) mod c = (a div c * c + a mod c + b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 211 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 212 | also have "\<dots> = (a mod c + b + a div c * c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 213 | by (simp only: add_ac) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 214 | also have "\<dots> = (a mod c + b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 215 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 216 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 217 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 218 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 219 | lemma mod_add_right_eq: "(a + b) mod c = (a + b mod c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 220 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 221 | have "(a + b) mod c = (a + (b div c * c + b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 222 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 223 | also have "\<dots> = (a + b mod c + b div c * c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 224 | by (simp only: add_ac) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 225 | also have "\<dots> = (a + b mod c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 226 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 227 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 228 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 229 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 230 | lemma mod_add_eq: "(a + b) mod c = (a mod c + b mod c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 231 | by (rule trans [OF mod_add_left_eq mod_add_right_eq]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 232 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 233 | lemma mod_add_cong: | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 234 | assumes "a mod c = a' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 235 | assumes "b mod c = b' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 236 | shows "(a + b) mod c = (a' + b') mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 237 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 238 | have "(a mod c + b mod c) mod c = (a' mod c + b' mod c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 239 | unfolding assms .. | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 240 | thus ?thesis | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 241 | by (simp only: mod_add_eq [symmetric]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 242 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 243 | |
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 244 | lemma div_add [simp]: "z dvd x \<Longrightarrow> z dvd y | 
| 30837 
3d4832d9f7e4
added strong_setprod_cong[cong] (in analogy with setsum)
 nipkow parents: 
30729diff
changeset | 245 | \<Longrightarrow> (x + y) div z = x div z + y div z" | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 246 | by (cases "z = 0", simp, unfold dvd_def, auto simp add: algebra_simps) | 
| 30837 
3d4832d9f7e4
added strong_setprod_cong[cong] (in analogy with setsum)
 nipkow parents: 
30729diff
changeset | 247 | |
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 248 | text {* Multiplication respects modular equivalence. *}
 | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 249 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 250 | lemma mod_mult_left_eq: "(a * b) mod c = ((a mod c) * b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 251 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 252 | have "(a * b) mod c = ((a div c * c + a mod c) * b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 253 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 254 | also have "\<dots> = (a mod c * b + a div c * b * c) mod c" | 
| 29667 | 255 | by (simp only: algebra_simps) | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 256 | also have "\<dots> = (a mod c * b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 257 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 258 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 259 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 260 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 261 | lemma mod_mult_right_eq: "(a * b) mod c = (a * (b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 262 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 263 | have "(a * b) mod c = (a * (b div c * c + b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 264 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 265 | also have "\<dots> = (a * (b mod c) + a * (b div c) * c) mod c" | 
| 29667 | 266 | by (simp only: algebra_simps) | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 267 | also have "\<dots> = (a * (b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 268 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 269 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 270 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 271 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 272 | lemma mod_mult_eq: "(a * b) mod c = ((a mod c) * (b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 273 | by (rule trans [OF mod_mult_left_eq mod_mult_right_eq]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 274 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 275 | lemma mod_mult_cong: | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 276 | assumes "a mod c = a' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 277 | assumes "b mod c = b' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 278 | shows "(a * b) mod c = (a' * b') mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 279 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 280 | have "(a mod c * (b mod c)) mod c = (a' mod c * (b' mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 281 | unfolding assms .. | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 282 | thus ?thesis | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 283 | by (simp only: mod_mult_eq [symmetric]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 284 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 285 | |
| 29404 | 286 | lemma mod_mod_cancel: | 
| 287 | assumes "c dvd b" | |
| 288 | shows "a mod b mod c = a mod c" | |
| 289 | proof - | |
| 290 | from `c dvd b` obtain k where "b = c * k" | |
| 291 | by (rule dvdE) | |
| 292 | have "a mod b mod c = a mod (c * k) mod c" | |
| 293 | by (simp only: `b = c * k`) | |
| 294 | also have "\<dots> = (a mod (c * k) + a div (c * k) * k * c) mod c" | |
| 295 | by (simp only: mod_mult_self1) | |
| 296 | also have "\<dots> = (a div (c * k) * (c * k) + a mod (c * k)) mod c" | |
| 297 | by (simp only: add_ac mult_ac) | |
| 298 | also have "\<dots> = a mod c" | |
| 299 | by (simp only: mod_div_equality) | |
| 300 | finally show ?thesis . | |
| 301 | qed | |
| 302 | ||
| 30930 | 303 | lemma div_mult_div_if_dvd: | 
| 304 | "y dvd x \<Longrightarrow> z dvd w \<Longrightarrow> (x div y) * (w div z) = (x * w) div (y * z)" | |
| 305 | apply (cases "y = 0", simp) | |
| 306 | apply (cases "z = 0", simp) | |
| 307 | apply (auto elim!: dvdE simp add: algebra_simps) | |
| 30476 | 308 | apply (subst mult_assoc [symmetric]) | 
| 309 | apply (simp add: no_zero_divisors) | |
| 30930 | 310 | done | 
| 311 | ||
| 35367 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 312 | lemma div_mult_swap: | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 313 | assumes "c dvd b" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 314 | shows "a * (b div c) = (a * b) div c" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 315 | proof - | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 316 | from assms have "b div c * (a div 1) = b * a div (c * 1)" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 317 | by (simp only: div_mult_div_if_dvd one_dvd) | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 318 | then show ?thesis by (simp add: mult_commute) | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 319 | qed | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 320 | |
| 30930 | 321 | lemma div_mult_mult2 [simp]: | 
| 322 | "c \<noteq> 0 \<Longrightarrow> (a * c) div (b * c) = a div b" | |
| 323 | by (drule div_mult_mult1) (simp add: mult_commute) | |
| 324 | ||
| 325 | lemma div_mult_mult1_if [simp]: | |
| 326 | "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" | |
| 327 | by simp_all | |
| 30476 | 328 | |
| 30930 | 329 | lemma mod_mult_mult1: | 
| 330 | "(c * a) mod (c * b) = c * (a mod b)" | |
| 331 | proof (cases "c = 0") | |
| 332 | case True then show ?thesis by simp | |
| 333 | next | |
| 334 | case False | |
| 335 | from mod_div_equality | |
| 336 | have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . | |
| 337 | with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) | |
| 338 | = c * a + c * (a mod b)" by (simp add: algebra_simps) | |
| 339 | with mod_div_equality show ?thesis by simp | |
| 340 | qed | |
| 341 | ||
| 342 | lemma mod_mult_mult2: | |
| 343 | "(a * c) mod (b * c) = (a mod b) * c" | |
| 344 | using mod_mult_mult1 [of c a b] by (simp add: mult_commute) | |
| 345 | ||
| 31662 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 346 | lemma dvd_mod: "k dvd m \<Longrightarrow> k dvd n \<Longrightarrow> k dvd (m mod n)" | 
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 347 | unfolding dvd_def by (auto simp add: mod_mult_mult1) | 
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 348 | |
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 349 | lemma dvd_mod_iff: "k dvd n \<Longrightarrow> k dvd (m mod n) \<longleftrightarrow> k dvd m" | 
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 350 | by (blast intro: dvd_mod_imp_dvd dvd_mod) | 
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 351 | |
| 31009 
41fd307cab30
dropped reference to class recpower and lemma duplicate
 haftmann parents: 
30934diff
changeset | 352 | lemma div_power: | 
| 31661 
1e252b8b2334
move lemma div_power into semiring_div context; class ring_div inherits from idom
 huffman parents: 
31009diff
changeset | 353 | "y dvd x \<Longrightarrow> (x div y) ^ n = x ^ n div y ^ n" | 
| 30476 | 354 | apply (induct n) | 
| 355 | apply simp | |
| 356 | apply(simp add: div_mult_div_if_dvd dvd_power_same) | |
| 357 | done | |
| 358 | ||
| 35367 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 359 | lemma dvd_div_eq_mult: | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 360 | assumes "a \<noteq> 0" and "a dvd b" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 361 | shows "b div a = c \<longleftrightarrow> b = c * a" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 362 | proof | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 363 | assume "b = c * a" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 364 | then show "b div a = c" by (simp add: assms) | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 365 | next | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 366 | assume "b div a = c" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 367 | then have "b div a * a = c * a" by simp | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 368 | moreover from `a dvd b` have "b div a * a = b" by (simp add: dvd_div_mult_self) | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 369 | ultimately show "b = c * a" by simp | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 370 | qed | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 371 | |
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 372 | lemma dvd_div_div_eq_mult: | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 373 | assumes "a \<noteq> 0" "c \<noteq> 0" and "a dvd b" "c dvd d" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 374 | shows "b div a = d div c \<longleftrightarrow> b * c = a * d" | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 375 | using assms by (auto simp add: mult_commute [of _ a] dvd_div_mult_self dvd_div_eq_mult div_mult_swap intro: sym) | 
| 
45a193f0ed0c
lemma div_mult_swap, dvd_div_eq_mult, dvd_div_div_eq_mult
 haftmann parents: 
35216diff
changeset | 376 | |
| 31661 
1e252b8b2334
move lemma div_power into semiring_div context; class ring_div inherits from idom
 huffman parents: 
31009diff
changeset | 377 | end | 
| 
1e252b8b2334
move lemma div_power into semiring_div context; class ring_div inherits from idom
 huffman parents: 
31009diff
changeset | 378 | |
| 35673 | 379 | class ring_div = semiring_div + comm_ring_1 | 
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 380 | begin | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 381 | |
| 36634 | 382 | subclass ring_1_no_zero_divisors .. | 
| 383 | ||
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 384 | text {* Negation respects modular equivalence. *}
 | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 385 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 386 | lemma mod_minus_eq: "(- a) mod b = (- (a mod b)) mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 387 | proof - | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 388 | have "(- a) mod b = (- (a div b * b + a mod b)) mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 389 | by (simp only: mod_div_equality) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 390 | also have "\<dots> = (- (a mod b) + - (a div b) * b) mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 391 | by (simp only: minus_add_distrib minus_mult_left add_ac) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 392 | also have "\<dots> = (- (a mod b)) mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 393 | by (rule mod_mult_self1) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 394 | finally show ?thesis . | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 395 | qed | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 396 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 397 | lemma mod_minus_cong: | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 398 | assumes "a mod b = a' mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 399 | shows "(- a) mod b = (- a') mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 400 | proof - | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 401 | have "(- (a mod b)) mod b = (- (a' mod b)) mod b" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 402 | unfolding assms .. | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 403 | thus ?thesis | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 404 | by (simp only: mod_minus_eq [symmetric]) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 405 | qed | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 406 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 407 | text {* Subtraction respects modular equivalence. *}
 | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 408 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 409 | lemma mod_diff_left_eq: "(a - b) mod c = (a mod c - b) mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 410 | unfolding diff_minus | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 411 | by (intro mod_add_cong mod_minus_cong) simp_all | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 412 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 413 | lemma mod_diff_right_eq: "(a - b) mod c = (a - b mod c) mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 414 | unfolding diff_minus | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 415 | by (intro mod_add_cong mod_minus_cong) simp_all | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 416 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 417 | lemma mod_diff_eq: "(a - b) mod c = (a mod c - b mod c) mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 418 | unfolding diff_minus | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 419 | by (intro mod_add_cong mod_minus_cong) simp_all | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 420 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 421 | lemma mod_diff_cong: | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 422 | assumes "a mod c = a' mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 423 | assumes "b mod c = b' mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 424 | shows "(a - b) mod c = (a' - b') mod c" | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 425 | unfolding diff_minus using assms | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 426 | by (intro mod_add_cong mod_minus_cong) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 427 | |
| 30180 | 428 | lemma dvd_neg_div: "y dvd x \<Longrightarrow> -x div y = - (x div y)" | 
| 429 | apply (case_tac "y = 0") apply simp | |
| 430 | apply (auto simp add: dvd_def) | |
| 431 | apply (subgoal_tac "-(y * k) = y * - k") | |
| 432 | apply (erule ssubst) | |
| 433 | apply (erule div_mult_self1_is_id) | |
| 434 | apply simp | |
| 435 | done | |
| 436 | ||
| 437 | lemma dvd_div_neg: "y dvd x \<Longrightarrow> x div -y = - (x div y)" | |
| 438 | apply (case_tac "y = 0") apply simp | |
| 439 | apply (auto simp add: dvd_def) | |
| 440 | apply (subgoal_tac "y * k = -y * -k") | |
| 441 | apply (erule ssubst) | |
| 442 | apply (rule div_mult_self1_is_id) | |
| 443 | apply simp | |
| 444 | apply simp | |
| 445 | done | |
| 446 | ||
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 447 | end | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 448 | |
| 25942 | 449 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 450 | subsection {* Division on @{typ nat} *}
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 451 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 452 | text {*
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 453 |   We define @{const div} and @{const mod} on @{typ nat} by means
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 454 | of a characteristic relation with two input arguments | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 455 |   @{term "m\<Colon>nat"}, @{term "n\<Colon>nat"} and two output arguments
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 456 |   @{term "q\<Colon>nat"}(uotient) and @{term "r\<Colon>nat"}(emainder).
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 457 | *} | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 458 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 459 | definition divmod_nat_rel :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat \<Rightarrow> bool" where | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 460 | "divmod_nat_rel m n qr \<longleftrightarrow> | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 461 | m = fst qr * n + snd qr \<and> | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 462 | (if n = 0 then fst qr = 0 else if n > 0 then 0 \<le> snd qr \<and> snd qr < n else n < snd qr \<and> snd qr \<le> 0)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 463 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 464 | text {* @{const divmod_nat_rel} is total: *}
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 465 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 466 | lemma divmod_nat_rel_ex: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 467 | obtains q r where "divmod_nat_rel m n (q, r)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 468 | proof (cases "n = 0") | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 469 | case True with that show thesis | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 470 | by (auto simp add: divmod_nat_rel_def) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 471 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 472 | case False | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 473 | have "\<exists>q r. m = q * n + r \<and> r < n" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 474 | proof (induct m) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 475 | case 0 with `n \<noteq> 0` | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 476 | have "(0\<Colon>nat) = 0 * n + 0 \<and> 0 < n" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 477 | then show ?case by blast | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 478 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 479 | case (Suc m) then obtain q' r' | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 480 | where m: "m = q' * n + r'" and n: "r' < n" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 481 | then show ?case proof (cases "Suc r' < n") | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 482 | case True | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 483 | from m n have "Suc m = q' * n + Suc r'" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 484 | with True show ?thesis by blast | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 485 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 486 | case False then have "n \<le> Suc r'" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 487 | moreover from n have "Suc r' \<le> n" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 488 | ultimately have "n = Suc r'" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 489 | with m have "Suc m = Suc q' * n + 0" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 490 | with `n \<noteq> 0` show ?thesis by blast | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 491 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 492 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 493 | with that show thesis | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 494 | using `n \<noteq> 0` by (auto simp add: divmod_nat_rel_def) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 495 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 496 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 497 | text {* @{const divmod_nat_rel} is injective: *}
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 498 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 499 | lemma divmod_nat_rel_unique: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 500 | assumes "divmod_nat_rel m n qr" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 501 | and "divmod_nat_rel m n qr'" | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 502 | shows "qr = qr'" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 503 | proof (cases "n = 0") | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 504 | case True with assms show ?thesis | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 505 | by (cases qr, cases qr') | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 506 | (simp add: divmod_nat_rel_def) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 507 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 508 | case False | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 509 | have aux: "\<And>q r q' r'. q' * n + r' = q * n + r \<Longrightarrow> r < n \<Longrightarrow> q' \<le> (q\<Colon>nat)" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 510 | apply (rule leI) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 511 | apply (subst less_iff_Suc_add) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 512 | apply (auto simp add: add_mult_distrib) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 513 | done | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 514 | from `n \<noteq> 0` assms have "fst qr = fst qr'" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 515 | by (auto simp add: divmod_nat_rel_def intro: order_antisym dest: aux sym) | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 516 | moreover from this assms have "snd qr = snd qr'" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 517 | by (simp add: divmod_nat_rel_def) | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 518 | ultimately show ?thesis by (cases qr, cases qr') simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 519 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 520 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 521 | text {*
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 522 | We instantiate divisibility on the natural numbers by | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 523 |   means of @{const divmod_nat_rel}:
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 524 | *} | 
| 25942 | 525 | |
| 526 | instantiation nat :: semiring_div | |
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25162diff
changeset | 527 | begin | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25162diff
changeset | 528 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 529 | definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" where | 
| 37767 | 530 | "divmod_nat m n = (THE qr. divmod_nat_rel m n qr)" | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 531 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 532 | lemma divmod_nat_rel_divmod_nat: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 533 | "divmod_nat_rel m n (divmod_nat m n)" | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 534 | proof - | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 535 | from divmod_nat_rel_ex | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 536 | obtain qr where rel: "divmod_nat_rel m n qr" . | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 537 | then show ?thesis | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 538 | by (auto simp add: divmod_nat_def intro: theI elim: divmod_nat_rel_unique) | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 539 | qed | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 540 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 541 | lemma divmod_nat_eq: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 542 | assumes "divmod_nat_rel m n qr" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 543 | shows "divmod_nat m n = qr" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 544 | using assms by (auto intro: divmod_nat_rel_unique divmod_nat_rel_divmod_nat) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 545 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 546 | definition div_nat where | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 547 | "m div n = fst (divmod_nat m n)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 548 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 549 | definition mod_nat where | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 550 | "m mod n = snd (divmod_nat m n)" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25162diff
changeset | 551 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 552 | lemma divmod_nat_div_mod: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 553 | "divmod_nat m n = (m div n, m mod n)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 554 | unfolding div_nat_def mod_nat_def by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 555 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 556 | lemma div_eq: | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 557 | assumes "divmod_nat_rel m n (q, r)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 558 | shows "m div n = q" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 559 | using assms by (auto dest: divmod_nat_eq simp add: divmod_nat_div_mod) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 560 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 561 | lemma mod_eq: | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 562 | assumes "divmod_nat_rel m n (q, r)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 563 | shows "m mod n = r" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 564 | using assms by (auto dest: divmod_nat_eq simp add: divmod_nat_div_mod) | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25162diff
changeset | 565 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 566 | lemma divmod_nat_rel: "divmod_nat_rel m n (m div n, m mod n)" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 567 | by (simp add: div_nat_def mod_nat_def divmod_nat_rel_divmod_nat) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 568 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 569 | lemma divmod_nat_zero: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 570 | "divmod_nat m 0 = (0, m)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 571 | proof - | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 572 | from divmod_nat_rel [of m 0] show ?thesis | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 573 | unfolding divmod_nat_div_mod divmod_nat_rel_def by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 574 | qed | 
| 25942 | 575 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 576 | lemma divmod_nat_base: | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 577 | assumes "m < n" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 578 | shows "divmod_nat m n = (0, m)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 579 | proof - | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 580 | from divmod_nat_rel [of m n] show ?thesis | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 581 | unfolding divmod_nat_div_mod divmod_nat_rel_def | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 582 | using assms by (cases "m div n = 0") | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 583 | (auto simp add: gr0_conv_Suc [of "m div n"]) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 584 | qed | 
| 25942 | 585 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 586 | lemma divmod_nat_step: | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 587 | assumes "0 < n" and "n \<le> m" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 588 | shows "divmod_nat m n = (Suc ((m - n) div n), (m - n) mod n)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 589 | proof - | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 590 | from divmod_nat_rel have divmod_nat_m_n: "divmod_nat_rel m n (m div n, m mod n)" . | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 591 | with assms have m_div_n: "m div n \<ge> 1" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 592 | by (cases "m div n") (auto simp add: divmod_nat_rel_def) | 
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 593 | have "divmod_nat_rel (m - n) n (m div n - Suc 0, m mod n)" | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 594 | proof - | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 595 | from assms have | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 596 | "n \<noteq> 0" | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 597 | "\<And>k. m = Suc k * n + m mod n ==> m - n = (Suc k - Suc 0) * n + m mod n" | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 598 | by simp_all | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 599 | then show ?thesis using assms divmod_nat_m_n | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 600 | by (cases "m div n") | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 601 | (simp_all only: divmod_nat_rel_def fst_conv snd_conv, simp_all) | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 602 | qed | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 603 | with divmod_nat_eq have "divmod_nat (m - n) n = (m div n - Suc 0, m mod n)" by simp | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 604 | moreover from divmod_nat_div_mod have "divmod_nat (m - n) n = ((m - n) div n, (m - n) mod n)" . | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 605 | ultimately have "m div n = Suc ((m - n) div n)" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 606 | and "m mod n = (m - n) mod n" using m_div_n by simp_all | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 607 | then show ?thesis using divmod_nat_div_mod by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 608 | qed | 
| 25942 | 609 | |
| 26300 | 610 | text {* The ''recursion'' equations for @{const div} and @{const mod} *}
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 611 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 612 | lemma div_less [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 613 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 614 | assumes "m < n" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 615 | shows "m div n = 0" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 616 | using assms divmod_nat_base divmod_nat_div_mod by simp | 
| 25942 | 617 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 618 | lemma le_div_geq: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 619 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 620 | assumes "0 < n" and "n \<le> m" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 621 | shows "m div n = Suc ((m - n) div n)" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 622 | using assms divmod_nat_step divmod_nat_div_mod by simp | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 623 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 624 | lemma mod_less [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 625 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 626 | assumes "m < n" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 627 | shows "m mod n = m" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 628 | using assms divmod_nat_base divmod_nat_div_mod by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 629 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 630 | lemma le_mod_geq: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 631 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 632 | assumes "n \<le> m" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 633 | shows "m mod n = (m - n) mod n" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 634 | using assms divmod_nat_step divmod_nat_div_mod by (cases "n = 0") simp_all | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 635 | |
| 30930 | 636 | instance proof - | 
| 637 | have [simp]: "\<And>n::nat. n div 0 = 0" | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 638 | by (simp add: div_nat_def divmod_nat_zero) | 
| 30930 | 639 | have [simp]: "\<And>n::nat. 0 div n = 0" | 
| 640 | proof - | |
| 641 | fix n :: nat | |
| 642 | show "0 div n = 0" | |
| 643 | by (cases "n = 0") simp_all | |
| 644 | qed | |
| 645 | show "OFCLASS(nat, semiring_div_class)" proof | |
| 646 | fix m n :: nat | |
| 647 | show "m div n * n + m mod n = m" | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 648 | using divmod_nat_rel [of m n] by (simp add: divmod_nat_rel_def) | 
| 30930 | 649 | next | 
| 650 | fix m n q :: nat | |
| 651 | assume "n \<noteq> 0" | |
| 652 | then show "(q + m * n) div n = m + q div n" | |
| 653 | by (induct m) (simp_all add: le_div_geq) | |
| 654 | next | |
| 655 | fix m n q :: nat | |
| 656 | assume "m \<noteq> 0" | |
| 657 | then show "(m * n) div (m * q) = n div q" | |
| 658 | proof (cases "n \<noteq> 0 \<and> q \<noteq> 0") | |
| 659 | case False then show ?thesis by auto | |
| 660 | next | |
| 661 | case True with `m \<noteq> 0` | |
| 662 | have "m > 0" and "n > 0" and "q > 0" by auto | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 663 | then have "\<And>a b. divmod_nat_rel n q (a, b) \<Longrightarrow> divmod_nat_rel (m * n) (m * q) (a, m * b)" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 664 | by (auto simp add: divmod_nat_rel_def) (simp_all add: algebra_simps) | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 665 | moreover from divmod_nat_rel have "divmod_nat_rel n q (n div q, n mod q)" . | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 666 | ultimately have "divmod_nat_rel (m * n) (m * q) (n div q, m * (n mod q))" . | 
| 30930 | 667 | then show ?thesis by (simp add: div_eq) | 
| 668 | qed | |
| 669 | qed simp_all | |
| 25942 | 670 | qed | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 671 | |
| 25942 | 672 | end | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 673 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 674 | lemma divmod_nat_if [code]: "divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 675 | let (q, r) = divmod_nat (m - n) n in (Suc q, r))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 676 | by (simp add: divmod_nat_zero divmod_nat_base divmod_nat_step) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 677 | (simp add: divmod_nat_div_mod) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 678 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 679 | text {* Simproc for cancelling @{const div} and @{const mod} *}
 | 
| 25942 | 680 | |
| 30934 | 681 | ML {*
 | 
| 682 | local | |
| 683 | ||
| 684 | structure CancelDivMod = CancelDivModFun(struct | |
| 25942 | 685 | |
| 30934 | 686 |   val div_name = @{const_name div};
 | 
| 687 |   val mod_name = @{const_name mod};
 | |
| 688 | val mk_binop = HOLogic.mk_binop; | |
| 689 | val mk_sum = Nat_Arith.mk_sum; | |
| 690 | val dest_sum = Nat_Arith.dest_sum; | |
| 25942 | 691 | |
| 30934 | 692 |   val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}];
 | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 693 | |
| 30934 | 694 | val trans = trans; | 
| 25942 | 695 | |
| 30934 | 696 | val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac | 
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34982diff
changeset | 697 |     (@{thm add_0_left} :: @{thm add_0_right} :: @{thms add_ac}))
 | 
| 25942 | 698 | |
| 30934 | 699 | end) | 
| 25942 | 700 | |
| 30934 | 701 | in | 
| 25942 | 702 | |
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
37767diff
changeset | 703 | val cancel_div_mod_nat_proc = Simplifier.simproc_global @{theory}
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 704 | "cancel_div_mod" ["(m::nat) + n"] (K CancelDivMod.proc); | 
| 25942 | 705 | |
| 30934 | 706 | val _ = Addsimprocs [cancel_div_mod_nat_proc]; | 
| 707 | ||
| 708 | end | |
| 25942 | 709 | *} | 
| 710 | ||
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 711 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 712 | subsubsection {* Quotient *}
 | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 713 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 714 | lemma div_geq: "0 < n \<Longrightarrow> \<not> m < n \<Longrightarrow> m div n = Suc ((m - n) div n)" | 
| 29667 | 715 | by (simp add: le_div_geq linorder_not_less) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 716 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 717 | lemma div_if: "0 < n \<Longrightarrow> m div n = (if m < n then 0 else Suc ((m - n) div n))" | 
| 29667 | 718 | by (simp add: div_geq) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 719 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 720 | lemma div_mult_self_is_m [simp]: "0<n ==> (m*n) div n = (m::nat)" | 
| 29667 | 721 | by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 722 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 723 | lemma div_mult_self1_is_m [simp]: "0<n ==> (n*m) div n = (m::nat)" | 
| 29667 | 724 | by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 725 | |
| 25942 | 726 | |
| 727 | subsubsection {* Remainder *}
 | |
| 728 | ||
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 729 | lemma mod_less_divisor [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 730 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 731 | assumes "n > 0" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 732 | shows "m mod n < (n::nat)" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 733 | using assms divmod_nat_rel [of m n] unfolding divmod_nat_rel_def by auto | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 734 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 735 | lemma mod_less_eq_dividend [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 736 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 737 | shows "m mod n \<le> m" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 738 | proof (rule add_leD2) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 739 | from mod_div_equality have "m div n * n + m mod n = m" . | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 740 | then show "m div n * n + m mod n \<le> m" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 741 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 742 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 743 | lemma mod_geq: "\<not> m < (n\<Colon>nat) \<Longrightarrow> m mod n = (m - n) mod n" | 
| 29667 | 744 | by (simp add: le_mod_geq linorder_not_less) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 745 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 746 | lemma mod_if: "m mod (n\<Colon>nat) = (if m < n then m else (m - n) mod n)" | 
| 29667 | 747 | by (simp add: le_mod_geq) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 748 | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 749 | lemma mod_1 [simp]: "m mod Suc 0 = 0" | 
| 29667 | 750 | by (induct m) (simp_all add: mod_geq) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 751 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 752 | lemma mod_mult_distrib: "(m mod n) * (k\<Colon>nat) = (m * k) mod (n * k)" | 
| 22718 | 753 | apply (cases "n = 0", simp) | 
| 754 | apply (cases "k = 0", simp) | |
| 755 | apply (induct m rule: nat_less_induct) | |
| 756 | apply (subst mod_if, simp) | |
| 757 | apply (simp add: mod_geq diff_mult_distrib) | |
| 758 | done | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 759 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 760 | lemma mod_mult_distrib2: "(k::nat) * (m mod n) = (k*m) mod (k*n)" | 
| 29667 | 761 | by (simp add: mult_commute [of k] mod_mult_distrib) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 762 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 763 | (* a simple rearrangement of mod_div_equality: *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 764 | lemma mult_div_cancel: "(n::nat) * (m div n) = m - (m mod n)" | 
| 29667 | 765 | by (cut_tac a = m and b = n in mod_div_equality2, arith) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 766 | |
| 15439 | 767 | lemma mod_le_divisor[simp]: "0 < n \<Longrightarrow> m mod n \<le> (n::nat)" | 
| 22718 | 768 | apply (drule mod_less_divisor [where m = m]) | 
| 769 | apply simp | |
| 770 | done | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 771 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 772 | subsubsection {* Quotient and Remainder *}
 | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 773 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 774 | lemma divmod_nat_rel_mult1_eq: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 775 | "divmod_nat_rel b c (q, r) \<Longrightarrow> c > 0 | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 776 | \<Longrightarrow> divmod_nat_rel (a * b) c (a * q + a * r div c, a * r mod c)" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 777 | by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 778 | |
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 779 | lemma div_mult1_eq: | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 780 | "(a * b) div c = a * (b div c) + a * (b mod c) div (c::nat)" | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 781 | apply (cases "c = 0", simp) | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 782 | apply (blast intro: divmod_nat_rel [THEN divmod_nat_rel_mult1_eq, THEN div_eq]) | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 783 | done | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 784 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 785 | lemma divmod_nat_rel_add1_eq: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 786 | "divmod_nat_rel a c (aq, ar) \<Longrightarrow> divmod_nat_rel b c (bq, br) \<Longrightarrow> c > 0 | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 787 | \<Longrightarrow> divmod_nat_rel (a + b) c (aq + bq + (ar + br) div c, (ar + br) mod c)" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 788 | by (auto simp add: split_ifs divmod_nat_rel_def algebra_simps) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 789 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 790 | (*NOT suitable for rewriting: the RHS has an instance of the LHS*) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 791 | lemma div_add1_eq: | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 792 | "(a+b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 793 | apply (cases "c = 0", simp) | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 794 | apply (blast intro: divmod_nat_rel_add1_eq [THEN div_eq] divmod_nat_rel) | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 795 | done | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 796 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 797 | lemma mod_lemma: "[| (0::nat) < c; r < b |] ==> b * (q mod c) + r < b * c" | 
| 22718 | 798 | apply (cut_tac m = q and n = c in mod_less_divisor) | 
| 799 | apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto) | |
| 800 | apply (erule_tac P = "%x. ?lhs < ?rhs x" in ssubst) | |
| 801 | apply (simp add: add_mult_distrib2) | |
| 802 | done | |
| 10559 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 803 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 804 | lemma divmod_nat_rel_mult2_eq: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 805 | "divmod_nat_rel a b (q, r) \<Longrightarrow> 0 < b \<Longrightarrow> 0 < c | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 806 | \<Longrightarrow> divmod_nat_rel a (b * c) (q div c, b *(q mod c) + r)" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 807 | by (auto simp add: mult_ac divmod_nat_rel_def add_mult_distrib2 [symmetric] mod_lemma) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 808 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 809 | lemma div_mult2_eq: "a div (b*c) = (a div b) div (c::nat)" | 
| 22718 | 810 | apply (cases "b = 0", simp) | 
| 811 | apply (cases "c = 0", simp) | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 812 | apply (force simp add: divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN div_eq]) | 
| 22718 | 813 | done | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 814 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 815 | lemma mod_mult2_eq: "a mod (b*c) = b*(a div b mod c) + a mod (b::nat)" | 
| 22718 | 816 | apply (cases "b = 0", simp) | 
| 817 | apply (cases "c = 0", simp) | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 818 | apply (auto simp add: mult_commute divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN mod_eq]) | 
| 22718 | 819 | done | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 820 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 821 | |
| 25942 | 822 | subsubsection{*Further Facts about Quotient and Remainder*}
 | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 823 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 824 | lemma div_1 [simp]: "m div Suc 0 = m" | 
| 29667 | 825 | by (induct m) (simp_all add: div_geq) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 826 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 827 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 828 | (* Monotonicity of div in first argument *) | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 829 | lemma div_le_mono [rule_format (no_asm)]: | 
| 22718 | 830 | "\<forall>m::nat. m \<le> n --> (m div k) \<le> (n div k)" | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 831 | apply (case_tac "k=0", simp) | 
| 15251 | 832 | apply (induct "n" rule: nat_less_induct, clarify) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 833 | apply (case_tac "n<k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 834 | (* 1 case n<k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 835 | apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 836 | (* 2 case n >= k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 837 | apply (case_tac "m<k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 838 | (* 2.1 case m<k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 839 | apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 840 | (* 2.2 case m>=k *) | 
| 15439 | 841 | apply (simp add: div_geq diff_le_mono) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 842 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 843 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 844 | (* Antimonotonicity of div in second argument *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 845 | lemma div_le_mono2: "!!m::nat. [| 0<m; m\<le>n |] ==> (k div n) \<le> (k div m)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 846 | apply (subgoal_tac "0<n") | 
| 22718 | 847 | prefer 2 apply simp | 
| 15251 | 848 | apply (induct_tac k rule: nat_less_induct) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 849 | apply (rename_tac "k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 850 | apply (case_tac "k<n", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 851 | apply (subgoal_tac "~ (k<m) ") | 
| 22718 | 852 | prefer 2 apply simp | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 853 | apply (simp add: div_geq) | 
| 15251 | 854 | apply (subgoal_tac "(k-n) div n \<le> (k-m) div n") | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 855 | prefer 2 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 856 | apply (blast intro: div_le_mono diff_le_mono2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 857 | apply (rule le_trans, simp) | 
| 15439 | 858 | apply (simp) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 859 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 860 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 861 | lemma div_le_dividend [simp]: "m div n \<le> (m::nat)" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 862 | apply (case_tac "n=0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 863 | apply (subgoal_tac "m div n \<le> m div 1", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 864 | apply (rule div_le_mono2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 865 | apply (simp_all (no_asm_simp)) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 866 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 867 | |
| 22718 | 868 | (* Similar for "less than" *) | 
| 17085 | 869 | lemma div_less_dividend [rule_format]: | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 870 | "!!n::nat. 1<n ==> 0 < m --> m div n < m" | 
| 15251 | 871 | apply (induct_tac m rule: nat_less_induct) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 872 | apply (rename_tac "m") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 873 | apply (case_tac "m<n", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 874 | apply (subgoal_tac "0<n") | 
| 22718 | 875 | prefer 2 apply simp | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 876 | apply (simp add: div_geq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 877 | apply (case_tac "n<m") | 
| 15251 | 878 | apply (subgoal_tac "(m-n) div n < (m-n) ") | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 879 | apply (rule impI less_trans_Suc)+ | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 880 | apply assumption | 
| 15439 | 881 | apply (simp_all) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 882 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 883 | |
| 17085 | 884 | declare div_less_dividend [simp] | 
| 885 | ||
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 886 | text{*A fact for the mutilated chess board*}
 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 887 | lemma mod_Suc: "Suc(m) mod n = (if Suc(m mod n) = n then 0 else Suc(m mod n))" | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 888 | apply (case_tac "n=0", simp) | 
| 15251 | 889 | apply (induct "m" rule: nat_less_induct) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 890 | apply (case_tac "Suc (na) <n") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 891 | (* case Suc(na) < n *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 892 | apply (frule lessI [THEN less_trans], simp add: less_not_refl3) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 893 | (* case n \<le> Suc(na) *) | 
| 16796 | 894 | apply (simp add: linorder_not_less le_Suc_eq mod_geq) | 
| 15439 | 895 | apply (auto simp add: Suc_diff_le le_mod_geq) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 896 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 897 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 898 | lemma mod_eq_0_iff: "(m mod d = 0) = (\<exists>q::nat. m = d*q)" | 
| 29667 | 899 | by (auto simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def) | 
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
16796diff
changeset | 900 | |
| 22718 | 901 | lemmas mod_eq_0D [dest!] = mod_eq_0_iff [THEN iffD1] | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 902 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 903 | (*Loses information, namely we also have r<d provided d is nonzero*) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 904 | lemma mod_eqD: "(m mod d = r) ==> \<exists>q::nat. m = r + q*d" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 905 | apply (cut_tac a = m in mod_div_equality) | 
| 22718 | 906 | apply (simp only: add_ac) | 
| 907 | apply (blast intro: sym) | |
| 908 | done | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 909 | |
| 13152 | 910 | lemma split_div: | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 911 | "P(n div k :: nat) = | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 912 | ((k = 0 \<longrightarrow> P 0) \<and> (k \<noteq> 0 \<longrightarrow> (!i. !j<k. n = k*i + j \<longrightarrow> P i)))" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 913 | (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))") | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 914 | proof | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 915 | assume P: ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 916 | show ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 917 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 918 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 919 | with P show ?Q by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 920 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 921 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 922 | thus ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 923 | proof (simp, intro allI impI) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 924 | fix i j | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 925 | assume n: "n = k*i + j" and j: "j < k" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 926 | show "P i" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 927 | proof (cases) | 
| 22718 | 928 | assume "i = 0" | 
| 929 | with n j P show "P i" by simp | |
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 930 | next | 
| 22718 | 931 | assume "i \<noteq> 0" | 
| 932 | with not0 n j P show "P i" by(simp add:add_ac) | |
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 933 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 934 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 935 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 936 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 937 | assume Q: ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 938 | show ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 939 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 940 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 941 | with Q show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 942 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 943 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 944 | with Q have R: ?R by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 945 | from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] | 
| 13517 | 946 | show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 947 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 948 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 949 | |
| 13882 | 950 | lemma split_div_lemma: | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 951 | assumes "0 < n" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 952 | shows "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> q = ((m\<Colon>nat) div n)" (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 953 | proof | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 954 | assume ?rhs | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 955 | with mult_div_cancel have nq: "n * q = m - (m mod n)" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 956 | then have A: "n * q \<le> m" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 957 | have "n - (m mod n) > 0" using mod_less_divisor assms by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 958 | then have "m < m + (n - (m mod n))" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 959 | then have "m < n + (m - (m mod n))" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 960 | with nq have "m < n + n * q" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 961 | then have B: "m < n * Suc q" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 962 | from A B show ?lhs .. | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 963 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 964 | assume P: ?lhs | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 965 | then have "divmod_nat_rel m n (q, m - n * q)" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 966 | unfolding divmod_nat_rel_def by (auto simp add: mult_ac) | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 967 | with divmod_nat_rel_unique divmod_nat_rel [of m n] | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 968 | have "(q, m - n * q) = (m div n, m mod n)" by auto | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 969 | then show ?rhs by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 970 | qed | 
| 13882 | 971 | |
| 972 | theorem split_div': | |
| 973 | "P ((m::nat) div n) = ((n = 0 \<and> P 0) \<or> | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 974 | (\<exists>q. (n * q \<le> m \<and> m < n * (Suc q)) \<and> P q))" | 
| 13882 | 975 | apply (case_tac "0 < n") | 
| 976 | apply (simp only: add: split_div_lemma) | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 977 | apply simp_all | 
| 13882 | 978 | done | 
| 979 | ||
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 980 | lemma split_mod: | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 981 | "P(n mod k :: nat) = | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 982 | ((k = 0 \<longrightarrow> P n) \<and> (k \<noteq> 0 \<longrightarrow> (!i. !j<k. n = k*i + j \<longrightarrow> P j)))" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 983 | (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))") | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 984 | proof | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 985 | assume P: ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 986 | show ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 987 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 988 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 989 | with P show ?Q by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 990 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 991 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 992 | thus ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 993 | proof (simp, intro allI impI) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 994 | fix i j | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 995 | assume "n = k*i + j" "j < k" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 996 | thus "P j" using not0 P by(simp add:add_ac mult_ac) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 997 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 998 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 999 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1000 | assume Q: ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1001 | show ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1002 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1003 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 1004 | with Q show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1005 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1006 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1007 | with Q have R: ?R by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1008 | from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] | 
| 13517 | 1009 | show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1010 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1011 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1012 | |
| 13882 | 1013 | theorem mod_div_equality': "(m::nat) mod n = m - (m div n) * n" | 
| 1014 | apply (rule_tac P="%x. m mod n = x - (m div n) * n" in | |
| 1015 | subst [OF mod_div_equality [of _ n]]) | |
| 1016 | apply arith | |
| 1017 | done | |
| 1018 | ||
| 22800 | 1019 | lemma div_mod_equality': | 
| 1020 | fixes m n :: nat | |
| 1021 | shows "m div n * n = m - m mod n" | |
| 1022 | proof - | |
| 1023 | have "m mod n \<le> m mod n" .. | |
| 1024 | from div_mod_equality have | |
| 1025 | "m div n * n + m mod n - m mod n = m - m mod n" by simp | |
| 1026 | with diff_add_assoc [OF `m mod n \<le> m mod n`, of "m div n * n"] have | |
| 1027 | "m div n * n + (m mod n - m mod n) = m - m mod n" | |
| 1028 | by simp | |
| 1029 | then show ?thesis by simp | |
| 1030 | qed | |
| 1031 | ||
| 1032 | ||
| 25942 | 1033 | subsubsection {*An ``induction'' law for modulus arithmetic.*}
 | 
| 14640 | 1034 | |
| 1035 | lemma mod_induct_0: | |
| 1036 | assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)" | |
| 1037 | and base: "P i" and i: "i<p" | |
| 1038 | shows "P 0" | |
| 1039 | proof (rule ccontr) | |
| 1040 | assume contra: "\<not>(P 0)" | |
| 1041 | from i have p: "0<p" by simp | |
| 1042 | have "\<forall>k. 0<k \<longrightarrow> \<not> P (p-k)" (is "\<forall>k. ?A k") | |
| 1043 | proof | |
| 1044 | fix k | |
| 1045 | show "?A k" | |
| 1046 | proof (induct k) | |
| 1047 | show "?A 0" by simp -- "by contradiction" | |
| 1048 | next | |
| 1049 | fix n | |
| 1050 | assume ih: "?A n" | |
| 1051 | show "?A (Suc n)" | |
| 1052 | proof (clarsimp) | |
| 22718 | 1053 | assume y: "P (p - Suc n)" | 
| 1054 | have n: "Suc n < p" | |
| 1055 | proof (rule ccontr) | |
| 1056 | assume "\<not>(Suc n < p)" | |
| 1057 | hence "p - Suc n = 0" | |
| 1058 | by simp | |
| 1059 | with y contra show "False" | |
| 1060 | by simp | |
| 1061 | qed | |
| 1062 | hence n2: "Suc (p - Suc n) = p-n" by arith | |
| 1063 | from p have "p - Suc n < p" by arith | |
| 1064 | with y step have z: "P ((Suc (p - Suc n)) mod p)" | |
| 1065 | by blast | |
| 1066 | show "False" | |
| 1067 | proof (cases "n=0") | |
| 1068 | case True | |
| 1069 | with z n2 contra show ?thesis by simp | |
| 1070 | next | |
| 1071 | case False | |
| 1072 | with p have "p-n < p" by arith | |
| 1073 | with z n2 False ih show ?thesis by simp | |
| 1074 | qed | |
| 14640 | 1075 | qed | 
| 1076 | qed | |
| 1077 | qed | |
| 1078 | moreover | |
| 1079 | from i obtain k where "0<k \<and> i+k=p" | |
| 1080 | by (blast dest: less_imp_add_positive) | |
| 1081 | hence "0<k \<and> i=p-k" by auto | |
| 1082 | moreover | |
| 1083 | note base | |
| 1084 | ultimately | |
| 1085 | show "False" by blast | |
| 1086 | qed | |
| 1087 | ||
| 1088 | lemma mod_induct: | |
| 1089 | assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)" | |
| 1090 | and base: "P i" and i: "i<p" and j: "j<p" | |
| 1091 | shows "P j" | |
| 1092 | proof - | |
| 1093 | have "\<forall>j<p. P j" | |
| 1094 | proof | |
| 1095 | fix j | |
| 1096 | show "j<p \<longrightarrow> P j" (is "?A j") | |
| 1097 | proof (induct j) | |
| 1098 | from step base i show "?A 0" | |
| 22718 | 1099 | by (auto elim: mod_induct_0) | 
| 14640 | 1100 | next | 
| 1101 | fix k | |
| 1102 | assume ih: "?A k" | |
| 1103 | show "?A (Suc k)" | |
| 1104 | proof | |
| 22718 | 1105 | assume suc: "Suc k < p" | 
| 1106 | hence k: "k<p" by simp | |
| 1107 | with ih have "P k" .. | |
| 1108 | with step k have "P (Suc k mod p)" | |
| 1109 | by blast | |
| 1110 | moreover | |
| 1111 | from suc have "Suc k mod p = Suc k" | |
| 1112 | by simp | |
| 1113 | ultimately | |
| 1114 | show "P (Suc k)" by simp | |
| 14640 | 1115 | qed | 
| 1116 | qed | |
| 1117 | qed | |
| 1118 | with j show ?thesis by blast | |
| 1119 | qed | |
| 1120 | ||
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1121 | lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1122 | by (auto simp add: numeral_2_eq_2 le_div_geq) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1123 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1124 | lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1125 | by (simp add: nat_mult_2 [symmetric]) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1126 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1127 | lemma mod2_Suc_Suc [simp]: "Suc(Suc(m)) mod 2 = m mod 2" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1128 | apply (subgoal_tac "m mod 2 < 2") | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1129 | apply (erule less_2_cases [THEN disjE]) | 
| 35216 | 1130 | apply (simp_all (no_asm_simp) add: Let_def mod_Suc) | 
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1131 | done | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1132 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1133 | lemma mod2_gr_0 [simp]: "0 < (m\<Colon>nat) mod 2 \<longleftrightarrow> m mod 2 = 1" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1134 | proof - | 
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1135 |   { fix n :: nat have  "(n::nat) < 2 \<Longrightarrow> n = 0 \<or> n = 1" by (cases n) simp_all }
 | 
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1136 | moreover have "m mod 2 < 2" by simp | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1137 | ultimately have "m mod 2 = 0 \<or> m mod 2 = 1" . | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1138 | then show ?thesis by auto | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1139 | qed | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1140 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1141 | text{*These lemmas collapse some needless occurrences of Suc:
 | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1142 | at least three Sucs, since two and fewer are rewritten back to Suc again! | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1143 | We already have some rules to simplify operands smaller than 3.*} | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1144 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1145 | lemma div_Suc_eq_div_add3 [simp]: "m div (Suc (Suc (Suc n))) = m div (3+n)" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1146 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1147 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1148 | lemma mod_Suc_eq_mod_add3 [simp]: "m mod (Suc (Suc (Suc n))) = m mod (3+n)" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1149 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1150 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1151 | lemma Suc_div_eq_add3_div: "(Suc (Suc (Suc m))) div n = (3+m) div n" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1152 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1153 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1154 | lemma Suc_mod_eq_add3_mod: "(Suc (Suc (Suc m))) mod n = (3+m) mod n" | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1155 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1156 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1157 | lemmas Suc_div_eq_add3_div_number_of = | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1158 | Suc_div_eq_add3_div [of _ "number_of v", standard] | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1159 | declare Suc_div_eq_add3_div_number_of [simp] | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1160 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1161 | lemmas Suc_mod_eq_add3_mod_number_of = | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1162 | Suc_mod_eq_add3_mod [of _ "number_of v", standard] | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1163 | declare Suc_mod_eq_add3_mod_number_of [simp] | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1164 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1165 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1166 | lemma Suc_times_mod_eq: "1<k ==> Suc (k * m) mod k = 1" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1167 | apply (induct "m") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1168 | apply (simp_all add: mod_Suc) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1169 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1170 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1171 | declare Suc_times_mod_eq [of "number_of w", standard, simp] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1172 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1173 | lemma [simp]: "n div k \<le> (Suc n) div k" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1174 | by (simp add: div_le_mono) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1175 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1176 | lemma Suc_n_div_2_gt_zero [simp]: "(0::nat) < n ==> 0 < (n + 1) div 2" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1177 | by (cases n) simp_all | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1178 | |
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1179 | lemma div_2_gt_zero [simp]: assumes A: "(1::nat) < n" shows "0 < n div 2" | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1180 | proof - | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1181 | from A have B: "0 < n - 1" and C: "n - 1 + 1 = n" by simp_all | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1182 | from Suc_n_div_2_gt_zero [OF B] C show ?thesis by simp | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1183 | qed | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1184 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1185 | (* Potential use of algebra : Equality modulo n*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1186 | lemma mod_mult_self3 [simp]: "(k*n + m) mod n = m mod (n::nat)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1187 | by (simp add: mult_ac add_ac) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1188 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1189 | lemma mod_mult_self4 [simp]: "Suc (k*n + m) mod n = Suc m mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1190 | proof - | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1191 | have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1192 | also have "... = Suc m mod n" by (rule mod_mult_self3) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1193 | finally show ?thesis . | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1194 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1195 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1196 | lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1197 | apply (subst mod_Suc [of m]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1198 | apply (subst mod_Suc [of "m mod n"], simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1199 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1200 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1201 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1202 | subsection {* Division on @{typ int} *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1203 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1204 | definition divmod_int_rel :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1205 |     --{*definition of quotient and remainder*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1206 | [code]: "divmod_int_rel a b = (\<lambda>(q, r). a = b * q + r \<and> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1207 | (if 0 < b then 0 \<le> r \<and> r < b else b < r \<and> r \<le> 0))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1208 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1209 | definition adjust :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1210 |     --{*for the division algorithm*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1211 | [code]: "adjust b = (\<lambda>(q, r). if 0 \<le> r - b then (2 * q + 1, r - b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1212 | else (2 * q, r))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1213 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1214 | text{*algorithm for the case @{text "a\<ge>0, b>0"}*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1215 | function posDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1216 | "posDivAlg a b = (if a < b \<or> b \<le> 0 then (0, a) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1217 | else adjust b (posDivAlg a (2 * b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1218 | by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1219 | termination by (relation "measure (\<lambda>(a, b). nat (a - b + 1))") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1220 | (auto simp add: mult_2) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1221 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1222 | text{*algorithm for the case @{text "a<0, b>0"}*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1223 | function negDivAlg :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1224 | "negDivAlg a b = (if 0 \<le>a + b \<or> b \<le> 0 then (-1, a + b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1225 | else adjust b (negDivAlg a (2 * b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1226 | by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1227 | termination by (relation "measure (\<lambda>(a, b). nat (- a - b))") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1228 | (auto simp add: mult_2) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1229 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1230 | text{*algorithm for the general case @{term "b\<noteq>0"}*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1231 | definition negateSnd :: "int \<times> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1232 | [code_unfold]: "negateSnd = apsnd uminus" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1233 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1234 | definition divmod_int :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1235 |     --{*The full division algorithm considers all possible signs for a, b
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1236 |        including the special case @{text "a=0, b<0"} because 
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1237 |        @{term negDivAlg} requires @{term "a<0"}.*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1238 | "divmod_int a b = (if 0 \<le> a then if 0 \<le> b then posDivAlg a b | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1239 | else if a = 0 then (0, 0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1240 | else negateSnd (negDivAlg (-a) (-b)) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1241 | else | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1242 | if 0 < b then negDivAlg a b | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1243 | else negateSnd (posDivAlg (-a) (-b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1244 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1245 | instantiation int :: Divides.div | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1246 | begin | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1247 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1248 | definition | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1249 | "a div b = fst (divmod_int a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1250 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1251 | definition | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1252 | "a mod b = snd (divmod_int a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1253 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1254 | instance .. | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1255 | |
| 3366 | 1256 | end | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1257 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1258 | lemma divmod_int_mod_div: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1259 | "divmod_int p q = (p div q, p mod q)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1260 | by (auto simp add: div_int_def mod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1261 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1262 | text{*
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1263 | Here is the division algorithm in ML: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1264 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1265 | \begin{verbatim}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1266 | fun posDivAlg (a,b) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1267 | if a<b then (0,a) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1268 | else let val (q,r) = posDivAlg(a, 2*b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1269 | in if 0\<le>r-b then (2*q+1, r-b) else (2*q, r) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1270 | end | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1271 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1272 | fun negDivAlg (a,b) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1273 | if 0\<le>a+b then (~1,a+b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1274 | else let val (q,r) = negDivAlg(a, 2*b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1275 | in if 0\<le>r-b then (2*q+1, r-b) else (2*q, r) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1276 | end; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1277 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1278 | fun negateSnd (q,r:int) = (q,~r); | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1279 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1280 | fun divmod (a,b) = if 0\<le>a then | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1281 | if b>0 then posDivAlg (a,b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1282 | else if a=0 then (0,0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1283 | else negateSnd (negDivAlg (~a,~b)) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1284 | else | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1285 | if 0<b then negDivAlg (a,b) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1286 | else negateSnd (posDivAlg (~a,~b)); | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1287 | \end{verbatim}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1288 | *} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1289 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1290 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1291 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1292 | subsubsection{*Uniqueness and Monotonicity of Quotients and Remainders*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1293 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1294 | lemma unique_quotient_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1295 | "[| b*q' + r' \<le> b*q + r; 0 \<le> r'; r' < b; r < b |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1296 | ==> q' \<le> (q::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1297 | apply (subgoal_tac "r' + b * (q'-q) \<le> r") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1298 | prefer 2 apply (simp add: right_diff_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1299 | apply (subgoal_tac "0 < b * (1 + q - q') ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1300 | apply (erule_tac [2] order_le_less_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1301 | prefer 2 apply (simp add: right_diff_distrib right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1302 | apply (subgoal_tac "b * q' < b * (1 + q) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1303 | prefer 2 apply (simp add: right_diff_distrib right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1304 | apply (simp add: mult_less_cancel_left) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1305 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1306 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1307 | lemma unique_quotient_lemma_neg: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1308 | "[| b*q' + r' \<le> b*q + r; r \<le> 0; b < r; b < r' |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1309 | ==> q \<le> (q'::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1310 | by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma, | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1311 | auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1312 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1313 | lemma unique_quotient: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1314 | "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r'); b \<noteq> 0 |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1315 | ==> q = q'" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1316 | apply (simp add: divmod_int_rel_def linorder_neq_iff split: split_if_asm) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1317 | apply (blast intro: order_antisym | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1318 | dest: order_eq_refl [THEN unique_quotient_lemma] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1319 | order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1320 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1321 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1322 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1323 | lemma unique_remainder: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1324 | "[| divmod_int_rel a b (q, r); divmod_int_rel a b (q', r'); b \<noteq> 0 |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1325 | ==> r = r'" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1326 | apply (subgoal_tac "q = q'") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1327 | apply (simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1328 | apply (blast intro: unique_quotient) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1329 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1330 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1331 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1332 | subsubsection{*Correctness of @{term posDivAlg}, the Algorithm for Non-Negative Dividends*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1333 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1334 | text{*And positive divisors*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1335 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1336 | lemma adjust_eq [simp]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1337 | "adjust b (q,r) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1338 | (let diff = r-b in | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1339 | if 0 \<le> diff then (2*q + 1, diff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1340 | else (2*q, r))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1341 | by (simp add: Let_def adjust_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1342 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1343 | declare posDivAlg.simps [simp del] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1344 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1345 | text{*use with a simproc to avoid repeatedly proving the premise*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1346 | lemma posDivAlg_eqn: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1347 | "0 < b ==> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1348 | posDivAlg a b = (if a<b then (0,a) else adjust b (posDivAlg a (2*b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1349 | by (rule posDivAlg.simps [THEN trans], simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1350 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1351 | text{*Correctness of @{term posDivAlg}: it computes quotients correctly*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1352 | theorem posDivAlg_correct: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1353 | assumes "0 \<le> a" and "0 < b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1354 | shows "divmod_int_rel a b (posDivAlg a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1355 | using prems apply (induct a b rule: posDivAlg.induct) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1356 | apply auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1357 | apply (simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1358 | apply (subst posDivAlg_eqn, simp add: right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1359 | apply (case_tac "a < b") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1360 | apply simp_all | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1361 | apply (erule splitE) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1362 | apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1363 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1364 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1365 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1366 | subsubsection{*Correctness of @{term negDivAlg}, the Algorithm for Negative Dividends*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1367 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1368 | text{*And positive divisors*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1369 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1370 | declare negDivAlg.simps [simp del] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1371 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1372 | text{*use with a simproc to avoid repeatedly proving the premise*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1373 | lemma negDivAlg_eqn: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1374 | "0 < b ==> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1375 | negDivAlg a b = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1376 | (if 0\<le>a+b then (-1,a+b) else adjust b (negDivAlg a (2*b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1377 | by (rule negDivAlg.simps [THEN trans], simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1378 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1379 | (*Correctness of negDivAlg: it computes quotients correctly | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1380 | It doesn't work if a=0 because the 0/b equals 0, not -1*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1381 | lemma negDivAlg_correct: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1382 | assumes "a < 0" and "b > 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1383 | shows "divmod_int_rel a b (negDivAlg a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1384 | using prems apply (induct a b rule: negDivAlg.induct) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1385 | apply (auto simp add: linorder_not_le) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1386 | apply (simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1387 | apply (subst negDivAlg_eqn, assumption) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1388 | apply (case_tac "a + b < (0\<Colon>int)") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1389 | apply simp_all | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1390 | apply (erule splitE) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1391 | apply (auto simp add: right_distrib Let_def mult_ac mult_2_right) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1392 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1393 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1394 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1395 | subsubsection{*Existence Shown by Proving the Division Algorithm to be Correct*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1396 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1397 | (*the case a=0*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1398 | lemma divmod_int_rel_0: "b \<noteq> 0 ==> divmod_int_rel 0 b (0, 0)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1399 | by (auto simp add: divmod_int_rel_def linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1400 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1401 | lemma posDivAlg_0 [simp]: "posDivAlg 0 b = (0, 0)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1402 | by (subst posDivAlg.simps, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1403 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1404 | lemma negDivAlg_minus1 [simp]: "negDivAlg -1 b = (-1, b - 1)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1405 | by (subst negDivAlg.simps, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1406 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1407 | lemma negateSnd_eq [simp]: "negateSnd(q,r) = (q,-r)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1408 | by (simp add: negateSnd_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1409 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1410 | lemma divmod_int_rel_neg: "divmod_int_rel (-a) (-b) qr ==> divmod_int_rel a b (negateSnd qr)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1411 | by (auto simp add: split_ifs divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1412 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1413 | lemma divmod_int_correct: "b \<noteq> 0 ==> divmod_int_rel a b (divmod_int a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1414 | by (force simp add: linorder_neq_iff divmod_int_rel_0 divmod_int_def divmod_int_rel_neg | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1415 | posDivAlg_correct negDivAlg_correct) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1416 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1417 | text{*Arbitrary definitions for division by zero.  Useful to simplify 
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1418 | certain equations.*} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1419 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1420 | lemma DIVISION_BY_ZERO [simp]: "a div (0::int) = 0 & a mod (0::int) = a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1421 | by (simp add: div_int_def mod_int_def divmod_int_def posDivAlg.simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1422 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1423 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1424 | text{*Basic laws about division and remainder*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1425 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1426 | lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1427 | apply (case_tac "b = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1428 | apply (cut_tac a = a and b = b in divmod_int_correct) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1429 | apply (auto simp add: divmod_int_rel_def div_int_def mod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1430 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1431 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1432 | lemma zdiv_zmod_equality: "(b * (a div b) + (a mod b)) + k = (a::int)+k" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1433 | by(simp add: zmod_zdiv_equality[symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1434 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1435 | lemma zdiv_zmod_equality2: "((a div b) * b + (a mod b)) + k = (a::int)+k" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1436 | by(simp add: mult_commute zmod_zdiv_equality[symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1437 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1438 | text {* Tool setup *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1439 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1440 | ML {*
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1441 | local | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1442 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1443 | structure CancelDivMod = CancelDivModFun(struct | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1444 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1445 |   val div_name = @{const_name div};
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1446 |   val mod_name = @{const_name mod};
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1447 | val mk_binop = HOLogic.mk_binop; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1448 | val mk_sum = Arith_Data.mk_sum HOLogic.intT; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1449 | val dest_sum = Arith_Data.dest_sum; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1450 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1451 |   val div_mod_eqs = map mk_meta_eq [@{thm zdiv_zmod_equality}, @{thm zdiv_zmod_equality2}];
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1452 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1453 | val trans = trans; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1454 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1455 | val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1456 |     (@{thm diff_minus} :: @{thms add_0s} @ @{thms add_ac}))
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1457 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1458 | end) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1459 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1460 | in | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1461 | |
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
37767diff
changeset | 1462 | val cancel_div_mod_int_proc = Simplifier.simproc_global @{theory}
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1463 | "cancel_zdiv_zmod" ["(k::int) + l"] (K CancelDivMod.proc); | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1464 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1465 | val _ = Addsimprocs [cancel_div_mod_int_proc]; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1466 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1467 | end | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1468 | *} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1469 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1470 | lemma pos_mod_conj : "(0::int) < b ==> 0 \<le> a mod b & a mod b < b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1471 | apply (cut_tac a = a and b = b in divmod_int_correct) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1472 | apply (auto simp add: divmod_int_rel_def mod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1473 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1474 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1475 | lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1, standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1476 | and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2, standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1477 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1478 | lemma neg_mod_conj : "b < (0::int) ==> a mod b \<le> 0 & b < a mod b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1479 | apply (cut_tac a = a and b = b in divmod_int_correct) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1480 | apply (auto simp add: divmod_int_rel_def div_int_def mod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1481 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1482 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1483 | lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1, standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1484 | and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2, standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1485 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1486 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1487 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1488 | subsubsection{*General Properties of div and mod*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1489 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1490 | lemma divmod_int_rel_div_mod: "b \<noteq> 0 ==> divmod_int_rel a b (a div b, a mod b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1491 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1492 | apply (force simp add: divmod_int_rel_def linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1493 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1494 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1495 | lemma divmod_int_rel_div: "[| divmod_int_rel a b (q, r); b \<noteq> 0 |] ==> a div b = q" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1496 | by (simp add: divmod_int_rel_div_mod [THEN unique_quotient]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1497 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1498 | lemma divmod_int_rel_mod: "[| divmod_int_rel a b (q, r); b \<noteq> 0 |] ==> a mod b = r" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1499 | by (simp add: divmod_int_rel_div_mod [THEN unique_remainder]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1500 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1501 | lemma div_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a div b = 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1502 | apply (rule divmod_int_rel_div) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1503 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1504 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1505 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1506 | lemma div_neg_neg_trivial: "[| a \<le> (0::int); b < a |] ==> a div b = 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1507 | apply (rule divmod_int_rel_div) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1508 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1509 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1510 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1511 | lemma div_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a div b = -1" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1512 | apply (rule divmod_int_rel_div) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1513 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1514 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1515 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1516 | (*There is no div_neg_pos_trivial because 0 div b = 0 would supersede it*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1517 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1518 | lemma mod_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a mod b = a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1519 | apply (rule_tac q = 0 in divmod_int_rel_mod) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1520 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1521 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1522 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1523 | lemma mod_neg_neg_trivial: "[| a \<le> (0::int); b < a |] ==> a mod b = a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1524 | apply (rule_tac q = 0 in divmod_int_rel_mod) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1525 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1526 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1527 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1528 | lemma mod_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a mod b = a+b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1529 | apply (rule_tac q = "-1" in divmod_int_rel_mod) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1530 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1531 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1532 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1533 | text{*There is no @{text mod_neg_pos_trivial}.*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1534 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1535 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1536 | (*Simpler laws such as -a div b = -(a div b) FAIL, but see just below*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1537 | lemma zdiv_zminus_zminus [simp]: "(-a) div (-b) = a div (b::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1538 | apply (case_tac "b = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1539 | apply (simp add: divmod_int_rel_div_mod [THEN divmod_int_rel_neg, simplified, | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1540 | THEN divmod_int_rel_div, THEN sym]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1541 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1542 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1543 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1544 | (*Simpler laws such as -a mod b = -(a mod b) FAIL, but see just below*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1545 | lemma zmod_zminus_zminus [simp]: "(-a) mod (-b) = - (a mod (b::int))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1546 | apply (case_tac "b = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1547 | apply (subst divmod_int_rel_div_mod [THEN divmod_int_rel_neg, simplified, THEN divmod_int_rel_mod], | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1548 | auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1549 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1550 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1551 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1552 | subsubsection{*Laws for div and mod with Unary Minus*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1553 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1554 | lemma zminus1_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1555 | "divmod_int_rel a b (q, r) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1556 | ==> divmod_int_rel (-a) b (if r=0 then -q else -q - 1, | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1557 | if r=0 then 0 else b-r)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1558 | by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_diff_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1559 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1560 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1561 | lemma zdiv_zminus1_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1562 | "b \<noteq> (0::int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1563 | ==> (-a) div b = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1564 | (if a mod b = 0 then - (a div b) else - (a div b) - 1)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1565 | by (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN divmod_int_rel_div]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1566 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1567 | lemma zmod_zminus1_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1568 | "(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1569 | apply (case_tac "b = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1570 | apply (blast intro: divmod_int_rel_div_mod [THEN zminus1_lemma, THEN divmod_int_rel_mod]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1571 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1572 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1573 | lemma zmod_zminus1_not_zero: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1574 | fixes k l :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1575 | shows "- k mod l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1576 | unfolding zmod_zminus1_eq_if by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1577 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1578 | lemma zdiv_zminus2: "a div (-b) = (-a::int) div b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1579 | by (cut_tac a = "-a" in zdiv_zminus_zminus, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1580 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1581 | lemma zmod_zminus2: "a mod (-b) = - ((-a::int) mod b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1582 | by (cut_tac a = "-a" and b = b in zmod_zminus_zminus, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1583 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1584 | lemma zdiv_zminus2_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1585 | "b \<noteq> (0::int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1586 | ==> a div (-b) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1587 | (if a mod b = 0 then - (a div b) else - (a div b) - 1)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1588 | by (simp add: zdiv_zminus1_eq_if zdiv_zminus2) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1589 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1590 | lemma zmod_zminus2_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1591 | "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1592 | by (simp add: zmod_zminus1_eq_if zmod_zminus2) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1593 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1594 | lemma zmod_zminus2_not_zero: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1595 | fixes k l :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1596 | shows "k mod - l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1597 | unfolding zmod_zminus2_eq_if by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1598 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1599 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1600 | subsubsection{*Division of a Number by Itself*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1601 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1602 | lemma self_quotient_aux1: "[| (0::int) < a; a = r + a*q; r < a |] ==> 1 \<le> q" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1603 | apply (subgoal_tac "0 < a*q") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1604 | apply (simp add: zero_less_mult_iff, arith) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1605 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1606 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1607 | lemma self_quotient_aux2: "[| (0::int) < a; a = r + a*q; 0 \<le> r |] ==> q \<le> 1" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1608 | apply (subgoal_tac "0 \<le> a* (1-q) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1609 | apply (simp add: zero_le_mult_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1610 | apply (simp add: right_diff_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1611 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1612 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1613 | lemma self_quotient: "[| divmod_int_rel a a (q, r); a \<noteq> (0::int) |] ==> q = 1" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1614 | apply (simp add: split_ifs divmod_int_rel_def linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1615 | apply (rule order_antisym, safe, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1616 | apply (rule_tac [3] a = "-a" and r = "-r" in self_quotient_aux1) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1617 | apply (rule_tac a = "-a" and r = "-r" in self_quotient_aux2) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1618 | apply (force intro: self_quotient_aux1 self_quotient_aux2 simp add: add_commute)+ | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1619 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1620 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1621 | lemma self_remainder: "[| divmod_int_rel a a (q, r); a \<noteq> (0::int) |] ==> r = 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1622 | apply (frule self_quotient, assumption) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1623 | apply (simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1624 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1625 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1626 | lemma zdiv_self [simp]: "a \<noteq> 0 ==> a div a = (1::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1627 | by (simp add: divmod_int_rel_div_mod [THEN self_quotient]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1628 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1629 | (*Here we have 0 mod 0 = 0, also assumed by Knuth (who puts m mod 0 = 0) *) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1630 | lemma zmod_self [simp]: "a mod a = (0::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1631 | apply (case_tac "a = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1632 | apply (simp add: divmod_int_rel_div_mod [THEN self_remainder]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1633 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1634 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1635 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1636 | subsubsection{*Computation of Division and Remainder*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1637 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1638 | lemma zdiv_zero [simp]: "(0::int) div b = 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1639 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1640 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1641 | lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1642 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1643 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1644 | lemma zmod_zero [simp]: "(0::int) mod b = 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1645 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1646 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1647 | lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1648 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1649 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1650 | text{*a positive, b positive *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1651 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1652 | lemma div_pos_pos: "[| 0 < a; 0 \<le> b |] ==> a div b = fst (posDivAlg a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1653 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1654 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1655 | lemma mod_pos_pos: "[| 0 < a; 0 \<le> b |] ==> a mod b = snd (posDivAlg a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1656 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1657 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1658 | text{*a negative, b positive *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1659 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1660 | lemma div_neg_pos: "[| a < 0; 0 < b |] ==> a div b = fst (negDivAlg a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1661 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1662 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1663 | lemma mod_neg_pos: "[| a < 0; 0 < b |] ==> a mod b = snd (negDivAlg a b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1664 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1665 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1666 | text{*a positive, b negative *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1667 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1668 | lemma div_pos_neg: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1669 | "[| 0 < a; b < 0 |] ==> a div b = fst (negateSnd (negDivAlg (-a) (-b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1670 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1671 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1672 | lemma mod_pos_neg: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1673 | "[| 0 < a; b < 0 |] ==> a mod b = snd (negateSnd (negDivAlg (-a) (-b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1674 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1675 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1676 | text{*a negative, b negative *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1677 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1678 | lemma div_neg_neg: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1679 | "[| a < 0; b \<le> 0 |] ==> a div b = fst (negateSnd (posDivAlg (-a) (-b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1680 | by (simp add: div_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1681 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1682 | lemma mod_neg_neg: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1683 | "[| a < 0; b \<le> 0 |] ==> a mod b = snd (negateSnd (posDivAlg (-a) (-b)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1684 | by (simp add: mod_int_def divmod_int_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1685 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1686 | text {*Simplify expresions in which div and mod combine numerical constants*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1687 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1688 | lemma divmod_int_relI: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1689 | "\<lbrakk>a == b * q + r; if 0 < b then 0 \<le> r \<and> r < b else b < r \<and> r \<le> 0\<rbrakk> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1690 | \<Longrightarrow> divmod_int_rel a b (q, r)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1691 | unfolding divmod_int_rel_def by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1692 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1693 | lemmas divmod_int_rel_div_eq = divmod_int_relI [THEN divmod_int_rel_div, THEN eq_reflection] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1694 | lemmas divmod_int_rel_mod_eq = divmod_int_relI [THEN divmod_int_rel_mod, THEN eq_reflection] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1695 | lemmas arithmetic_simps = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1696 | arith_simps | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1697 | add_special | 
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34982diff
changeset | 1698 | add_0_left | 
| 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34982diff
changeset | 1699 | add_0_right | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1700 | mult_zero_left | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1701 | mult_zero_right | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1702 | mult_1_left | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1703 | mult_1_right | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1704 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1705 | (* simprocs adapted from HOL/ex/Binary.thy *) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1706 | ML {*
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1707 | local | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1708 | val mk_number = HOLogic.mk_number HOLogic.intT; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1709 |   fun mk_cert u k l = @{term "plus :: int \<Rightarrow> int \<Rightarrow> int"} $
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1710 |     (@{term "times :: int \<Rightarrow> int \<Rightarrow> int"} $ u $ mk_number k) $
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1711 | mk_number l; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1712 | fun prove ctxt prop = Goal.prove ctxt [] [] prop | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1713 |     (K (ALLGOALS (full_simp_tac (HOL_basic_ss addsimps @{thms arithmetic_simps}))));
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1714 | fun binary_proc proc ss ct = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1715 | (case Thm.term_of ct of | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1716 | _ $ t $ u => | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1717 | (case try (pairself (`(snd o HOLogic.dest_number))) (t, u) of | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1718 | SOME args => proc (Simplifier.the_context ss) args | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1719 | | NONE => NONE) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1720 | | _ => NONE); | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1721 | in | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1722 | fun divmod_proc rule = binary_proc (fn ctxt => fn ((m, t), (n, u)) => | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1723 | if n = 0 then NONE | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1724 | else let val (k, l) = Integer.div_mod m n; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1725 | in SOME (rule OF [prove ctxt (Logic.mk_equals (t, mk_cert u k l))]) end); | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1726 | end | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1727 | *} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1728 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1729 | simproc_setup binary_int_div ("number_of m div number_of n :: int") =
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1730 |   {* K (divmod_proc (@{thm divmod_int_rel_div_eq})) *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1731 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1732 | simproc_setup binary_int_mod ("number_of m mod number_of n :: int") =
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1733 |   {* K (divmod_proc (@{thm divmod_int_rel_mod_eq})) *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1734 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1735 | lemmas posDivAlg_eqn_number_of [simp] = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1736 | posDivAlg_eqn [of "number_of v" "number_of w", standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1737 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1738 | lemmas negDivAlg_eqn_number_of [simp] = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1739 | negDivAlg_eqn [of "number_of v" "number_of w", standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1740 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1741 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1742 | text{*Special-case simplification *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1743 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1744 | lemma zmod_minus1_right [simp]: "a mod (-1::int) = 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1745 | apply (cut_tac a = a and b = "-1" in neg_mod_sign) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1746 | apply (cut_tac [2] a = a and b = "-1" in neg_mod_bound) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1747 | apply (auto simp del: neg_mod_sign neg_mod_bound) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1748 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1749 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1750 | lemma zdiv_minus1_right [simp]: "a div (-1::int) = -a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1751 | by (cut_tac a = a and b = "-1" in zmod_zdiv_equality, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1752 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1753 | (** The last remaining special cases for constant arithmetic: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1754 | 1 div z and 1 mod z **) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1755 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1756 | lemmas div_pos_pos_1_number_of [simp] = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1757 | div_pos_pos [OF int_0_less_1, of "number_of w", standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1758 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1759 | lemmas div_pos_neg_1_number_of [simp] = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1760 | div_pos_neg [OF int_0_less_1, of "number_of w", standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1761 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1762 | lemmas mod_pos_pos_1_number_of [simp] = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1763 | mod_pos_pos [OF int_0_less_1, of "number_of w", standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1764 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1765 | lemmas mod_pos_neg_1_number_of [simp] = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1766 | mod_pos_neg [OF int_0_less_1, of "number_of w", standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1767 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1768 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1769 | lemmas posDivAlg_eqn_1_number_of [simp] = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1770 | posDivAlg_eqn [of concl: 1 "number_of w", standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1771 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1772 | lemmas negDivAlg_eqn_1_number_of [simp] = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1773 | negDivAlg_eqn [of concl: 1 "number_of w", standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1774 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1775 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1776 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1777 | subsubsection{*Monotonicity in the First Argument (Dividend)*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1778 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1779 | lemma zdiv_mono1: "[| a \<le> a'; 0 < (b::int) |] ==> a div b \<le> a' div b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1780 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1781 | apply (cut_tac a = a' and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1782 | apply (rule unique_quotient_lemma) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1783 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1784 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1785 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1786 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1787 | lemma zdiv_mono1_neg: "[| a \<le> a'; (b::int) < 0 |] ==> a' div b \<le> a div b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1788 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1789 | apply (cut_tac a = a' and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1790 | apply (rule unique_quotient_lemma_neg) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1791 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1792 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1793 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1794 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1795 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1796 | subsubsection{*Monotonicity in the Second Argument (Divisor)*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1797 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1798 | lemma q_pos_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1799 | "[| 0 \<le> b'*q' + r'; r' < b'; 0 < b' |] ==> 0 \<le> (q'::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1800 | apply (subgoal_tac "0 < b'* (q' + 1) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1801 | apply (simp add: zero_less_mult_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1802 | apply (simp add: right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1803 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1804 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1805 | lemma zdiv_mono2_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1806 | "[| b*q + r = b'*q' + r'; 0 \<le> b'*q' + r'; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1807 | r' < b'; 0 \<le> r; 0 < b'; b' \<le> b |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1808 | ==> q \<le> (q'::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1809 | apply (frule q_pos_lemma, assumption+) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1810 | apply (subgoal_tac "b*q < b* (q' + 1) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1811 | apply (simp add: mult_less_cancel_left) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1812 | apply (subgoal_tac "b*q = r' - r + b'*q'") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1813 | prefer 2 apply simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1814 | apply (simp (no_asm_simp) add: right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1815 | apply (subst add_commute, rule zadd_zless_mono, arith) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1816 | apply (rule mult_right_mono, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1817 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1818 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1819 | lemma zdiv_mono2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1820 | "[| (0::int) \<le> a; 0 < b'; b' \<le> b |] ==> a div b \<le> a div b'" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1821 | apply (subgoal_tac "b \<noteq> 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1822 | prefer 2 apply arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1823 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1824 | apply (cut_tac a = a and b = b' in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1825 | apply (rule zdiv_mono2_lemma) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1826 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1827 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1828 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1829 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1830 | lemma q_neg_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1831 | "[| b'*q' + r' < 0; 0 \<le> r'; 0 < b' |] ==> q' \<le> (0::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1832 | apply (subgoal_tac "b'*q' < 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1833 | apply (simp add: mult_less_0_iff, arith) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1834 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1835 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1836 | lemma zdiv_mono2_neg_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1837 | "[| b*q + r = b'*q' + r'; b'*q' + r' < 0; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1838 | r < b; 0 \<le> r'; 0 < b'; b' \<le> b |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1839 | ==> q' \<le> (q::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1840 | apply (frule q_neg_lemma, assumption+) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1841 | apply (subgoal_tac "b*q' < b* (q + 1) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1842 | apply (simp add: mult_less_cancel_left) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1843 | apply (simp add: right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1844 | apply (subgoal_tac "b*q' \<le> b'*q'") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1845 | prefer 2 apply (simp add: mult_right_mono_neg, arith) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1846 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1847 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1848 | lemma zdiv_mono2_neg: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1849 | "[| a < (0::int); 0 < b'; b' \<le> b |] ==> a div b' \<le> a div b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1850 | apply (cut_tac a = a and b = b in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1851 | apply (cut_tac a = a and b = b' in zmod_zdiv_equality) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1852 | apply (rule zdiv_mono2_neg_lemma) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1853 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1854 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1855 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1856 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1857 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1858 | subsubsection{*More Algebraic Laws for div and mod*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1859 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1860 | text{*proving (a*b) div c = a * (b div c) + a * (b mod c) *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1861 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1862 | lemma zmult1_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1863 | "[| divmod_int_rel b c (q, r); c \<noteq> 0 |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1864 | ==> divmod_int_rel (a * b) c (a*q + a*r div c, a*r mod c)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1865 | by (auto simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib mult_ac) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1866 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1867 | lemma zdiv_zmult1_eq: "(a*b) div c = a*(b div c) + a*(b mod c) div (c::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1868 | apply (case_tac "c = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1869 | apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN divmod_int_rel_div]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1870 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1871 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1872 | lemma zmod_zmult1_eq: "(a*b) mod c = a*(b mod c) mod (c::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1873 | apply (case_tac "c = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1874 | apply (blast intro: divmod_int_rel_div_mod [THEN zmult1_lemma, THEN divmod_int_rel_mod]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1875 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1876 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1877 | lemma zmod_zdiv_trivial: "(a mod b) div b = (0::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1878 | apply (case_tac "b = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1879 | apply (auto simp add: linorder_neq_iff div_pos_pos_trivial div_neg_neg_trivial) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1880 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1881 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1882 | text{*proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c) *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1883 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1884 | lemma zadd1_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1885 | "[| divmod_int_rel a c (aq, ar); divmod_int_rel b c (bq, br); c \<noteq> 0 |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1886 | ==> divmod_int_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1887 | by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1888 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1889 | (*NOT suitable for rewriting: the RHS has an instance of the LHS*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1890 | lemma zdiv_zadd1_eq: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1891 | "(a+b) div (c::int) = a div c + b div c + ((a mod c + b mod c) div c)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1892 | apply (case_tac "c = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1893 | apply (blast intro: zadd1_lemma [OF divmod_int_rel_div_mod divmod_int_rel_div_mod] divmod_int_rel_div) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1894 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1895 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1896 | instance int :: ring_div | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1897 | proof | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1898 | fix a b c :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1899 | assume not0: "b \<noteq> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1900 | show "(a + c * b) div b = c + a div b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1901 | unfolding zdiv_zadd1_eq [of a "c * b"] using not0 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1902 | by (simp add: zmod_zmult1_eq zmod_zdiv_trivial zdiv_zmult1_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1903 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1904 | fix a b c :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1905 | assume "a \<noteq> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1906 | then show "(a * b) div (a * c) = b div c" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1907 | proof (cases "b \<noteq> 0 \<and> c \<noteq> 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1908 | case False then show ?thesis by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1909 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1910 | case True then have "b \<noteq> 0" and "c \<noteq> 0" by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1911 | with `a \<noteq> 0` | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1912 | have "\<And>q r. divmod_int_rel b c (q, r) \<Longrightarrow> divmod_int_rel (a * b) (a * c) (q, a * r)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1913 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1914 | apply (auto simp add: algebra_simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1915 | apply (auto simp add: zero_less_mult_iff zero_le_mult_iff mult_le_0_iff mult_commute [of a] mult_less_cancel_right) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1916 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1917 | moreover with `c \<noteq> 0` divmod_int_rel_div_mod have "divmod_int_rel b c (b div c, b mod c)" by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1918 | ultimately have "divmod_int_rel (a * b) (a * c) (b div c, a * (b mod c))" . | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1919 | moreover from `a \<noteq> 0` `c \<noteq> 0` have "a * c \<noteq> 0" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1920 | ultimately show ?thesis by (rule divmod_int_rel_div) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1921 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1922 | qed auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1923 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1924 | lemma posDivAlg_div_mod: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1925 | assumes "k \<ge> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1926 | and "l \<ge> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1927 | shows "posDivAlg k l = (k div l, k mod l)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1928 | proof (cases "l = 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1929 | case True then show ?thesis by (simp add: posDivAlg.simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1930 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1931 | case False with assms posDivAlg_correct | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1932 | have "divmod_int_rel k l (fst (posDivAlg k l), snd (posDivAlg k l))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1933 | by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1934 | from divmod_int_rel_div [OF this `l \<noteq> 0`] divmod_int_rel_mod [OF this `l \<noteq> 0`] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1935 | show ?thesis by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1936 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1937 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1938 | lemma negDivAlg_div_mod: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1939 | assumes "k < 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1940 | and "l > 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1941 | shows "negDivAlg k l = (k div l, k mod l)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1942 | proof - | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1943 | from assms have "l \<noteq> 0" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1944 | from assms negDivAlg_correct | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1945 | have "divmod_int_rel k l (fst (negDivAlg k l), snd (negDivAlg k l))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1946 | by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1947 | from divmod_int_rel_div [OF this `l \<noteq> 0`] divmod_int_rel_mod [OF this `l \<noteq> 0`] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1948 | show ?thesis by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1949 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1950 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1951 | lemma zmod_eq_0_iff: "(m mod d = 0) = (EX q::int. m = d*q)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1952 | by (simp add: dvd_eq_mod_eq_0 [symmetric] dvd_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1953 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1954 | (* REVISIT: should this be generalized to all semiring_div types? *) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1955 | lemmas zmod_eq_0D [dest!] = zmod_eq_0_iff [THEN iffD1] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1956 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1957 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1958 | subsubsection{*Proving  @{term "a div (b*c) = (a div b) div c"} *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1959 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1960 | (*The condition c>0 seems necessary. Consider that 7 div ~6 = ~2 but | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1961 | 7 div 2 div ~3 = 3 div ~3 = ~1. The subcase (a div b) mod c = 0 seems | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1962 | to cause particular problems.*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1963 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1964 | text{*first, four lemmas to bound the remainder for the cases b<0 and b>0 *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1965 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1966 | lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \<le> 0 |] ==> b*c < b*(q mod c) + r" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1967 | apply (subgoal_tac "b * (c - q mod c) < r * 1") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1968 | apply (simp add: algebra_simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1969 | apply (rule order_le_less_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1970 | apply (erule_tac [2] mult_strict_right_mono) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1971 | apply (rule mult_left_mono_neg) | 
| 35216 | 1972 | using add1_zle_eq[of "q mod c"]apply(simp add: algebra_simps) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1973 | apply (simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1974 | apply (simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1975 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1976 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1977 | lemma zmult2_lemma_aux2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1978 | "[| (0::int) < c; b < r; r \<le> 0 |] ==> b * (q mod c) + r \<le> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1979 | apply (subgoal_tac "b * (q mod c) \<le> 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1980 | apply arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1981 | apply (simp add: mult_le_0_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1982 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1983 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1984 | lemma zmult2_lemma_aux3: "[| (0::int) < c; 0 \<le> r; r < b |] ==> 0 \<le> b * (q mod c) + r" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1985 | apply (subgoal_tac "0 \<le> b * (q mod c) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1986 | apply arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1987 | apply (simp add: zero_le_mult_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1988 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1989 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1990 | lemma zmult2_lemma_aux4: "[| (0::int) < c; 0 \<le> r; r < b |] ==> b * (q mod c) + r < b * c" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1991 | apply (subgoal_tac "r * 1 < b * (c - q mod c) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1992 | apply (simp add: right_diff_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1993 | apply (rule order_less_le_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1994 | apply (erule mult_strict_right_mono) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1995 | apply (rule_tac [2] mult_left_mono) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1996 | apply simp | 
| 35216 | 1997 | using add1_zle_eq[of "q mod c"] apply (simp add: algebra_simps) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1998 | apply simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1999 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2000 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2001 | lemma zmult2_lemma: "[| divmod_int_rel a b (q, r); b \<noteq> 0; 0 < c |] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2002 | ==> divmod_int_rel a (b * c) (q div c, b*(q mod c) + r)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2003 | by (auto simp add: mult_ac divmod_int_rel_def linorder_neq_iff | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2004 | zero_less_mult_iff right_distrib [symmetric] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2005 | zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2006 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2007 | lemma zdiv_zmult2_eq: "(0::int) < c ==> a div (b*c) = (a div b) div c" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2008 | apply (case_tac "b = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2009 | apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_div]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2010 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2011 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2012 | lemma zmod_zmult2_eq: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2013 | "(0::int) < c ==> a mod (b*c) = b*(a div b mod c) + a mod b" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2014 | apply (case_tac "b = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2015 | apply (force simp add: divmod_int_rel_div_mod [THEN zmult2_lemma, THEN divmod_int_rel_mod]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2016 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2017 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2018 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2019 | subsubsection {*Splitting Rules for div and mod*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2020 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2021 | text{*The proofs of the two lemmas below are essentially identical*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2022 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2023 | lemma split_pos_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2024 | "0<k ==> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2025 | P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i j)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2026 | apply (rule iffI, clarify) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2027 | apply (erule_tac P="P ?x ?y" in rev_mp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2028 | apply (subst mod_add_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2029 | apply (subst zdiv_zadd1_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2030 | apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2031 | txt{*converse direction*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2032 | apply (drule_tac x = "n div k" in spec) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2033 | apply (drule_tac x = "n mod k" in spec, simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2034 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2035 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2036 | lemma split_neg_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2037 | "k<0 ==> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2038 | P(n div k :: int)(n mod k) = (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i j)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2039 | apply (rule iffI, clarify) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2040 | apply (erule_tac P="P ?x ?y" in rev_mp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2041 | apply (subst mod_add_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2042 | apply (subst zdiv_zadd1_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2043 | apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2044 | txt{*converse direction*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2045 | apply (drule_tac x = "n div k" in spec) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2046 | apply (drule_tac x = "n mod k" in spec, simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2047 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2048 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2049 | lemma split_zdiv: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2050 | "P(n div k :: int) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2051 | ((k = 0 --> P 0) & | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2052 | (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i)) & | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2053 | (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P i)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2054 | apply (case_tac "k=0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2055 | apply (simp only: linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2056 | apply (erule disjE) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2057 | apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2058 | split_neg_lemma [of concl: "%x y. P x"]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2059 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2060 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2061 | lemma split_zmod: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2062 | "P(n mod k :: int) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2063 | ((k = 0 --> P n) & | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2064 | (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P j)) & | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2065 | (k<0 --> (\<forall>i j. k<j & j\<le>0 & n = k*i + j --> P j)))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2066 | apply (case_tac "k=0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2067 | apply (simp only: linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2068 | apply (erule disjE) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2069 | apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2070 | split_neg_lemma [of concl: "%x y. P y"]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2071 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2072 | |
| 33730 
1755ca4ec022
Fixed splitting of div and mod on integers (split theorem differed from implementation).
 webertj parents: 
33728diff
changeset | 2073 | text {* Enable (lin)arith to deal with @{const div} and @{const mod}
 | 
| 
1755ca4ec022
Fixed splitting of div and mod on integers (split theorem differed from implementation).
 webertj parents: 
33728diff
changeset | 2074 | when these are applied to some constant that is of the form | 
| 
1755ca4ec022
Fixed splitting of div and mod on integers (split theorem differed from implementation).
 webertj parents: 
33728diff
changeset | 2075 |   @{term "number_of k"}: *}
 | 
| 33728 
cb4235333c30
Fixed splitting of div and mod on integers (split theorem differed from implementation).
 webertj parents: 
33364diff
changeset | 2076 | declare split_zdiv [of _ _ "number_of k", standard, arith_split] | 
| 
cb4235333c30
Fixed splitting of div and mod on integers (split theorem differed from implementation).
 webertj parents: 
33364diff
changeset | 2077 | declare split_zmod [of _ _ "number_of k", standard, arith_split] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2078 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2079 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2080 | subsubsection{*Speeding up the Division Algorithm with Shifting*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2081 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2082 | text{*computing div by shifting *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2083 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2084 | lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2085 | proof cases | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2086 | assume "a=0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2087 | thus ?thesis by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2088 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2089 | assume "a\<noteq>0" and le_a: "0\<le>a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2090 | hence a_pos: "1 \<le> a" by arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2091 | hence one_less_a2: "1 < 2 * a" by arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2092 | hence le_2a: "2 * (1 + b mod a) \<le> 2 * a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2093 | unfolding mult_le_cancel_left | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2094 | by (simp add: add1_zle_eq add_commute [of 1]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2095 | with a_pos have "0 \<le> b mod a" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2096 | hence le_addm: "0 \<le> 1 mod (2*a) + 2*(b mod a)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2097 | by (simp add: mod_pos_pos_trivial one_less_a2) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2098 | with le_2a | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2099 | have "(1 mod (2*a) + 2*(b mod a)) div (2*a) = 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2100 | by (simp add: div_pos_pos_trivial le_addm mod_pos_pos_trivial one_less_a2 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2101 | right_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2102 | thus ?thesis | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2103 | by (subst zdiv_zadd1_eq, | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2104 | simp add: mod_mult_mult1 one_less_a2 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2105 | div_pos_pos_trivial) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2106 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2107 | |
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2108 | lemma neg_zdiv_mult_2: | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2109 | assumes A: "a \<le> (0::int)" shows "(1 + 2*b) div (2*a) = (b+1) div a" | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2110 | proof - | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2111 | have R: "1 + - (2 * (b + 1)) = - (1 + 2 * b)" by simp | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2112 | have "(1 + 2 * (-b - 1)) div (2 * (-a)) = (-b - 1) div (-a)" | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2113 | by (rule pos_zdiv_mult_2, simp add: A) | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2114 | thus ?thesis | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2115 | by (simp only: R zdiv_zminus_zminus diff_minus | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2116 | minus_add_distrib [symmetric] mult_minus_right) | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2117 | qed | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2118 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2119 | lemma zdiv_number_of_Bit0 [simp]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2120 | "number_of (Int.Bit0 v) div number_of (Int.Bit0 w) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2121 | number_of v div (number_of w :: int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2122 | by (simp only: number_of_eq numeral_simps) (simp add: mult_2 [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2123 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2124 | lemma zdiv_number_of_Bit1 [simp]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2125 | "number_of (Int.Bit1 v) div number_of (Int.Bit0 w) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2126 | (if (0::int) \<le> number_of w | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2127 | then number_of v div (number_of w) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2128 | else (number_of v + (1::int)) div (number_of w))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2129 | apply (simp only: number_of_eq numeral_simps UNIV_I split: split_if) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2130 | apply (simp add: pos_zdiv_mult_2 neg_zdiv_mult_2 add_ac mult_2 [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2131 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2132 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2133 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2134 | subsubsection{*Computing mod by Shifting (proofs resemble those for div)*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2135 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2136 | lemma pos_zmod_mult_2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2137 | fixes a b :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2138 | assumes "0 \<le> a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2139 | shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2140 | proof (cases "0 < a") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2141 | case False with assms show ?thesis by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2142 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2143 | case True | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2144 | then have "b mod a < a" by (rule pos_mod_bound) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2145 | then have "1 + b mod a \<le> a" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2146 | then have A: "2 * (1 + b mod a) \<le> 2 * a" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2147 | from `0 < a` have "0 \<le> b mod a" by (rule pos_mod_sign) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2148 | then have B: "0 \<le> 1 + 2 * (b mod a)" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2149 | have "((1\<Colon>int) mod ((2\<Colon>int) * a) + (2\<Colon>int) * b mod ((2\<Colon>int) * a)) mod ((2\<Colon>int) * a) = (1\<Colon>int) + (2\<Colon>int) * (b mod a)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2150 | using `0 < a` and A | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2151 | by (auto simp add: mod_mult_mult1 mod_pos_pos_trivial ring_distribs intro!: mod_pos_pos_trivial B) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2152 | then show ?thesis by (subst mod_add_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2153 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2154 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2155 | lemma neg_zmod_mult_2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2156 | fixes a b :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2157 | assumes "a \<le> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2158 | shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2159 | proof - | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2160 | from assms have "0 \<le> - a" by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2161 | then have "(1 + 2 * (- b - 1)) mod (2 * (- a)) = 1 + 2 * ((- b - 1) mod (- a))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2162 | by (rule pos_zmod_mult_2) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2163 | then show ?thesis by (simp add: zmod_zminus2 algebra_simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2164 | (simp add: diff_minus add_ac) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2165 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2166 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2167 | lemma zmod_number_of_Bit0 [simp]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2168 | "number_of (Int.Bit0 v) mod number_of (Int.Bit0 w) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2169 | (2::int) * (number_of v mod number_of w)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2170 | apply (simp only: number_of_eq numeral_simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2171 | apply (simp add: mod_mult_mult1 pos_zmod_mult_2 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2172 | neg_zmod_mult_2 add_ac mult_2 [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2173 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2174 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2175 | lemma zmod_number_of_Bit1 [simp]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2176 | "number_of (Int.Bit1 v) mod number_of (Int.Bit0 w) = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2177 | (if (0::int) \<le> number_of w | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2178 | then 2 * (number_of v mod number_of w) + 1 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2179 | else 2 * ((number_of v + (1::int)) mod number_of w) - 1)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2180 | apply (simp only: number_of_eq numeral_simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2181 | apply (simp add: mod_mult_mult1 pos_zmod_mult_2 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2182 | neg_zmod_mult_2 add_ac mult_2 [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2183 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2184 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2185 | |
| 39489 | 2186 | lemma zdiv_eq_0_iff: | 
| 2187 | "(i::int) div k = 0 \<longleftrightarrow> k=0 \<or> 0\<le>i \<and> i<k \<or> i\<le>0 \<and> k<i" (is "?L = ?R") | |
| 2188 | proof | |
| 2189 | assume ?L | |
| 2190 | have "?L \<longrightarrow> ?R" by (rule split_zdiv[THEN iffD2]) simp | |
| 2191 | with `?L` show ?R by blast | |
| 2192 | next | |
| 2193 | assume ?R thus ?L | |
| 2194 | by(auto simp: div_pos_pos_trivial div_neg_neg_trivial) | |
| 2195 | qed | |
| 2196 | ||
| 2197 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2198 | subsubsection{*Quotients of Signs*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2199 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2200 | lemma div_neg_pos_less0: "[| a < (0::int); 0 < b |] ==> a div b < 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2201 | apply (subgoal_tac "a div b \<le> -1", force) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2202 | apply (rule order_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2203 | apply (rule_tac a' = "-1" in zdiv_mono1) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2204 | apply (auto simp add: div_eq_minus1) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2205 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2206 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2207 | lemma div_nonneg_neg_le0: "[| (0::int) \<le> a; b < 0 |] ==> a div b \<le> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2208 | by (drule zdiv_mono1_neg, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2209 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2210 | lemma div_nonpos_pos_le0: "[| (a::int) \<le> 0; b > 0 |] ==> a div b \<le> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2211 | by (drule zdiv_mono1, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2212 | |
| 33804 | 2213 | text{* Now for some equivalences of the form @{text"a div b >=< 0 \<longleftrightarrow> \<dots>"}
 | 
| 2214 | conditional upon the sign of @{text a} or @{text b}. There are many more.
 | |
| 2215 | They should all be simp rules unless that causes too much search. *} | |
| 2216 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2217 | lemma pos_imp_zdiv_nonneg_iff: "(0::int) < b ==> (0 \<le> a div b) = (0 \<le> a)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2218 | apply auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2219 | apply (drule_tac [2] zdiv_mono1) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2220 | apply (auto simp add: linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2221 | apply (simp (no_asm_use) add: linorder_not_less [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2222 | apply (blast intro: div_neg_pos_less0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2223 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2224 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2225 | lemma neg_imp_zdiv_nonneg_iff: | 
| 33804 | 2226 | "b < (0::int) ==> (0 \<le> a div b) = (a \<le> (0::int))" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2227 | apply (subst zdiv_zminus_zminus [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2228 | apply (subst pos_imp_zdiv_nonneg_iff, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2229 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2230 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2231 | (*But not (a div b \<le> 0 iff a\<le>0); consider a=1, b=2 when a div b = 0.*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2232 | lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2233 | by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2234 | |
| 39489 | 2235 | lemma pos_imp_zdiv_pos_iff: | 
| 2236 | "0<k \<Longrightarrow> 0 < (i::int) div k \<longleftrightarrow> k \<le> i" | |
| 2237 | using pos_imp_zdiv_nonneg_iff[of k i] zdiv_eq_0_iff[of i k] | |
| 2238 | by arith | |
| 2239 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2240 | (*Again the law fails for \<le>: consider a = -1, b = -2 when a div b = 0*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2241 | lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2242 | by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2243 | |
| 33804 | 2244 | lemma nonneg1_imp_zdiv_pos_iff: | 
| 2245 | "(0::int) <= a \<Longrightarrow> (a div b > 0) = (a >= b & b>0)" | |
| 2246 | apply rule | |
| 2247 | apply rule | |
| 2248 | using div_pos_pos_trivial[of a b]apply arith | |
| 2249 | apply(cases "b=0")apply simp | |
| 2250 | using div_nonneg_neg_le0[of a b]apply arith | |
| 2251 | using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b]apply simp | |
| 2252 | done | |
| 2253 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2254 | |
| 39489 | 2255 | lemma zmod_le_nonneg_dividend: "(m::int) \<ge> 0 ==> m mod k \<le> m" | 
| 2256 | apply (rule split_zmod[THEN iffD2]) | |
| 2257 | apply(fastsimp dest: q_pos_lemma intro: split_mult_pos_le) | |
| 2258 | done | |
| 2259 | ||
| 2260 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2261 | subsubsection {* The Divides Relation *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2262 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2263 | lemmas zdvd_iff_zmod_eq_0_number_of [simp] = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2264 | dvd_eq_mod_eq_0 [of "number_of x::int" "number_of y::int", standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2265 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2266 | lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2267 | by (rule dvd_mod) (* TODO: remove *) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2268 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2269 | lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2270 | by (rule dvd_mod_imp_dvd) (* TODO: remove *) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2271 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2272 | lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2273 | using zmod_zdiv_equality[where a="m" and b="n"] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2274 | by (simp add: algebra_simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2275 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2276 | lemma zpower_zmod: "((x::int) mod m)^y mod m = x^y mod m" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2277 | apply (induct "y", auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2278 | apply (rule zmod_zmult1_eq [THEN trans]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2279 | apply (simp (no_asm_simp)) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2280 | apply (rule mod_mult_eq [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2281 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2282 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2283 | lemma zdiv_int: "int (a div b) = (int a) div (int b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2284 | apply (subst split_div, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2285 | apply (subst split_zdiv, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2286 | apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in unique_quotient) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2287 | apply (auto simp add: divmod_int_rel_def of_nat_mult) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2288 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2289 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2290 | lemma zmod_int: "int (a mod b) = (int a) mod (int b)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2291 | apply (subst split_mod, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2292 | apply (subst split_zmod, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2293 | apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2294 | in unique_remainder) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2295 | apply (auto simp add: divmod_int_rel_def of_nat_mult) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2296 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2297 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2298 | lemma abs_div: "(y::int) dvd x \<Longrightarrow> abs (x div y) = abs x div abs y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2299 | by (unfold dvd_def, cases "y=0", auto simp add: abs_mult) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2300 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2301 | lemma zdvd_mult_div_cancel:"(n::int) dvd m \<Longrightarrow> n * (m div n) = m" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2302 | apply (subgoal_tac "m mod n = 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2303 | apply (simp add: zmult_div_cancel) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2304 | apply (simp only: dvd_eq_mod_eq_0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2305 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2306 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2307 | text{*Suggested by Matthias Daum*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2308 | lemma int_power_div_base: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2309 | "\<lbrakk>0 < m; 0 < k\<rbrakk> \<Longrightarrow> k ^ m div k = (k::int) ^ (m - Suc 0)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2310 | apply (subgoal_tac "k ^ m = k ^ ((m - Suc 0) + Suc 0)") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2311 | apply (erule ssubst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2312 | apply (simp only: power_add) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2313 | apply simp_all | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2314 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2315 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2316 | text {* by Brian Huffman *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2317 | lemma zminus_zmod: "- ((x::int) mod m) mod m = - x mod m" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2318 | by (rule mod_minus_eq [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2319 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2320 | lemma zdiff_zmod_left: "(x mod m - y) mod m = (x - y) mod (m::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2321 | by (rule mod_diff_left_eq [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2322 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2323 | lemma zdiff_zmod_right: "(x - y mod m) mod m = (x - y) mod (m::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2324 | by (rule mod_diff_right_eq [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2325 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2326 | lemmas zmod_simps = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2327 | mod_add_left_eq [symmetric] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2328 | mod_add_right_eq [symmetric] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2329 | zmod_zmult1_eq [symmetric] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2330 | mod_mult_left_eq [symmetric] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2331 | zpower_zmod | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2332 | zminus_zmod zdiff_zmod_left zdiff_zmod_right | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2333 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2334 | text {* Distributive laws for function @{text nat}. *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2335 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2336 | lemma nat_div_distrib: "0 \<le> x \<Longrightarrow> nat (x div y) = nat x div nat y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2337 | apply (rule linorder_cases [of y 0]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2338 | apply (simp add: div_nonneg_neg_le0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2339 | apply simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2340 | apply (simp add: nat_eq_iff pos_imp_zdiv_nonneg_iff zdiv_int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2341 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2342 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2343 | (*Fails if y<0: the LHS collapses to (nat z) but the RHS doesn't*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2344 | lemma nat_mod_distrib: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2345 | "\<lbrakk>0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> nat (x mod y) = nat x mod nat y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2346 | apply (case_tac "y = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2347 | apply (simp add: nat_eq_iff zmod_int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2348 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2349 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2350 | text  {* transfer setup *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2351 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2352 | lemma transfer_nat_int_functions: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2353 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) div (nat y) = nat (x div y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2354 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) mod (nat y) = nat (x mod y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2355 | by (auto simp add: nat_div_distrib nat_mod_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2356 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2357 | lemma transfer_nat_int_function_closures: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2358 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x div y >= 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2359 | "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x mod y >= 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2360 | apply (cases "y = 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2361 | apply (auto simp add: pos_imp_zdiv_nonneg_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2362 | apply (cases "y = 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2363 | apply auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2364 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2365 | |
| 35644 | 2366 | declare transfer_morphism_nat_int [transfer add return: | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2367 | transfer_nat_int_functions | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2368 | transfer_nat_int_function_closures | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2369 | ] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2370 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2371 | lemma transfer_int_nat_functions: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2372 | "(int x) div (int y) = int (x div y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2373 | "(int x) mod (int y) = int (x mod y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2374 | by (auto simp add: zdiv_int zmod_int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2375 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2376 | lemma transfer_int_nat_function_closures: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2377 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x div y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2378 | "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x mod y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2379 | by (simp_all only: is_nat_def transfer_nat_int_function_closures) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2380 | |
| 35644 | 2381 | declare transfer_morphism_int_nat [transfer add return: | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2382 | transfer_int_nat_functions | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2383 | transfer_int_nat_function_closures | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2384 | ] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2385 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2386 | text{*Suggested by Matthias Daum*}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2387 | lemma int_div_less_self: "\<lbrakk>0 < x; 1 < k\<rbrakk> \<Longrightarrow> x div k < (x::int)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2388 | apply (subgoal_tac "nat x div nat k < nat x") | 
| 34225 | 2389 | apply (simp add: nat_div_distrib [symmetric]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2390 | apply (rule Divides.div_less_dividend, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2391 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2392 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2393 | lemma zmod_eq_dvd_iff: "(x::int) mod n = y mod n \<longleftrightarrow> n dvd x - y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2394 | proof | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2395 | assume H: "x mod n = y mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2396 | hence "x mod n - y mod n = 0" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2397 | hence "(x mod n - y mod n) mod n = 0" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2398 | hence "(x - y) mod n = 0" by (simp add: mod_diff_eq[symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2399 | thus "n dvd x - y" by (simp add: dvd_eq_mod_eq_0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2400 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2401 | assume H: "n dvd x - y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2402 | then obtain k where k: "x-y = n*k" unfolding dvd_def by blast | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2403 | hence "x = n*k + y" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2404 | hence "x mod n = (n*k + y) mod n" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2405 | thus "x mod n = y mod n" by (simp add: mod_add_left_eq) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2406 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2407 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2408 | lemma nat_mod_eq_lemma: assumes xyn: "(x::nat) mod n = y mod n" and xy:"y \<le> x" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2409 | shows "\<exists>q. x = y + n * q" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2410 | proof- | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2411 | from xy have th: "int x - int y = int (x - y)" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2412 | from xyn have "int x mod int n = int y mod int n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2413 | by (simp add: zmod_int[symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2414 | hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2415 | hence "n dvd x - y" by (simp add: th zdvd_int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2416 | then show ?thesis using xy unfolding dvd_def apply clarsimp apply (rule_tac x="k" in exI) by arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2417 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2418 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2419 | lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2420 | (is "?lhs = ?rhs") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2421 | proof | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2422 | assume H: "x mod n = y mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2423 |   {assume xy: "x \<le> y"
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2424 | from H have th: "y mod n = x mod n" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2425 | from nat_mod_eq_lemma[OF th xy] have ?rhs | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2426 | apply clarify apply (rule_tac x="q" in exI) by (rule exI[where x="0"], simp)} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2427 | moreover | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2428 |   {assume xy: "y \<le> x"
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2429 | from nat_mod_eq_lemma[OF H xy] have ?rhs | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2430 | apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)} | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2431 | ultimately show ?rhs using linear[of x y] by blast | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2432 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2433 | assume ?rhs then obtain q1 q2 where q12: "x + n * q1 = y + n * q2" by blast | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2434 | hence "(x + n * q1) mod n = (y + n * q2) mod n" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2435 | thus ?lhs by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2436 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2437 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2438 | lemma div_nat_number_of [simp]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2439 | "(number_of v :: nat) div number_of v' = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2440 | (if neg (number_of v :: int) then 0 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2441 | else nat (number_of v div number_of v'))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2442 | unfolding nat_number_of_def number_of_is_id neg_def | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2443 | by (simp add: nat_div_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2444 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2445 | lemma one_div_nat_number_of [simp]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2446 | "Suc 0 div number_of v' = nat (1 div number_of v')" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2447 | by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2448 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2449 | lemma mod_nat_number_of [simp]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2450 | "(number_of v :: nat) mod number_of v' = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2451 | (if neg (number_of v :: int) then 0 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2452 | else if neg (number_of v' :: int) then number_of v | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2453 | else nat (number_of v mod number_of v'))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2454 | unfolding nat_number_of_def number_of_is_id neg_def | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2455 | by (simp add: nat_mod_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2456 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2457 | lemma one_mod_nat_number_of [simp]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2458 | "Suc 0 mod number_of v' = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2459 | (if neg (number_of v' :: int) then Suc 0 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2460 | else nat (1 mod number_of v'))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2461 | by (simp del: nat_numeral_1_eq_1 add: numeral_1_eq_Suc_0 [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2462 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2463 | lemmas dvd_eq_mod_eq_0_number_of = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2464 | dvd_eq_mod_eq_0 [of "number_of x" "number_of y", standard] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2465 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2466 | declare dvd_eq_mod_eq_0_number_of [simp] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2467 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2468 | |
| 34126 | 2469 | subsubsection {* Nitpick *}
 | 
| 2470 | ||
| 2471 | lemma zmod_zdiv_equality': | |
| 2472 | "(m\<Colon>int) mod n = m - (m div n) * n" | |
| 2473 | by (rule_tac P="%x. m mod n = x - (m div n) * n" | |
| 2474 | in subst [OF mod_div_equality [of _ n]]) | |
| 2475 | arith | |
| 2476 | ||
| 34982 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34225diff
changeset | 2477 | lemmas [nitpick_def] = dvd_eq_mod_eq_0 [THEN eq_reflection] | 
| 
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
 blanchet parents: 
34225diff
changeset | 2478 | mod_div_equality' [THEN eq_reflection] | 
| 34126 | 2479 | zmod_zdiv_equality' [THEN eq_reflection] | 
| 2480 | ||
| 35673 | 2481 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2482 | subsubsection {* Code generation *}
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2483 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2484 | definition pdivmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2485 | "pdivmod k l = (\<bar>k\<bar> div \<bar>l\<bar>, \<bar>k\<bar> mod \<bar>l\<bar>)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2486 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2487 | lemma pdivmod_posDivAlg [code]: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2488 | "pdivmod k l = (if l = 0 then (0, \<bar>k\<bar>) else posDivAlg \<bar>k\<bar> \<bar>l\<bar>)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2489 | by (subst posDivAlg_div_mod) (simp_all add: pdivmod_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2490 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2491 | lemma divmod_int_pdivmod: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2492 | apsnd ((op *) (sgn l)) (if 0 < l \<and> 0 \<le> k \<or> l < 0 \<and> k < 0 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2493 | then pdivmod k l | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2494 | else (let (r, s) = pdivmod k l in | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2495 | if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2496 | proof - | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2497 | have aux: "\<And>q::int. - k = l * q \<longleftrightarrow> k = l * - q" by auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2498 | show ?thesis | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2499 | by (simp add: divmod_int_mod_div pdivmod_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2500 | (auto simp add: aux not_less not_le zdiv_zminus1_eq_if | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2501 | zmod_zminus1_eq_if zdiv_zminus2_eq_if zmod_zminus2_eq_if) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2502 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2503 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2504 | lemma divmod_int_code [code]: "divmod_int k l = (if k = 0 then (0, 0) else if l = 0 then (0, k) else | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2505 | apsnd ((op *) (sgn l)) (if sgn k = sgn l | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2506 | then pdivmod k l | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2507 | else (let (r, s) = pdivmod k l in | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2508 | if s = 0 then (- r, 0) else (- r - 1, \<bar>l\<bar> - s))))" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2509 | proof - | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2510 | have "k \<noteq> 0 \<Longrightarrow> l \<noteq> 0 \<Longrightarrow> 0 < l \<and> 0 \<le> k \<or> l < 0 \<and> k < 0 \<longleftrightarrow> sgn k = sgn l" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2511 | by (auto simp add: not_less sgn_if) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2512 | then show ?thesis by (simp add: divmod_int_pdivmod) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2513 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2514 | |
| 35673 | 2515 | context ring_1 | 
| 2516 | begin | |
| 2517 | ||
| 2518 | lemma of_int_num [code]: | |
| 2519 | "of_int k = (if k = 0 then 0 else if k < 0 then | |
| 2520 | - of_int (- k) else let | |
| 2521 | (l, m) = divmod_int k 2; | |
| 2522 | l' = of_int l | |
| 2523 | in if m = 0 then l' + l' else l' + l' + 1)" | |
| 2524 | proof - | |
| 2525 | have aux1: "k mod (2\<Colon>int) \<noteq> (0\<Colon>int) \<Longrightarrow> | |
| 2526 | of_int k = of_int (k div 2 * 2 + 1)" | |
| 2527 | proof - | |
| 2528 | have "k mod 2 < 2" by (auto intro: pos_mod_bound) | |
| 2529 | moreover have "0 \<le> k mod 2" by (auto intro: pos_mod_sign) | |
| 2530 | moreover assume "k mod 2 \<noteq> 0" | |
| 2531 | ultimately have "k mod 2 = 1" by arith | |
| 2532 | moreover have "of_int k = of_int (k div 2 * 2 + k mod 2)" by simp | |
| 2533 | ultimately show ?thesis by auto | |
| 2534 | qed | |
| 2535 | have aux2: "\<And>x. of_int 2 * x = x + x" | |
| 2536 | proof - | |
| 2537 | fix x | |
| 2538 | have int2: "(2::int) = 1 + 1" by arith | |
| 2539 | show "of_int 2 * x = x + x" | |
| 2540 | unfolding int2 of_int_add left_distrib by simp | |
| 2541 | qed | |
| 2542 | have aux3: "\<And>x. x * of_int 2 = x + x" | |
| 2543 | proof - | |
| 2544 | fix x | |
| 2545 | have int2: "(2::int) = 1 + 1" by arith | |
| 2546 | show "x * of_int 2 = x + x" | |
| 2547 | unfolding int2 of_int_add right_distrib by simp | |
| 2548 | qed | |
| 2549 | from aux1 show ?thesis by (auto simp add: divmod_int_mod_div Let_def aux2 aux3) | |
| 2550 | qed | |
| 2551 | ||
| 2552 | end | |
| 2553 | ||
| 33364 | 2554 | code_modulename SML | 
| 2555 | Divides Arith | |
| 2556 | ||
| 2557 | code_modulename OCaml | |
| 2558 | Divides Arith | |
| 2559 | ||
| 2560 | code_modulename Haskell | |
| 2561 | Divides Arith | |
| 2562 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2563 | end |