| author | wenzelm | 
| Mon, 24 Jun 2019 16:26:25 +0200 | |
| changeset 70359 | 470d4f145e4c | 
| parent 70340 | 7383930fc946 | 
| child 71138 | 9de7f1067520 | 
| 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 | 
| 66817 | 9 | imports Parity | 
| 15131 | 10 | begin | 
| 3366 | 11 | |
| 66817 | 12 | subsection \<open>More on division\<close> | 
| 60758 | 13 | |
| 64635 | 14 | inductive eucl_rel_int :: "int \<Rightarrow> int \<Rightarrow> int \<times> int \<Rightarrow> bool" | 
| 15 | where eucl_rel_int_by0: "eucl_rel_int k 0 (0, k)" | |
| 16 | | eucl_rel_int_dividesI: "l \<noteq> 0 \<Longrightarrow> k = q * l \<Longrightarrow> eucl_rel_int k l (q, 0)" | |
| 17 | | eucl_rel_int_remainderI: "sgn r = sgn l \<Longrightarrow> \<bar>r\<bar> < \<bar>l\<bar> | |
| 18 | \<Longrightarrow> k = q * l + r \<Longrightarrow> eucl_rel_int k l (q, r)" | |
| 19 | ||
| 20 | lemma eucl_rel_int_iff: | |
| 21 | "eucl_rel_int k l (q, r) \<longleftrightarrow> | |
| 22 | k = l * q + r \<and> | |
| 23 | (if 0 < l then 0 \<le> r \<and> r < l else if l < 0 then l < r \<and> r \<le> 0 else q = 0)" | |
| 24 | by (cases "r = 0") | |
| 25 | (auto elim!: eucl_rel_int.cases intro: eucl_rel_int_by0 eucl_rel_int_dividesI eucl_rel_int_remainderI | |
| 26 | 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 | 27 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 28 | lemma unique_quotient_lemma: | 
| 68626 | 29 | assumes "b * q' + r' \<le> b * q + r" "0 \<le> r'" "r' < b" "r < b" shows "q' \<le> (q::int)" | 
| 30 | proof - | |
| 31 | have "r' + b * (q'-q) \<le> r" | |
| 32 | using assms by (simp add: right_diff_distrib) | |
| 33 | moreover have "0 < b * (1 + q - q') " | |
| 34 | using assms by (simp add: right_diff_distrib distrib_left) | |
| 35 | moreover have "b * q' < b * (1 + q)" | |
| 36 | using assms by (simp add: right_diff_distrib distrib_left) | |
| 37 | ultimately show ?thesis | |
| 38 | using assms by (simp add: mult_less_cancel_left) | |
| 39 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 40 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 41 | lemma unique_quotient_lemma_neg: | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 42 | "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 | 43 | 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 | 44 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 45 | lemma unique_quotient: | 
| 64635 | 46 | "eucl_rel_int a b (q, r) \<Longrightarrow> eucl_rel_int a b (q', r') \<Longrightarrow> q = q'" | 
| 68626 | 47 | apply (rule order_antisym) | 
| 48 | apply (simp_all add: eucl_rel_int_iff linorder_neq_iff split: if_split_asm) | |
| 49 | apply (blast intro: order_eq_refl [THEN unique_quotient_lemma] order_eq_refl [THEN unique_quotient_lemma_neg] sym)+ | |
| 64635 | 50 | done | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 51 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 52 | lemma unique_remainder: | 
| 64635 | 53 | "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 | 54 | apply (subgoal_tac "q = q'") | 
| 64635 | 55 | apply (simp add: eucl_rel_int_iff) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 56 | apply (blast intro: unique_quotient) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 57 | done | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 58 | |
| 64635 | 59 | lemma eucl_rel_int: | 
| 60 | "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 | 61 | proof (cases k rule: int_cases3) | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 62 | case zero | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 63 | then show ?thesis | 
| 64635 | 64 | 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 | 65 | next | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 66 | case (pos n) | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 67 | then show ?thesis | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 68 | using div_mult_mod_eq [of n] | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 69 | by (cases l rule: int_cases3) | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 70 | (auto simp del: of_nat_mult of_nat_add | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 71 | simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps | 
| 67118 | 72 | 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 | 73 | next | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 74 | case (neg n) | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 75 | then show ?thesis | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 76 | using div_mult_mod_eq [of n] | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 77 | by (cases l rule: int_cases3) | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 78 | (auto simp del: of_nat_mult of_nat_add | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 79 | simp add: mod_greater_zero_iff_not_dvd of_nat_mult [symmetric] of_nat_add [symmetric] algebra_simps | 
| 67118 | 80 | 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 | 81 | qed | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 82 | |
| 47141 
02d6b816e4b3
move int::ring_div instance upward, simplify several proofs
 huffman parents: 
47140diff
changeset | 83 | lemma divmod_int_unique: | 
| 64635 | 84 | assumes "eucl_rel_int k l (q, r)" | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 85 | shows div_int_unique: "k div l = q" and mod_int_unique: "k mod l = r" | 
| 64635 | 86 | using assms eucl_rel_int [of k l] | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 87 | using unique_quotient [of k l] unique_remainder [of k l] | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 88 | by auto | 
| 64592 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 89 | |
| 64715 | 90 | lemma div_abs_eq_div_nat: | 
| 91 | "\<bar>k\<bar> div \<bar>l\<bar> = int (nat \<bar>k\<bar> div nat \<bar>l\<bar>)" | |
| 92 | by (simp add: divide_int_def) | |
| 93 | ||
| 94 | lemma mod_abs_eq_div_nat: | |
| 95 | "\<bar>k\<bar> mod \<bar>l\<bar> = int (nat \<bar>k\<bar> mod nat \<bar>l\<bar>)" | |
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 96 | by (simp add: modulo_int_def) | 
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 97 | |
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 98 | lemma zdiv_int: | 
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 99 | "int (a div b) = int a div int b" | 
| 69695 | 100 | by (simp add: divide_int_def) | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 101 | |
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 102 | lemma zmod_int: | 
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 103 | "int (a mod b) = int a mod int b" | 
| 69695 | 104 | by (simp add: modulo_int_def) | 
| 64715 | 105 | |
| 106 | lemma div_sgn_abs_cancel: | |
| 107 | fixes k l v :: int | |
| 108 | assumes "v \<noteq> 0" | |
| 109 | shows "(sgn v * \<bar>k\<bar>) div (sgn v * \<bar>l\<bar>) = \<bar>k\<bar> div \<bar>l\<bar>" | |
| 110 | proof - | |
| 111 | from assms have "sgn v = - 1 \<or> sgn v = 1" | |
| 112 | by (cases "v \<ge> 0") auto | |
| 113 | then show ?thesis | |
| 66630 | 114 | using assms unfolding divide_int_def [of "sgn v * \<bar>k\<bar>" "sgn v * \<bar>l\<bar>"] | 
| 115 | by (fastforce simp add: not_less div_abs_eq_div_nat) | |
| 64715 | 116 | qed | 
| 117 | ||
| 118 | lemma div_eq_sgn_abs: | |
| 119 | fixes k l v :: int | |
| 120 | assumes "sgn k = sgn l" | |
| 121 | shows "k div l = \<bar>k\<bar> div \<bar>l\<bar>" | |
| 122 | proof (cases "l = 0") | |
| 123 | case True | |
| 124 | then show ?thesis | |
| 125 | by simp | |
| 126 | next | |
| 127 | case False | |
| 128 | with assms have "(sgn k * \<bar>k\<bar>) div (sgn l * \<bar>l\<bar>) = \<bar>k\<bar> div \<bar>l\<bar>" | |
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 129 | using div_sgn_abs_cancel [of l k l] by simp | 
| 64715 | 130 | then show ?thesis | 
| 131 | by (simp add: sgn_mult_abs) | |
| 132 | qed | |
| 133 | ||
| 134 | lemma div_dvd_sgn_abs: | |
| 135 | fixes k l :: int | |
| 136 | assumes "l dvd k" | |
| 137 | shows "k div l = (sgn k * sgn l) * (\<bar>k\<bar> div \<bar>l\<bar>)" | |
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 138 | proof (cases "k = 0 \<or> l = 0") | 
| 64715 | 139 | case True | 
| 140 | then show ?thesis | |
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 141 | by auto | 
| 64715 | 142 | next | 
| 143 | case False | |
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 144 | then have "k \<noteq> 0" and "l \<noteq> 0" | 
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 145 | by auto | 
| 64715 | 146 | show ?thesis | 
| 147 | proof (cases "sgn l = sgn k") | |
| 148 | case True | |
| 149 | then show ?thesis | |
| 150 | by (simp add: div_eq_sgn_abs) | |
| 151 | next | |
| 152 | case False | |
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 153 | with \<open>k \<noteq> 0\<close> \<open>l \<noteq> 0\<close> | 
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 154 | have "sgn l * sgn k = - 1" | 
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 155 | by (simp add: sgn_if split: if_splits) | 
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 156 | with assms show ?thesis | 
| 64715 | 157 | unfolding divide_int_def [of k l] | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 158 | by (auto simp add: zdiv_int ac_simps) | 
| 64715 | 159 | qed | 
| 160 | qed | |
| 161 | ||
| 162 | lemma div_noneq_sgn_abs: | |
| 163 | fixes k l :: int | |
| 164 | assumes "l \<noteq> 0" | |
| 165 | assumes "sgn k \<noteq> sgn l" | |
| 166 | shows "k div l = - (\<bar>k\<bar> div \<bar>l\<bar>) - of_bool (\<not> l dvd k)" | |
| 167 | using assms | |
| 168 | by (simp only: divide_int_def [of k l], auto simp add: not_less zdiv_int) | |
| 169 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 170 | |
| 60758 | 171 | 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 | 172 | |
| 66886 | 173 | lemma div_pos_pos_trivial [simp]: | 
| 174 | "k div l = 0" if "k \<ge> 0" and "k < l" for k l :: int | |
| 69695 | 175 | using that by (simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) | 
| 176 | ||
| 177 | lemma mod_pos_pos_trivial [simp]: | |
| 178 | "k mod l = k" if "k \<ge> 0" and "k < l" for k l :: int | |
| 179 | using that by (simp add: mod_eq_self_iff_div_eq_0) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 180 | |
| 66886 | 181 | lemma div_neg_neg_trivial [simp]: | 
| 182 | "k div l = 0" if "k \<le> 0" and "l < k" for k l :: int | |
| 69695 | 183 | using that by (cases "k = 0") (simp, simp add: unique_euclidean_semiring_class.div_eq_0_iff division_segment_int_def) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 184 | |
| 66886 | 185 | lemma mod_neg_neg_trivial [simp]: | 
| 186 | "k mod l = k" if "k \<le> 0" and "l < k" for k l :: int | |
| 69695 | 187 | using that by (simp add: mod_eq_self_iff_div_eq_0) | 
| 188 | ||
| 189 | lemma div_pos_neg_trivial: | |
| 190 | "k div l = - 1" if "0 < k" and "k + l \<le> 0" for k l :: int | |
| 191 | apply (rule div_int_unique [of _ _ _ "k + l"]) | |
| 192 | apply (use that in \<open>auto simp add: eucl_rel_int_iff\<close>) | |
| 193 | done | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 194 | |
| 69695 | 195 | lemma mod_pos_neg_trivial: | 
| 196 | "k mod l = k + l" if "0 < k" and "k + l \<le> 0" for k l :: int | |
| 197 | apply (rule mod_int_unique [of _ _ "- 1"]) | |
| 198 | apply (use that in \<open>auto simp add: eucl_rel_int_iff\<close>) | |
| 199 | done | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 200 | |
| 69695 | 201 | text \<open>There is neither \<open>div_neg_pos_trivial\<close> nor \<open>mod_neg_pos_trivial\<close> | 
| 202 | because \<^term>\<open>0 div l = 0\<close> would supersede it.\<close> | |
| 203 | ||
| 60758 | 204 | |
| 205 | ||
| 206 | 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 | 207 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 208 | lemma zminus1_lemma: | 
| 64635 | 209 | "eucl_rel_int a b (q, r) ==> b \<noteq> 0 | 
| 210 | ==> 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 | 211 | if r=0 then 0 else b-r)" | 
| 66630 | 212 | by (force simp add: eucl_rel_int_iff right_diff_distrib) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 213 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 214 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 215 | 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 | 216 | "b \<noteq> (0::int) | 
| 68631 | 217 | \<Longrightarrow> (-a) div b = (if a mod b = 0 then - (a div b) else - (a div b) - 1)" | 
| 64635 | 218 | 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 | 219 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 220 | lemma zmod_zminus1_eq_if: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 221 | "(-a::int) mod b = (if a mod b = 0 then 0 else b - (a mod b))" | 
| 68631 | 222 | proof (cases "b = 0") | 
| 223 | case False | |
| 224 | then show ?thesis | |
| 225 | by (blast intro: eucl_rel_int [THEN zminus1_lemma, THEN mod_int_unique]) | |
| 226 | qed auto | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 227 | |
| 64593 
50c715579715
reoriented congruence rules in non-explosive direction
 haftmann parents: 
64592diff
changeset | 228 | lemma zmod_zminus1_not_zero: | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 229 | fixes k l :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 230 | 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 | 231 | by (simp add: mod_eq_0_iff_dvd) | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 232 | |
| 64593 
50c715579715
reoriented congruence rules in non-explosive direction
 haftmann parents: 
64592diff
changeset | 233 | lemma zmod_zminus2_not_zero: | 
| 64592 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 234 | fixes k l :: int | 
| 
7759f1766189
more fine-grained type class hierarchy for div and mod
 haftmann parents: 
64250diff
changeset | 235 | 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 | 236 | by (simp add: mod_eq_0_iff_dvd) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 237 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 238 | lemma zdiv_zminus2_eq_if: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 239 | "b \<noteq> (0::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 | 240 | ==> a div (-b) = | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 241 | (if a mod b = 0 then - (a div b) else - (a div b) - 1)" | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 242 | by (auto 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 | 243 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 244 | lemma zmod_zminus2_eq_if: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 245 | "a mod (-b::int) = (if a mod b = 0 then 0 else (a mod b) - b)" | 
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 246 | by (auto 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 | 247 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 248 | |
| 60758 | 249 | 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 | 250 | |
| 68631 | 251 | lemma zdiv_mono1: | 
| 252 | fixes b::int | |
| 253 | assumes "a \<le> a'" "0 < b" shows "a div b \<le> a' div b" | |
| 254 | proof (rule unique_quotient_lemma) | |
| 255 | show "b * (a div b) + a mod b \<le> b * (a' div b) + a' mod b" | |
| 256 | using assms(1) by auto | |
| 257 | qed (use assms in auto) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 258 | |
| 68631 | 259 | lemma zdiv_mono1_neg: | 
| 260 | fixes b::int | |
| 261 | assumes "a \<le> a'" "b < 0" shows "a' div b \<le> a div b" | |
| 262 | proof (rule unique_quotient_lemma_neg) | |
| 263 | show "b * (a div b) + a mod b \<le> b * (a' div b) + a' mod b" | |
| 264 | using assms(1) by auto | |
| 265 | qed (use assms in auto) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 266 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 267 | |
| 60758 | 268 | 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 | 269 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 270 | lemma q_pos_lemma: | 
| 68631 | 271 | fixes q'::int | 
| 272 | assumes "0 \<le> b'*q' + r'" "r' < b'" "0 < b'" | |
| 273 | shows "0 \<le> q'" | |
| 274 | proof - | |
| 275 | have "0 < b'* (q' + 1)" | |
| 276 | using assms by (simp add: distrib_left) | |
| 277 | with assms show ?thesis | |
| 278 | by (simp add: zero_less_mult_iff) | |
| 279 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 280 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 281 | lemma zdiv_mono2_lemma: | 
| 68631 | 282 | fixes q'::int | 
| 283 | assumes eq: "b*q + r = b'*q' + r'" and le: "0 \<le> b'*q' + r'" and "r' < b'" "0 \<le> r" "0 < b'" "b' \<le> b" | |
| 284 | shows "q \<le> q'" | |
| 285 | proof - | |
| 286 | have "0 \<le> q'" | |
| 287 | using q_pos_lemma le \<open>r' < b'\<close> \<open>0 < b'\<close> by blast | |
| 288 | moreover have "b*q = r' - r + b'*q'" | |
| 289 | using eq by linarith | |
| 290 | ultimately have "b*q < b* (q' + 1)" | |
| 291 | using mult_right_mono assms unfolding distrib_left by fastforce | |
| 292 | with assms show ?thesis | |
| 293 | by (simp add: mult_less_cancel_left_pos) | |
| 294 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 295 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 296 | lemma zdiv_mono2: | 
| 68631 | 297 | fixes a::int | 
| 298 | assumes "0 \<le> a" "0 < b'" "b' \<le> b" shows "a div b \<le> a div b'" | |
| 299 | proof (rule zdiv_mono2_lemma) | |
| 300 | have "b \<noteq> 0" | |
| 301 | using assms by linarith | |
| 302 | show "b * (a div b) + a mod b = b' * (a div b') + a mod b'" | |
| 303 | by simp | |
| 304 | qed (use assms in auto) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 305 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 306 | lemma zdiv_mono2_neg_lemma: | 
| 68631 | 307 | fixes q'::int | 
| 308 | assumes "b*q + r = b'*q' + r'" "b'*q' + r' < 0" "r < b" "0 \<le> r'" "0 < b'" "b' \<le> b" | |
| 309 | shows "q' \<le> q" | |
| 310 | proof - | |
| 311 | have "b'*q' < 0" | |
| 312 | using assms by linarith | |
| 313 | with assms have "q' \<le> 0" | |
| 314 | by (simp add: mult_less_0_iff) | |
| 315 | have "b*q' \<le> b'*q'" | |
| 316 | by (simp add: \<open>q' \<le> 0\<close> assms(6) mult_right_mono_neg) | |
| 317 | then have "b*q' < b* (q + 1)" | |
| 318 | using assms by (simp add: distrib_left) | |
| 319 | then show ?thesis | |
| 320 | using assms by (simp add: mult_less_cancel_left) | |
| 321 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 322 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 323 | lemma zdiv_mono2_neg: | 
| 68631 | 324 | fixes a::int | 
| 325 | assumes "a < 0" "0 < b'" "b' \<le> b" shows "a div b' \<le> a div b" | |
| 326 | proof (rule zdiv_mono2_neg_lemma) | |
| 327 | have "b \<noteq> 0" | |
| 328 | using assms by linarith | |
| 329 | show "b * (a div b) + a mod b = b' * (a div b') + a mod b'" | |
| 330 | by simp | |
| 331 | qed (use assms in auto) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 332 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 333 | |
| 69593 | 334 | subsubsection \<open>Proving \<^term>\<open>a div (b * c) = (a div b) div c\<close>\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 335 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 336 | (*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 | 337 | 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 | 338 | to cause particular problems.*) | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 339 | |
| 60758 | 340 | 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 | 341 | |
| 68631 | 342 | lemma zmult2_lemma_aux1: | 
| 343 | fixes c::int | |
| 344 | assumes "0 < c" "b < r" "r \<le> 0" | |
| 345 | shows "b * c < b * (q mod c) + r" | |
| 346 | proof - | |
| 347 | have "b * (c - q mod c) \<le> b * 1" | |
| 348 | by (rule mult_left_mono_neg) (use assms in \<open>auto simp: int_one_le_iff_zero_less\<close>) | |
| 349 | also have "... < r * 1" | |
| 350 | by (simp add: \<open>b < r\<close>) | |
| 351 | finally show ?thesis | |
| 352 | by (simp add: algebra_simps) | |
| 353 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 354 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 355 | lemma zmult2_lemma_aux2: | 
| 68631 | 356 | fixes c::int | 
| 357 | assumes "0 < c" "b < r" "r \<le> 0" | |
| 358 | shows "b * (q mod c) + r \<le> 0" | |
| 359 | proof - | |
| 360 | have "b * (q mod c) \<le> 0" | |
| 361 | using assms by (simp add: mult_le_0_iff) | |
| 362 | with assms show ?thesis | |
| 363 | by arith | |
| 364 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 365 | |
| 68631 | 366 | lemma zmult2_lemma_aux3: | 
| 367 | fixes c::int | |
| 368 | assumes "0 < c" "0 \<le> r" "r < b" | |
| 369 | shows "0 \<le> b * (q mod c) + r" | |
| 370 | proof - | |
| 371 | have "0 \<le> b * (q mod c)" | |
| 372 | using assms by (simp add: mult_le_0_iff) | |
| 373 | with assms show ?thesis | |
| 374 | by arith | |
| 375 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 376 | |
| 68631 | 377 | lemma zmult2_lemma_aux4: | 
| 378 | fixes c::int | |
| 379 | assumes "0 < c" "0 \<le> r" "r < b" | |
| 380 | shows "b * (q mod c) + r < b * c" | |
| 381 | proof - | |
| 382 | have "r * 1 < b * 1" | |
| 383 | by (simp add: \<open>r < b\<close>) | |
| 384 | also have "\<dots> \<le> b * (c - q mod c) " | |
| 385 | by (rule mult_left_mono) (use assms in \<open>auto simp: int_one_le_iff_zero_less\<close>) | |
| 386 | finally show ?thesis | |
| 387 | by (simp add: algebra_simps) | |
| 388 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 389 | |
| 64635 | 390 | lemma zmult2_lemma: "[| eucl_rel_int a b (q, r); 0 < c |] | 
| 391 | ==> eucl_rel_int a (b * c) (q div c, b*(q mod c) + r)" | |
| 392 | 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 | 393 | zero_less_mult_iff distrib_left [symmetric] | 
| 62390 | 394 | 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 | 395 | |
| 53068 | 396 | lemma zdiv_zmult2_eq: | 
| 397 | fixes a b c :: int | |
| 68631 | 398 | assumes "0 \<le> c" | 
| 399 | shows "a div (b * c) = (a div b) div c" | |
| 400 | proof (cases "b = 0") | |
| 401 | case False | |
| 402 | with assms show ?thesis | |
| 403 | by (force simp add: le_less eucl_rel_int [THEN zmult2_lemma, THEN div_int_unique]) | |
| 404 | qed auto | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 405 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 406 | lemma zmod_zmult2_eq: | 
| 53068 | 407 | fixes a b c :: int | 
| 68631 | 408 | assumes "0 \<le> c" | 
| 409 | shows "a mod (b * c) = b * (a div b mod c) + a mod b" | |
| 410 | proof (cases "b = 0") | |
| 411 | case False | |
| 412 | with assms show ?thesis | |
| 413 | by (force simp add: le_less eucl_rel_int [THEN zmult2_lemma, THEN mod_int_unique]) | |
| 414 | qed auto | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 415 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 416 | lemma div_pos_geq: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 417 | fixes k l :: int | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 418 | assumes "0 < l" and "l \<le> k" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 419 | shows "k div l = (k - l) div l + 1" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 420 | proof - | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 421 | have "k = (k - l) + l" by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 422 | 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 | 423 | with assms show ?thesis by (simp add: div_add_self2) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 424 | qed | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 425 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 426 | lemma mod_pos_geq: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 427 | fixes k l :: int | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 428 | assumes "0 < l" and "l \<le> k" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 429 | shows "k mod l = (k - l) mod l" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 430 | proof - | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 431 | have "k = (k - l) + l" by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 432 | then obtain j where k: "k = j + l" .. | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 433 | with assms show ?thesis by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 434 | qed | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 435 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 436 | |
| 60758 | 437 | subsubsection \<open>Splitting Rules for div and mod\<close> | 
| 438 | ||
| 439 | 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 | 440 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 441 | lemma split_pos_lemma: | 
| 67091 | 442 | "0<k \<Longrightarrow> | 
| 443 | P(n div k :: int)(n mod k) = (\<forall>i j. 0\<le>j \<and> j<k \<and> n = k*i + j \<longrightarrow> P i j)" | |
| 66886 | 444 | by auto | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 445 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 446 | lemma split_neg_lemma: | 
| 67091 | 447 | "k<0 \<Longrightarrow> | 
| 448 | P(n div k :: int)(n mod k) = (\<forall>i j. k<j \<and> j\<le>0 \<and> n = k*i + j \<longrightarrow> P i j)" | |
| 66886 | 449 | by auto | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 450 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 451 | lemma split_zdiv: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 452 | "P(n div k :: int) = | 
| 67091 | 453 | ((k = 0 \<longrightarrow> P 0) \<and> | 
| 454 | (0<k \<longrightarrow> (\<forall>i j. 0\<le>j \<and> j<k \<and> n = k*i + j \<longrightarrow> P i)) \<and> | |
| 455 | (k<0 \<longrightarrow> (\<forall>i j. k<j \<and> j\<le>0 \<and> n = k*i + j \<longrightarrow> P i)))" | |
| 68631 | 456 | proof (cases "k = 0") | 
| 457 | case False | |
| 458 | then show ?thesis | |
| 459 | unfolding linorder_neq_iff | |
| 460 | by (auto simp add: split_pos_lemma [of concl: "\<lambda>x y. P x"] split_neg_lemma [of concl: "\<lambda>x y. P x"]) | |
| 461 | qed auto | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 462 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 463 | lemma split_zmod: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 464 | "P(n mod k :: int) = | 
| 67091 | 465 | ((k = 0 \<longrightarrow> P n) \<and> | 
| 466 | (0<k \<longrightarrow> (\<forall>i j. 0\<le>j \<and> j<k \<and> n = k*i + j \<longrightarrow> P j)) \<and> | |
| 467 | (k<0 \<longrightarrow> (\<forall>i j. k<j \<and> j\<le>0 \<and> n = k*i + j \<longrightarrow> P j)))" | |
| 68631 | 468 | proof (cases "k = 0") | 
| 469 | case False | |
| 470 | then show ?thesis | |
| 471 | unfolding linorder_neq_iff | |
| 472 | by (auto simp add: split_pos_lemma [of concl: "\<lambda>x y. P y"] split_neg_lemma [of concl: "\<lambda>x y. P y"]) | |
| 473 | qed auto | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 474 | |
| 69593 | 475 | text \<open>Enable (lin)arith to deal with \<^const>\<open>divide\<close> and \<^const>\<open>modulo\<close> | 
| 33730 
1755ca4ec022
Fixed splitting of div and mod on integers (split theorem differed from implementation).
 webertj parents: 
33728diff
changeset | 476 | when these are applied to some constant that is of the form | 
| 69593 | 477 | \<^term>\<open>numeral k\<close>:\<close> | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 478 | declare split_zdiv [of _ _ "numeral k", arith_split] for k | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 479 | 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 | 480 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 481 | |
| 61799 | 482 | subsubsection \<open>Computing \<open>div\<close> and \<open>mod\<close> with shifting\<close> | 
| 47166 | 483 | |
| 64635 | 484 | lemma pos_eucl_rel_int_mult_2: | 
| 47166 | 485 | assumes "0 \<le> b" | 
| 64635 | 486 | assumes "eucl_rel_int a b (q, r)" | 
| 487 | shows "eucl_rel_int (1 + 2*a) (2*b) (q, 1 + 2*r)" | |
| 488 | using assms unfolding eucl_rel_int_iff by auto | |
| 489 | ||
| 490 | lemma neg_eucl_rel_int_mult_2: | |
| 47166 | 491 | assumes "b \<le> 0" | 
| 64635 | 492 | assumes "eucl_rel_int (a + 1) b (q, r)" | 
| 493 | shows "eucl_rel_int (1 + 2*a) (2*b) (q, 2*r - 1)" | |
| 494 | 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 | 495 | |
| 60758 | 496 | text\<open>computing div by shifting\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 497 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 498 | lemma pos_zdiv_mult_2: "(0::int) \<le> a ==> (1 + 2*b) div (2*a) = b div a" | 
| 64635 | 499 | using pos_eucl_rel_int_mult_2 [OF _ eucl_rel_int] | 
| 47166 | 500 | by (rule div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 501 | |
| 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 | 502 | lemma neg_zdiv_mult_2: | 
| 35815 
10e723e54076
tuned proofs (to avoid linarith error message caused by bootstrapping of HOL)
 boehmes parents: 
35673diff
changeset | 503 | assumes A: "a \<le> (0::int)" shows "(1 + 2*b) div (2*a) = (b+1) div a" | 
| 64635 | 504 | using neg_eucl_rel_int_mult_2 [OF A eucl_rel_int] | 
| 47166 | 505 | by (rule div_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 506 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 507 | lemma zdiv_numeral_Bit0 [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 508 | "numeral (Num.Bit0 v) div numeral (Num.Bit0 w) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 509 | numeral v div (numeral w :: int)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 510 | unfolding numeral.simps unfolding mult_2 [symmetric] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 511 | by (rule div_mult_mult1, simp) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 512 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 513 | 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 | 514 | "numeral (Num.Bit1 v) div numeral (Num.Bit0 w) = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 515 | (numeral v div (numeral w :: int))" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 516 | unfolding numeral.simps | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 517 | unfolding mult_2 [symmetric] add.commute [of _ 1] | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 518 | by (rule pos_zdiv_mult_2, simp) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 519 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 520 | lemma pos_zmod_mult_2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 521 | fixes a b :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 522 | assumes "0 \<le> a" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 523 | shows "(1 + 2 * b) mod (2 * a) = 1 + 2 * (b mod a)" | 
| 64635 | 524 | using pos_eucl_rel_int_mult_2 [OF assms eucl_rel_int] | 
| 47166 | 525 | by (rule mod_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 526 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 527 | lemma neg_zmod_mult_2: | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 528 | fixes a b :: int | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 529 | assumes "a \<le> 0" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 530 | shows "(1 + 2 * b) mod (2 * a) = 2 * ((b + 1) mod a) - 1" | 
| 64635 | 531 | using neg_eucl_rel_int_mult_2 [OF assms eucl_rel_int] | 
| 47166 | 532 | by (rule mod_int_unique) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 533 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 534 | 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 | 535 | "numeral (Num.Bit0 v) mod numeral (Num.Bit0 w) = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 536 | (2::int) * (numeral v mod numeral w)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 537 | unfolding numeral_Bit0 [of v] numeral_Bit0 [of w] | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 538 | unfolding mult_2 [symmetric] by (rule mod_mult_mult1) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 539 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 540 | lemma zmod_numeral_Bit1 [simp]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 541 | "numeral (Num.Bit1 v) mod numeral (Num.Bit0 w) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 542 | 2 * (numeral v mod numeral w) + (1::int)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 543 | 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 | 544 | unfolding mult_2 [symmetric] add.commute [of _ 1] | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46560diff
changeset | 545 | by (rule pos_zmod_mult_2, simp) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 546 | |
| 39489 | 547 | lemma zdiv_eq_0_iff: | 
| 66886 | 548 | "i div k = 0 \<longleftrightarrow> k = 0 \<or> 0 \<le> i \<and> i < k \<or> i \<le> 0 \<and> k < i" (is "?L = ?R") | 
| 549 | for i k :: int | |
| 39489 | 550 | proof | 
| 551 | assume ?L | |
| 66886 | 552 | moreover have "?L \<longrightarrow> ?R" | 
| 553 | by (rule split_zdiv [THEN iffD2]) simp | |
| 554 | ultimately show ?R | |
| 555 | by blast | |
| 39489 | 556 | next | 
| 66886 | 557 | assume ?R then show ?L | 
| 558 | by auto | |
| 39489 | 559 | qed | 
| 560 | ||
| 63947 | 561 | lemma zmod_trival_iff: | 
| 562 | fixes i k :: int | |
| 563 | shows "i mod k = i \<longleftrightarrow> k = 0 \<or> 0 \<le> i \<and> i < k \<or> i \<le> 0 \<and> k < i" | |
| 564 | proof - | |
| 565 | 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 | 566 | by safe (insert div_mult_mod_eq [of i k], auto) | 
| 63947 | 567 | with zdiv_eq_0_iff | 
| 568 | show ?thesis | |
| 569 | by simp | |
| 570 | qed | |
| 39489 | 571 | |
| 64785 | 572 | |
| 60758 | 573 | subsubsection \<open>Quotients of Signs\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 574 | |
| 67083 | 575 | lemma div_eq_minus1: "0 < b \<Longrightarrow> - 1 div b = - 1" for b :: int | 
| 576 | by (simp add: divide_int_def) | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 577 | |
| 67083 | 578 | lemma zmod_minus1: "0 < b \<Longrightarrow> - 1 mod b = b - 1" for b :: int | 
| 579 | by (auto simp add: modulo_int_def) | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 580 | |
| 68631 | 581 | lemma div_neg_pos_less0: | 
| 582 | fixes a::int | |
| 583 | assumes "a < 0" "0 < b" | |
| 584 | shows "a div b < 0" | |
| 585 | proof - | |
| 586 | have "a div b \<le> - 1 div b" | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
68631diff
changeset | 587 | using zdiv_mono1 assms by auto | 
| 68631 | 588 | also have "... \<le> -1" | 
| 589 | by (simp add: assms(2) div_eq_minus1) | |
| 590 | finally show ?thesis | |
| 591 | by force | |
| 592 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 593 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 594 | lemma div_nonneg_neg_le0: "[| (0::int) \<le> a; b < 0 |] ==> a div b \<le> 0" | 
| 68631 | 595 | by (drule zdiv_mono1_neg, auto) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 596 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 597 | lemma div_nonpos_pos_le0: "[| (a::int) \<le> 0; b > 0 |] ==> a div b \<le> 0" | 
| 68631 | 598 | by (drule zdiv_mono1, auto) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 599 | |
| 61799 | 600 | text\<open>Now for some equivalences of the form \<open>a div b >=< 0 \<longleftrightarrow> \<dots>\<close> | 
| 601 | conditional upon the sign of \<open>a\<close> or \<open>b\<close>. There are many more. | |
| 60758 | 602 | They should all be simp rules unless that causes too much search.\<close> | 
| 33804 | 603 | |
| 68631 | 604 | lemma pos_imp_zdiv_nonneg_iff: | 
| 605 | fixes a::int | |
| 606 | assumes "0 < b" | |
| 607 | shows "(0 \<le> a div b) = (0 \<le> a)" | |
| 608 | proof | |
| 609 | show "0 \<le> a div b \<Longrightarrow> 0 \<le> a" | |
| 610 | using assms | |
| 611 | by (simp add: linorder_not_less [symmetric]) (blast intro: div_neg_pos_less0) | |
| 612 | next | |
| 613 | assume "0 \<le> a" | |
| 614 | then have "0 div b \<le> a div b" | |
| 615 | using zdiv_mono1 assms by blast | |
| 616 | then show "0 \<le> a div b" | |
| 617 | by auto | |
| 618 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 619 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 620 | lemma pos_imp_zdiv_pos_iff: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 621 | "0<k \<Longrightarrow> 0 < (i::int) div k \<longleftrightarrow> k \<le> i" | 
| 68631 | 622 | using pos_imp_zdiv_nonneg_iff[of k i] zdiv_eq_0_iff[of i k] by arith | 
| 623 | ||
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 624 | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 625 | lemma neg_imp_zdiv_nonneg_iff: | 
| 68631 | 626 | fixes a::int | 
| 627 | assumes "b < 0" | |
| 628 | shows "(0 \<le> a div b) = (a \<le> 0)" | |
| 629 | using assms by (simp add: div_minus_minus [of a, symmetric] pos_imp_zdiv_nonneg_iff del: div_minus_minus) | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 630 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 631 | (*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 | 632 | lemma pos_imp_zdiv_neg_iff: "(0::int) < b ==> (a div b < 0) = (a < 0)" | 
| 68631 | 633 | by (simp add: linorder_not_le [symmetric] pos_imp_zdiv_nonneg_iff) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 634 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 635 | (*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 | 636 | lemma neg_imp_zdiv_neg_iff: "b < (0::int) ==> (a div b < 0) = (0 < a)" | 
| 68631 | 637 | by (simp add: linorder_not_le [symmetric] neg_imp_zdiv_nonneg_iff) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 638 | |
| 33804 | 639 | lemma nonneg1_imp_zdiv_pos_iff: | 
| 68631 | 640 | fixes a::int | 
| 641 | assumes "0 \<le> a" | |
| 642 | shows "a div b > 0 \<longleftrightarrow> a \<ge> b \<and> b>0" | |
| 643 | proof - | |
| 644 | have "0 < a div b \<Longrightarrow> b \<le> a" | |
| 645 | using div_pos_pos_trivial[of a b] assms by arith | |
| 646 | moreover have "0 < a div b \<Longrightarrow> b > 0" | |
| 647 | using assms div_nonneg_neg_le0[of a b] by(cases "b=0"; force) | |
| 648 | moreover have "b \<le> a \<and> 0 < b \<Longrightarrow> 0 < a div b" | |
| 649 | using int_one_le_iff_zero_less[of "a div b"] zdiv_mono1[of b a b] by simp | |
| 650 | ultimately show ?thesis | |
| 651 | by blast | |
| 652 | qed | |
| 33804 | 653 | |
| 68631 | 654 | lemma zmod_le_nonneg_dividend: "(m::int) \<ge> 0 \<Longrightarrow> m mod k \<le> m" | 
| 655 | by (rule split_zmod[THEN iffD2]) (fastforce dest: q_pos_lemma intro: split_mult_pos_le) | |
| 60930 | 656 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 657 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 658 | subsubsection \<open>Further properties\<close> | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 659 | |
| 66817 | 660 | lemma div_int_pos_iff: | 
| 661 | "k div l \<ge> 0 \<longleftrightarrow> k = 0 \<or> l = 0 \<or> k \<ge> 0 \<and> l \<ge> 0 | |
| 662 | \<or> k < 0 \<and> l < 0" | |
| 663 | for k l :: int | |
| 68631 | 664 | proof (cases "k = 0 \<or> l = 0") | 
| 665 | case False | |
| 666 | then show ?thesis | |
| 66817 | 667 | apply (auto simp add: pos_imp_zdiv_nonneg_iff neg_imp_zdiv_nonneg_iff) | 
| 68631 | 668 | by (meson neg_imp_zdiv_neg_iff not_le not_less_iff_gr_or_eq) | 
| 669 | qed auto | |
| 66817 | 670 | |
| 671 | lemma mod_int_pos_iff: | |
| 672 | "k mod l \<ge> 0 \<longleftrightarrow> l dvd k \<or> l = 0 \<and> k \<ge> 0 \<or> l > 0" | |
| 673 | for k l :: int | |
| 68631 | 674 | proof (cases "l > 0") | 
| 675 | case False | |
| 676 | then show ?thesis | |
| 69695 | 677 | by (simp add: dvd_eq_mod_eq_0) (use neg_mod_sign [of l k] in \<open>auto simp add: le_less not_less\<close>) | 
| 68631 | 678 | qed auto | 
| 66817 | 679 | |
| 68631 | 680 | text \<open>Simplify expressions in which div and mod combine numerical constants\<close> | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 681 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 682 | lemma int_div_pos_eq: "\<lbrakk>(a::int) = b * q + r; 0 \<le> r; r < b\<rbrakk> \<Longrightarrow> a div b = q" | 
| 64635 | 683 | 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 | 684 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 685 | 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 | 686 | by (rule div_int_unique [of a b q r], | 
| 64635 | 687 | simp add: eucl_rel_int_iff) | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 688 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 689 | 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 | 690 | by (rule mod_int_unique [of a b q r], | 
| 64635 | 691 | simp add: eucl_rel_int_iff) | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 692 | |
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 693 | 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 | 694 | by (rule mod_int_unique [of a b q r], | 
| 64635 | 695 | simp add: eucl_rel_int_iff) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 696 | |
| 61944 | 697 | lemma abs_div: "(y::int) dvd x \<Longrightarrow> \<bar>x div y\<bar> = \<bar>x\<bar> div \<bar>y\<bar>" | 
| 68631 | 698 | unfolding dvd_def by (cases "y=0") (auto simp add: abs_mult) | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 699 | |
| 60758 | 700 | text\<open>Suggested by Matthias Daum\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 701 | lemma int_power_div_base: | 
| 68631 | 702 | fixes k :: int | 
| 703 | assumes "0 < m" "0 < k" | |
| 704 | shows "k ^ m div k = (k::int) ^ (m - Suc 0)" | |
| 705 | proof - | |
| 706 | have eq: "k ^ m = k ^ ((m - Suc 0) + Suc 0)" | |
| 707 | by (simp add: assms) | |
| 708 | show ?thesis | |
| 709 | using assms by (simp only: power_add eq) auto | |
| 710 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 711 | |
| 61799 | 712 | 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 | 713 | |
| 68631 | 714 | lemma nat_div_distrib: | 
| 715 | assumes "0 \<le> x" | |
| 716 | shows "nat (x div y) = nat x div nat y" | |
| 717 | proof (cases y "0::int" rule: linorder_cases) | |
| 718 | case less | |
| 719 | with assms show ?thesis | |
| 720 | using div_nonneg_neg_le0 by auto | |
| 721 | next | |
| 722 | case greater | |
| 723 | then show ?thesis | |
| 724 | by (simp add: nat_eq_iff pos_imp_zdiv_nonneg_iff zdiv_int) | |
| 725 | qed auto | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 726 | |
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 727 | (*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 | 728 | lemma nat_mod_distrib: | 
| 68631 | 729 | assumes "0 \<le> x" "0 \<le> y" | 
| 730 | shows "nat (x mod y) = nat x mod nat y" | |
| 731 | proof (cases "y = 0") | |
| 732 | case False | |
| 733 | with assms show ?thesis | |
| 734 | by (simp add: nat_eq_iff zmod_int) | |
| 735 | qed auto | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 736 | |
| 60758 | 737 | text\<open>Suggested by Matthias Daum\<close> | 
| 68631 | 738 | lemma int_div_less_self: | 
| 739 | fixes x::int | |
| 740 | assumes "0 < x" "1 < k" | |
| 741 | shows "x div k < x" | |
| 742 | proof - | |
| 743 | have "nat x div nat k < nat x" | |
| 744 | by (simp add: assms) | |
| 745 | with assms show ?thesis | |
| 746 | by (simp add: nat_div_distrib [symmetric]) | |
| 747 | qed | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 748 | |
| 66837 | 749 | lemma mod_eq_dvd_iff_nat: | 
| 750 | "m mod q = n mod q \<longleftrightarrow> q dvd m - n" if "m \<ge> n" for m n q :: nat | |
| 751 | proof - | |
| 752 | have "int m mod int q = int n mod int q \<longleftrightarrow> int q dvd int m - int n" | |
| 753 | by (simp add: mod_eq_dvd_iff) | |
| 754 | with that have "int (m mod q) = int (n mod q) \<longleftrightarrow> int q dvd int (m - n)" | |
| 755 | by (simp only: of_nat_mod of_nat_diff) | |
| 756 | then show ?thesis | |
| 67118 | 757 | by simp | 
| 66837 | 758 | qed | 
| 759 | ||
| 760 | lemma mod_eq_nat1E: | |
| 761 | fixes m n q :: nat | |
| 762 | assumes "m mod q = n mod q" and "m \<ge> n" | |
| 763 | obtains s where "m = n + q * s" | |
| 764 | proof - | |
| 765 | from assms have "q dvd m - n" | |
| 766 | by (simp add: mod_eq_dvd_iff_nat) | |
| 767 | then obtain s where "m - n = q * s" .. | |
| 768 | with \<open>m \<ge> n\<close> have "m = n + q * s" | |
| 769 | by simp | |
| 770 | with that show thesis . | |
| 771 | qed | |
| 772 | ||
| 773 | lemma mod_eq_nat2E: | |
| 774 | fixes m n q :: nat | |
| 775 | assumes "m mod q = n mod q" and "n \<ge> m" | |
| 776 | obtains s where "n = m + q * s" | |
| 777 | using assms mod_eq_nat1E [of n q m] by (auto simp add: ac_simps) | |
| 778 | ||
| 779 | lemma nat_mod_eq_lemma: | |
| 780 | assumes "(x::nat) mod n = y mod n" and "y \<le> x" | |
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 781 | shows "\<exists>q. x = y + n * q" | 
| 66837 | 782 | using assms by (rule mod_eq_nat1E) rule | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 783 | |
| 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 | 784 | 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 | 785 | (is "?lhs = ?rhs") | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 786 | proof | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 787 | assume H: "x mod n = y mod n" | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 788 |   {assume xy: "x \<le> y"
 | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 789 | 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 | 790 | 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 | 791 | 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 | 792 | moreover | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 793 |   {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 | 794 | 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 | 795 | 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 | 796 | 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 | 797 | next | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 798 | 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 | 799 | 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 | 800 | thus ?lhs by simp | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 801 | qed | 
| 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 802 | |
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66806diff
changeset | 803 | |
| 68253 | 804 | subsection \<open>Numeral division with a pragmatic type class\<close> | 
| 805 | ||
| 806 | text \<open> | |
| 807 | The following type class contains everything necessary to formulate | |
| 808 | a division algorithm in ring structures with numerals, restricted | |
| 809 | to its positive segments. This is its primary motivation, and it | |
| 810 | could surely be formulated using a more fine-grained, more algebraic | |
| 811 | and less technical class hierarchy. | |
| 812 | \<close> | |
| 813 | ||
| 70340 | 814 | class unique_euclidean_semiring_numeral = unique_euclidean_semiring_with_nat + linordered_semidom + | 
| 68253 | 815 | assumes div_less: "0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a div b = 0" | 
| 816 | and mod_less: " 0 \<le> a \<Longrightarrow> a < b \<Longrightarrow> a mod b = a" | |
| 817 | and div_positive: "0 < b \<Longrightarrow> b \<le> a \<Longrightarrow> a div b > 0" | |
| 818 | and mod_less_eq_dividend: "0 \<le> a \<Longrightarrow> a mod b \<le> a" | |
| 819 | and pos_mod_bound: "0 < b \<Longrightarrow> a mod b < b" | |
| 820 | and pos_mod_sign: "0 < b \<Longrightarrow> 0 \<le> a mod b" | |
| 821 | and mod_mult2_eq: "0 \<le> c \<Longrightarrow> a mod (b * c) = b * (a div b mod c) + a mod b" | |
| 822 | and div_mult2_eq: "0 \<le> c \<Longrightarrow> a div (b * c) = a div b div c" | |
| 823 | assumes discrete: "a < b \<longleftrightarrow> a + 1 \<le> b" | |
| 824 | fixes divmod :: "num \<Rightarrow> num \<Rightarrow> 'a \<times> 'a" | |
| 825 | and divmod_step :: "num \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<times> 'a" | |
| 826 | assumes divmod_def: "divmod m n = (numeral m div numeral n, numeral m mod numeral n)" | |
| 827 | and divmod_step_def: "divmod_step l qr = (let (q, r) = qr | |
| 828 | in if r \<ge> numeral l then (2 * q + 1, r - numeral l) | |
| 829 | else (2 * q, r))" | |
| 830 | \<comment> \<open>These are conceptually definitions but force generated code | |
| 831 | to be monomorphic wrt. particular instances of this class which | |
| 832 | yields a significant speedup.\<close> | |
| 833 | begin | |
| 834 | ||
| 835 | lemma divmod_digit_1: | |
| 836 | assumes "0 \<le> a" "0 < b" and "b \<le> a mod (2 * b)" | |
| 837 | shows "2 * (a div (2 * b)) + 1 = a div b" (is "?P") | |
| 838 | and "a mod (2 * b) - b = a mod b" (is "?Q") | |
| 839 | proof - | |
| 840 | from assms mod_less_eq_dividend [of a "2 * b"] have "b \<le> a" | |
| 841 | by (auto intro: trans) | |
| 842 | with \<open>0 < b\<close> have "0 < a div b" by (auto intro: div_positive) | |
| 843 | then have [simp]: "1 \<le> a div b" by (simp add: discrete) | |
| 844 | with \<open>0 < b\<close> have mod_less: "a mod b < b" by (simp add: pos_mod_bound) | |
| 845 | define w where "w = a div b mod 2" | |
| 846 | then have w_exhaust: "w = 0 \<or> w = 1" by auto | |
| 847 | have mod_w: "a mod (2 * b) = a mod b + b * w" | |
| 848 | by (simp add: w_def mod_mult2_eq ac_simps) | |
| 849 | from assms w_exhaust have "w = 1" | |
| 850 | by (auto simp add: mod_w) (insert mod_less, auto) | |
| 851 | with mod_w have mod: "a mod (2 * b) = a mod b + b" by simp | |
| 852 | have "2 * (a div (2 * b)) = a div b - w" | |
| 853 | by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) | |
| 854 | with \<open>w = 1\<close> have div: "2 * (a div (2 * b)) = a div b - 1" by simp | |
| 855 | then show ?P and ?Q | |
| 856 | by (simp_all add: div mod add_implies_diff [symmetric]) | |
| 857 | qed | |
| 858 | ||
| 859 | lemma divmod_digit_0: | |
| 860 | assumes "0 < b" and "a mod (2 * b) < b" | |
| 861 | shows "2 * (a div (2 * b)) = a div b" (is "?P") | |
| 862 | and "a mod (2 * b) = a mod b" (is "?Q") | |
| 863 | proof - | |
| 864 | define w where "w = a div b mod 2" | |
| 865 | then have w_exhaust: "w = 0 \<or> w = 1" by auto | |
| 866 | have mod_w: "a mod (2 * b) = a mod b + b * w" | |
| 867 | by (simp add: w_def mod_mult2_eq ac_simps) | |
| 868 | moreover have "b \<le> a mod b + b" | |
| 869 | proof - | |
| 870 | from \<open>0 < b\<close> pos_mod_sign have "0 \<le> a mod b" by blast | |
| 871 | then have "0 + b \<le> a mod b + b" by (rule add_right_mono) | |
| 872 | then show ?thesis by simp | |
| 873 | qed | |
| 874 | moreover note assms w_exhaust | |
| 875 | ultimately have "w = 0" by auto | |
| 876 | with mod_w have mod: "a mod (2 * b) = a mod b" by simp | |
| 877 | have "2 * (a div (2 * b)) = a div b - w" | |
| 878 | by (simp add: w_def div_mult2_eq minus_mod_eq_mult_div ac_simps) | |
| 879 | with \<open>w = 0\<close> have div: "2 * (a div (2 * b)) = a div b" by simp | |
| 880 | then show ?P and ?Q | |
| 881 | by (simp_all add: div mod) | |
| 882 | qed | |
| 883 | ||
| 69785 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 884 | lemma mod_double_modulus: | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 885 | assumes "m > 0" "x \<ge> 0" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 886 | shows "x mod (2 * m) = x mod m \<or> x mod (2 * m) = x mod m + m" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 887 | proof (cases "x mod (2 * m) < m") | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 888 | case True | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 889 | thus ?thesis using assms using divmod_digit_0(2)[of m x] by auto | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 890 | next | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 891 | case False | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 892 | hence *: "x mod (2 * m) - m = x mod m" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 893 | using assms by (intro divmod_digit_1) auto | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 894 | hence "x mod (2 * m) = x mod m + m" | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 895 | by (subst * [symmetric], subst le_add_diff_inverse2) (use False in auto) | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 896 | thus ?thesis by simp | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 897 | qed | 
| 
9e326f6f8a24
More material for HOL-Number_Theory: ord, Carmichael's function, primitive roots
 Manuel Eberl <eberlm@in.tum.de> parents: 
69695diff
changeset | 898 | |
| 68253 | 899 | lemma fst_divmod: | 
| 900 | "fst (divmod m n) = numeral m div numeral n" | |
| 901 | by (simp add: divmod_def) | |
| 902 | ||
| 903 | lemma snd_divmod: | |
| 904 | "snd (divmod m n) = numeral m mod numeral n" | |
| 905 | by (simp add: divmod_def) | |
| 906 | ||
| 907 | text \<open> | |
| 908 | This is a formulation of one step (referring to one digit position) | |
| 909 | in school-method division: compare the dividend at the current | |
| 910 | digit position with the remainder from previous division steps | |
| 911 | and evaluate accordingly. | |
| 912 | \<close> | |
| 913 | ||
| 914 | lemma divmod_step_eq [simp]: | |
| 915 | "divmod_step l (q, r) = (if numeral l \<le> r | |
| 916 | then (2 * q + 1, r - numeral l) else (2 * q, r))" | |
| 917 | by (simp add: divmod_step_def) | |
| 918 | ||
| 919 | text \<open> | |
| 920 | This is a formulation of school-method division. | |
| 921 | If the divisor is smaller than the dividend, terminate. | |
| 922 | If not, shift the dividend to the right until termination | |
| 923 | occurs and then reiterate single division steps in the | |
| 924 | opposite direction. | |
| 925 | \<close> | |
| 926 | ||
| 927 | lemma divmod_divmod_step: | |
| 928 | "divmod m n = (if m < n then (0, numeral m) | |
| 929 | else divmod_step n (divmod m (Num.Bit0 n)))" | |
| 930 | proof (cases "m < n") | |
| 931 | case True then have "numeral m < numeral n" by simp | |
| 932 | then show ?thesis | |
| 933 | by (simp add: prod_eq_iff div_less mod_less fst_divmod snd_divmod) | |
| 934 | next | |
| 935 | case False | |
| 936 | have "divmod m n = | |
| 937 | divmod_step n (numeral m div (2 * numeral n), | |
| 938 | numeral m mod (2 * numeral n))" | |
| 939 | proof (cases "numeral n \<le> numeral m mod (2 * numeral n)") | |
| 940 | case True | |
| 941 | with divmod_step_eq | |
| 942 | have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) = | |
| 943 | (2 * (numeral m div (2 * numeral n)) + 1, numeral m mod (2 * numeral n) - numeral n)" | |
| 944 | by simp | |
| 945 | moreover from True divmod_digit_1 [of "numeral m" "numeral n"] | |
| 946 | have "2 * (numeral m div (2 * numeral n)) + 1 = numeral m div numeral n" | |
| 947 | and "numeral m mod (2 * numeral n) - numeral n = numeral m mod numeral n" | |
| 948 | by simp_all | |
| 949 | ultimately show ?thesis by (simp only: divmod_def) | |
| 950 | next | |
| 951 | case False then have *: "numeral m mod (2 * numeral n) < numeral n" | |
| 952 | by (simp add: not_le) | |
| 953 | with divmod_step_eq | |
| 954 | have "divmod_step n (numeral m div (2 * numeral n), numeral m mod (2 * numeral n)) = | |
| 955 | (2 * (numeral m div (2 * numeral n)), numeral m mod (2 * numeral n))" | |
| 956 | by auto | |
| 957 | moreover from * divmod_digit_0 [of "numeral n" "numeral m"] | |
| 958 | have "2 * (numeral m div (2 * numeral n)) = numeral m div numeral n" | |
| 959 | and "numeral m mod (2 * numeral n) = numeral m mod numeral n" | |
| 960 | by (simp_all only: zero_less_numeral) | |
| 961 | ultimately show ?thesis by (simp only: divmod_def) | |
| 962 | qed | |
| 963 | then have "divmod m n = | |
| 964 | divmod_step n (numeral m div numeral (Num.Bit0 n), | |
| 965 | numeral m mod numeral (Num.Bit0 n))" | |
| 966 | by (simp only: numeral.simps distrib mult_1) | |
| 967 | then have "divmod m n = divmod_step n (divmod m (Num.Bit0 n))" | |
| 968 | by (simp add: divmod_def) | |
| 969 | with False show ?thesis by simp | |
| 970 | qed | |
| 971 | ||
| 972 | text \<open>The division rewrite proper -- first, trivial results involving \<open>1\<close>\<close> | |
| 973 | ||
| 974 | lemma divmod_trivial [simp]: | |
| 975 | "divmod Num.One Num.One = (numeral Num.One, 0)" | |
| 976 | "divmod (Num.Bit0 m) Num.One = (numeral (Num.Bit0 m), 0)" | |
| 977 | "divmod (Num.Bit1 m) Num.One = (numeral (Num.Bit1 m), 0)" | |
| 978 | "divmod num.One (num.Bit0 n) = (0, Numeral1)" | |
| 979 | "divmod num.One (num.Bit1 n) = (0, Numeral1)" | |
| 980 | using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def) | |
| 981 | ||
| 982 | text \<open>Division by an even number is a right-shift\<close> | |
| 983 | ||
| 984 | lemma divmod_cancel [simp]: | |
| 985 | "divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r))" (is ?P) | |
| 986 | "divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r + 1))" (is ?Q) | |
| 987 | proof - | |
| 988 | have *: "\<And>q. numeral (Num.Bit0 q) = 2 * numeral q" | |
| 989 | "\<And>q. numeral (Num.Bit1 q) = 2 * numeral q + 1" | |
| 990 | by (simp_all only: numeral_mult numeral.simps distrib) simp_all | |
| 991 | have "1 div 2 = 0" "1 mod 2 = 1" by (auto intro: div_less mod_less) | |
| 992 | then show ?P and ?Q | |
| 993 | by (simp_all add: fst_divmod snd_divmod prod_eq_iff split_def * [of m] * [of n] mod_mult_mult1 | |
| 994 | div_mult2_eq [of _ _ 2] mod_mult2_eq [of _ _ 2] | |
| 995 | add.commute del: numeral_times_numeral) | |
| 996 | qed | |
| 997 | ||
| 998 | text \<open>The really hard work\<close> | |
| 999 | ||
| 1000 | lemma divmod_steps [simp]: | |
| 1001 | "divmod (num.Bit0 m) (num.Bit1 n) = | |
| 1002 | (if m \<le> n then (0, numeral (num.Bit0 m)) | |
| 1003 | else divmod_step (num.Bit1 n) | |
| 1004 | (divmod (num.Bit0 m) | |
| 1005 | (num.Bit0 (num.Bit1 n))))" | |
| 1006 | "divmod (num.Bit1 m) (num.Bit1 n) = | |
| 1007 | (if m < n then (0, numeral (num.Bit1 m)) | |
| 1008 | else divmod_step (num.Bit1 n) | |
| 1009 | (divmod (num.Bit1 m) | |
| 1010 | (num.Bit0 (num.Bit1 n))))" | |
| 1011 | by (simp_all add: divmod_divmod_step) | |
| 1012 | ||
| 1013 | lemmas divmod_algorithm_code = divmod_step_eq divmod_trivial divmod_cancel divmod_steps | |
| 1014 | ||
| 1015 | text \<open>Special case: divisibility\<close> | |
| 1016 | ||
| 1017 | definition divides_aux :: "'a \<times> 'a \<Rightarrow> bool" | |
| 1018 | where | |
| 1019 | "divides_aux qr \<longleftrightarrow> snd qr = 0" | |
| 1020 | ||
| 1021 | lemma divides_aux_eq [simp]: | |
| 1022 | "divides_aux (q, r) \<longleftrightarrow> r = 0" | |
| 1023 | by (simp add: divides_aux_def) | |
| 1024 | ||
| 1025 | lemma dvd_numeral_simp [simp]: | |
| 1026 | "numeral m dvd numeral n \<longleftrightarrow> divides_aux (divmod n m)" | |
| 1027 | by (simp add: divmod_def mod_eq_0_iff_dvd) | |
| 1028 | ||
| 1029 | text \<open>Generic computation of quotient and remainder\<close> | |
| 1030 | ||
| 1031 | lemma numeral_div_numeral [simp]: | |
| 1032 | "numeral k div numeral l = fst (divmod k l)" | |
| 1033 | by (simp add: fst_divmod) | |
| 1034 | ||
| 1035 | lemma numeral_mod_numeral [simp]: | |
| 1036 | "numeral k mod numeral l = snd (divmod k l)" | |
| 1037 | by (simp add: snd_divmod) | |
| 1038 | ||
| 1039 | lemma one_div_numeral [simp]: | |
| 1040 | "1 div numeral n = fst (divmod num.One n)" | |
| 1041 | by (simp add: fst_divmod) | |
| 1042 | ||
| 1043 | lemma one_mod_numeral [simp]: | |
| 1044 | "1 mod numeral n = snd (divmod num.One n)" | |
| 1045 | by (simp add: snd_divmod) | |
| 1046 | ||
| 1047 | text \<open>Computing congruences modulo \<open>2 ^ q\<close>\<close> | |
| 1048 | ||
| 1049 | lemma cong_exp_iff_simps: | |
| 1050 | "numeral n mod numeral Num.One = 0 | |
| 1051 | \<longleftrightarrow> True" | |
| 1052 | "numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) = 0 | |
| 1053 | \<longleftrightarrow> numeral n mod numeral q = 0" | |
| 1054 | "numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) = 0 | |
| 1055 | \<longleftrightarrow> False" | |
| 1056 | "numeral m mod numeral Num.One = (numeral n mod numeral Num.One) | |
| 1057 | \<longleftrightarrow> True" | |
| 1058 | "numeral Num.One mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) | |
| 1059 | \<longleftrightarrow> True" | |
| 1060 | "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) | |
| 1061 | \<longleftrightarrow> False" | |
| 1062 | "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) | |
| 1063 | \<longleftrightarrow> (numeral n mod numeral q) = 0" | |
| 1064 | "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) | |
| 1065 | \<longleftrightarrow> False" | |
| 1066 | "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) | |
| 1067 | \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)" | |
| 1068 | "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) | |
| 1069 | \<longleftrightarrow> False" | |
| 1070 | "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) | |
| 1071 | \<longleftrightarrow> (numeral m mod numeral q) = 0" | |
| 1072 | "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) | |
| 1073 | \<longleftrightarrow> False" | |
| 1074 | "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) | |
| 1075 | \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)" | |
| 1076 | by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even]) | |
| 1077 | ||
| 1078 | end | |
| 1079 | ||
| 1080 | hide_fact (open) div_less mod_less mod_less_eq_dividend mod_mult2_eq div_mult2_eq | |
| 1081 | ||
| 1082 | instantiation nat :: unique_euclidean_semiring_numeral | |
| 1083 | begin | |
| 1084 | ||
| 1085 | definition divmod_nat :: "num \<Rightarrow> num \<Rightarrow> nat \<times> nat" | |
| 1086 | where | |
| 1087 | divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)" | |
| 1088 | ||
| 1089 | definition divmod_step_nat :: "num \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat" | |
| 1090 | where | |
| 1091 | "divmod_step_nat l qr = (let (q, r) = qr | |
| 1092 | in if r \<ge> numeral l then (2 * q + 1, r - numeral l) | |
| 1093 | else (2 * q, r))" | |
| 1094 | ||
| 1095 | instance by standard | |
| 1096 | (auto simp add: divmod'_nat_def divmod_step_nat_def div_greater_zero_iff div_mult2_eq mod_mult2_eq) | |
| 1097 | ||
| 1098 | end | |
| 1099 | ||
| 1100 | declare divmod_algorithm_code [where ?'a = nat, code] | |
| 1101 | ||
| 1102 | lemma Suc_0_div_numeral [simp]: | |
| 1103 | fixes k l :: num | |
| 1104 | shows "Suc 0 div numeral k = fst (divmod Num.One k)" | |
| 1105 | by (simp_all add: fst_divmod) | |
| 1106 | ||
| 1107 | lemma Suc_0_mod_numeral [simp]: | |
| 1108 | fixes k l :: num | |
| 1109 | shows "Suc 0 mod numeral k = snd (divmod Num.One k)" | |
| 1110 | by (simp_all add: snd_divmod) | |
| 1111 | ||
| 1112 | instantiation int :: unique_euclidean_semiring_numeral | |
| 1113 | begin | |
| 1114 | ||
| 1115 | definition divmod_int :: "num \<Rightarrow> num \<Rightarrow> int \<times> int" | |
| 1116 | where | |
| 1117 | "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)" | |
| 1118 | ||
| 1119 | definition divmod_step_int :: "num \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int" | |
| 1120 | where | |
| 1121 | "divmod_step_int l qr = (let (q, r) = qr | |
| 1122 | in if r \<ge> numeral l then (2 * q + 1, r - numeral l) | |
| 1123 | else (2 * q, r))" | |
| 1124 | ||
| 1125 | instance | |
| 1126 | by standard (auto intro: zmod_le_nonneg_dividend simp add: divmod_int_def divmod_step_int_def | |
| 1127 | pos_imp_zdiv_pos_iff zmod_zmult2_eq zdiv_zmult2_eq) | |
| 1128 | ||
| 1129 | end | |
| 1130 | ||
| 1131 | declare divmod_algorithm_code [where ?'a = int, code] | |
| 1132 | ||
| 1133 | context | |
| 1134 | begin | |
| 1135 | ||
| 1136 | qualified definition adjust_div :: "int \<times> int \<Rightarrow> int" | |
| 1137 | where | |
| 1138 | "adjust_div qr = (let (q, r) = qr in q + of_bool (r \<noteq> 0))" | |
| 1139 | ||
| 1140 | qualified lemma adjust_div_eq [simp, code]: | |
| 1141 | "adjust_div (q, r) = q + of_bool (r \<noteq> 0)" | |
| 1142 | by (simp add: adjust_div_def) | |
| 1143 | ||
| 1144 | qualified definition adjust_mod :: "int \<Rightarrow> int \<Rightarrow> int" | |
| 1145 | where | |
| 1146 | [simp]: "adjust_mod l r = (if r = 0 then 0 else l - r)" | |
| 1147 | ||
| 1148 | lemma minus_numeral_div_numeral [simp]: | |
| 1149 | "- numeral m div numeral n = - (adjust_div (divmod m n) :: int)" | |
| 1150 | proof - | |
| 1151 | have "int (fst (divmod m n)) = fst (divmod m n)" | |
| 1152 | by (simp only: fst_divmod divide_int_def) auto | |
| 1153 | then show ?thesis | |
| 1154 | by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) | |
| 1155 | qed | |
| 1156 | ||
| 1157 | lemma minus_numeral_mod_numeral [simp]: | |
| 1158 | "- numeral m mod numeral n = adjust_mod (numeral n) (snd (divmod m n) :: int)" | |
| 1159 | proof (cases "snd (divmod m n) = (0::int)") | |
| 1160 | case True | |
| 1161 | then show ?thesis | |
| 1162 | by (simp add: mod_eq_0_iff_dvd divides_aux_def) | |
| 1163 | next | |
| 1164 | case False | |
| 1165 | then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)" | |
| 1166 | by (simp only: snd_divmod modulo_int_def) auto | |
| 1167 | then show ?thesis | |
| 1168 | by (simp add: divides_aux_def adjust_div_def) | |
| 1169 | (simp add: divides_aux_def modulo_int_def) | |
| 1170 | qed | |
| 1171 | ||
| 1172 | lemma numeral_div_minus_numeral [simp]: | |
| 1173 | "numeral m div - numeral n = - (adjust_div (divmod m n) :: int)" | |
| 1174 | proof - | |
| 1175 | have "int (fst (divmod m n)) = fst (divmod m n)" | |
| 1176 | by (simp only: fst_divmod divide_int_def) auto | |
| 1177 | then show ?thesis | |
| 1178 | by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) | |
| 1179 | qed | |
| 1180 | ||
| 1181 | lemma numeral_mod_minus_numeral [simp]: | |
| 1182 | "numeral m mod - numeral n = - adjust_mod (numeral n) (snd (divmod m n) :: int)" | |
| 1183 | proof (cases "snd (divmod m n) = (0::int)") | |
| 1184 | case True | |
| 1185 | then show ?thesis | |
| 1186 | by (simp add: mod_eq_0_iff_dvd divides_aux_def) | |
| 1187 | next | |
| 1188 | case False | |
| 1189 | then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)" | |
| 1190 | by (simp only: snd_divmod modulo_int_def) auto | |
| 1191 | then show ?thesis | |
| 1192 | by (simp add: divides_aux_def adjust_div_def) | |
| 1193 | (simp add: divides_aux_def modulo_int_def) | |
| 1194 | qed | |
| 1195 | ||
| 1196 | lemma minus_one_div_numeral [simp]: | |
| 1197 | "- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)" | |
| 1198 | using minus_numeral_div_numeral [of Num.One n] by simp | |
| 1199 | ||
| 1200 | lemma minus_one_mod_numeral [simp]: | |
| 1201 | "- 1 mod numeral n = adjust_mod (numeral n) (snd (divmod Num.One n) :: int)" | |
| 1202 | using minus_numeral_mod_numeral [of Num.One n] by simp | |
| 1203 | ||
| 1204 | lemma one_div_minus_numeral [simp]: | |
| 1205 | "1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)" | |
| 1206 | using numeral_div_minus_numeral [of Num.One n] by simp | |
| 1207 | ||
| 1208 | lemma one_mod_minus_numeral [simp]: | |
| 1209 | "1 mod - numeral n = - adjust_mod (numeral n) (snd (divmod Num.One n) :: int)" | |
| 1210 | using numeral_mod_minus_numeral [of Num.One n] by simp | |
| 1211 | ||
| 1212 | end | |
| 1213 | ||
| 1214 | lemma div_positive_int: | |
| 1215 | "k div l > 0" if "k \<ge> l" and "l > 0" for k l :: int | |
| 1216 | using that div_positive [of l k] by blast | |
| 1217 | ||
| 1218 | ||
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1219 | subsubsection \<open>Dedicated simproc for calculation\<close> | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1220 | |
| 60758 | 1221 | text \<open> | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1222 | There is space for improvement here: the calculation itself | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66806diff
changeset | 1223 | could be carried out outside the logic, and a generic simproc | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1224 | (simplifier setup) for generic calculation would be helpful. | 
| 60758 | 1225 | \<close> | 
| 53067 
ee0b7c2315d2
type class for generic division algorithm on numerals
 haftmann parents: 
