| author | wenzelm | 
| Mon, 13 Apr 2020 17:40:44 +0200 | |
| changeset 71749 | 77232ff6b8f6 | 
| parent 71535 | b612edee9b0c | 
| child 71755 | 318695613bb7 | 
| 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 | |
| 66815 | 9 | imports Euclidean_Division | 
| 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 + | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 15 | assumes even_iff_mod_2_eq_zero: "2 dvd a \<longleftrightarrow> a mod 2 = 0" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 16 | and odd_iff_mod_2_eq_one: "\<not> 2 dvd a \<longleftrightarrow> a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 17 | and odd_one [simp]: "\<not> 2 dvd 1" | 
| 66839 | 18 | begin | 
| 19 | ||
| 58740 | 20 | abbreviation even :: "'a \<Rightarrow> bool" | 
| 63654 | 21 | where "even a \<equiv> 2 dvd a" | 
| 54228 | 22 | |
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 23 | abbreviation odd :: "'a \<Rightarrow> bool" | 
| 63654 | 24 | where "odd a \<equiv> \<not> 2 dvd a" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 25 | |
| 66815 | 26 | lemma parity_cases [case_names even odd]: | 
| 27 | assumes "even a \<Longrightarrow> a mod 2 = 0 \<Longrightarrow> P" | |
| 28 | assumes "odd a \<Longrightarrow> a mod 2 = 1 \<Longrightarrow> P" | |
| 29 | shows P | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 30 | using assms by (cases "even a") | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 31 | (simp_all add: even_iff_mod_2_eq_zero [symmetric] odd_iff_mod_2_eq_one [symmetric]) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 32 | |
| 71181 | 33 | lemma odd_of_bool_self [simp]: | 
| 34 | \<open>odd (of_bool p) \<longleftrightarrow> p\<close> | |
| 35 | by (cases p) simp_all | |
| 36 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 37 | lemma not_mod_2_eq_0_eq_1 [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 38 | "a mod 2 \<noteq> 0 \<longleftrightarrow> a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 39 | by (cases a rule: parity_cases) simp_all | 
| 66815 | 40 | |
| 41 | lemma not_mod_2_eq_1_eq_0 [simp]: | |
| 42 | "a mod 2 \<noteq> 1 \<longleftrightarrow> a mod 2 = 0" | |
| 43 | by (cases a rule: parity_cases) simp_all | |
| 44 | ||
| 58690 | 45 | lemma evenE [elim?]: | 
| 46 | assumes "even a" | |
| 47 | obtains b where "a = 2 * b" | |
| 58740 | 48 | using assms by (rule dvdE) | 
| 58690 | 49 | |
| 58681 | 50 | lemma oddE [elim?]: | 
| 58680 | 51 | assumes "odd a" | 
| 52 | obtains b where "a = 2 * b + 1" | |
| 58787 | 53 | proof - | 
| 66815 | 54 | have "a = 2 * (a div 2) + a mod 2" | 
| 55 | by (simp add: mult_div_mod_eq) | |
| 56 | with assms have "a = 2 * (a div 2) + 1" | |
| 57 | by (simp add: odd_iff_mod_2_eq_one) | |
| 58 | then show ?thesis .. | |
| 59 | qed | |
| 60 | ||
| 61 | lemma mod_2_eq_odd: | |
| 62 | "a mod 2 = of_bool (odd a)" | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 63 | by (auto elim: oddE simp add: even_iff_mod_2_eq_zero) | 
| 66815 | 64 | |
| 67816 | 65 | lemma of_bool_odd_eq_mod_2: | 
| 66 | "of_bool (odd a) = a mod 2" | |
| 67 | by (simp add: mod_2_eq_odd) | |
| 68 | ||
| 71426 | 69 | lemma even_mod_2_iff [simp]: | 
| 70 | \<open>even (a mod 2) \<longleftrightarrow> even a\<close> | |
| 71 | by (simp add: mod_2_eq_odd) | |
| 72 | ||
| 73 | lemma mod2_eq_if: | |
| 74 | "a mod 2 = (if even a then 0 else 1)" | |
| 75 | by (simp add: mod_2_eq_odd) | |
| 76 | ||
| 66815 | 77 | lemma even_zero [simp]: | 
| 78 | "even 0" | |
| 79 | by (fact dvd_0_right) | |
| 80 | ||
| 81 | lemma odd_even_add: | |
| 82 | "even (a + b)" if "odd a" and "odd b" | |
| 83 | proof - | |
| 84 | from that obtain c d where "a = 2 * c + 1" and "b = 2 * d + 1" | |
| 85 | by (blast elim: oddE) | |
| 86 | then have "a + b = 2 * c + 2 * d + (1 + 1)" | |
| 87 | by (simp only: ac_simps) | |
| 88 | also have "\<dots> = 2 * (c + d + 1)" | |
| 89 | by (simp add: algebra_simps) | |
| 90 | finally show ?thesis .. | |
| 91 | qed | |
| 92 | ||
| 93 | lemma even_add [simp]: | |
| 94 | "even (a + b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)" | |
| 95 | by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add) | |
| 96 | ||
| 97 | lemma odd_add [simp]: | |
| 98 | "odd (a + b) \<longleftrightarrow> \<not> (odd a \<longleftrightarrow> odd b)" | |
| 99 | by simp | |
| 100 | ||
| 101 | lemma even_plus_one_iff [simp]: | |
| 102 | "even (a + 1) \<longleftrightarrow> odd a" | |
| 103 | by (auto simp add: dvd_add_right_iff intro: odd_even_add) | |
| 104 | ||
| 105 | lemma even_mult_iff [simp]: | |
| 106 | "even (a * b) \<longleftrightarrow> even a \<or> even b" (is "?P \<longleftrightarrow> ?Q") | |
| 107 | proof | |
| 108 | assume ?Q | |
| 109 | then show ?P | |
| 110 | by auto | |
| 111 | next | |
| 112 | assume ?P | |
| 113 | show ?Q | |
| 114 | proof (rule ccontr) | |
| 115 | assume "\<not> (even a \<or> even b)" | |
| 116 | then have "odd a" and "odd b" | |
| 117 | by auto | |
| 118 | then obtain r s where "a = 2 * r + 1" and "b = 2 * s + 1" | |
| 119 | by (blast elim: oddE) | |
| 120 | then have "a * b = (2 * r + 1) * (2 * s + 1)" | |
| 121 | by simp | |
| 122 | also have "\<dots> = 2 * (2 * r * s + r + s) + 1" | |
| 123 | by (simp add: algebra_simps) | |
| 124 | finally have "odd (a * b)" | |
| 125 | by simp | |
| 126 | with \<open>?P\<close> show False | |
| 127 | by auto | |
| 128 | qed | |
| 129 | qed | |
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 130 | |
| 63654 | 131 | lemma even_numeral [simp]: "even (numeral (Num.Bit0 n))" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 132 | proof - | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 133 | have "even (2 * numeral n)" | 
| 66815 | 134 | unfolding even_mult_iff by simp | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 135 | then have "even (numeral n + numeral n)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 136 | unfolding mult_2 . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 137 | then show ?thesis | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 138 | unfolding numeral.simps . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 139 | qed | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 140 | |
| 63654 | 141 | lemma odd_numeral [simp]: "odd (numeral (Num.Bit1 n))" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 142 | proof | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 143 | assume "even (numeral (num.Bit1 n))" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 144 | then have "even (numeral n + numeral n + 1)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 145 | unfolding numeral.simps . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 146 | then have "even (2 * numeral n + 1)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 147 | unfolding mult_2 . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 148 | then have "2 dvd numeral n * 2 + 1" | 
| 58740 | 149 | by (simp add: ac_simps) | 
| 63654 | 150 | then have "2 dvd 1" | 
| 151 | 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 | 152 | then show False by simp | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 153 | qed | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 154 | |
| 63654 | 155 | lemma even_power [simp]: "even (a ^ n) \<longleftrightarrow> even a \<and> n > 0" | 
| 58680 | 156 | by (induct n) auto | 
| 157 | ||
| 71412 | 158 | lemma mask_eq_sum_exp: | 
| 159 |   \<open>2 ^ n - 1 = (\<Sum>m\<in>{q. q < n}. 2 ^ m)\<close>
 | |
| 160 | proof - | |
| 161 |   have *: \<open>{q. q < Suc m} = insert m {q. q < m}\<close> for m
 | |
| 162 | by auto | |
| 163 |   have \<open>2 ^ n = (\<Sum>m\<in>{q. q < n}. 2 ^ m) + 1\<close>
 | |
| 164 | by (induction n) (simp_all add: ac_simps mult_2 *) | |
| 165 |   then have \<open>2 ^ n - 1 = (\<Sum>m\<in>{q. q < n}. 2 ^ m) + 1 - 1\<close>
 | |
