| author | wenzelm | 
| Sat, 17 Aug 2019 11:23:20 +0200 | |
| changeset 70557 | 5d6e9c65ea67 | 
| parent 70365 | 4df0628e8545 | 
| child 70911 | 38298c04c12e | 
| 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 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 33 | 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 | 34 | "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 | 35 | by (cases a rule: parity_cases) simp_all | 
| 66815 | 36 | |
| 37 | lemma not_mod_2_eq_1_eq_0 [simp]: | |
| 38 | "a mod 2 \<noteq> 1 \<longleftrightarrow> a mod 2 = 0" | |
| 39 | by (cases a rule: parity_cases) simp_all | |
| 40 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 41 | lemma mod2_eq_if: "a mod 2 = (if 2 dvd a then 0 else 1)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 42 | by (simp add: even_iff_mod_2_eq_zero odd_iff_mod_2_eq_one) | 
| 58787 | 43 | |
| 58690 | 44 | lemma evenE [elim?]: | 
| 45 | assumes "even a" | |
| 46 | obtains b where "a = 2 * b" | |
| 58740 | 47 | using assms by (rule dvdE) | 
| 58690 | 48 | |
| 58681 | 49 | lemma oddE [elim?]: | 
| 58680 | 50 | assumes "odd a" | 
| 51 | obtains b where "a = 2 * b + 1" | |
| 58787 | 52 | proof - | 
| 66815 | 53 | have "a = 2 * (a div 2) + a mod 2" | 
| 54 | by (simp add: mult_div_mod_eq) | |
| 55 | with assms have "a = 2 * (a div 2) + 1" | |
| 56 | by (simp add: odd_iff_mod_2_eq_one) | |
| 57 | then show ?thesis .. | |
| 58 | qed | |
| 59 | ||
| 60 | lemma mod_2_eq_odd: | |
| 61 | "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 | 62 | by (auto elim: oddE simp add: even_iff_mod_2_eq_zero) | 
| 66815 | 63 | |
| 67816 | 64 | lemma of_bool_odd_eq_mod_2: | 
| 65 | "of_bool (odd a) = a mod 2" | |
| 66 | by (simp add: mod_2_eq_odd) | |
| 67 | ||
| 66815 | 68 | lemma even_zero [simp]: | 
| 69 | "even 0" | |
| 70 | by (fact dvd_0_right) | |
| 71 | ||
| 72 | lemma odd_even_add: | |
| 73 | "even (a + b)" if "odd a" and "odd b" | |
| 74 | proof - | |
| 75 | from that obtain c d where "a = 2 * c + 1" and "b = 2 * d + 1" | |
| 76 | by (blast elim: oddE) | |
| 77 | then have "a + b = 2 * c + 2 * d + (1 + 1)" | |
| 78 | by (simp only: ac_simps) | |
| 79 | also have "\<dots> = 2 * (c + d + 1)" | |
| 80 | by (simp add: algebra_simps) | |
| 81 | finally show ?thesis .. | |
| 82 | qed | |
| 83 | ||
| 84 | lemma even_add [simp]: | |
| 85 | "even (a + b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)" | |
| 86 | by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add) | |
| 87 | ||
| 88 | lemma odd_add [simp]: | |
| 89 | "odd (a + b) \<longleftrightarrow> \<not> (odd a \<longleftrightarrow> odd b)" | |
| 90 | by simp | |
| 91 | ||
| 92 | lemma even_plus_one_iff [simp]: | |
| 93 | "even (a + 1) \<longleftrightarrow> odd a" | |
| 94 | by (auto simp add: dvd_add_right_iff intro: odd_even_add) | |
| 95 | ||
| 96 | lemma even_mult_iff [simp]: | |
| 97 | "even (a * b) \<longleftrightarrow> even a \<or> even b" (is "?P \<longleftrightarrow> ?Q") | |
| 98 | proof | |
| 99 | assume ?Q | |
| 100 | then show ?P | |
| 101 | by auto | |
| 102 | next | |
| 103 | assume ?P | |
| 104 | show ?Q | |
| 105 | proof (rule ccontr) | |
| 106 | assume "\<not> (even a \<or> even b)" | |
| 107 | then have "odd a" and "odd b" | |
| 108 | by auto | |
| 109 | then obtain r s where "a = 2 * r + 1" and "b = 2 * s + 1" | |
| 110 | by (blast elim: oddE) | |
| 111 | then have "a * b = (2 * r + 1) * (2 * s + 1)" | |
| 112 | by simp | |
| 113 | also have "\<dots> = 2 * (2 * r * s + r + s) + 1" | |
| 114 | by (simp add: algebra_simps) | |
| 115 | finally have "odd (a * b)" | |
| 116 | by simp | |
| 117 | with \<open>?P\<close> show False | |
| 118 | by auto | |
| 119 | qed | |
| 120 | qed | |
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 121 | |
| 63654 | 122 | lemma even_numeral [simp]: "even (numeral (Num.Bit0 n))" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 123 | proof - | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 124 | have "even (2 * numeral n)" | 
| 66815 | 125 | unfolding even_mult_iff by simp | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 126 | then have "even (numeral n + numeral n)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 127 | unfolding mult_2 . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 128 | then show ?thesis | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 129 | unfolding numeral.simps . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 130 | qed | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 131 | |
| 63654 | 132 | lemma odd_numeral [simp]: "odd (numeral (Num.Bit1 n))" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 133 | proof | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 134 | assume "even (numeral (num.Bit1 n))" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 135 | then have "even (numeral n + numeral n + 1)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 136 | unfolding numeral.simps . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 137 | then have "even (2 * numeral n + 1)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 138 | unfolding mult_2 . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 139 | then have "2 dvd numeral n * 2 + 1" | 
| 58740 | 140 | by (simp add: ac_simps) | 
| 63654 | 141 | then have "2 dvd 1" | 
| 142 | 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 | 143 | then show False by simp | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 144 | qed | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 145 | |
| 63654 | 146 | lemma even_power [simp]: "even (a ^ n) \<longleftrightarrow> even a \<and> n > 0" | 
| 58680 | 147 | by (induct n) auto | 
| 148 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 149 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 150 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 151 | 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 | 152 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 153 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 154 | subclass comm_ring_1 .. | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 155 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 156 | lemma even_minus: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 157 | "even (- a) \<longleftrightarrow> even a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 158 | by (fact dvd_minus_iff) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 159 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 160 | lemma even_diff [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 161 | "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 | 162 | 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 | 163 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 164 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 165 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 166 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 167 | 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 | 168 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 169 | class unique_euclidean_semiring_with_nat = semidom + semiring_char_0 + unique_euclidean_semiring + | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 170 | assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 171 | and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 172 | and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 173 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 174 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 175 | lemma division_segment_eq_iff: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 176 | "a = b" if "division_segment a = division_segment b" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 177 | and "euclidean_size a = euclidean_size b" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 178 | using that division_segment_euclidean_size [of a] by simp | 
| 
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 | lemma euclidean_size_of_nat [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 181 | "euclidean_size (of_nat n) = n" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 182 | proof - | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 183 | have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 184 | by (fact division_segment_euclidean_size) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 185 | then show ?thesis by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 186 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 187 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 188 | lemma of_nat_euclidean_size: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 189 | "of_nat (euclidean_size a) = a div division_segment a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 190 | proof - | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 191 | have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 192 | by (subst nonzero_mult_div_cancel_left) simp_all | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 193 | also have "\<dots> = a div division_segment a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 194 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 195 | finally show ?thesis . | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 196 | qed | 
| 
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 | lemma division_segment_1 [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 199 | "division_segment 1 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 200 | using division_segment_of_nat [of 1] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 201 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 202 | lemma division_segment_numeral [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 203 | "division_segment (numeral k) = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 204 | using division_segment_of_nat [of "numeral k"] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 205 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 206 | lemma euclidean_size_1 [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 207 | "euclidean_size 1 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 208 | using euclidean_size_of_nat [of 1] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 209 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 210 | lemma euclidean_size_numeral [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 211 | "euclidean_size (numeral k) = numeral k" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 212 | using euclidean_size_of_nat [of "numeral k"] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 213 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 214 | lemma of_nat_dvd_iff: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 215 | "of_nat m dvd of_nat n \<longleftrightarrow> m dvd n" (is "?P \<longleftrightarrow> ?Q") | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 216 | proof (cases "m = 0") | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 217 | case True | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 218 | then show ?thesis | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 219 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 220 | next | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 221 | case False | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 222 | show ?thesis | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 223 | proof | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 224 | assume ?Q | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 225 | then show ?P | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 226 | by auto | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 227 | next | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 228 | assume ?P | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 229 | with False have "of_nat n = of_nat n div of_nat m * of_nat m" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 230 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 231 | then have "of_nat n = of_nat (n div m * m)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 232 | by (simp add: of_nat_div) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 233 | then have "n = n div m * m" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 234 | by (simp only: of_nat_eq_iff) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 235 | then have "n = m * (n div m)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 236 | by (simp add: ac_simps) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 237 | then show ?Q .. | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 238 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 239 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 240 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 241 | lemma of_nat_mod: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 242 | "of_nat (m mod n) = of_nat m mod of_nat n" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 243 | proof - | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 244 | have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 245 | by (simp add: div_mult_mod_eq) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 246 | also have "of_nat m = of_nat (m div n * n + m mod n)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 247 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 248 | finally show ?thesis | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 249 | by (simp only: of_nat_div of_nat_mult of_nat_add) simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 250 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 251 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 252 | lemma one_div_two_eq_zero [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 253 | "1 div 2 = 0" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 254 | proof - | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 255 | from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 256 | by (simp only:) simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 257 | then show ?thesis | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 258 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 259 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 260 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 261 | lemma one_mod_two_eq_one [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 262 | "1 mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 263 | proof - | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 264 | from of_nat_mod [symmetric] have "of_nat 1 mod of_nat 2 = of_nat 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 265 | by (simp only:) simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 266 | then show ?thesis | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 267 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 268 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 269 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 270 | subclass semiring_parity | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 271 | proof | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 272 | 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 | 273 | 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 | 274 | 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 | 275 | proof | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 276 | assume "a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 277 | 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 | 278 | by auto | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 279 | next | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 280 | assume "\<not> 2 dvd a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 281 | 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 | 282 | proof (rule order_antisym) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 283 | 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 | 284 | 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 | 285 | 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 | 286 | 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 | 287 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 288 | 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 | 289 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 290 | 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 | 291 | 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 | 292 | 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 | 293 | 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 | 294 | 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 | 295 | 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 | 296 | 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 | 297 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 298 | 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 | 299 | 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 | 300 | 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 | 301 | show "a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 302 | 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 | 303 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 304 | show "\<not> is_unit 2" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 305 | proof (rule notI) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 306 | assume "is_unit 2" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 307 | 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 | 308 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 309 | 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 | 310 | 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 | 311 | then show False | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 312 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 313 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 314 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 315 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 316 | lemma even_of_nat [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 317 | "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 | 318 | proof - | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 319 | 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 | 320 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 321 | 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 | 322 | 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 | 323 | finally show ?thesis . | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 324 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 325 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 326 | lemma one_mod_2_pow_eq [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 327 | "1 mod (2 ^ n) = of_bool (n > 0)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 328 | proof - | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 329 | have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 330 | using of_nat_mod [of 1 "2 ^ n"] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 331 | also have "\<dots> = of_bool (n > 0)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 332 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 333 | finally show ?thesis . | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 334 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 335 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 336 | lemma one_div_2_pow_eq [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 337 | "1 div (2 ^ n) = of_bool (n = 0)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 338 | using div_mult_mod_eq [of 1 "2 ^ n"] by auto | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 339 | |
| 66815 | 340 | lemma even_succ_div_two [simp]: | 
| 341 | "even a \<Longrightarrow> (a + 1) div 2 = a div 2" | |
| 342 | by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero) | |
| 343 | ||
| 344 | lemma odd_succ_div_two [simp]: | |
| 345 | "odd a \<Longrightarrow> (a + 1) div 2 = a div 2 + 1" | |
| 346 | by (auto elim!: oddE simp add: add.assoc) | |
| 347 | ||
| 348 | lemma even_two_times_div_two: | |
| 349 | "even a \<Longrightarrow> 2 * (a div 2) = a" | |
| 350 | by (fact dvd_mult_div_cancel) | |
| 351 | ||
| 352 | lemma odd_two_times_div_two_succ [simp]: | |
| 353 | "odd a \<Longrightarrow> 2 * (a div 2) + 1 = a" | |
| 354 | using mult_div_mod_eq [of 2 a] | |
| 355 | by (simp add: even_iff_mod_2_eq_zero) | |
| 356 | ||
| 67051 | 357 | lemma coprime_left_2_iff_odd [simp]: | 
| 358 | "coprime 2 a \<longleftrightarrow> odd a" | |
| 359 | proof | |
| 360 | assume "odd a" | |
| 361 | show "coprime 2 a" | |
| 362 | proof (rule coprimeI) | |
| 363 | fix b | |
| 364 | assume "b dvd 2" "b dvd a" | |
| 365 | then have "b dvd a mod 2" | |
| 366 | by (auto intro: dvd_mod) | |
| 367 | with \<open>odd a\<close> show "is_unit b" | |
| 368 | by (simp add: mod_2_eq_odd) | |
| 369 | qed | |
| 370 | next | |
| 371 | assume "coprime 2 a" | |
| 372 | show "odd a" | |
| 373 | proof (rule notI) | |
| 374 | assume "even a" | |
| 375 | then obtain b where "a = 2 * b" .. | |
| 376 | with \<open>coprime 2 a\<close> have "coprime 2 (2 * b)" | |
| 377 | by simp | |
| 378 | moreover have "\<not> coprime 2 (2 * b)" | |
| 379 | by (rule not_coprimeI [of 2]) simp_all | |
| 380 | ultimately show False | |
| 381 | by blast | |
| 382 | qed | |
| 383 | qed | |
| 384 | ||
| 385 | lemma coprime_right_2_iff_odd [simp]: | |
| 386 | "coprime a 2 \<longleftrightarrow> odd a" | |
| 387 | using coprime_left_2_iff_odd [of a] by (simp add: ac_simps) | |
| 388 | ||
| 67828 | 389 | lemma div_mult2_eq': | 
| 390 | "a div (of_nat m * of_nat n) = a div of_nat m div of_nat n" | |
| 391 | proof (cases a "of_nat m * of_nat n" rule: divmod_cases) | |
| 392 | case (divides q) | |
| 393 | then show ?thesis | |
| 394 | using nonzero_mult_div_cancel_right [of "of_nat m" "q * of_nat n"] | |
| 395 | by (simp add: ac_simps) | |
| 396 | next | |
| 397 | case (remainder q r) | |
| 398 | then have "division_segment r = 1" | |
| 399 | using division_segment_of_nat [of "m * n"] by simp | |
| 400 | with division_segment_euclidean_size [of r] | |
| 401 | have "of_nat (euclidean_size r) = r" | |
| 402 | by simp | |
| 67908 | 403 | have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0" | 
| 404 | by simp | |
| 405 | with remainder(6) have "r div (of_nat m * of_nat n) = 0" | |
| 67828 | 406 | by simp | 
| 67908 | 407 | with \<open>of_nat (euclidean_size r) = r\<close> | 
| 408 | have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0" | |
| 409 | by simp | |
| 410 | then have "of_nat (euclidean_size r div (m * n)) = 0" | |
| 67828 | 411 | by (simp add: of_nat_div) | 
| 67908 | 412 | then have "of_nat (euclidean_size r div m div n) = 0" | 
| 413 | by (simp add: div_mult2_eq) | |
| 414 | with \<open>of_nat (euclidean_size r) = r\<close> have "r div of_nat m div of_nat n = 0" | |
| 415 | by (simp add: of_nat_div) | |
| 67828 | 416 | with remainder(1) | 
| 417 | have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n" | |
| 418 | by simp | |
| 67908 | 419 | with remainder(5) remainder(7) show ?thesis | 
| 67828 | 420 | using div_plus_div_distrib_dvd_right [of "of_nat m" "q * (of_nat m * of_nat n)" r] | 
| 421 | by (simp add: ac_simps) | |
| 422 | next | |
| 423 | case by0 | |
| 424 | then show ?thesis | |
| 425 | by auto | |
| 426 | qed | |
| 427 | ||
| 428 | lemma mod_mult2_eq': | |
| 429 | "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m" | |
| 430 | proof - | |
| 431 | have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)" | |
| 432 | by (simp add: combine_common_factor div_mult_mod_eq) | |
| 433 | moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)" | |
| 434 | by (simp add: ac_simps) | |
| 435 | ultimately show ?thesis | |
| 436 | by (simp add: div_mult2_eq' mult_commute) | |
| 437 | qed | |
| 438 | ||
| 68028 | 439 | lemma div_mult2_numeral_eq: | 
| 440 | "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B") | |
| 441 | proof - | |
| 442 | have "?A = a div of_nat (numeral k) div of_nat (numeral l)" | |
| 443 | by simp | |
| 444 | also have "\<dots> = a div (of_nat (numeral k) * of_nat (numeral l))" | |
| 445 | by (fact div_mult2_eq' [symmetric]) | |
| 446 | also have "\<dots> = ?B" | |
| 447 | by simp | |
| 448 | finally show ?thesis . | |
| 449 | qed | |
| 450 | ||
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 451 | end | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 452 | |
| 70340 | 453 | class unique_euclidean_ring_with_nat = ring + unique_euclidean_semiring_with_nat | 
| 58679 | 454 | begin | 
| 455 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 456 | subclass ring_parity .. | 
| 58680 | 457 | |
| 67906 | 458 | lemma minus_1_mod_2_eq [simp]: | 
| 459 | "- 1 mod 2 = 1" | |
| 460 | by (simp add: mod_2_eq_odd) | |
| 461 | ||
| 462 | lemma minus_1_div_2_eq [simp]: | |
| 463 | "- 1 div 2 = - 1" | |
| 464 | proof - | |
| 465 | from div_mult_mod_eq [of "- 1" 2] | |
| 466 | 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 | 467 | using add_implies_diff by fastforce | 
| 67906 | 468 | then show ?thesis | 
| 469 | using mult_right_cancel [of 2 "- 1 div 2" "- 1"] by simp | |
| 470 | qed | |
| 471 | ||
| 58679 | 472 | end | 
| 473 | ||
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 474 | |
| 69593 | 475 | subsection \<open>Instance for \<^typ>\<open>nat\<close>\<close> | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 476 | |
| 70340 | 477 | instance nat :: unique_euclidean_semiring_with_nat | 
| 66815 | 478 | by standard (simp_all add: dvd_eq_mod_eq_0) | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 479 | |
| 66815 | 480 | lemma even_Suc_Suc_iff [simp]: | 
| 481 | "even (Suc (Suc n)) \<longleftrightarrow> even n" | |
| 58787 | 482 | using dvd_add_triv_right_iff [of 2 n] by simp | 
| 58687 | 483 | |
| 66815 | 484 | lemma even_Suc [simp]: "even (Suc n) \<longleftrightarrow> odd n" | 
| 485 | using even_plus_one_iff [of n] by simp | |
| 58787 | 486 | |
| 66815 | 487 | lemma even_diff_nat [simp]: | 
| 488 | "even (m - n) \<longleftrightarrow> m < n \<or> even (m + n)" for m n :: nat | |
| 58787 | 489 | proof (cases "n \<le> m") | 
| 490 | case True | |
| 491 | then have "m - n + n * 2 = m + n" by (simp add: mult_2_right) | |
| 66815 | 492 | moreover have "even (m - n) \<longleftrightarrow> even (m - n + n * 2)" by simp | 
| 493 | ultimately have "even (m - n) \<longleftrightarrow> even (m + n)" by (simp only:) | |
| 58787 | 494 | then show ?thesis by auto | 
| 495 | next | |
| 496 | case False | |
| 497 | then show ?thesis by simp | |
| 63654 | 498 | qed | 
| 499 | ||
| 66815 | 500 | lemma odd_pos: | 
| 501 | "odd n \<Longrightarrow> 0 < n" for n :: nat | |
| 58690 | 502 | by (auto elim: oddE) | 
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 503 | |
| 66815 | 504 | lemma Suc_double_not_eq_double: | 
| 505 | "Suc (2 * m) \<noteq> 2 * n" | |
| 62597 | 506 | proof | 
| 507 | assume "Suc (2 * m) = 2 * n" | |
| 508 | moreover have "odd (Suc (2 * m))" and "even (2 * n)" | |
| 509 | by simp_all | |
| 510 | ultimately show False by simp | |
| 511 | qed | |
| 512 | ||
| 66815 | 513 | lemma double_not_eq_Suc_double: | 
| 514 | "2 * m \<noteq> Suc (2 * n)" | |
| 62597 | 515 | using Suc_double_not_eq_double [of n m] by simp | 
| 516 | ||
| 66815 | 517 | lemma odd_Suc_minus_one [simp]: "odd n \<Longrightarrow> Suc (n - Suc 0) = n" | 
| 518 | by (auto elim: oddE) | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 519 | |
| 66815 | 520 | lemma even_Suc_div_two [simp]: | 
| 521 | "even n \<Longrightarrow> Suc n div 2 = n div 2" | |
| 522 | using even_succ_div_two [of n] by simp | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 523 | |
| 66815 | 524 | lemma odd_Suc_div_two [simp]: | 
| 525 | "odd n \<Longrightarrow> Suc n div 2 = Suc (n div 2)" | |
| 526 | using odd_succ_div_two [of n] by simp | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 527 | |
| 66815 | 528 | lemma odd_two_times_div_two_nat [simp]: | 
| 529 | assumes "odd n" | |
| 530 | shows "2 * (n div 2) = n - (1 :: nat)" | |
| 531 | proof - | |
| 532 | from assms have "2 * (n div 2) + 1 = n" | |
| 533 | by (rule odd_two_times_div_two_succ) | |
| 534 | then have "Suc (2 * (n div 2)) - 1 = n - 1" | |
| 58787 | 535 | by simp | 
| 66815 | 536 | then show ?thesis | 
| 537 | by simp | |
| 58787 | 538 | qed | 
| 58680 | 539 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 540 | 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 | 541 | "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 | 542 | 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 | 543 | |
| 70353 | 544 | lemma nat_bit_induct [case_names zero even odd]: | 
| 70226 | 545 | "P n" if zero: "P 0" | 
| 546 | and even: "\<And>n. P n \<Longrightarrow> n > 0 \<Longrightarrow> P (2 * n)" | |
| 547 | and odd: "\<And>n. P n \<Longrightarrow> P (Suc (2 * n))" | |
| 548 | proof (induction n rule: less_induct) | |
| 66815 | 549 | case (less n) | 
| 550 | show "P n" | |
| 551 | proof (cases "n = 0") | |
| 552 | case True with zero show ?thesis by simp | |
| 553 | next | |
| 554 | case False | |
| 555 | with less have hyp: "P (n div 2)" by simp | |
| 556 | show ?thesis | |
| 557 | proof (cases "even n") | |
| 558 | case True | |
| 70226 | 559 | then have "n \<noteq> 1" | 
| 560 | by auto | |
| 561 | with \<open>n \<noteq> 0\<close> have "n div 2 > 0" | |
| 562 | by simp | |
| 563 | with \<open>even n\<close> hyp even [of "n div 2"] show ?thesis | |
| 66815 | 564 | by simp | 
| 565 | next | |
| 566 | case False | |
| 567 | with hyp odd [of "n div 2"] show ?thesis | |
| 568 | by simp | |
| 569 | qed | |
| 570 | qed | |
| 571 | qed | |
| 58687 | 572 | |
| 69502 | 573 | lemma odd_card_imp_not_empty: | 
| 574 |   \<open>A \<noteq> {}\<close> if \<open>odd (card A)\<close>
 | |
| 575 | using that by auto | |
| 576 | ||
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 577 | 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 | 578 | 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 | 579 | shows "P n" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 580 | proof (induct n rule: less_induct) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 581 | case (less n) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 582 | show ?case | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 583 | proof (cases "n < Suc (Suc 0)") | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 584 | case True | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 585 | then show ?thesis | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 586 | 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 | 587 | next | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 588 | case False | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 589 | 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 | 590 | 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 | 591 | then have "k<n" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 592 | by simp | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 593 | 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 | 594 | by blast | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 595 | then show ?thesis | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 596 | by (simp add: k) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 597 | qed | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 598 | qed | 
| 58687 | 599 | |
| 60758 | 600 | subsection \<open>Parity and powers\<close> | 
| 58689 | 601 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
60867diff
changeset | 602 | context ring_1 | 
| 58689 | 603 | begin | 
| 604 | ||
| 63654 | 605 | lemma power_minus_even [simp]: "even n \<Longrightarrow> (- a) ^ n = a ^ n" | 
| 58690 | 606 | by (auto elim: evenE) | 
| 58689 | 607 | |
| 63654 | 608 | lemma power_minus_odd [simp]: "odd n \<Longrightarrow> (- a) ^ n = - (a ^ n)" | 
| 58690 | 609 | by (auto elim: oddE) | 
| 610 | ||
| 66815 | 611 | lemma uminus_power_if: | 
| 612 | "(- a) ^ n = (if even n then a ^ n else - (a ^ n))" | |
| 613 | by auto | |
| 614 | ||
| 63654 | 615 | lemma neg_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1" | 
| 58690 | 616 | by simp | 
| 58689 | 617 | |
| 63654 | 618 | lemma neg_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1" | 
| 58690 | 619 | by simp | 
| 58689 | 620 | |
| 66582 | 621 | lemma neg_one_power_add_eq_neg_one_power_diff: "k \<le> n \<Longrightarrow> (- 1) ^ (n + k) = (- 1) ^ (n - k)" | 
| 622 | by (cases "even (n + k)") auto | |
| 623 | ||
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 624 | 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 | 625 | by (induct n) auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 626 | |
| 63654 | 627 | end | 
| 58689 | 628 | |
| 629 | context linordered_idom | |
| 630 | begin | |
| 631 | ||
| 63654 | 632 | lemma zero_le_even_power: "even n \<Longrightarrow> 0 \<le> a ^ n" | 
| 58690 | 633 | by (auto elim: evenE) | 
| 58689 | 634 | |
| 63654 | 635 | lemma zero_le_odd_power: "odd n \<Longrightarrow> 0 \<le> a ^ n \<longleftrightarrow> 0 \<le> a" | 
| 58689 | 636 | by (auto simp add: power_even_eq zero_le_mult_iff elim: oddE) | 
| 637 | ||
| 63654 | 638 | lemma zero_le_power_eq: "0 \<le> a ^ n \<longleftrightarrow> even n \<or> odd n \<and> 0 \<le> a" | 
| 58787 | 639 | by (auto simp add: zero_le_even_power zero_le_odd_power) | 
| 63654 | 640 | |
| 641 | 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 | 642 | proof - | 
| 643 | have [simp]: "0 = a ^ n \<longleftrightarrow> a = 0 \<and> n > 0" | |
| 58787 | 644 | unfolding power_eq_0_iff [of a n, symmetric] by blast | 
| 58689 | 645 | show ?thesis | 
| 63654 | 646 | unfolding less_le zero_le_power_eq by auto | 
| 58689 | 647 | qed | 
| 648 | ||
| 63654 | 649 | lemma power_less_zero_eq [simp]: "a ^ n < 0 \<longleftrightarrow> odd n \<and> a < 0" | 
| 58689 | 650 | unfolding not_le [symmetric] zero_le_power_eq by auto | 
| 651 | ||
| 63654 | 652 | 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)" | 
| 653 | unfolding not_less [symmetric] zero_less_power_eq by auto | |
| 654 | ||
| 655 | lemma power_even_abs: "even n \<Longrightarrow> \<bar>a\<bar> ^ n = a ^ n" | |
| 58689 | 656 | using power_abs [of a n] by (simp add: zero_le_even_power) | 
| 657 | ||
| 658 | lemma power_mono_even: | |
| 659 | assumes "even n" and "\<bar>a\<bar> \<le> \<bar>b\<bar>" | |
| 660 | shows "a ^ n \<le> b ^ n" | |
| 661 | proof - | |
| 662 | have "0 \<le> \<bar>a\<bar>" by auto | |
| 63654 | 663 | with \<open>\<bar>a\<bar> \<le> \<bar>b\<bar>\<close> have "\<bar>a\<bar> ^ n \<le> \<bar>b\<bar> ^ n" | 
| 664 | by (rule power_mono) | |
| 665 | with \<open>even n\<close> show ?thesis | |
| 666 | by (simp add: power_even_abs) | |
| 58689 | 667 | qed | 
| 668 | ||
| 669 | lemma power_mono_odd: | |
| 670 | assumes "odd n" and "a \<le> b" | |
| 671 | shows "a ^ n \<le> b ^ n" | |
| 672 | proof (cases "b < 0") | |
| 63654 | 673 | case True | 
| 674 | with \<open>a \<le> b\<close> have "- b \<le> - a" and "0 \<le> - b" by auto | |
| 675 | then have "(- b) ^ n \<le> (- a) ^ n" by (rule power_mono) | |
| 60758 | 676 | with \<open>odd n\<close> show ?thesis by simp | 
| 58689 | 677 | next | 
| 63654 | 678 | case False | 
| 679 | then have "0 \<le> b" by auto | |
| 58689 | 680 | show ?thesis | 
| 681 | proof (cases "a < 0") | |
| 63654 | 682 | case True | 
| 683 | then have "n \<noteq> 0" and "a \<le> 0" using \<open>odd n\<close> [THEN odd_pos] by auto | |
| 60758 | 684 | then have "a ^ n \<le> 0" unfolding power_le_zero_eq using \<open>odd n\<close> by auto | 
| 63654 | 685 | moreover from \<open>0 \<le> b\<close> have "0 \<le> b ^ n" by auto | 
| 58689 | 686 | ultimately show ?thesis by auto | 
| 687 | next | |
| 63654 | 688 | case False | 
| 689 | then have "0 \<le> a" by auto | |
| 690 | with \<open>a \<le> b\<close> show ?thesis | |
| 691 | using power_mono by auto | |
| 58689 | 692 | qed | 
| 693 | qed | |
| 62083 | 694 | |
| 60758 | 695 | text \<open>Simplify, when the exponent is a numeral\<close> | 
| 58689 | 696 | |
| 697 | lemma zero_le_power_eq_numeral [simp]: | |
| 698 | "0 \<le> a ^ numeral w \<longleftrightarrow> even (numeral w :: nat) \<or> odd (numeral w :: nat) \<and> 0 \<le> a" | |
| 699 | by (fact zero_le_power_eq) | |
| 700 | ||
| 701 | lemma zero_less_power_eq_numeral [simp]: | |
| 63654 | 702 | "0 < a ^ numeral w \<longleftrightarrow> | 
| 703 | numeral w = (0 :: nat) \<or> | |
| 704 | even (numeral w :: nat) \<and> a \<noteq> 0 \<or> | |
| 705 | odd (numeral w :: nat) \<and> 0 < a" | |
| 58689 | 706 | by (fact zero_less_power_eq) | 
| 707 | ||
| 708 | lemma power_le_zero_eq_numeral [simp]: | |
| 63654 | 709 | "a ^ numeral w \<le> 0 \<longleftrightarrow> | 
| 710 | (0 :: nat) < numeral w \<and> | |
| 711 | (odd (numeral w :: nat) \<and> a \<le> 0 \<or> even (numeral w :: nat) \<and> a = 0)" | |
| 58689 | 712 | by (fact power_le_zero_eq) | 
| 713 | ||
| 714 | lemma power_less_zero_eq_numeral [simp]: | |
| 715 | "a ^ numeral w < 0 \<longleftrightarrow> odd (numeral w :: nat) \<and> a < 0" | |
| 716 | by (fact power_less_zero_eq) | |
| 717 | ||
| 718 | lemma power_even_abs_numeral [simp]: | |
| 719 | "even (numeral w :: nat) \<Longrightarrow> \<bar>a\<bar> ^ numeral w = a ^ numeral w" | |
| 720 | by (fact power_even_abs) | |
| 721 | ||
| 722 | end | |
| 723 | ||
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 724 | |
| 69593 | 725 | subsection \<open>Instance for \<^typ>\<open>int\<close>\<close> | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 726 | |
| 70340 | 727 | instance int :: unique_euclidean_ring_with_nat | 
| 66839 | 728 | by standard (simp_all add: dvd_eq_mod_eq_0 divide_int_def division_segment_int_def) | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 729 | |
| 67816 | 730 | lemma even_diff_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 731 | "even (k - l) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 732 | by (fact even_diff) | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 733 | |
| 67816 | 734 | lemma even_abs_add_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 735 | "even (\<bar>k\<bar> + l) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 736 | by simp | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 737 | |
| 67816 | 738 | lemma even_add_abs_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 739 | "even (k + \<bar>l\<bar>) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 740 | by simp | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 741 | |
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 742 | 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 | 743 | 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 | 744 | |
| 70353 | 745 | lemma int_bit_induct [case_names zero minus even odd]: | 
| 70338 | 746 | "P k" if zero_int: "P 0" | 
| 747 | and minus_int: "P (- 1)" | |
| 748 | and even_int: "\<And>k. P k \<Longrightarrow> k \<noteq> 0 \<Longrightarrow> P (k * 2)" | |
| 749 | and odd_int: "\<And>k. P k \<Longrightarrow> k \<noteq> - 1 \<Longrightarrow> P (1 + (k * 2))" for k :: int | |
| 750 | proof (cases "k \<ge> 0") | |
| 751 | case True | |
| 752 | define n where "n = nat k" | |
| 753 | with True have "k = int n" | |
| 754 | by simp | |
| 755 | then show "P k" | |
| 70353 | 756 | proof (induction n arbitrary: k rule: nat_bit_induct) | 
| 70338 | 757 | case zero | 
| 758 | then show ?case | |
| 759 | by (simp add: zero_int) | |
| 760 | next | |
| 761 | case (even n) | |
| 762 | have "P (int n * 2)" | |
| 763 | by (rule even_int) (use even in simp_all) | |
| 764 | with even show ?case | |
| 765 | by (simp add: ac_simps) | |
| 766 | next | |
| 767 | case (odd n) | |
| 768 | have "P (1 + (int n * 2))" | |
| 769 | by (rule odd_int) (use odd in simp_all) | |
| 770 | with odd show ?case | |
| 771 | by (simp add: ac_simps) | |
| 772 | qed | |
| 773 | next | |
| 774 | case False | |
| 775 | define n where "n = nat (- k - 1)" | |
| 776 | with False have "k = - int n - 1" | |
| 777 | by simp | |
| 778 | then show "P k" | |
| 70353 | 779 | proof (induction n arbitrary: k rule: nat_bit_induct) | 
| 70338 | 780 | case zero | 
| 781 | then show ?case | |
| 782 | by (simp add: minus_int) | |
| 783 | next | |
| 784 | case (even n) | |
| 785 | have "P (1 + (- int (Suc n) * 2))" | |
| 786 | by (rule odd_int) (use even in \<open>simp_all add: algebra_simps\<close>) | |
| 787 | also have "\<dots> = - int (2 * n) - 1" | |
| 788 | by (simp add: algebra_simps) | |
| 789 | finally show ?case | |
| 790 | using even by simp | |
| 791 | next | |
| 792 | case (odd n) | |
| 793 | have "P (- int (Suc n) * 2)" | |
| 794 | by (rule even_int) (use odd in \<open>simp_all add: algebra_simps\<close>) | |
| 795 | also have "\<dots> = - int (Suc (2 * n)) - 1" | |
| 796 | by (simp add: algebra_simps) | |
| 797 | finally show ?case | |
| 798 | using odd by simp | |
| 799 | qed | |
| 800 | qed | |
| 801 | ||
| 67816 | 802 | |
| 67828 | 803 | subsection \<open>Abstract bit operations\<close> | 
| 804 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 805 | context semiring_parity | 
| 67816 | 806 | begin | 
| 807 | ||
| 808 | text \<open>The primary purpose of the following operations is | |
| 69593 | 809 | to avoid ad-hoc simplification of concrete expressions \<^term>\<open>2 ^ n\<close>\<close> | 
| 67816 | 810 | |
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 811 | definition push_bit :: "nat \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 812 | where push_bit_eq_mult: "push_bit n a = a * 2 ^ n" | 
| 67816 | 813 | |
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 814 | definition take_bit :: "nat \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 68010 | 815 | where take_bit_eq_mod: "take_bit n a = a mod 2 ^ n" | 
| 67816 | 816 | |
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 817 | definition drop_bit :: "nat \<Rightarrow> 'a \<Rightarrow> 'a" | 
| 68010 | 818 | where drop_bit_eq_div: "drop_bit n a = a div 2 ^ n" | 
| 67816 | 819 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 820 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 821 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 822 | context unique_euclidean_semiring_with_nat | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 823 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 824 | |
| 67816 | 825 | lemma bit_ident: | 
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 826 | "push_bit n (drop_bit n a) + take_bit n a = a" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 827 | using div_mult_mod_eq by (simp add: push_bit_eq_mult take_bit_eq_mod drop_bit_eq_div) | 
| 67816 | 828 | |
| 67960 | 829 | lemma push_bit_push_bit [simp]: | 
| 830 | "push_bit m (push_bit n a) = push_bit (m + n) a" | |
| 831 | by (simp add: push_bit_eq_mult power_add ac_simps) | |
| 832 | ||
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 833 | lemma take_bit_take_bit [simp]: | 
| 67960 | 834 | "take_bit m (take_bit n a) = take_bit (min m n) a" | 
| 835 | proof (cases "m \<le> n") | |
| 836 | case True | |
| 837 | then show ?thesis | |
| 838 | by (simp add: take_bit_eq_mod not_le min_def mod_mod_cancel le_imp_power_dvd) | |
| 839 | next | |
| 840 | case False | |
| 841 | then have "n < m" and "min m n = n" | |
| 842 | by simp_all | |
| 843 | then have "2 ^ m = of_nat (2 ^ n) * of_nat (2 ^ (m - n))" | |
| 844 | by (simp add: power_add [symmetric]) | |
| 845 | then have "a mod 2 ^ n mod 2 ^ m = a mod 2 ^ n mod (of_nat (2 ^ n) * of_nat (2 ^ (m - n)))" | |
| 846 | by simp | |
| 847 | also have "\<dots> = of_nat (2 ^ n) * (a mod 2 ^ n div of_nat (2 ^ n) mod of_nat (2 ^ (m - n))) + a mod 2 ^ n mod of_nat (2 ^ n)" | |
| 848 | by (simp only: mod_mult2_eq') | |
| 849 | finally show ?thesis | |
| 850 | using \<open>min m n = n\<close> by (simp add: take_bit_eq_mod) | |
| 851 | qed | |
| 852 | ||
| 853 | lemma drop_bit_drop_bit [simp]: | |
| 854 | "drop_bit m (drop_bit n a) = drop_bit (m + n) a" | |
| 855 | proof - | |
| 856 | have "a div (2 ^ m * 2 ^ n) = a div (of_nat (2 ^ n) * of_nat (2 ^ m))" | |
| 857 | by (simp add: ac_simps) | |
| 858 | also have "\<dots> = a div of_nat (2 ^ n) div of_nat (2 ^ m)" | |
| 859 | by (simp only: div_mult2_eq') | |
| 860 | finally show ?thesis | |
| 861 | by (simp add: drop_bit_eq_div power_add) | |
| 862 | qed | |
| 863 | ||
| 864 | lemma push_bit_take_bit: | |
| 865 | "push_bit m (take_bit n a) = take_bit (m + n) (push_bit m a)" | |
| 866 | by (simp add: push_bit_eq_mult take_bit_eq_mod power_add mult_mod_right ac_simps) | |
| 867 | ||
| 868 | lemma take_bit_push_bit: | |
| 869 | "take_bit m (push_bit n a) = push_bit n (take_bit (m - n) a)" | |
| 870 | proof (cases "m \<le> n") | |
| 871 | case True | |
| 872 | then show ?thesis | |
| 873 | by (simp_all add: push_bit_eq_mult take_bit_eq_mod mod_eq_0_iff_dvd dvd_power_le) | |
| 874 | next | |
| 875 | case False | |
| 876 | then show ?thesis | |
| 877 | using push_bit_take_bit [of n "m - n" a] | |
| 878 | by simp | |
| 879 | qed | |
| 880 | ||
| 881 | lemma take_bit_drop_bit: | |
| 882 | "take_bit m (drop_bit n a) = drop_bit n (take_bit (m + n) a)" | |
| 883 | using mod_mult2_eq' [of a "2 ^ n" "2 ^ m"] | |
| 884 | by (simp add: drop_bit_eq_div take_bit_eq_mod power_add ac_simps) | |
| 885 | ||
| 886 | lemma drop_bit_take_bit: | |
| 887 | "drop_bit m (take_bit n a) = take_bit (n - m) (drop_bit m a)" | |
| 888 | proof (cases "m \<le> n") | |
| 889 | case True | |
| 890 | then show ?thesis | |
| 891 | using take_bit_drop_bit [of "n - m" m a] by simp | |
| 892 | next | |
| 893 | case False | |
| 894 | then have "a mod 2 ^ n div 2 ^ m = a mod 2 ^ n div 2 ^ (n + (m - n))" | |
| 895 | by simp | |
| 896 | also have "\<dots> = a mod 2 ^ n div (2 ^ n * 2 ^ (m - n))" | |
| 897 | by (simp add: power_add) | |
| 898 | also have "\<dots> = a mod 2 ^ n div (of_nat (2 ^ n) * of_nat (2 ^ (m - n)))" | |
| 899 | by simp | |
| 900 | also have "\<dots> = a mod 2 ^ n div of_nat (2 ^ n) div of_nat (2 ^ (m - n))" | |
| 901 | by (simp only: div_mult2_eq') | |
| 902 | finally show ?thesis | |
| 903 | using False by (simp add: take_bit_eq_mod drop_bit_eq_div) | |
| 904 | qed | |
| 905 | ||
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 906 | lemma push_bit_0_id [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 907 | "push_bit 0 = id" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 908 | by (simp add: fun_eq_iff push_bit_eq_mult) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 909 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 910 | lemma push_bit_of_0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 911 | "push_bit n 0 = 0" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 912 | by (simp add: push_bit_eq_mult) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 913 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 914 | lemma push_bit_of_1: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 915 | "push_bit n 1 = 2 ^ n" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 916 | by (simp add: push_bit_eq_mult) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 917 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 918 | lemma push_bit_Suc [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 919 | "push_bit (Suc n) a = push_bit n (a * 2)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 920 | by (simp add: push_bit_eq_mult ac_simps) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 921 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 922 | lemma push_bit_double: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 923 | "push_bit n (a * 2) = push_bit n a * 2" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 924 | by (simp add: push_bit_eq_mult ac_simps) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 925 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 926 | lemma push_bit_eq_0_iff [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 927 | "push_bit n a = 0 \<longleftrightarrow> a = 0" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 928 | by (simp add: push_bit_eq_mult) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 929 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 930 | lemma push_bit_add: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 931 | "push_bit n (a + b) = push_bit n a + push_bit n b" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 932 | by (simp add: push_bit_eq_mult algebra_simps) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 933 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 934 | lemma push_bit_numeral [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 935 | "push_bit (numeral l) (numeral k) = push_bit (pred_numeral l) (numeral (Num.Bit0 k))" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 936 | by (simp only: numeral_eq_Suc power_Suc numeral_Bit0 [of k] mult_2 [symmetric]) (simp add: ac_simps) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 937 | |
| 68010 | 938 | lemma push_bit_of_nat: | 
| 939 | "push_bit n (of_nat m) = of_nat (push_bit n m)" | |
| 940 | by (simp add: push_bit_eq_mult Parity.push_bit_eq_mult) | |
| 941 | ||
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 942 | lemma take_bit_0 [simp]: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 943 | "take_bit 0 a = 0" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 944 | by (simp add: take_bit_eq_mod) | 
| 67816 | 945 | |
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 946 | lemma take_bit_Suc [simp]: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 947 | "take_bit (Suc n) a = take_bit n (a div 2) * 2 + of_bool (odd a)" | 
| 67816 | 948 | proof - | 
| 949 | have "1 + 2 * (a div 2) mod (2 * 2 ^ n) = (a div 2 * 2 + a mod 2) mod (2 * 2 ^ n)" | |
| 950 | if "odd a" | |
| 951 | using that mod_mult2_eq' [of "1 + 2 * (a div 2)" 2 "2 ^ n"] | |
| 952 | by (simp add: ac_simps odd_iff_mod_2_eq_one mult_mod_right) | |
| 953 | also have "\<dots> = a mod (2 * 2 ^ n)" | |
| 954 | by (simp only: div_mult_mod_eq) | |
| 955 | finally show ?thesis | |
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 956 | by (simp add: take_bit_eq_mod algebra_simps mult_mod_right) | 
| 67816 | 957 | qed | 
| 958 | ||
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 959 | lemma take_bit_of_0 [simp]: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 960 | "take_bit n 0 = 0" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 961 | by (simp add: take_bit_eq_mod) | 
| 67816 | 962 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 963 | lemma take_bit_of_1 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 964 | "take_bit n 1 = of_bool (n > 0)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 965 | by (simp add: take_bit_eq_mod) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 966 | |
| 67961 | 967 | lemma take_bit_add: | 
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 968 | "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 | 969 | by (simp add: take_bit_eq_mod mod_simps) | 
| 67816 | 970 | |
| 67961 | 971 | lemma take_bit_eq_0_iff: | 
| 972 | "take_bit n a = 0 \<longleftrightarrow> 2 ^ n dvd a" | |
| 973 | by (simp add: take_bit_eq_mod mod_eq_0_iff_dvd) | |
| 974 | ||
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 975 | lemma take_bit_of_1_eq_0_iff [simp]: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 976 | "take_bit n 1 = 0 \<longleftrightarrow> n = 0" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 977 | by (simp add: take_bit_eq_mod) | 
| 67816 | 978 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 979 | lemma even_take_bit_eq [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 980 | "even (take_bit n a) \<longleftrightarrow> n = 0 \<or> even a" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 981 | by (cases n) (simp_all add: take_bit_eq_mod dvd_mod_iff) | 
| 67816 | 982 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 983 | lemma take_bit_numeral_bit0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 984 | "take_bit (numeral l) (numeral (Num.Bit0 k)) = take_bit (pred_numeral l) (numeral k) * 2" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 985 | by (simp only: numeral_eq_Suc power_Suc numeral_Bit0 [of k] mult_2 [symmetric] take_bit_Suc | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 986 | ac_simps even_mult_iff nonzero_mult_div_cancel_right [OF numeral_neq_zero]) simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 987 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 988 | lemma take_bit_numeral_bit1 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 989 | "take_bit (numeral l) (numeral (Num.Bit1 k)) = take_bit (pred_numeral l) (numeral k) * 2 + 1" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 990 | by (simp only: numeral_eq_Suc power_Suc numeral_Bit1 [of k] mult_2 [symmetric] take_bit_Suc | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 991 | ac_simps even_add even_mult_iff div_mult_self1 [OF numeral_neq_zero]) (simp add: ac_simps) | 
| 67961 | 992 | |
| 68010 | 993 | lemma take_bit_of_nat: | 
| 994 | "take_bit n (of_nat m) = of_nat (take_bit n m)" | |
| 995 | by (simp add: take_bit_eq_mod Parity.take_bit_eq_mod of_nat_mod [of m "2 ^ n"]) | |
| 996 | ||
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 997 | lemma drop_bit_0 [simp]: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 998 | "drop_bit 0 = id" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 999 | by (simp add: fun_eq_iff drop_bit_eq_div) | 
| 67816 | 1000 | |
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1001 | lemma drop_bit_of_0 [simp]: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1002 | "drop_bit n 0 = 0" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1003 | by (simp add: drop_bit_eq_div) | 
| 67816 | 1004 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1005 | lemma drop_bit_of_1 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1006 | "drop_bit n 1 = of_bool (n = 0)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1007 | by (simp add: drop_bit_eq_div) | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1008 | |
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1009 | lemma drop_bit_Suc [simp]: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1010 | "drop_bit (Suc n) a = drop_bit n (a div 2)" | 
| 67816 | 1011 | proof (cases "even a") | 
| 1012 | case True | |
| 1013 | then obtain b where "a = 2 * b" .. | |
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1014 | moreover have "drop_bit (Suc n) (2 * b) = drop_bit n b" | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1015 | by (simp add: drop_bit_eq_div) | 
| 67816 | 1016 | ultimately show ?thesis | 
| 1017 | by simp | |
| 1018 | next | |
| 1019 | case False | |
| 1020 | then obtain b where "a = 2 * b + 1" .. | |
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1021 | moreover have "drop_bit (Suc n) (2 * b + 1) = drop_bit n b" | 
| 67816 | 1022 | using div_mult2_eq' [of "1 + b * 2" 2 "2 ^ n"] | 
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1023 | by (auto simp add: drop_bit_eq_div ac_simps) | 
| 67816 | 1024 | ultimately show ?thesis | 
| 1025 | by simp | |
| 1026 | qed | |
| 1027 | ||
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1028 | lemma drop_bit_half: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1029 | "drop_bit n (a div 2) = drop_bit n a div 2" | 
| 67816 | 1030 | by (induction n arbitrary: a) simp_all | 
| 1031 | ||
| 67907 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1032 | lemma drop_bit_of_bool [simp]: | 
| 
02a14c1cb917
prefer convention to place operation name before type name
 haftmann parents: 
67906diff
changeset | 1033 | "drop_bit n (of_bool d) = of_bool (n = 0 \<and> d)" | 
| 67816 | 1034 | by (cases n) simp_all | 
| 1035 | ||
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1036 | lemma drop_bit_numeral_bit0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1037 | "drop_bit (numeral l) (numeral (Num.Bit0 k)) = drop_bit (pred_numeral l) (numeral k)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1038 | by (simp only: numeral_eq_Suc power_Suc numeral_Bit0 [of k] mult_2 [symmetric] drop_bit_Suc | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1039 | nonzero_mult_div_cancel_left [OF numeral_neq_zero]) | 
| 67816 | 1040 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1041 | lemma drop_bit_numeral_bit1 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1042 | "drop_bit (numeral l) (numeral (Num.Bit1 k)) = drop_bit (pred_numeral l) (numeral k)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1043 | by (simp only: numeral_eq_Suc power_Suc numeral_Bit1 [of k] mult_2 [symmetric] drop_bit_Suc | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1044 | div_mult_self4 [OF numeral_neq_zero]) simp | 
| 67816 | 1045 | |
| 68010 | 1046 | lemma drop_bit_of_nat: | 
| 1047 | "drop_bit n (of_nat m) = of_nat (drop_bit n m)" | |
| 68389 | 1048 | by (simp add: drop_bit_eq_div Parity.drop_bit_eq_div of_nat_div [of m "2 ^ n"]) | 
| 68010 | 1049 | |
| 58770 | 1050 | end | 
| 67816 | 1051 | |
| 67988 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1052 | lemma push_bit_of_Suc_0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1053 | "push_bit n (Suc 0) = 2 ^ n" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1054 | using push_bit_of_1 [where ?'a = nat] by simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1055 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1056 | lemma take_bit_of_Suc_0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1057 | "take_bit n (Suc 0) = of_bool (0 < n)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1058 | using take_bit_of_1 [where ?'a = nat] by simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1059 | |
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1060 | lemma drop_bit_of_Suc_0 [simp]: | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1061 | "drop_bit n (Suc 0) = of_bool (n = 0)" | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1062 | using drop_bit_of_1 [where ?'a = nat] by simp | 
| 
01c651412081
explicit simp rules for computing abstract bit operations
 haftmann parents: 
67961diff
changeset | 1063 | |
| 67816 | 1064 | end |