53066diff
changeset | 1226 | |
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1227 | simproc_setup numeral_divmod | 
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66801diff
changeset | 1228 |   ("0 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 0 :: 'a :: unique_euclidean_semiring_numeral" |
 | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66801diff
changeset | 1229 | "0 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "0 mod 1 :: 'a :: unique_euclidean_semiring_numeral" | | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1230 | "0 div - 1 :: int" | "0 mod - 1 :: int" | | 
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66801diff
changeset | 1231 | "0 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "0 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1232 | "0 div - numeral b :: int" | "0 mod - numeral b :: int" | | 
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66801diff
changeset | 1233 | "1 div 0 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 0 :: 'a :: unique_euclidean_semiring_numeral" | | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66801diff
changeset | 1234 | "1 div 1 :: 'a :: unique_euclidean_semiring_numeral" | "1 mod 1 :: 'a :: unique_euclidean_semiring_numeral" | | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1235 | "1 div - 1 :: int" | "1 mod - 1 :: int" | | 
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66801diff
changeset | 1236 | "1 div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "1 mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1237 | "1 div - numeral b :: int" |"1 mod - numeral b :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1238 | "- 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 | 1239 | "- 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 | 1240 | "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" | | 
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66801diff
changeset | 1241 | "numeral a div 0 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 0 :: 'a :: unique_euclidean_semiring_numeral" | | 
| 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66801diff
changeset | 1242 | "numeral a div 1 :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod 1 :: 'a :: unique_euclidean_semiring_numeral" | | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1243 | "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" | | 
| 66806 
a4e82b58d833
abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
 haftmann parents: 
66801diff
changeset | 1244 | "numeral a div numeral b :: 'a :: unique_euclidean_semiring_numeral" | "numeral a mod numeral b :: 'a :: unique_euclidean_semiring_numeral" | | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1245 | "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 | 1246 | "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1247 | "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1248 | "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" | | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1249 | "- 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 | 1250 | "- 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 | 1251 | \<open> let | 
| 69593 | 1252 | val if_cong = the (Code.get_case_cong \<^theory> \<^const_name>\<open>If\<close>); | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1253 | fun successful_rewrite ctxt ct = | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1254 | let | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1255 | val thm = Simplifier.rewrite ctxt ct | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1256 | 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 | 1257 | in fn phi => | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1258 | let | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1259 |       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 | 1260 | 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 | 1261 | one_div_minus_numeral one_mod_minus_numeral | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1262 | 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 | 1263 | numeral_div_minus_numeral numeral_mod_minus_numeral | 
| 60930 | 1264 | 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 | 1265 | 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 | 1266 | divmod_cancel divmod_steps divmod_step_eq fst_conv snd_conv numeral_One | 
| 60930 | 1267 | 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 | 1268 | minus_minus numeral_times_numeral mult_zero_right mult_1_right} | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1269 |         @ [@{lemma "0 = 0 \<longleftrightarrow> True" by simp}]);
 | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1270 | fun prepare_simpset ctxt = HOL_ss |> Simplifier.simpset_map ctxt | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1271 | (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps) | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1272 | in fn ctxt => successful_rewrite (Simplifier.put_simpset (prepare_simpset ctxt) ctxt) end | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
