| author | wenzelm | 
| Sat, 10 Oct 2020 21:12:20 +0200 | |
| changeset 72425 | d0937d55eb90 | 
| parent 72281 | beeadb35e357 | 
| child 72512 | 83b5911c0164 | 
| 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 | |
| 71755 | 155 | lemma odd_numeral_BitM [simp]: | 
| 156 | \<open>odd (numeral (Num.BitM w))\<close> | |
| 157 | by (cases w) simp_all | |
| 158 | ||
| 63654 | 159 | lemma even_power [simp]: "even (a ^ n) \<longleftrightarrow> even a \<and> n > 0" | 
| 58680 | 160 | by (induct n) auto | 
| 161 | ||
| 71412 | 162 | lemma mask_eq_sum_exp: | 
| 163 |   \<open>2 ^ n - 1 = (\<Sum>m\<in>{q. q < n}. 2 ^ m)\<close>
 | |
| 164 | proof - | |
| 165 |   have *: \<open>{q. q < Suc m} = insert m {q. q < m}\<close> for m
 | |
| 166 | by auto | |
| 167 |   have \<open>2 ^ n = (\<Sum>m\<in>{q. q < n}. 2 ^ m) + 1\<close>
 | |
| 168 | by (induction n) (simp_all add: ac_simps mult_2 *) | |
| 169 |   then have \<open>2 ^ n - 1 = (\<Sum>m\<in>{q. q < n}. 2 ^ m) + 1 - 1\<close>
 | |
| 170 | by simp | |
| 171 | then show ?thesis | |
| 172 | by simp | |
| 173 | qed | |
| 174 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 175 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 176 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 177 | 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 | 178 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 179 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 180 | subclass comm_ring_1 .. | 
| 
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 | lemma even_minus: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 183 | "even (- a) \<longleftrightarrow> even a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 184 | by (fact dvd_minus_iff) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 185 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 186 | lemma even_diff [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 187 | "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 | 188 | 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 | 189 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 190 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 191 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 192 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 193 | 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 | 194 | |
| 71157 | 195 | 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 | 196 | begin | 
| 
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 | subclass semiring_parity | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 199 | proof | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 200 | 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 | 201 | 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 | 202 | 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 | 203 | proof | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 204 | assume "a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 205 | 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 | 206 | by auto | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 207 | next | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 208 | assume "\<not> 2 dvd a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 209 | 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 | 210 | proof (rule order_antisym) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 211 | 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 | 212 | 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 | 213 | 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 | 214 | 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 | 215 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 216 | 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 | 217 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 218 | 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 | 219 | 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 | 220 | 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 | 221 | 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 | 222 | 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 | 223 | 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 | 224 | 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 | 225 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 226 | 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 | 227 | 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 | 228 | 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 | 229 | show "a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 230 | 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 | 231 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 232 | show "\<not> is_unit 2" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 233 | proof (rule notI) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 234 | assume "is_unit 2" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 235 | 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 | 236 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 237 | 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 | 238 | 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 | 239 | then show False | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 240 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 241 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 242 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 243 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 244 | lemma even_of_nat [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 245 | "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 | 246 | proof - | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 247 | 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 | 248 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 249 | 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 | 250 | 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 | 251 | finally show ?thesis . | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 252 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 253 | |
| 66815 | 254 | lemma even_succ_div_two [simp]: | 
| 255 | "even a \<Longrightarrow> (a + 1) div 2 = a div 2" | |
| 256 | by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero) | |
| 257 | ||
| 258 | lemma odd_succ_div_two [simp]: | |
| 259 | "odd a \<Longrightarrow> (a + 1) div 2 = a div 2 + 1" | |
| 260 | by (auto elim!: oddE simp add: add.assoc) | |
| 261 | ||
| 262 | lemma even_two_times_div_two: | |
| 263 | "even a \<Longrightarrow> 2 * (a div 2) = a" | |
| 264 | by (fact dvd_mult_div_cancel) | |
| 265 | ||
| 266 | lemma odd_two_times_div_two_succ [simp]: | |
| 267 | "odd a \<Longrightarrow> 2 * (a div 2) + 1 = a" | |
| 268 | using mult_div_mod_eq [of 2 a] | |
| 269 | by (simp add: even_iff_mod_2_eq_zero) | |
| 270 | ||
| 67051 | 271 | lemma coprime_left_2_iff_odd [simp]: | 
| 272 | "coprime 2 a \<longleftrightarrow> odd a" | |
| 273 | proof | |
| 274 | assume "odd a" | |
| 275 | show "coprime 2 a" | |
| 276 | proof (rule coprimeI) | |
| 277 | fix b | |
| 278 | assume "b dvd 2" "b dvd a" | |
| 279 | then have "b dvd a mod 2" | |
| 280 | by (auto intro: dvd_mod) | |
| 281 | with \<open>odd a\<close> show "is_unit b" | |
| 282 | by (simp add: mod_2_eq_odd) | |
| 283 | qed | |
| 284 | next | |
| 285 | assume "coprime 2 a" | |
| 286 | show "odd a" | |
| 287 | proof (rule notI) | |
| 288 | assume "even a" | |
| 289 | then obtain b where "a = 2 * b" .. | |
| 290 | with \<open>coprime 2 a\<close> have "coprime 2 (2 * b)" | |
| 291 | by simp | |
| 292 | moreover have "\<not> coprime 2 (2 * b)" | |
| 293 | by (rule not_coprimeI [of 2]) simp_all | |
| 294 | ultimately show False | |
| 295 | by blast | |
| 296 | qed | |
| 297 | qed | |
| 298 | ||
| 299 | lemma coprime_right_2_iff_odd [simp]: | |
| 300 | "coprime a 2 \<longleftrightarrow> odd a" | |
| 301 | using coprime_left_2_iff_odd [of a] by (simp add: ac_simps) | |
| 302 | ||
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 303 | end | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 304 | |
| 71157 | 305 | context unique_euclidean_ring_with_nat | 
| 58679 | 306 | begin | 
| 307 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 308 | subclass ring_parity .. | 
| 58680 | 309 | |
| 67906 | 310 | lemma minus_1_mod_2_eq [simp]: | 
| 311 | "- 1 mod 2 = 1" | |
| 312 | by (simp add: mod_2_eq_odd) | |
| 313 | ||
| 314 | lemma minus_1_div_2_eq [simp]: | |
| 315 | "- 1 div 2 = - 1" | |
| 316 | proof - | |
| 317 | from div_mult_mod_eq [of "- 1" 2] | |
| 318 | 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 | 319 | using add_implies_diff by fastforce | 
| 67906 | 320 | then show ?thesis | 
| 321 | using mult_right_cancel [of 2 "- 1 div 2" "- 1"] by simp | |
| 322 | qed | |
| 323 | ||
| 58679 | 324 | end | 
| 325 | ||
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 326 | |
| 69593 | 327 | subsection \<open>Instance for \<^typ>\<open>nat\<close>\<close> | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 328 | |
| 70340 | 329 | instance nat :: unique_euclidean_semiring_with_nat | 
| 66815 | 330 | by standard (simp_all add: dvd_eq_mod_eq_0) | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 331 | |
| 66815 | 332 | lemma even_Suc_Suc_iff [simp]: | 
| 333 | "even (Suc (Suc n)) \<longleftrightarrow> even n" | |
| 58787 | 334 | using dvd_add_triv_right_iff [of 2 n] by simp | 
| 58687 | 335 | |
| 66815 | 336 | lemma even_Suc [simp]: "even (Suc n) \<longleftrightarrow> odd n" | 
| 337 | using even_plus_one_iff [of n] by simp | |
| 58787 | 338 | |
| 66815 | 339 | lemma even_diff_nat [simp]: | 
| 340 | "even (m - n) \<longleftrightarrow> m < n \<or> even (m + n)" for m n :: nat | |
| 58787 | 341 | proof (cases "n \<le> m") | 
| 342 | case True | |
| 343 | then have "m - n + n * 2 = m + n" by (simp add: mult_2_right) | |
| 66815 | 344 | moreover have "even (m - n) \<longleftrightarrow> even (m - n + n * 2)" by simp | 
| 345 | ultimately have "even (m - n) \<longleftrightarrow> even (m + n)" by (simp only:) | |
| 58787 | 346 | then show ?thesis by auto | 
| 347 | next | |
| 348 | case False | |
| 349 | then show ?thesis by simp | |
| 63654 | 350 | qed | 
| 351 | ||
| 66815 | 352 | lemma odd_pos: | 
| 353 | "odd n \<Longrightarrow> 0 < n" for n :: nat | |
| 58690 | 354 | by (auto elim: oddE) | 
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 355 | |
| 66815 | 356 | lemma Suc_double_not_eq_double: | 
| 357 | "Suc (2 * m) \<noteq> 2 * n" | |
| 62597 | 358 | proof | 
| 359 | assume "Suc (2 * m) = 2 * n" | |
| 360 | moreover have "odd (Suc (2 * m))" and "even (2 * n)" | |
| 361 | by simp_all | |
| 362 | ultimately show False by simp | |
| 363 | qed | |
| 364 | ||
| 66815 | 365 | lemma double_not_eq_Suc_double: | 
| 366 | "2 * m \<noteq> Suc (2 * n)" | |
| 62597 | 367 | using Suc_double_not_eq_double [of n m] by simp | 
| 368 | ||
| 66815 | 369 | lemma odd_Suc_minus_one [simp]: "odd n \<Longrightarrow> Suc (n - Suc 0) = n" | 
| 370 | by (auto elim: oddE) | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 371 | |
| 66815 | 372 | lemma even_Suc_div_two [simp]: | 
| 373 | "even n \<Longrightarrow> Suc n div 2 = n div 2" | |
| 374 | using even_succ_div_two [of n] by simp | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 375 | |
| 66815 | 376 | lemma odd_Suc_div_two [simp]: | 
| 377 | "odd n \<Longrightarrow> Suc n div 2 = Suc (n div 2)" | |
| 378 | using odd_succ_div_two [of n] by simp | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 379 | |
| 66815 | 380 | lemma odd_two_times_div_two_nat [simp]: | 
| 381 | assumes "odd n" | |
| 382 | shows "2 * (n div 2) = n - (1 :: nat)" | |
| 383 | proof - | |
| 384 | from assms have "2 * (n div 2) + 1 = n" | |
| 385 | by (rule odd_two_times_div_two_succ) | |
| 386 | then have "Suc (2 * (n div 2)) - 1 = n - 1" | |
| 58787 | 387 | by simp | 
| 66815 | 388 | then show ?thesis | 
| 389 | by simp | |
| 58787 | 390 | qed | 
| 58680 | 391 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 392 | 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 | 393 | "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 | 394 | 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 | 395 | |
| 69502 | 396 | lemma odd_card_imp_not_empty: | 
| 397 |   \<open>A \<noteq> {}\<close> if \<open>odd (card A)\<close>
 | |
| 398 | using that by auto | |
| 399 | ||
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 400 | 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 | 401 | 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 | 402 | shows "P n" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 403 | proof (induct n rule: less_induct) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 404 | case (less n) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 405 | show ?case | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 406 | proof (cases "n < Suc (Suc 0)") | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 407 | case True | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 408 | then show ?thesis | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 409 | 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 | 410 | next | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 411 | case False | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 412 | 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 | 413 | 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 | 414 | then have "k<n" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 415 | by simp | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 416 | 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 | 417 | by blast | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 418 | then show ?thesis | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 419 | by (simp add: k) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 420 | qed | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 421 | qed | 
| 58687 | 422 | |
| 71413 | 423 | lemma mask_eq_sum_exp_nat: | 
| 424 |   \<open>2 ^ n - Suc 0 = (\<Sum>m\<in>{q. q < n}. 2 ^ m)\<close>
 | |
