| author | wenzelm | 
| Sun, 27 Oct 2024 12:32:40 +0100 | |
| changeset 81275 | 5ed639c16ce7 | 
| parent 79585 | dafb3d343cd6 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Parity.thy | 
| 2 | Author: Jeremy Avigad | |
| 3 | Author: Jacques D. Fleuriot | |
| 21256 | 4 | *) | 
| 5 | ||
| 60758 | 6 | section \<open>Parity in rings and semirings\<close> | 
| 21256 | 7 | |
| 8 | theory Parity | |
| 77061 
5de3772609ea
generalized theory name: euclidean division denotes one particular division definition on integers
 haftmann parents: 
76387diff
changeset | 9 | imports Euclidean_Rings | 
| 21256 | 10 | begin | 
| 11 | ||
| 61799 | 12 | subsection \<open>Ring structures with parity and \<open>even\<close>/\<open>odd\<close> predicates\<close> | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 13 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 14 | class semiring_parity = comm_semiring_1 + semiring_modulo + | 
| 79118 | 15 | assumes mod_2_eq_odd: \<open>a mod 2 = of_bool (\<not> 2 dvd a)\<close> | 
| 16 | and odd_one [simp]: \<open>\<not> 2 dvd 1\<close> | |
| 79555 | 17 | and even_half_succ_eq [simp]: \<open>2 dvd a \<Longrightarrow> (1 + a) div 2 = a div 2\<close> | 
| 66839 | 18 | begin | 
| 19 | ||
| 58740 | 20 | abbreviation even :: "'a \<Rightarrow> bool" | 
| 79118 | 21 | where \<open>even a \<equiv> 2 dvd a\<close> | 
| 54228 | 22 | |
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 23 | abbreviation odd :: "'a \<Rightarrow> bool" | 
| 79118 | 24 | where \<open>odd a \<equiv> \<not> 2 dvd a\<close> | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 25 | |
| 76387 | 26 | end | 
| 27 | ||
| 28 | class ring_parity = ring + semiring_parity | |
| 29 | begin | |
| 30 | ||
| 31 | subclass comm_ring_1 .. | |
| 32 | ||
| 33 | end | |
| 34 | ||
| 35 | instance nat :: semiring_parity | |
| 79555 | 36 | by standard (auto simp add: dvd_eq_mod_eq_0) | 
| 76387 | 37 | |
| 38 | instance int :: ring_parity | |
| 39 | by standard (auto simp add: dvd_eq_mod_eq_0) | |
| 40 | ||
| 41 | context semiring_parity | |
| 42 | begin | |
| 43 | ||
| 58690 | 44 | lemma evenE [elim?]: | 
| 79118 | 45 | assumes \<open>even a\<close> | 
| 46 | obtains b where \<open>a = 2 * b\<close> | |
| 58740 | 47 | using assms by (rule dvdE) | 
| 58690 | 48 | |
| 58681 | 49 | lemma oddE [elim?]: | 
| 79118 | 50 | assumes \<open>odd a\<close> | 
| 51 | obtains b where \<open>a = 2 * b + 1\<close> | |
| 58787 | 52 | proof - | 
| 79118 | 53 | have \<open>a = 2 * (a div 2) + a mod 2\<close> | 
| 66815 | 54 | by (simp add: mult_div_mod_eq) | 
| 79118 | 55 | with assms have \<open>a = 2 * (a div 2) + 1\<close> | 
| 56 | by (simp add: mod_2_eq_odd) | |
| 57 | then show thesis .. | |
| 66815 | 58 | qed | 
| 59 | ||
| 67816 | 60 | lemma of_bool_odd_eq_mod_2: | 
| 79117 | 61 | \<open>of_bool (odd a) = a mod 2\<close> | 
| 62 | by (simp add: mod_2_eq_odd) | |
| 63 | ||
| 64 | lemma odd_of_bool_self [simp]: | |
| 65 | \<open>odd (of_bool p) \<longleftrightarrow> p\<close> | |
| 66 | by (cases p) simp_all | |
| 67 | ||
| 68 | lemma not_mod_2_eq_0_eq_1 [simp]: | |
| 69 | \<open>a mod 2 \<noteq> 0 \<longleftrightarrow> a mod 2 = 1\<close> | |
| 70 | by (simp add: mod_2_eq_odd) | |
| 71 | ||
| 72 | lemma not_mod_2_eq_1_eq_0 [simp]: | |
| 73 | \<open>a mod 2 \<noteq> 1 \<longleftrightarrow> a mod 2 = 0\<close> | |
| 67816 | 74 | by (simp add: mod_2_eq_odd) | 
| 75 | ||
| 79118 | 76 | lemma even_iff_mod_2_eq_zero: | 
| 77 | \<open>2 dvd a \<longleftrightarrow> a mod 2 = 0\<close> | |
| 78 | by (simp add: mod_2_eq_odd) | |
| 79 | ||
| 80 | lemma odd_iff_mod_2_eq_one: | |
| 81 | \<open>\<not> 2 dvd a \<longleftrightarrow> a mod 2 = 1\<close> | |
| 82 | by (simp add: mod_2_eq_odd) | |
| 83 | ||
| 71426 | 84 | lemma even_mod_2_iff [simp]: | 
| 85 | \<open>even (a mod 2) \<longleftrightarrow> even a\<close> | |
| 86 | by (simp add: mod_2_eq_odd) | |
| 87 | ||
| 88 | lemma mod2_eq_if: | |
| 89 | "a mod 2 = (if even a then 0 else 1)" | |
| 90 | by (simp add: mod_2_eq_odd) | |
| 91 | ||
| 79117 | 92 | lemma zero_mod_two_eq_zero [simp]: | 
| 93 | \<open>0 mod 2 = 0\<close> | |
| 94 | by (simp add: mod_2_eq_odd) | |
| 95 | ||
| 96 | lemma one_mod_two_eq_one [simp]: | |
| 97 | \<open>1 mod 2 = 1\<close> | |
| 98 | by (simp add: mod_2_eq_odd) | |
| 99 | ||
| 100 | lemma parity_cases [case_names even odd]: | |
| 101 | assumes \<open>even a \<Longrightarrow> a mod 2 = 0 \<Longrightarrow> P\<close> | |
| 102 | assumes \<open>odd a \<Longrightarrow> a mod 2 = 1 \<Longrightarrow> P\<close> | |
| 103 | shows P | |
| 104 | using assms by (auto simp add: mod_2_eq_odd) | |
| 105 | ||
| 66815 | 106 | lemma even_zero [simp]: | 
| 79117 | 107 | \<open>even 0\<close> | 
| 66815 | 108 | by (fact dvd_0_right) | 
| 109 | ||
| 110 | lemma odd_even_add: | |
| 111 | "even (a + b)" if "odd a" and "odd b" | |
| 112 | proof - | |
| 113 | from that obtain c d where "a = 2 * c + 1" and "b = 2 * d + 1" | |
| 114 | by (blast elim: oddE) | |
| 115 | then have "a + b = 2 * c + 2 * d + (1 + 1)" | |
| 116 | by (simp only: ac_simps) | |
| 117 | also have "\<dots> = 2 * (c + d + 1)" | |
| 118 | by (simp add: algebra_simps) | |
| 119 | finally show ?thesis .. | |
| 120 | qed | |
| 121 | ||
| 122 | lemma even_add [simp]: | |
| 123 | "even (a + b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)" | |
| 124 | by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add) | |
| 125 | ||
| 126 | lemma odd_add [simp]: | |
| 127 | "odd (a + b) \<longleftrightarrow> \<not> (odd a \<longleftrightarrow> odd b)" | |
| 128 | by simp | |
| 129 | ||
| 130 | lemma even_plus_one_iff [simp]: | |
| 131 | "even (a + 1) \<longleftrightarrow> odd a" | |
| 132 | by (auto simp add: dvd_add_right_iff intro: odd_even_add) | |
| 133 | ||
| 134 | lemma even_mult_iff [simp]: | |
| 135 | "even (a * b) \<longleftrightarrow> even a \<or> even b" (is "?P \<longleftrightarrow> ?Q") | |
| 136 | proof | |
| 137 | assume ?Q | |
| 138 | then show ?P | |
| 139 | by auto | |
| 140 | next | |
| 141 | assume ?P | |
| 142 | show ?Q | |
| 143 | proof (rule ccontr) | |
| 144 | assume "\<not> (even a \<or> even b)" | |
| 145 | then have "odd a" and "odd b" | |
| 146 | by auto | |
| 147 | then obtain r s where "a = 2 * r + 1" and "b = 2 * s + 1" | |
| 148 | by (blast elim: oddE) | |
| 149 | then have "a * b = (2 * r + 1) * (2 * s + 1)" | |
| 150 | by simp | |
| 151 | also have "\<dots> = 2 * (2 * r * s + r + s) + 1" | |
| 152 | by (simp add: algebra_simps) | |
| 153 | finally have "odd (a * b)" | |
| 154 | by simp | |
| 155 | with \<open>?P\<close> show False | |
| 156 | by auto | |
| 157 | qed | |
| 158 | qed | |
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 159 | |
| 63654 | 160 | lemma even_numeral [simp]: "even (numeral (Num.Bit0 n))" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 161 | proof - | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 162 | have "even (2 * numeral n)" | 
| 66815 | 163 | unfolding even_mult_iff by simp | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 164 | then have "even (numeral n + numeral n)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 165 | unfolding mult_2 . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 166 | then show ?thesis | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 167 | unfolding numeral.simps . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 168 | qed | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 169 | |
| 63654 | 170 | lemma odd_numeral [simp]: "odd (numeral (Num.Bit1 n))" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 171 | proof | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 172 | assume "even (numeral (num.Bit1 n))" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 173 | then have "even (numeral n + numeral n + 1)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 174 | unfolding numeral.simps . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 175 | then have "even (2 * numeral n + 1)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 176 | unfolding mult_2 . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 177 | then have "2 dvd numeral n * 2 + 1" | 
| 58740 | 178 | by (simp add: ac_simps) | 
| 63654 | 179 | then have "2 dvd 1" | 
| 180 | using dvd_add_times_triv_left_iff [of 2 "numeral n" 1] by simp | |
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 181 | then show False by simp | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 182 | qed | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 183 | |
| 71755 | 184 | lemma odd_numeral_BitM [simp]: | 
| 185 | \<open>odd (numeral (Num.BitM w))\<close> | |
| 186 | by (cases w) simp_all | |
| 187 | ||
| 63654 | 188 | lemma even_power [simp]: "even (a ^ n) \<longleftrightarrow> even a \<and> n > 0" | 
| 58680 | 189 | by (induct n) auto | 
| 190 | ||
| 76387 | 191 | lemma even_prod_iff: | 
| 192 | \<open>even (prod f A) \<longleftrightarrow> (\<exists>a\<in>A. even (f a))\<close> if \<open>finite A\<close> | |
| 193 | using that by (induction A) simp_all | |
| 194 | ||
| 79585 | 195 | lemma even_half_maybe_succ_eq [simp]: | 
| 196 | \<open>even a \<Longrightarrow> (of_bool b + a) div 2 = a div 2\<close> | |
| 197 | by simp | |
| 198 | ||
| 199 | lemma even_half_maybe_succ'_eq [simp]: | |
| 200 | \<open>even a \<Longrightarrow> (b mod 2 + a) div 2 = a div 2\<close> | |
| 201 | by (simp add: mod2_eq_if) | |
| 202 | ||
| 71412 | 203 | lemma mask_eq_sum_exp: | 
| 204 |   \<open>2 ^ n - 1 = (\<Sum>m\<in>{q. q < n}. 2 ^ m)\<close>
 | |