68644diff
changeset | 1273 | end | 
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1274 | \<close> | 
| 34126 | 1275 | |
| 35673 | 1276 | |
| 60758 | 1277 | subsubsection \<open>Code generation\<close> | 
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1278 | |
| 68253 | 1279 | definition divmod_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" | 
| 1280 | where "divmod_nat m n = (m div n, m mod n)" | |
| 1281 | ||
| 1282 | lemma fst_divmod_nat [simp]: | |
| 1283 | "fst (divmod_nat m n) = m div n" | |
| 1284 | by (simp add: divmod_nat_def) | |
| 1285 | ||
| 1286 | lemma snd_divmod_nat [simp]: | |
| 1287 | "snd (divmod_nat m n) = m mod n" | |
| 1288 | by (simp add: divmod_nat_def) | |
| 1289 | ||
| 1290 | lemma divmod_nat_if [code]: | |
| 1291 | "Divides.divmod_nat m n = (if n = 0 \<or> m < n then (0, m) else | |
| 1292 | let (q, r) = Divides.divmod_nat (m - n) n in (Suc q, r))" | |
| 1293 | by (simp add: prod_eq_iff case_prod_beta not_less le_div_geq le_mod_geq) | |
| 1294 | ||
| 1295 | lemma [code]: | |
| 1296 | "m div n = fst (divmod_nat m n)" | |
| 1297 | "m mod n = snd (divmod_nat m n)" | |
| 1298 | by simp_all | |
| 1299 | ||
| 60868 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1300 | lemma [code]: | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1301 | fixes k :: int | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1302 | shows | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1303 | "k div 0 = 0" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1304 | "k mod 0 = k" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1305 | "0 div k = 0" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1306 | "0 mod k = 0" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1307 | "k div Int.Pos Num.One = k" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1308 | "k mod Int.Pos Num.One = 0" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1309 | "k div Int.Neg Num.One = - k" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1310 | "k mod Int.Neg Num.One = 0" | 
| 
dd18c33c001e
direct bootstrap of integer division from natural division
 haftmann parents: 