| 166 | by simp | |
| 167 | then show ?thesis | |
| 168 | by simp | |
| 169 | qed | |
| 170 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 171 | lemma mask_eq_seq_sum: | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 172 | \<open>2 ^ n - 1 = ((\<lambda>k. 1 + k * 2) ^^ n) 0\<close> | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 173 | proof - | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 174 | have \<open>2 ^ n = ((\<lambda>k. 1 + k * 2) ^^ n) 0 + 1\<close> | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 175 | by (induction n) (simp_all add: ac_simps mult_2) | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 176 | then show ?thesis | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 177 | by simp | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 178 | qed | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 179 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 180 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 181 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 182 | class ring_parity = ring + semiring_parity | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 183 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 184 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 185 | subclass comm_ring_1 .. | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 186 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 187 | lemma even_minus: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 188 | "even (- a) \<longleftrightarrow> even a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 189 | by (fact dvd_minus_iff) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 190 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 191 | lemma even_diff [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 192 | "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 | 193 | 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 | 194 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 195 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 196 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 197 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 198 | subsection \<open>Special case: euclidean rings containing the natural numbers\<close> | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 199 | |
| 71157 | 200 | context unique_euclidean_semiring_with_nat | 
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 201 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 202 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 203 | subclass semiring_parity | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 204 | proof | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 205 | show "2 dvd a \<longleftrightarrow> a mod 2 = 0" for a | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 206 | by (fact dvd_eq_mod_eq_0) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 207 | show "\<not> 2 dvd a \<longleftrightarrow> a mod 2 = 1" for a | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 208 | proof | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 209 | assume "a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 210 | then show "\<not> 2 dvd a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 211 | by auto | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 212 | next | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 213 | assume "\<not> 2 dvd a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 214 | have eucl: "euclidean_size (a mod 2) = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 215 | proof (rule order_antisym) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 216 | show "euclidean_size (a mod 2) \<le> 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 217 | using mod_size_less [of 2 a] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 218 | show "1 \<le> euclidean_size (a mod 2)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 219 | using \<open>\<not> 2 dvd a\<close> by (simp add: Suc_le_eq dvd_eq_mod_eq_0) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 220 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 221 | from \<open>\<not> 2 dvd a\<close> have "\<not> of_nat 2 dvd division_segment a * of_nat (euclidean_size a)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 222 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 223 | then have "\<not> of_nat 2 dvd of_nat (euclidean_size a)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 224 | by (auto simp only: dvd_mult_unit_iff' is_unit_division_segment) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 225 | then have "\<not> 2 dvd euclidean_size a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 226 | using of_nat_dvd_iff [of 2] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 227 | then have "euclidean_size a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 228 | by (simp add: semidom_modulo_class.dvd_eq_mod_eq_0) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 229 | then have "of_nat (euclidean_size a mod 2) = of_nat 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 230 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 231 | then have "of_nat (euclidean_size a) mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 232 | by (simp add: of_nat_mod) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 233 | from \<open>\<not> 2 dvd a\<close> eucl | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 234 | show "a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 235 | by (auto intro: division_segment_eq_iff simp add: division_segment_mod) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 236 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 237 | show "\<not> is_unit 2" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 238 | proof (rule notI) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 239 | assume "is_unit 2" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 240 | then have "of_nat 2 dvd of_nat 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 241 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 242 | then have "is_unit (2::nat)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 243 | by (simp only: of_nat_dvd_iff) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 244 | then show False | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 245 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 246 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 247 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 248 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 249 | lemma even_of_nat [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 250 | "even (of_nat a) \<longleftrightarrow> even a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 251 | proof - | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 252 | have "even (of_nat a) \<longleftrightarrow> of_nat 2 dvd of_nat a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 253 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 254 | also have "\<dots> \<longleftrightarrow> even a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 255 | by (simp only: of_nat_dvd_iff) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 256 | finally show ?thesis . | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 257 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 258 | |
| 66815 | 259 | lemma even_succ_div_two [simp]: | 
| 260 | "even a \<Longrightarrow> (a + 1) div 2 = a div 2" | |
| 261 | by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero) | |
| 262 | ||
| 263 | lemma odd_succ_div_two [simp]: | |
| 264 | "odd a \<Longrightarrow> (a + 1) div 2 = a div 2 + 1" | |
| 265 | by (auto elim!: oddE simp add: add.assoc) | |
| 266 | ||
| 267 | lemma even_two_times_div_two: | |
| 268 | "even a \<Longrightarrow> 2 * (a div 2) = a" | |
| 269 | by (fact dvd_mult_div_cancel) | |
| 270 | ||
| 271 | lemma odd_two_times_div_two_succ [simp]: | |
| 272 | "odd a \<Longrightarrow> 2 * (a div 2) + 1 = a" | |
| 273 | using mult_div_mod_eq [of 2 a] | |
| 274 | by (simp add: even_iff_mod_2_eq_zero) | |
| 275 | ||
| 67051 | 276 | lemma coprime_left_2_iff_odd [simp]: | 
| 277 | "coprime 2 a \<longleftrightarrow> odd a" | |
| 278 | proof | |
| 279 | assume "odd a" | |
| 280 | show "coprime 2 a" | |
| 281 | proof (rule coprimeI) | |
| 282 | fix b | |
| 283 | assume "b dvd 2" "b dvd a" | |
| 284 | then have "b dvd a mod 2" | |
| 285 | by (auto intro: dvd_mod) | |
| 286 | with \<open>odd a\<close> show "is_unit b" | |
| 287 | by (simp add: mod_2_eq_odd) | |
| 288 | qed | |
| 289 | next | |
| 290 | assume "coprime 2 a" | |
| 291 | show "odd a" | |
| 292 | proof (rule notI) | |
| 293 | assume "even a" | |
| 294 | then obtain b where "a = 2 * b" .. | |
| 295 | with \<open>coprime 2 a\<close> have "coprime 2 (2 * b)" | |
| 296 | by simp | |
| 297 | moreover have "\<not> coprime 2 (2 * b)" | |
| 298 | by (rule not_coprimeI [of 2]) simp_all | |
| 299 | ultimately show False | |
| 300 | by blast | |
| 301 | qed | |
| 302 | qed | |
| 303 | ||
| 304 | lemma coprime_right_2_iff_odd [simp]: | |
| 305 | "coprime a 2 \<longleftrightarrow> odd a" | |
| 306 | using coprime_left_2_iff_odd [of a] by (simp add: ac_simps) | |
| 307 | ||
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 308 | end | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 309 | |
| 71157 | 310 | context unique_euclidean_ring_with_nat | 
| 58679 | 311 | begin | 
| 312 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 313 | subclass ring_parity .. | 
| 58680 | 314 | |
| 67906 | 315 | lemma minus_1_mod_2_eq [simp]: | 
| 316 | "- 1 mod 2 = 1" | |
| 317 | by (simp add: mod_2_eq_odd) | |
| 318 | ||
| 319 | lemma minus_1_div_2_eq [simp]: | |
| 320 | "- 1 div 2 = - 1" | |
| 321 | proof - | |
| 322 | from div_mult_mod_eq [of "- 1" 2] | |
| 323 | have "- 1 div 2 * 2 = - 1 * 2" | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 324 | using add_implies_diff by fastforce | 
| 67906 | 325 | then show ?thesis | 
| 326 | using mult_right_cancel [of 2 "- 1 div 2" "- 1"] by simp | |
| 327 | qed | |
| 328 | ||
| 58679 | 329 | end | 
| 330 | ||
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 331 | |
| 69593 | 332 | subsection \<open>Instance for \<^typ>\<open>nat\<close>\<close> | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 333 | |
| 70340 | 334 | instance nat :: unique_euclidean_semiring_with_nat | 
| 66815 | 335 | by standard (simp_all add: dvd_eq_mod_eq_0) | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 336 | |
| 66815 | 337 | lemma even_Suc_Suc_iff [simp]: | 
| 338 | "even (Suc (Suc n)) \<longleftrightarrow> even n" | |
| 58787 | 339 | using dvd_add_triv_right_iff [of 2 n] by simp | 
| 58687 | 340 | |
| 66815 | 341 | lemma even_Suc [simp]: "even (Suc n) \<longleftrightarrow> odd n" | 
| 342 | using even_plus_one_iff [of n] by simp | |
| 58787 | 343 | |
| 66815 | 344 | lemma even_diff_nat [simp]: | 
| 345 | "even (m - n) \<longleftrightarrow> m < n \<or> even (m + n)" for m n :: nat | |
| 58787 | 346 | proof (cases "n \<le> m") | 
| 347 | case True | |
| 348 | then have "m - n + n * 2 = m + n" by (simp add: mult_2_right) | |
| 66815 | 349 | moreover have "even (m - n) \<longleftrightarrow> even (m - n + n * 2)" by simp | 
| 350 | ultimately have "even (m - n) \<longleftrightarrow> even (m + n)" by (simp only:) | |
| 58787 | 351 | then show ?thesis by auto | 
| 352 | next | |
| 353 | case False | |
| 354 | then show ?thesis by simp | |
| 63654 | 355 | qed | 
| 356 | ||
| 66815 | 357 | lemma odd_pos: | 
| 358 | "odd n \<Longrightarrow> 0 < n" for n :: nat | |
| 58690 | 359 | by (auto elim: oddE) | 
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 360 | |
| 66815 | 361 | lemma Suc_double_not_eq_double: | 
| 362 | "Suc (2 * m) \<noteq> 2 * n" | |
| 62597 | 363 | proof | 
| 364 | assume "Suc (2 * m) = 2 * n" | |
| 365 | moreover have "odd (Suc (2 * m))" and "even (2 * n)" | |
| 366 | by simp_all | |
| 367 | ultimately show False by simp | |
| 368 | qed | |
| 369 | ||
| 66815 | 370 | lemma double_not_eq_Suc_double: | 
| 371 | "2 * m \<noteq> Suc (2 * n)" | |
| 62597 | 372 | using Suc_double_not_eq_double [of n m] by simp | 
| 373 | ||
| 66815 | 374 | lemma odd_Suc_minus_one [simp]: "odd n \<Longrightarrow> Suc (n - Suc 0) = n" | 
| 375 | by (auto elim: oddE) | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 376 | |
| 66815 | 377 | lemma even_Suc_div_two [simp]: | 
| 378 | "even n \<Longrightarrow> Suc n div 2 = n div 2" | |
| 379 | using even_succ_div_two [of n] by simp | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 380 | |
| 66815 | 381 | lemma odd_Suc_div_two [simp]: | 
| 382 | "odd n \<Longrightarrow> Suc n div 2 = Suc (n div 2)" | |
| 383 | using odd_succ_div_two [of n] by simp | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 384 | |
| 66815 | 385 | lemma odd_two_times_div_two_nat [simp]: | 
| 386 | assumes "odd n" | |
| 387 | shows "2 * (n div 2) = n - (1 :: nat)" | |
| 388 | proof - | |
| 389 | from assms have "2 * (n div 2) + 1 = n" | |
| 390 | by (rule odd_two_times_div_two_succ) | |
| 391 | then have "Suc (2 * (n div 2)) - 1 = n - 1" | |
| 58787 | 392 | by simp | 
| 66815 | 393 | then show ?thesis | 
| 394 | by simp | |
| 58787 | 395 | qed | 
| 58680 | 396 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 397 | 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 | 398 | "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 | 399 | 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 | 400 | |
| 69502 | 401 | lemma odd_card_imp_not_empty: | 
| 402 |   \<open>A \<noteq> {}\<close> if \<open>odd (card A)\<close>
 | |
| 403 | using that by auto | |
| 404 | ||
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 405 | 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 | 406 | 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 | 407 | shows "P n" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 408 | proof (induct n rule: less_induct) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 409 | case (less n) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 410 | show ?case | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 411 | proof (cases "n < Suc (Suc 0)") | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 412 | case True | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 413 | then show ?thesis | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 414 | 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 | 415 | next | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 416 | case False | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 417 | 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 | 418 | 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 | 419 | then have "k<n" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 420 | by simp | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 421 | 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 | 422 | by blast | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 423 | then show ?thesis | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 424 | by (simp add: k) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 425 | qed | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 426 | qed | 
| 58687 | 427 | |
| 71413 | 428 | lemma mask_eq_sum_exp_nat: | 
| 429 |   \<open>2 ^ n - Suc 0 = (\<Sum>m\<in>{q. q < n}. 2 ^ m)\<close>
 | |
