| author | wenzelm | 
| Wed, 14 Oct 2015 17:24:21 +0200 | |
| changeset 61441 | 20ff1d5c74e1 | 
| parent 61275 | 053ec04ea866 | 
| child 61433 | a4c0de1df3d8 | 
| 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 | |
| 60758 | 6 | section \<open>The division operators div and mod\<close> | 
| 3366 | 7 | |
| 15131 | 8 | theory Divides | 
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 9 | imports Parity | 
| 15131 | 10 | begin | 
| 3366 | 11 | |
| 60758 | 12 | subsection \<open>Abstract division in commutative semirings.\<close> | 
| 25942 | 13 | |
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 14 | class div = dvd + divide + | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 15 | fixes mod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "mod" 70) | 
| 25942 | 16 | |
| 59833 
ab828c2c5d67
clarified no_zero_devisors: makes only sense in a semiring;
 haftmann parents: 
59816diff
changeset | 17 | class semiring_div = semidom + div + | 
| 25942 | 18 | 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 | 19 | 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 | 20 | 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 | 21 | and div_mult_self1 [simp]: "b \<noteq> 0 \<Longrightarrow> (a + c * b) div b = c + a div b" | 
| 30930 | 22 | and div_mult_mult1 [simp]: "c \<noteq> 0 \<Longrightarrow> (c * a) div (c * b) = a div b" | 
| 25942 | 23 | begin | 
| 24 | ||
| 60517 
f16e4fb20652
separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
 haftmann parents: 
60516diff
changeset | 25 | subclass algebraic_semidom | 
| 60353 
838025c6e278
implicit partial divison operation in integral domains
 haftmann parents: 
60352diff
changeset | 26 | proof | 
| 
838025c6e278
implicit partial divison operation in integral domains
 haftmann parents: 
60352diff
changeset | 27 | fix b a | 
| 
838025c6e278
implicit partial divison operation in integral domains
 haftmann parents: 
60352diff
changeset | 28 | assume "b \<noteq> 0" | 
| 
838025c6e278
implicit partial divison operation in integral domains
 haftmann parents: 
60352diff
changeset | 29 | then show "a * b div b = a" | 
| 
838025c6e278
implicit partial divison operation in integral domains
 haftmann parents: 
60352diff
changeset | 30 | using div_mult_self1 [of b 0 a] by (simp add: ac_simps) | 
| 
838025c6e278
implicit partial divison operation in integral domains
 haftmann parents: 
60352diff
changeset | 31 | qed simp | 
| 58953 | 32 | |
| 60867 | 33 | lemma div_by_1: | 
| 34 | "a div 1 = a" | |
| 35 | by (fact divide_1) | |
| 36 | ||
| 37 | lemma div_mult_self1_is_id: | |
| 38 | "b \<noteq> 0 \<Longrightarrow> b * a div b = a" | |
| 39 | by (fact nonzero_mult_divide_cancel_left) | |
| 40 | ||
| 41 | lemma div_mult_self2_is_id: | |
| 42 | "b \<noteq> 0 \<Longrightarrow> a * b div b = a" | |
| 43 | by (fact nonzero_mult_divide_cancel_right) | |
| 59009 
348561aa3869
generalized lemmas (particularly concerning dvd) as far as appropriate
 haftmann parents: 