| 425 | using mask_eq_sum_exp [where ?'a = nat] by simp | |
| 426 | ||
| 71412 | 427 | context semiring_parity | 
| 428 | begin | |
| 429 | ||
| 430 | lemma even_sum_iff: | |
| 431 |   \<open>even (sum f A) \<longleftrightarrow> even (card {a\<in>A. odd (f a)})\<close> if \<open>finite A\<close>
 | |
| 432 | using that proof (induction A) | |
| 433 | case empty | |
| 434 | then show ?case | |
| 435 | by simp | |
| 436 | next | |
| 437 | case (insert a A) | |
| 438 |   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>
 | |
| 439 | by auto | |
| 440 | ultimately show ?case | |
| 441 | by simp | |
| 442 | qed | |
| 443 | ||
| 444 | lemma even_prod_iff: | |
| 445 | \<open>even (prod f A) \<longleftrightarrow> (\<exists>a\<in>A. even (f a))\<close> if \<open>finite A\<close> | |
| 446 | using that by (induction A) simp_all | |
| 447 | ||
| 448 | lemma even_mask_iff [simp]: | |
| 449 | \<open>even (2 ^ n - 1) \<longleftrightarrow> n = 0\<close> | |
| 450 | proof (cases \<open>n = 0\<close>) | |
| 451 | case True | |
| 452 | then show ?thesis | |
| 453 | by simp | |
| 454 | next | |
| 455 | case False | |
| 456 |   then have \<open>{a. a = 0 \<and> a < n} = {0}\<close>
 | |
| 457 | by auto | |
| 458 | then show ?thesis | |
| 459 | by (auto simp add: mask_eq_sum_exp even_sum_iff) | |
| 460 | qed | |
| 461 | ||
| 462 | end | |
| 463 | ||
| 71138 | 464 | |
| 60758 | 465 | subsection \<open>Parity and powers\<close> | 
| 58689 | 466 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
60867diff
changeset | 467 | context ring_1 | 
| 58689 | 468 | begin | 
| 469 | ||
| 63654 | 470 | lemma power_minus_even [simp]: "even n \<Longrightarrow> (- a) ^ n = a ^ n" | 
| 58690 | 471 | by (auto elim: evenE) | 
| 58689 | 472 | |
| 63654 | 473 | lemma power_minus_odd [simp]: "odd n \<Longrightarrow> (- a) ^ n = - (a ^ n)" | 
| 58690 | 474 | by (auto elim: oddE) | 
| 475 | ||
| 66815 | 476 | lemma uminus_power_if: | 
| 477 | "(- a) ^ n = (if even n then a ^ n else - (a ^ n))" | |
| 478 | by auto | |
| 479 | ||
| 63654 | 480 | lemma neg_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1" | 
| 58690 | 481 | by simp | 
| 58689 | 482 | |
| 63654 | 483 | lemma neg_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1" | 
| 58690 | 484 | by simp | 
| 58689 | 485 | |
| 66582 | 486 | lemma neg_one_power_add_eq_neg_one_power_diff: "k \<le> n \<Longrightarrow> (- 1) ^ (n + k) = (- 1) ^ (n - k)" | 
| 487 | by (cases "even (n + k)") auto | |
| 488 | ||
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 489 | 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 | 490 | by (induct n) auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 491 | |
| 63654 | 492 | end | 
| 58689 | 493 | |
| 494 | context linordered_idom | |
| 495 | begin | |
| 496 | ||
| 63654 | 497 | lemma zero_le_even_power: "even n \<Longrightarrow> 0 \<le> a ^ n" | 
| 58690 | 498 | by (auto elim: evenE) | 
| 58689 | 499 | |
| 63654 | 500 | lemma zero_le_odd_power: "odd n \<Longrightarrow> 0 \<le> a ^ n \<longleftrightarrow> 0 \<le> a" | 
| 58689 | 501 | by (auto simp add: power_even_eq zero_le_mult_iff elim: oddE) | 
| 502 | ||
| 63654 | 503 | lemma zero_le_power_eq: "0 \<le> a ^ n \<longleftrightarrow> even n \<or> odd n \<and> 0 \<le> a" | 
| 58787 | 504 | by (auto simp add: zero_le_even_power zero_le_odd_power) | 
| 63654 | 505 | |
| 506 | 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 | 507 | proof - | 
| 508 | have [simp]: "0 = a ^ n \<longleftrightarrow> a = 0 \<and> n > 0" | |
| 58787 | 509 | unfolding power_eq_0_iff [of a n, symmetric] by blast | 
| 58689 | 510 | show ?thesis | 
| 63654 | 511 | unfolding less_le zero_le_power_eq by auto | 
| 58689 | 512 | qed | 
| 513 | ||
| 63654 | 514 | lemma power_less_zero_eq [simp]: "a ^ n < 0 \<longleftrightarrow> odd n \<and> a < 0" | 
| 58689 | 515 | unfolding not_le [symmetric] zero_le_power_eq by auto | 
| 516 | ||
| 63654 | 517 | 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)" | 
| 518 | unfolding not_less [symmetric] zero_less_power_eq by auto | |
| 519 | ||
| 520 | lemma power_even_abs: "even n \<Longrightarrow> \<bar>a\<bar> ^ n = a ^ n" | |
| 58689 | 521 | using power_abs [of a n] by (simp add: zero_le_even_power) | 
| 522 | ||
| 523 | lemma power_mono_even: | |
| 524 | assumes "even n" and "\<bar>a\<bar> \<le> \<bar>b\<bar>" | |
| 525 | shows "a ^ n \<le> b ^ n" | |
| 526 | proof - | |
| 527 | have "0 \<le> \<bar>a\<bar>" by auto | |
| 63654 | 528 | with \<open>\<bar>a\<bar> \<le> \<bar>b\<bar>\<close> have "\<bar>a\<bar> ^ n \<le> \<bar>b\<bar> ^ n" | 
| 529 | by (rule power_mono) | |
| 530 | with \<open>even n\<close> show ?thesis | |
| 531 | by (simp add: power_even_abs) | |
| 58689 | 532 | qed | 
| 533 | ||
| 534 | lemma power_mono_odd: | |
| 535 | assumes "odd n" and "a \<le> b" | |
| 536 | shows "a ^ n \<le> b ^ n" | |
| 537 | proof (cases "b < 0") | |
| 63654 | 538 | case True | 
| 539 | with \<open>a \<le> b\<close> have "- b \<le> - a" and "0 \<le> - b" by auto | |
| 540 | then have "(- b) ^ n \<le> (- a) ^ n" by (rule power_mono) | |
| 60758 | 541 | with \<open>odd n\<close> show ?thesis by simp | 
| 58689 | 542 | next | 
| 63654 | 543 | case False | 
| 544 | then have "0 \<le> b" by auto | |
| 58689 | 545 | show ?thesis | 
| 546 | proof (cases "a < 0") | |
| 63654 | 547 | case True | 
| 548 | then have "n \<noteq> 0" and "a \<le> 0" using \<open>odd n\<close> [THEN odd_pos] by auto | |
| 60758 | 549 | then have "a ^ n \<le> 0" unfolding power_le_zero_eq using \<open>odd n\<close> by auto | 
| 63654 | 550 | moreover from \<open>0 \<le> b\<close> have "0 \<le> b ^ n" by auto | 
| 58689 | 551 | ultimately show ?thesis by auto | 
| 552 | next | |
| 63654 | 553 | case False | 
| 554 | then have "0 \<le> a" by auto | |
| 555 | with \<open>a \<le> b\<close> show ?thesis | |
| 556 | using power_mono by auto | |
| 58689 | 557 | qed | 
| 558 | qed | |
| 62083 | 559 | |
| 60758 | 560 | text \<open>Simplify, when the exponent is a numeral\<close> | 
| 58689 | 561 | |
| 562 | lemma zero_le_power_eq_numeral [simp]: | |
| 563 | "0 \<le> a ^ numeral w \<longleftrightarrow> even (numeral w :: nat) \<or> odd (numeral w :: nat) \<and> 0 \<le> a" | |
| 564 | by (fact zero_le_power_eq) | |
| 565 | ||
| 566 | lemma zero_less_power_eq_numeral [simp]: | |
| 63654 | 567 | "0 < a ^ numeral w \<longleftrightarrow> | 
| 568 | numeral w = (0 :: nat) \<or> | |
| 569 | even (numeral w :: nat) \<and> a \<noteq> 0 \<or> | |
| 570 | odd (numeral w :: nat) \<and> 0 < a" | |
| 58689 | 571 | by (fact zero_less_power_eq) | 
| 572 | ||
| 573 | lemma power_le_zero_eq_numeral [simp]: | |
| 63654 | 574 | "a ^ numeral w \<le> 0 \<longleftrightarrow> | 
| 575 | (0 :: nat) < numeral w \<and> | |
| 576 | (odd (numeral w :: nat) \<and> a \<le> 0 \<or> even (numeral w :: nat) \<and> a = 0)" | |
| 58689 | 577 | by (fact power_le_zero_eq) | 
| 578 | ||
| 579 | lemma power_less_zero_eq_numeral [simp]: | |
| 580 | "a ^ numeral w < 0 \<longleftrightarrow> odd (numeral w :: nat) \<and> a < 0" | |
| 581 | by (fact power_less_zero_eq) | |
| 582 | ||
| 583 | lemma power_even_abs_numeral [simp]: | |
| 584 | "even (numeral w :: nat) \<Longrightarrow> \<bar>a\<bar> ^ numeral w = a ^ numeral w" | |
| 585 | by (fact power_even_abs) | |
| 586 | ||
| 587 | end | |
| 588 | ||
| 71413 | 589 | context unique_euclidean_semiring_with_nat | 
| 590 | begin | |
| 591 | ||
| 592 | lemma even_mask_div_iff': | |
| 593 | \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> m \<le> n\<close> | |
| 594 | proof - | |
| 595 | have \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> even (of_nat ((2 ^ m - Suc 0) div 2 ^ n))\<close> | |
| 596 | by (simp only: of_nat_div) (simp add: of_nat_diff) | |
| 597 | also have \<open>\<dots> \<longleftrightarrow> even ((2 ^ m - Suc 0) div 2 ^ n)\<close> | |
| 598 | by simp | |
| 599 | also have \<open>\<dots> \<longleftrightarrow> m \<le> n\<close> | |
| 600 | proof (cases \<open>m \<le> n\<close>) | |
| 601 | case True | |
| 602 | then show ?thesis | |
| 603 | by (simp add: Suc_le_lessD) | |
| 604 | next | |
| 605 | case False | |
| 606 | then obtain r where r: \<open>m = n + Suc r\<close> | |
| 607 | using less_imp_Suc_add by fastforce | |
| 608 |     from r have \<open>{q. q < m} \<inter> {q. 2 ^ n dvd (2::nat) ^ q} = {q. n \<le> q \<and> q < m}\<close>
 | |
| 609 | by (auto simp add: dvd_power_iff_le) | |
| 610 |     moreover from r have \<open>{q. q < m} \<inter> {q. \<not> 2 ^ n dvd (2::nat) ^ q} = {q. q < n}\<close>
 | |
| 611 | by (auto simp add: dvd_power_iff_le) | |
| 612 |     moreover from False have \<open>{q. n \<le> q \<and> q < m \<and> q \<le> n} = {n}\<close>
 | |
| 613 | by auto | |
| 614 |     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>
 | |
| 615 | 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]) | |
| 616 |     ultimately have \<open>odd (sum ((^) (2::nat)) {q. q < m} div 2 ^ n)\<close>
 | |
| 617 | by (subst euclidean_semiring_cancel_class.sum_div_partition) simp_all | |
| 618 | with False show ?thesis | |
| 619 | by (simp add: mask_eq_sum_exp_nat) | |
| 620 | qed | |
| 621 | finally show ?thesis . | |
| 622 | qed | |
| 623 | ||
| 624 | end | |
| 625 | ||
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 626 | |
| 69593 | 627 | subsection \<open>Instance for \<^typ>\<open>int\<close>\<close> | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 628 | |
| 67816 | 629 | lemma even_diff_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 630 | "even (k - l) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 631 | by (fact even_diff) | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 632 | |
| 67816 | 633 | lemma even_abs_add_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 634 | "even (\<bar>k\<bar> + l) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 635 | by simp | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 636 | |
| 67816 | 637 | lemma even_add_abs_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 638 | "even (k + \<bar>l\<bar>) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 639 | by simp | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 640 | |
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 641 | 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 | 642 | 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 | 643 | |
| 71138 | 644 | lemma zdiv_zmult2_eq: | 
| 645 | \<open>a div (b * c) = (a div b) div c\<close> if \<open>c \<ge> 0\<close> for a b c :: int | |
| 646 | proof (cases \<open>b \<ge> 0\<close>) | |
| 647 | case True | |
| 648 | with that show ?thesis | |
| 649 | using div_mult2_eq' [of a \<open>nat b\<close> \<open>nat c\<close>] by simp | |
| 650 | next | |
| 651 | case False | |
| 652 | with that show ?thesis | |
| 653 | using div_mult2_eq' [of \<open>- a\<close> \<open>nat (- b)\<close> \<open>nat c\<close>] by simp | |
| 654 | qed | |
| 655 | ||
| 656 | lemma zmod_zmult2_eq: | |
| 657 | \<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 | |
| 658 | proof (cases \<open>b \<ge> 0\<close>) | |
| 659 | case True | |
| 660 | with that show ?thesis | |
| 661 | using mod_mult2_eq' [of a \<open>nat b\<close> \<open>nat c\<close>] by simp | |
| 662 | next | |
| 663 | case False | |
| 664 | with that show ?thesis | |
| 665 | using mod_mult2_eq' [of \<open>- a\<close> \<open>nat (- b)\<close> \<open>nat c\<close>] by simp | |
| 666 | qed | |
| 667 | ||
| 71837 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 668 | context | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 669 |   assumes "SORT_CONSTRAINT('a::division_ring)"
 | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 670 | begin | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 671 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 672 | lemma power_int_minus_left: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 673 | "power_int (-a :: 'a) n = (if even n then power_int a n else -power_int a n)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 674 | by (auto simp: power_int_def minus_one_power_iff even_nat_iff) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 675 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 676 | lemma power_int_minus_left_even [simp]: "even n \<Longrightarrow> power_int (-a :: 'a) n = power_int a n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 677 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 678 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 679 | lemma power_int_minus_left_odd [simp]: "odd n \<Longrightarrow> power_int (-a :: 'a) n = -power_int a n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 680 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 681 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 682 | lemma power_int_minus_left_distrib: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 683 | "NO_MATCH (-1) x \<Longrightarrow> power_int (-a :: 'a) n = power_int (-1) n * power_int a n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 684 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 685 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 686 | lemma power_int_minus_one_minus: "power_int (-1 :: 'a) (-n) = power_int (-1) n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 687 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 688 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 689 | lemma power_int_minus_one_diff_commute: "power_int (-1 :: 'a) (a - b) = power_int (-1) (b - a)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 690 | by (subst power_int_minus_one_minus [symmetric]) auto | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 691 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 692 | lemma power_int_minus_one_mult_self [simp]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 693 | "power_int (-1 :: 'a) m * power_int (-1) m = 1" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 694 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 695 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 696 | lemma power_int_minus_one_mult_self' [simp]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 697 | "power_int (-1 :: 'a) m * (power_int (-1) m * b) = b" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 698 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 699 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 700 | end | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 701 | |