60867diff
changeset | 1311 | "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 | 1312 | "Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)" | 
| 60930 | 1313 | "Int.Neg m div Int.Pos n = - (Divides.adjust_div (divmod m n) :: int)" | 
| 1314 | "Int.Neg m mod Int.Pos n = Divides.adjust_mod (Int.Pos n) (snd (divmod m n) :: int)" | |
| 1315 | "Int.Pos m div Int.Neg n = - (Divides.adjust_div (divmod m n) :: int)" | |
| 1316 | "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 | 1317 | "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 | 1318 | "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 | 1319 | by simp_all | 
| 53069 
d165213e3924
execution of int division by class semiring_numeral_div, replacing pdivmod by divmod_abs
 haftmann parents: 
53068diff
changeset | 1320 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52398diff
changeset | 1321 | code_identifier | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52398diff
changeset | 1322 | code_module Divides \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith | 
| 33364 | 1323 | |
| 64246 | 1324 | |
| 68253 | 1325 | subsection \<open>Lemmas of doubtful value\<close> | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66806diff
changeset | 1326 | |
| 68631 | 1327 | lemma div_geq: "m div n = Suc ((m - n) div n)" if "0 < n" and " \<not> m < n" for m n :: nat | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66806diff
changeset | 1328 | by (rule le_div_geq) (use that in \<open>simp_all add: not_less\<close>) | 
| 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66806diff
changeset | 1329 | |
| 68631 | 1330 | lemma mod_geq: "m mod n = (m - n) mod n" if "\<not> m < n" for m n :: nat | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66806diff
changeset | 1331 | by (rule le_mod_geq) (use that in \<open>simp add: not_less\<close>) | 
| 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66806diff
changeset | 1332 | |
| 68631 | 1333 | lemma mod_eq_0D: "\<exists>q. m = d * q" if "m mod d = 0" for m d :: nat | 
| 1334 | using that by (auto simp add: mod_eq_0_iff_dvd) | |
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 1335 | |
| 69695 | 1336 | lemma pos_mod_conj: "0 < b \<Longrightarrow> 0 \<le> a mod b \<and> a mod b < b" for a b :: int | 
| 1337 | by simp | |
| 1338 | ||
| 1339 | lemma neg_mod_conj: "b < 0 \<Longrightarrow> a mod b \<le> 0 \<and> b < a mod b" for a b :: int | |
| 1340 | by simp | |
| 1341 | ||
| 1342 | lemma zmod_eq_0_iff: "m mod d = 0 \<longleftrightarrow> (\<exists>q. m = d * q)" for m d :: int | |
| 1343 | by (auto simp add: mod_eq_0_iff_dvd) | |
| 1344 | ||
| 1345 | (* REVISIT: should this be generalized to all semiring_div types? *) | |
| 1346 | lemma zmod_eq_0D [dest!]: "\<exists>q. m = d * q" if "m mod d = 0" for m d :: int | |
| 1347 | using that by auto | |
| 1348 | ||
| 33361 
1f18de40b43f
combined former theories Divides and IntDiv to one theory Divides
 haftmann parents: 
33340diff
changeset | 1349 | end |