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