| 71094 | 702 | |
| 71181 | 703 | subsection \<open>Abstract bit structures\<close> | 
| 71094 | 704 | |
| 705 | class semiring_bits = semiring_parity + | |
| 71195 | 706 | assumes bits_induct [case_names stable rec]: | 
| 71094 | 707 | \<open>(\<And>a. a div 2 = a \<Longrightarrow> P a) | 
| 708 | \<Longrightarrow> (\<And>a b. P a \<Longrightarrow> (of_bool b + 2 * a) div 2 = a \<Longrightarrow> P (of_bool b + 2 * a)) | |
| 709 | \<Longrightarrow> P a\<close> | |
| 71138 | 710 | assumes bits_div_0 [simp]: \<open>0 div a = 0\<close> | 
| 711 | and bits_div_by_1 [simp]: \<open>a div 1 = a\<close> | |
| 71195 | 712 | and bits_mod_div_trivial [simp]: \<open>a mod b div b = 0\<close> | 
| 71138 | 713 | and even_succ_div_2 [simp]: \<open>even a \<Longrightarrow> (1 + a) div 2 = a div 2\<close> | 
| 71413 | 714 | and even_mask_div_iff: \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> 2 ^ n = 0 \<or> m \<le> n\<close> | 
| 71182 | 715 | 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 | 716 | and div_exp_eq: \<open>a div 2 ^ m div 2 ^ n = a div 2 ^ (m + n)\<close> | 
| 717 | and mod_exp_eq: \<open>a mod 2 ^ m mod 2 ^ n = a mod 2 ^ min m n\<close> | |
| 718 | 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> | |
| 719 | and div_exp_mod_exp_eq: \<open>a div 2 ^ n mod 2 ^ m = a mod (2 ^ (n + m)) div 2 ^ n\<close> | |
| 71424 | 720 | 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> | 
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 721 | fixes bit :: \<open>'a \<Rightarrow> nat \<Rightarrow> bool\<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 722 | assumes bit_iff_odd: \<open>bit a n \<longleftrightarrow> odd (a div 2 ^ n)\<close> | 
| 71138 | 723 | begin | 
| 724 | ||
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 725 | text \<open> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 726 | Having \<^const>\<open>bit\<close> as definitional class operation | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 727 | takes into account that specific instances can be implemented | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 728 | differently wrt. code generation. | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 729 | \<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 730 | |
| 71195 | 731 | lemma bits_div_by_0 [simp]: | 
| 732 | \<open>a div 0 = 0\<close> | |
| 733 | by (metis add_cancel_right_right bits_mod_div_trivial mod_mult_div_eq mult_not_zero) | |
| 734 | ||
| 71138 | 735 | lemma bits_1_div_2 [simp]: | 
| 736 | \<open>1 div 2 = 0\<close> | |
| 737 | using even_succ_div_2 [of 0] by simp | |
| 738 | ||
| 739 | lemma bits_1_div_exp [simp]: | |
| 740 | \<open>1 div 2 ^ n = of_bool (n = 0)\<close> | |
| 741 | using div_exp_eq [of 1 1] by (cases n) simp_all | |
| 742 | ||
| 743 | lemma even_succ_div_exp [simp]: | |
| 744 | \<open>(1 + a) div 2 ^ n = a div 2 ^ n\<close> if \<open>even a\<close> and \<open>n > 0\<close> | |
| 745 | proof (cases n) | |
| 746 | case 0 | |
| 747 | with that show ?thesis | |
| 748 | by simp | |
| 749 | next | |
| 750 | case (Suc n) | |
| 751 | with \<open>even a\<close> have \<open>(1 + a) div 2 ^ Suc n = a div 2 ^ Suc n\<close> | |
| 752 | proof (induction n) | |
| 753 | case 0 | |
| 754 | then show ?case | |
| 755 | by simp | |
| 756 | next | |
| 757 | case (Suc n) | |
| 758 | then show ?case | |
| 759 | using div_exp_eq [of _ 1 \<open>Suc n\<close>, symmetric] | |
| 760 | by simp | |
| 761 | qed | |
| 762 | with Suc show ?thesis | |
| 763 | by simp | |
| 764 | qed | |
| 765 | ||
| 766 | lemma even_succ_mod_exp [simp]: | |
| 767 | \<open>(1 + a) mod 2 ^ n = 1 + (a mod 2 ^ n)\<close> if \<open>even a\<close> and \<open>n > 0\<close> | |
| 768 | using div_mult_mod_eq [of \<open>1 + a\<close> \<open>2 ^ n\<close>] that | |
| 769 | apply simp | |
| 770 | by (metis local.add.left_commute local.add_left_cancel local.div_mult_mod_eq) | |
| 771 | ||
| 772 | lemma bits_mod_by_1 [simp]: | |
| 773 | \<open>a mod 1 = 0\<close> | |
| 774 | using div_mult_mod_eq [of a 1] by simp | |
| 775 | ||
| 776 | lemma bits_mod_0 [simp]: | |
| 777 | \<open>0 mod a = 0\<close> | |
| 778 | using div_mult_mod_eq [of 0 a] by simp | |
| 779 | ||
| 71195 | 780 | lemma bits_one_mod_two_eq_one [simp]: | 
| 71138 | 781 | \<open>1 mod 2 = 1\<close> | 
| 782 | by (simp add: mod2_eq_if) | |
| 783 | ||
| 71181 | 784 | lemma bit_0 [simp]: | 
| 785 | \<open>bit a 0 \<longleftrightarrow> odd a\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 786 | by (simp add: bit_iff_odd) | 
| 71181 | 787 | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 788 | lemma bit_Suc: | 
| 71181 | 789 | \<open>bit a (Suc n) \<longleftrightarrow> bit (a div 2) n\<close> | 
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 790 | using div_exp_eq [of a 1 n] by (simp add: bit_iff_odd) | 
| 71181 | 791 | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 792 | lemma bit_rec: | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 793 | \<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 | 794 | 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 | 795 | |
| 71195 | 796 | lemma bit_0_eq [simp]: | 
| 797 | \<open>bit 0 = bot\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 798 | by (simp add: fun_eq_iff bit_iff_odd) | 
| 71195 | 799 | |
| 71181 | 800 | context | 
| 801 | fixes a | |
| 802 | assumes stable: \<open>a div 2 = a\<close> | |
| 803 | begin | |
| 804 | ||
| 71195 | 805 | lemma bits_stable_imp_add_self: | 
| 71181 | 806 | \<open>a + a mod 2 = 0\<close> | 
| 807 | proof - | |
| 808 | have \<open>a div 2 * 2 + a mod 2 = a\<close> | |
| 809 | by (fact div_mult_mod_eq) | |
| 810 | then have \<open>a * 2 + a mod 2 = a\<close> | |
| 811 | by (simp add: stable) | |
| 812 | then show ?thesis | |
| 813 | by (simp add: mult_2_right ac_simps) | |
| 814 | qed | |
| 815 | ||
| 816 | lemma stable_imp_bit_iff_odd: | |
| 817 | \<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 | 818 | by (induction n) (simp_all add: stable bit_Suc) | 
| 71181 | 819 | |
| 820 | end | |
| 821 | ||
| 822 | lemma bit_iff_idd_imp_stable: | |
| 823 | \<open>a div 2 = a\<close> if \<open>\<And>n. bit a n \<longleftrightarrow> odd a\<close> | |
| 71195 | 824 | using that proof (induction a rule: bits_induct) | 
| 71181 | 825 | case (stable a) | 
| 826 | then show ?case | |
| 827 | by simp | |
| 828 | next | |
| 829 | case (rec a b) | |
| 830 | 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 | 831 | by (simp add: rec.hyps bit_Suc) | 
| 71181 | 832 | from rec.hyps have hyp: \<open>(of_bool (odd a) + 2 * a) div 2 = a\<close> | 
| 833 | by simp | |
| 834 | 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 | 835 | using rec.prems [of \<open>Suc n\<close>] by (simp add: hyp bit_Suc) | 
| 71181 | 836 | then have \<open>a div 2 = a\<close> | 
| 837 | by (rule rec.IH) | |
| 838 | then have \<open>of_bool (odd a) + 2 * a = 2 * (a div 2) + of_bool (odd a)\<close> | |
| 839 | by (simp add: ac_simps) | |
| 840 | also have \<open>\<dots> = a\<close> | |
| 841 | using mult_div_mod_eq [of 2 a] | |
| 842 | by (simp add: of_bool_odd_eq_mod_2) | |
| 843 | finally show ?case | |
| 844 | using \<open>a div 2 = a\<close> by (simp add: hyp) | |
| 845 | qed | |
| 846 | ||
| 71418 | 847 | lemma exp_eq_0_imp_not_bit: | 
| 848 | \<open>\<not> bit a n\<close> if \<open>2 ^ n = 0\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 849 | using that by (simp add: bit_iff_odd) | 
| 71418 | 850 | |
| 71181 | 851 | lemma bit_eqI: | 
| 71418 | 852 | \<open>a = b\<close> if \<open>\<And>n. 2 ^ n \<noteq> 0 \<Longrightarrow> bit a n \<longleftrightarrow> bit b n\<close> | 
| 853 | proof - | |
| 854 | have \<open>bit a n \<longleftrightarrow> bit b n\<close> for n | |
| 855 | proof (cases \<open>2 ^ n = 0\<close>) | |
| 856 | case True | |
| 857 | then show ?thesis | |
| 858 | by (simp add: exp_eq_0_imp_not_bit) | |
| 859 | next | |
| 860 | case False | |
| 861 | then show ?thesis | |
| 862 | by (rule that) | |
| 71181 | 863 | qed | 
| 71418 | 864 | then show ?thesis proof (induction a arbitrary: b rule: bits_induct) | 
| 865 | case (stable a) | |
| 866 | from stable(2) [of 0] have **: \<open>even b \<longleftrightarrow> even a\<close> | |
| 867 | by simp | |
| 868 | have \<open>b div 2 = b\<close> | |
| 869 | proof (rule bit_iff_idd_imp_stable) | |
| 870 | fix n | |
| 871 | from stable have *: \<open>bit b n \<longleftrightarrow> bit a n\<close> | |
| 872 | by simp | |
| 873 | also have \<open>bit a n \<longleftrightarrow> odd a\<close> | |
| 874 | using stable by (simp add: stable_imp_bit_iff_odd) | |
| 875 | finally show \<open>bit b n \<longleftrightarrow> odd b\<close> | |
| 876 | by (simp add: **) | |
| 877 | qed | |
| 878 | from ** have \<open>a mod 2 = b mod 2\<close> | |
| 879 | by (simp add: mod2_eq_if) | |
| 880 | then have \<open>a mod 2 + (a + b) = b mod 2 + (a + b)\<close> | |
| 881 | by simp | |
| 882 | then have \<open>a + a mod 2 + b = b + b mod 2 + a\<close> | |
| 883 | by (simp add: ac_simps) | |
| 884 | with \<open>a div 2 = a\<close> \<open>b div 2 = b\<close> show ?case | |
| 885 | by (simp add: bits_stable_imp_add_self) | |
| 886 | next | |
| 887 | case (rec a p) | |
| 888 | from rec.prems [of 0] have [simp]: \<open>p = odd b\<close> | |
| 889 | by simp | |
| 890 | 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 | 891 | using rec.prems [of \<open>Suc n\<close>] by (simp add: bit_Suc) | 
| 71418 | 892 | then have \<open>a = b div 2\<close> | 
| 893 | by (rule rec.IH) | |
| 894 | then have \<open>2 * a = 2 * (b div 2)\<close> | |
| 895 | by simp | |
| 896 | then have \<open>b mod 2 + 2 * a = b mod 2 + 2 * (b div 2)\<close> | |
| 897 | by simp | |
| 898 | also have \<open>\<dots> = b\<close> | |
| 899 | by (fact mod_mult_div_eq) | |
| 900 | finally show ?case | |
| 901 | by (auto simp add: mod2_eq_if) | |
| 902 | qed | |
| 71181 | 903 | qed | 
| 904 | ||
| 905 | lemma bit_eq_iff: | |
| 906 | \<open>a = b \<longleftrightarrow> (\<forall>n. bit a n \<longleftrightarrow> bit b n)\<close> | |
| 907 | by (auto intro: bit_eqI) | |
| 908 | ||
| 71182 | 909 | lemma bit_exp_iff: | 
| 910 | \<open>bit (2 ^ m) n \<longleftrightarrow> 2 ^ m \<noteq> 0 \<and> m = n\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 911 | by (auto simp add: bit_iff_odd exp_div_exp_eq) | 
| 71182 | 912 | |
| 71408 | 913 | lemma bit_1_iff: | 
| 914 | \<open>bit 1 n \<longleftrightarrow> 1 \<noteq> 0 \<and> n = 0\<close> | |
| 915 | using bit_exp_iff [of 0 n] by simp | |
| 916 | ||
| 917 | lemma bit_2_iff: | |
| 918 | \<open>bit 2 n \<longleftrightarrow> 2 \<noteq> 0 \<and> n = 1\<close> | |
| 919 | using bit_exp_iff [of 1 n] by auto | |
| 920 | ||
| 71418 | 921 | lemma even_bit_succ_iff: | 
| 922 | \<open>bit (1 + a) n \<longleftrightarrow> bit a n \<or> n = 0\<close> if \<open>even a\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 923 | using that by (cases \<open>n = 0\<close>) (simp_all add: bit_iff_odd) | 
| 71418 | 924 | |
| 925 | lemma odd_bit_iff_bit_pred: | |
| 926 | \<open>bit a n \<longleftrightarrow> bit (a - 1) n \<or> n = 0\<close> if \<open>odd a\<close> | |
| 927 | proof - | |
| 928 | from \<open>odd a\<close> obtain b where \<open>a = 2 * b + 1\<close> .. | |
| 929 | moreover have \<open>bit (2 * b) n \<or> n = 0 \<longleftrightarrow> bit (1 + 2 * b) n\<close> | |
| 930 | using even_bit_succ_iff by simp | |
| 931 | ultimately show ?thesis by (simp add: ac_simps) | |
| 932 | qed | |
| 933 | ||
| 71426 | 934 | lemma bit_double_iff: | 
| 935 | \<open>bit (2 * a) n \<longleftrightarrow> bit a (n - 1) \<and> n \<noteq> 0 \<and> 2 ^ n \<noteq> 0\<close> | |
| 936 | using even_mult_exp_div_exp_iff [of a 1 n] | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 937 | by (cases n, auto simp add: bit_iff_odd ac_simps) | 
| 71426 | 938 | |
| 939 | lemma bit_eq_rec: | |
| 71441 | 940 | \<open>a = b \<longleftrightarrow> (even a \<longleftrightarrow> even b) \<and> a div 2 = b div 2\<close> (is \<open>?P = ?Q\<close>) | 
| 941 | proof | |
| 942 | assume ?P | |
| 943 | then show ?Q | |
| 944 | by simp | |
| 945 | next | |
| 946 | assume ?Q | |
| 947 | then have \<open>even a \<longleftrightarrow> even b\<close> and \<open>a div 2 = b div 2\<close> | |
| 948 | by simp_all | |
| 949 | show ?P | |
| 950 | proof (rule bit_eqI) | |
| 951 | fix n | |
| 952 | show \<open>bit a n \<longleftrightarrow> bit b n\<close> | |
| 953 | proof (cases n) | |
| 954 | case 0 | |
| 955 | with \<open>even a \<longleftrightarrow> even b\<close> show ?thesis | |
| 956 | by simp | |
| 957 | next | |
| 958 | case (Suc n) | |
| 959 | moreover from \<open>a div 2 = b div 2\<close> have \<open>bit (a div 2) n = bit (b div 2) n\<close> | |
| 960 | by simp | |
| 961 | ultimately show ?thesis | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 962 | by (simp add: bit_Suc) | 
| 71441 | 963 | qed | 
| 964 | qed | |
| 965 | qed | |
| 71426 | 966 | |
| 71822 | 967 | lemma bit_mod_2_iff [simp]: | 
| 968 | \<open>bit (a mod 2) n \<longleftrightarrow> n = 0 \<and> odd a\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 969 | by (cases a rule: parity_cases) (simp_all add: bit_iff_odd) | 
| 71822 | 970 | |
| 71418 | 971 | lemma bit_mask_iff: | 
| 972 | \<open>bit (2 ^ m - 1) n \<longleftrightarrow> 2 ^ n \<noteq> 0 \<and> n < m\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 973 | by (simp add: bit_iff_odd even_mask_div_iff not_le) | 
| 71418 | 974 | |
| 71757 | 975 | lemma bit_Numeral1_iff [simp]: | 
| 976 | \<open>bit (numeral Num.One) n \<longleftrightarrow> n = 0\<close> | |
| 977 | by (simp add: bit_rec) | |
| 978 | ||
| 72227 | 979 | lemma exp_add_not_zero_imp: | 
| 980 | \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close> if \<open>2 ^ (m + n) \<noteq> 0\<close> | |
| 981 | proof - | |
| 982 | have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close> | |
| 983 | proof (rule notI) | |
| 984 | assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close> | |
| 985 | then have \<open>2 ^ (m + n) = 0\<close> | |
| 986 | by (rule disjE) (simp_all add: power_add) | |
| 987 | with that show False .. | |
| 988 | qed | |
| 989 | then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close> | |
| 990 | by simp_all | |
| 991 | qed | |
| 992 | ||
| 72239 | 993 | lemma bit_disjunctive_add_iff: | 
| 994 | \<open>bit (a + b) n \<longleftrightarrow> bit a n \<or> bit b n\<close> | |
| 995 | if \<open>\<And>n. \<not> bit a n \<or> \<not> bit b n\<close> | |
| 996 | proof (cases \<open>2 ^ n = 0\<close>) | |
| 997 | case True | |
| 998 | then show ?thesis | |
| 999 | by (simp add: exp_eq_0_imp_not_bit) | |
| 1000 | next | |
| 1001 | case False | |
| 1002 | with that show ?thesis proof (induction n arbitrary: a b) | |
| 1003 | case 0 | |
| 1004 | from "0.prems"(1) [of 0] show ?case | |
| 1005 | by auto | |
| 1006 | next | |
| 1007 | case (Suc n) | |
| 1008 | from Suc.prems(1) [of 0] have even: \<open>even a \<or> even b\<close> | |
| 1009 | by auto | |
| 1010 | have bit: \<open>\<not> bit (a div 2) n \<or> \<not> bit (b div 2) n\<close> for n | |
| 1011 | using Suc.prems(1) [of \<open>Suc n\<close>] by (simp add: bit_Suc) | |
| 1012 | from Suc.prems(2) have \<open>2 * 2 ^ n \<noteq> 0\<close> \<open>2 ^ n \<noteq> 0\<close> | |
| 1013 | by (auto simp add: mult_2) | |
| 1014 | have \<open>a + b = (a div 2 * 2 + a mod 2) + (b div 2 * 2 + b mod 2)\<close> | |
| 1015 | using div_mult_mod_eq [of a 2] div_mult_mod_eq [of b 2] by simp | |
| 1016 | also have \<open>\<dots> = of_bool (odd a \<or> odd b) + 2 * (a div 2 + b div 2)\<close> | |
| 1017 | using even by (auto simp add: algebra_simps mod2_eq_if) | |
| 1018 | finally have \<open>bit ((a + b) div 2) n \<longleftrightarrow> bit (a div 2 + b div 2) n\<close> | |
| 1019 | using \<open>2 * 2 ^ n \<noteq> 0\<close> by simp (simp flip: bit_Suc add: bit_double_iff) | |
| 1020 | also have \<open>\<dots> \<longleftrightarrow> bit (a div 2) n \<or> bit (b div 2) n\<close> | |
| 1021 | using bit \<open>2 ^ n \<noteq> 0\<close> by (rule Suc.IH) | |
| 1022 | finally show ?case | |
| 1023 | by (simp add: bit_Suc) | |
| 1024 | qed | |
| 1025 | qed | |
| 1026 | ||
| 72261 | 1027 | lemma | 
| 1028 | exp_add_not_zero_imp_left: \<open>2 ^ m \<noteq> 0\<close> | |
| 1029 | and exp_add_not_zero_imp_right: \<open>2 ^ n \<noteq> 0\<close> | |
| 1030 | if \<open>2 ^ (m + n) \<noteq> 0\<close> | |
| 1031 | proof - | |
| 1032 | have \<open>\<not> (2 ^ m = 0 \<or> 2 ^ n = 0)\<close> | |
| 1033 | proof (rule notI) | |
| 1034 | assume \<open>2 ^ m = 0 \<or> 2 ^ n = 0\<close> | |
| 1035 | then have \<open>2 ^ (m + n) = 0\<close> | |
| 1036 | by (rule disjE) (simp_all add: power_add) | |
| 1037 | with that show False .. | |
| 1038 | qed | |
| 1039 | then show \<open>2 ^ m \<noteq> 0\<close> and \<open>2 ^ n \<noteq> 0\<close> | |
| 1040 | by simp_all | |
| 1041 | qed | |
| 1042 | ||
| 1043 | lemma exp_not_zero_imp_exp_diff_not_zero: | |
| 1044 | \<open>2 ^ (n - m) \<noteq> 0\<close> if \<open>2 ^ n \<noteq> 0\<close> | |
| 1045 | proof (cases \<open>m \<le> n\<close>) | |
| 1046 | case True | |
| 1047 | moreover define q where \<open>q = n - m\<close> | |
| 1048 | ultimately have \<open>n = m + q\<close> | |
| 1049 | by simp | |
| 1050 | with that show ?thesis | |
| 1051 | by (simp add: exp_add_not_zero_imp_right) | |
| 1052 | next | |
| 1053 | case False | |
| 1054 | with that show ?thesis | |
| 1055 | by simp | |
| 1056 | qed | |
| 1057 | ||
| 71138 | 1058 | end | 
| 71094 | 1059 | |
| 1060 | lemma nat_bit_induct [case_names zero even odd]: | |
| 1061 | "P n" if zero: "P 0" | |
| 1062 | and even: "\<And>n. P n \<Longrightarrow> n > 0 \<Longrightarrow> P (2 * n)" | |
| 1063 | and odd: "\<And>n. P n \<Longrightarrow> P (Suc (2 * n))" | |
| 1064 | proof (induction n rule: less_induct) | |
| 1065 | case (less n) | |
| 1066 | show "P n" | |
| 1067 | proof (cases "n = 0") | |
| 1068 | case True with zero show ?thesis by simp | |
| 1069 | next | |
| 1070 | case False | |
| 1071 | with less have hyp: "P (n div 2)" by simp | |
| 1072 | show ?thesis | |
| 1073 | proof (cases "even n") | |
| 1074 | case True | |
| 1075 | then have "n \<noteq> 1" | |
| 1076 | by auto | |
| 1077 | with \<open>n \<noteq> 0\<close> have "n div 2 > 0" | |
| 1078 | by simp | |
| 1079 | with \<open>even n\<close> hyp even [of "n div 2"] show ?thesis | |
| 1080 | by simp | |
| 1081 | next | |
| 1082 | case False | |
| 1083 | with hyp odd [of "n div 2"] show ?thesis | |
| 1084 | by simp | |
| 1085 | qed | |
| 1086 | qed | |
| 1087 | qed | |
| 1088 | ||
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1089 | instantiation nat :: semiring_bits | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1090 | begin | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1091 | |
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1092 | definition bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> bool\<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1093 | where \<open>bit_nat m n \<longleftrightarrow> odd (m div 2 ^ n)\<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1094 | |
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1095 | instance | 
| 71094 | 1096 | proof | 
| 1097 | show \<open>P n\<close> if stable: \<open>\<And>n. n div 2 = n \<Longrightarrow> P n\<close> | |
| 1098 | 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> | |
| 1099 | for P and n :: nat | |
| 1100 | proof (induction n rule: nat_bit_induct) | |
| 1101 | case zero | |
| 1102 | from stable [of 0] show ?case | |
| 1103 | by simp | |
| 1104 | next | |
| 1105 | case (even n) | |
| 1106 | with rec [of n False] show ?case | |
| 1107 | by simp | |
| 1108 | next | |
| 1109 | case (odd n) | |
| 1110 | with rec [of n True] show ?case | |
| 1111 | by simp | |
| 1112 | qed | |
| 71138 | 1113 | show \<open>q mod 2 ^ m mod 2 ^ n = q mod 2 ^ min m n\<close> | 
| 1114 | for q m n :: nat | |
| 1115 | apply (auto simp add: less_iff_Suc_add power_add mod_mod_cancel split: split_min_lin) | |
| 1116 | apply (metis div_mult2_eq mod_div_trivial mod_eq_self_iff_div_eq_0 mod_mult_self2_is_0 power_commutes) | |
| 1117 | done | |
| 1118 | show \<open>(q * 2 ^ m) mod (2 ^ n) = (q mod 2 ^ (n - m)) * 2 ^ m\<close> if \<open>m \<le> n\<close> | |
| 1119 | for q m n :: nat | |
| 1120 | using that | |
| 1121 | apply (auto simp add: mod_mod_cancel div_mult2_eq power_add mod_mult2_eq le_iff_add split: split_min_lin) | |
| 1122 | apply (simp add: mult.commute) | |
| 1123 | done | |
| 71413 | 1124 | show \<open>even ((2 ^ m - (1::nat)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::nat) \<or> m \<le> n\<close> | 
| 1125 | for m n :: nat | |
| 1126 | using even_mask_div_iff' [where ?'a = nat, of m n] by simp | |
| 71424 | 1127 | 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> | 
| 1128 | for m n q r :: nat | |
| 1129 | apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex) | |
| 1130 | apply (metis (full_types) dvd_mult dvd_mult_imp_div dvd_power_iff_le not_less not_less_eq order_refl power_Suc) | |
| 1131 | done | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1132 | qed (auto simp add: div_mult2_eq mod_mult2_eq power_add power_diff bit_nat_def) | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1133 | |
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1134 | end | 
| 71094 | 1135 | |
| 70353 | 1136 | lemma int_bit_induct [case_names zero minus even odd]: | 
| 70338 | 1137 | "P k" if zero_int: "P 0" | 
| 1138 | and minus_int: "P (- 1)" | |
| 1139 | and even_int: "\<And>k. P k \<Longrightarrow> k \<noteq> 0 \<Longrightarrow> P (k * 2)" | |
| 1140 | and odd_int: "\<And>k. P k \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> P (1 + (k * 2))" for k :: int | |
| 1141 | proof (cases "k \<ge> 0") | |
| 1142 | case True | |
| 1143 | define n where "n = nat k" | |
| 1144 | with True have "k = int n" | |
| 1145 | by simp | |
| 1146 | then show "P k" | |
| 70353 | 1147 | proof (induction n arbitrary: k rule: nat_bit_induct) | 
| 70338 | 1148 | case zero | 
| 1149 | then show ?case | |
| 1150 | by (simp add: zero_int) | |
| 1151 | next | |
| 1152 | case (even n) | |
| 1153 | have "P (int n * 2)" | |
| 1154 | by (rule even_int) (use even in simp_all) | |
| 1155 | with even show ?case | |
| 1156 | by (simp add: ac_simps) | |
| 1157 | next | |
| 1158 | case (odd n) | |
| 1159 | have "P (1 + (int n * 2))" | |
| 1160 | by (rule odd_int) (use odd in simp_all) | |
| 1161 | with odd show ?case | |
| 1162 | by (simp add: ac_simps) | |
| 1163 | qed | |
| 1164 | next | |
| 1165 | case False | |
| 1166 | define n where "n = nat (- k - 1)" | |
| 1167 | with False have "k = - int n - 1" | |
| 1168 | by simp | |
| 1169 | then show "P k" | |
| 70353 | 1170 | proof (induction n arbitrary: k rule: nat_bit_induct) | 
| 70338 | 1171 | case zero | 
| 1172 | then show ?case | |
| 1173 | by (simp add: minus_int) | |
| 1174 | next | |
| 1175 | case (even n) | |
| 1176 | have "P (1 + (- int (Suc n) * 2))" | |
| 1177 | by (rule odd_int) (use even in \<open>simp_all add: algebra_simps\<close>) | |
| 1178 | also have "\<dots> = - int (2 * n) - 1" | |
| 1179 | by (simp add: algebra_simps) | |
| 1180 | finally show ?case | |
| 72261 | 1181 | using even.prems by simp | 
| 70338 | 1182 | next | 
| 1183 | case (odd n) | |
| 1184 | have "P (- int (Suc n) * 2)" | |
| 1185 | by (rule even_int) (use odd in \<open>simp_all add: algebra_simps\<close>) | |
| 1186 | also have "\<dots> = - int (Suc (2 * n)) - 1" | |
| 1187 | by (simp add: algebra_simps) | |
| 1188 | finally show ?case | |
| 72261 | 1189 | using odd.prems by simp | 
| 70338 | 1190 | qed | 
| 1191 | qed | |
| 1192 | ||
| 72227 | 1193 | context semiring_bits | 
| 1194 | begin | |
| 1195 | ||
| 72281 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 haftmann parents: 
72262diff
changeset | 1196 | lemma bit_of_bool_iff: | 
| 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 haftmann parents: 
72262diff
changeset | 1197 | \<open>bit (of_bool b) n \<longleftrightarrow> b \<and> n = 0\<close> | 
| 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 haftmann parents: 
72262diff
changeset | 1198 | by (simp add: bit_1_iff) | 
| 
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
 haftmann parents: 