| 430 | using mask_eq_sum_exp [where ?'a = nat] by simp | |
| 431 | ||
| 71412 | 432 | context semiring_parity | 
| 433 | begin | |
| 434 | ||
| 435 | lemma even_sum_iff: | |
| 436 |   \<open>even (sum f A) \<longleftrightarrow> even (card {a\<in>A. odd (f a)})\<close> if \<open>finite A\<close>
 | |
| 437 | using that proof (induction A) | |
| 438 | case empty | |
| 439 | then show ?case | |
| 440 | by simp | |
| 441 | next | |
| 442 | case (insert a A) | |
| 443 |   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>
 | |
| 444 | by auto | |
| 445 | ultimately show ?case | |
| 446 | by simp | |
| 447 | qed | |
| 448 | ||
| 449 | lemma even_prod_iff: | |
| 450 | \<open>even (prod f A) \<longleftrightarrow> (\<exists>a\<in>A. even (f a))\<close> if \<open>finite A\<close> | |
| 451 | using that by (induction A) simp_all | |
| 452 | ||
| 453 | lemma even_mask_iff [simp]: | |
| 454 | \<open>even (2 ^ n - 1) \<longleftrightarrow> n = 0\<close> | |
| 455 | proof (cases \<open>n = 0\<close>) | |
| 456 | case True | |
| 457 | then show ?thesis | |
| 458 | by simp | |
| 459 | next | |
| 460 | case False | |
| 461 |   then have \<open>{a. a = 0 \<and> a < n} = {0}\<close>
 | |
| 462 | by auto | |
| 463 | then show ?thesis | |
| 464 | by (auto simp add: mask_eq_sum_exp even_sum_iff) | |
| 465 | qed | |
| 466 | ||
| 467 | end | |
| 468 | ||
| 71138 | 469 | |
| 60758 | 470 | subsection \<open>Parity and powers\<close> | 
| 58689 | 471 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
60867diff
changeset | 472 | context ring_1 | 
| 58689 | 473 | begin | 
| 474 | ||
| 63654 | 475 | lemma power_minus_even [simp]: "even n \<Longrightarrow> (- a) ^ n = a ^ n" | 
| 58690 | 476 | by (auto elim: evenE) | 
| 58689 | 477 | |
| 63654 | 478 | lemma power_minus_odd [simp]: "odd n \<Longrightarrow> (- a) ^ n = - (a ^ n)" | 
| 58690 | 479 | by (auto elim: oddE) | 
| 480 | ||
| 66815 | 481 | lemma uminus_power_if: | 
| 482 | "(- a) ^ n = (if even n then a ^ n else - (a ^ n))" | |
| 483 | by auto | |
| 484 | ||
| 63654 | 485 | lemma neg_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1" | 
| 58690 | 486 | by simp | 
| 58689 | 487 | |
| 63654 | 488 | lemma neg_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1" | 
| 58690 | 489 | by simp | 
| 58689 | 490 | |
| 66582 | 491 | lemma neg_one_power_add_eq_neg_one_power_diff: "k \<le> n \<Longrightarrow> (- 1) ^ (n + k) = (- 1) ^ (n - k)" | 
| 492 | by (cases "even (n + k)") auto | |
| 493 | ||
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 494 | 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 | 495 | by (induct n) auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 496 | |
| 63654 | 497 | end | 
| 58689 | 498 | |
| 499 | context linordered_idom | |
| 500 | begin | |
| 501 | ||
| 63654 | 502 | lemma zero_le_even_power: "even n \<Longrightarrow> 0 \<le> a ^ n" | 
| 58690 | 503 | by (auto elim: evenE) | 
| 58689 | 504 | |
| 63654 | 505 | lemma zero_le_odd_power: "odd n \<Longrightarrow> 0 \<le> a ^ n \<longleftrightarrow> 0 \<le> a" | 
| 58689 | 506 | by (auto simp add: power_even_eq zero_le_mult_iff elim: oddE) | 
| 507 | ||
| 63654 | 508 | lemma zero_le_power_eq: "0 \<le> a ^ n \<longleftrightarrow> even n \<or> odd n \<and> 0 \<le> a" | 
| 58787 | 509 | by (auto simp add: zero_le_even_power zero_le_odd_power) | 
| 63654 | 510 | |
| 511 | 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 | 512 | proof - | 
| 513 | have [simp]: "0 = a ^ n \<longleftrightarrow> a = 0 \<and> n > 0" | |
| 58787 | 514 | unfolding power_eq_0_iff [of a n, symmetric] by blast | 
| 58689 | 515 | show ?thesis | 
| 63654 | 516 | unfolding less_le zero_le_power_eq by auto | 
| 58689 | 517 | qed | 
| 518 | ||
| 63654 | 519 | lemma power_less_zero_eq [simp]: "a ^ n < 0 \<longleftrightarrow> odd n \<and> a < 0" | 
| 58689 | 520 | unfolding not_le [symmetric] zero_le_power_eq by auto | 
| 521 | ||
| 63654 | 522 | 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)" | 
| 523 | unfolding not_less [symmetric] zero_less_power_eq by auto | |
| 524 | ||
| 525 | lemma power_even_abs: "even n \<Longrightarrow> \<bar>a\<bar> ^ n = a ^ n" | |
| 58689 | 526 | using power_abs [of a n] by (simp add: zero_le_even_power) | 
| 527 | ||
| 528 | lemma power_mono_even: | |
| 529 | assumes "even n" and "\<bar>a\<bar> \<le> \<bar>b\<bar>" | |
| 530 | shows "a ^ n \<le> b ^ n" | |
| 531 | proof - | |
| 532 | have "0 \<le> \<bar>a\<bar>" by auto | |
| 63654 | 533 | with \<open>\<bar>a\<bar> \<le> \<bar>b\<bar>\<close> have "\<bar>a\<bar> ^ n \<le> \<bar>b\<bar> ^ n" | 
| 534 | by (rule power_mono) | |
| 535 | with \<open>even n\<close> show ?thesis | |
| 536 | by (simp add: power_even_abs) | |
| 58689 | 537 | qed | 
| 538 | ||
| 539 | lemma power_mono_odd: | |
| 540 | assumes "odd n" and "a \<le> b" | |
| 541 | shows "a ^ n \<le> b ^ n" | |
| 542 | proof (cases "b < 0") | |
| 63654 | 543 | case True | 
| 544 | with \<open>a \<le> b\<close> have "- b \<le> - a" and "0 \<le> - b" by auto | |
| 545 | then have "(- b) ^ n \<le> (- a) ^ n" by (rule power_mono) | |
| 60758 | 546 | with \<open>odd n\<close> show ?thesis by simp | 
| 58689 | 547 | next | 
| 63654 | 548 | case False | 
| 549 | then have "0 \<le> b" by auto | |
| 58689 | 550 | show ?thesis | 
| 551 | proof (cases "a < 0") | |
| 63654 | 552 | case True | 
| 553 | then have "n \<noteq> 0" and "a \<le> 0" using \<open>odd n\<close> [THEN odd_pos] by auto | |
| 60758 | 554 | then have "a ^ n \<le> 0" unfolding power_le_zero_eq using \<open>odd n\<close> by auto | 
| 63654 | 555 | moreover from \<open>0 \<le> b\<close> have "0 \<le> b ^ n" by auto | 
| 58689 | 556 | ultimately show ?thesis by auto | 
| 557 | next | |
| 63654 | 558 | case False | 
| 559 | then have "0 \<le> a" by auto | |
| 560 | with \<open>a \<le> b\<close> show ?thesis | |
| 561 | using power_mono by auto | |
| 58689 | 562 | qed | 
| 563 | qed | |
| 62083 | 564 | |
| 60758 | 565 | text \<open>Simplify, when the exponent is a numeral\<close> | 
| 58689 | 566 | |
| 567 | lemma zero_le_power_eq_numeral [simp]: | |
| 568 | "0 \<le> a ^ numeral w \<longleftrightarrow> even (numeral w :: nat) \<or> odd (numeral w :: nat) \<and> 0 \<le> a" | |
| 569 | by (fact zero_le_power_eq) | |
| 570 | ||
| 571 | lemma zero_less_power_eq_numeral [simp]: | |
| 63654 | 572 | "0 < a ^ numeral w \<longleftrightarrow> | 
| 573 | numeral w = (0 :: nat) \<or> | |
| 574 | even (numeral w :: nat) \<and> a \<noteq> 0 \<or> | |
| 575 | odd (numeral w :: nat) \<and> 0 < a" | |
| 58689 | 576 | by (fact zero_less_power_eq) | 
| 577 | ||
| 578 | lemma power_le_zero_eq_numeral [simp]: | |
| 63654 | 579 | "a ^ numeral w \<le> 0 \<longleftrightarrow> | 
| 580 | (0 :: nat) < numeral w \<and> | |
| 581 | (odd (numeral w :: nat) \<and> a \<le> 0 \<or> even (numeral w :: nat) \<and> a = 0)" | |
| 58689 | 582 | by (fact power_le_zero_eq) | 
| 583 | ||
| 584 | lemma power_less_zero_eq_numeral [simp]: | |
| 585 | "a ^ numeral w < 0 \<longleftrightarrow> odd (numeral w :: nat) \<and> a < 0" | |
| 586 | by (fact power_less_zero_eq) | |
| 587 | ||
| 588 | lemma power_even_abs_numeral [simp]: | |
| 589 | "even (numeral w :: nat) \<Longrightarrow> \<bar>a\<bar> ^ numeral w = a ^ numeral w" | |
| 590 | by (fact power_even_abs) | |
| 591 | ||
| 592 | end | |
| 593 | ||
| 71413 | 594 | context unique_euclidean_semiring_with_nat | 
| 595 | begin | |
| 596 | ||
| 597 | lemma even_mask_div_iff': | |
| 598 | \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> m \<le> n\<close> | |
| 599 | proof - | |
| 600 | have \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> even (of_nat ((2 ^ m - Suc 0) div 2 ^ n))\<close> | |
| 601 | by (simp only: of_nat_div) (simp add: of_nat_diff) | |
| 602 | also have \<open>\<dots> \<longleftrightarrow> even ((2 ^ m - Suc 0) div 2 ^ n)\<close> | |
| 603 | by simp | |
| 604 | also have \<open>\<dots> \<longleftrightarrow> m \<le> n\<close> | |
| 605 | proof (cases \<open>m \<le> n\<close>) | |
| 606 | case True | |
| 607 | then show ?thesis | |
| 608 | by (simp add: Suc_le_lessD) | |
| 609 | next | |
| 610 | case False | |
| 611 | then obtain r where r: \<open>m = n + Suc r\<close> | |
| 612 | using less_imp_Suc_add by fastforce | |
| 613 |     from r have \<open>{q. q < m} \<inter> {q. 2 ^ n dvd (2::nat) ^ q} = {q. n \<le> q \<and> q < m}\<close>
 | |
| 614 | by (auto simp add: dvd_power_iff_le) | |
| 615 |     moreover from r have \<open>{q. q < m} \<inter> {q. \<not> 2 ^ n dvd (2::nat) ^ q} = {q. q < n}\<close>
 | |
| 616 | by (auto simp add: dvd_power_iff_le) | |
| 617 |     moreover from False have \<open>{q. n \<le> q \<and> q < m \<and> q \<le> n} = {n}\<close>
 | |
| 618 | by auto | |
| 619 |     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>
 | |
| 620 | by (simp_all add: euclidean_semiring_cancel_class.power_diff_power_eq semiring_parity_class.even_sum_iff not_less mask_eq_sum_exp_nat [symmetric]) | |
| 621 |     ultimately have \<open>odd (sum ((^) (2::nat)) {q. q < m} div 2 ^ n)\<close>
 | |