| 205 | proof - | |
| 206 |   have *: \<open>{q. q < Suc m} = insert m {q. q < m}\<close> for m
 | |
| 207 | by auto | |
| 208 |   have \<open>2 ^ n = (\<Sum>m\<in>{q. q < n}. 2 ^ m) + 1\<close>
 | |
| 209 | by (induction n) (simp_all add: ac_simps mult_2 *) | |
| 210 |   then have \<open>2 ^ n - 1 = (\<Sum>m\<in>{q. q < n}. 2 ^ m) + 1 - 1\<close>
 | |
| 211 | by simp | |
| 212 | then show ?thesis | |
| 213 | by simp | |
| 214 | qed | |
| 215 | ||
| 76387 | 216 | lemma (in -) mask_eq_sum_exp_nat: | 
| 217 |   \<open>2 ^ n - Suc 0 = (\<Sum>m\<in>{q. q < n}. 2 ^ m)\<close>
 | |
| 218 | using mask_eq_sum_exp [where ?'a = nat] by simp | |
| 219 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 220 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 221 | |
| 76387 | 222 | context ring_parity | 
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 223 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 224 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 225 | lemma even_minus: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 226 | "even (- a) \<longleftrightarrow> even a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 227 | by (fact dvd_minus_iff) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 228 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 229 | lemma even_diff [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 230 | "even (a - b) \<longleftrightarrow> even (a + b)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 231 | using even_add [of a "- b"] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 232 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 233 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 234 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 235 | |
| 69593 | 236 | subsection \<open>Instance for \<^typ>\<open>nat\<close>\<close> | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 237 | |
| 66815 | 238 | lemma even_Suc_Suc_iff [simp]: | 
| 239 | "even (Suc (Suc n)) \<longleftrightarrow> even n" | |
| 58787 | 240 | using dvd_add_triv_right_iff [of 2 n] by simp | 
| 58687 | 241 | |
| 66815 | 242 | lemma even_Suc [simp]: "even (Suc n) \<longleftrightarrow> odd n" | 
| 243 | using even_plus_one_iff [of n] by simp | |
| 58787 | 244 | |
| 66815 | 245 | lemma even_diff_nat [simp]: | 
| 246 | "even (m - n) \<longleftrightarrow> m < n \<or> even (m + n)" for m n :: nat | |
| 58787 | 247 | proof (cases "n \<le> m") | 
| 248 | case True | |
| 249 | then have "m - n + n * 2 = m + n" by (simp add: mult_2_right) | |
| 66815 | 250 | moreover have "even (m - n) \<longleftrightarrow> even (m - n + n * 2)" by simp | 
| 251 | ultimately have "even (m - n) \<longleftrightarrow> even (m + n)" by (simp only:) | |
| 58787 | 252 | then show ?thesis by auto | 
| 253 | next | |
| 254 | case False | |
| 255 | then show ?thesis by simp | |
| 63654 | 256 | qed | 
| 257 | ||
| 66815 | 258 | lemma odd_pos: | 
| 259 | "odd n \<Longrightarrow> 0 < n" for n :: nat | |
| 58690 | 260 | by (auto elim: oddE) | 
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 261 | |
| 66815 | 262 | lemma Suc_double_not_eq_double: | 
| 263 | "Suc (2 * m) \<noteq> 2 * n" | |
| 62597 | 264 | proof | 
| 265 | assume "Suc (2 * m) = 2 * n" | |
| 266 | moreover have "odd (Suc (2 * m))" and "even (2 * n)" | |
| 267 | by simp_all | |
| 268 | ultimately show False by simp | |
| 269 | qed | |
| 270 | ||
| 66815 | 271 | lemma double_not_eq_Suc_double: | 
| 272 | "2 * m \<noteq> Suc (2 * n)" | |
| 62597 | 273 | using Suc_double_not_eq_double [of n m] by simp | 
| 274 | ||
| 66815 | 275 | lemma odd_Suc_minus_one [simp]: "odd n \<Longrightarrow> Suc (n - Suc 0) = n" | 
| 276 | by (auto elim: oddE) | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 277 | |
| 66815 | 278 | lemma even_Suc_div_two [simp]: | 
| 279 | "even n \<Longrightarrow> Suc n div 2 = n div 2" | |
| 76387 | 280 | by auto | 
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 281 | |
| 66815 | 282 | lemma odd_Suc_div_two [simp]: | 
| 283 | "odd n \<Longrightarrow> Suc n div 2 = Suc (n div 2)" | |
| 76387 | 284 | by (auto elim: oddE) | 
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 285 | |
| 66815 | 286 | lemma odd_two_times_div_two_nat [simp]: | 
| 287 | assumes "odd n" | |
| 288 | shows "2 * (n div 2) = n - (1 :: nat)" | |
| 289 | proof - | |
| 290 | from assms have "2 * (n div 2) + 1 = n" | |
| 76387 | 291 | by (auto elim: oddE) | 
| 66815 | 292 | then have "Suc (2 * (n div 2)) - 1 = n - 1" | 
| 58787 | 293 | by simp | 
| 66815 | 294 | then show ?thesis | 
| 295 | by simp | |
| 58787 | 296 | qed | 
| 58680 | 297 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 298 | lemma not_mod2_eq_Suc_0_eq_0 [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 299 | "n mod 2 \<noteq> Suc 0 \<longleftrightarrow> n mod 2 = 0" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 300 | using not_mod_2_eq_1_eq_0 [of n] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 301 | |
| 69502 | 302 | lemma odd_card_imp_not_empty: | 
| 303 |   \<open>A \<noteq> {}\<close> if \<open>odd (card A)\<close>
 | |
| 304 | using that by auto | |
| 305 | ||
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 306 | lemma nat_induct2 [case_names 0 1 step]: | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 307 | assumes "P 0" "P 1" and step: "\<And>n::nat. P n \<Longrightarrow> P (n + 2)" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 308 | shows "P n" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 309 | proof (induct n rule: less_induct) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 310 | case (less n) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 311 | show ?case | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 312 | proof (cases "n < Suc (Suc 0)") | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 313 | case True | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 314 | then show ?thesis | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 315 | using assms by (auto simp: less_Suc_eq) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 316 | next | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 317 | case False | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 318 | then obtain k where k: "n = Suc (Suc k)" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 319 | by (force simp: not_less nat_le_iff_add) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 320 | then have "k<n" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 321 | by simp | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 322 | with less assms have "P (k+2)" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 323 | by blast | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 324 | then show ?thesis | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 325 | by (simp add: k) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 326 | qed | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 327 | qed | 
| 58687 | 328 | |
| 78668 | 329 | lemma mod_double_nat: | 
| 330 | \<open>n mod (2 * m) = n mod m \<or> n mod (2 * m) = n mod m + m\<close> | |
| 331 | for m n :: nat | |
| 332 | by (cases \<open>even (n div m)\<close>) | |
| 333 | (simp_all add: mod_mult2_eq ac_simps even_iff_mod_2_eq_zero) | |
| 334 | ||
| 71412 | 335 | context semiring_parity | 
| 336 | begin | |
| 337 | ||
| 338 | lemma even_sum_iff: | |
| 339 |   \<open>even (sum f A) \<longleftrightarrow> even (card {a\<in>A. odd (f a)})\<close> if \<open>finite A\<close>
 | |
| 340 | using that proof (induction A) | |
| 341 | case empty | |
| 342 | then show ?case | |
| 343 | by simp | |
| 344 | next | |
| 345 | case (insert a A) | |
| 346 |   moreover have \<open>{b \<in> insert a A. odd (f b)} = (if odd (f a) then {a} else {}) \<union> {b \<in> A. odd (f b)}\<close>
 | |
| 347 | by auto | |
| 348 | ultimately show ?case | |
| 349 | by simp | |
| 350 | qed | |
| 351 | ||
| 352 | lemma even_mask_iff [simp]: | |
| 353 | \<open>even (2 ^ n - 1) \<longleftrightarrow> n = 0\<close> | |
| 354 | proof (cases \<open>n = 0\<close>) | |
| 355 | case True | |
| 356 | then show ?thesis | |
| 357 | by simp | |
| 358 | next | |
| 359 | case False | |
| 360 |   then have \<open>{a. a = 0 \<and> a < n} = {0}\<close>
 | |
| 361 | by auto | |
| 362 | then show ?thesis | |
| 363 | by (auto simp add: mask_eq_sum_exp even_sum_iff) | |
| 364 | qed | |
| 365 | ||
| 76387 | 366 | lemma even_of_nat_iff [simp]: | 
| 367 | "even (of_nat n) \<longleftrightarrow> even n" | |
| 368 | by (induction n) simp_all | |
| 369 | ||
| 71412 | 370 | end | 
| 371 | ||
| 71138 | 372 | |
| 60758 | 373 | subsection \<open>Parity and powers\<close> | 
| 58689 | 374 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
60867diff
changeset | 375 | context ring_1 | 
| 58689 | 376 | begin | 
| 377 | ||
| 63654 | 378 | lemma power_minus_even [simp]: "even n \<Longrightarrow> (- a) ^ n = a ^ n" | 
| 58690 | 379 | by (auto elim: evenE) | 
| 58689 | 380 | |
| 63654 | 381 | lemma power_minus_odd [simp]: "odd n \<Longrightarrow> (- a) ^ n = - (a ^ n)" | 
| 58690 | 382 | by (auto elim: oddE) | 
| 383 | ||
| 66815 | 384 | lemma uminus_power_if: | 
| 385 | "(- a) ^ n = (if even n then a ^ n else - (a ^ n))" | |
| 386 | by auto | |
| 387 | ||
| 63654 | 388 | lemma neg_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1" | 
| 58690 | 389 | by simp | 
| 58689 | 390 | |
| 63654 | 391 | lemma neg_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1" | 
| 58690 | 392 | by simp | 
| 58689 | 393 | |
| 66582 | 394 | lemma neg_one_power_add_eq_neg_one_power_diff: "k \<le> n \<Longrightarrow> (- 1) ^ (n + k) = (- 1) ^ (n - k)" | 
| 395 | by (cases "even (n + k)") auto | |
| 396 | ||
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 397 | lemma minus_one_power_iff: "(- 1) ^ n = (if even n then 1 else - 1)" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 398 | by (induct n) auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 399 | |
| 63654 | 400 | end | 
| 58689 | 401 | |
| 402 | context linordered_idom | |
| 403 | begin | |
| 404 | ||
| 63654 | 405 | lemma zero_le_even_power: "even n \<Longrightarrow> 0 \<le> a ^ n" | 
| 58690 | 406 | by (auto elim: evenE) | 
| 58689 | 407 | |
| 63654 | 408 | lemma zero_le_odd_power: "odd n \<Longrightarrow> 0 \<le> a ^ n \<longleftrightarrow> 0 \<le> a" | 
| 58689 | 409 | by (auto simp add: power_even_eq zero_le_mult_iff elim: oddE) | 
| 410 | ||
| 63654 | 411 | lemma zero_le_power_eq: "0 \<le> a ^ n \<longleftrightarrow> even n \<or> odd n \<and> 0 \<le> a" | 
| 58787 | 412 | by (auto simp add: zero_le_even_power zero_le_odd_power) | 
| 63654 | 413 | |
| 414 | lemma zero_less_power_eq: "0 < a ^ n \<longleftrightarrow> n = 0 \<or> even n \<and> a \<noteq> 0 \<or> odd n \<and> 0 < a" | |
| 58689 | 415 | proof - | 
| 416 | have [simp]: "0 = a ^ n \<longleftrightarrow> a = 0 \<and> n > 0" | |
| 58787 | 417 | unfolding power_eq_0_iff [of a n, symmetric] by blast | 
| 58689 | 418 | show ?thesis | 
| 63654 | 419 | unfolding less_le zero_le_power_eq by auto | 
| 58689 | 420 | qed | 
| 421 | ||
| 63654 | 422 | lemma power_less_zero_eq [simp]: "a ^ n < 0 \<longleftrightarrow> odd n \<and> a < 0" | 
| 58689 | 423 | unfolding not_le [symmetric] zero_le_power_eq by auto | 
| 424 | ||
| 63654 | 425 | lemma power_le_zero_eq: "a ^ n \<le> 0 \<longleftrightarrow> n > 0 \<and> (odd n \<and> a \<le> 0 \<or> even n \<and> a = 0)" | 
| 426 | unfolding not_less [symmetric] zero_less_power_eq by auto | |
| 427 | ||
| 428 | lemma power_even_abs: "even n \<Longrightarrow> \<bar>a\<bar> ^ n = a ^ n" | |
| 58689 | 429 | using power_abs [of a n] by (simp add: zero_le_even_power) | 
| 430 | ||
| 431 | lemma power_mono_even: | |
| 432 | assumes "even n" and "\<bar>a\<bar> \<le> \<bar>b\<bar>" | |
| 433 | shows "a ^ n \<le> b ^ n" | |
| 434 | proof - | |
| 435 | have "0 \<le> \<bar>a\<bar>" by auto | |
| 63654 | 436 | with \<open>\<bar>a\<bar> \<le> \<bar>b\<bar>\<close> have "\<bar>a\<bar> ^ n \<le> \<bar>b\<bar> ^ n" | 
| 437 | by (rule power_mono) | |
| 438 | with \<open>even n\<close> show ?thesis | |
| 439 | by (simp add: power_even_abs) | |
| 58689 | 440 | qed | 
| 441 | ||
| 442 | lemma power_mono_odd: | |
| 443 | assumes "odd n" and "a \<le> b" | |
| 444 | shows "a ^ n \<le> b ^ n" | |
| 445 | proof (cases "b < 0") | |
| 63654 | 446 | case True | 
| 447 | with \<open>a \<le> b\<close> have "- b \<le> - a" and "0 \<le> - b" by auto | |
| 448 | then have "(- b) ^ n \<le> (- a) ^ n" by (rule power_mono) | |
| 60758 | 449 | with \<open>odd n\<close> show ?thesis by simp | 
| 58689 | 450 | next | 
| 63654 | 451 | case False | 
| 452 | then have "0 \<le> b" by auto | |
| 58689 | 453 | show ?thesis | 
| 454 | proof (cases "a < 0") | |
| 63654 | 455 | case True | 
| 456 | then have "n \<noteq> 0" and "a \<le> 0" using \<open>odd n\<close> [THEN odd_pos] by auto | |
| 60758 | 457 | then have "a ^ n \<le> 0" unfolding power_le_zero_eq using \<open>odd n\<close> by auto | 
| 63654 | 458 | moreover from \<open>0 \<le> b\<close> have "0 \<le> b ^ n" by auto | 
| 58689 | 459 | ultimately show ?thesis by auto | 
| 460 | next | |
| 63654 | 461 | case False | 
| 462 | then have "0 \<le> a" by auto | |
| 463 | with \<open>a \<le> b\<close> show ?thesis | |
| 464 | using power_mono by auto | |
| 58689 | 465 | qed | 
| 466 | qed | |
| 62083 | 467 | |
| 60758 | 468 | text \<open>Simplify, when the exponent is a numeral\<close> | 
| 58689 | 469 | |
| 470 | lemma zero_le_power_eq_numeral [simp]: | |
| 471 | "0 \<le> a ^ numeral w \<longleftrightarrow> even (numeral w :: nat) \<or> odd (numeral w :: nat) \<and> 0 \<le> a" | |
| 472 | by (fact zero_le_power_eq) | |
| 473 | ||
| 474 | lemma zero_less_power_eq_numeral [simp]: | |
| 63654 | 475 | "0 < a ^ numeral w \<longleftrightarrow> | 
| 476 | numeral w = (0 :: nat) \<or> | |
| 477 | even (numeral w :: nat) \<and> a \<noteq> 0 \<or> | |
| 478 | odd (numeral w :: nat) \<and> 0 < a" | |
| 58689 | 479 | by (fact zero_less_power_eq) | 
| 480 | ||
| 481 | lemma power_le_zero_eq_numeral [simp]: | |
| 63654 | 482 | "a ^ numeral w \<le> 0 \<longleftrightarrow> | 
| 483 | (0 :: nat) < numeral w \<and> | |
| 484 | (odd (numeral w :: nat) \<and> a \<le> 0 \<or> even (numeral w :: nat) \<and> a = 0)" | |
| 58689 | 485 | by (fact power_le_zero_eq) | 
| 486 | ||
| 487 | lemma power_less_zero_eq_numeral [simp]: | |
| 488 | "a ^ numeral w < 0 \<longleftrightarrow> odd (numeral w :: nat) \<and> a < 0" | |
| 489 | by (fact power_less_zero_eq) | |
| 490 | ||
| 491 | lemma power_even_abs_numeral [simp]: | |
| 492 | "even (numeral w :: nat) \<Longrightarrow> \<bar>a\<bar> ^ numeral w = a ^ numeral w" | |
| 493 | by (fact power_even_abs) | |
| 494 | ||
| 495 | end | |
| 496 | ||
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 497 | |
| 69593 | 498 | subsection \<open>Instance for \<^typ>\<open>int\<close>\<close> | 
| 76387 | 499 | |
| 67816 | 500 | lemma even_diff_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 501 | "even (k - l) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 502 | by (fact even_diff) | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 503 | |
| 67816 | 504 | lemma even_abs_add_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 505 | "even (\<bar>k\<bar> + l) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 506 | by simp | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 507 | |
| 67816 | 508 | lemma even_add_abs_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 509 | "even (k + \<bar>l\<bar>) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 510 | by simp | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 511 | |
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 512 | lemma even_nat_iff: "0 \<le> k \<Longrightarrow> even (nat k) \<longleftrightarrow> even k" | 
| 74592 | 513 | by (simp add: even_of_nat_iff [of "nat k", where ?'a = int, symmetric]) | 
| 71138 | 514 | |
| 71837 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 515 | context | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 516 |   assumes "SORT_CONSTRAINT('a::division_ring)"
 | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 517 | begin | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 518 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 519 | lemma power_int_minus_left: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 520 | "power_int (-a :: 'a) n = (if even n then power_int a n else -power_int a n)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 521 | by (auto simp: power_int_def minus_one_power_iff even_nat_iff) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 522 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 523 | lemma power_int_minus_left_even [simp]: "even n \<Longrightarrow> power_int (-a :: 'a) n = power_int a n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 524 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 525 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 526 | lemma power_int_minus_left_odd [simp]: "odd n \<Longrightarrow> power_int (-a :: 'a) n = -power_int a n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 527 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 528 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 529 | lemma power_int_minus_left_distrib: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 530 | "NO_MATCH (-1) x \<Longrightarrow> power_int (-a :: 'a) n = power_int (-1) n * power_int a n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 531 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 532 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 533 | lemma power_int_minus_one_minus: "power_int (-1 :: 'a) (-n) = power_int (-1) n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 534 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 535 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 536 | lemma power_int_minus_one_diff_commute: "power_int (-1 :: 'a) (a - b) = power_int (-1) (b - a)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 537 | by (subst power_int_minus_one_minus [symmetric]) auto | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 538 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 539 | lemma power_int_minus_one_mult_self [simp]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 540 | "power_int (-1 :: 'a) m * power_int (-1) m = 1" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 541 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 542 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 543 | lemma power_int_minus_one_mult_self' [simp]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 544 | "power_int (-1 :: 'a) m * (power_int (-1) m * b) = b" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 545 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 546 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 547 | end | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 548 | |
| 75937 | 549 | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 550 | subsection \<open>Special case: euclidean rings structurally containing the natural numbers\<close> | 
| 76387 | 551 | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 552 | class linordered_euclidean_semiring = discrete_linordered_semidom + unique_euclidean_semiring + | 
| 76387 | 553 | assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n" | 
| 554 | and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1" | |
| 555 | and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a" | |
| 556 | begin | |
| 557 | ||
| 558 | lemma division_segment_eq_iff: | |
| 559 | "a = b" if "division_segment a = division_segment b" | |
| 560 | and "euclidean_size a = euclidean_size b" | |
| 561 | using that division_segment_euclidean_size [of a] by simp | |
| 562 | ||
| 563 | lemma euclidean_size_of_nat [simp]: | |
| 564 | "euclidean_size (of_nat n) = n" | |
| 565 | proof - | |
| 566 | have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n" | |
| 567 | by (fact division_segment_euclidean_size) | |
| 568 | then show ?thesis by simp | |
| 569 | qed | |
| 570 | ||
| 571 | lemma of_nat_euclidean_size: | |
| 572 | "of_nat (euclidean_size a) = a div division_segment a" | |
| 573 | proof - | |
| 574 | have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a" | |
| 575 | by (subst nonzero_mult_div_cancel_left) simp_all | |
| 576 | also have "\<dots> = a div division_segment a" | |
| 577 | by simp | |
| 578 | finally show ?thesis . | |
| 579 | qed | |
| 580 | ||
| 581 | lemma division_segment_1 [simp]: | |
| 582 | "division_segment 1 = 1" | |
| 583 | using division_segment_of_nat [of 1] by simp | |
| 584 | ||
| 585 | lemma division_segment_numeral [simp]: | |
| 586 | "division_segment (numeral k) = 1" | |
| 587 | using division_segment_of_nat [of "numeral k"] by simp | |
| 588 | ||
| 589 | lemma euclidean_size_1 [simp]: | |
| 590 | "euclidean_size 1 = 1" | |
| 591 | using euclidean_size_of_nat [of 1] by simp | |
| 592 | ||
| 593 | lemma euclidean_size_numeral [simp]: | |
| 594 | "euclidean_size (numeral k) = numeral k" | |
| 595 | using euclidean_size_of_nat [of "numeral k"] by simp | |
| 596 | ||
| 597 | lemma of_nat_dvd_iff: | |
| 598 | "of_nat m dvd of_nat n \<longleftrightarrow> m dvd n" (is "?P \<longleftrightarrow> ?Q") | |
| 599 | proof (cases "m = 0") | |
| 600 | case True | |
| 601 | then show ?thesis | |
| 602 | by simp | |
| 603 | next | |
| 604 | case False | |
| 605 | show ?thesis | |
| 606 | proof | |
| 607 | assume ?Q | |
| 608 | then show ?P | |
| 609 | by auto | |
| 610 | next | |
| 611 | assume ?P | |
| 612 | with False have "of_nat n = of_nat n div of_nat m * of_nat m" | |
| 613 | by simp | |
| 614 | then have "of_nat n = of_nat (n div m * m)" | |
| 615 | by (simp add: of_nat_div) | |
| 616 | then have "n = n div m * m" | |
| 617 | by (simp only: of_nat_eq_iff) | |
| 618 | then have "n = m * (n div m)" | |
| 619 | by (simp add: ac_simps) | |
| 620 | then show ?Q .. | |
| 621 | qed | |
| 622 | qed | |
| 623 | ||
| 624 | lemma of_nat_mod: | |
| 625 | "of_nat (m mod n) = of_nat m mod of_nat n" | |
| 626 | proof - | |
| 627 | have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m" | |
| 628 | by (simp add: div_mult_mod_eq) | |
| 629 | also have "of_nat m = of_nat (m div n * n + m mod n)" | |
| 630 | by simp | |
| 631 | finally show ?thesis | |
| 632 | by (simp only: of_nat_div of_nat_mult of_nat_add) simp | |
| 633 | qed | |
| 634 | ||
| 635 | lemma one_div_two_eq_zero [simp]: | |
| 636 | "1 div 2 = 0" | |
| 637 | proof - | |
| 638 | from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0" | |
| 639 | by (simp only:) simp | |
| 640 | then show ?thesis | |
| 641 | by simp | |
| 642 | qed | |
| 643 | ||
| 644 | lemma one_mod_2_pow_eq [simp]: | |
| 645 | "1 mod (2 ^ n) = of_bool (n > 0)" | |
| 646 | proof - | |
| 647 | have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))" | |
| 648 | using of_nat_mod [of 1 "2 ^ n"] by simp | |
| 649 | also have "\<dots> = of_bool (n > 0)" | |
| 650 | by simp | |
| 651 | finally show ?thesis . | |
| 652 | qed | |
| 653 | ||
| 654 | lemma one_div_2_pow_eq [simp]: | |
| 655 | "1 div (2 ^ n) = of_bool (n = 0)" | |
| 656 | using div_mult_mod_eq [of 1 "2 ^ n"] by auto | |
| 657 | ||
| 658 | lemma div_mult2_eq': | |
| 659 | \<open>a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\<close> | |
| 660 | proof (cases \<open>m = 0 \<or> n = 0\<close>) | |
| 661 | case True | |
| 662 | then show ?thesis | |
| 663 | by auto | |
| 664 | next | |
| 665 | case False | |
| 666 | then have \<open>m > 0\<close> \<open>n > 0\<close> | |
| 667 | by simp_all | |
| 668 | show ?thesis | |
| 669 | proof (cases \<open>of_nat m * of_nat n dvd a\<close>) | |
| 670 | case True | |
| 671 | then obtain b where \<open>a = (of_nat m * of_nat n) * b\<close> .. | |
| 672 | then have \<open>a = of_nat m * (of_nat n * b)\<close> | |
| 673 | by (simp add: ac_simps) | |
| 674 | then show ?thesis | |
| 675 | by simp | |
| 676 | next | |
| 677 | case False | |
| 678 | define q where \<open>q = a div (of_nat m * of_nat n)\<close> | |
| 679 | define r where \<open>r = a mod (of_nat m * of_nat n)\<close> | |
| 680 | from \<open>m > 0\<close> \<open>n > 0\<close> \<open>\<not> of_nat m * of_nat n dvd a\<close> r_def have "division_segment r = 1" | |
| 681 | using division_segment_of_nat [of "m * n"] by (simp add: division_segment_mod) | |
| 682 | with division_segment_euclidean_size [of r] | |
| 683 | have "of_nat (euclidean_size r) = r" | |
| 684 | by simp | |
| 685 | have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0" | |
| 686 | by simp | |
| 687 | with \<open>m > 0\<close> \<open>n > 0\<close> r_def have "r div (of_nat m * of_nat n) = 0" | |
| 688 | by simp | |
| 689 | with \<open>of_nat (euclidean_size r) = r\<close> | |
| 690 | have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0" | |
| 691 | by simp | |
| 692 | then have "of_nat (euclidean_size r div (m * n)) = 0" | |
| 693 | by (simp add: of_nat_div) | |
| 694 | then have "of_nat (euclidean_size r div m div n) = 0" | |
| 695 | by (simp add: div_mult2_eq) | |
| 696 | with \<open>of_nat (euclidean_size r) = r\<close> have "r div of_nat m div of_nat n = 0" | |
| 697 | by (simp add: of_nat_div) | |
| 698 | with \<open>m > 0\<close> \<open>n > 0\<close> q_def | |
| 699 | have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n" | |
| 700 | by simp | |
| 701 | moreover have \<open>a = q * (of_nat m * of_nat n) + r\<close> | |
| 702 | by (simp add: q_def r_def div_mult_mod_eq) | |
| 703 | ultimately show \<open>a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\<close> | |
| 704 | using q_def [symmetric] div_plus_div_distrib_dvd_right [of \<open>of_nat m\<close> \<open>q * (of_nat m * of_nat n)\<close> r] | |
| 705 | by (simp add: ac_simps) | |
| 706 | qed | |
| 707 | qed | |
| 708 | ||
| 709 | lemma mod_mult2_eq': | |
| 710 | "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m" | |
| 711 | proof - | |
| 712 | have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)" | |
| 713 | by (simp add: combine_common_factor div_mult_mod_eq) | |
| 714 | moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)" | |
| 715 | by (simp add: ac_simps) | |
| 716 | ultimately show ?thesis | |
| 717 | by (simp add: div_mult2_eq' mult_commute) | |
| 718 | qed | |
| 719 | ||
| 720 | lemma div_mult2_numeral_eq: | |
| 721 | "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B") | |
| 722 | proof - | |
| 723 | have "?A = a div of_nat (numeral k) div of_nat (numeral l)" | |
| 724 | by simp | |
| 725 | also have "\<dots> = a div (of_nat (numeral k) * of_nat (numeral l))" | |
| 726 | by (fact div_mult2_eq' [symmetric]) | |
| 727 | also have "\<dots> = ?B" | |
| 728 | by simp | |
| 729 | finally show ?thesis . | |
| 730 | qed | |
| 731 | ||
| 732 | lemma numeral_Bit0_div_2: | |
| 733 | "numeral (num.Bit0 n) div 2 = numeral n" | |
| 734 | proof - | |
| 735 | have "numeral (num.Bit0 n) = numeral n + numeral n" | |
| 736 | by (simp only: numeral.simps) | |
| 737 | also have "\<dots> = numeral n * 2" | |
| 738 | by (simp add: mult_2_right) | |
| 739 | finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2" | |
| 740 | by simp | |
| 741 | also have "\<dots> = numeral n" | |
| 742 | by (rule nonzero_mult_div_cancel_right) simp | |
| 743 | finally show ?thesis . | |
| 744 | qed | |
| 745 | ||
| 746 | lemma numeral_Bit1_div_2: | |
| 747 | "numeral (num.Bit1 n) div 2 = numeral n" | |
| 748 | proof - | |
| 749 | have "numeral (num.Bit1 n) = numeral n + numeral n + 1" | |
| 750 | by (simp only: numeral.simps) | |
| 751 | also have "\<dots> = numeral n * 2 + 1" | |
| 752 | by (simp add: mult_2_right) | |
| 753 | finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2" | |
| 754 | by simp | |
| 755 | also have "\<dots> = numeral n * 2 div 2 + 1 div 2" | |
| 756 | using dvd_triv_right by (rule div_plus_div_distrib_dvd_left) | |
| 757 | also have "\<dots> = numeral n * 2 div 2" | |
| 758 | by simp | |
| 759 | also have "\<dots> = numeral n" | |
| 760 | by (rule nonzero_mult_div_cancel_right) simp | |
| 761 | finally show ?thesis . | |
| 762 | qed | |
| 763 | ||
| 764 | lemma exp_mod_exp: | |
| 765 | \<open>2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\<close> | |
| 766 | proof - | |
| 767 | have \<open>(2::nat) ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\<close> (is \<open>?lhs = ?rhs\<close>) | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 768 | by (auto simp add: linorder_class.not_less monoid_mult_class.power_add dest!: le_Suc_ex) | 
| 76387 | 769 | then have \<open>of_nat ?lhs = of_nat ?rhs\<close> | 
| 770 | by simp | |
| 771 | then show ?thesis | |
| 772 | by (simp add: of_nat_mod) | |
| 773 | qed | |
| 774 | ||
| 775 | lemma mask_mod_exp: | |
| 776 | \<open>(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\<close> | |
| 777 | proof - | |
| 778 | have \<open>(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - (1::nat)\<close> (is \<open>?lhs = ?rhs\<close>) | |
| 779 | proof (cases \<open>n \<le> m\<close>) | |
| 780 | case True | |
| 781 | then show ?thesis | |
| 782 | by (simp add: Suc_le_lessD) | |
| 783 | next | |
| 784 | case False | |
| 785 | then have \<open>m < n\<close> | |
| 786 | by simp | |
| 787 | then obtain q where n: \<open>n = Suc q + m\<close> | |
| 788 | by (auto dest: less_imp_Suc_add) | |
| 789 | then have \<open>min m n = m\<close> | |
| 790 | by simp | |
| 791 | moreover have \<open>(2::nat) ^ m \<le> 2 * 2 ^ q * 2 ^ m\<close> | |
| 792 | using mult_le_mono1 [of 1 \<open>2 * 2 ^ q\<close> \<open>2 ^ m\<close>] by simp | |
| 793 | with n have \<open>2 ^ n - 1 = (2 ^ Suc q - 1) * 2 ^ m + (2 ^ m - (1::nat))\<close> | |
| 794 | by (simp add: monoid_mult_class.power_add algebra_simps) | |
| 795 | ultimately show ?thesis | |
| 796 | by (simp only: euclidean_semiring_cancel_class.mod_mult_self3) simp | |
| 797 | qed | |
| 798 | then have \<open>of_nat ?lhs = of_nat ?rhs\<close> | |
| 799 | by simp | |
| 800 | then show ?thesis | |
| 801 | by (simp add: of_nat_mod of_nat_diff) | |
| 802 | qed | |
| 803 | ||
| 804 | lemma of_bool_half_eq_0 [simp]: | |
| 805 | \<open>of_bool b div 2 = 0\<close> | |
| 806 | by simp | |
| 807 | ||
| 78668 | 808 | lemma of_nat_mod_double: | 
| 809 | \<open>of_nat n mod (2 * of_nat m) = of_nat n mod of_nat m \<or> of_nat n mod (2 * of_nat m) = of_nat n mod of_nat m + of_nat m\<close> | |
| 810 | by (simp add: mod_double_nat flip: of_nat_mod of_nat_add of_nat_mult of_nat_numeral) | |
| 811 | ||
| 76387 | 812 | end | 
| 813 | ||
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 814 | instance nat :: linordered_euclidean_semiring | 
| 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 815 | by standard simp_all | 
| 76387 | 816 | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 817 | instance int :: linordered_euclidean_semiring | 
| 76387 | 818 | by standard (auto simp add: divide_int_def division_segment_int_def elim: contrapos_np) | 
| 819 | ||
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 820 | context linordered_euclidean_semiring | 
| 76387 | 821 | begin | 
| 822 | ||
| 823 | subclass semiring_parity | |
| 824 | proof | |
| 79118 | 825 | show \<open>a mod 2 = of_bool (\<not> 2 dvd a)\<close> for a | 
| 826 | proof (cases \<open>2 dvd a\<close>) | |
| 827 | case True | |
| 828 | then show ?thesis | |
| 829 | by (simp add: dvd_eq_mod_eq_0) | |
| 76387 | 830 | next | 
| 79118 | 831 | case False | 
| 76387 | 832 | have eucl: "euclidean_size (a mod 2) = 1" | 
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 833 | proof (rule Orderings.order_antisym) | 
| 76387 | 834 | show "euclidean_size (a mod 2) \<le> 1" | 
| 835 | using mod_size_less [of 2 a] by simp | |
| 836 | show "1 \<le> euclidean_size (a mod 2)" | |
| 837 | using \<open>\<not> 2 dvd a\<close> by (simp add: Suc_le_eq dvd_eq_mod_eq_0) | |
| 838 | qed | |
| 839 | from \<open>\<not> 2 dvd a\<close> have "\<not> of_nat 2 dvd division_segment a * of_nat (euclidean_size a)" | |
| 840 | by simp | |
| 841 | then have "\<not> of_nat 2 dvd of_nat (euclidean_size a)" | |
| 842 | by (auto simp only: dvd_mult_unit_iff' is_unit_division_segment) | |
| 843 | then have "\<not> 2 dvd euclidean_size a" | |
| 844 | using of_nat_dvd_iff [of 2] by simp | |
| 845 | then have "euclidean_size a mod 2 = 1" | |
| 846 | by (simp add: semidom_modulo_class.dvd_eq_mod_eq_0) | |
| 847 | then have "of_nat (euclidean_size a mod 2) = of_nat 1" | |
| 848 | by simp | |
| 849 | then have "of_nat (euclidean_size a) mod 2 = 1" | |
| 850 | by (simp add: of_nat_mod) | |
| 851 | from \<open>\<not> 2 dvd a\<close> eucl | |
| 79118 | 852 | have "a mod 2 = 1" | 
| 76387 | 853 | by (auto intro: division_segment_eq_iff simp add: division_segment_mod) | 
| 79118 | 854 | with \<open>\<not> 2 dvd a\<close> show ?thesis | 
| 855 | by simp | |
| 76387 | 856 | qed | 
| 79118 | 857 | show \<open>\<not> is_unit 2\<close> | 
| 858 | proof | |
| 859 | assume \<open>is_unit 2\<close> | |
| 860 | then have \<open>of_nat 2 dvd of_nat 1\<close> | |
| 76387 | 861 | by simp | 
| 79118 | 862 | then have \<open>is_unit (2::nat)\<close> | 
| 76387 | 863 | by (simp only: of_nat_dvd_iff) | 
| 864 | then show False | |
| 865 | by simp | |
| 866 | qed | |
| 79555 | 867 | show \<open>(1 + a) div 2 = a div 2\<close> if \<open>2 dvd a\<close> for a | 
| 868 | using that by auto | |
| 76387 | 869 | qed | 
| 870 | ||
| 871 | lemma even_succ_div_two [simp]: | |
| 872 | "even a \<Longrightarrow> (a + 1) div 2 = a div 2" | |
| 873 | by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero) | |
| 874 | ||
| 875 | lemma odd_succ_div_two [simp]: | |
| 876 | "odd a \<Longrightarrow> (a + 1) div 2 = a div 2 + 1" | |
| 877 | by (auto elim!: oddE simp add: add.assoc) | |
| 878 | ||
| 879 | lemma even_two_times_div_two: | |
| 880 | "even a \<Longrightarrow> 2 * (a div 2) = a" | |
| 881 | by (fact dvd_mult_div_cancel) | |
| 882 | ||
| 883 | lemma odd_two_times_div_two_succ [simp]: | |
| 884 | "odd a \<Longrightarrow> 2 * (a div 2) + 1 = a" | |
| 885 | using mult_div_mod_eq [of 2 a] | |
| 886 | by (simp add: even_iff_mod_2_eq_zero) | |
| 887 | ||
| 888 | lemma coprime_left_2_iff_odd [simp]: | |
| 889 | "coprime 2 a \<longleftrightarrow> odd a" | |
| 890 | proof | |
| 891 | assume "odd a" | |
| 892 | show "coprime 2 a" | |
| 893 | proof (rule coprimeI) | |
| 894 | fix b | |
| 895 | assume "b dvd 2" "b dvd a" | |
| 896 | then have "b dvd a mod 2" | |
| 897 | by (auto intro: dvd_mod) | |
| 898 | with \<open>odd a\<close> show "is_unit b" | |
| 899 | by (simp add: mod_2_eq_odd) | |
| 900 | qed | |
| 901 | next | |
| 902 | assume "coprime 2 a" | |
| 903 | show "odd a" | |
| 904 | proof (rule notI) | |
| 905 | assume "even a" | |
| 906 | then obtain b where "a = 2 * b" .. | |
| 907 | with \<open>coprime 2 a\<close> have "coprime 2 (2 * b)" | |
| 908 | by simp | |
| 909 | moreover have "\<not> coprime 2 (2 * b)" | |
| 910 | by (rule not_coprimeI [of 2]) simp_all | |
| 911 | ultimately show False | |
| 912 | by blast | |
| 913 | qed | |
| 914 | qed | |
| 915 | ||
| 916 | lemma coprime_right_2_iff_odd [simp]: | |
| 917 | "coprime a 2 \<longleftrightarrow> odd a" | |
| 918 | using coprime_left_2_iff_odd [of a] by (simp add: ac_simps) | |
| 919 | ||
| 920 | end | |
| 921 | ||
| 922 | lemma minus_1_mod_2_eq [simp]: | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 923 | \<open>- 1 mod 2 = (1::int)\<close> | 
| 76387 | 924 | by (simp add: mod_2_eq_odd) | 
| 925 | ||
| 926 | lemma minus_1_div_2_eq [simp]: | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 927 | "- 1 div 2 = - (1::int)" | 
| 76387 | 928 | proof - | 
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 929 | from div_mult_mod_eq [of "- 1 :: int" 2] | 
| 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 930 | have "- 1 div 2 * 2 = - 1 * (2 :: int)" | 
| 76387 | 931 | using add_implies_diff by fastforce | 
| 932 | then show ?thesis | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 933 | using mult_right_cancel [of 2 "- 1 div 2" "- (1 :: int)"] by simp | 
| 76387 | 934 | qed | 
| 935 | ||
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 936 | context linordered_euclidean_semiring | 
| 76387 | 937 | begin | 
| 938 | ||
| 79531 
22a137a6de44
rearranged and reformulated abstract classes for bit structures and operations
 haftmann parents: 
79118diff
changeset | 939 | lemma even_decr_exp_div_exp_iff': | 
| 76387 | 940 | \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> m \<le> n\<close> | 
| 941 | proof - | |
| 942 | have \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> even (of_nat ((2 ^ m - Suc 0) div 2 ^ n))\<close> | |
| 943 | by (simp only: of_nat_div) (simp add: of_nat_diff) | |
| 944 | also have \<open>\<dots> \<longleftrightarrow> even ((2 ^ m - Suc 0) div 2 ^ n)\<close> | |
| 945 | by simp | |
| 946 | also have \<open>\<dots> \<longleftrightarrow> m \<le> n\<close> | |
| 947 | proof (cases \<open>m \<le> n\<close>) | |
| 948 | case True | |
| 949 | then show ?thesis | |
| 950 | by (simp add: Suc_le_lessD) | |
| 951 | next | |
| 952 | case False | |
| 953 | then obtain r where r: \<open>m = n + Suc r\<close> | |
| 954 | using less_imp_Suc_add by fastforce | |
| 955 |     from r have \<open>{q. q < m} \<inter> {q. 2 ^ n dvd (2::nat) ^ q} = {q. n \<le> q \<and> q < m}\<close>
 | |
| 956 | by (auto simp add: dvd_power_iff_le) | |
| 957 |     moreover from r have \<open>{q. q < m} \<inter> {q. \<not> 2 ^ n dvd (2::nat) ^ q} = {q. q < n}\<close>
 | |
| 958 | by (auto simp add: dvd_power_iff_le) | |
| 959 |     moreover from False have \<open>{q. n \<le> q \<and> q < m \<and> q \<le> n} = {n}\<close>
 | |
| 960 | by auto | |
| 961 |     then have \<open>odd ((\<Sum>a\<in>{q. n \<le> q \<and> q < m}. 2 ^ a div (2::nat) ^ n) + sum ((^) 2) {q. q < n} div 2 ^ n)\<close>
 | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 962 | by (simp_all add: euclidean_semiring_cancel_class.power_diff_power_eq semiring_parity_class.even_sum_iff | 
| 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 963 | linorder_class.not_less mask_eq_sum_exp_nat [symmetric]) | 
| 76387 | 964 |     ultimately have \<open>odd (sum ((^) (2::nat)) {q. q < m} div 2 ^ n)\<close>
 | 
| 965 | by (subst euclidean_semiring_cancel_class.sum_div_partition) simp_all | |
| 966 | with False show ?thesis | |
| 967 | by (simp add: mask_eq_sum_exp_nat) | |
| 968 | qed | |
| 969 | finally show ?thesis . | |
| 970 | qed | |
| 971 | ||
| 972 | end | |
| 973 | ||
| 974 | ||
| 975 | subsection \<open>Generic symbolic computations\<close> | |
| 976 | ||
| 977 | text \<open> | |
| 978 | The following type class contains everything necessary to formulate | |
| 979 | a division algorithm in ring structures with numerals, restricted | |
| 980 | to its positive segments. | |
| 981 | \<close> | |
| 982 | ||
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 983 | class linordered_euclidean_semiring_division = linordered_euclidean_semiring + | 
| 76387 | 984 | fixes divmod :: \<open>num \<Rightarrow> num \<Rightarrow> 'a \<times> 'a\<close> | 
| 985 | and divmod_step :: \<open>'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<times> 'a\<close> \<comment> \<open> | |
| 986 | These are conceptually definitions but force generated code | |
| 987 | to be monomorphic wrt. particular instances of this class which | |
| 988 | yields a significant speedup.\<close> | |
| 989 | assumes divmod_def: \<open>divmod m n = (numeral m div numeral n, numeral m mod numeral n)\<close> | |
| 990 | and divmod_step_def [simp]: \<open>divmod_step l (q, r) = | |
| 991 | (if euclidean_size l \<le> euclidean_size r then (2 * q + 1, r - l) | |
| 992 | else (2 * q, r))\<close> \<comment> \<open> | |
| 993 | This is a formulation of one step (referring to one digit position) | |
| 994 | in school-method division: compare the dividend at the current | |
| 995 | digit position with the remainder from previous division steps | |
| 996 | and evaluate accordingly.\<close> | |
| 997 | begin | |
| 998 | ||
| 999 | lemma fst_divmod: | |
| 1000 | \<open>fst (divmod m n) = numeral m div numeral n\<close> | |
| 1001 | by (simp add: divmod_def) | |
| 1002 | ||
| 1003 | lemma snd_divmod: | |
| 1004 | \<open>snd (divmod m n) = numeral m mod numeral n\<close> | |
| 1005 | by (simp add: divmod_def) | |
| 1006 | ||
| 1007 | text \<open> | |
| 1008 | Following a formulation of school-method division. | |
| 1009 | If the divisor is smaller than the dividend, terminate. | |
| 1010 | If not, shift the dividend to the right until termination | |
| 1011 | occurs and then reiterate single division steps in the | |
| 1012 | opposite direction. | |
| 1013 | \<close> | |
| 1014 | ||
| 1015 | lemma divmod_divmod_step: | |
| 1016 | \<open>divmod m n = (if m < n then (0, numeral m) | |
| 1017 | else divmod_step (numeral n) (divmod m (Num.Bit0 n)))\<close> | |
| 1018 | proof (cases \<open>m < n\<close>) | |
| 1019 | case True | |
| 1020 | then show ?thesis | |
| 1021 | by (simp add: prod_eq_iff fst_divmod snd_divmod flip: of_nat_numeral of_nat_div of_nat_mod) | |
| 1022 | next | |
| 1023 | case False | |
| 1024 | define r s t where \<open>r = (numeral m :: nat)\<close> \<open>s = (numeral n :: nat)\<close> \<open>t = 2 * s\<close> | |
| 1025 | then have *: \<open>numeral m = of_nat r\<close> \<open>numeral n = of_nat s\<close> \<open>numeral (num.Bit0 n) = of_nat t\<close> | |
| 1026 | and \<open>\<not> s \<le> r mod s\<close> | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1027 | by (simp_all add: linorder_class.not_le) | 
| 76387 | 1028 | have t: \<open>2 * (r div t) = r div s - r div s mod 2\<close> | 
| 1029 | \<open>r mod t = s * (r div s mod 2) + r mod s\<close> | |
| 77061 
5de3772609ea
generalized theory name: euclidean division denotes one particular division definition on integers
 haftmann parents: 
76387diff
changeset | 1030 | by (simp add: Rings.minus_mod_eq_mult_div Groups.mult.commute [of 2] Euclidean_Rings.div_mult2_eq \<open>t = 2 * s\<close>) | 
| 76387 | 1031 | (use mod_mult2_eq [of r s 2] in \<open>simp add: ac_simps \<open>t = 2 * s\<close>\<close>) | 
| 1032 | have rs: \<open>r div s mod 2 = 0 \<or> r div s mod 2 = Suc 0\<close> | |
| 1033 | by auto | |
| 1034 | from \<open>\<not> s \<le> r mod s\<close> have \<open>s \<le> r mod t \<Longrightarrow> | |
| 1035 | r div s = Suc (2 * (r div t)) \<and> | |
| 1036 | r mod s = r mod t - s\<close> | |
| 1037 | using rs | |
| 1038 | by (auto simp add: t) | |
| 1039 | moreover have \<open>r mod t < s \<Longrightarrow> | |
| 1040 | r div s = 2 * (r div t) \<and> | |
| 1041 | r mod s = r mod t\<close> | |
| 1042 | using rs | |
| 1043 | by (auto simp add: t) | |
| 1044 | ultimately show ?thesis | |
| 1045 | by (simp add: divmod_def prod_eq_iff split_def Let_def | |
| 1046 | not_less mod_eq_0_iff_dvd Rings.mod_eq_0_iff_dvd False not_le *) | |
| 1047 | (simp add: flip: of_nat_numeral of_nat_mult add.commute [of 1] of_nat_div of_nat_mod of_nat_Suc of_nat_diff) | |
| 1048 | qed | |
| 1049 | ||
| 1050 | text \<open>The division rewrite proper -- first, trivial results involving \<open>1\<close>\<close> | |
| 1051 | ||
| 1052 | lemma divmod_trivial [simp]: | |
| 1053 | "divmod m Num.One = (numeral m, 0)" | |
| 1054 | "divmod num.One (num.Bit0 n) = (0, Numeral1)" | |
| 1055 | "divmod num.One (num.Bit1 n) = (0, Numeral1)" | |
| 1056 | using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def) | |
| 1057 | ||
| 1058 | text \<open>Division by an even number is a right-shift\<close> | |
| 1059 | ||
| 1060 | lemma divmod_cancel [simp]: | |
| 1061 | \<open>divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r))\<close> (is ?P) | |
| 1062 | \<open>divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r + 1))\<close> (is ?Q) | |
| 1063 | proof - | |
| 1064 | define r s where \<open>r = (numeral m :: nat)\<close> \<open>s = (numeral n :: nat)\<close> | |
| 1065 | then have *: \<open>numeral m = of_nat r\<close> \<open>numeral n = of_nat s\<close> | |
| 1066 | \<open>numeral (num.Bit0 m) = of_nat (2 * r)\<close> \<open>numeral (num.Bit0 n) = of_nat (2 * s)\<close> | |
| 1067 | \<open>numeral (num.Bit1 m) = of_nat (Suc (2 * r))\<close> | |
| 1068 | by simp_all | |
| 1069 | have **: \<open>Suc (2 * r) div 2 = r\<close> | |
| 1070 | by simp | |
| 1071 | show ?P and ?Q | |
| 1072 | by (simp_all add: divmod_def *) | |
| 1073 | (simp_all flip: of_nat_numeral of_nat_div of_nat_mod of_nat_mult add.commute [of 1] of_nat_Suc | |
| 77061 
5de3772609ea
generalized theory name: euclidean division denotes one particular division definition on integers
 haftmann parents: 
76387diff
changeset | 1074 | add: Euclidean_Rings.mod_mult_mult1 div_mult2_eq [of _ 2] mod_mult2_eq [of _ 2] **) | 
| 76387 | 1075 | qed | 
| 1076 | ||
| 1077 | text \<open>The really hard work\<close> | |
| 1078 | ||
| 1079 | lemma divmod_steps [simp]: | |
| 1080 | "divmod (num.Bit0 m) (num.Bit1 n) = | |
| 1081 | (if m \<le> n then (0, numeral (num.Bit0 m)) | |
| 1082 | else divmod_step (numeral (num.Bit1 n)) | |
| 1083 | (divmod (num.Bit0 m) | |
| 1084 | (num.Bit0 (num.Bit1 n))))" | |
| 1085 | "divmod (num.Bit1 m) (num.Bit1 n) = | |
| 1086 | (if m < n then (0, numeral (num.Bit1 m)) | |
| 1087 | else divmod_step (numeral (num.Bit1 n)) | |
| 1088 | (divmod (num.Bit1 m) | |
| 1089 | (num.Bit0 (num.Bit1 n))))" | |
| 1090 | by (simp_all add: divmod_divmod_step) | |
| 1091 | ||
| 1092 | lemmas divmod_algorithm_code = divmod_trivial divmod_cancel divmod_steps | |
| 1093 | ||
| 1094 | text \<open>Special case: divisibility\<close> | |
| 1095 | ||
| 1096 | definition divides_aux :: "'a \<times> 'a \<Rightarrow> bool" | |
| 1097 | where | |
| 1098 | "divides_aux qr \<longleftrightarrow> snd qr = 0" | |
| 1099 | ||
| 1100 | lemma divides_aux_eq [simp]: | |
| 1101 | "divides_aux (q, r) \<longleftrightarrow> r = 0" | |
| 1102 | by (simp add: divides_aux_def) | |
| 1103 | ||
| 1104 | lemma dvd_numeral_simp [simp]: | |
| 1105 | "numeral m dvd numeral n \<longleftrightarrow> divides_aux (divmod n m)" | |
| 1106 | by (simp add: divmod_def mod_eq_0_iff_dvd) | |
| 1107 | ||
| 1108 | text \<open>Generic computation of quotient and remainder\<close> | |
| 1109 | ||
| 1110 | lemma numeral_div_numeral [simp]: | |
| 1111 | "numeral k div numeral l = fst (divmod k l)" | |
| 1112 | by (simp add: fst_divmod) | |
| 1113 | ||
| 1114 | lemma numeral_mod_numeral [simp]: | |
| 1115 | "numeral k mod numeral l = snd (divmod k l)" | |
| 1116 | by (simp add: snd_divmod) | |
| 1117 | ||
| 1118 | lemma one_div_numeral [simp]: | |
| 1119 | "1 div numeral n = fst (divmod num.One n)" | |
| 1120 | by (simp add: fst_divmod) | |
| 1121 | ||
| 1122 | lemma one_mod_numeral [simp]: | |
| 1123 | "1 mod numeral n = snd (divmod num.One n)" | |
| 1124 | by (simp add: snd_divmod) | |
| 1125 | ||
| 1126 | end | |
| 1127 | ||
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1128 | instantiation nat :: linordered_euclidean_semiring_division | 
| 76387 | 1129 | begin | 
| 1130 | ||
| 1131 | definition divmod_nat :: "num \<Rightarrow> num \<Rightarrow> nat \<times> nat" | |
| 1132 | where | |
| 1133 | divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)" | |
| 1134 | ||
| 1135 | definition divmod_step_nat :: "nat \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat" | |
| 1136 | where | |
| 1137 | "divmod_step_nat l qr = (let (q, r) = qr | |
| 1138 | in if r \<ge> l then (2 * q + 1, r - l) | |
| 1139 | else (2 * q, r))" | |
| 1140 | ||
| 1141 | instance | |
| 1142 | by standard (simp_all add: divmod'_nat_def divmod_step_nat_def) | |
| 1143 | ||
| 1144 | end | |
| 1145 | ||
| 1146 | declare divmod_algorithm_code [where ?'a = nat, code] | |
| 1147 | ||
| 1148 | lemma Suc_0_div_numeral [simp]: | |
| 1149 | \<open>Suc 0 div numeral Num.One = 1\<close> | |
| 1150 | \<open>Suc 0 div numeral (Num.Bit0 n) = 0\<close> | |
| 1151 | \<open>Suc 0 div numeral (Num.Bit1 n) = 0\<close> | |
| 1152 | by simp_all | |
| 1153 | ||
| 1154 | lemma Suc_0_mod_numeral [simp]: | |
| 1155 | \<open>Suc 0 mod numeral Num.One = 0\<close> | |
| 1156 | \<open>Suc 0 mod numeral (Num.Bit0 n) = 1\<close> | |
| 1157 | \<open>Suc 0 mod numeral (Num.Bit1 n) = 1\<close> | |
| 1158 | by simp_all | |
| 1159 | ||
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1160 | instantiation int :: linordered_euclidean_semiring_division | 
| 76387 | 1161 | begin | 
| 1162 | ||
| 1163 | definition divmod_int :: "num \<Rightarrow> num \<Rightarrow> int \<times> int" | |
| 1164 | where | |
| 1165 | "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)" | |
| 1166 | ||
| 1167 | definition divmod_step_int :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int" | |
| 1168 | where | |
| 1169 | "divmod_step_int l qr = (let (q, r) = qr | |
| 1170 | in if \<bar>l\<bar> \<le> \<bar>r\<bar> then (2 * q + 1, r - l) | |
| 1171 | else (2 * q, r))" | |
| 1172 | ||
| 1173 | instance | |
| 1174 | by standard (auto simp add: divmod_int_def divmod_step_int_def) | |
| 1175 | ||
| 1176 | end | |
| 1177 | ||
| 1178 | declare divmod_algorithm_code [where ?'a = int, code] | |
| 1179 | ||
| 1180 | context | |
| 1181 | begin | |
| 1182 | ||
| 1183 | qualified definition adjust_div :: "int \<times> int \<Rightarrow> int" | |
| 1184 | where | |
| 1185 | "adjust_div qr = (let (q, r) = qr in q + of_bool (r \<noteq> 0))" | |
| 1186 | ||
| 1187 | qualified lemma adjust_div_eq [simp, code]: | |
| 1188 | "adjust_div (q, r) = q + of_bool (r \<noteq> 0)" | |
| 1189 | by (simp add: adjust_div_def) | |
| 1190 | ||
| 1191 | qualified definition adjust_mod :: "num \<Rightarrow> int \<Rightarrow> int" | |
| 1192 | where | |
| 1193 | [simp]: "adjust_mod l r = (if r = 0 then 0 else numeral l - r)" | |
| 1194 | ||
| 1195 | lemma minus_numeral_div_numeral [simp]: | |
| 1196 | "- numeral m div numeral n = - (adjust_div (divmod m n) :: int)" | |
| 1197 | proof - | |
| 1198 | have "int (fst (divmod m n)) = fst (divmod m n)" | |
| 1199 | by (simp only: fst_divmod divide_int_def) auto | |
| 1200 | then show ?thesis | |
| 1201 | by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) | |
| 1202 | qed | |
| 1203 | ||
| 1204 | lemma minus_numeral_mod_numeral [simp]: | |
| 1205 | "- numeral m mod numeral n = adjust_mod n (snd (divmod m n) :: int)" | |
| 1206 | proof (cases "snd (divmod m n) = (0::int)") | |
| 1207 | case True | |
| 1208 | then show ?thesis | |
| 1209 | by (simp add: mod_eq_0_iff_dvd divides_aux_def) | |
| 1210 | next | |
| 1211 | case False | |
| 1212 | then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)" | |
| 1213 | by (simp only: snd_divmod modulo_int_def) auto | |
| 1214 | then show ?thesis | |
| 1215 | by (simp add: divides_aux_def adjust_div_def) | |
| 1216 | (simp add: divides_aux_def modulo_int_def) | |
| 1217 | qed | |
| 1218 | ||
| 1219 | lemma numeral_div_minus_numeral [simp]: | |
| 1220 | "numeral m div - numeral n = - (adjust_div (divmod m n) :: int)" | |
| 1221 | proof - | |
| 1222 | have "int (fst (divmod m n)) = fst (divmod m n)" | |
| 1223 | by (simp only: fst_divmod divide_int_def) auto | |
| 1224 | then show ?thesis | |
| 1225 | by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def) | |
| 1226 | qed | |
| 1227 | ||
| 1228 | lemma numeral_mod_minus_numeral [simp]: | |
| 1229 | "numeral m mod - numeral n = - adjust_mod n (snd (divmod m n) :: int)" | |
| 1230 | proof (cases "snd (divmod m n) = (0::int)") | |
| 1231 | case True | |
| 1232 | then show ?thesis | |
| 1233 | by (simp add: mod_eq_0_iff_dvd divides_aux_def) | |
| 1234 | next | |
| 1235 | case False | |
| 1236 | then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)" | |
| 1237 | by (simp only: snd_divmod modulo_int_def) auto | |
| 1238 | then show ?thesis | |
| 1239 | by (simp add: divides_aux_def adjust_div_def) | |
| 1240 | (simp add: divides_aux_def modulo_int_def) | |
| 1241 | qed | |
| 1242 | ||
| 1243 | lemma minus_one_div_numeral [simp]: | |
| 1244 | "- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)" | |
| 1245 | using minus_numeral_div_numeral [of Num.One n] by simp | |
| 1246 | ||
| 1247 | lemma minus_one_mod_numeral [simp]: | |
| 1248 | "- 1 mod numeral n = adjust_mod n (snd (divmod Num.One n) :: int)" | |
| 1249 | using minus_numeral_mod_numeral [of Num.One n] by simp | |
| 1250 | ||
| 1251 | lemma one_div_minus_numeral [simp]: | |
| 1252 | "1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)" | |
| 1253 | using numeral_div_minus_numeral [of Num.One n] by simp | |
| 1254 | ||
| 1255 | lemma one_mod_minus_numeral [simp]: | |
| 1256 | "1 mod - numeral n = - adjust_mod n (snd (divmod Num.One n) :: int)" | |
| 1257 | using numeral_mod_minus_numeral [of Num.One n] by simp | |
| 1258 | ||
| 1259 | lemma [code]: | |
| 1260 | fixes k :: int | |
| 1261 | shows | |
| 1262 | "k div 0 = 0" | |
| 1263 | "k mod 0 = k" | |
| 1264 | "0 div k = 0" | |
| 1265 | "0 mod k = 0" | |
| 1266 | "k div Int.Pos Num.One = k" | |
| 1267 | "k mod Int.Pos Num.One = 0" | |
| 1268 | "k div Int.Neg Num.One = - k" | |
| 1269 | "k mod Int.Neg Num.One = 0" | |
| 1270 | "Int.Pos m div Int.Pos n = (fst (divmod m n) :: int)" | |
| 1271 | "Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)" | |
| 1272 | "Int.Neg m div Int.Pos n = - (adjust_div (divmod m n) :: int)" | |
| 1273 | "Int.Neg m mod Int.Pos n = adjust_mod n (snd (divmod m n) :: int)" | |
| 1274 | "Int.Pos m div Int.Neg n = - (adjust_div (divmod m n) :: int)" | |
| 1275 | "Int.Pos m mod Int.Neg n = - adjust_mod n (snd (divmod m n) :: int)" | |
| 1276 | "Int.Neg m div Int.Neg n = (fst (divmod m n) :: int)" | |
| 1277 | "Int.Neg m mod Int.Neg n = - (snd (divmod m n) :: int)" | |
| 1278 | by simp_all | |
| 1279 | ||
| 1280 | end | |
| 1281 | ||
| 1282 | lemma divmod_BitM_2_eq [simp]: | |
| 1283 | \<open>divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\<close> | |
| 1284 | by (cases m) simp_all | |
| 1285 | ||
| 1286 | ||
| 1287 | subsubsection \<open>Computation by simplification\<close> | |
| 1288 | ||
| 1289 | lemma euclidean_size_nat_less_eq_iff: | |
| 1290 | \<open>euclidean_size m \<le> euclidean_size n \<longleftrightarrow> m \<le> n\<close> for m n :: nat | |
| 1291 | by simp | |
| 1292 | ||
| 1293 | lemma euclidean_size_int_less_eq_iff: | |
| 1294 | \<open>euclidean_size k \<le> euclidean_size l \<longleftrightarrow> \<bar>k\<bar> \<le> \<bar>l\<bar>\<close> for k l :: int | |
| 1295 | by auto | |
| 1296 | ||
| 1297 | simproc_setup numeral_divmod | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1298 |   ("0 div 0 :: 'a :: linordered_euclidean_semiring_division" | "0 mod 0 :: 'a :: linordered_euclidean_semiring_division" |
 | 
| 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1299 | "0 div 1 :: 'a :: linordered_euclidean_semiring_division" | "0 mod 1 :: 'a :: linordered_euclidean_semiring_division" | | 
| 76387 | 1300 | "0 div - 1 :: int" | "0 mod - 1 :: int" | | 
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1301 | "0 div numeral b :: 'a :: linordered_euclidean_semiring_division" | "0 mod numeral b :: 'a :: linordered_euclidean_semiring_division" | | 
| 76387 | 1302 | "0 div - numeral b :: int" | "0 mod - numeral b :: int" | | 
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1303 | "1 div 0 :: 'a :: linordered_euclidean_semiring_division" | "1 mod 0 :: 'a :: linordered_euclidean_semiring_division" | | 
| 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1304 | "1 div 1 :: 'a :: linordered_euclidean_semiring_division" | "1 mod 1 :: 'a :: linordered_euclidean_semiring_division" | | 
| 76387 | 1305 | "1 div - 1 :: int" | "1 mod - 1 :: int" | | 
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1306 | "1 div numeral b :: 'a :: linordered_euclidean_semiring_division" | "1 mod numeral b :: 'a :: linordered_euclidean_semiring_division" | | 
| 76387 | 1307 | "1 div - numeral b :: int" |"1 mod - numeral b :: int" | | 
| 1308 | "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" | | |
| 1309 | "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" | | |
| 1310 | "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" | | |
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1311 | "numeral a div 0 :: 'a :: linordered_euclidean_semiring_division" | "numeral a mod 0 :: 'a :: linordered_euclidean_semiring_division" | | 
| 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1312 | "numeral a div 1 :: 'a :: linordered_euclidean_semiring_division" | "numeral a mod 1 :: 'a :: linordered_euclidean_semiring_division" | | 
| 76387 | 1313 | "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" | | 
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1314 | "numeral a div numeral b :: 'a :: linordered_euclidean_semiring_division" | "numeral a mod numeral b :: 'a :: linordered_euclidean_semiring_division" | | 
| 76387 | 1315 | "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" | | 
| 1316 | "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" | | |
| 1317 | "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" | | |
| 1318 | "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" | | |
| 1319 | "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" | | |
| 1320 | "- numeral a div - numeral b :: int" | "- numeral a mod - numeral b :: int") = \<open> | |
| 1321 | let | |
| 1322 | val if_cong = the (Code.get_case_cong \<^theory> \<^const_name>\<open>If\<close>); | |
| 1323 | fun successful_rewrite ctxt ct = | |
| 1324 | let | |
| 1325 | val thm = Simplifier.rewrite ctxt ct | |
| 1326 | in if Thm.is_reflexive thm then NONE else SOME thm end; | |
| 78082 | 1327 |     val simps = @{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1
 | 
| 1328 | one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral | |
| 1329 | one_div_minus_numeral one_mod_minus_numeral | |
| 1330 | numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral | |
| 1331 | numeral_div_minus_numeral numeral_mod_minus_numeral | |
| 1332 | div_minus_minus mod_minus_minus Parity.adjust_div_eq of_bool_eq one_neq_zero | |
| 1333 | numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial | |
| 1334 | divmod_cancel divmod_steps divmod_step_def fst_conv snd_conv numeral_One | |
| 1335 | case_prod_beta rel_simps Parity.adjust_mod_def div_minus1_right mod_minus1_right | |
| 1336 | minus_minus numeral_times_numeral mult_zero_right mult_1_right | |
| 1337 | euclidean_size_nat_less_eq_iff euclidean_size_int_less_eq_iff diff_nat_numeral nat_numeral} | |
| 1338 |       @ [@{lemma "0 = 0 \<longleftrightarrow> True" by simp}];
 | |