72262diff
changeset | 1199 | |
| 72227 | 1200 | lemma even_of_nat_iff: | 
| 1201 | \<open>even (of_nat n) \<longleftrightarrow> even n\<close> | |
| 1202 | by (induction n rule: nat_bit_induct) simp_all | |
| 1203 | ||
| 1204 | lemma bit_of_nat_iff: | |
| 1205 | \<open>bit (of_nat m) n \<longleftrightarrow> (2::'a) ^ n \<noteq> 0 \<and> bit m n\<close> | |
| 1206 | proof (cases \<open>(2::'a) ^ n = 0\<close>) | |
| 1207 | case True | |
| 1208 | then show ?thesis | |
| 1209 | by (simp add: exp_eq_0_imp_not_bit) | |
| 1210 | next | |
| 1211 | case False | |
| 1212 | then have \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close> | |
| 1213 | proof (induction m arbitrary: n rule: nat_bit_induct) | |
| 1214 | case zero | |
| 1215 | then show ?case | |
| 1216 | by simp | |
| 1217 | next | |
| 1218 | case (even m) | |
| 1219 | then show ?case | |
| 1220 | by (cases n) | |
| 1221 | (auto simp add: bit_double_iff Parity.bit_double_iff dest: mult_not_zero) | |
| 1222 | next | |
| 1223 | case (odd m) | |
| 1224 | then show ?case | |
| 1225 | by (cases n) | |
| 1226 | (auto simp add: bit_double_iff even_bit_succ_iff Parity.bit_Suc dest: mult_not_zero) | |
| 1227 | qed | |
| 1228 | with False show ?thesis | |
| 1229 | by simp | |
| 1230 | qed | |
| 1231 | ||
| 1232 | end | |
| 1233 | ||
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1234 | instantiation int :: semiring_bits | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1235 | begin | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1236 | |
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1237 | definition bit_int :: \<open>int \<Rightarrow> nat \<Rightarrow> bool\<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1238 | where \<open>bit_int k n \<longleftrightarrow> odd (k div 2 ^ n)\<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1239 | |
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1240 | instance | 
| 71094 | 1241 | proof | 
| 1242 | show \<open>P k\<close> if stable: \<open>\<And>k. k div 2 = k \<Longrightarrow> P k\<close> | |
| 1243 | 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> | |
| 1244 | for P and k :: int | |
| 1245 | proof (induction k rule: int_bit_induct) | |
| 1246 | case zero | |
| 1247 | from stable [of 0] show ?case | |
| 1248 | by simp | |
| 1249 | next | |
| 1250 | case minus | |
| 1251 | from stable [of \<open>- 1\<close>] show ?case | |
| 1252 | by simp | |
| 1253 | next | |
| 1254 | case (even k) | |
| 1255 | with rec [of k False] show ?case | |
| 1256 | by (simp add: ac_simps) | |
| 1257 | next | |
| 1258 | case (odd k) | |
| 1259 | with rec [of k True] show ?case | |
| 1260 | by (simp add: ac_simps) | |
| 1261 | qed | |
| 71182 | 1262 | show \<open>(2::int) ^ m div 2 ^ n = of_bool ((2::int) ^ m \<noteq> 0 \<and> n \<le> m) * 2 ^ (m - n)\<close> | 
| 1263 | for m n :: nat | |
| 1264 | proof (cases \<open>m < n\<close>) | |
| 1265 | case True | |
| 1266 | then have \<open>n = m + (n - m)\<close> | |
| 1267 | by simp | |
| 1268 | then have \<open>(2::int) ^ m div 2 ^ n = (2::int) ^ m div 2 ^ (m + (n - m))\<close> | |
| 1269 | by simp | |
| 1270 | also have \<open>\<dots> = (2::int) ^ m div (2 ^ m * 2 ^ (n - m))\<close> | |
| 1271 | by (simp add: power_add) | |
| 1272 | also have \<open>\<dots> = (2::int) ^ m div 2 ^ m div 2 ^ (n - m)\<close> | |
| 1273 | by (simp add: zdiv_zmult2_eq) | |
| 1274 | finally show ?thesis using \<open>m < n\<close> by simp | |
| 1275 | next | |
| 1276 | case False | |
| 1277 | then show ?thesis | |
| 1278 | by (simp add: power_diff) | |
| 1279 | qed | |
| 71138 | 1280 | show \<open>k mod 2 ^ m mod 2 ^ n = k mod 2 ^ min m n\<close> | 
| 1281 | for m n :: nat and k :: int | |
| 1282 | using mod_exp_eq [of \<open>nat k\<close> m n] | |
| 1283 | apply (auto simp add: mod_mod_cancel zdiv_zmult2_eq power_add zmod_zmult2_eq le_iff_add split: split_min_lin) | |
| 1284 | apply (auto simp add: less_iff_Suc_add mod_mod_cancel power_add) | |
| 1285 | apply (simp only: flip: mult.left_commute [of \<open>2 ^ m\<close>]) | |
| 1286 | apply (subst zmod_zmult2_eq) apply simp_all | |
| 1287 | done | |
| 1288 | show \<open>(k * 2 ^ m) mod (2 ^ n) = (k mod 2 ^ (n - m)) * 2 ^ m\<close> | |
| 1289 | if \<open>m \<le> n\<close> for m n :: nat and k :: int | |
| 1290 | using that | |
| 1291 | apply (auto simp add: power_add zmod_zmult2_eq le_iff_add split: split_min_lin) | |
| 1292 | apply (simp add: ac_simps) | |
| 1293 | done | |
| 71413 | 1294 | show \<open>even ((2 ^ m - (1::int)) div 2 ^ n) \<longleftrightarrow> 2 ^ n = (0::int) \<or> m \<le> n\<close> | 
| 1295 | for m n :: nat | |
| 1296 | using even_mask_div_iff' [where ?'a = int, of m n] by simp | |
| 71424 | 1297 | 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> | 
| 1298 | for m n :: nat and k l :: int | |
| 1299 | apply (auto simp add: not_less power_add ac_simps dest!: le_Suc_ex) | |
| 1300 | apply (metis Suc_leI dvd_mult dvd_mult_imp_div dvd_power_le dvd_refl power.simps(2)) | |
| 1301 | done | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1302 | qed (auto simp add: zdiv_zmult2_eq zmod_zmult2_eq power_add power_diff not_le bit_int_def) | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1303 | |
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1304 | end | 
| 67816 | 1305 | |
| 71094 | 1306 | class semiring_bit_shifts = semiring_bits + | 
| 1307 | fixes push_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | |
| 1308 | assumes push_bit_eq_mult: \<open>push_bit n a = a * 2 ^ n\<close> | |
| 1309 | fixes drop_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | |
| 1310 | assumes drop_bit_eq_div: \<open>drop_bit n a = a div 2 ^ n\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1311 | fixes take_bit :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'a\<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1312 | assumes take_bit_eq_mod: \<open>take_bit n a = a mod 2 ^ n\<close> | 
| 67816 | 1313 | begin | 
| 1314 | ||
| 71094 | 1315 | text \<open> | 
| 1316 | Logically, \<^const>\<open>push_bit\<close>, | |
| 1317 | \<^const>\<open>drop_bit\<close> and \<^const>\<open>take_bit\<close> are just aliases; having them | |
| 1318 | as separate operations makes proofs easier, otherwise proof automation | |
| 1319 | would fiddle with concrete expressions \<^term>\<open>2 ^ n\<close> in a way obfuscating the basic | |
| 1320 | algebraic relationships between those operations. | |
| 1321 | Having | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1322 | them as definitional class operations | 
| 71094 | 1323 | takes into account that specific instances of these can be implemented | 
| 1324 | differently wrt. code generation. | |
| 1325 | \<close> | |
| 67816 | 1326 | |
| 71408 | 1327 | lemma bit_iff_odd_drop_bit: | 
| 1328 | \<open>bit a n \<longleftrightarrow> odd (drop_bit n a)\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1329 | by (simp add: bit_iff_odd drop_bit_eq_div) | 
| 71408 | 1330 | |
| 1331 | lemma even_drop_bit_iff_not_bit: | |
| 1332 | \<open>even (drop_bit n a) \<longleftrightarrow> \<not> bit a n\<close> | |
| 1333 | by (simp add: bit_iff_odd_drop_bit) | |
| 1334 | ||
| 71423 | 1335 | lemma div_push_bit_of_1_eq_drop_bit: | 
| 1336 | \<open>a div push_bit n 1 = drop_bit n a\<close> | |
| 1337 | by (simp add: push_bit_eq_mult drop_bit_eq_div) | |
| 1338 | ||
| 71195 | 1339 | lemma bits_ident: | 
| 71138 | 1340 | "push_bit n (drop_bit n a) + take_bit n a = a" | 
| 1341 | using div_mult_mod_eq by (simp add: push_bit_eq_mult take_bit_eq_mod drop_bit_eq_div) | |
| 1342 | ||
| 1343 | lemma push_bit_push_bit [simp]: | |
| 1344 | "push_bit m (push_bit n a) = push_bit (m + n) a" | |
| 1345 | by (simp add: push_bit_eq_mult power_add ac_simps) | |
| 1346 | ||
| 1347 | lemma push_bit_0_id [simp]: | |
| 1348 | "push_bit 0 = id" | |
| 1349 | by (simp add: fun_eq_iff push_bit_eq_mult) | |
| 1350 | ||
| 1351 | lemma push_bit_of_0 [simp]: | |
| 1352 | "push_bit n 0 = 0" | |
| 1353 | by (simp add: push_bit_eq_mult) | |
| 1354 | ||
| 1355 | lemma push_bit_of_1: | |
| 1356 | "push_bit n 1 = 2 ^ n" | |
| 1357 | by (simp add: push_bit_eq_mult) | |
| 1358 | ||
| 1359 | lemma push_bit_Suc [simp]: | |
| 1360 | "push_bit (Suc n) a = push_bit n (a * 2)" | |
| 1361 | by (simp add: push_bit_eq_mult ac_simps) | |
| 1362 | ||
| 1363 | lemma push_bit_double: | |
| 1364 | "push_bit n (a * 2) = push_bit n a * 2" | |
| 1365 | by (simp add: push_bit_eq_mult ac_simps) | |
| 1366 | ||
| 1367 | lemma push_bit_add: | |
| 1368 | "push_bit n (a + b) = push_bit n a + push_bit n b" | |
| 1369 | by (simp add: push_bit_eq_mult algebra_simps) | |
| 1370 | ||
| 71991 | 1371 | lemma push_bit_numeral [simp]: | 
| 1372 | \<open>push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))\<close> | |
| 1373 | by (simp add: numeral_eq_Suc mult_2_right) (simp add: numeral_Bit0) | |
| 1374 | ||
| 71138 | 1375 | lemma take_bit_0 [simp]: | 
| 1376 | "take_bit 0 a = 0" | |
| 1377 | by (simp add: take_bit_eq_mod) | |
| 1378 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1379 | lemma take_bit_Suc: | 
| 71822 | 1380 | \<open>take_bit (Suc n) a = take_bit n (a div 2) * 2 + a mod 2\<close> | 
| 71138 | 1381 | proof - | 
| 1382 | 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> | |
| 1383 | using even_succ_mod_exp [of \<open>2 * (a div 2)\<close> \<open>Suc n\<close>] | |
| 1384 | mult_exp_mod_exp_eq [of 1 \<open>Suc n\<close> \<open>a div 2\<close>] | |
| 1385 | by (auto simp add: take_bit_eq_mod ac_simps) | |
| 1386 | then show ?thesis | |
| 1387 | using div_mult_mod_eq [of a 2] by (simp add: mod_2_eq_odd) | |
| 1388 | qed | |
| 1389 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1390 | lemma take_bit_rec: | 
| 71822 | 1391 | \<open>take_bit n a = (if n = 0 then 0 else take_bit (n - 1) (a div 2) * 2 + a mod 2)\<close> | 
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1392 | 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 | 1393 | |
| 71759 | 1394 | lemma take_bit_Suc_0 [simp]: | 
| 1395 | \<open>take_bit (Suc 0) a = a mod 2\<close> | |
| 1396 | by (simp add: take_bit_eq_mod) | |
| 1397 | ||
| 71138 | 1398 | lemma take_bit_of_0 [simp]: | 
| 1399 | "take_bit n 0 = 0" | |
| 1400 | by (simp add: take_bit_eq_mod) | |
| 1401 | ||
| 1402 | lemma take_bit_of_1 [simp]: | |
| 1403 | "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 | 1404 | by (cases n) (simp_all add: take_bit_Suc) | 
| 71138 | 1405 | |
| 1406 | lemma drop_bit_of_0 [simp]: | |
| 1407 | "drop_bit n 0 = 0" | |
| 1408 | by (simp add: drop_bit_eq_div) | |
| 1409 | ||
| 1410 | lemma drop_bit_of_1 [simp]: | |
| 1411 | "drop_bit n 1 = of_bool (n = 0)" | |
| 1412 | by (simp add: drop_bit_eq_div) | |
| 1413 | ||
| 1414 | lemma drop_bit_0 [simp]: | |
| 1415 | "drop_bit 0 = id" | |
| 1416 | by (simp add: fun_eq_iff drop_bit_eq_div) | |
| 1417 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1418 | lemma drop_bit_Suc: | 
| 71138 | 1419 | "drop_bit (Suc n) a = drop_bit n (a div 2)" | 
| 1420 | using div_exp_eq [of a 1] by (simp add: drop_bit_eq_div) | |
| 1421 | ||
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1422 | lemma drop_bit_rec: | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1423 | "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 | 1424 | 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 | 1425 | |
| 71138 | 1426 | lemma drop_bit_half: | 
| 1427 | "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 | 1428 | by (induction n arbitrary: a) (simp_all add: drop_bit_Suc) | 
| 71138 | 1429 | |
| 1430 | 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 | 1431 | "drop_bit n (of_bool b) = of_bool (n = 0 \<and> b)" | 
| 71138 | 1432 | by (cases n) simp_all | 
| 1433 | ||
| 1434 | lemma even_take_bit_eq [simp]: | |
| 1435 | \<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 | 1436 | by (simp add: take_bit_rec [of n a]) | 
| 71138 | 1437 | |
| 1438 | lemma take_bit_take_bit [simp]: | |
| 1439 | "take_bit m (take_bit n a) = take_bit (min m n) a" | |
| 1440 | by (simp add: take_bit_eq_mod mod_exp_eq ac_simps) | |
| 1441 | ||
| 1442 | lemma drop_bit_drop_bit [simp]: | |
| 1443 | "drop_bit m (drop_bit n a) = drop_bit (m + n) a" | |
| 1444 | by (simp add: drop_bit_eq_div power_add div_exp_eq ac_simps) | |
| 1445 | ||
| 1446 | lemma push_bit_take_bit: | |
| 1447 | "push_bit m (take_bit n a) = take_bit (m + n) (push_bit m a)" | |
| 1448 | apply (simp add: push_bit_eq_mult take_bit_eq_mod power_add ac_simps) | |
| 1449 | using mult_exp_mod_exp_eq [of m \<open>m + n\<close> a] apply (simp add: ac_simps power_add) | |
| 1450 | done | |
| 1451 | ||
| 1452 | lemma take_bit_push_bit: | |
| 1453 | "take_bit m (push_bit n a) = push_bit n (take_bit (m - n) a)" | |
| 1454 | proof (cases "m \<le> n") | |
| 1455 | case True | |
| 1456 | then show ?thesis | |
| 1457 | apply (simp add:) | |
| 1458 | apply (simp_all add: push_bit_eq_mult take_bit_eq_mod) | |
| 1459 | apply (auto dest!: le_Suc_ex simp add: power_add ac_simps) | |
| 1460 | using mult_exp_mod_exp_eq [of m m \<open>a * 2 ^ n\<close> for n] | |
| 1461 | apply (simp add: ac_simps) | |
| 1462 | done | |
| 1463 | next | |
| 1464 | case False | |
| 1465 | then show ?thesis | |
| 1466 | using push_bit_take_bit [of n "m - n" a] | |
| 1467 | by simp | |
| 1468 | qed | |
| 1469 | ||
| 1470 | lemma take_bit_drop_bit: | |
| 1471 | "take_bit m (drop_bit n a) = drop_bit n (take_bit (m + n) a)" | |
| 1472 | by (simp add: drop_bit_eq_div take_bit_eq_mod ac_simps div_exp_mod_exp_eq) | |
| 1473 | ||
| 1474 | lemma drop_bit_take_bit: | |
| 1475 | "drop_bit m (take_bit n a) = take_bit (n - m) (drop_bit m a)" | |
| 1476 | proof (cases "m \<le> n") | |
| 1477 | case True | |
| 1478 | then show ?thesis | |
| 1479 | using take_bit_drop_bit [of "n - m" m a] by simp | |
| 1480 | next | |
| 1481 | case False | |
| 1482 | then obtain q where \<open>m = n + q\<close> | |
| 1483 | by (auto simp add: not_le dest: less_imp_Suc_add) | |
| 1484 | then have \<open>drop_bit m (take_bit n a) = 0\<close> | |
| 1485 | using div_exp_eq [of \<open>a mod 2 ^ n\<close> n q] | |
| 1486 | by (simp add: take_bit_eq_mod drop_bit_eq_div) | |
| 1487 | with False show ?thesis | |
| 1488 | by simp | |
| 1489 | qed | |
| 1490 | ||
| 71424 | 1491 | lemma even_push_bit_iff [simp]: | 
| 1492 | \<open>even (push_bit n a) \<longleftrightarrow> n \<noteq> 0 \<or> even a\<close> | |
| 1493 | by (simp add: push_bit_eq_mult) auto | |
| 1494 | ||
| 1495 | lemma bit_push_bit_iff: | |
| 71991 | 1496 | \<open>bit (push_bit m a) n \<longleftrightarrow> m \<le> n \<and> 2 ^ n \<noteq> 0 \<and> bit a (n - m)\<close> | 
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1497 | by (auto simp add: bit_iff_odd push_bit_eq_mult even_mult_exp_div_exp_iff) | 
| 71424 | 1498 | |
| 71181 | 1499 | lemma bit_drop_bit_eq: | 
| 1500 | \<open>bit (drop_bit n a) = bit a \<circ> (+) n\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1501 | by (simp add: bit_iff_odd fun_eq_iff ac_simps flip: drop_bit_eq_div) | 
| 71181 | 1502 | |
| 1503 | lemma bit_take_bit_iff: | |
| 1504 | \<open>bit (take_bit m a) n \<longleftrightarrow> n < m \<and> bit a n\<close> | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1505 | by (simp add: bit_iff_odd drop_bit_take_bit not_le flip: drop_bit_eq_div) | 
| 71181 | 1506 | |
| 71535 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1507 | lemma stable_imp_drop_bit_eq: | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1508 | \<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 | 1509 | 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 | 1510 | 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 | 1511 | |
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1512 | lemma stable_imp_take_bit_eq: | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1513 | \<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 | 1514 | 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 | 1515 | proof (rule bit_eqI) | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1516 | fix m | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1517 | 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 | 1518 | 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 | 1519 | 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 | 1520 | qed | 
| 
b612edee9b0c
more frugal simp rules for bit operations; more pervasive use of bit selector
 haftmann parents: 