| 622 | by (subst euclidean_semiring_cancel_class.sum_div_partition) simp_all | |
| 623 | with False show ?thesis | |
| 624 | by (simp add: mask_eq_sum_exp_nat) | |
| 625 | qed | |
| 626 | finally show ?thesis . | |
| 627 | qed | |
| 628 | ||
| 629 | end | |
| 630 | ||
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 631 | |
| 69593 | 632 | subsection \<open>Instance for \<^typ>\<open>int\<close>\<close> | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 633 | |
| 67816 | 634 | lemma even_diff_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 635 | "even (k - l) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 636 | by (fact even_diff) | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 637 | |
| 67816 | 638 | lemma even_abs_add_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 639 | "even (\<bar>k\<bar> + l) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 640 | by simp | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 641 | |
| 67816 | 642 | lemma even_add_abs_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 643 | "even (k + \<bar>l\<bar>) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 644 | by simp | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 645 | |
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 646 | lemma even_nat_iff: "0 \<le> k \<Longrightarrow> even (nat k) \<longleftrightarrow> even k" | 
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 647 | by (simp add: even_of_nat [of "nat k", where ?'a = int, symmetric]) | 
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 648 | |
| 71138 | 649 | lemma zdiv_zmult2_eq: | 
| 650 | \<open>a div (b * c) = (a div b) div c\<close> if \<open>c \<ge> 0\<close> for a b c :: int | |
| 651 | proof (cases \<open>b \<ge> 0\<close>) | |
| 652 | case True | |
| 653 | with that show ?thesis | |
| 654 | using div_mult2_eq' [of a \<open>nat b\<close> \<open>nat c\<close>] by simp | |
| 655 | next | |
| 656 | case False | |
| 657 | with that show ?thesis | |
| 658 | using div_mult2_eq' [of \<open>- a\<close> \<open>nat (- b)\<close> \<open>nat c\<close>] by simp | |
| 659 | qed | |
| 660 | ||
| 661 | lemma zmod_zmult2_eq: | |
| 662 | \<open>a mod (b * c) = b * (a div b mod c) + a mod b\<close> if \<open>c \<ge> 0\<close> for a b c :: int | |
| 663 | proof (cases \<open>b \<ge> 0\<close>) | |
| 664 | case True | |
| 665 | with that show ?thesis | |
| 666 | using mod_mult2_eq' [of a \<open>nat b\<close> \<open>nat c\<close>] by simp | |
| 667 | next | |
| 668 | case False | |
| 669 | with that show ?thesis | |
| 670 | using mod_mult2_eq' [of \<open>- a\<close> \<open>nat (- b)\<close> \<open>nat c\<close>] by simp | |
| 671 | qed | |
| 672 | ||
| 71094 | 673 | |
| 71181 | 674 | subsection \<open>Abstract bit structures\<close> | 
| 71094 | 675 | |
| 676 | class semiring_bits = semiring_parity + | |
| 71195 | 677 | assumes bits_induct [case_names stable rec]: | 
| 71094 | 678 | \<open>(\<And>a. a div 2 = a \<Longrightarrow> P a) | 
| 679 | \<Longrightarrow> (\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a)) | |
| 680 | \<Longrightarrow> P a\<close> | |
| 71138 | 681 | assumes bits_div_0 [simp]: \<open>0 div a = 0\<close> | 
| 682 | and bits_div_by_1 [simp]: \<open>a div 1 = a\<close> | |
| 71195 | 683 | and bits_mod_div_trivial [simp]: \<open>a mod b div b = 0\<close> | 
| 71138 | 684 | and even_succ_div_2 [simp]: \<open>even a \<Longrightarrow> (1 + a) div 2 = a div 2\<close> | 
| 71413 | 685 | and even_mask_div_iff: \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> 2 ^ n = 0 \<or> m \<le> n\<close> | 
| 71182 | 686 | and exp_div_exp_eq: \<open>2 ^ m div 2 ^ n = of_bool (2 ^ m \<noteq> 0 \<and> m \<ge> n) * 2 ^ (m - n)\<close> | 
| 71138 | 687 | and div_exp_eq: \<open>a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)\<close> | 
| 688 | and mod_exp_eq: \<open>a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n\<close> | |
| 689 | and mult_exp_mod_exp_eq: \<open>m \<le> n \<Longrightarrow> (a * 2 ^ m) mod (2 ^ n) = (a mod 2 ^ (n - m)) * 2 ^ m\<close> | |
| 690 | and div_exp_mod_exp_eq: \<open>a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\<close> | |
| 71424 | 691 | and even_mult_exp_div_exp_iff: \<open>even (a * 2 ^ m div 2 ^ n) \<longleftrightarrow> m > n \<or> 2 ^ n = 0 \<or> (m \<le> n \<and> even (a div 2 ^ (n - m)))\<close> | 
| 71138 | 692 | begin | 
| 693 | ||
| 71195 | 694 | lemma bits_div_by_0 [simp]: | 
| 695 | \<open>a div 0 = 0\<close> | |
| 696 | by (metis add_cancel_right_right bits_mod_div_trivial mod_mult_div_eq mult_not_zero) | |
| 697 | ||
| 71138 | 698 | lemma bits_1_div_2 [simp]: | 
| 699 | \<open>1 div 2 = 0\<close> | |
| 700 | using even_succ_div_2 [of 0] by simp | |
| 701 | ||
| 702 | lemma bits_1_div_exp [simp]: | |
| 703 | \<open>1 div 2 ^ n = of_bool (n = 0)\<close> | |
| 704 | using div_exp_eq [of 1 1] by (cases n) simp_all | |
| 705 | ||
| 706 | lemma even_succ_div_exp [simp]: | |
| 707 | \<open>(1 + a) div 2 ^ n = a div 2 ^ n\<close> if \<open>even a\<close> and \<open>n > 0\<close> | |
| 708 | proof (cases n) | |
| 709 | case 0 | |
| 710 | with that show ?thesis | |
| 711 | by simp | |
| 712 | next | |
| 713 | case (Suc n) | |
| 714 | with \<open>even a\<close> have \<open>(1 + a) div 2 ^ Suc n = a div 2 ^ Suc n\<close> | |
| 715 | proof (induction n) | |
| 716 | case 0 | |
| 717 | then show ?case | |
| 718 | by simp | |
| 719 | next | |
| 720 | case (Suc n) | |
| 721 | then show ?case | |
| 722 | using div_exp_eq [of _ 1 \<open>Suc n\<close>, symmetric] | |
| 723 | by simp | |
| 724 | qed | |
| 725 | with Suc show ?thesis | |
| 726 | by simp | |
| 727 | qed | |
| 728 | ||
| 729 | lemma even_succ_mod_exp [simp]: | |
| 730 | \<open>(1 + a) mod 2 ^ n = 1 + (a mod 2 ^ n)\<close> if \<open>even a\<close> and \<open>n > 0\<close> | |
| 731 | using div_mult_mod_eq [of \<open>1 + a\<close> \<open>2 ^ n\<close>] that | |
| 732 | apply simp | |
| 733 | by (metis local.add.left_commute local.add_left_cancel local.div_mult_mod_eq) | |
| 734 | ||
| 735 | lemma bits_mod_by_1 [simp]: | |
| 736 | \<open>a mod 1 = 0\<close> | |
| 737 | using div_mult_mod_eq [of a 1] by simp | |
| 738 | ||
| 739 | lemma bits_mod_0 [simp]: | |
| 740 | \<open>0 mod a = 0\<close> | |
| 741 | using div_mult_mod_eq [of 0 a] by simp | |
| 742 | ||
| 71195 | 743 | lemma bits_one_mod_two_eq_one [simp]: | 
| 71138 | 744 | \<open>1 mod 2 = 1\<close> | 
| 745 | by (simp add: mod2_eq_if) | |
| 746 | ||
| 71181 | 747 | definition bit :: \<open>'a \<Rightarrow> nat \<Rightarrow> bool\<close> | 
| 748 | where \<open>bit a n \<longleftrightarrow> odd (a div 2 ^ n)\<close> | |
| 749 | ||
| 750 | lemma bit_0 [simp]: | |
| 751 | \<open>bit a 0 \<longleftrightarrow> odd a\<close> | |
| 752 | by (simp add: bit_def) | |
| 753 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 754 | lemma bit_Suc: | 
| 71181 | 755 | \<open>bit a (Suc n) \<longleftrightarrow> bit (a div 2) n\<close> | 
| 756 | using div_exp_eq [of a 1 n] by (simp add: bit_def) | |
| 757 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 758 | lemma bit_rec: | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 759 | \<open>bit a n \<longleftrightarrow> (if n = 0 then odd a else bit (a div 2) (n - 1))\<close> | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 760 | by (cases n) (simp_all add: bit_Suc) | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 761 | |
| 71195 | 762 | lemma bit_0_eq [simp]: | 
| 763 | \<open>bit 0 = bot\<close> | |
| 764 | by (simp add: fun_eq_iff bit_def) | |
| 765 | ||
| 71181 | 766 | context | 
| 767 | fixes a | |
| 768 | assumes stable: \<open>a div 2 = a\<close> | |
| 769 | begin | |
| 770 | ||
| 71195 | 771 | lemma bits_stable_imp_add_self: | 
| 71181 | 772 | \<open>a + a mod 2 = 0\<close> | 
| 773 | proof - | |
| 774 | have \<open>a div 2 * 2 + a mod 2 = a\<close> | |
| 775 | by (fact div_mult_mod_eq) | |
| 776 | then have \<open>a * 2 + a mod 2 = a\<close> | |
| 777 | by (simp add: stable) | |
| 778 | then show ?thesis | |
| 779 | by (simp add: mult_2_right ac_simps) | |
| 780 | qed | |
| 781 | ||
| 782 | lemma stable_imp_bit_iff_odd: | |
| 783 | \<open>bit a n \<longleftrightarrow> odd a\<close> | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 784 | by (induction n) (simp_all add: stable bit_Suc) | 
| 71181 | 785 | |
| 786 | end | |
| 787 | ||
| 788 | lemma bit_iff_idd_imp_stable: | |
| 789 | \<open>a div 2 = a\<close> if \<open>\<And>n. bit a n \<longleftrightarrow> odd a\<close> | |
| 71195 | 790 | using that proof (induction a rule: bits_induct) | 
| 71181 | 791 | case (stable a) | 
| 792 | then show ?case | |
| 793 | by simp | |
| 794 | next | |
| 795 | case (rec a b) | |
| 796 | from rec.prems [of 1] have [simp]: \<open>b = odd a\<close> | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 797 | by (simp add: rec.hyps bit_Suc) | 
| 71181 | 798 | from rec.hyps have hyp: \<open>(of_bool (odd a) + 2 * a) div 2 = a\<close> | 
| 799 | by simp | |
| 800 | have \<open>bit a n \<longleftrightarrow> odd a\<close> for n | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 801 | using rec.prems [of \<open>Suc n\<close>] by (simp add: hyp bit_Suc) | 
| 71181 | 802 | then have \<open>a div 2 = a\<close> | 
| 803 | by (rule rec.IH) | |
| 804 | then have \<open>of_bool (odd a) + 2 * a = 2 * (a div 2) + of_bool (odd a)\<close> | |
| 805 | by (simp add: ac_simps) | |
| 806 | also have \<open>\<dots> = a\<close> | |
| 807 | using mult_div_mod_eq [of 2 a] | |
| 808 | by (simp add: of_bool_odd_eq_mod_2) | |
| 809 | finally show ?case | |
| 810 | using \<open>a div 2 = a\<close> by (simp add: hyp) | |
| 811 | qed | |
| 812 | ||
| 71418 | 813 | lemma exp_eq_0_imp_not_bit: | 
| 814 | \<open>\<not> bit a n\<close> if \<open>2 ^ n = 0\<close> | |
| 815 | using that by (simp add: bit_def) | |
| 816 | ||
| 71181 | 817 | lemma bit_eqI: | 
| 71418 | 818 | \<open>a = b\<close> if \<open>\<And>n. 2 ^ n \<noteq> 0 \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close> | 
| 819 | proof - | |
| 820 | have \<open>bit a n \<longleftrightarrow> bit b n\<close> for n | |
| 821 | proof (cases \<open>2 ^ n = 0\<close>) | |
| 822 | case True | |
| 823 | then show ?thesis | |
| 824 | by (simp add: exp_eq_0_imp_not_bit) | |
| 825 | next | |
| 826 | case False | |
| 827 | then show ?thesis | |
| 828 | by (rule that) | |
| 71181 | 829 | qed | 
| 71418 | 830 | then show ?thesis proof (induction a arbitrary: b rule: bits_induct) | 
| 831 | case (stable a) | |
| 832 | from stable(2) [of 0] have **: \<open>even b \<longleftrightarrow> even a\<close> | |
| 833 | by simp | |
| 834 | have \<open>b div 2 = b\<close> | |
| 835 | proof (rule bit_iff_idd_imp_stable) | |
| 836 | fix n | |
| 837 | from stable have *: \<open>bit b n \<longleftrightarrow> bit a n\<close> | |
| 838 | by simp | |
| 839 | also have \<open>bit a n \<longleftrightarrow> odd a\<close> | |
| 840 | using stable by (simp add: stable_imp_bit_iff_odd) | |
| 841 | finally show \<open>bit b n \<longleftrightarrow> odd b\<close> | |
| 842 | by (simp add: **) | |
| 843 | qed | |
| 844 | from ** have \<open>a mod 2 = b mod 2\<close> | |
| 845 | by (simp add: mod2_eq_if) | |
| 846 | then have \<open>a mod 2 + (a + b) = b mod 2 + (a + b)\<close> | |
| 847 | by simp | |
| 848 | then have \<open>a + a mod 2 + b = b + b mod 2 + a\<close> | |
| 849 | by (simp add: ac_simps) | |
| 850 | with \<open>a div 2 = a\<close> \<open>b div 2 = b\<close> show ?case | |
| 851 | by (simp add: bits_stable_imp_add_self) | |
| 852 | next | |
| 853 | case (rec a p) | |
| 854 | from rec.prems [of 0] have [simp]: \<open>p = odd b\<close> | |
| 855 | by simp | |
| 856 | from rec.hyps have \<open>bit a n \<longleftrightarrow> bit (b div 2) n\<close> for n | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 857 | using rec.prems [of \<open>Suc n\<close>] by (simp add: bit_Suc) | 
| 71418 | 858 | then have \<open>a = b div 2\<close> | 
| 859 | by (rule rec.IH) | |
| 860 | then have \<open>2 * a = 2 * (b div 2)\<close> | |
| 861 | by simp | |
| 862 | then have \<open>b mod 2 + 2 * a = b mod 2 + 2 * (b div 2)\<close> | |
| 863 | by simp | |
| 864 | also have \<open>\<dots> = b\<close> | |
| 865 | by (fact mod_mult_div_eq) | |
| 866 | finally show ?case | |
| 867 | by (auto simp add: mod2_eq_if) | |
| 868 | qed | |
| 71181 | 869 | qed | 
| 870 | ||
| 871 | lemma bit_eq_iff: | |
| 872 | \<open>a = b \<longleftrightarrow> (\<forall>n. bit a n \<longleftrightarrow> bit b n)\<close> | |
| 873 | by (auto intro: bit_eqI) | |
| 874 | ||
| 71182 | 875 | lemma bit_exp_iff: | 
| 876 | \<open>bit (2 ^ m) n \<longleftrightarrow> 2 ^ m \<noteq> 0 \<and> m = n\<close> | |
| 877 | by (auto simp add: bit_def exp_div_exp_eq) | |
| 878 | ||
| 71408 | 879 | lemma bit_1_iff: | 
| 880 | \<open>bit 1 n \<longleftrightarrow> 1 \<noteq> 0 \<and> n = 0\<close> | |
| 881 | using bit_exp_iff [of 0 n] by simp | |
| 882 | ||
| 883 | lemma bit_2_iff: | |
| 884 | \<open>bit 2 n \<longleftrightarrow> 2 \<noteq> 0 \<and> n = 1\<close> | |
| 885 | using bit_exp_iff [of 1 n] by auto | |
| 886 | ||
| 71418 | 887 | lemma even_bit_succ_iff: | 
| 888 | \<open>bit (1 + a) n \<longleftrightarrow> bit a n \<or> n = 0\<close> if \<open>even a\<close> | |
| 889 | using that by (cases \<open>n = 0\<close>) (simp_all add: bit_def) | |
| 890 | ||
| 891 | lemma odd_bit_iff_bit_pred: | |
| 892 | \<open>bit a n \<longleftrightarrow> bit (a - 1) n \<or> n = 0\<close> if \<open>odd a\<close> | |
| 893 | proof - | |
| 894 | from \<open>odd a\<close> obtain b where \<open>a = 2 * b + 1\<close> .. | |
| 895 | moreover have \<open>bit (2 * b) n \<or> n = 0 \<longleftrightarrow> bit (1 + 2 * b) n\<close> | |
| 896 | using even_bit_succ_iff by simp | |
| 897 | ultimately show ?thesis by (simp add: ac_simps) | |
| 898 | qed | |
| 899 | ||
| 71426 | 900 | lemma bit_double_iff: | 
| 901 | \<open>bit (2 * a) n \<longleftrightarrow> bit a (n - 1) \<and> n \<noteq> 0 \<and> 2 ^ n \<noteq> 0\<close> | |
| 902 | using even_mult_exp_div_exp_iff [of a 1 n] | |
| 903 | by (cases n, auto simp add: bit_def ac_simps) | |
| 904 | ||
| 905 | lemma bit_eq_rec: | |
| 71441 | 906 | \<open>a = b \<longleftrightarrow> (even a \<longleftrightarrow> even b) \<and> a div 2 = b div 2\<close> (is \<open>?P = ?Q\<close>) | 
| 907 | proof | |
| 908 | assume ?P | |
| 909 | then show ?Q | |
| 910 | by simp | |
| 911 | next | |
| 912 | assume ?Q | |
| 913 | then have \<open>even a \<longleftrightarrow> even b\<close> and \<open>a div 2 = b div 2\<close> | |
| 914 | by simp_all | |
| 915 | show ?P | |
| 916 | proof (rule bit_eqI) | |
| 917 | fix n | |
| 918 | show \<open>bit a n \<longleftrightarrow> bit b n\<close> | |
| 919 | proof (cases n) | |
| 920 | case 0 | |
| 921 | with \<open>even a \<longleftrightarrow> even b\<close> show ?thesis | |
| 922 | by simp | |
| 923 | next | |
| 924 | case (Suc n) | |
| 925 | moreover from \<open>a div 2 = b div 2\<close> have \<open>bit (a div 2) n = bit (b div 2) n\<close> | |
| 926 | by simp | |
| 927 | ultimately show ?thesis | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 928 | by (simp add: bit_Suc) | 
| 71441 | 929 | qed | 
| 930 | qed | |
| 931 | qed | |
| 71426 | 932 | |
| 71418 | 933 | lemma bit_mask_iff: | 
| 934 | \<open>bit (2 ^ m - 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n < m\<close> | |
| 935 | by (simp add: bit_def even_mask_div_iff not_le) | |
| 936 | ||
| 71138 | 937 | end | 
| 71094 | 938 | |
| 939 | lemma nat_bit_induct [case_names zero even odd]: | |
| 940 | "P n" if zero: "P 0" | |
| 941 | and even: "\<And>n. P n \<Longrightarrow> n > 0 \<Longrightarrow> P (2 * n)" | |
| 942 | and odd: "\<And>n. P n \<Longrightarrow> P (Suc (2 * n))" | |
| 943 | proof (induction n rule: less_induct) | |
| 944 | case (less n) | |
| 945 | show "P n" | |
| 946 | proof (cases "n = 0") | |
| 947 | case True with zero show ?thesis by simp | |
| 948 | next | |
| 949 | case False | |
| 950 | with less have hyp: "P (n div 2)" by simp | |
| 951 | show ?thesis | |
| 952 | proof (cases "even n") | |
| 953 | case True | |
| 954 | then have "n \<noteq> 1" | |
| 955 | by auto | |
| 956 | with \<open>n \<noteq> 0\<close> have "n div 2 > 0" | |
| 957 | by simp | |
| 958 | with \<open>even n\<close> hyp even [of "n div 2"] show ?thesis | |
| 959 | by simp | |
| 960 | next | |
| 961 | case False | |
| 962 | with hyp odd [of "n div 2"] show ?thesis | |
| 963 | by simp | |
| 964 | qed | |
| 965 | qed | |
| 966 | qed | |
| 967 | ||
| 968 | instance nat :: semiring_bits | |
| 969 | proof | |
| 970 | show \<open>P n\<close> if stable: \<open>\<And>n. n div 2 = n \<Longrightarrow> P n\<close> | |
| 971 | and rec: \<open>\<And>n b. P n \<Longrightarrow> (of_bool b + 2 * n) div 2 = n \<Longrightarrow> P (of_bool b + 2 * n)\<close> | |
| 972 | for P and n :: nat | |
| 973 | proof (induction n rule: nat_bit_induct) | |
| 974 | case zero | |
| 975 | from stable [of 0] show ?case | |
| 976 | by simp | |
| 977 | next | |
| 978 | case (even n) | |
| 979 | with rec [of n False] show ?case | |
| 980 | by simp | |
| 981 | next | |
| 982 | case (odd n) | |
| 983 | with rec [of n True] show ?case | |
| 984 | by simp | |
| 985 | qed | |
| 71138 | 986 | show \<open>q mod 2 ^ m mod 2 ^ n = q mod 2 ^ min m n\<close> | 
| 987 | for q m n :: nat | |
| 988 | apply (auto simp add: less_iff_Suc_add power_add mod_mod_cancel split: split_min_lin) | |
| 989 | apply (metis div_mult2_eq mod_div_trivial mod_eq_self_iff_div_eq_0 mod_mult_self2_is_0 power_commutes) | |
| 990 | done | |
| 991 | show \<open>(q * 2 ^ m) mod (2 ^ n) = (q mod 2 ^ (n - m)) * 2 ^ m\<close> if \<open>m \<le> n\<close> | |
| 992 | for q m n :: nat | |
| 993 | using that | |
| 994 | apply (auto simp add: mod_mod_cancel div_mult2_eq power_add mod_mult2_eq le_iff_add split: split_min_lin) | |
| 995 | apply (simp add: mult.commute) | |
| 996 | done | |
| 71413 | 997 | show \<open>even ((2 ^ m - (1::nat)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::nat) \<or> m \<le> n\<close> | 
| 998 | for m n :: nat | |
| 999 | using even_mask_div_iff' [where ?'a = nat, of m n] by simp | |
| 71424 | 1000 | show \<open>even (q * 2 ^ m div 2 ^ n) \<longleftrightarrow> n < m \<or> (2::nat) ^ n = 0 \<or> m \<le> n \<and> even (q div 2 ^ (n - m))\<close> | 
| 1001 | for m n q r :: nat | |
| 1002 | apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex) | |
| 1003 | apply (metis (full_types) dvd_mult dvd_mult_imp_div dvd_power_iff_le not_less not_less_eq order_refl power_Suc) | |
| 1004 | done | |
| 71182 | 1005 | qed (auto simp add: div_mult2_eq mod_mult2_eq power_add power_diff) | 
| 71094 | 1006 | |
| 70353 | 1007 | lemma int_bit_induct [case_names zero minus even odd]: | 
| 70338 | 1008 | "P k" if zero_int: "P 0" | 
| 1009 | and minus_int: "P (- 1)" | |
| 1010 | and even_int: "\<And>k. P k \<Longrightarrow> k \<noteq> 0 \<Longrightarrow> P (k * 2)" | |
| 1011 | and odd_int: "\<And>k. P k \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> P (1 + (k * 2))" for k :: int | |
| 1012 | proof (cases "k \<ge> 0") | |
| 1013 | case True | |
| 1014 | define n where "n = nat k" | |
| 1015 | with True have "k = int n" | |
| 1016 | by simp | |
| 1017 | then show "P k" | |
| 70353 | 1018 | proof (induction n arbitrary: k rule: nat_bit_induct) | 
| 70338 | 1019 | case zero | 
| 1020 | then show ?case | |
| 1021 | by (simp add: zero_int) | |
| 1022 | next | |
| 1023 | case (even n) | |
| 1024 | have "P (int n * 2)" | |
| 1025 | by (rule even_int) (use even in simp_all) | |
| 1026 | with even show ?case | |
| 1027 | by (simp add: ac_simps) | |
| 1028 | next | |
| 1029 | case (odd n) | |
| 1030 | have "P (1 + (int n * 2))" | |
| 1031 | by (rule odd_int) (use odd in simp_all) | |
| 1032 | with odd show ?case | |
| 1033 | by (simp add: ac_simps) | |
| 1034 | qed | |
| 1035 | next | |
| 1036 | case False | |
| 1037 | define n where "n = nat (- k - 1)" | |
| 1038 | with False have "k = - int n - 1" | |
| 1039 | by simp | |
| 1040 | then show "P k" | |
| 70353 | 1041 | proof (induction n arbitrary: k rule: nat_bit_induct) | 
| 70338 | 1042 | case zero | 
| 1043 | then show ?case | |
| 1044 | by (simp add: minus_int) | |
| 1045 | next | |
| 1046 | case (even n) | |
| 1047 | have "P (1 + (- int (Suc n) * 2))" | |
| 1048 | by (rule odd_int) (use even in \<open>simp_all add: algebra_simps\<close>) | |
| 1049 | also have "\<dots> = - int (2 * n) - 1" | |
| 1050 | by (simp add: algebra_simps) | |
| 1051 | finally show ?case | |
| 1052 | using even by simp | |
| 1053 | next | |
| 1054 | case (odd n) | |
| 1055 | have "P (- int (Suc n) * 2)" | |
| 1056 | by (rule even_int) (use odd in \<open>simp_all add: algebra_simps\<close>) | |
| 1057 | also have "\<dots> = - int (Suc (2 * n)) - 1" | |
| 1058 | by (simp add: algebra_simps) | |
| 1059 | finally show ?case | |
| 1060 | using odd by simp | |
| 1061 | qed | |
| 1062 | qed | |
| 1063 | ||
| 71094 | 1064 | instance int :: semiring_bits | 
| 1065 | proof | |
| 1066 | show \<open>P k\<close> if stable: \<open>\<And>k. k div 2 = k \<Longrightarrow> P k\<close> | |
| 1067 | and rec: \<open>\<And>k b. P k \<Longrightarrow> (of_bool b + 2 * k) div 2 = k \<Longrightarrow> P (of_bool b + 2 * k)\<close> | |
| 1068 | for P and k :: int | |
| 1069 | proof (induction k rule: int_bit_induct) | |
| 1070 | case zero | |
| 1071 | from stable [of 0] show ?case | |
| 1072 | by simp | |
| 1073 | next | |
| 1074 | case minus | |
| 1075 | from stable [of \<open>- 1\<close>] show ?case | |
| 1076 | by simp | |
| 1077 | next | |
| 1078 | case (even k) | |
| 1079 | with rec [of k False] show ?case | |
| 1080 | by (simp add: ac_simps) | |
| 1081 | next | |
| 1082 | case (odd k) | |
| 1083 | with rec [of k True] show ?case | |
| 1084 | by (simp add: ac_simps) | |
| 1085 | qed | |
| 71182 | 1086 | show \<open>(2::int) ^ m div 2 ^ n = of_bool ((2::int) ^ m \<noteq> 0 \<and> n \<le> m) * 2 ^ (m - n)\<close> | 
| 1087 | for m n :: nat | |
| 1088 | proof (cases \<open>m < n\<close>) | |
| 1089 | case True | |
| 1090 | then have \<open>n = m + (n - m)\<close> | |
| 1091 | by simp | |
| 1092 | then have \<open>(2::int) ^ m div 2 ^ n = (2::int) ^ m div 2 ^ (m + (n - m))\<close> | |
| 1093 | by simp | |
| 1094 | also have \<open>\<dots> = (2::int) ^ m div (2 ^ m * 2 ^ (n - m))\<close> | |
| 1095 | by (simp add: power_add) | |
| 1096 | also have \<open>\<dots> = (2::int) ^ m div 2 ^ m div 2 ^ (n - m)\<close> | |
| 1097 | by (simp add: zdiv_zmult2_eq) | |
| 1098 | finally show ?thesis using \<open>m < n\<close> by simp | |
| 1099 | next | |
| 1100 | case False | |
| 1101 | then show ?thesis | |
| 1102 | by (simp add: power_diff) | |
| 1103 | qed | |
| 71138 | 1104 | show \<open>k mod 2 ^ m mod 2 ^ n = k mod 2 ^ min m n\<close> | 
| 1105 | for m n :: nat and k :: int | |
| 1106 | using mod_exp_eq [of \<open>nat k\<close> m n] | |
| 1107 | apply (auto simp add: mod_mod_cancel zdiv_zmult2_eq power_add zmod_zmult2_eq le_iff_add split: split_min_lin) | |
| 1108 | apply (auto simp add: less_iff_Suc_add mod_mod_cancel power_add) | |
| 1109 | apply (simp only: flip: mult.left_commute [of \<open>2 ^ m\<close>]) | |
| 1110 | apply (subst zmod_zmult2_eq) apply simp_all | |
| 1111 | done | |
| 1112 | show \<open>(k * 2 ^ m) mod (2 ^ n) = (k mod 2 ^ (n - m)) * 2 ^ m\<close> | |
| 1113 | if \<open>m \<le> n\<close> for m n :: nat and k :: int | |
| 1114 | using that | |
| 1115 | apply (auto simp add: power_add zmod_zmult2_eq le_iff_add split: split_min_lin) | |
| 1116 | apply (simp add: ac_simps) | |
| 1117 | done | |
| 71413 | 1118 | show \<open>even ((2 ^ m - (1::int)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::int) \<or> m \<le> n\<close> | 
| 1119 | for m n :: nat | |
| 1120 | using even_mask_div_iff' [where ?'a = int, of m n] by simp | |
| 71424 | 1121 | show \<open>even (k * 2 ^ m div 2 ^ n) \<longleftrightarrow> n < m \<or> (2::int) ^ n = 0 \<or> m \<le> n \<and> even (k div 2 ^ (n - m))\<close> | 
| 1122 | for m n :: nat and k l :: int | |
| 1123 | apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex) | |
| 1124 | apply (metis Suc_leI dvd_mult dvd_mult_imp_div dvd_power_le dvd_refl power.simps(2)) | |
| 1125 | done | |
| 71182 | 1126 | qed (auto simp add: zdiv_zmult2_eq zmod_zmult2_eq power_add power_diff not_le) | 
| 67816 | 1127 | |
| 71094 | 1128 | class semiring_bit_shifts = semiring_bits + | 
| 1129 | fixes push_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | |
| 1130 | assumes push_bit_eq_mult: \<open>push_bit n a = a * 2 ^ n\<close> | |
| 1131 | fixes drop_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | |
| 1132 | assumes drop_bit_eq_div: \<open>drop_bit n a = a div 2 ^ n\<close> | |
| 67816 | 1133 | begin | 
| 1134 | ||
| 71094 | 1135 | definition take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | 
| 1136 | where take_bit_eq_mod: \<open>take_bit n a = a mod 2 ^ n\<close> | |
| 67816 | 1137 | |
| 71094 | 1138 | text \<open> | 
| 1139 | Logically, \<^const>\<open>push_bit\<close>, | |
| 1140 | \<^const>\<open>drop_bit\<close> and \<^const>\<open>take_bit\<close> are just aliases; having them | |
| 1141 | as separate operations makes proofs easier, otherwise proof automation | |
| 1142 | would fiddle with concrete expressions \<^term>\<open>2 ^ n\<close> in a way obfuscating the basic | |
| 1143 | algebraic relationships between those operations. | |
| 1144 | Having | |
| 1145 | \<^const>\<open>push_bit\<close> and \<^const>\<open>drop_bit\<close> as definitional class operations | |
| 1146 | takes into account that specific instances of these can be implemented | |
| 1147 | differently wrt. code generation. | |
| 1148 | \<close> | |
| 67816 | 1149 | |
| 71408 | 1150 | lemma bit_iff_odd_drop_bit: | 
| 1151 | \<open>bit a n \<longleftrightarrow> odd (drop_bit n a)\<close> | |
| 1152 | by (simp add: bit_def drop_bit_eq_div) | |
| 1153 | ||
| 1154 | lemma even_drop_bit_iff_not_bit: | |
| 1155 | \<open>even (drop_bit n a) \<longleftrightarrow> \<not> bit a n\<close> | |
| 1156 | by (simp add: bit_iff_odd_drop_bit) | |
| 1157 | ||
| 71423 | 1158 | lemma div_push_bit_of_1_eq_drop_bit: | 
| 1159 | \<open>a div push_bit n 1 = drop_bit n a\<close> | |
| 1160 | by (simp add: push_bit_eq_mult drop_bit_eq_div) | |
| 1161 | ||
| 71195 | 1162 | lemma bits_ident: | 
| 71138 | 1163 | "push_bit n (drop_bit n a) + take_bit n a = a" | 
| 1164 | using div_mult_mod_eq by (simp add: push_bit_eq_mult take_bit_eq_mod drop_bit_eq_div) | |
| 1165 | ||
| 1166 | lemma push_bit_push_bit [simp]: | |
| 1167 | "push_bit m (push_bit n a) = push_bit (m + n) a" | |
| 1168 | by (simp add: push_bit_eq_mult power_add ac_simps) | |
| 1169 | ||
| 1170 | lemma push_bit_0_id [simp]: | |
| 1171 | "push_bit 0 = id" | |
| 1172 | by (simp add: fun_eq_iff push_bit_eq_mult) | |
| 1173 | ||
| 1174 | lemma push_bit_of_0 [simp]: | |
| 1175 | "push_bit n 0 = 0" | |
| 1176 | by (simp add: push_bit_eq_mult) | |
| 1177 | ||
| 1178 | lemma push_bit_of_1: | |
| 1179 | "push_bit n 1 = 2 ^ n" | |
| 1180 | by (simp add: push_bit_eq_mult) | |
| 1181 | ||
| 1182 | lemma push_bit_Suc [simp]: | |
| 1183 | "push_bit (Suc n) a = push_bit n (a * 2)" | |
| 1184 | by (simp add: push_bit_eq_mult ac_simps) | |
| 1185 | ||
| 1186 | lemma push_bit_double: | |
| 1187 | "push_bit n (a * 2) = push_bit n a * 2" | |
| 1188 | by (simp add: push_bit_eq_mult ac_simps) | |
| 1189 | ||
| 1190 | lemma push_bit_add: | |
| 1191 | "push_bit n (a + b) = push_bit n a + push_bit n b" | |
| 1192 | by (simp add: push_bit_eq_mult algebra_simps) | |
| 1193 | ||
| 1194 | lemma take_bit_0 [simp]: | |
| 1195 | "take_bit 0 a = 0" | |
| 1196 | by (simp add: take_bit_eq_mod) | |
| 1197 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1198 | lemma take_bit_Suc: | 
| 71138 | 1199 | \<open>take_bit (Suc n) a = take_bit n (a div 2) * 2 + of_bool (odd a)\<close> | 
| 1200 | proof - | |
| 1201 | have \<open>take_bit (Suc n) (a div 2 * 2 + of_bool (odd a)) = take_bit n (a div 2) * 2 + of_bool (odd a)\<close> | |
| 1202 | using even_succ_mod_exp [of \<open>2 * (a div 2)\<close> \<open>Suc n\<close>] | |
| 1203 | mult_exp_mod_exp_eq [of 1 \<open>Suc n\<close> \<open>a div 2\<close>] | |
| 1204 | by (auto simp add: take_bit_eq_mod ac_simps) | |
| 1205 | then show ?thesis | |
| 1206 | using div_mult_mod_eq [of a 2] by (simp add: mod_2_eq_odd) | |
| 1207 | qed | |
| 1208 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1209 | lemma take_bit_rec: | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1210 | \<open>take_bit n a = (if n = 0 then 0 else take_bit (n - 1) (a div 2) * 2 + of_bool (odd a))\<close> | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1211 | by (cases n) (simp_all add: take_bit_Suc) | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1212 | |
| 71138 | 1213 | lemma take_bit_of_0 [simp]: | 
| 1214 | "take_bit n 0 = 0" | |
| 1215 | by (simp add: take_bit_eq_mod) | |
| 1216 | ||
| 1217 | lemma take_bit_of_1 [simp]: | |
| 1218 | "take_bit n 1 = of_bool (n > 0)" | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1219 | by (cases n) (simp_all add: take_bit_Suc) | 
| 71138 | 1220 | |
| 1221 | lemma drop_bit_of_0 [simp]: | |
| 1222 | "drop_bit n 0 = 0" | |
| 1223 | by (simp add: drop_bit_eq_div) | |
| 1224 | ||
| 1225 | lemma drop_bit_of_1 [simp]: | |
| 1226 | "drop_bit n 1 = of_bool (n = 0)" | |
| 1227 | by (simp add: drop_bit_eq_div) | |
| 1228 | ||
| 1229 | lemma drop_bit_0 [simp]: | |
| 1230 | "drop_bit 0 = id" | |
| 1231 | by (simp add: fun_eq_iff drop_bit_eq_div) | |
| 1232 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1233 | lemma drop_bit_Suc: | 
| 71138 | 1234 | "drop_bit (Suc n) a = drop_bit n (a div 2)" | 
| 1235 | using div_exp_eq [of a 1] by (simp add: drop_bit_eq_div) | |
| 1236 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1237 | lemma drop_bit_rec: | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1238 | "drop_bit n a = (if n = 0 then a else drop_bit (n - 1) (a div 2))" | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1239 | by (cases n) (simp_all add: drop_bit_Suc) | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1240 | |
| 71138 | 1241 | lemma drop_bit_half: | 
| 1242 | "drop_bit n (a div 2) = drop_bit n a div 2" | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1243 | by (induction n arbitrary: a) (simp_all add: drop_bit_Suc) | 
| 71138 | 1244 | |
| 1245 | lemma drop_bit_of_bool [simp]: | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1246 | "drop_bit n (of_bool b) = of_bool (n = 0 \<and> b)" | 
| 71138 | 1247 | by (cases n) simp_all | 
| 1248 | ||
| 1249 | lemma take_bit_eq_0_imp_dvd: | |
| 1250 | "take_bit n a = 0 \<Longrightarrow> 2 ^ n dvd a" | |
| 1251 | by (simp add: take_bit_eq_mod mod_0_imp_dvd) | |
| 1252 | ||
| 1253 | lemma even_take_bit_eq [simp]: | |
| 1254 | \<open>even (take_bit n a) \<longleftrightarrow> n = 0 \<or> even a\<close> | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1255 | by (simp add: take_bit_rec [of n a]) | 
| 71138 | 1256 | |
| 1257 | lemma take_bit_take_bit [simp]: | |
| 1258 | "take_bit m (take_bit n a) = take_bit (min m n) a" | |
| 1259 | by (simp add: take_bit_eq_mod mod_exp_eq ac_simps) | |
| 1260 | ||
| 1261 | lemma drop_bit_drop_bit [simp]: | |
| 1262 | "drop_bit m (drop_bit n a) = drop_bit (m + n) a" | |
| 1263 | by (simp add: drop_bit_eq_div power_add div_exp_eq ac_simps) | |
| 1264 | ||
| 1265 | lemma push_bit_take_bit: | |
| 1266 | "push_bit m (take_bit n a) = take_bit (m + n) (push_bit m a)" | |
| 1267 | apply (simp add: push_bit_eq_mult take_bit_eq_mod power_add ac_simps) | |
| 1268 | using mult_exp_mod_exp_eq [of m \<open>m + n\<close> a] apply (simp add: ac_simps power_add) | |
| 1269 | done | |
| 1270 | ||
| 1271 | lemma take_bit_push_bit: | |
| 1272 | "take_bit m (push_bit n a) = push_bit n (take_bit (m - n) a)" | |
| 1273 | proof (cases "m \<le> n") | |
| 1274 | case True | |
| 1275 | then show ?thesis | |
| 1276 | apply (simp add:) | |
| 1277 | apply (simp_all add: push_bit_eq_mult take_bit_eq_mod) | |
| 1278 | apply (auto dest!: le_Suc_ex simp add: power_add ac_simps) | |
| 1279 | using mult_exp_mod_exp_eq [of m m \<open>a * 2 ^ n\<close> for n] | |
| 1280 | apply (simp add: ac_simps) | |
| 1281 | done | |
| 1282 | next | |
| 1283 | case False | |
| 1284 | then show ?thesis | |
| 1285 | using push_bit_take_bit [of n "m - n" a] | |
| 1286 | by simp | |
| 1287 | qed | |
| 1288 | ||
| 1289 | lemma take_bit_drop_bit: | |
| 1290 | "take_bit m (drop_bit n a) = drop_bit n (take_bit (m + n) a)" | |
| 1291 | by (simp add: drop_bit_eq_div take_bit_eq_mod ac_simps div_exp_mod_exp_eq) | |
| 1292 | ||
| 1293 | lemma drop_bit_take_bit: | |
| 1294 | "drop_bit m (take_bit n a) = take_bit (n - m) (drop_bit m a)" | |
| 1295 | proof (cases "m \<le> n") | |
| 1296 | case True | |
| 1297 | then show ?thesis | |
| 1298 | using take_bit_drop_bit [of "n - m" m a] by simp | |
| 1299 | next | |
| 1300 | case False | |
| 1301 | then obtain q where \<open>m = n + q\<close> | |
| 1302 | by (auto simp add: not_le dest: less_imp_Suc_add) | |
| 1303 | then have \<open>drop_bit m (take_bit n a) = 0\<close> | |
| 1304 | using div_exp_eq [of \<open>a mod 2 ^ n\<close> n q] | |
| 1305 | by (simp add: take_bit_eq_mod drop_bit_eq_div) | |
| 1306 | with False show ?thesis | |
| 1307 | by simp | |
| 1308 | qed | |
| 1309 | ||
| 71424 | 1310 | lemma even_push_bit_iff [simp]: | 
| 1311 | \<open>even (push_bit n a) \<longleftrightarrow> n \<noteq> 0 \<or> even a\<close> | |
| 1312 | by (simp add: push_bit_eq_mult) auto | |
| 1313 | ||
| 1314 | lemma bit_push_bit_iff: | |
| 1315 | \<open>bit (push_bit m a) n \<longleftrightarrow> n \<ge> m \<and> 2 ^ n \<noteq> 0 \<and> (n < m \<or> bit a (n - m))\<close> | |
| 1316 | by (auto simp add: bit_def push_bit_eq_mult even_mult_exp_div_exp_iff) | |
| 1317 | ||
| 71181 | 1318 | lemma bit_drop_bit_eq: | 
| 1319 | \<open>bit (drop_bit n a) = bit a \<circ> (+) n\<close> | |
| 1320 | by (simp add: bit_def fun_eq_iff ac_simps flip: drop_bit_eq_div) | |
| 1321 | ||
| 1322 | lemma bit_take_bit_iff: | |
| 1323 | \<open>bit (take_bit m a) n \<longleftrightarrow> n < m \<and> bit a n\<close> | |
| 1324 | by (simp add: bit_def drop_bit_take_bit not_le flip: drop_bit_eq_div) | |
| 1325 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1326 | lemma stable_imp_drop_bit_eq: | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1327 | \<open>drop_bit n a = a\<close> | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1328 | if \<open>a div 2 = a\<close> | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1329 | by (induction n) (simp_all add: that drop_bit_Suc) | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1330 | |
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1331 | lemma stable_imp_take_bit_eq: | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1332 | \<open>take_bit n a = (if even a then 0 else 2 ^ n - 1)\<close> | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1333 | if \<open>a div 2 = a\<close> | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1334 | proof (rule bit_eqI) | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1335 | fix m | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1336 | assume \<open>2 ^ m \<noteq> 0\<close> | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1337 | with that show \<open>bit (take_bit n a) m \<longleftrightarrow> bit (if even a then 0 else 2 ^ n - 1) m\<close> | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1338 | by (simp add: bit_take_bit_iff bit_mask_iff stable_imp_bit_iff_odd) | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1339 | qed | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1340 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 1341 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 1342 | |
| 71094 | 1343 | instantiation nat :: semiring_bit_shifts | 
| 1344 | begin | |
| 1345 | ||
| 1346 | definition push_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | |
| 1347 | where \<open>push_bit_nat n m = m * 2 ^ n\<close> | |
| 1348 | ||
| 1349 | definition drop_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | |
| 1350 | where \<open>drop_bit_nat n m = m div 2 ^ n\<close> | |
| 1351 | ||
| 1352 | instance proof | |
| 1353 | show \<open>push_bit n m = m * 2 ^ n\<close> for n m :: nat | |
| 1354 | by (simp add: push_bit_nat_def) | |
| 1355 | show \<open>drop_bit n m = m div 2 ^ n\<close> for n m :: nat | |
| 1356 | by (simp add: drop_bit_nat_def) | |
| 1357 | qed | |
| 1358 | ||
| 1359 | end | |
| 1360 | ||
| 1361 | instantiation int :: semiring_bit_shifts | |
| 1362 | begin | |
| 1363 | ||
| 1364 | definition push_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | |
| 1365 | where \<open>push_bit_int n k = k * 2 ^ n\<close> | |
| 1366 | ||
| 1367 | definition drop_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | |
| 1368 | where \<open>drop_bit_int n k = k div 2 ^ n\<close> | |
| 1369 | ||
| 1370 | instance proof | |
| 1371 | show \<open>push_bit n k = k * 2 ^ n\<close> for n :: nat and k :: int | |
| 1372 | by (simp add: push_bit_int_def) | |
| 1373 | show \<open>drop_bit n k = k div 2 ^ n\<close> for n :: nat and k :: int | |
| 1374 | by (simp add: drop_bit_int_def) | |
| 1375 | qed | |
| 1376 | ||
| 1377 | end | |
| 1378 | ||
| 71412 | 1379 | lemma bit_push_bit_iff_nat: | 
| 1380 | \<open>bit (push_bit m q) n \<longleftrightarrow> m \<le> n \<and> bit q (n - m)\<close> for q :: nat | |
| 71424 | 1381 | by (auto simp add: bit_push_bit_iff) | 
| 71412 | 1382 | |
| 1383 | lemma bit_push_bit_iff_int: | |
| 1384 | \<open>bit (push_bit m k) n \<longleftrightarrow> m \<le> n \<and> bit k (n - m)\<close> for k :: int | |
| 71424 | 1385 | by (auto simp add: bit_push_bit_iff) | 
| 71412 | 1386 | |
| 71094 | 1387 | class unique_euclidean_semiring_with_bit_shifts = | 
| 1388 | unique_euclidean_semiring_with_nat + semiring_bit_shifts | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 1389 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 1390 | |
| 71138 | 1391 | lemma take_bit_of_exp [simp]: | 
| 1392 | \<open>take_bit m (2 ^ n) = of_bool (n < m) * 2 ^ n\<close> | |
| 1393 | by (simp add: take_bit_eq_mod exp_mod_exp) | |
| 67960 | 1394 | |
| 71138 | 1395 | lemma take_bit_of_2 [simp]: | 
| 1396 | \<open>take_bit n 2 = of_bool (2 \<le> n) * 2\<close> | |
| 1397 | using take_bit_of_exp [of n 1] by simp | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1398 | |
| 71412 | 1399 | lemma take_bit_of_mask: | 
| 71408 | 1400 | \<open>take_bit m (2 ^ n - 1) = 2 ^ min m n - 1\<close> | 
| 71412 | 1401 | by (simp add: take_bit_eq_mod mask_mod_exp) | 
| 71408 | 1402 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1403 | lemma push_bit_eq_0_iff [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1404 | "push_bit n a = 0 \<longleftrightarrow> a = 0" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1405 | by (simp add: push_bit_eq_mult) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1406 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1407 | lemma push_bit_numeral [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1408 | "push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1409 | by (simp only: numeral_eq_Suc power_Suc numeral_Bit0 [of k] mult_2 [symmetric]) (simp add: ac_simps) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1410 | |
| 68010 | 1411 | lemma push_bit_of_nat: | 
| 1412 | "push_bit n (of_nat m) = of_nat (push_bit n m)" | |
| 1413 | by (simp add: push_bit_eq_mult Parity.push_bit_eq_mult) | |
| 1414 | ||
| 67961 | 1415 | lemma take_bit_add: | 
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1416 | "take_bit n (take_bit n a + take_bit n b) = take_bit n (a + b)" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1417 | by (simp add: take_bit_eq_mod mod_simps) | 
| 67816 | 1418 | |
| 67961 | 1419 | lemma take_bit_eq_0_iff: | 
| 1420 | "take_bit n a = 0 \<longleftrightarrow> 2 ^ n dvd a" | |
| 1421 | by (simp add: take_bit_eq_mod mod_eq_0_iff_dvd) | |
| 1422 | ||
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1423 | lemma take_bit_of_1_eq_0_iff [simp]: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1424 | "take_bit n 1 = 0 \<longleftrightarrow> n = 0" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1425 | by (simp add: take_bit_eq_mod) | 
| 67816 | 1426 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1427 | lemma take_bit_numeral_bit0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1428 | "take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1429 | by (simp only: numeral_eq_Suc power_Suc numeral_Bit0 [of k] mult_2 [symmetric] take_bit_Suc | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1430 | ac_simps even_mult_iff nonzero_mult_div_cancel_right [OF numeral_neq_zero]) simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1431 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1432 | lemma take_bit_numeral_bit1 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1433 | "take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1434 | by (simp only: numeral_eq_Suc power_Suc numeral_Bit1 [of k] mult_2 [symmetric] take_bit_Suc | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1435 | ac_simps even_add even_mult_iff div_mult_self1 [OF numeral_neq_zero]) (simp add: ac_simps) | 
| 67961 | 1436 | |
| 68010 | 1437 | lemma take_bit_of_nat: | 
| 1438 | "take_bit n (of_nat m) = of_nat (take_bit n m)" | |
| 1439 | by (simp add: take_bit_eq_mod Parity.take_bit_eq_mod of_nat_mod [of m "2 ^ n"]) | |
| 1440 | ||
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1441 | lemma drop_bit_numeral_bit0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1442 | "drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1443 | by (simp only: numeral_eq_Suc power_Suc numeral_Bit0 [of k] mult_2 [symmetric] drop_bit_Suc | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1444 | nonzero_mult_div_cancel_left [OF numeral_neq_zero]) | 
| 67816 | 1445 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1446 | lemma drop_bit_numeral_bit1 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1447 | "drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1448 | by (simp only: numeral_eq_Suc power_Suc numeral_Bit1 [of k] mult_2 [symmetric] drop_bit_Suc | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1449 | div_mult_self4 [OF numeral_neq_zero]) simp | 
| 67816 | 1450 | |
| 68010 | 1451 | lemma drop_bit_of_nat: | 
| 1452 | "drop_bit n (of_nat m) = of_nat (drop_bit n m)" | |
| 68389 | 1453 | by (simp add: drop_bit_eq_div Parity.drop_bit_eq_div of_nat_div [of m "2 ^ n"]) | 
| 68010 | 1454 | |
| 71412 | 1455 | lemma bit_of_nat_iff_bit [simp]: | 
| 1456 | \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close> | |
| 1457 | proof - | |
| 1458 | have \<open>even (m div 2 ^ n) \<longleftrightarrow> even (of_nat (m div 2 ^ n))\<close> | |
| 1459 | by simp | |
| 1460 | also have \<open>of_nat (m div 2 ^ n) = of_nat m div of_nat (2 ^ n)\<close> | |
| 1461 | by (simp add: of_nat_div) | |
| 1462 | finally show ?thesis | |
| 1463 | by (simp add: bit_def semiring_bits_class.bit_def) | |
| 1464 | qed | |
| 1465 | ||
| 1466 | lemma of_nat_push_bit: | |
| 1467 | \<open>of_nat (push_bit m n) = push_bit m (of_nat n)\<close> | |
| 1468 | by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult) | |
| 1469 | ||
| 1470 | lemma of_nat_drop_bit: | |
| 1471 | \<open>of_nat (drop_bit m n) = drop_bit m (of_nat n)\<close> | |
| 1472 | by (simp add: drop_bit_eq_div semiring_bit_shifts_class.drop_bit_eq_div of_nat_div) | |
| 1473 | ||
| 1474 | lemma of_nat_take_bit: | |
| 1475 | \<open>of_nat (take_bit m n) = take_bit m (of_nat n)\<close> | |
| 1476 | by (simp add: take_bit_eq_mod semiring_bit_shifts_class.take_bit_eq_mod of_nat_mod) | |
| 1477 | ||
| 1478 | lemma bit_push_bit_iff_of_nat_iff: | |
| 1479 | \<open>bit (push_bit m (of_nat r)) n \<longleftrightarrow> m \<le> n \<and> bit (of_nat r) (n - m)\<close> | |
| 71424 | 1480 | by (auto simp add: bit_push_bit_iff) | 
| 71412 | 1481 | |
| 58770 | 1482 | end | 
| 67816 | 1483 | |
| 71094 | 1484 | instance nat :: unique_euclidean_semiring_with_bit_shifts .. | 
| 1485 | ||
| 1486 | instance int :: unique_euclidean_semiring_with_bit_shifts .. | |
| 1487 | ||
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1488 | lemma push_bit_of_Suc_0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1489 | "push_bit n (Suc 0) = 2 ^ n" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1490 | using push_bit_of_1 [where ?'a = nat] by simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1491 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1492 | lemma take_bit_of_Suc_0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1493 | "take_bit n (Suc 0) = of_bool (0 < n)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1494 | using take_bit_of_1 [where ?'a = nat] by simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1495 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1496 | lemma drop_bit_of_Suc_0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1497 | "drop_bit n (Suc 0) = of_bool (n = 0)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1498 | using drop_bit_of_1 [where ?'a = nat] by simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1499 | |
| 70973 | 1500 | lemma take_bit_eq_self: | 
| 1501 | \<open>take_bit n m = m\<close> if \<open>m < 2 ^ n\<close> for n m :: nat | |
| 1502 | using that by (simp add: take_bit_eq_mod) | |
| 1503 | ||
| 70911 | 1504 | lemma push_bit_minus_one: | 
| 1505 | "push_bit n (- 1 :: int) = - (2 ^ n)" | |
| 1506 | by (simp add: push_bit_eq_mult) | |
| 1507 | ||
| 71195 | 1508 | lemma minus_1_div_exp_eq_int: | 
| 1509 | \<open>- 1 div (2 :: int) ^ n = - 1\<close> | |
| 1510 | by (induction n) (use div_exp_eq [symmetric, of \<open>- 1 :: int\<close> 1] in \<open>simp_all add: ac_simps\<close>) | |
| 1511 | ||
| 1512 | lemma drop_bit_minus_one [simp]: | |
| 1513 | \<open>drop_bit n (- 1 :: int) = - 1\<close> | |
| 1514 | by (simp add: drop_bit_eq_div minus_1_div_exp_eq_int) | |
| 1515 | ||
| 71424 | 1516 | lemma take_bit_minus: | 
| 71195 | 1517 | "take_bit n (- (take_bit n k)) = take_bit n (- k)" | 
| 1518 | for k :: int | |
| 1519 | by (simp add: take_bit_eq_mod mod_minus_eq) | |
| 1520 | ||
| 71424 | 1521 | lemma take_bit_diff: | 
| 71195 | 1522 | "take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)" | 
| 1523 | for k l :: int | |
| 1524 | by (simp add: take_bit_eq_mod mod_diff_eq) | |
| 1525 | ||
| 1526 | lemma take_bit_nonnegative [simp]: | |
| 1527 | "take_bit n k \<ge> 0" | |
| 1528 | for k :: int | |
| 1529 | by (simp add: take_bit_eq_mod) | |
| 1530 | ||
| 71424 | 1531 | lemma drop_bit_push_bit_int: | 
| 1532 | \<open>drop_bit m (push_bit n k) = drop_bit (m - n) (push_bit (n - m) k)\<close> for k :: int | |
| 1533 | by (cases \<open>m \<le> n\<close>) (auto simp add: mult.left_commute [of _ \<open>2 ^ n\<close>] mult.commute [of _ \<open>2 ^ n\<close>] mult.assoc | |
| 1534 | mult.commute [of k] drop_bit_eq_div push_bit_eq_mult not_le power_add dest!: le_Suc_ex less_imp_Suc_add) | |
| 1535 | ||
| 67816 | 1536 | end |