| 78083 | 1339 | val simpset = | 
| 1340 | HOL_ss |> Simplifier.simpset_map \<^context> | |
| 1341 | (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps); | |
| 1342 | in K (fn ctxt => successful_rewrite (Simplifier.put_simpset simpset ctxt)) end | |
| 76387 | 1343 | \<close> \<comment> \<open> | 
| 1344 | There is space for improvement here: the calculation itself | |
| 1345 | could be carried out outside the logic, and a generic simproc | |
| 1346 | (simplifier setup) for generic calculation would be helpful. | |
| 1347 | \<close> | |
| 1348 | ||
| 1349 | ||
| 75937 | 1350 | subsection \<open>Computing congruences modulo \<open>2 ^ q\<close>\<close> | 
| 1351 | ||
| 78937 
5e6b195eee83
slightly less technical formulation of very specific type class
 haftmann parents: 
78668diff
changeset | 1352 | context linordered_euclidean_semiring_division | 
| 75937 | 1353 | begin | 
| 1354 | ||
| 1355 | lemma cong_exp_iff_simps: | |
| 1356 | "numeral n mod numeral Num.One = 0 | |
| 1357 | \<longleftrightarrow> True" | |
| 1358 | "numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) = 0 | |
| 1359 | \<longleftrightarrow> numeral n mod numeral q = 0" | |
| 1360 | "numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) = 0 | |
| 1361 | \<longleftrightarrow> False" | |
| 1362 | "numeral m mod numeral Num.One = (numeral n mod numeral Num.One) | |
| 1363 | \<longleftrightarrow> True" | |
| 1364 | "numeral Num.One mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) | |
| 1365 | \<longleftrightarrow> True" | |
| 1366 | "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) | |
| 1367 | \<longleftrightarrow> False" | |
| 1368 | "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) | |
| 1369 | \<longleftrightarrow> (numeral n mod numeral q) = 0" | |
| 1370 | "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) | |
| 1371 | \<longleftrightarrow> False" | |
| 1372 | "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) | |
| 1373 | \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)" | |
| 1374 | "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) | |
| 1375 | \<longleftrightarrow> False" | |
| 1376 | "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) | |
| 1377 | \<longleftrightarrow> (numeral m mod numeral q) = 0" | |
| 1378 | "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) | |
| 1379 | \<longleftrightarrow> False" | |
| 1380 | "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) | |
| 1381 | \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)" | |
| 1382 | by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even]) | |
| 1383 | ||
| 1384 | end | |
| 1385 | ||
| 1386 | ||
| 71853 | 1387 | code_identifier | 
| 1388 | code_module Parity \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith | |
| 1389 | ||
| 74592 | 1390 | lemmas even_of_nat = even_of_nat_iff | 
| 1391 | ||
| 67816 | 1392 | end |