71441diff
changeset | 1521 | |
| 71958 | 1522 | lemma exp_dvdE: | 
| 1523 | assumes \<open>2 ^ n dvd a\<close> | |
| 1524 | obtains b where \<open>a = push_bit n b\<close> | |
| 1525 | proof - | |
| 1526 | from assms obtain b where \<open>a = 2 ^ n * b\<close> .. | |
| 1527 | then have \<open>a = push_bit n b\<close> | |
| 1528 | by (simp add: push_bit_eq_mult ac_simps) | |
| 1529 | with that show thesis . | |
| 1530 | qed | |
| 1531 | ||
| 1532 | lemma take_bit_eq_0_iff: | |
| 1533 | \<open>take_bit n a = 0 \<longleftrightarrow> 2 ^ n dvd a\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 1534 | proof | |
| 1535 | assume ?P | |
| 1536 | then show ?Q | |
| 1537 | by (simp add: take_bit_eq_mod mod_0_imp_dvd) | |
| 1538 | next | |
| 1539 | assume ?Q | |
| 1540 | then obtain b where \<open>a = push_bit n b\<close> | |
| 1541 | by (rule exp_dvdE) | |
| 1542 | then show ?P | |
| 1543 | by (simp add: take_bit_push_bit) | |
| 1544 | qed | |
| 1545 | ||
| 72079 | 1546 | lemma take_bit_tightened: | 
| 1547 | \<open>take_bit m a = take_bit m b\<close> if \<open>take_bit n a = take_bit n b\<close> and \<open>m \<le> n\<close> | |
| 1548 | proof - | |
| 1549 | from that have \<open>take_bit m (take_bit n a) = take_bit m (take_bit n b)\<close> | |
| 1550 | by simp | |
| 1551 | then have \<open>take_bit (min m n) a = take_bit (min m n) b\<close> | |
| 1552 | by simp | |
| 1553 | with that show ?thesis | |
| 1554 | by (simp add: min_def) | |
| 1555 | qed | |
| 1556 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 1557 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 1558 | |
| 71094 | 1559 | instantiation nat :: semiring_bit_shifts | 
| 1560 | begin | |
| 1561 | ||
| 1562 | definition push_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | |
| 1563 | where \<open>push_bit_nat n m = m * 2 ^ n\<close> | |
| 1564 | ||
| 1565 | definition drop_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | |
| 1566 | where \<open>drop_bit_nat n m = m div 2 ^ n\<close> | |
| 1567 | ||
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1568 | definition take_bit_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1569 | where \<open>take_bit_nat n m = m mod 2 ^ n\<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1570 | |
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1571 | instance | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1572 | by standard (simp_all add: push_bit_nat_def drop_bit_nat_def take_bit_nat_def) | 
| 71094 | 1573 | |
| 1574 | end | |
| 1575 | ||
| 72227 | 1576 | context semiring_bit_shifts | 
| 1577 | begin | |
| 1578 | ||
| 1579 | lemma push_bit_of_nat: | |
| 1580 | \<open>push_bit n (of_nat m) = of_nat (push_bit n m)\<close> | |
| 1581 | by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult) | |
| 1582 | ||
| 1583 | lemma of_nat_push_bit: | |
| 1584 | \<open>of_nat (push_bit m n) = push_bit m (of_nat n)\<close> | |
| 1585 | by (simp add: push_bit_eq_mult semiring_bit_shifts_class.push_bit_eq_mult) | |
| 1586 | ||
| 1587 | lemma take_bit_of_nat: | |
| 1588 | \<open>take_bit n (of_nat m) = of_nat (take_bit n m)\<close> | |
| 1589 | by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_nat_iff) | |
| 1590 | ||
| 1591 | lemma of_nat_take_bit: | |
| 1592 | \<open>of_nat (take_bit n m) = take_bit n (of_nat m)\<close> | |
| 1593 | by (rule bit_eqI) (simp add: bit_take_bit_iff Parity.bit_take_bit_iff bit_of_nat_iff) | |
| 1594 | ||
| 1595 | end | |
| 1596 | ||
| 71094 | 1597 | instantiation int :: semiring_bit_shifts | 
| 1598 | begin | |
| 1599 | ||
| 1600 | definition push_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | |
| 1601 | where \<open>push_bit_int n k = k * 2 ^ n\<close> | |
| 1602 | ||
| 1603 | definition drop_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | |
| 1604 | where \<open>drop_bit_int n k = k div 2 ^ n\<close> | |
| 1605 | ||
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1606 | definition take_bit_int :: \<open>nat \<Rightarrow> int \<Rightarrow> int\<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1607 | where \<open>take_bit_int n k = k mod 2 ^ n\<close> | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1608 | |
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1609 | instance | 
| 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1610 | by standard (simp_all add: push_bit_int_def drop_bit_int_def take_bit_int_def) | 
| 71094 | 1611 | |
| 1612 | end | |
| 1613 | ||
| 71412 | 1614 | lemma bit_push_bit_iff_nat: | 
| 1615 | \<open>bit (push_bit m q) n \<longleftrightarrow> m \<le> n \<and> bit q (n - m)\<close> for q :: nat | |
| 71424 | 1616 | by (auto simp add: bit_push_bit_iff) | 
| 71412 | 1617 | |
| 1618 | lemma bit_push_bit_iff_int: | |
| 1619 | \<open>bit (push_bit m k) n \<longleftrightarrow> m \<le> n \<and> bit k (n - m)\<close> for k :: int | |
| 71424 | 1620 | by (auto simp add: bit_push_bit_iff) | 
| 71412 | 1621 | |
| 72262 | 1622 | lemma take_bit_nat_less_exp [simp]: | 
| 72261 | 1623 | \<open>take_bit n m < 2 ^ n\<close> for n m ::nat | 
| 1624 | by (simp add: take_bit_eq_mod) | |
| 1625 | ||
| 72262 | 1626 | lemma take_bit_nonnegative [simp]: | 
| 1627 | \<open>take_bit n k \<ge> 0\<close> for k :: int | |
| 1628 | by (simp add: take_bit_eq_mod) | |
| 1629 | ||
| 1630 | lemma not_take_bit_negative [simp]: | |
| 1631 | \<open>\<not> take_bit n k < 0\<close> for k :: int | |
| 1632 | by (simp add: not_less) | |
| 1633 | ||
| 1634 | lemma take_bit_int_less_exp [simp]: | |
| 1635 | \<open>take_bit n k < 2 ^ n\<close> for k :: int | |
| 1636 | by (simp add: take_bit_eq_mod) | |
| 1637 | ||
| 72261 | 1638 | lemma take_bit_nat_eq_self_iff: | 
| 1639 | \<open>take_bit n m = m \<longleftrightarrow> m < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 1640 | for n m :: nat | |
| 1641 | proof | |
| 1642 | assume ?P | |
| 1643 | moreover note take_bit_nat_less_exp [of n m] | |
| 1644 | ultimately show ?Q | |
| 1645 | by simp | |
| 1646 | next | |
| 1647 | assume ?Q | |
| 1648 | then show ?P | |
| 1649 | by (simp add: take_bit_eq_mod) | |
| 1650 | qed | |
| 1651 | ||
| 72262 | 1652 | lemma take_bit_nat_eq_self: | 
| 1653 | \<open>take_bit n m = m\<close> if \<open>m < 2 ^ n\<close> for m n :: nat | |
| 1654 | using that by (simp add: take_bit_nat_eq_self_iff) | |
| 72261 | 1655 | |
| 1656 | lemma take_bit_int_eq_self_iff: | |
| 1657 | \<open>take_bit n k = k \<longleftrightarrow> 0 \<le> k \<and> k < 2 ^ n\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 1658 | for k :: int | |
| 1659 | proof | |
| 1660 | assume ?P | |
| 1661 | moreover note take_bit_int_less_exp [of n k] take_bit_nonnegative [of n k] | |
| 1662 | ultimately show ?Q | |
| 1663 | by simp | |
| 1664 | next | |
| 1665 | assume ?Q | |
| 1666 | then show ?P | |
| 1667 | by (simp add: take_bit_eq_mod) | |
| 1668 | qed | |
| 1669 | ||
| 72262 | 1670 | lemma take_bit_int_eq_self: | 
| 1671 | \<open>take_bit n k = k\<close> if \<open>0 \<le> k\<close> \<open>k < 2 ^ n\<close> for k :: int | |
| 1672 | using that by (simp add: take_bit_int_eq_self_iff) | |
| 1673 | ||
| 1674 | lemma take_bit_nat_less_eq_self [simp]: | |
| 1675 | \<open>take_bit n m \<le> m\<close> for n m :: nat | |
| 1676 | by (simp add: take_bit_eq_mod) | |
| 1677 | ||
| 1678 | lemma take_bit_nat_less_self_iff: | |
| 1679 | \<open>take_bit n m < m \<longleftrightarrow> 2 ^ n \<le> m\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 1680 | for m n :: nat | |
| 1681 | proof | |
| 1682 | assume ?P | |
| 1683 | then have \<open>take_bit n m \<noteq> m\<close> | |
| 1684 | by simp | |
| 1685 | then show \<open>?Q\<close> | |
| 1686 | by (simp add: take_bit_nat_eq_self_iff) | |
| 1687 | next | |
| 1688 | have \<open>take_bit n m < 2 ^ n\<close> | |
| 1689 | by (fact take_bit_nat_less_exp) | |
| 1690 | also assume ?Q | |
| 1691 | finally show ?P . | |
| 1692 | qed | |
| 1693 | ||
| 71094 | 1694 | class unique_euclidean_semiring_with_bit_shifts = | 
| 1695 | 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 | 1696 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 1697 | |
| 71138 | 1698 | lemma take_bit_of_exp [simp]: | 
| 1699 | \<open>take_bit m (2 ^ n) = of_bool (n < m) * 2 ^ n\<close> | |
| 1700 | by (simp add: take_bit_eq_mod exp_mod_exp) | |
| 67960 | 1701 | |
| 71138 | 1702 | lemma take_bit_of_2 [simp]: | 
| 1703 | \<open>take_bit n 2 = of_bool (2 \<le> n) * 2\<close> | |
| 1704 | using take_bit_of_exp [of n 1] by simp | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1705 | |
| 71412 | 1706 | lemma take_bit_of_mask: | 
| 71408 | 1707 | \<open>take_bit m (2 ^ n - 1) = 2 ^ min m n - 1\<close> | 
| 71412 | 1708 | by (simp add: take_bit_eq_mod mask_mod_exp) | 
| 71408 | 1709 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1710 | lemma push_bit_eq_0_iff [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1711 | "push_bit n a = 0 \<longleftrightarrow> a = 0" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1712 | by (simp add: push_bit_eq_mult) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1713 | |
| 67961 | 1714 | lemma take_bit_add: | 
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1715 | "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 | 1716 | by (simp add: take_bit_eq_mod mod_simps) | 
| 67816 | 1717 | |
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1718 | lemma take_bit_of_1_eq_0_iff [simp]: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1719 | "take_bit n 1 = 0 \<longleftrightarrow> n = 0" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1720 | by (simp add: take_bit_eq_mod) | 
| 67816 | 1721 | |
| 72026 | 1722 | lemma take_bit_Suc_1 [simp]: | 
| 1723 | \<open>take_bit (Suc n) 1 = 1\<close> | |
| 1724 | by (simp add: take_bit_Suc) | |
| 1725 | ||
| 71799 | 1726 | lemma take_bit_Suc_bit0 [simp]: | 
| 1727 | \<open>take_bit (Suc n) (numeral (Num.Bit0 k)) = take_bit n (numeral k) * 2\<close> | |
| 1728 | by (simp add: take_bit_Suc numeral_Bit0_div_2) | |
| 1729 | ||
| 1730 | lemma take_bit_Suc_bit1 [simp]: | |
| 1731 | \<open>take_bit (Suc n) (numeral (Num.Bit1 k)) = take_bit n (numeral k) * 2 + 1\<close> | |
| 71822 | 1732 | by (simp add: take_bit_Suc numeral_Bit1_div_2 mod_2_eq_odd) | 
| 71799 | 1733 | |
| 72026 | 1734 | lemma take_bit_numeral_1 [simp]: | 
| 1735 | \<open>take_bit (numeral l) 1 = 1\<close> | |
| 1736 | by (simp add: take_bit_rec [of \<open>numeral l\<close> 1]) | |
| 1737 | ||
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1738 | lemma take_bit_numeral_bit0 [simp]: | 
| 71799 | 1739 | \<open>take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2\<close> | 
| 1740 | by (simp add: take_bit_rec numeral_Bit0_div_2) | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1741 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1742 | lemma take_bit_numeral_bit1 [simp]: | 
| 71799 | 1743 | \<open>take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1\<close> | 
| 71822 | 1744 | by (simp add: take_bit_rec numeral_Bit1_div_2 mod_2_eq_odd) | 
| 67961 | 1745 | |
| 71799 | 1746 | lemma drop_bit_Suc_bit0 [simp]: | 
| 1747 | \<open>drop_bit (Suc n) (numeral (Num.Bit0 k)) = drop_bit n (numeral k)\<close> | |
| 1748 | by (simp add: drop_bit_Suc numeral_Bit0_div_2) | |
| 1749 | ||
| 1750 | lemma drop_bit_Suc_bit1 [simp]: | |
| 1751 | \<open>drop_bit (Suc n) (numeral (Num.Bit1 k)) = drop_bit n (numeral k)\<close> | |
| 1752 | by (simp add: drop_bit_Suc numeral_Bit1_div_2) | |
| 1753 | ||
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1754 | lemma drop_bit_numeral_bit0 [simp]: | 
| 71799 | 1755 | \<open>drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)\<close> | 
| 1756 | by (simp add: drop_bit_rec numeral_Bit0_div_2) | |
| 67816 | 1757 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1758 | lemma drop_bit_numeral_bit1 [simp]: | 
| 71799 | 1759 | \<open>drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)\<close> | 
| 1760 | by (simp add: drop_bit_rec numeral_Bit1_div_2) | |
| 67816 | 1761 | |
| 68010 | 1762 | lemma drop_bit_of_nat: | 
| 1763 | "drop_bit n (of_nat m) = of_nat (drop_bit n m)" | |
| 68389 | 1764 | by (simp add: drop_bit_eq_div Parity.drop_bit_eq_div of_nat_div [of m "2 ^ n"]) | 
| 68010 | 1765 | |
| 71412 | 1766 | lemma bit_of_nat_iff_bit [simp]: | 
| 1767 | \<open>bit (of_nat m) n \<longleftrightarrow> bit m n\<close> | |
| 1768 | proof - | |
| 1769 | have \<open>even (m div 2 ^ n) \<longleftrightarrow> even (of_nat (m div 2 ^ n))\<close> | |
| 1770 | by simp | |
| 1771 | also have \<open>of_nat (m div 2 ^ n) = of_nat m div of_nat (2 ^ n)\<close> | |
| 1772 | by (simp add: of_nat_div) | |
| 1773 | finally show ?thesis | |
| 71965 
d45f5d4c41bd
more class operations for the sake of efficient generated code
 haftmann parents: 
