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