59000diff
changeset | 44 | |
| 60758 | 45 | text \<open>@{const divide} and @{const mod}\<close>
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 46 | |
| 26062 | 47 | lemma mod_div_equality2: "b * (a div b) + a mod b = a" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 48 | unfolding mult.commute [of b] | 
| 26062 | 49 | by (rule mod_div_equality) | 
| 50 | ||
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 51 | 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 | 52 | using mod_div_equality [of a b] | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 53 | by (simp only: ac_simps) | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 54 | |
| 26062 | 55 | lemma div_mod_equality: "((a div b) * b + a mod b) + c = a + c" | 
| 30934 | 56 | by (simp add: mod_div_equality) | 
| 26062 | 57 | |
| 58 | lemma div_mod_equality2: "(b * (a div b) + a mod b) + c = a + c" | |
| 30934 | 59 | by (simp add: mod_div_equality2) | 
| 26062 | 60 | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 61 | lemma mod_by_0 [simp]: "a mod 0 = a" | 
| 30934 | 62 | 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 | 63 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 64 | lemma mod_0 [simp]: "0 mod a = 0" | 
| 30934 | 65 | 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 | 66 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 67 | lemma div_mult_self2 [simp]: | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 68 | assumes "b \<noteq> 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 69 | shows "(a + b * c) div b = c + a div b" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 70 | 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 | 71 | |
| 54221 | 72 | lemma div_mult_self3 [simp]: | 
| 73 | assumes "b \<noteq> 0" | |
| 74 | shows "(c * b + a) div b = c + a div b" | |
| 75 | using assms by (simp add: add.commute) | |
| 76 | ||
| 77 | lemma div_mult_self4 [simp]: | |
| 78 | assumes "b \<noteq> 0" | |
| 79 | shows "(b * c + a) div b = c + a div b" | |
| 80 | using assms by (simp add: add.commute) | |
| 81 | ||
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 82 | 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 | 83 | proof (cases "b = 0") | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 84 | 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 | 85 | next | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 86 | case False | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 87 | 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 | 88 | by (simp add: mod_div_equality) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 89 | 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 | 90 | "\<dots> = (c + a div b) * b + (a + c * b) mod b" | 
| 29667 | 91 | by (simp add: algebra_simps) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 92 | finally have "a = a div b * b + (a + c * b) mod b" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 93 | by (simp add: add.commute [of a] add.assoc distrib_right) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 94 | 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 | 95 | by (simp add: mod_div_equality) | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 96 | then show ?thesis by simp | 
| 
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 | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 99 | lemma mod_mult_self2 [simp]: | 
| 54221 | 100 | "(a + b * c) mod b = a mod b" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 101 | 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 | 102 | |
| 54221 | 103 | lemma mod_mult_self3 [simp]: | 
| 104 | "(c * b + a) mod b = a mod b" | |
| 105 | by (simp add: add.commute) | |
| 106 | ||
| 107 | lemma mod_mult_self4 [simp]: | |
| 108 | "(b * c + a) mod b = a mod b" | |
| 109 | by (simp add: add.commute) | |
| 110 | ||
| 60867 | 111 | lemma mod_mult_self1_is_0 [simp]: | 
| 112 | "b * a mod b = 0" | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 113 | 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 | 114 | |
| 60867 | 115 | lemma mod_mult_self2_is_0 [simp]: | 
| 116 | "a * b mod b = 0" | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 117 | using mod_mult_self1 [of 0 a b] by simp | 
| 26062 | 118 | |
| 60867 | 119 | lemma mod_by_1 [simp]: | 
| 120 | "a mod 1 = 0" | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 121 | proof - | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 122 | 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 | 123 | 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 | 124 | 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 | 125 | qed | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 126 | |
| 60867 | 127 | lemma mod_self [simp]: | 
| 128 | "a mod a = 0" | |
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 129 | 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 | 130 | |
| 27676 | 131 | lemma div_add_self1 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 132 | assumes "b \<noteq> 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 133 | shows "(b + a) div b = a div b + 1" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 134 | using assms div_mult_self1 [of b a 1] by (simp add: add.commute) | 
| 26062 | 135 | |
| 27676 | 136 | lemma div_add_self2 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 137 | assumes "b \<noteq> 0" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 138 | shows "(a + b) div b = a div b + 1" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 139 | using assms div_add_self1 [of b a] by (simp add: add.commute) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 140 | |
| 27676 | 141 | lemma mod_add_self1 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 142 | "(b + a) mod b = a mod b" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 143 | using mod_mult_self1 [of a 1 b] by (simp add: add.commute) | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 144 | |
| 27676 | 145 | lemma mod_add_self2 [simp]: | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 146 | "(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 | 147 | 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 | 148 | |
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 149 | lemma mod_div_decomp: | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 150 | fixes a b | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 151 | 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 | 152 | and "a = q * b + r" | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 153 | proof - | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 154 | 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 | 155 | 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 | 156 | 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 | 157 | 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 | 158 | qed | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 159 | |
| 58834 | 160 | lemma dvd_imp_mod_0 [simp]: | 
| 161 | assumes "a dvd b" | |
| 162 | shows "b mod a = 0" | |
| 163 | proof - | |
| 164 | from assms obtain c where "b = a * c" .. | |
| 165 | then have "b mod a = a * c mod a" by simp | |
| 166 | then show "b mod a = 0" by simp | |
| 167 | qed | |
| 58911 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 168 | |
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 169 | lemma mod_eq_0_iff_dvd: | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 170 | "a mod b = 0 \<longleftrightarrow> b dvd a" | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 171 | proof | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 172 | assume "b dvd a" | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 173 | then show "a mod b = 0" by simp | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 174 | next | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 175 | assume "a mod b = 0" | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 176 | with mod_div_equality [of a b] have "a div b * b = a" by simp | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 177 | then have "a = b * (a div b)" by (simp add: ac_simps) | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 178 | then show "b dvd a" .. | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 179 | qed | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 180 | |
| 60867 | 181 | lemma dvd_eq_mod_eq_0 [nitpick_unfold, code]: | 
| 58834 | 182 | "a dvd b \<longleftrightarrow> b mod a = 0" | 
| 58911 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 183 | by (simp add: mod_eq_0_iff_dvd) | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 184 | |
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 185 | lemma mod_div_trivial [simp]: | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 186 | "a mod b div b = 0" | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 187 | proof (cases "b = 0") | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 188 | assume "b = 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 189 | thus ?thesis by simp | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 190 | next | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 191 | assume "b \<noteq> 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 192 | 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 | 193 | by (rule div_mult_self1 [symmetric]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 194 | also have "\<dots> = a div b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 195 | by (simp only: mod_div_equality') | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 196 | also have "\<dots> = a div b + 0" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 197 | by simp | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 198 | finally show ?thesis | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 199 | by (rule add_left_imp_eq) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 200 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 201 | |
| 58911 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 202 | lemma mod_mod_trivial [simp]: | 
| 
2cf595ee508b
proper oriented equivalence of dvd predicate and mod
 haftmann parents: 
58889diff
changeset | 203 | "a mod b mod b = a mod b" | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 204 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 205 | 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 | 206 | by (simp only: mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 207 | also have "\<dots> = a mod b" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 208 | by (simp only: mod_div_equality') | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 209 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 210 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 211 | |
| 58834 | 212 | lemma dvd_mod_imp_dvd: | 
| 213 | assumes "k dvd m mod n" and "k dvd n" | |
| 214 | shows "k dvd m" | |
| 215 | proof - | |
| 216 | from assms have "k dvd (m div n) * n + m mod n" | |
| 217 | by (simp only: dvd_add dvd_mult) | |
| 218 | then show ?thesis by (simp add: mod_div_equality) | |
| 219 | qed | |
| 30078 
beee83623cc9
move lemma dvd_mod_imp_dvd into class semiring_div
 huffman parents: 
30052diff
changeset | 220 | |
| 60758 | 221 | text \<open>Addition respects modular equivalence.\<close> | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 222 | |
| 60867 | 223 | lemma mod_add_left_eq: -- \<open>FIXME reorient\<close> | 
| 224 | "(a + b) mod c = (a mod c + b) mod c" | |
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 225 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 226 | 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 | 227 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 228 | also have "\<dots> = (a mod c + b + a div c * c) mod c" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 229 | by (simp only: ac_simps) | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 230 | also have "\<dots> = (a mod c + b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 231 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 232 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 233 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 234 | |
| 60867 | 235 | lemma mod_add_right_eq: -- \<open>FIXME reorient\<close> | 
| 236 | "(a + b) mod c = (a + b mod c) mod c" | |
| 29403 
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 + 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 | 239 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 240 | also have "\<dots> = (a + b mod c + b div c * c) mod c" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 241 | by (simp only: ac_simps) | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 242 | also have "\<dots> = (a + b mod c) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 243 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 244 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 245 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 246 | |
| 60867 | 247 | lemma mod_add_eq: -- \<open>FIXME reorient\<close> | 
| 248 | "(a + b) mod c = (a mod c + b mod c) mod c" | |
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 249 | 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 | 250 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 251 | lemma mod_add_cong: | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 252 | assumes "a mod c = a' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 253 | assumes "b mod c = b' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 254 | shows "(a + b) mod c = (a' + b') mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 255 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 256 | 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 | 257 | unfolding assms .. | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 258 | thus ?thesis | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 259 | by (simp only: mod_add_eq [symmetric]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 260 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 261 | |
| 60758 | 262 | text \<open>Multiplication respects modular equivalence.\<close> | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 263 | |
| 60867 | 264 | lemma mod_mult_left_eq: -- \<open>FIXME reorient\<close> | 
| 265 | "(a * b) mod c = ((a mod c) * b) mod c" | |
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 266 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 267 | 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 | 268 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 269 | also have "\<dots> = (a mod c * b + a div c * b * c) mod c" | 
| 29667 | 270 | by (simp only: algebra_simps) | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 271 | also have "\<dots> = (a mod c * b) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 272 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 273 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 274 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 275 | |
| 60867 | 276 | lemma mod_mult_right_eq: -- \<open>FIXME reorient\<close> | 
| 277 | "(a * b) mod c = (a * (b mod c)) mod c" | |
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 278 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 279 | 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 | 280 | by (simp only: mod_div_equality) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 281 | also have "\<dots> = (a * (b mod c) + a * (b div c) * c) mod c" | 
| 29667 | 282 | by (simp only: algebra_simps) | 
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 283 | also have "\<dots> = (a * (b mod c)) mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 284 | by (rule mod_mult_self1) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 285 | finally show ?thesis . | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 286 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 287 | |
| 60867 | 288 | lemma mod_mult_eq: -- \<open>FIXME reorient\<close> | 
| 289 | "(a * b) mod c = ((a mod c) * (b mod c)) mod c" | |
| 29403 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 290 | 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 | 291 | |
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 292 | lemma mod_mult_cong: | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 293 | assumes "a mod c = a' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 294 | assumes "b mod c = b' mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 295 | shows "(a * b) mod c = (a' * b') mod c" | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 296 | proof - | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 297 | 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 | 298 | unfolding assms .. | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 299 | thus ?thesis | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 300 | by (simp only: mod_mult_eq [symmetric]) | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 301 | qed | 
| 
fe17df4e4ab3
generalize some div/mod lemmas; remove type-specific proofs
 huffman parents: 
29252diff
changeset | 302 | |
| 60758 | 303 | text \<open>Exponentiation respects modular equivalence.\<close> | 
| 47164 | 304 | |
| 60867 | 305 | lemma power_mod: "(a mod b) ^ n mod b = a ^ n mod b" | 
| 47164 | 306 | apply (induct n, simp_all) | 
| 307 | apply (rule mod_mult_right_eq [THEN trans]) | |
| 308 | apply (simp (no_asm_simp)) | |
| 309 | apply (rule mod_mult_eq [symmetric]) | |
| 310 | done | |
| 311 | ||
| 29404 | 312 | lemma mod_mod_cancel: | 
| 313 | assumes "c dvd b" | |
| 314 | shows "a mod b mod c = a mod c" | |
| 315 | proof - | |
| 60758 | 316 | from \<open>c dvd b\<close> obtain k where "b = c * k" | 
| 29404 | 317 | by (rule dvdE) | 
| 318 | have "a mod b mod c = a mod (c * k) mod c" | |
| 60758 | 319 | by (simp only: \<open>b = c * k\<close>) | 
| 29404 | 320 | also have "\<dots> = (a mod (c * k) + a div (c * k) * k * c) mod c" | 
| 321 | by (simp only: mod_mult_self1) | |
| 322 | also have "\<dots> = (a div (c * k) * (c * k) + a mod (c * k)) mod c" | |
| 58786 | 323 | by (simp only: ac_simps) | 
| 29404 | 324 | also have "\<dots> = a mod c" | 
| 325 | by (simp only: mod_div_equality) | |
| 326 | finally show ?thesis . | |
| 327 | qed | |
| 328 | ||
| 30930 | 329 | lemma div_mult_mult2 [simp]: | 
| 330 | "c \<noteq> 0 \<Longrightarrow> (a * c) div (b * c) = a div b" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 331 | by (drule div_mult_mult1) (simp add: mult.commute) | 
| 30930 | 332 | |
| 333 | lemma div_mult_mult1_if [simp]: | |
| 334 | "(c * a) div (c * b) = (if c = 0 then 0 else a div b)" | |
| 335 | by simp_all | |
| 30476 | 336 | |
| 30930 | 337 | lemma mod_mult_mult1: | 
| 338 | "(c * a) mod (c * b) = c * (a mod b)" | |
| 339 | proof (cases "c = 0") | |
| 340 | case True then show ?thesis by simp | |
| 341 | next | |
| 342 | case False | |
| 343 | from mod_div_equality | |
| 344 | have "((c * a) div (c * b)) * (c * b) + (c * a) mod (c * b) = c * a" . | |
| 345 | with False have "c * ((a div b) * b + a mod b) + (c * a) mod (c * b) | |
| 346 | = c * a + c * (a mod b)" by (simp add: algebra_simps) | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 347 | with mod_div_equality show ?thesis by simp | 
| 30930 | 348 | qed | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 349 | |
| 30930 | 350 | lemma mod_mult_mult2: | 
| 351 | "(a * c) mod (b * c) = (a mod b) * c" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 352 | using mod_mult_mult1 [of c a b] by (simp add: mult.commute) | 
| 30930 | 353 | |
| 47159 | 354 | lemma mult_mod_left: "(a mod b) * c = (a * c) mod (b * c)" | 
| 355 | by (fact mod_mult_mult2 [symmetric]) | |
| 356 | ||
| 357 | lemma mult_mod_right: "c * (a mod b) = (c * a) mod (c * b)" | |
| 358 | by (fact mod_mult_mult1 [symmetric]) | |
| 359 | ||
| 31662 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 360 | 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 | 361 | 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 | 362 | |
| 
57f7ef0dba8e
generalize lemmas dvd_mod and dvd_mod_iff to class semiring_div
 huffman parents: 
31661diff
changeset | 363 | 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 | 364 | 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 | 365 | |
| 31661 
1e252b8b2334
move lemma div_power into semiring_div context; class ring_div inherits from idom
 huffman parents: 
31009diff
changeset | 366 | end | 
| 
1e252b8b2334
move lemma div_power into semiring_div context; class ring_div inherits from idom
 huffman parents: 
31009diff
changeset | 367 | |
| 59833 
ab828c2c5d67
clarified no_zero_devisors: makes only sense in a semiring;
 haftmann parents: 
59816diff
changeset | 368 | class ring_div = comm_ring_1 + semiring_div | 
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 369 | begin | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 370 | |
| 60353 
838025c6e278
implicit partial divison operation in integral domains
 haftmann parents: 
60352diff
changeset | 371 | subclass idom_divide .. | 
| 36634 | 372 | |
| 60758 | 373 | text \<open>Negation respects modular equivalence.\<close> | 
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 374 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 375 | 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 | 376 | proof - | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 377 | 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 | 378 | by (simp only: mod_div_equality) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 379 | also have "\<dots> = (- (a mod b) + - (a div b) * b) mod b" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 380 | by (simp add: ac_simps) | 
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 381 | 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 | 382 | by (rule mod_mult_self1) | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 383 | finally show ?thesis . | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 384 | qed | 
| 
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_cong: | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 387 | 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 | 388 | 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 | 389 | proof - | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 390 | 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 | 391 | unfolding assms .. | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 392 | thus ?thesis | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 393 | 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 | 394 | qed | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 395 | |
| 60758 | 396 | text \<open>Subtraction respects modular equivalence.\<close> | 
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 397 | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 398 | lemma mod_diff_left_eq: | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 399 | "(a - b) mod c = (a mod c - b) mod c" | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 400 | using mod_add_cong [of a c "a mod c" "- b" "- b"] by simp | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 401 | |
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 402 | lemma mod_diff_right_eq: | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 403 | "(a - b) mod c = (a - b mod c) mod c" | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 404 | using mod_add_cong [of a c a "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 405 | |
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 406 | lemma mod_diff_eq: | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 407 | "(a - b) mod c = (a mod c - b mod c) mod c" | 
| 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 408 | using mod_add_cong [of a c "a mod c" "- b" "- (b mod c)"] mod_minus_cong [of "b mod c" c b] by simp | 
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 409 | |
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 410 | lemma mod_diff_cong: | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 411 | 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 | 412 | 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 | 413 | shows "(a - b) mod c = (a' - b') mod c" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 414 | using assms mod_add_cong [of a c a' "- b" "- b'"] mod_minus_cong [of b c "b'"] by simp | 
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 415 | |
| 30180 | 416 | lemma dvd_neg_div: "y dvd x \<Longrightarrow> -x div y = - (x div y)" | 
| 417 | apply (case_tac "y = 0") apply simp | |
| 418 | apply (auto simp add: dvd_def) | |
| 419 | apply (subgoal_tac "-(y * k) = y * - k") | |
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
55440diff
changeset | 420 | apply (simp only:) | 
| 30180 | 421 | apply (erule div_mult_self1_is_id) | 
| 422 | apply simp | |
| 423 | done | |
| 424 | ||
| 425 | lemma dvd_div_neg: "y dvd x \<Longrightarrow> x div -y = - (x div y)" | |
| 426 | apply (case_tac "y = 0") apply simp | |
| 427 | apply (auto simp add: dvd_def) | |
| 428 | apply (subgoal_tac "y * k = -y * -k") | |
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
55440diff
changeset | 429 | apply (erule ssubst, rule div_mult_self1_is_id) | 
| 30180 | 430 | apply simp | 
| 431 | apply simp | |
| 432 | done | |
| 433 | ||
| 60867 | 434 | lemma div_diff [simp]: | 
| 435 | "z dvd x \<Longrightarrow> z dvd y \<Longrightarrow> (x - y) div z = x div z - y div z" | |
| 436 | using div_add [of _ _ "- y"] by (simp add: dvd_neg_div) | |
| 59380 | 437 | |
| 47159 | 438 | lemma div_minus_minus [simp]: "(-a) div (-b) = a div b" | 
| 439 | using div_mult_mult1 [of "- 1" a b] | |
| 440 | unfolding neg_equal_0_iff_equal by simp | |
| 441 | ||
| 442 | lemma mod_minus_minus [simp]: "(-a) mod (-b) = - (a mod b)" | |
| 443 | using mod_mult_mult1 [of "- 1" a b] by simp | |
| 444 | ||
| 445 | lemma div_minus_right: "a div (-b) = (-a) div b" | |
| 446 | using div_minus_minus [of "-a" b] by simp | |
| 447 | ||
| 448 | lemma mod_minus_right: "a mod (-b) = - ((-a) mod b)" | |
| 449 | using mod_minus_minus [of "-a" b] by simp | |
| 450 | ||
| 47160 | 451 | lemma div_minus1_right [simp]: "a div (-1) = -a" | 
| 452 | using div_minus_right [of a 1] by simp | |
| 453 | ||
| 454 | lemma mod_minus1_right [simp]: "a mod (-1) = 0" | |
| 455 | using mod_minus_right [of a 1] by simp | |
| 456 | ||
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 457 | lemma minus_mod_self2 [simp]: | 
| 54221 | 458 | "(a - b) mod b = a mod b" | 
| 459 | by (simp add: mod_diff_right_eq) | |
| 460 | ||
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 461 | lemma minus_mod_self1 [simp]: | 
| 54221 | 462 | "(b - a) mod b = - a mod b" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
54227diff
changeset | 463 | using mod_add_self2 [of "- a" b] by simp | 
| 54221 | 464 | |
| 29405 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 465 | end | 
| 
98ab21b14f09
add class ring_div; generalize mod/diff/minus proofs for class ring_div
 huffman parents: 
29404diff
changeset | 466 | |
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 467 | |
| 60758 | 468 | subsubsection \<open>Parity and division\<close> | 
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 469 | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 470 | class semiring_div_parity = semiring_div + comm_semiring_1_cancel + numeral + | 
| 54226 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 471 | assumes parity: "a mod 2 = 0 \<or> a mod 2 = 1" | 
| 58786 | 472 | assumes one_mod_two_eq_one [simp]: "1 mod 2 = 1" | 
| 58710 
7216a10d69ba
augmented and tuned facts on even/odd and division
 haftmann parents: 
58646diff
changeset | 473 | assumes zero_not_eq_two: "0 \<noteq> 2" | 
| 54226 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 474 | begin | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 475 | |
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 476 | lemma parity_cases [case_names even odd]: | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 477 | assumes "a mod 2 = 0 \<Longrightarrow> P" | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 478 | assumes "a mod 2 = 1 \<Longrightarrow> P" | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 479 | shows P | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 480 | using assms parity by blast | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 481 | |
| 58786 | 482 | lemma one_div_two_eq_zero [simp]: | 
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 483 | "1 div 2 = 0" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 484 | proof (cases "2 = 0") | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 485 | case True then show ?thesis by simp | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 486 | next | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 487 | case False | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 488 | from mod_div_equality have "1 div 2 * 2 + 1 mod 2 = 1" . | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 489 | with one_mod_two_eq_one have "1 div 2 * 2 + 1 = 1" by simp | 
| 58953 | 490 | then have "1 div 2 * 2 = 0" by (simp add: ac_simps add_left_imp_eq del: mult_eq_0_iff) | 
| 491 | then have "1 div 2 = 0 \<or> 2 = 0" by simp | |
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 492 | with False show ?thesis by auto | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 493 | qed | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 494 | |
| 58786 | 495 | lemma not_mod_2_eq_0_eq_1 [simp]: | 
| 496 | "a mod 2 \<noteq> 0 \<longleftrightarrow> a mod 2 = 1" | |
| 497 | by (cases a rule: parity_cases) simp_all | |
| 498 | ||
| 499 | lemma not_mod_2_eq_1_eq_0 [simp]: | |
| 500 | "a mod 2 \<noteq> 1 \<longleftrightarrow> a mod 2 = 0" | |
| 501 | by (cases a rule: parity_cases) simp_all | |
| 502 | ||
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 503 | subclass semiring_parity | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 504 | proof (unfold_locales, unfold dvd_eq_mod_eq_0 not_mod_2_eq_0_eq_1) | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 505 | show "1 mod 2 = 1" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 506 | by (fact one_mod_two_eq_one) | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 507 | next | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 508 | fix a b | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 509 | assume "a mod 2 = 1" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 510 | moreover assume "b mod 2 = 1" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 511 | ultimately show "(a + b) mod 2 = 0" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 512 | using mod_add_eq [of a b 2] by simp | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 513 | next | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 514 | fix a b | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 515 | assume "(a * b) mod 2 = 0" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 516 | then have "(a mod 2) * (b mod 2) = 0" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 517 | by (cases "a mod 2 = 0") (simp_all add: mod_mult_eq [of a b 2]) | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 518 | then show "a mod 2 = 0 \<or> b mod 2 = 0" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 519 | by (rule divisors_zero) | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 520 | next | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 521 | fix a | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 522 | assume "a mod 2 = 1" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 523 | then have "a = a div 2 * 2 + 1" using mod_div_equality [of a 2] by simp | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 524 | then show "\<exists>b. a = b + 1" .. | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 525 | qed | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 526 | |
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 527 | lemma even_iff_mod_2_eq_zero: | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 528 | "even a \<longleftrightarrow> a mod 2 = 0" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 529 | by (fact dvd_eq_mod_eq_0) | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 530 | |
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 531 | lemma even_succ_div_two [simp]: | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 532 | "even a \<Longrightarrow> (a + 1) div 2 = a div 2" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 533 | by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero) | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 534 | |
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 535 | lemma odd_succ_div_two [simp]: | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 536 | "odd a \<Longrightarrow> (a + 1) div 2 = a div 2 + 1" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 537 | by (auto elim!: oddE simp add: zero_not_eq_two [symmetric] add.assoc) | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 538 | |
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 539 | lemma even_two_times_div_two: | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 540 | "even a \<Longrightarrow> 2 * (a div 2) = a" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 541 | by (fact dvd_mult_div_cancel) | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 542 | |
| 58834 | 543 | lemma odd_two_times_div_two_succ [simp]: | 
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 544 | "odd a \<Longrightarrow> 2 * (a div 2) + 1 = a" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 545 | using mod_div_equality2 [of 2 a] by (simp add: even_iff_mod_2_eq_zero) | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 546 | |
| 54226 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 547 | end | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 548 | |
| 25942 | 549 | |
| 60758 | 550 | subsection \<open>Generic numeral division with a pragmatic type class\<close> | 
| 551 | ||
| 552 | text \<open> | |
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 553 | The following type class contains everything necessary to formulate | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 554 | a division algorithm in ring structures with numerals, restricted | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 555 | to its positive segments. This is its primary motiviation, and it | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 556 | could surely be formulated using a more fine-grained, more algebraic | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 557 | and less technical class hierarchy. | 
| 60758 | 558 | \<close> | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 559 | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 560 | class semiring_numeral_div = semiring_div + comm_semiring_1_cancel + linordered_semidom + | 
| 59816 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 561 | assumes div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0" | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 562 | and mod_less: " 0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a mod b = a" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 563 | and div_positive: "0 < b \<Longrightarrow> b \<le> a \<Longrightarrow> a div b > 0" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 564 | and mod_less_eq_dividend: "0 \<le> a \<Longrightarrow> a mod b \<le> a" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 565 | and pos_mod_bound: "0 < b \<Longrightarrow> a mod b < b" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 566 | and pos_mod_sign: "0 < b \<Longrightarrow> 0 \<le> a mod b" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 567 | and mod_mult2_eq: "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 568 | and div_mult2_eq: "0 \<le> c \<Longrightarrow> a div (b * c) = a div b div c" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 569 | assumes discrete: "a < b \<longleftrightarrow> a + 1 \<le> b" | 
| 61275 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 570 | fixes divmod :: "num \<Rightarrow> num \<Rightarrow> 'a \<times> 'a" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 571 | and divmod_step :: "num \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<times> 'a" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 572 | assumes divmod_def: "divmod m n = (numeral m div numeral n, numeral m mod numeral n)" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 573 | and divmod_step_def: "divmod_step l qr = (let (q, r) = qr | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 574 | in if r \<ge> numeral l then (2 * q + 1, r - numeral l) | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 575 | else (2 * q, r))" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 576 | -- \<open>These are conceptually definitions but force generated code | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 577 | to be monomorphic wrt. particular instances of this class which | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 578 | yields a significant speedup.\<close> | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 579 | |
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 580 | begin | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 581 | |
| 59816 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 582 | lemma mult_div_cancel: | 
| 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 583 | "b * (a div b) = a - a mod b" | 
| 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 584 | proof - | 
| 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 585 | have "b * (a div b) + a mod b = a" | 
| 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 586 | using mod_div_equality [of a b] by (simp add: ac_simps) | 
| 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 587 | then have "b * (a div b) + a mod b - a mod b = a - a mod b" | 
| 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 588 | by simp | 
| 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 589 | then show ?thesis | 
| 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 590 | by simp | 
| 
034b13f4efae
distributivity of partial minus establishes desired properties of dvd in semirings
 haftmann parents: 
59807diff
changeset | 591 | qed | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 592 | |
| 54226 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 593 | subclass semiring_div_parity | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 594 | proof | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 595 | fix a | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 596 | show "a mod 2 = 0 \<or> a mod 2 = 1" | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 597 | proof (rule ccontr) | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 598 | assume "\<not> (a mod 2 = 0 \<or> a mod 2 = 1)" | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 599 | then have "a mod 2 \<noteq> 0" and "a mod 2 \<noteq> 1" by simp_all | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 600 | have "0 < 2" by simp | 
| 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 601 | with pos_mod_bound pos_mod_sign have "0 \<le> a mod 2" "a mod 2 < 2" by simp_all | 
| 60758 | 602 | with \<open>a mod 2 \<noteq> 0\<close> have "0 < a mod 2" by simp | 
| 54226 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 603 | with discrete have "1 \<le> a mod 2" by simp | 
| 60758 | 604 | with \<open>a mod 2 \<noteq> 1\<close> have "1 < a mod 2" by simp | 
| 54226 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 605 | with discrete have "2 \<le> a mod 2" by simp | 
| 60758 | 606 | with \<open>a mod 2 < 2\<close> show False by simp | 
| 54226 
e3df2a4e02fc
explicit type class for modelling even/odd parity
 haftmann parents: 
54221diff
changeset | 607 | qed | 
| 58646 
cd63a4b12a33
specialized specification: avoid trivial instances
 haftmann parents: 
58511diff
changeset | 608 | next | 
| 
cd63a4b12a33
specialized specification: avoid trivial instances
 haftmann parents: 
58511diff
changeset | 609 | show "1 mod 2 = 1" | 
| 
cd63a4b12a33
specialized specification: avoid trivial instances
 haftmann parents: 
58511diff
changeset | 610 | by (rule mod_less) simp_all | 
| 58710 
7216a10d69ba
augmented and tuned facts on even/odd and division
 haftmann parents: 
58646diff
changeset | 611 | next | 
| 
7216a10d69ba
augmented and tuned facts on even/odd and division
 haftmann parents: 
58646diff
changeset | 612 | show "0 \<noteq> 2" | 
| 
7216a10d69ba
augmented and tuned facts on even/odd and division
 haftmann parents: 
58646diff
changeset | 613 | by simp | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 614 | qed | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 615 | |
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 616 | lemma divmod_digit_1: | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 617 | assumes "0 \<le> a" "0 < b" and "b \<le> a mod (2 * b)" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 618 | shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P") | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 619 | and "a mod (2 * b) - b = a mod b" (is "?Q") | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 620 | proof - | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 621 | from assms mod_less_eq_dividend [of a "2 * b"] have "b \<le> a" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 622 | by (auto intro: trans) | 
| 60758 | 623 | with \<open>0 < b\<close> have "0 < a div b" by (auto intro: div_positive) | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 624 | then have [simp]: "1 \<le> a div b" by (simp add: discrete) | 
| 60758 | 625 | with \<open>0 < b\<close> have mod_less: "a mod b < b" by (simp add: pos_mod_bound) | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 626 | def w \<equiv> "a div b mod 2" with parity have w_exhaust: "w = 0 \<or> w = 1" by auto | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 627 | have mod_w: "a mod (2 * b) = a mod b + b * w" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 628 | by (simp add: w_def mod_mult2_eq ac_simps) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 629 | from assms w_exhaust have "w = 1" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 630 | by (auto simp add: mod_w) (insert mod_less, auto) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 631 | with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 632 | have "2 * (a div (2 * b)) = a div b - w" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 633 | by (simp add: w_def div_mult2_eq mult_div_cancel ac_simps) | 
| 60758 | 634 | with \<open>w = 1\<close> have div: "2 * (a div (2 * b)) = a div b - 1" by simp | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 635 | then show ?P and ?Q | 
| 60867 | 636 | by (simp_all add: div mod add_implies_diff [symmetric]) | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 637 | qed | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 638 | |
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 639 | lemma divmod_digit_0: | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 640 | assumes "0 < b" and "a mod (2 * b) < b" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 641 | shows "2 * (a div (2 * b)) = a div b" (is "?P") | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 642 | and "a mod (2 * b) = a mod b" (is "?Q") | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 643 | proof - | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 644 | def w \<equiv> "a div b mod 2" with parity have w_exhaust: "w = 0 \<or> w = 1" by auto | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 645 | have mod_w: "a mod (2 * b) = a mod b + b * w" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 646 | by (simp add: w_def mod_mult2_eq ac_simps) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 647 | moreover have "b \<le> a mod b + b" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 648 | proof - | 
| 60758 | 649 | from \<open>0 < b\<close> pos_mod_sign have "0 \<le> a mod b" by blast | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 650 | then have "0 + b \<le> a mod b + b" by (rule add_right_mono) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 651 | then show ?thesis by simp | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 652 | qed | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 653 | moreover note assms w_exhaust | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 654 | ultimately have "w = 0" by auto | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 655 | with mod_w have mod: "a mod (2 * b) = a mod b" by simp | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 656 | have "2 * (a div (2 * b)) = a div b - w" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 657 | by (simp add: w_def div_mult2_eq mult_div_cancel ac_simps) | 
| 60758 | 658 | with \<open>w = 0\<close> have div: "2 * (a div (2 * b)) = a div b" by simp | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 659 | then show ?P and ?Q | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 660 | by (simp_all add: div mod) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 661 | qed | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 662 | |
| 60867 | 663 | lemma fst_divmod: | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 664 | "fst (divmod m n) = numeral m div numeral n" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 665 | by (simp add: divmod_def) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 666 | |
| 60867 | 667 | lemma snd_divmod: | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 668 | "snd (divmod m n) = numeral m mod numeral n" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 669 | by (simp add: divmod_def) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 670 | |
| 60758 | 671 | text \<open> | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 672 | This is a formulation of one step (referring to one digit position) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 673 | in school-method division: compare the dividend at the current | 
| 53070 | 674 | digit position with the remainder from previous division steps | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 675 | and evaluate accordingly. | 
| 60758 | 676 | \<close> | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 677 | |
| 61275 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 678 | lemma divmod_step_eq [simp]: | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 679 | "divmod_step l (q, r) = (if numeral l \<le> r | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 680 | then (2 * q + 1, r - numeral l) else (2 * q, r))" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 681 | by (simp add: divmod_step_def) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 682 | |
| 60758 | 683 | text \<open> | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 684 | This is a formulation of school-method division. | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 685 | If the divisor is smaller than the dividend, terminate. | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 686 | If not, shift the dividend to the right until termination | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 687 | occurs and then reiterate single division steps in the | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 688 | opposite direction. | 
| 60758 | 689 | \<close> | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 690 | |
| 60867 | 691 | lemma divmod_divmod_step: | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 692 | "divmod m n = (if m < n then (0, numeral m) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 693 | else divmod_step n (divmod m (Num.Bit0 n)))" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 694 | proof (cases "m < n") | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 695 | case True then have "numeral m < numeral n" by simp | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 696 | then show ?thesis | 
| 60867 | 697 | by (simp add: prod_eq_iff div_less mod_less fst_divmod snd_divmod) | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 698 | next | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 699 | case False | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 700 | have "divmod m n = | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 701 | divmod_step n (numeral m div (2 * numeral n), | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 702 | numeral m mod (2 * numeral n))" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 703 | proof (cases "numeral n \<le> numeral m mod (2 * numeral n)") | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 704 | case True | 
| 60867 | 705 | with divmod_step_eq | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 706 | have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) = | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 707 | (2 * (numeral m div (2 * numeral n)) + 1, numeral m mod (2 * numeral n) - numeral n)" | 
| 60867 | 708 | by simp | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 709 | moreover from True divmod_digit_1 [of "numeral m" "numeral n"] | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 710 | have "2 * (numeral m div (2 * numeral n)) + 1 = numeral m div numeral n" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 711 | and "numeral m mod (2 * numeral n) - numeral n = numeral m mod numeral n" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 712 | by simp_all | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 713 | ultimately show ?thesis by (simp only: divmod_def) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 714 | next | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 715 | case False then have *: "numeral m mod (2 * numeral n) < numeral n" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 716 | by (simp add: not_le) | 
| 60867 | 717 | with divmod_step_eq | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 718 | have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) = | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 719 | (2 * (numeral m div (2 * numeral n)), numeral m mod (2 * numeral n))" | 
| 60867 | 720 | by auto | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 721 | moreover from * divmod_digit_0 [of "numeral n" "numeral m"] | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 722 | have "2 * (numeral m div (2 * numeral n)) = numeral m div numeral n" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 723 | and "numeral m mod (2 * numeral n) = numeral m mod numeral n" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 724 | by (simp_all only: zero_less_numeral) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 725 | ultimately show ?thesis by (simp only: divmod_def) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 726 | qed | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 727 | then have "divmod m n = | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 728 | divmod_step n (numeral m div numeral (Num.Bit0 n), | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 729 | numeral m mod numeral (Num.Bit0 n))" | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 730 | by (simp only: numeral.simps distrib mult_1) | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 731 | then have "divmod m n = divmod_step n (divmod m (Num.Bit0 n))" | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 732 | by (simp add: divmod_def) | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 733 | with False show ?thesis by simp | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 734 | qed | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 735 | |
| 61201 | 736 | text \<open>The division rewrite proper -- first, trivial results involving @{text 1}\<close>
 | 
| 60867 | 737 | |
| 61275 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 738 | lemma divmod_trivial [simp]: | 
| 60867 | 739 | "divmod Num.One Num.One = (numeral Num.One, 0)" | 
| 740 | "divmod (Num.Bit0 m) Num.One = (numeral (Num.Bit0 m), 0)" | |
| 741 | "divmod (Num.Bit1 m) Num.One = (numeral (Num.Bit1 m), 0)" | |
| 742 | "divmod num.One (num.Bit0 n) = (0, Numeral1)" | |
| 743 | "divmod num.One (num.Bit1 n) = (0, Numeral1)" | |
| 744 | using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def) | |
| 745 | ||
| 746 | text \<open>Division by an even number is a right-shift\<close> | |
| 58953 | 747 | |
| 61275 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 748 | lemma divmod_cancel [simp]: | 
| 53069 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 749 | "divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r))" (is ?P) | 
| 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 750 | "divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r + 1))" (is ?Q) | 
| 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 751 | proof - | 
| 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 752 | have *: "\<And>q. numeral (Num.Bit0 q) = 2 * numeral q" | 
| 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 753 | "\<And>q. numeral (Num.Bit1 q) = 2 * numeral q + 1" | 
| 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 754 | by (simp_all only: numeral_mult numeral.simps distrib) simp_all | 
| 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 755 | have "1 div 2 = 0" "1 mod 2 = 1" by (auto intro: div_less mod_less) | 
| 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 756 | then show ?P and ?Q | 
| 60867 | 757 | by (simp_all add: fst_divmod snd_divmod prod_eq_iff split_def * [of m] * [of n] mod_mult_mult1 | 
| 758 | div_mult2_eq [of _ _ 2] mod_mult2_eq [of _ _ 2] | |
| 759 | add.commute del: numeral_times_numeral) | |
| 58953 | 760 | qed | 
| 761 | ||
| 60867 | 762 | text \<open>The really hard work\<close> | 
| 763 | ||
| 61275 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 764 | lemma divmod_steps [simp]: | 
| 60867 | 765 | "divmod (num.Bit0 m) (num.Bit1 n) = | 
| 766 | (if m \<le> n then (0, numeral (num.Bit0 m)) | |
| 767 | else divmod_step (num.Bit1 n) | |
| 768 | (divmod (num.Bit0 m) | |
| 769 | (num.Bit0 (num.Bit1 n))))" | |
| 770 | "divmod (num.Bit1 m) (num.Bit1 n) = | |
| 771 | (if m < n then (0, numeral (num.Bit1 m)) | |
| 772 | else divmod_step (num.Bit1 n) | |
| 773 | (divmod (num.Bit1 m) | |
| 774 | (num.Bit0 (num.Bit1 n))))" | |
| 775 | by (simp_all add: divmod_divmod_step) | |
| 776 | ||
| 61275 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 777 | lemmas divmod_algorithm_code = divmod_step_eq divmod_trivial divmod_cancel divmod_steps | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 778 | |
| 60758 | 779 | text \<open>Special case: divisibility\<close> | 
| 58953 | 780 | |
| 781 | definition divides_aux :: "'a \<times> 'a \<Rightarrow> bool" | |
| 782 | where | |
| 783 | "divides_aux qr \<longleftrightarrow> snd qr = 0" | |
| 784 | ||
| 785 | lemma divides_aux_eq [simp]: | |
| 786 | "divides_aux (q, r) \<longleftrightarrow> r = 0" | |
| 787 | by (simp add: divides_aux_def) | |
| 788 | ||
| 789 | lemma dvd_numeral_simp [simp]: | |
| 790 | "numeral m dvd numeral n \<longleftrightarrow> divides_aux (divmod n m)" | |
| 791 | by (simp add: divmod_def mod_eq_0_iff_dvd) | |
| 53069 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 792 | |
| 60867 | 793 | text \<open>Generic computation of quotient and remainder\<close> | 
| 794 | ||
| 795 | lemma numeral_div_numeral [simp]: | |
| 796 | "numeral k div numeral l = fst (divmod k l)" | |
| 797 | by (simp add: fst_divmod) | |
| 798 | ||
| 799 | lemma numeral_mod_numeral [simp]: | |
| 800 | "numeral k mod numeral l = snd (divmod k l)" | |
| 801 | by (simp add: snd_divmod) | |
| 802 | ||
| 803 | lemma one_div_numeral [simp]: | |
| 804 | "1 div numeral n = fst (divmod num.One n)" | |
| 805 | by (simp add: fst_divmod) | |
| 806 | ||
| 807 | lemma one_mod_numeral [simp]: | |
| 808 | "1 mod numeral n = snd (divmod num.One n)" | |
| 809 | by (simp add: snd_divmod) | |
| 810 | ||
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 811 | end | 
| 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 812 | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 813 | |
| 60758 | 814 | subsection \<open>Division on @{typ nat}\<close>
 | 
| 815 | ||
| 816 | text \<open> | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60353diff
changeset | 817 |   We define @{const divide} and @{const mod} on @{typ nat} by means
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 818 | of a characteristic relation with two input arguments | 
| 61076 | 819 |   @{term "m::nat"}, @{term "n::nat"} and two output arguments
 | 
| 820 |   @{term "q::nat"}(uotient) and @{term "r::nat"}(emainder).
 | |
| 60758 | 821 | \<close> | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 822 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 823 | 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 | 824 | "divmod_nat_rel m n qr \<longleftrightarrow> | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 825 | m = fst qr * n + snd qr \<and> | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 826 | (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 | 827 | |
| 60758 | 828 | text \<open>@{const divmod_nat_rel} is total:\<close>
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 829 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 830 | lemma divmod_nat_rel_ex: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 831 | 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 | 832 | proof (cases "n = 0") | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 833 | case True with that show thesis | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 834 | by (auto simp add: divmod_nat_rel_def) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 835 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 836 | case False | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 837 | 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 | 838 | proof (induct m) | 
| 60758 | 839 | case 0 with \<open>n \<noteq> 0\<close> | 
| 61076 | 840 | have "(0::nat) = 0 * n + 0 \<and> 0 < n" by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 841 | then show ?case by blast | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 842 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 843 | case (Suc m) then obtain q' r' | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 844 | 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 | 845 | then show ?case proof (cases "Suc r' < n") | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 846 | case True | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 847 | 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 | 848 | with True show ?thesis by blast | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 849 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 850 | 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 | 851 | 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 | 852 | ultimately have "n = Suc r'" by auto | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 853 | with m have "Suc m = Suc q' * n + 0" by simp | 
| 60758 | 854 | with \<open>n \<noteq> 0\<close> show ?thesis by blast | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 855 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 856 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 857 | with that show thesis | 
| 60758 | 858 | using \<open>n \<noteq> 0\<close> by (auto simp add: divmod_nat_rel_def) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 859 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 860 | |
| 60758 | 861 | text \<open>@{const divmod_nat_rel} is injective:\<close>
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 862 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 863 | lemma divmod_nat_rel_unique: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 864 | assumes "divmod_nat_rel m n qr" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 865 | and "divmod_nat_rel m n qr'" | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 866 | shows "qr = qr'" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 867 | proof (cases "n = 0") | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 868 | case True with assms show ?thesis | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 869 | by (cases qr, cases qr') | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 870 | (simp add: divmod_nat_rel_def) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 871 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 872 | case False | 
| 61076 | 873 | have aux: "\<And>q r q' r'. q' * n + r' = q * n + r \<Longrightarrow> r < n \<Longrightarrow> q' \<le> (q::nat)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 874 | apply (rule leI) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 875 | apply (subst less_iff_Suc_add) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 876 | apply (auto simp add: add_mult_distrib) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 877 | done | 
| 60758 | 878 | from \<open>n \<noteq> 0\<close> assms have *: "fst qr = fst qr'" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 879 | by (auto simp add: divmod_nat_rel_def intro: order_antisym dest: aux sym) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53199diff
changeset | 880 | with assms have "snd qr = snd qr'" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 881 | by (simp add: divmod_nat_rel_def) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53199diff
changeset | 882 | with * show ?thesis by (cases qr, cases qr') simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 883 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 884 | |
| 60758 | 885 | text \<open> | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 886 | We instantiate divisibility on the natural numbers by | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 887 |   means of @{const divmod_nat_rel}:
 | 
| 60758 | 888 | \<close> | 
| 25942 | 889 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 890 | definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" where | 
| 37767 | 891 | "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 | 892 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 893 | lemma divmod_nat_rel_divmod_nat: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 894 | "divmod_nat_rel m n (divmod_nat m n)" | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 895 | proof - | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 896 | from divmod_nat_rel_ex | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 897 | 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 | 898 | then show ?thesis | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 899 | 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 | 900 | qed | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 901 | |
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 902 | lemma divmod_nat_unique: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 903 | assumes "divmod_nat_rel m n qr" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 904 | shows "divmod_nat m n = qr" | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 905 | 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 | 906 | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60353diff
changeset | 907 | instantiation nat :: semiring_div | 
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 908 | begin | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 909 | |
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 910 | definition divide_nat where | 
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60353diff
changeset | 911 | div_nat_def: "m div n = fst (divmod_nat m n)" | 
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 912 | |
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 913 | definition mod_nat where | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 914 | "m mod n = snd (divmod_nat m n)" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 915 | |
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 916 | lemma fst_divmod_nat [simp]: | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 917 | "fst (divmod_nat m n) = m div n" | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 918 | by (simp add: div_nat_def) | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 919 | |
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 920 | lemma snd_divmod_nat [simp]: | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 921 | "snd (divmod_nat m n) = m mod n" | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 922 | by (simp add: mod_nat_def) | 
| 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 923 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 924 | lemma divmod_nat_div_mod: | 
| 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 925 | "divmod_nat m n = (m div n, m mod n)" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 926 | by (simp add: prod_eq_iff) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 927 | |
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 928 | lemma div_nat_unique: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 929 | assumes "divmod_nat_rel m n (q, r)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 930 | shows "m div n = q" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 931 | using assms by (auto dest!: divmod_nat_unique simp add: prod_eq_iff) | 
| 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 932 | |
| 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 933 | lemma mod_nat_unique: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 934 | assumes "divmod_nat_rel m n (q, r)" | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 935 | shows "m mod n = r" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 936 | using assms by (auto dest!: divmod_nat_unique simp add: prod_eq_iff) | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25162diff
changeset | 937 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 938 | lemma divmod_nat_rel: "divmod_nat_rel m n (m div n, m mod n)" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 939 | using divmod_nat_rel_divmod_nat by (simp add: divmod_nat_div_mod) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 940 | |
| 47136 | 941 | lemma divmod_nat_zero: "divmod_nat m 0 = (0, m)" | 
| 942 | by (simp add: divmod_nat_unique divmod_nat_rel_def) | |
| 943 | ||
| 944 | lemma divmod_nat_zero_left: "divmod_nat 0 n = (0, 0)" | |
| 945 | by (simp add: divmod_nat_unique divmod_nat_rel_def) | |
| 25942 | 946 | |
| 47137 | 947 | lemma divmod_nat_base: "m < n \<Longrightarrow> divmod_nat m n = (0, m)" | 
| 948 | by (simp add: divmod_nat_unique divmod_nat_rel_def) | |
| 25942 | 949 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 950 | lemma divmod_nat_step: | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 951 | assumes "0 < n" and "n \<le> m" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 952 | shows "divmod_nat m n = (Suc ((m - n) div n), (m - n) mod n)" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 953 | proof (rule divmod_nat_unique) | 
| 47134 | 954 | have "divmod_nat_rel (m - n) n ((m - n) div n, (m - n) mod n)" | 
| 955 | by (rule divmod_nat_rel) | |
| 956 | thus "divmod_nat_rel m n (Suc ((m - n) div n), (m - n) mod n)" | |
| 957 | unfolding divmod_nat_rel_def using assms by auto | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 958 | qed | 
| 25942 | 959 | |
| 60758 | 960 | text \<open>The ''recursion'' equations for @{const divide} and @{const mod}\<close>
 | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 961 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 962 | lemma div_less [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 963 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 964 | assumes "m < n" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 965 | shows "m div n = 0" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 966 | using assms divmod_nat_base by (simp add: prod_eq_iff) | 
| 25942 | 967 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 968 | lemma le_div_geq: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 969 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 970 | assumes "0 < n" and "n \<le> m" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 971 | shows "m div n = Suc ((m - n) div n)" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 972 | using assms divmod_nat_step by (simp add: prod_eq_iff) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 973 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 974 | lemma mod_less [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 975 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 976 | assumes "m < n" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 977 | shows "m mod n = m" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 978 | using assms divmod_nat_base by (simp add: prod_eq_iff) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 979 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 980 | lemma le_mod_geq: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 981 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 982 | assumes "n \<le> m" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 983 | shows "m mod n = (m - n) mod n" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 984 | using assms divmod_nat_step by (cases "n = 0") (simp_all add: prod_eq_iff) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 985 | |
| 47136 | 986 | instance proof | 
| 987 | fix m n :: nat | |
| 988 | show "m div n * n + m mod n = m" | |
| 989 | using divmod_nat_rel [of m n] by (simp add: divmod_nat_rel_def) | |
| 990 | next | |
| 991 | fix m n q :: nat | |
| 992 | assume "n \<noteq> 0" | |
| 993 | then show "(q + m * n) div n = m + q div n" | |
| 994 | by (induct m) (simp_all add: le_div_geq) | |
| 995 | next | |
| 996 | fix m n q :: nat | |
| 997 | assume "m \<noteq> 0" | |
| 998 | hence "\<And>a b. divmod_nat_rel n q (a, b) \<Longrightarrow> divmod_nat_rel (m * n) (m * q) (a, m * b)" | |
| 999 | unfolding divmod_nat_rel_def | |
| 1000 | by (auto split: split_if_asm, simp_all add: algebra_simps) | |
| 1001 | moreover from divmod_nat_rel have "divmod_nat_rel n q (n div q, n mod q)" . | |
| 1002 | ultimately have "divmod_nat_rel (m * n) (m * q) (n div q, m * (n mod q))" . | |
| 1003 | thus "(m * n) div (m * q) = n div q" by (rule div_nat_unique) | |
| 1004 | next | |
| 1005 | fix n :: nat show "n div 0 = 0" | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 1006 | by (simp add: div_nat_def divmod_nat_zero) | 
| 47136 | 1007 | next | 
| 1008 | fix n :: nat show "0 div n = 0" | |
| 1009 | by (simp add: div_nat_def divmod_nat_zero_left) | |
| 25942 | 1010 | qed | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1011 | |
| 25942 | 1012 | end | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1013 | |
| 60685 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1014 | instantiation nat :: normalization_semidom | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1015 | begin | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1016 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1017 | definition normalize_nat | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1018 | where [simp]: "normalize = (id :: nat \<Rightarrow> nat)" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1019 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1020 | definition unit_factor_nat | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1021 | where "unit_factor n = (if n = 0 then 0 else 1 :: nat)" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1022 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1023 | lemma unit_factor_simps [simp]: | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1024 | "unit_factor 0 = (0::nat)" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1025 | "unit_factor (Suc n) = 1" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1026 | by (simp_all add: unit_factor_nat_def) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1027 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1028 | instance | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1029 | by standard (simp_all add: unit_factor_nat_def) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1030 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1031 | end | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1032 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1033 | 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 | 1034 | let (q, r) = divmod_nat (m - n) n in (Suc q, r))" | 
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55172diff
changeset | 1035 | by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1036 | |
| 60758 | 1037 | text \<open>Simproc for cancelling @{const divide} and @{const mod}\<close>
 | 
| 25942 | 1038 | |
| 51299 
30b014246e21
proper place for cancel_div_mod.ML (see also ee729dbd1b7f and ec7f10155389);
 wenzelm parents: 
51173diff
changeset | 1039 | ML_file "~~/src/Provers/Arith/cancel_div_mod.ML" | 
| 
30b014246e21
proper place for cancel_div_mod.ML (see also ee729dbd1b7f and ec7f10155389);
 wenzelm parents: 
51173diff
changeset | 1040 | |
| 60758 | 1041 | ML \<open> | 
| 43594 | 1042 | structure Cancel_Div_Mod_Nat = Cancel_Div_Mod | 
| 41550 | 1043 | ( | 
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1044 |   val div_name = @{const_name divide};
 | 
| 30934 | 1045 |   val mod_name = @{const_name mod};
 | 
| 1046 | val mk_binop = HOLogic.mk_binop; | |
| 48561 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1047 |   val mk_plus = HOLogic.mk_binop @{const_name Groups.plus};
 | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1048 |   val dest_plus = HOLogic.dest_bin @{const_name Groups.plus} HOLogic.natT;
 | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1049 | fun mk_sum [] = HOLogic.zero | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1050 | | mk_sum [t] = t | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1051 | | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1052 | fun dest_sum tm = | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1053 | if HOLogic.is_zero tm then [] | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1054 | else | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1055 | (case try HOLogic.dest_Suc tm of | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1056 | SOME t => HOLogic.Suc_zero :: dest_sum t | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1057 | | NONE => | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1058 | (case try dest_plus tm of | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1059 | SOME (t, u) => dest_sum t @ dest_sum u | 
| 
12aa0cb2b447
move ML functions from nat_arith.ML to Divides.thy, which is the only place they are used
 huffman parents: 
47268diff
changeset | 1060 | | NONE => [tm])); | 
| 25942 | 1061 | |
| 30934 | 1062 |   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 | 1063 | |
| 30934 | 1064 | val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1065 |     (@{thm add_0_left} :: @{thm add_0_right} :: @{thms ac_simps}))
 | 
| 41550 | 1066 | ) | 
| 60758 | 1067 | \<close> | 
| 1068 | ||
| 1069 | simproc_setup cancel_div_mod_nat ("(m::nat) + n") = \<open>K Cancel_Div_Mod_Nat.proc\<close>
 | |
| 1070 | ||
| 1071 | ||
| 1072 | subsubsection \<open>Quotient\<close> | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1073 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1074 | lemma div_geq: "0 < n \<Longrightarrow> \<not> m < n \<Longrightarrow> m div n = Suc ((m - n) div n)" | 
| 29667 | 1075 | 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 | 1076 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1077 | lemma div_if: "0 < n \<Longrightarrow> m div n = (if m < n then 0 else Suc ((m - n) div n))" | 
| 29667 | 1078 | by (simp add: div_geq) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1079 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1080 | lemma div_mult_self_is_m [simp]: "0<n ==> (m*n) div n = (m::nat)" | 
| 29667 | 1081 | by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1082 | |
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1083 | lemma div_mult_self1_is_m [simp]: "0<n ==> (n*m) div n = (m::nat)" | 
| 29667 | 1084 | by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1085 | |
| 53066 | 1086 | lemma div_positive: | 
| 1087 | fixes m n :: nat | |
| 1088 | assumes "n > 0" | |
| 1089 | assumes "m \<ge> n" | |
| 1090 | shows "m div n > 0" | |
| 1091 | proof - | |
| 60758 | 1092 | from \<open>m \<ge> n\<close> obtain q where "m = n + q" | 
| 53066 | 1093 | by (auto simp add: le_iff_add) | 
| 60758 | 1094 | with \<open>n > 0\<close> show ?thesis by simp | 
| 53066 | 1095 | qed | 
| 1096 | ||
| 59000 | 1097 | lemma div_eq_0_iff: "(a div b::nat) = 0 \<longleftrightarrow> a < b \<or> b = 0" | 
| 1098 | by (metis div_less div_positive div_by_0 gr0I less_numeral_extra(3) not_less) | |
| 25942 | 1099 | |
| 60758 | 1100 | subsubsection \<open>Remainder\<close> | 
| 25942 | 1101 | |
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1102 | lemma mod_less_divisor [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1103 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1104 | assumes "n > 0" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1105 | shows "m mod n < (n::nat)" | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 1106 | 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 | 1107 | |
| 51173 | 1108 | lemma mod_Suc_le_divisor [simp]: | 
| 1109 | "m mod Suc n \<le> n" | |
| 1110 | using mod_less_divisor [of "Suc n" m] by arith | |
| 1111 | ||
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1112 | lemma mod_less_eq_dividend [simp]: | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1113 | fixes m n :: nat | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1114 | shows "m mod n \<le> m" | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1115 | proof (rule add_leD2) | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1116 | 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 | 1117 | 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 | 1118 | qed | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1119 | |
| 61076 | 1120 | lemma mod_geq: "\<not> m < (n::nat) \<Longrightarrow> m mod n = (m - n) mod n" | 
| 29667 | 1121 | 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 | 1122 | |
| 61076 | 1123 | lemma mod_if: "m mod (n::nat) = (if m < n then m else (m - n) mod n)" | 
| 29667 | 1124 | by (simp add: le_mod_geq) | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1125 | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1126 | lemma mod_1 [simp]: "m mod Suc 0 = 0" | 
| 29667 | 1127 | 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 | 1128 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1129 | (* a simple rearrangement of mod_div_equality: *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1130 | lemma mult_div_cancel: "(n::nat) * (m div n) = m - (m mod n)" | 
| 47138 | 1131 | using mod_div_equality2 [of n m] by arith | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1132 | |
| 15439 | 1133 | lemma mod_le_divisor[simp]: "0 < n \<Longrightarrow> m mod n \<le> (n::nat)" | 
| 22718 | 1134 | apply (drule mod_less_divisor [where m = m]) | 
| 1135 | apply simp | |
| 1136 | done | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1137 | |
| 60758 | 1138 | subsubsection \<open>Quotient and Remainder\<close> | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1139 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 1140 | lemma divmod_nat_rel_mult1_eq: | 
| 46552 | 1141 | "divmod_nat_rel b c (q, r) | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 1142 | \<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 | 1143 | 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 | 1144 | |
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 1145 | lemma div_mult1_eq: | 
| 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 1146 | "(a * b) div c = a * (b div c) + a * (b mod c) div (c::nat)" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 1147 | by (blast intro: divmod_nat_rel_mult1_eq [THEN div_nat_unique] divmod_nat_rel) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1148 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 1149 | lemma divmod_nat_rel_add1_eq: | 
| 46552 | 1150 | "divmod_nat_rel a c (aq, ar) \<Longrightarrow> divmod_nat_rel b c (bq, br) | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 1151 | \<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 | 1152 | 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 | 1153 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1154 | (*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 | 1155 | lemma div_add1_eq: | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 1156 | "(a+b) div (c::nat) = a div c + b div c + ((a mod c + b mod c) div c)" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 1157 | by (blast intro: divmod_nat_rel_add1_eq [THEN div_nat_unique] divmod_nat_rel) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1158 | |
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 1159 | lemma divmod_nat_rel_mult2_eq: | 
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1160 | assumes "divmod_nat_rel a b (q, r)" | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1161 | shows "divmod_nat_rel a (b * c) (q div c, b *(q mod c) + r)" | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1162 | proof - | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1163 |   { assume "r < b" and "0 < c"
 | 
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1164 | then have "b * (q mod c) + r < b * c" | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1165 | apply (cut_tac m = q and n = c in mod_less_divisor) | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1166 | apply (drule_tac [2] m = "q mod c" in less_imp_Suc_add, auto) | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1167 | apply (erule_tac P = "%x. lhs < rhs x" for lhs rhs in ssubst) | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1168 | apply (simp add: add_mult_distrib2) | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1169 | done | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1170 | then have "r + b * (q mod c) < b * c" | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1171 | by (simp add: ac_simps) | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1172 | } with assms show ?thesis | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1173 | by (auto simp add: divmod_nat_rel_def algebra_simps add_mult_distrib2 [symmetric]) | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1174 | qed | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1175 | |
| 55085 
0e8e4dc55866
moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
 blanchet parents: 
54489diff
changeset | 1176 | lemma div_mult2_eq: "a div (b * c) = (a div b) div (c::nat)" | 
| 47135 
fb67b596067f
rename lemmas {div,mod}_eq -> {div,mod}_nat_unique, for consistency with minus_unique, inverse_unique, etc.
 huffman parents: 
47134diff
changeset | 1177 | by (force simp add: divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN div_nat_unique]) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1178 | |
| 55085 
0e8e4dc55866
moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
 blanchet parents: 
54489diff
changeset | 1179 | lemma mod_mult2_eq: "a mod (b * c) = b * (a div b mod c) + a mod (b::nat)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1180 | by (auto simp add: mult.commute divmod_nat_rel [THEN divmod_nat_rel_mult2_eq, THEN mod_nat_unique]) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1181 | |
| 61275 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1182 | instantiation nat :: semiring_numeral_div | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1183 | begin | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1184 | |
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1185 | definition divmod_nat :: "num \<Rightarrow> num \<Rightarrow> nat \<times> nat" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1186 | where | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1187 | divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1188 | |
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1189 | definition divmod_step_nat :: "num \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1190 | where | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1191 | "divmod_step_nat l qr = (let (q, r) = qr | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1192 | in if r \<ge> numeral l then (2 * q + 1, r - numeral l) | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1193 | else (2 * q, r))" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1194 | |
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1195 | instance | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1196 | by standard (auto intro: div_positive simp add: divmod'_nat_def divmod_step_nat_def mod_mult2_eq div_mult2_eq) | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1197 | |
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1198 | end | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1199 | |
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1200 | declare divmod_algorithm_code [where ?'a = nat, code] | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 1201 | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1202 | |
| 60758 | 1203 | subsubsection \<open>Further Facts about Quotient and Remainder\<close> | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1204 | |
| 58786 | 1205 | lemma div_1 [simp]: | 
| 1206 | "m div Suc 0 = m" | |
| 1207 | using div_by_1 [of m] by simp | |
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1208 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1209 | (* Monotonicity of div in first argument *) | 
| 30923 
2697a1d1d34a
more coherent developement in Divides.thy and IntDiv.thy
 haftmann parents: 
30840diff
changeset | 1210 | lemma div_le_mono [rule_format (no_asm)]: | 
| 22718 | 1211 | "\<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 | 1212 | apply (case_tac "k=0", simp) | 
| 15251 | 1213 | 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 | 1214 | apply (case_tac "n<k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1215 | (* 1 case n<k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1216 | apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1217 | (* 2 case n >= k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1218 | apply (case_tac "m<k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1219 | (* 2.1 case m<k *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1220 | apply simp | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1221 | (* 2.2 case m>=k *) | 
| 15439 | 1222 | 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 | 1223 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1224 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1225 | (* Antimonotonicity of div in second argument *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1226 | 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 | 1227 | apply (subgoal_tac "0<n") | 
| 22718 | 1228 | prefer 2 apply simp | 
| 15251 | 1229 | 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 | 1230 | apply (rename_tac "k") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1231 | apply (case_tac "k<n", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1232 | apply (subgoal_tac "~ (k<m) ") | 
| 22718 | 1233 | prefer 2 apply simp | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1234 | apply (simp add: div_geq) | 
| 15251 | 1235 | 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 | 1236 | prefer 2 | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1237 | 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 | 1238 | apply (rule le_trans, simp) | 
| 15439 | 1239 | apply (simp) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1240 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1241 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1242 | 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 | 1243 | apply (case_tac "n=0", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1244 | 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 | 1245 | apply (rule div_le_mono2) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1246 | apply (simp_all (no_asm_simp)) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1247 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1248 | |
| 22718 | 1249 | (* Similar for "less than" *) | 
| 47138 | 1250 | lemma div_less_dividend [simp]: | 
| 1251 | "\<lbrakk>(1::nat) < n; 0 < m\<rbrakk> \<Longrightarrow> m div n < m" | |
| 1252 | 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 | 1253 | apply (rename_tac "m") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1254 | apply (case_tac "m<n", simp) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1255 | apply (subgoal_tac "0<n") | 
| 22718 | 1256 | prefer 2 apply simp | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1257 | apply (simp add: div_geq) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1258 | apply (case_tac "n<m") | 
| 15251 | 1259 | 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 | 1260 | apply (rule impI less_trans_Suc)+ | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1261 | apply assumption | 
| 15439 | 1262 | apply (simp_all) | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1263 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1264 | |
| 60758 | 1265 | text\<open>A fact for the mutilated chess board\<close> | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1266 | 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 | 1267 | apply (case_tac "n=0", simp) | 
| 15251 | 1268 | 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 | 1269 | apply (case_tac "Suc (na) <n") | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1270 | (* case Suc(na) < n *) | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1271 | 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 | 1272 | (* case n \<le> Suc(na) *) | 
| 16796 | 1273 | apply (simp add: linorder_not_less le_Suc_eq mod_geq) | 
| 15439 | 1274 | 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 | 1275 | done | 
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1276 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1277 | lemma mod_eq_0_iff: "(m mod d = 0) = (\<exists>q::nat. m = d*q)" | 
| 29667 | 1278 | 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 | 1279 | |
| 22718 | 1280 | 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 | 1281 | |
| 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1282 | (*Loses information, namely we also have r<d provided d is nonzero*) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1283 | lemma mod_eqD: | 
| 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1284 | fixes m d r q :: nat | 
| 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1285 | assumes "m mod d = r" | 
| 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1286 | shows "\<exists>q. m = r + q * d" | 
| 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1287 | proof - | 
| 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1288 | from mod_div_equality obtain q where "q * d + m mod d = m" by blast | 
| 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1289 | with assms have "m = r + q * d" by simp | 
| 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1290 | then show ?thesis .. | 
| 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1291 | qed | 
| 14267 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 paulson parents: 
14208diff
changeset | 1292 | |
| 13152 | 1293 | lemma split_div: | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1294 | "P(n div k :: nat) = | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1295 | ((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 | 1296 | (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))") | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1297 | proof | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1298 | assume P: ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1299 | show ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1300 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1301 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 1302 | with P show ?Q by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1303 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1304 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1305 | thus ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1306 | proof (simp, intro allI impI) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1307 | fix i j | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1308 | 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 | 1309 | show "P i" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1310 | proof (cases) | 
| 22718 | 1311 | assume "i = 0" | 
| 1312 | 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 | 1313 | next | 
| 22718 | 1314 | assume "i \<noteq> 0" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1315 | with not0 n j P show "P i" by(simp add:ac_simps) | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1316 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1317 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1318 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1319 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1320 | assume Q: ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1321 | show ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1322 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1323 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 1324 | with Q show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1325 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1326 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1327 | with Q have R: ?R by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1328 | from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] | 
| 13517 | 1329 | show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1330 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1331 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1332 | |
| 13882 | 1333 | lemma split_div_lemma: | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1334 | assumes "0 < n" | 
| 61076 | 1335 | shows "n * q \<le> m \<and> m < n * Suc q \<longleftrightarrow> q = ((m::nat) div n)" (is "?lhs \<longleftrightarrow> ?rhs") | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1336 | proof | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1337 | assume ?rhs | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1338 | 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 | 1339 | then have A: "n * q \<le> m" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1340 | 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 | 1341 | 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 | 1342 | 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 | 1343 | with nq have "m < n + n * q" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1344 | then have B: "m < n * Suc q" by simp | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1345 | from A B show ?lhs .. | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1346 | next | 
| 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1347 | assume P: ?lhs | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 1348 | then have "divmod_nat_rel m n (q, m - n * q)" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1349 | unfolding divmod_nat_rel_def by (auto simp add: ac_simps) | 
| 33340 
a165b97f3658
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
33318diff
changeset | 1350 | 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 | 1351 | 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 | 1352 | then show ?rhs by simp | 
| 26100 
fbc60cd02ae2
using only an relation predicate to construct div and mod
 haftmann parents: 
26072diff
changeset | 1353 | qed | 
| 13882 | 1354 | |
| 1355 | theorem split_div': | |
| 1356 | "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 | 1357 | (\<exists>q. (n * q \<le> m \<and> m < n * (Suc q)) \<and> P q))" | 
| 13882 | 1358 | apply (case_tac "0 < n") | 
| 1359 | 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 | 1360 | apply simp_all | 
| 13882 | 1361 | done | 
| 1362 | ||
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1363 | lemma split_mod: | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1364 | "P(n mod k :: nat) = | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1365 | ((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 | 1366 | (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))") | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1367 | proof | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1368 | assume P: ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1369 | show ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1370 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1371 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 1372 | with P show ?Q by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1373 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1374 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1375 | thus ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1376 | proof (simp, intro allI impI) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1377 | fix i j | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1378 | assume "n = k*i + j" "j < k" | 
| 58786 | 1379 | thus "P j" using not0 P by (simp add: ac_simps) | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1380 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1381 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1382 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1383 | assume Q: ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1384 | show ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1385 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1386 | assume "k = 0" | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
27540diff
changeset | 1387 | with Q show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1388 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1389 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1390 | with Q have R: ?R by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1391 | from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] | 
| 13517 | 1392 | show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1393 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1394 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 1395 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1396 | theorem mod_div_equality' [nitpick_unfold]: "(m::nat) mod n = m - (m div n) * n" | 
| 47138 | 1397 | using mod_div_equality [of m n] by arith | 
| 1398 | ||
| 1399 | lemma div_mod_equality': "(m::nat) div n * n = m - m mod n" | |
| 1400 | using mod_div_equality [of m n] by arith | |
| 1401 | (* FIXME: very similar to mult_div_cancel *) | |
| 22800 | 1402 | |
| 52398 | 1403 | lemma div_eq_dividend_iff: "a \<noteq> 0 \<Longrightarrow> (a :: nat) div b = a \<longleftrightarrow> b = 1" | 
| 1404 | apply rule | |
| 1405 | apply (cases "b = 0") | |
| 1406 | apply simp_all | |
| 1407 | apply (metis (full_types) One_nat_def Suc_lessI div_less_dividend less_not_refl3) | |
| 1408 | done | |
| 1409 | ||
| 22800 | 1410 | |
| 60758 | 1411 | subsubsection \<open>An ``induction'' law for modulus arithmetic.\<close> | 
| 14640 | 1412 | |
| 1413 | lemma mod_induct_0: | |
| 1414 | assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)" | |
| 1415 | and base: "P i" and i: "i<p" | |
| 1416 | shows "P 0" | |
| 1417 | proof (rule ccontr) | |
| 1418 | assume contra: "\<not>(P 0)" | |
| 1419 | from i have p: "0<p" by simp | |
| 1420 | have "\<forall>k. 0<k \<longrightarrow> \<not> P (p-k)" (is "\<forall>k. ?A k") | |
| 1421 | proof | |
| 1422 | fix k | |
| 1423 | show "?A k" | |
| 1424 | proof (induct k) | |
| 1425 | show "?A 0" by simp -- "by contradiction" | |
| 1426 | next | |
| 1427 | fix n | |
| 1428 | assume ih: "?A n" | |
| 1429 | show "?A (Suc n)" | |
| 1430 | proof (clarsimp) | |
| 22718 | 1431 | assume y: "P (p - Suc n)" | 
| 1432 | have n: "Suc n < p" | |
| 1433 | proof (rule ccontr) | |
| 1434 | assume "\<not>(Suc n < p)" | |
| 1435 | hence "p - Suc n = 0" | |
| 1436 | by simp | |
| 1437 | with y contra show "False" | |
| 1438 | by simp | |
| 1439 | qed | |
| 1440 | hence n2: "Suc (p - Suc n) = p-n" by arith | |
| 1441 | from p have "p - Suc n < p" by arith | |
| 1442 | with y step have z: "P ((Suc (p - Suc n)) mod p)" | |
| 1443 | by blast | |
| 1444 | show "False" | |
| 1445 | proof (cases "n=0") | |
| 1446 | case True | |
| 1447 | with z n2 contra show ?thesis by simp | |
| 1448 | next | |
| 1449 | case False | |
| 1450 | with p have "p-n < p" by arith | |
| 1451 | with z n2 False ih show ?thesis by simp | |
| 1452 | qed | |
| 14640 | 1453 | qed | 
| 1454 | qed | |
| 1455 | qed | |
| 1456 | moreover | |
| 1457 | from i obtain k where "0<k \<and> i+k=p" | |
| 1458 | by (blast dest: less_imp_add_positive) | |
| 1459 | hence "0<k \<and> i=p-k" by auto | |
| 1460 | moreover | |
| 1461 | note base | |
| 1462 | ultimately | |
| 1463 | show "False" by blast | |
| 1464 | qed | |
| 1465 | ||
| 1466 | lemma mod_induct: | |
| 1467 | assumes step: "\<forall>i<p. P i \<longrightarrow> P ((Suc i) mod p)" | |
| 1468 | and base: "P i" and i: "i<p" and j: "j<p" | |
| 1469 | shows "P j" | |
| 1470 | proof - | |
| 1471 | have "\<forall>j<p. P j" | |
| 1472 | proof | |
| 1473 | fix j | |
| 1474 | show "j<p \<longrightarrow> P j" (is "?A j") | |
| 1475 | proof (induct j) | |
| 1476 | from step base i show "?A 0" | |
| 22718 | 1477 | by (auto elim: mod_induct_0) | 
| 14640 | 1478 | next | 
| 1479 | fix k | |
| 1480 | assume ih: "?A k" | |
| 1481 | show "?A (Suc k)" | |
| 1482 | proof | |
| 22718 | 1483 | assume suc: "Suc k < p" | 
| 1484 | hence k: "k<p" by simp | |
| 1485 | with ih have "P k" .. | |
| 1486 | with step k have "P (Suc k mod p)" | |
| 1487 | by blast | |
| 1488 | moreover | |
| 1489 | from suc have "Suc k mod p = Suc k" | |
| 1490 | by simp | |
| 1491 | ultimately | |
| 1492 | show "P (Suc k)" by simp | |
| 14640 | 1493 | qed | 
| 1494 | qed | |
| 1495 | qed | |
| 1496 | with j show ?thesis by blast | |
| 1497 | qed | |
| 1498 | ||
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1499 | lemma div2_Suc_Suc [simp]: "Suc (Suc m) div 2 = Suc (m div 2)" | 
| 47138 | 1500 | by (simp add: numeral_2_eq_2 le_div_geq) | 
| 1501 | ||
| 1502 | lemma mod2_Suc_Suc [simp]: "Suc (Suc m) mod 2 = m mod 2" | |
| 1503 | by (simp add: numeral_2_eq_2 le_mod_geq) | |
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1504 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1505 | lemma add_self_div_2 [simp]: "(m + m) div 2 = (m::nat)" | 
| 47217 
501b9bbd0d6e
removed redundant nat-specific copies of theorems
 huffman parents: 
47167diff
changeset | 1506 | by (simp add: mult_2 [symmetric]) | 
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1507 | |
| 61076 | 1508 | lemma mod2_gr_0 [simp]: "0 < (m::nat) mod 2 \<longleftrightarrow> m mod 2 = 1" | 
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1509 | proof - | 
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1510 |   { 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 | 1511 | moreover have "m mod 2 < 2" by simp | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1512 | 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 | 1513 | then show ?thesis by auto | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1514 | qed | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1515 | |
| 60758 | 1516 | text\<open>These lemmas collapse some needless occurrences of Suc: | 
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1517 | at least three Sucs, since two and fewer are rewritten back to Suc again! | 
| 60758 | 1518 | We already have some rules to simplify operands smaller than 3.\<close> | 
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1519 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1520 | 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 | 1521 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1522 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1523 | 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 | 1524 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1525 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1526 | 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 | 1527 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1528 | |
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1529 | 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 | 1530 | by (simp add: Suc3_eq_add_3) | 
| 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1531 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1532 | lemmas Suc_div_eq_add3_div_numeral [simp] = Suc_div_eq_add3_div [of _ "numeral v"] for v | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1533 | lemmas Suc_mod_eq_add3_mod_numeral [simp] = Suc_mod_eq_add3_mod [of _ "numeral v"] for v | 
| 33296 
a3924d1069e5
moved theory Divides after theory Nat_Numeral; tuned some proof texts
 haftmann parents: 
33274diff
changeset | 1534 | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1535 | lemma Suc_times_mod_eq: "1<k ==> Suc (k * m) mod k = 1" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1536 | apply (induct "m") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1537 | apply (simp_all add: mod_Suc) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1538 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1539 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1540 | declare Suc_times_mod_eq [of "numeral w", simp] for w | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1541 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1542 | lemma mod_greater_zero_iff_not_dvd: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1543 | fixes m n :: nat | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1544 | shows "m mod n > 0 \<longleftrightarrow> \<not> n dvd m" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1545 | by (simp add: dvd_eq_mod_eq_0) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1546 | |
| 47138 | 1547 | lemma Suc_div_le_mono [simp]: "n div k \<le> (Suc n) div k" | 
| 1548 | by (simp add: div_le_mono) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1549 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1550 | 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 | 1551 | by (cases n) simp_all | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1552 | |
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1553 | 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 | 1554 | proof - | 
| 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1555 | from A have B: "0 < n - 1" and C: "n - 1 + 1 = n" by simp_all | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1556 | from Suc_n_div_2_gt_zero [OF B] C show ?thesis by simp | 
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 1557 | qed | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1558 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1559 | 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 | 1560 | proof - | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1561 | have "Suc (k * n + m) mod n = (k * n + Suc m) mod n" by simp | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1562 | also have "... = Suc m mod n" by (rule mod_mult_self3) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1563 | finally show ?thesis . | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1564 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1565 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1566 | lemma mod_Suc_eq_Suc_mod: "Suc m mod n = Suc (m mod n) mod n" | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1567 | apply (subst mod_Suc [of m]) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1568 | apply (subst mod_Suc [of "m mod n"], simp) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1569 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1570 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1571 | lemma mod_2_not_eq_zero_eq_one_nat: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1572 | fixes n :: nat | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 1573 | shows "n mod 2 \<noteq> 0 \<longleftrightarrow> n mod 2 = 1" | 
| 58786 | 1574 | by (fact not_mod_2_eq_0_eq_1) | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1575 | |
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1576 | lemma even_Suc_div_two [simp]: | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1577 | "even n \<Longrightarrow> Suc n div 2 = n div 2" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1578 | using even_succ_div_two [of n] by simp | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1579 | |
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1580 | lemma odd_Suc_div_two [simp]: | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1581 | "odd n \<Longrightarrow> Suc n div 2 = Suc (n div 2)" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1582 | using odd_succ_div_two [of n] by simp | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1583 | |
| 58834 | 1584 | lemma odd_two_times_div_two_nat [simp]: | 
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1585 | assumes "odd n" | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1586 | shows "2 * (n div 2) = n - (1 :: nat)" | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1587 | proof - | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1588 | from assms have "2 * (n div 2) + 1 = n" | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1589 | by (rule odd_two_times_div_two_succ) | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1590 | then have "Suc (2 * (n div 2)) - 1 = n - 1" | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1591 | by simp | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1592 | then show ?thesis | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1593 | by simp | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1594 | qed | 
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1595 | |
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1596 | lemma parity_induct [case_names zero even odd]: | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1597 | assumes zero: "P 0" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1598 | assumes even: "\<And>n. P n \<Longrightarrow> P (2 * n)" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1599 | assumes odd: "\<And>n. P n \<Longrightarrow> P (Suc (2 * n))" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1600 | shows "P n" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1601 | proof (induct n rule: less_induct) | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1602 | case (less n) | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1603 | show "P n" | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1604 | proof (cases "n = 0") | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1605 | case True with zero show ?thesis by simp | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1606 | next | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1607 | case False | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1608 | with less have hyp: "P (n div 2)" by simp | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1609 | show ?thesis | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1610 | proof (cases "even n") | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1611 | case True | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1612 | with hyp even [of "n div 2"] show ?thesis | 
| 58834 | 1613 | by simp | 
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1614 | next | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1615 | case False | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1616 | with hyp odd [of "n div 2"] show ?thesis | 
| 58834 | 1617 | by simp | 
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1618 | qed | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1619 | qed | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1620 | qed | 
| 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58710diff
changeset | 1621 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1622 | lemma Suc_0_div_numeral [simp]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1623 | fixes k l :: num | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1624 | shows "Suc 0 div numeral k = fst (divmod Num.One k)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1625 | by (simp_all add: fst_divmod) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1626 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1627 | lemma Suc_0_mod_numeral [simp]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1628 | fixes k l :: num | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1629 | shows "Suc 0 mod numeral k = snd (divmod Num.One k)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1630 | by (simp_all add: snd_divmod) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1631 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1632 | |
| 60758 | 1633 | subsection \<open>Division on @{typ int}\<close>
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1634 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1635 | definition divmod_int_rel :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" -- \<open>definition of quotient and remainder\<close> | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1636 | where "divmod_int_rel a b = (\<lambda>(q, r). a = b * q + r \<and> | 
| 47139 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1637 | (if 0 < b then 0 \<le> r \<and> r < b else if b < 0 then b < r \<and> r \<le> 0 else q = 0))" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1638 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1639 | lemma unique_quotient_lemma: | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1640 | "b * q' + r' \<le> b * q + r \<Longrightarrow> 0 \<le> r' \<Longrightarrow> r' < b \<Longrightarrow> r < b \<Longrightarrow> q' \<le> (q::int)" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1641 | apply (subgoal_tac "r' + b * (q'-q) \<le> r") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1642 | prefer 2 apply (simp add: right_diff_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1643 | apply (subgoal_tac "0 < b * (1 + q - q') ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1644 | apply (erule_tac [2] order_le_less_trans) | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
48891diff
changeset | 1645 | prefer 2 apply (simp add: right_diff_distrib distrib_left) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1646 | apply (subgoal_tac "b * q' < b * (1 + q) ") | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
48891diff
changeset | 1647 | prefer 2 apply (simp add: right_diff_distrib distrib_left) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1648 | apply (simp add: mult_less_cancel_left) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1649 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1650 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1651 | lemma unique_quotient_lemma_neg: | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1652 | "b * q' + r' \<le> b*q + r \<Longrightarrow> r \<le> 0 \<Longrightarrow> b < r \<Longrightarrow> b < r' \<Longrightarrow> q \<le> (q'::int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1653 | by (rule_tac b = "-b" and r = "-r'" and r' = "-r" in unique_quotient_lemma) auto | 
| 33361 
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 unique_quotient: | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1656 | "divmod_int_rel a b (q, r) \<Longrightarrow> divmod_int_rel a b (q', r') \<Longrightarrow> q = q'" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1657 | 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 | 1658 | apply (blast intro: order_antisym | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1659 | dest: order_eq_refl [THEN unique_quotient_lemma] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1660 | order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1661 | done | 
| 
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 unique_remainder: | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1664 | "divmod_int_rel a b (q, r) \<Longrightarrow> divmod_int_rel a b (q', r') \<Longrightarrow> r = r'" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1665 | apply (subgoal_tac "q = q'") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1666 | apply (simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1667 | apply (blast intro: unique_quotient) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1668 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1669 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1670 | instantiation int :: Divides.div | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1671 | begin | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1672 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1673 | definition divide_int | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1674 | where "k div l = (if l = 0 \<or> k = 0 then 0 | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1675 | else if k > 0 \<and> l > 0 \<or> k < 0 \<and> l < 0 | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1676 | then int (nat \<bar>k\<bar> div nat \<bar>l\<bar>) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1677 | else | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1678 | if l dvd k then - int (nat \<bar>k\<bar> div nat \<bar>l\<bar>) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1679 | else - int (Suc (nat \<bar>k\<bar> div nat \<bar>l\<bar>)))" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1680 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1681 | definition mod_int | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1682 | where "k mod l = (if l = 0 then k else if l dvd k then 0 | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1683 | else if k > 0 \<and> l > 0 \<or> k < 0 \<and> l < 0 | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1684 | then sgn l * int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1685 | else sgn l * (\<bar>l\<bar> - int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)))" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1686 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1687 | instance .. | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1688 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1689 | end | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1690 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1691 | lemma divmod_int_rel: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1692 | "divmod_int_rel k l (k div l, k mod l)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1693 | apply (cases k rule: int_cases3) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1694 | apply (simp add: mod_greater_zero_iff_not_dvd not_le divmod_int_rel_def divide_int_def mod_int_def algebra_simps int_dvd_iff of_nat_mult [symmetric]) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1695 | apply (cases l rule: int_cases3) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1696 | apply (simp add: mod_greater_zero_iff_not_dvd not_le divmod_int_rel_def divide_int_def mod_int_def algebra_simps int_dvd_iff of_nat_mult [symmetric]) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1697 | apply (simp add: mod_greater_zero_iff_not_dvd not_le divmod_int_rel_def divide_int_def mod_int_def algebra_simps int_dvd_iff of_nat_mult [symmetric]) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1698 | apply (simp add: of_nat_add [symmetric]) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1699 | apply (simp add: mod_greater_zero_iff_not_dvd not_le divmod_int_rel_def divide_int_def mod_int_def algebra_simps int_dvd_iff of_nat_mult [symmetric]) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1700 | apply (simp add: of_nat_add [symmetric]) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1701 | apply (cases l rule: int_cases3) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1702 | apply (simp_all add: mod_greater_zero_iff_not_dvd not_le divmod_int_rel_def divide_int_def mod_int_def algebra_simps int_dvd_iff of_nat_mult [symmetric]) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1703 | apply (simp_all add: of_nat_add [symmetric]) | 
| 41550 | 1704 | done | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1705 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1706 | instantiation int :: ring_div | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1707 | begin | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1708 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1709 | subsubsection \<open>Uniqueness and Monotonicity of Quotients and Remainders\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1710 | |
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1711 | lemma divmod_int_unique: | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1712 | assumes "divmod_int_rel k l (q, r)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1713 | shows div_int_unique: "k div l = q" and mod_int_unique: "k mod l = r" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1714 | using assms divmod_int_rel [of k l] | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1715 | using unique_quotient [of k l] unique_remainder [of k l] | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1716 | by auto | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1717 | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60353diff
changeset | 1718 | instance | 
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1719 | proof | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1720 | fix a b :: int | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1721 | show "a div b * b + a mod b = a" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1722 | using divmod_int_rel [of a b] | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1723 | unfolding divmod_int_rel_def by (simp add: mult.commute) | 
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1724 | next | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1725 | fix a b c :: int | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1726 | assume "b \<noteq> 0" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1727 | hence "divmod_int_rel (a + c * b) b (c + a div b, a mod b)" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1728 | using divmod_int_rel [of a b] | 
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1729 | unfolding divmod_int_rel_def by (auto simp: algebra_simps) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1730 | thus "(a + c * b) div b = c + a div b" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1731 | by (rule div_int_unique) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1732 | next | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1733 | fix a b c :: int | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1734 | assume "c \<noteq> 0" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1735 | hence "\<And>q r. divmod_int_rel a b (q, r) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1736 | \<Longrightarrow> divmod_int_rel (c * a) (c * b) (q, c * r)" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1737 | unfolding divmod_int_rel_def | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1738 | by - (rule linorder_cases [of 0 b], auto simp: algebra_simps | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1739 | mult_less_0_iff zero_less_mult_iff mult_strict_right_mono | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1740 | mult_strict_right_mono_neg zero_le_mult_iff mult_le_0_iff) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1741 | hence "divmod_int_rel (c * a) (c * b) (a div b, c * (a mod b))" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1742 | using divmod_int_rel [of a b] . | 
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1743 | thus "(c * a) div (c * b) = a div b" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1744 | by (rule div_int_unique) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1745 | next | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1746 | fix a :: int show "a div 0 = 0" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1747 | by (rule div_int_unique, simp add: divmod_int_rel_def) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1748 | next | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1749 | fix a :: int show "0 div a = 0" | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1750 | by (rule div_int_unique, auto simp add: divmod_int_rel_def) | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1751 | qed | 
| 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1752 | |
| 60429 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60353diff
changeset | 1753 | end | 
| 
d3d1e185cd63
uniform _ div _ as infix syntax for ring division
 haftmann parents: 
60353diff
changeset | 1754 | |
| 60517 
f16e4fb20652
separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
 haftmann parents: 
60516diff
changeset | 1755 | lemma is_unit_int: | 
| 
f16e4fb20652
separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
 haftmann parents: 
60516diff
changeset | 1756 | "is_unit (k::int) \<longleftrightarrow> k = 1 \<or> k = - 1" | 
| 
f16e4fb20652
separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
 haftmann parents: 
60516diff
changeset | 1757 | by auto | 
| 
f16e4fb20652
separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
 haftmann parents: 
60516diff
changeset | 1758 | |
| 60685 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1759 | instantiation int :: normalization_semidom | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1760 | begin | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1761 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1762 | definition normalize_int | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1763 | where [simp]: "normalize = (abs :: int \<Rightarrow> int)" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1764 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1765 | definition unit_factor_int | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1766 | where [simp]: "unit_factor = (sgn :: int \<Rightarrow> int)" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1767 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1768 | instance | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1769 | proof | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1770 | fix k :: int | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1771 | assume "k \<noteq> 0" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1772 | then have "\<bar>sgn k\<bar> = 1" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1773 | by (cases "0::int" k rule: linorder_cases) simp_all | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1774 | then show "is_unit (unit_factor k)" | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1775 | by simp | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1776 | qed (simp_all add: sgn_times mult_sgn_abs) | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1777 | |
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1778 | end | 
| 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 haftmann parents: 
60562diff
changeset | 1779 | |
| 60758 | 1780 | text\<open>Basic laws about division and remainder\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1781 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1782 | lemma zmod_zdiv_equality: "(a::int) = b * (a div b) + (a mod b)" | 
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1783 | by (fact mod_div_equality2 [symmetric]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1784 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1785 | lemma zdiv_int: "int (a div b) = int a div int b" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1786 | by (simp add: divide_int_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1787 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1788 | lemma zmod_int: "int (a mod b) = int a mod int b" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1789 | by (simp add: mod_int_def int_dvd_iff) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1790 | |
| 60758 | 1791 | text \<open>Tool setup\<close> | 
| 1792 | ||
| 1793 | ML \<open> | |
| 43594 | 1794 | structure Cancel_Div_Mod_Int = Cancel_Div_Mod | 
| 41550 | 1795 | ( | 
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59833diff
changeset | 1796 |   val div_name = @{const_name Rings.divide};
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1797 |   val mod_name = @{const_name mod};
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1798 | val mk_binop = HOLogic.mk_binop; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1799 | val mk_sum = Arith_Data.mk_sum HOLogic.intT; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1800 | val dest_sum = Arith_Data.dest_sum; | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1801 | |
| 47165 | 1802 |   val div_mod_eqs = map mk_meta_eq [@{thm div_mod_equality}, @{thm div_mod_equality2}];
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1803 | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1804 | val prove_eq_sums = Arith_Data.prove_conv2 all_tac (Arith_Data.simp_all_tac | 
| 59556 | 1805 |     (@{thm diff_conv_add_uminus} :: @{thms add_0_left add_0_right} @ @{thms ac_simps}))
 | 
| 41550 | 1806 | ) | 
| 60758 | 1807 | \<close> | 
| 1808 | ||
| 1809 | simproc_setup cancel_div_mod_int ("(k::int) + l") = \<open>K Cancel_Div_Mod_Int.proc\<close>
 | |
| 43594 | 1810 | |
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1811 | lemma pos_mod_conj: "(0::int) < b \<Longrightarrow> 0 \<le> a mod b \<and> a mod b < b" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1812 | using divmod_int_rel [of a b] | 
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1813 | by (auto simp add: divmod_int_rel_def prod_eq_iff) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1814 | |
| 45607 | 1815 | lemmas pos_mod_sign [simp] = pos_mod_conj [THEN conjunct1] | 
| 1816 | and pos_mod_bound [simp] = pos_mod_conj [THEN conjunct2] | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1817 | |
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1818 | lemma neg_mod_conj: "b < (0::int) \<Longrightarrow> a mod b \<le> 0 \<and> b < a mod b" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1819 | using divmod_int_rel [of a b] | 
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 1820 | by (auto simp add: divmod_int_rel_def prod_eq_iff) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1821 | |
| 45607 | 1822 | lemmas neg_mod_sign [simp] = neg_mod_conj [THEN conjunct1] | 
| 1823 | and neg_mod_bound [simp] = neg_mod_conj [THEN conjunct2] | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1824 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1825 | |
| 60758 | 1826 | subsubsection \<open>General Properties of div and mod\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1827 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1828 | lemma div_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a div b = 0" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1829 | apply (rule div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1830 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1831 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1832 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1833 | lemma div_neg_neg_trivial: "[| a \<le> (0::int); b < a |] ==> a div b = 0" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1834 | apply (rule div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1835 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1836 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1837 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1838 | lemma div_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a div b = -1" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1839 | apply (rule div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1840 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1841 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1842 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1843 | (*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 | 1844 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1845 | lemma mod_pos_pos_trivial: "[| (0::int) \<le> a; a < b |] ==> a mod b = a" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1846 | apply (rule_tac q = 0 in mod_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1847 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1848 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1849 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1850 | lemma mod_neg_neg_trivial: "[| a \<le> (0::int); b < a |] ==> a mod b = a" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1851 | apply (rule_tac q = 0 in mod_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1852 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1853 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1854 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1855 | lemma mod_pos_neg_trivial: "[| (0::int) < a; a+b \<le> 0 |] ==> a mod b = a+b" | 
| 47140 
97c3676c5c94
rename lemmas {divmod_int_rel_{div,mod} -> {div,mod}_int_unique, for consistency with nat lemma names
 huffman parents: 
47139diff
changeset | 1856 | apply (rule_tac q = "-1" in mod_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1857 | apply (auto simp add: divmod_int_rel_def) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1858 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1859 | |
| 60758 | 1860 | text\<open>There is no @{text mod_neg_pos_trivial}.\<close>
 | 
| 1861 | ||
| 1862 | ||
| 1863 | subsubsection \<open>Laws for div and mod with Unary Minus\<close> | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1864 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1865 | lemma zminus1_lemma: | 
| 47139 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 1866 | "divmod_int_rel a b (q, r) ==> b \<noteq> 0 | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1867 | ==> divmod_int_rel (-a) b (if r=0 then -q else -q - 1, | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1868 | if r=0 then 0 else b-r)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1869 | 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 | 1870 | |
| 
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 zdiv_zminus1_eq_if: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1873 | "b \<noteq> (0::int) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1874 | ==> (-a) div b = | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1875 | (if a mod b = 0 then - (a div b) else - (a div b) - 1)" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1876 | by (blast intro: divmod_int_rel [THEN zminus1_lemma, THEN div_int_unique]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1877 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1878 | lemma zmod_zminus1_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1879 | "(-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 | 1880 | apply (case_tac "b = 0", simp) | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1881 | apply (blast intro: divmod_int_rel [THEN zminus1_lemma, THEN mod_int_unique]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1882 | done | 
| 
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 zmod_zminus1_not_zero: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1885 | fixes k l :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1886 | 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 | 1887 | unfolding zmod_zminus1_eq_if by auto | 
| 
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 | lemma zdiv_zminus2_eq_if: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1890 | "b \<noteq> (0::int) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1891 | ==> a div (-b) = | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1892 | (if a mod b = 0 then - (a div b) else - (a div b) - 1)" | 
| 47159 | 1893 | by (simp add: zdiv_zminus1_eq_if div_minus_right) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1894 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1895 | lemma zmod_zminus2_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1896 | "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" | 
| 47159 | 1897 | by (simp add: zmod_zminus1_eq_if mod_minus_right) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1898 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1899 | lemma zmod_zminus2_not_zero: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1900 | fixes k l :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1901 | shows "k mod - l \<noteq> 0 \<Longrightarrow> k mod l \<noteq> 0" | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1902 | unfolding zmod_zminus2_eq_if by auto | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1903 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1904 | |
| 60758 | 1905 | subsubsection \<open>Monotonicity in the First Argument (Dividend)\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1906 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1907 | 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 | 1908 | 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 | 1909 | 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 | 1910 | apply (rule unique_quotient_lemma) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1911 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1912 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1913 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1914 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1915 | 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 | 1916 | 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 | 1917 | 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 | 1918 | apply (rule unique_quotient_lemma_neg) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1919 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1920 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1921 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1922 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1923 | |
| 60758 | 1924 | subsubsection \<open>Monotonicity in the Second Argument (Divisor)\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1925 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1926 | lemma q_pos_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1927 | "[| 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 | 1928 | apply (subgoal_tac "0 < b'* (q' + 1) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1929 | apply (simp add: zero_less_mult_iff) | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
48891diff
changeset | 1930 | apply (simp add: distrib_left) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1931 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1932 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1933 | lemma zdiv_mono2_lemma: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1934 | "[| b*q + r = b'*q' + r'; 0 \<le> b'*q' + r'; | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1935 | r' < b'; 0 \<le> r; 0 < b'; b' \<le> b |] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1936 | ==> q \<le> (q'::int)" | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1937 | apply (frule q_pos_lemma, assumption+) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1938 | apply (subgoal_tac "b*q < b* (q' + 1) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1939 | apply (simp add: mult_less_cancel_left) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1940 | apply (subgoal_tac "b*q = r' - r + b'*q'") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1941 | prefer 2 apply simp | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
48891diff
changeset | 1942 | apply (simp (no_asm_simp) add: distrib_left) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1943 | apply (subst add.commute, rule add_less_le_mono, arith) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1944 | apply (rule mult_right_mono, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1945 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1946 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1947 | lemma zdiv_mono2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1948 | "[| (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 | 1949 | apply (subgoal_tac "b \<noteq> 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1950 | prefer 2 apply arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1951 | 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 | 1952 | 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 | 1953 | apply (rule zdiv_mono2_lemma) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1954 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1955 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1956 | done | 
| 
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 | lemma q_neg_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1959 | "[| 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 | 1960 | apply (subgoal_tac "b'*q' < 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1961 | apply (simp add: mult_less_0_iff, arith) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1962 | done | 
| 
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 | lemma zdiv_mono2_neg_lemma: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1965 | "[| b*q + r = b'*q' + r'; b'*q' + r' < 0; | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1966 | r < b; 0 \<le> r'; 0 < b'; b' \<le> b |] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1967 | ==> q' \<le> (q::int)" | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1968 | apply (frule q_neg_lemma, assumption+) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1969 | apply (subgoal_tac "b*q' < b* (q + 1) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1970 | apply (simp add: mult_less_cancel_left) | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
48891diff
changeset | 1971 | apply (simp add: distrib_left) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1972 | apply (subgoal_tac "b*q' \<le> b'*q'") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1973 | 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 | 1974 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1975 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1976 | lemma zdiv_mono2_neg: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1977 | "[| 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 | 1978 | 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 | 1979 | 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 | 1980 | apply (rule zdiv_mono2_neg_lemma) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1981 | apply (erule subst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1982 | apply (erule subst, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1983 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1984 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1985 | |
| 60758 | 1986 | subsubsection \<open>More Algebraic Laws for div and mod\<close> | 
| 1987 | ||
| 1988 | text\<open>proving (a*b) div c = a * (b div c) + a * (b mod c)\<close> | |
| 33361 
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 zmult1_lemma: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 1991 | "[| divmod_int_rel b c (q, r) |] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1992 | ==> divmod_int_rel (a * b) c (a*q + a*r div c, a*r mod c)" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1993 | by (auto simp add: split_ifs divmod_int_rel_def linorder_neq_iff distrib_left ac_simps) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1994 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1995 | 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 | 1996 | apply (case_tac "c = 0", simp) | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1997 | apply (blast intro: divmod_int_rel [THEN zmult1_lemma, THEN div_int_unique]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1998 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1999 | |
| 60758 | 2000 | text\<open>proving (a+b) div c = a div c + b div c + ((a mod c + b mod c) div c)\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2001 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2002 | lemma zadd1_lemma: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2003 | "[| divmod_int_rel a c (aq, ar); divmod_int_rel b c (bq, br) |] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2004 | ==> divmod_int_rel (a+b) c (aq + bq + (ar+br) div c, (ar+br) mod c)" | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
48891diff
changeset | 2005 | by (force simp add: split_ifs divmod_int_rel_def linorder_neq_iff distrib_left) | 
| 33361 
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 | (*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 | 2008 | lemma zdiv_zadd1_eq: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2009 | "(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 | 2010 | apply (case_tac "c = 0", simp) | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2011 | apply (blast intro: zadd1_lemma [OF divmod_int_rel divmod_int_rel] div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2012 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2013 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2014 | 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 | 2015 | 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 | 2016 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2017 | (* 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 | 2018 | 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 | 2019 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2020 | lemma zmod_zdiv_equality' [nitpick_unfold]: | 
| 61076 | 2021 | "(m::int) mod n = m - (m div n) * n" | 
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 2022 | using mod_div_equality [of m n] by arith | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2023 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2024 | |
| 60758 | 2025 | subsubsection \<open>Proving  @{term "a div (b * c) = (a div b) div c"}\<close>
 | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2026 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2027 | (*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 | 2028 | 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 | 2029 | to cause particular problems.*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2030 | |
| 60758 | 2031 | text\<open>first, four lemmas to bound the remainder for the cases b<0 and b>0\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2032 | |
| 55085 
0e8e4dc55866
moved 'fundef_cong' attribute (and other basic 'fun' stuff) up the dependency chain
 blanchet parents: 
54489diff
changeset | 2033 | lemma zmult2_lemma_aux1: "[| (0::int) < c; b < r; r \<le> 0 |] ==> b * c < b * (q mod c) + r" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2034 | 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 | 2035 | apply (simp add: algebra_simps) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2036 | apply (rule order_le_less_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2037 | apply (erule_tac [2] mult_strict_right_mono) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2038 | apply (rule mult_left_mono_neg) | 
| 35216 | 2039 | 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 | 2040 | apply (simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2041 | apply (simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2042 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2043 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2044 | lemma zmult2_lemma_aux2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2045 | "[| (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 | 2046 | apply (subgoal_tac "b * (q mod c) \<le> 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2047 | apply arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2048 | apply (simp add: mult_le_0_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2049 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2050 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2051 | 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 | 2052 | apply (subgoal_tac "0 \<le> b * (q mod c) ") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2053 | apply arith | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2054 | apply (simp add: zero_le_mult_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2055 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2056 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2057 | 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 | 2058 | 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 | 2059 | apply (simp add: right_diff_distrib) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2060 | apply (rule order_less_le_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2061 | apply (erule mult_strict_right_mono) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2062 | apply (rule_tac [2] mult_left_mono) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2063 | apply simp | 
| 35216 | 2064 | 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 | 2065 | apply simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2066 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2067 | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2068 | lemma zmult2_lemma: "[| divmod_int_rel a b (q, r); 0 < c |] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2069 | ==> divmod_int_rel a (b * c) (q div c, b*(q mod c) + r)" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 2070 | by (auto simp add: mult.assoc divmod_int_rel_def linorder_neq_iff | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2071 | zero_less_mult_iff distrib_left [symmetric] | 
| 47139 
98bddfa0f717
extend definition of divmod_int_rel to handle denominator=0 case
 huffman parents: 
47138diff
changeset | 2072 | zmult2_lemma_aux1 zmult2_lemma_aux2 zmult2_lemma_aux3 zmult2_lemma_aux4 mult_less_0_iff split: split_if_asm) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2073 | |
| 53068 | 2074 | lemma zdiv_zmult2_eq: | 
| 2075 | fixes a b c :: int | |
| 2076 | shows "0 \<le> c \<Longrightarrow> a div (b * c) = (a div b) div c" | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2077 | apply (case_tac "b = 0", simp) | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2078 | apply (force simp add: le_less divmod_int_rel [THEN zmult2_lemma, THEN div_int_unique]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2079 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2080 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2081 | lemma zmod_zmult2_eq: | 
| 53068 | 2082 | fixes a b c :: int | 
| 2083 | shows "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b" | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2084 | apply (case_tac "b = 0", simp) | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2085 | apply (force simp add: le_less divmod_int_rel [THEN zmult2_lemma, THEN mod_int_unique]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2086 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2087 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2088 | lemma div_pos_geq: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2089 | fixes k l :: int | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2090 | assumes "0 < l" and "l \<le> k" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2091 | shows "k div l = (k - l) div l + 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2092 | proof - | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2093 | have "k = (k - l) + l" by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2094 | then obtain j where k: "k = j + l" .. | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2095 | with assms show ?thesis by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2096 | qed | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2097 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2098 | lemma mod_pos_geq: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2099 | fixes k l :: int | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2100 | assumes "0 < l" and "l \<le> k" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2101 | shows "k mod l = (k - l) mod l" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2102 | proof - | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2103 | have "k = (k - l) + l" by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2104 | then obtain j where k: "k = j + l" .. | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2105 | with assms show ?thesis by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2106 | qed | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2107 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2108 | |
| 60758 | 2109 | subsubsection \<open>Splitting Rules for div and mod\<close> | 
| 2110 | ||
| 2111 | text\<open>The proofs of the two lemmas below are essentially identical\<close> | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2112 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2113 | lemma split_pos_lemma: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2114 | "0<k ==> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2115 | 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 | 2116 | apply (rule iffI, clarify) | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2117 | apply (erule_tac P="P x y" for x y in rev_mp) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2118 | apply (subst mod_add_eq) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2119 | apply (subst zdiv_zadd1_eq) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2120 | apply (simp add: div_pos_pos_trivial mod_pos_pos_trivial) | 
| 60758 | 2121 | txt\<open>converse direction\<close> | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2122 | apply (drule_tac x = "n div k" in spec) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2123 | 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 | 2124 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2125 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2126 | lemma split_neg_lemma: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2127 | "k<0 ==> | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2128 | 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 | 2129 | apply (rule iffI, clarify) | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2130 | apply (erule_tac P="P x y" for x y in rev_mp) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2131 | apply (subst mod_add_eq) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2132 | apply (subst zdiv_zadd1_eq) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2133 | apply (simp add: div_neg_neg_trivial mod_neg_neg_trivial) | 
| 60758 | 2134 | txt\<open>converse direction\<close> | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2135 | apply (drule_tac x = "n div k" in spec) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2136 | 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 | 2137 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2138 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2139 | lemma split_zdiv: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2140 | "P(n div k :: int) = | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2141 | ((k = 0 --> P 0) & | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2142 | (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P i)) & | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2143 | (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 | 2144 | apply (case_tac "k=0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2145 | apply (simp only: linorder_neq_iff) | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2146 | apply (erule disjE) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2147 | apply (simp_all add: split_pos_lemma [of concl: "%x y. P x"] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2148 | split_neg_lemma [of concl: "%x y. P x"]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2149 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2150 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2151 | lemma split_zmod: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2152 | "P(n mod k :: int) = | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2153 | ((k = 0 --> P n) & | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2154 | (0<k --> (\<forall>i j. 0\<le>j & j<k & n = k*i + j --> P j)) & | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2155 | (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 | 2156 | apply (case_tac "k=0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2157 | apply (simp only: linorder_neq_iff) | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2158 | apply (erule disjE) | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2159 | apply (simp_all add: split_pos_lemma [of concl: "%x y. P y"] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2160 | split_neg_lemma [of concl: "%x y. P y"]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2161 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2162 | |
| 60758 | 2163 | text \<open>Enable (lin)arith to deal with @{const divide} and @{const mod}
 | 
| 33730 
1755ca4ec022
Fixed splitting of div and mod on integers (split theorem differed from implementation).
 webertj parents: 
33728diff
changeset | 2164 | when these are applied to some constant that is of the form | 
| 60758 | 2165 |   @{term "numeral k"}:\<close>
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2166 | declare split_zdiv [of _ _ "numeral k", arith_split] for k | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2167 | declare split_zmod [of _ _ "numeral k", arith_split] for k | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2168 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2169 | |
| 60758 | 2170 | subsubsection \<open>Computing @{text "div"} and @{text "mod"} with shifting\<close>
 | 
| 47166 | 2171 | |
| 2172 | lemma pos_divmod_int_rel_mult_2: | |
| 2173 | assumes "0 \<le> b" | |
| 2174 | assumes "divmod_int_rel a b (q, r)" | |
| 2175 | shows "divmod_int_rel (1 + 2*a) (2*b) (q, 1 + 2*r)" | |
| 2176 | using assms unfolding divmod_int_rel_def by auto | |
| 2177 | ||
| 60758 | 2178 | declaration \<open>K (Lin_Arith.add_simps @{thms uminus_numeral_One})\<close>
 | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 2179 | |
| 47166 | 2180 | lemma neg_divmod_int_rel_mult_2: | 
| 2181 | assumes "b \<le> 0" | |
| 2182 | assumes "divmod_int_rel (a + 1) b (q, r)" | |
| 2183 | shows "divmod_int_rel (1 + 2*a) (2*b) (q, 2*r - 1)" | |
| 2184 | using assms unfolding divmod_int_rel_def by auto | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2185 | |
| 60758 | 2186 | text\<open>computing div by shifting\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2187 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2188 | lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2189 | using pos_divmod_int_rel_mult_2 [OF _ divmod_int_rel] | 
| 47166 | 2190 | by (rule div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2191 | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2192 | lemma neg_zdiv_mult_2: | 
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 2193 | assumes A: "a \<le> (0::int)" shows "(1 + 2*b) div (2*a) = (b+1) div a" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2194 | using neg_divmod_int_rel_mult_2 [OF A divmod_int_rel] | 
| 47166 | 2195 | by (rule div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2196 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2197 | (* FIXME: add rules for negative numerals *) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2198 | lemma zdiv_numeral_Bit0 [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2199 | "numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2200 | numeral v div (numeral w :: int)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2201 | unfolding numeral.simps unfolding mult_2 [symmetric] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2202 | by (rule div_mult_mult1, simp) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2203 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2204 | lemma zdiv_numeral_Bit1 [simp]: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2205 | "numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2206 | (numeral v div (numeral w :: int))" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2207 | unfolding numeral.simps | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 2208 | unfolding mult_2 [symmetric] add.commute [of _ 1] | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2209 | by (rule pos_zdiv_mult_2, simp) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2210 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2211 | lemma pos_zmod_mult_2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2212 | fixes a b :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2213 | assumes "0 \<le> a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2214 | shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2215 | using pos_divmod_int_rel_mult_2 [OF assms divmod_int_rel] | 
| 47166 | 2216 | by (rule mod_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2217 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2218 | lemma neg_zmod_mult_2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2219 | fixes a b :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2220 | assumes "a \<le> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2221 | shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2222 | using neg_divmod_int_rel_mult_2 [OF assms divmod_int_rel] | 
| 47166 | 2223 | by (rule mod_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2224 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2225 | (* FIXME: add rules for negative numerals *) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2226 | lemma zmod_numeral_Bit0 [simp]: | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2227 | "numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2228 | (2::int) * (numeral v mod numeral w)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2229 | unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2230 | unfolding mult_2 [symmetric] by (rule mod_mult_mult1) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2231 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2232 | lemma zmod_numeral_Bit1 [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2233 | "numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2234 | 2 * (numeral v mod numeral w) + (1::int)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2235 | unfolding numeral_Bit1 [of v] numeral_Bit0 [of w] | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 2236 | unfolding mult_2 [symmetric] add.commute [of _ 1] | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 2237 | by (rule pos_zmod_mult_2, simp) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2238 | |
| 39489 | 2239 | lemma zdiv_eq_0_iff: | 
| 2240 | "(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") | |
| 2241 | proof | |
| 2242 | assume ?L | |
| 2243 | have "?L \<longrightarrow> ?R" by (rule split_zdiv[THEN iffD2]) simp | |
| 60758 | 2244 | with \<open>?L\<close> show ?R by blast | 
| 39489 | 2245 | next | 
| 2246 | assume ?R thus ?L | |
| 2247 | by(auto simp: div_pos_pos_trivial div_neg_neg_trivial) | |
| 2248 | qed | |
| 2249 | ||
| 2250 | ||
| 60758 | 2251 | subsubsection \<open>Quotients of Signs\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2252 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2253 | lemma div_eq_minus1: "(0::int) < b ==> -1 div b = -1" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2254 | by (simp add: divide_int_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2255 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2256 | lemma zmod_minus1: "(0::int) < b ==> -1 mod b = b - 1" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2257 | by (simp add: mod_int_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2258 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2259 | 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 | 2260 | apply (subgoal_tac "a div b \<le> -1", force) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2261 | apply (rule order_trans) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2262 | apply (rule_tac a' = "-1" in zdiv_mono1) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2263 | apply (auto simp add: div_eq_minus1) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2264 | done | 
| 
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 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 | 2267 | by (drule zdiv_mono1_neg, auto) | 
| 
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 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 | 2270 | by (drule zdiv_mono1, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2271 | |
| 60758 | 2272 | text\<open>Now for some equivalences of the form @{text"a div b >=< 0 \<longleftrightarrow> \<dots>"}
 | 
| 33804 | 2273 | conditional upon the sign of @{text a} or @{text b}. There are many more.
 | 
| 60758 | 2274 | They should all be simp rules unless that causes too much search.\<close> | 
| 33804 | 2275 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2276 | 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 | 2277 | apply auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2278 | apply (drule_tac [2] zdiv_mono1) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2279 | apply (auto simp add: linorder_neq_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2280 | 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 | 2281 | apply (blast intro: div_neg_pos_less0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2282 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2283 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2284 | lemma pos_imp_zdiv_pos_iff: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2285 | "0<k \<Longrightarrow> 0 < (i::int) div k \<longleftrightarrow> k \<le> i" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2286 | using pos_imp_zdiv_nonneg_iff[of k i] zdiv_eq_0_iff[of i k] | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2287 | by arith | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2288 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2289 | lemma neg_imp_zdiv_nonneg_iff: | 
| 33804 | 2290 | "b < (0::int) ==> (0 \<le> a div b) = (a \<le> (0::int))" | 
| 47159 | 2291 | apply (subst div_minus_minus [symmetric]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2292 | apply (subst pos_imp_zdiv_nonneg_iff, auto) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2293 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2294 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2295 | (*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 | 2296 | 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 | 2297 | 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 | 2298 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2299 | (*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 | 2300 | 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 | 2301 | 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 | 2302 | |
| 33804 | 2303 | lemma nonneg1_imp_zdiv_pos_iff: | 
| 2304 | "(0::int) <= a \<Longrightarrow> (a div b > 0) = (a >= b & b>0)" | |
| 2305 | apply rule | |
| 2306 | apply rule | |
| 2307 | using div_pos_pos_trivial[of a b]apply arith | |
| 2308 | apply(cases "b=0")apply simp | |
| 2309 | using div_nonneg_neg_le0[of a b]apply arith | |
| 2310 | using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b]apply simp | |
| 2311 | done | |
| 2312 | ||
| 39489 | 2313 | lemma zmod_le_nonneg_dividend: "(m::int) \<ge> 0 ==> m mod k \<le> m" | 
| 2314 | apply (rule split_zmod[THEN iffD2]) | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44766diff
changeset | 2315 | apply(fastforce dest: q_pos_lemma intro: split_mult_pos_le) | 
| 39489 | 2316 | done | 
| 2317 | ||
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2318 | lemma zmult_div_cancel: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2319 | "(n::int) * (m div n) = m - (m mod n)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2320 | using zmod_zdiv_equality [where a="m" and b="n"] | 
| 47142 | 2321 | by (simp add: algebra_simps) (* FIXME: generalize *) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2322 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2323 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2324 | subsubsection \<open>Computation of Division and Remainder\<close> | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2325 | |
| 61275 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2326 | instantiation int :: semiring_numeral_div | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2327 | begin | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2328 | |
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2329 | definition divmod_int :: "num \<Rightarrow> num \<Rightarrow> int \<times> int" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2330 | where | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2331 | "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2332 | |
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2333 | definition divmod_step_int :: "num \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2334 | where | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2335 | "divmod_step_int l qr = (let (q, r) = qr | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2336 | in if r \<ge> numeral l then (2 * q + 1, r - numeral l) | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2337 | else (2 * q, r))" | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2338 | |
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2339 | instance | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2340 | by standard (auto intro: zmod_le_nonneg_dividend simp add: divmod_int_def divmod_step_int_def | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2341 | pos_imp_zdiv_pos_iff div_pos_pos_trivial mod_pos_pos_trivial zmod_zmult2_eq zdiv_zmult2_eq) | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2342 | |
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2343 | end | 
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2344 | |
| 
053ec04ea866
monomorphization of divmod wrt. code generation avoids costly dictionary unpacking at runtime
 haftmann parents: 
61201diff
changeset | 2345 | declare divmod_algorithm_code [where ?'a = int, code] | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2346 | |
| 60930 | 2347 | context | 
| 2348 | begin | |
| 2349 | ||
| 2350 | qualified definition adjust_div :: "int \<times> int \<Rightarrow> int" | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2351 | where | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2352 | "adjust_div qr = (let (q, r) = qr in q + of_bool (r \<noteq> 0))" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2353 | |
| 60930 | 2354 | qualified lemma adjust_div_eq [simp, code]: | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2355 | "adjust_div (q, r) = q + of_bool (r \<noteq> 0)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2356 | by (simp add: adjust_div_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2357 | |
| 60930 | 2358 | qualified definition adjust_mod :: "int \<Rightarrow> int \<Rightarrow> int" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2359 | where | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2360 | [simp]: "adjust_mod l r = (if r = 0 then 0 else l - r)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2361 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2362 | lemma minus_numeral_div_numeral [simp]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2363 | "- numeral m div numeral n = - (adjust_div (divmod m n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2364 | proof - | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2365 | have "int (fst (divmod m n)) = fst (divmod m n)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2366 | by (simp only: fst_divmod divide_int_def) auto | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2367 | then show ?thesis | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2368 | by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2369 | qed | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2370 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2371 | lemma minus_numeral_mod_numeral [simp]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2372 | "- numeral m mod numeral n = adjust_mod (numeral n) (snd (divmod m n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2373 | proof - | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2374 | have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2375 | using that by (simp only: snd_divmod mod_int_def) auto | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2376 | then show ?thesis | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2377 | by (auto simp add: split_def Let_def adjust_div_def divides_aux_def mod_int_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2378 | qed | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2379 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2380 | lemma numeral_div_minus_numeral [simp]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2381 | "numeral m div - numeral n = - (adjust_div (divmod m n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2382 | proof - | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2383 | have "int (fst (divmod m n)) = fst (divmod m n)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2384 | by (simp only: fst_divmod divide_int_def) auto | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2385 | then show ?thesis | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2386 | by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2387 | qed | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2388 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2389 | lemma numeral_mod_minus_numeral [simp]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2390 | "numeral m mod - numeral n = - adjust_mod (numeral n) (snd (divmod m n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2391 | proof - | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2392 | have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2393 | using that by (simp only: snd_divmod mod_int_def) auto | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2394 | then show ?thesis | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2395 | by (auto simp add: split_def Let_def adjust_div_def divides_aux_def mod_int_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2396 | qed | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2397 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2398 | lemma minus_one_div_numeral [simp]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2399 | "- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2400 | using minus_numeral_div_numeral [of Num.One n] by simp | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2401 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2402 | lemma minus_one_mod_numeral [simp]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2403 | "- 1 mod numeral n = adjust_mod (numeral n) (snd (divmod Num.One n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2404 | using minus_numeral_mod_numeral [of Num.One n] by simp | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2405 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2406 | lemma one_div_minus_numeral [simp]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2407 | "1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2408 | using numeral_div_minus_numeral [of Num.One n] by simp | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2409 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2410 | lemma one_mod_minus_numeral [simp]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2411 | "1 mod - numeral n = - adjust_mod (numeral n) (snd (divmod Num.One n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2412 | using numeral_mod_minus_numeral [of Num.One n] by simp | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2413 | |
| 60930 | 2414 | end | 
| 2415 | ||
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2416 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2417 | subsubsection \<open>Further properties\<close> | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2418 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2419 | text \<open>Simplify expresions in which div and mod combine numerical constants\<close> | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2420 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2421 | lemma int_div_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a div b = q" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2422 | by (rule div_int_unique [of a b q r]) (simp add: divmod_int_rel_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2423 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2424 | lemma int_div_neg_eq: "\<lbrakk>(a::int) = b * q + r; r \<le> 0; b < r\<rbrakk> \<Longrightarrow> a div b = q" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2425 | by (rule div_int_unique [of a b q r], | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2426 | simp add: divmod_int_rel_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2427 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2428 | lemma int_mod_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a mod b = r" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2429 | by (rule mod_int_unique [of a b q r], | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2430 | simp add: divmod_int_rel_def) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2431 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2432 | lemma int_mod_neg_eq: "\<lbrakk>(a::int) = b * q + r; r \<le> 0; b < r\<rbrakk> \<Longrightarrow> a mod b = r" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2433 | by (rule mod_int_unique [of a b q r], | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2434 | simp add: divmod_int_rel_def) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2435 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2436 | 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 | 2437 | 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 | 2438 | |
| 60758 | 2439 | text\<open>Suggested by Matthias Daum\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2440 | lemma int_power_div_base: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2441 | "\<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 | 2442 | 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 | 2443 | apply (erule ssubst) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2444 | apply (simp only: power_add) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2445 | apply simp_all | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2446 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2447 | |
| 60758 | 2448 | text \<open>by Brian Huffman\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2449 | 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 | 2450 | by (rule mod_minus_eq [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2451 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2452 | 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 | 2453 | by (rule mod_diff_left_eq [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2454 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2455 | 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 | 2456 | by (rule mod_diff_right_eq [symmetric]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2457 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2458 | lemmas zmod_simps = | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2459 | mod_add_left_eq [symmetric] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2460 | mod_add_right_eq [symmetric] | 
| 47142 | 2461 | mod_mult_right_eq[symmetric] | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2462 | mod_mult_left_eq [symmetric] | 
| 47164 | 2463 | power_mod | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2464 | zminus_zmod zdiff_zmod_left zdiff_zmod_right | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2465 | |
| 60758 | 2466 | text \<open>Distributive laws for function @{text nat}.\<close>
 | 
| 33361 
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 | 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 | 2469 | apply (rule linorder_cases [of y 0]) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2470 | apply (simp add: div_nonneg_neg_le0) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2471 | apply simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2472 | 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 | 2473 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2474 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2475 | (*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 | 2476 | lemma nat_mod_distrib: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2477 | "\<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 | 2478 | apply (case_tac "y = 0", simp) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2479 | apply (simp add: nat_eq_iff zmod_int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2480 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2481 | |
| 60758 | 2482 | text \<open>transfer setup\<close> | 
| 33361 
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 | lemma transfer_nat_int_functions: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2485 | "(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 | 2486 | "(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 | 2487 | 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 | 2488 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2489 | lemma transfer_nat_int_function_closures: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2490 | "(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 | 2491 | "(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 | 2492 | apply (cases "y = 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2493 | apply (auto simp add: pos_imp_zdiv_nonneg_iff) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2494 | apply (cases "y = 0") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2495 | apply auto | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2496 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2497 | |
| 35644 | 2498 | declare transfer_morphism_nat_int [transfer add return: | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2499 | transfer_nat_int_functions | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2500 | transfer_nat_int_function_closures | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2501 | ] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2502 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2503 | lemma transfer_int_nat_functions: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2504 | "(int x) div (int y) = int (x div y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2505 | "(int x) mod (int y) = int (x mod y)" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2506 | by (auto simp add: zdiv_int zmod_int) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2507 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2508 | lemma transfer_int_nat_function_closures: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2509 | "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 | 2510 | "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 | 2511 | 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 | 2512 | |
| 35644 | 2513 | declare transfer_morphism_int_nat [transfer add return: | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2514 | transfer_int_nat_functions | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2515 | transfer_int_nat_function_closures | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2516 | ] | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2517 | |
| 60758 | 2518 | text\<open>Suggested by Matthias Daum\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2519 | 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 | 2520 | apply (subgoal_tac "nat x div nat k < nat x") | 
| 34225 | 2521 | apply (simp add: nat_div_distrib [symmetric]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2522 | apply (rule Divides.div_less_dividend, simp_all) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2523 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2524 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2525 | 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 | 2526 | proof | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2527 | assume H: "x mod n = y mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2528 | hence "x mod n - y mod n = 0" by simp | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2529 | hence "(x mod n - y mod n) mod n = 0" by simp | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2530 | 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 | 2531 | 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 | 2532 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2533 | assume H: "n dvd x - y" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2534 | 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 | 2535 | hence "x = n*k + y" by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2536 | 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 | 2537 | 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 | 2538 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2539 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2540 | 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 | 2541 | shows "\<exists>q. x = y + n * q" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2542 | proof- | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2543 | from xy have th: "int x - int y = int (x - y)" by simp | 
| 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2544 | from xyn have "int x mod int n = int y mod int n" | 
| 46551 
866bce5442a3
simplify projections on simultaneous computations of div and mod; tuned structure (from Florian Haftmann)
 huffman parents: 
46026diff
changeset | 2545 | by (simp add: zmod_int [symmetric]) | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2546 | hence "int n dvd int x - int y" by (simp only: zmod_eq_dvd_iff[symmetric]) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2547 | 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 | 2548 | 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 | 2549 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2550 | |
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2551 | lemma nat_mod_eq_iff: "(x::nat) mod n = y mod n \<longleftrightarrow> (\<exists>q1 q2. x + n * q1 = y + n * q2)" | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2552 | (is "?lhs = ?rhs") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2553 | proof | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2554 | assume H: "x mod n = y mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2555 |   {assume xy: "x \<le> y"
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2556 | from H have th: "y mod n = x mod n" by simp | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2557 | from nat_mod_eq_lemma[OF th xy] have ?rhs | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2558 | 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 | 2559 | moreover | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2560 |   {assume xy: "y \<le> x"
 | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2561 | from nat_mod_eq_lemma[OF H xy] have ?rhs | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2562 | apply clarify apply (rule_tac x="0" in exI) by (rule_tac x="q" in exI, simp)} | 
| 60562 
24af00b010cf
Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
 paulson <lp15@cam.ac.uk> parents: 
60517diff
changeset | 2563 | ultimately show ?rhs using linear[of x y] by blast | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2564 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2565 | 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 | 2566 | 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 | 2567 | thus ?lhs by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2568 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2569 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2570 | subsubsection \<open>Dedicated simproc for calculation\<close> | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2571 | |
| 60758 | 2572 | text \<open> | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2573 | There is space for improvement here: the calculation itself | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2574 | could be carried outside the logic, and a generic simproc | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2575 | (simplifier setup) for generic calculation would be helpful. | 
| 60758 | 2576 | \<close> | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 2577 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2578 | simproc_setup numeral_divmod | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2579 |   ("0 div 0 :: 'a :: semiring_numeral_div" | "0 mod 0 :: 'a :: semiring_numeral_div" |
 | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2580 | "0 div 1 :: 'a :: semiring_numeral_div" | "0 mod 1 :: 'a :: semiring_numeral_div" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2581 | "0 div - 1 :: int" | "0 mod - 1 :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2582 | "0 div numeral b :: 'a :: semiring_numeral_div" | "0 mod numeral b :: 'a :: semiring_numeral_div" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2583 | "0 div - numeral b :: int" | "0 mod - numeral b :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2584 | "1 div 0 :: 'a :: semiring_numeral_div" | "1 mod 0 :: 'a :: semiring_numeral_div" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2585 | "1 div 1 :: 'a :: semiring_numeral_div" | "1 mod 1 :: 'a :: semiring_numeral_div" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2586 | "1 div - 1 :: int" | "1 mod - 1 :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2587 | "1 div numeral b :: 'a :: semiring_numeral_div" | "1 mod numeral b :: 'a :: semiring_numeral_div" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2588 | "1 div - numeral b :: int" |"1 mod - numeral b :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2589 | "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2590 | "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2591 | "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2592 | "numeral a div 0 :: 'a :: semiring_numeral_div" | "numeral a mod 0 :: 'a :: semiring_numeral_div" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2593 | "numeral a div 1 :: 'a :: semiring_numeral_div" | "numeral a mod 1 :: 'a :: semiring_numeral_div" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2594 | "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2595 | "numeral a div numeral b :: 'a :: semiring_numeral_div" | "numeral a mod numeral b :: 'a :: semiring_numeral_div" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2596 | "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2597 | "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2598 | "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2599 | "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2600 | "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2601 | "- numeral a div - numeral b :: int" | "- numeral a mod - numeral b :: int") = | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2602 | \<open> let | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2603 |     val if_cong = the (Code.get_case_cong @{theory} @{const_name If});
 | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2604 | fun successful_rewrite ctxt ct = | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2605 | let | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2606 | val thm = Simplifier.rewrite ctxt ct | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2607 | in if Thm.is_reflexive thm then NONE else SOME thm end; | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2608 | in fn phi => | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2609 | let | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2610 |       val simps = Morphism.fact phi (@{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1
 | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2611 | one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2612 | one_div_minus_numeral one_mod_minus_numeral | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2613 | numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2614 | numeral_div_minus_numeral numeral_mod_minus_numeral | 
| 60930 | 2615 | div_minus_minus mod_minus_minus Divides.adjust_div_eq of_bool_eq one_neq_zero | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2616 | numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2617 | divmod_cancel divmod_steps divmod_step_eq fst_conv snd_conv numeral_One | 
| 60930 | 2618 | case_prod_beta rel_simps Divides.adjust_mod_def div_minus1_right mod_minus1_right | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2619 | minus_minus numeral_times_numeral mult_zero_right mult_1_right} | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2620 |         @ [@{lemma "0 = 0 \<longleftrightarrow> True" by simp}]);
 | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2621 | fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2622 | (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2623 | in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2624 | end; | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2625 | \<close> | 
| 34126 | 2626 | |
| 35673 | 2627 | |
| 60758 | 2628 | subsubsection \<open>Code generation\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2629 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2630 | lemma [code]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2631 | fixes k :: int | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2632 | shows | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2633 | "k div 0 = 0" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2634 | "k mod 0 = k" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2635 | "0 div k = 0" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2636 | "0 mod k = 0" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2637 | "k div Int.Pos Num.One = k" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2638 | "k mod Int.Pos Num.One = 0" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2639 | "k div Int.Neg Num.One = - k" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2640 | "k mod Int.Neg Num.One = 0" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2641 | "Int.Pos m div Int.Pos n = (fst (divmod m n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2642 | "Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)" | 
| 60930 | 2643 | "Int.Neg m div Int.Pos n = - (Divides.adjust_div (divmod m n) :: int)" | 
| 2644 | "Int.Neg m mod Int.Pos n = Divides.adjust_mod (Int.Pos n) (snd (divmod m n) :: int)" | |
| 2645 | "Int.Pos m div Int.Neg n = - (Divides.adjust_div (divmod m n) :: int)" | |
| 2646 | "Int.Pos m mod Int.Neg n = - Divides.adjust_mod (Int.Pos n) (snd (divmod m n) :: int)" | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2647 | "Int.Neg m div Int.Neg n = (fst (divmod m n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2648 | "Int.Neg m mod Int.Neg n = - (snd (divmod m n) :: int)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2649 | by simp_all | 
| 53069 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 2650 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52398diff
changeset | 2651 | code_identifier | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52398diff
changeset | 2652 | code_module Divides \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith | 
| 33364 | 2653 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2654 | lemma dvd_eq_mod_eq_0_numeral: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2655 | "numeral x dvd (numeral y :: 'a) \<longleftrightarrow> numeral y mod numeral x = (0 :: 'a::semiring_div)" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2656 | by (fact dvd_eq_mod_eq_0) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 2657 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 2658 | end |