71958diff
changeset | 1774 | by (simp add: bit_iff_odd semiring_bits_class.bit_iff_odd) | 
| 71412 | 1775 | qed | 
| 1776 | ||
| 1777 | lemma of_nat_drop_bit: | |
| 1778 | \<open>of_nat (drop_bit m n) = drop_bit m (of_nat n)\<close> | |
| 1779 | by (simp add: drop_bit_eq_div semiring_bit_shifts_class.drop_bit_eq_div of_nat_div) | |
| 1780 | ||
| 1781 | lemma bit_push_bit_iff_of_nat_iff: | |
| 1782 | \<open>bit (push_bit m (of_nat r)) n \<longleftrightarrow> m \<le> n \<and> bit (of_nat r) (n - m)\<close> | |
| 71424 | 1783 | by (auto simp add: bit_push_bit_iff) | 
| 71412 | 1784 | |
| 58770 | 1785 | end | 
| 67816 | 1786 | |
| 71094 | 1787 | instance nat :: unique_euclidean_semiring_with_bit_shifts .. | 
| 1788 | ||
| 1789 | instance int :: unique_euclidean_semiring_with_bit_shifts .. | |
| 1790 | ||
| 72227 | 1791 | lemma bit_nat_iff: | 
| 1792 | \<open>bit (nat k) n \<longleftrightarrow> k \<ge> 0 \<and> bit k n\<close> | |
| 1793 | proof (cases \<open>k \<ge> 0\<close>) | |
| 71804 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1794 | case True | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1795 | moreover define m where \<open>m = nat k\<close> | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1796 | ultimately have \<open>k = int m\<close> | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1797 | by simp | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1798 | then show ?thesis | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1799 | by (auto intro: ccontr) | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1800 | next | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1801 | case False | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1802 | then show ?thesis | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1803 | by simp | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1804 | qed | 
| 
6fd70ed18199
simplified construction of binary bit operations
 haftmann parents: 
71802diff
changeset | 1805 | |
| 72227 | 1806 | lemma push_bit_nat_eq: | 
| 1807 | \<open>push_bit n (nat k) = nat (push_bit n k)\<close> | |
| 1808 | by (cases \<open>k \<ge> 0\<close>) (simp_all add: push_bit_eq_mult nat_mult_distrib not_le mult_nonneg_nonpos2) | |
| 1809 | ||
| 1810 | lemma drop_bit_nat_eq: | |
| 1811 | \<open>drop_bit n (nat k) = nat (drop_bit n k)\<close> | |
| 1812 | apply (cases \<open>k \<ge> 0\<close>) | |
| 1813 | apply (simp_all add: drop_bit_eq_div nat_div_distrib nat_power_eq not_le) | |
| 1814 | apply (simp add: divide_int_def) | |
| 1815 | done | |
| 1816 | ||
| 1817 | lemma take_bit_nat_eq: | |
| 1818 | \<open>take_bit n (nat k) = nat (take_bit n k)\<close> if \<open>k \<ge> 0\<close> | |
| 1819 | using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) | |
| 1820 | ||
| 72187 | 1821 | lemma nat_take_bit_eq: | 
| 1822 | \<open>nat (take_bit n k) = take_bit n (nat k)\<close> | |
| 1823 | if \<open>k \<ge> 0\<close> | |
| 1824 | using that by (simp add: take_bit_eq_mod nat_mod_distrib nat_power_eq) | |
| 1825 | ||
| 71802 | 1826 | lemma not_exp_less_eq_0_int [simp]: | 
| 1827 | \<open>\<not> 2 ^ n \<le> (0::int)\<close> | |
| 1828 | by (simp add: power_le_zero_eq) | |
| 1829 | ||
| 1830 | lemma half_nonnegative_int_iff [simp]: | |
| 1831 | \<open>k div 2 \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int | |
| 1832 | proof (cases \<open>k \<ge> 0\<close>) | |
| 1833 | case True | |
| 1834 | then show ?thesis | |
| 1835 | by (auto simp add: divide_int_def sgn_1_pos) | |
| 1836 | next | |
| 1837 | case False | |
| 1838 | then show ?thesis | |
| 1839 | apply (auto simp add: divide_int_def not_le elim!: evenE) | |
| 1840 | apply (simp only: minus_mult_right) | |
| 1841 | apply (subst nat_mult_distrib) | |
| 1842 | apply simp_all | |
| 1843 | done | |
| 1844 | qed | |
| 1845 | ||
| 1846 | lemma half_negative_int_iff [simp]: | |
| 1847 | \<open>k div 2 < 0 \<longleftrightarrow> k < 0\<close> for k :: int | |
| 1848 | by (subst Not_eq_iff [symmetric]) (simp add: not_less) | |
| 1849 | ||
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1850 | lemma push_bit_of_Suc_0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1851 | "push_bit n (Suc 0) = 2 ^ n" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1852 | using push_bit_of_1 [where ?'a = nat] by simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1853 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1854 | lemma take_bit_of_Suc_0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1855 | "take_bit n (Suc 0) = of_bool (0 < n)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1856 | using take_bit_of_1 [where ?'a = nat] by simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1857 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1858 | lemma drop_bit_of_Suc_0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1859 | "drop_bit n (Suc 0) = of_bool (n = 0)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1860 | using drop_bit_of_1 [where ?'a = nat] by simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1861 | |
| 70911 | 1862 | lemma push_bit_minus_one: | 
| 1863 | "push_bit n (- 1 :: int) = - (2 ^ n)" | |
| 1864 | by (simp add: push_bit_eq_mult) | |
| 1865 | ||
| 71195 | 1866 | lemma minus_1_div_exp_eq_int: | 
| 1867 | \<open>- 1 div (2 :: int) ^ n = - 1\<close> | |
| 1868 | by (induction n) (use div_exp_eq [symmetric, of \<open>- 1 :: int\<close> 1] in \<open>simp_all add: ac_simps\<close>) | |
| 1869 | ||
| 1870 | lemma drop_bit_minus_one [simp]: | |
| 1871 | \<open>drop_bit n (- 1 :: int) = - 1\<close> | |
| 1872 | by (simp add: drop_bit_eq_div minus_1_div_exp_eq_int) | |
| 1873 | ||
| 72023 | 1874 | lemma take_bit_Suc_from_most: | 
| 1875 | \<open>take_bit (Suc n) k = 2 ^ n * of_bool (bit k n) + take_bit n k\<close> for k :: int | |
| 1876 | by (simp only: take_bit_eq_mod power_Suc2) (simp_all add: bit_iff_odd odd_iff_mod_2_eq_one zmod_zmult2_eq) | |
| 1877 | ||
| 71424 | 1878 | lemma take_bit_minus: | 
| 71966 | 1879 | \<open>take_bit n (- take_bit n k) = take_bit n (- k)\<close> | 
| 71195 | 1880 | for k :: int | 
| 1881 | by (simp add: take_bit_eq_mod mod_minus_eq) | |
| 1882 | ||
| 71424 | 1883 | lemma take_bit_diff: | 
| 71966 | 1884 | \<open>take_bit n (take_bit n k - take_bit n l) = take_bit n (k - l)\<close> | 
| 71195 | 1885 | for k l :: int | 
| 1886 | by (simp add: take_bit_eq_mod mod_diff_eq) | |
| 1887 | ||
| 72187 | 1888 | lemma bit_imp_take_bit_positive: | 
| 1889 | \<open>0 < take_bit m k\<close> if \<open>n < m\<close> and \<open>bit k n\<close> for k :: int | |
| 1890 | proof (rule ccontr) | |
| 1891 | assume \<open>\<not> 0 < take_bit m k\<close> | |
| 1892 | then have \<open>take_bit m k = 0\<close> | |
| 1893 | by (auto simp add: not_less intro: order_antisym) | |
| 1894 | then have \<open>bit (take_bit m k) n = bit 0 n\<close> | |
| 1895 | by simp | |
| 1896 | with that show False | |
| 1897 | by (simp add: bit_take_bit_iff) | |
| 1898 | qed | |
| 1899 | ||
| 1900 | lemma take_bit_mult: | |
| 1901 | \<open>take_bit n (take_bit n k * take_bit n l) = take_bit n (k * l)\<close> | |
| 1902 | for k l :: int | |
| 1903 | by (simp add: take_bit_eq_mod mod_mult_eq) | |
| 1904 | ||
| 71966 | 1905 | lemma (in ring_1) of_nat_nat_take_bit_eq [simp]: | 
| 1906 | \<open>of_nat (nat (take_bit n k)) = of_int (take_bit n k)\<close> | |
| 1907 | by simp | |
| 1908 | ||
| 71759 | 1909 | lemma take_bit_minus_small_eq: | 
| 1910 | \<open>take_bit n (- k) = 2 ^ n - k\<close> if \<open>0 < k\<close> \<open>k \<le> 2 ^ n\<close> for k :: int | |
| 1911 | proof - | |
| 1912 | define m where \<open>m = nat k\<close> | |
| 1913 | with that have \<open>k = int m\<close> and \<open>0 < m\<close> and \<open>m \<le> 2 ^ n\<close> | |
| 1914 | by simp_all | |
| 1915 | have \<open>(2 ^ n - m) mod 2 ^ n = 2 ^ n - m\<close> | |
| 1916 | using \<open>0 < m\<close> by simp | |
| 1917 | then have \<open>int ((2 ^ n - m) mod 2 ^ n) = int (2 ^ n - m)\<close> | |
| 1918 | by simp | |
| 1919 | then have \<open>(2 ^ n - int m) mod 2 ^ n = 2 ^ n - int m\<close> | |
| 1920 | using \<open>m \<le> 2 ^ n\<close> by (simp only: of_nat_mod of_nat_diff) simp | |
| 1921 | with \<open>k = int m\<close> have \<open>(2 ^ n - k) mod 2 ^ n = 2 ^ n - k\<close> | |
| 1922 | by simp | |
| 1923 | then show ?thesis | |
| 1924 | by (simp add: take_bit_eq_mod) | |
| 1925 | qed | |
| 1926 | ||
| 71424 | 1927 | lemma drop_bit_push_bit_int: | 
| 1928 | \<open>drop_bit m (push_bit n k) = drop_bit (m - n) (push_bit (n - m) k)\<close> for k :: int | |
| 1929 | 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 | |
| 71802 | 1930 | mult.commute [of k] drop_bit_eq_div push_bit_eq_mult not_le power_add dest!: le_Suc_ex less_imp_Suc_add) | 
| 1931 | ||
| 1932 | lemma push_bit_nonnegative_int_iff [simp]: | |
| 1933 | \<open>push_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int | |
| 1934 | by (simp add: push_bit_eq_mult zero_le_mult_iff) | |
| 1935 | ||
| 1936 | lemma push_bit_negative_int_iff [simp]: | |
| 1937 | \<open>push_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int | |
| 1938 | by (subst Not_eq_iff [symmetric]) (simp add: not_less) | |
| 1939 | ||
| 1940 | lemma drop_bit_nonnegative_int_iff [simp]: | |
| 1941 | \<open>drop_bit n k \<ge> 0 \<longleftrightarrow> k \<ge> 0\<close> for k :: int | |
| 1942 | by (induction n) (simp_all add: drop_bit_Suc drop_bit_half) | |
| 1943 | ||
| 1944 | lemma drop_bit_negative_int_iff [simp]: | |
| 1945 | \<open>drop_bit n k < 0 \<longleftrightarrow> k < 0\<close> for k :: int | |
| 1946 | by (subst Not_eq_iff [symmetric]) (simp add: not_less) | |
| 71424 | 1947 | |
| 71853 | 1948 | code_identifier | 
| 1949 | code_module Parity \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith | |
| 1950 | ||
| 67816 | 1951 | end |