| author | haftmann | 
| Thu, 29 Sep 2022 14:03:40 +0000 | |
| changeset 76224 | 64e8d4afcf10 | 
| parent 75937 | 02b18f59f903 | 
| child 76387 | 8cb141384682 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Parity.thy | 
| 2 | Author: Jeremy Avigad | |
| 3 | Author: Jacques D. Fleuriot | |
| 21256 | 4 | *) | 
| 5 | ||
| 60758 | 6 | section \<open>Parity in rings and semirings\<close> | 
| 21256 | 7 | |
| 8 | theory Parity | |
| 66815 | 9 | imports Euclidean_Division | 
| 21256 | 10 | begin | 
| 11 | ||
| 61799 | 12 | subsection \<open>Ring structures with parity and \<open>even\<close>/\<open>odd\<close> predicates\<close> | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 13 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 14 | class semiring_parity = comm_semiring_1 + semiring_modulo + | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 15 | assumes even_iff_mod_2_eq_zero: "2 dvd a \<longleftrightarrow> a mod 2 = 0" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 16 | and odd_iff_mod_2_eq_one: "\<not> 2 dvd a \<longleftrightarrow> a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 17 | and odd_one [simp]: "\<not> 2 dvd 1" | 
| 66839 | 18 | begin | 
| 19 | ||
| 58740 | 20 | abbreviation even :: "'a \<Rightarrow> bool" | 
| 63654 | 21 | where "even a \<equiv> 2 dvd a" | 
| 54228 | 22 | |
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 23 | abbreviation odd :: "'a \<Rightarrow> bool" | 
| 63654 | 24 | where "odd a \<equiv> \<not> 2 dvd a" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 25 | |
| 66815 | 26 | lemma parity_cases [case_names even odd]: | 
| 27 | assumes "even a \<Longrightarrow> a mod 2 = 0 \<Longrightarrow> P" | |
| 28 | assumes "odd a \<Longrightarrow> a mod 2 = 1 \<Longrightarrow> P" | |
| 29 | shows P | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 30 | using assms by (cases "even a") | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 31 | (simp_all add: even_iff_mod_2_eq_zero [symmetric] odd_iff_mod_2_eq_one [symmetric]) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 32 | |
| 71181 | 33 | lemma odd_of_bool_self [simp]: | 
| 34 | \<open>odd (of_bool p) \<longleftrightarrow> p\<close> | |
| 35 | by (cases p) simp_all | |
| 36 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 37 | lemma not_mod_2_eq_0_eq_1 [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 38 | "a mod 2 \<noteq> 0 \<longleftrightarrow> a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 39 | by (cases a rule: parity_cases) simp_all | 
| 66815 | 40 | |
| 41 | lemma not_mod_2_eq_1_eq_0 [simp]: | |
| 42 | "a mod 2 \<noteq> 1 \<longleftrightarrow> a mod 2 = 0" | |
| 43 | by (cases a rule: parity_cases) simp_all | |
| 44 | ||
| 58690 | 45 | lemma evenE [elim?]: | 
| 46 | assumes "even a" | |
| 47 | obtains b where "a = 2 * b" | |
| 58740 | 48 | using assms by (rule dvdE) | 
| 58690 | 49 | |
| 58681 | 50 | lemma oddE [elim?]: | 
| 58680 | 51 | assumes "odd a" | 
| 52 | obtains b where "a = 2 * b + 1" | |
| 58787 | 53 | proof - | 
| 66815 | 54 | have "a = 2 * (a div 2) + a mod 2" | 
| 55 | by (simp add: mult_div_mod_eq) | |
| 56 | with assms have "a = 2 * (a div 2) + 1" | |
| 57 | by (simp add: odd_iff_mod_2_eq_one) | |
| 58 | then show ?thesis .. | |
| 59 | qed | |
| 60 | ||
| 61 | lemma mod_2_eq_odd: | |
| 62 | "a mod 2 = of_bool (odd a)" | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 63 | by (auto elim: oddE simp add: even_iff_mod_2_eq_zero) | 
| 66815 | 64 | |
| 67816 | 65 | lemma of_bool_odd_eq_mod_2: | 
| 66 | "of_bool (odd a) = a mod 2" | |
| 67 | by (simp add: mod_2_eq_odd) | |
| 68 | ||
| 71426 | 69 | lemma even_mod_2_iff [simp]: | 
| 70 | \<open>even (a mod 2) \<longleftrightarrow> even a\<close> | |
| 71 | by (simp add: mod_2_eq_odd) | |
| 72 | ||
| 73 | lemma mod2_eq_if: | |
| 74 | "a mod 2 = (if even a then 0 else 1)" | |
| 75 | by (simp add: mod_2_eq_odd) | |
| 76 | ||
| 66815 | 77 | lemma even_zero [simp]: | 
| 78 | "even 0" | |
| 79 | by (fact dvd_0_right) | |
| 80 | ||
| 81 | lemma odd_even_add: | |
| 82 | "even (a + b)" if "odd a" and "odd b" | |
| 83 | proof - | |
| 84 | from that obtain c d where "a = 2 * c + 1" and "b = 2 * d + 1" | |
| 85 | by (blast elim: oddE) | |
| 86 | then have "a + b = 2 * c + 2 * d + (1 + 1)" | |
| 87 | by (simp only: ac_simps) | |
| 88 | also have "\<dots> = 2 * (c + d + 1)" | |
| 89 | by (simp add: algebra_simps) | |
| 90 | finally show ?thesis .. | |
| 91 | qed | |
| 92 | ||
| 93 | lemma even_add [simp]: | |
| 94 | "even (a + b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)" | |
| 95 | by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add) | |
| 96 | ||
| 97 | lemma odd_add [simp]: | |
| 98 | "odd (a + b) \<longleftrightarrow> \<not> (odd a \<longleftrightarrow> odd b)" | |
| 99 | by simp | |
| 100 | ||
| 101 | lemma even_plus_one_iff [simp]: | |
| 102 | "even (a + 1) \<longleftrightarrow> odd a" | |
| 103 | by (auto simp add: dvd_add_right_iff intro: odd_even_add) | |
| 104 | ||
| 105 | lemma even_mult_iff [simp]: | |
| 106 | "even (a * b) \<longleftrightarrow> even a \<or> even b" (is "?P \<longleftrightarrow> ?Q") | |
| 107 | proof | |
| 108 | assume ?Q | |
| 109 | then show ?P | |
| 110 | by auto | |
| 111 | next | |
| 112 | assume ?P | |
| 113 | show ?Q | |
| 114 | proof (rule ccontr) | |
| 115 | assume "\<not> (even a \<or> even b)" | |
| 116 | then have "odd a" and "odd b" | |
| 117 | by auto | |
| 118 | then obtain r s where "a = 2 * r + 1" and "b = 2 * s + 1" | |
| 119 | by (blast elim: oddE) | |
| 120 | then have "a * b = (2 * r + 1) * (2 * s + 1)" | |
| 121 | by simp | |
| 122 | also have "\<dots> = 2 * (2 * r * s + r + s) + 1" | |
| 123 | by (simp add: algebra_simps) | |
| 124 | finally have "odd (a * b)" | |
| 125 | by simp | |
| 126 | with \<open>?P\<close> show False | |
| 127 | by auto | |
| 128 | qed | |
| 129 | qed | |
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 130 | |
| 63654 | 131 | lemma even_numeral [simp]: "even (numeral (Num.Bit0 n))" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 132 | proof - | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 133 | have "even (2 * numeral n)" | 
| 66815 | 134 | unfolding even_mult_iff by simp | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 135 | then have "even (numeral n + numeral n)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 136 | unfolding mult_2 . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 137 | then show ?thesis | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 138 | unfolding numeral.simps . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 139 | qed | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 140 | |
| 63654 | 141 | lemma odd_numeral [simp]: "odd (numeral (Num.Bit1 n))" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 142 | proof | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 143 | assume "even (numeral (num.Bit1 n))" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 144 | then have "even (numeral n + numeral n + 1)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 145 | unfolding numeral.simps . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 146 | then have "even (2 * numeral n + 1)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 147 | unfolding mult_2 . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 148 | then have "2 dvd numeral n * 2 + 1" | 
| 58740 | 149 | by (simp add: ac_simps) | 
| 63654 | 150 | then have "2 dvd 1" | 
| 151 | using dvd_add_times_triv_left_iff [of 2 "numeral n" 1] by simp | |
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 152 | then show False by simp | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 153 | qed | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 154 | |
| 71755 | 155 | lemma odd_numeral_BitM [simp]: | 
| 156 | \<open>odd (numeral (Num.BitM w))\<close> | |
| 157 | by (cases w) simp_all | |
| 158 | ||
| 63654 | 159 | lemma even_power [simp]: "even (a ^ n) \<longleftrightarrow> even a \<and> n > 0" | 
| 58680 | 160 | by (induct n) auto | 
| 161 | ||
| 71412 | 162 | lemma mask_eq_sum_exp: | 
| 163 |   \<open>2 ^ n - 1 = (\<Sum>m\<in>{q. q < n}. 2 ^ m)\<close>
 | |
| 164 | proof - | |
| 165 |   have *: \<open>{q. q < Suc m} = insert m {q. q < m}\<close> for m
 | |
| 166 | by auto | |
| 167 |   have \<open>2 ^ n = (\<Sum>m\<in>{q. q < n}. 2 ^ m) + 1\<close>
 | |
| 168 | by (induction n) (simp_all add: ac_simps mult_2 *) | |
| 169 |   then have \<open>2 ^ n - 1 = (\<Sum>m\<in>{q. q < n}. 2 ^ m) + 1 - 1\<close>
 | |
| 170 | by simp | |
| 171 | then show ?thesis | |
| 172 | by simp | |
| 173 | qed | |
| 174 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 175 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 176 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 177 | class ring_parity = ring + semiring_parity | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 178 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 179 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 180 | subclass comm_ring_1 .. | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 181 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 182 | lemma even_minus: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 183 | "even (- a) \<longleftrightarrow> even a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 184 | by (fact dvd_minus_iff) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 185 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 186 | lemma even_diff [simp]: | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 187 | "even (a - b) \<longleftrightarrow> even (a + b)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 188 | using even_add [of a "- b"] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 189 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 190 | end | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 191 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 192 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 193 | subsection \<open>Special case: euclidean rings containing the natural numbers\<close> | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 194 | |
| 71157 | 195 | context unique_euclidean_semiring_with_nat | 
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 196 | begin | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 197 | |
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 198 | subclass semiring_parity | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 199 | proof | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 200 | show "2 dvd a \<longleftrightarrow> a mod 2 = 0" for a | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 201 | by (fact dvd_eq_mod_eq_0) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 202 | show "\<not> 2 dvd a \<longleftrightarrow> a mod 2 = 1" for a | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 203 | proof | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 204 | assume "a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 205 | then show "\<not> 2 dvd a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 206 | by auto | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 207 | next | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 208 | assume "\<not> 2 dvd a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 209 | have eucl: "euclidean_size (a mod 2) = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 210 | proof (rule order_antisym) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 211 | show "euclidean_size (a mod 2) \<le> 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 212 | using mod_size_less [of 2 a] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 213 | show "1 \<le> euclidean_size (a mod 2)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 214 | using \<open>\<not> 2 dvd a\<close> by (simp add: Suc_le_eq dvd_eq_mod_eq_0) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 215 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 216 | from \<open>\<not> 2 dvd a\<close> have "\<not> of_nat 2 dvd division_segment a * of_nat (euclidean_size a)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 217 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 218 | then have "\<not> of_nat 2 dvd of_nat (euclidean_size a)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 219 | by (auto simp only: dvd_mult_unit_iff' is_unit_division_segment) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 220 | then have "\<not> 2 dvd euclidean_size a" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 221 | using of_nat_dvd_iff [of 2] by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 222 | then have "euclidean_size a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 223 | by (simp add: semidom_modulo_class.dvd_eq_mod_eq_0) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 224 | then have "of_nat (euclidean_size a mod 2) = of_nat 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 225 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 226 | then have "of_nat (euclidean_size a) mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 227 | by (simp add: of_nat_mod) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 228 | from \<open>\<not> 2 dvd a\<close> eucl | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 229 | show "a mod 2 = 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 230 | by (auto intro: division_segment_eq_iff simp add: division_segment_mod) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 231 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 232 | show "\<not> is_unit 2" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 233 | proof (rule notI) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 234 | assume "is_unit 2" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 235 | then have "of_nat 2 dvd of_nat 1" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 236 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 237 | then have "is_unit (2::nat)" | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 238 | by (simp only: of_nat_dvd_iff) | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 239 | then show False | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 240 | by simp | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 241 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 242 | qed | 
| 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 243 | |
| 66815 | 244 | lemma even_succ_div_two [simp]: | 
| 245 | "even a \<Longrightarrow> (a + 1) div 2 = a div 2" | |
| 246 | by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero) | |
| 247 | ||
| 248 | lemma odd_succ_div_two [simp]: | |
| 249 | "odd a \<Longrightarrow> (a + 1) div 2 = a div 2 + 1" | |
| 250 | by (auto elim!: oddE simp add: add.assoc) | |
| 251 | ||
| 252 | lemma even_two_times_div_two: | |
| 253 | "even a \<Longrightarrow> 2 * (a div 2) = a" | |
| 254 | by (fact dvd_mult_div_cancel) | |
| 255 | ||
| 256 | lemma odd_two_times_div_two_succ [simp]: | |
| 257 | "odd a \<Longrightarrow> 2 * (a div 2) + 1 = a" | |
| 258 | using mult_div_mod_eq [of 2 a] | |
| 259 | by (simp add: even_iff_mod_2_eq_zero) | |
| 260 | ||
| 67051 | 261 | lemma coprime_left_2_iff_odd [simp]: | 
| 262 | "coprime 2 a \<longleftrightarrow> odd a" | |
| 263 | proof | |
| 264 | assume "odd a" | |
| 265 | show "coprime 2 a" | |
| 266 | proof (rule coprimeI) | |
| 267 | fix b | |
| 268 | assume "b dvd 2" "b dvd a" | |
| 269 | then have "b dvd a mod 2" | |
| 270 | by (auto intro: dvd_mod) | |
| 271 | with \<open>odd a\<close> show "is_unit b" | |
| 272 | by (simp add: mod_2_eq_odd) | |
| 273 | qed | |
| 274 | next | |
| 275 | assume "coprime 2 a" | |
| 276 | show "odd a" | |
| 277 | proof (rule notI) | |
| 278 | assume "even a" | |
| 279 | then obtain b where "a = 2 * b" .. | |
| 280 | with \<open>coprime 2 a\<close> have "coprime 2 (2 * b)" | |
| 281 | by simp | |
| 282 | moreover have "\<not> coprime 2 (2 * b)" | |
| 283 | by (rule not_coprimeI [of 2]) simp_all | |
| 284 | ultimately show False | |
| 285 | by blast | |
| 286 | qed | |
| 287 | qed | |
| 288 | ||
| 289 | lemma coprime_right_2_iff_odd [simp]: | |
| 290 | "coprime a 2 \<longleftrightarrow> odd a" | |
| 291 | using coprime_left_2_iff_odd [of a] by (simp add: ac_simps) | |
| 292 | ||
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 293 | end | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 294 | |
| 71157 | 295 | context unique_euclidean_ring_with_nat | 
| 58679 | 296 | begin | 
| 297 | ||
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 298 | subclass ring_parity .. | 
| 58680 | 299 | |
| 67906 | 300 | lemma minus_1_mod_2_eq [simp]: | 
| 301 | "- 1 mod 2 = 1" | |
| 302 | by (simp add: mod_2_eq_odd) | |
| 303 | ||
| 304 | lemma minus_1_div_2_eq [simp]: | |
| 305 | "- 1 div 2 = - 1" | |
| 306 | proof - | |
| 307 | from div_mult_mod_eq [of "- 1" 2] | |
| 308 | 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 | 309 | using add_implies_diff by fastforce | 
| 67906 | 310 | then show ?thesis | 
| 311 | using mult_right_cancel [of 2 "- 1 div 2" "- 1"] by simp | |
| 312 | qed | |
| 313 | ||
| 58679 | 314 | end | 
| 315 | ||
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 316 | |
| 69593 | 317 | subsection \<open>Instance for \<^typ>\<open>nat\<close>\<close> | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 318 | |
| 70340 | 319 | instance nat :: unique_euclidean_semiring_with_nat | 
| 66815 | 320 | by standard (simp_all add: dvd_eq_mod_eq_0) | 
| 66808 
1907167b6038
elementary definition of division on natural numbers
 haftmann parents: 
66582diff
changeset | 321 | |
| 66815 | 322 | lemma even_Suc_Suc_iff [simp]: | 
| 323 | "even (Suc (Suc n)) \<longleftrightarrow> even n" | |
| 58787 | 324 | using dvd_add_triv_right_iff [of 2 n] by simp | 
| 58687 | 325 | |
| 66815 | 326 | lemma even_Suc [simp]: "even (Suc n) \<longleftrightarrow> odd n" | 
| 327 | using even_plus_one_iff [of n] by simp | |
| 58787 | 328 | |
| 66815 | 329 | lemma even_diff_nat [simp]: | 
| 330 | "even (m - n) \<longleftrightarrow> m < n \<or> even (m + n)" for m n :: nat | |
| 58787 | 331 | proof (cases "n \<le> m") | 
| 332 | case True | |
| 333 | then have "m - n + n * 2 = m + n" by (simp add: mult_2_right) | |
| 66815 | 334 | moreover have "even (m - n) \<longleftrightarrow> even (m - n + n * 2)" by simp | 
| 335 | ultimately have "even (m - n) \<longleftrightarrow> even (m + n)" by (simp only:) | |
| 58787 | 336 | then show ?thesis by auto | 
| 337 | next | |
| 338 | case False | |
| 339 | then show ?thesis by simp | |
| 63654 | 340 | qed | 
| 341 | ||
| 66815 | 342 | lemma odd_pos: | 
| 343 | "odd n \<Longrightarrow> 0 < n" for n :: nat | |
| 58690 | 344 | by (auto elim: oddE) | 
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 345 | |
| 66815 | 346 | lemma Suc_double_not_eq_double: | 
| 347 | "Suc (2 * m) \<noteq> 2 * n" | |
| 62597 | 348 | proof | 
| 349 | assume "Suc (2 * m) = 2 * n" | |
| 350 | moreover have "odd (Suc (2 * m))" and "even (2 * n)" | |
| 351 | by simp_all | |
| 352 | ultimately show False by simp | |
| 353 | qed | |
| 354 | ||
| 66815 | 355 | lemma double_not_eq_Suc_double: | 
| 356 | "2 * m \<noteq> Suc (2 * n)" | |
| 62597 | 357 | using Suc_double_not_eq_double [of n m] by simp | 
| 358 | ||
| 66815 | 359 | lemma odd_Suc_minus_one [simp]: "odd n \<Longrightarrow> Suc (n - Suc 0) = n" | 
| 360 | by (auto elim: oddE) | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 361 | |
| 66815 | 362 | lemma even_Suc_div_two [simp]: | 
| 363 | "even n \<Longrightarrow> Suc n div 2 = n div 2" | |
| 364 | using even_succ_div_two [of n] by simp | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 365 | |
| 66815 | 366 | lemma odd_Suc_div_two [simp]: | 
| 367 | "odd n \<Longrightarrow> Suc n div 2 = Suc (n div 2)" | |
| 368 | using odd_succ_div_two [of n] by simp | |
| 60343 
063698416239
correct sort constraints for abbreviations in type classes
 haftmann parents: 
59816diff
changeset | 369 | |
| 66815 | 370 | lemma odd_two_times_div_two_nat [simp]: | 
| 371 | assumes "odd n" | |
| 372 | shows "2 * (n div 2) = n - (1 :: nat)" | |
| 373 | proof - | |
| 374 | from assms have "2 * (n div 2) + 1 = n" | |
| 375 | by (rule odd_two_times_div_two_succ) | |
| 376 | then have "Suc (2 * (n div 2)) - 1 = n - 1" | |
| 58787 | 377 | by simp | 
| 66815 | 378 | then show ?thesis | 
| 379 | by simp | |
| 58787 | 380 | qed | 
| 58680 | 381 | |
| 70341 
972c0c744e7c
generalized type classes for parity to cover word types also, which contain zero divisors
 haftmann parents: 
70340diff
changeset | 382 | 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 | 383 | "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 | 384 | 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 | 385 | |
| 69502 | 386 | lemma odd_card_imp_not_empty: | 
| 387 |   \<open>A \<noteq> {}\<close> if \<open>odd (card A)\<close>
 | |
| 388 | using that by auto | |
| 389 | ||
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 390 | 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 | 391 | 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 | 392 | shows "P n" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 393 | proof (induct n rule: less_induct) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 394 | case (less n) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 395 | show ?case | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 396 | proof (cases "n < Suc (Suc 0)") | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 397 | case True | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 398 | then show ?thesis | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 399 | 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 | 400 | next | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 401 | case False | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 402 | 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 | 403 | 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 | 404 | then have "k<n" | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 405 | by simp | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 406 | 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 | 407 | by blast | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 408 | then show ?thesis | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 409 | by (simp add: k) | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 410 | qed | 
| 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70353diff
changeset | 411 | qed | 
| 58687 | 412 | |
| 71413 | 413 | lemma mask_eq_sum_exp_nat: | 
| 414 |   \<open>2 ^ n - Suc 0 = (\<Sum>m\<in>{q. q < n}. 2 ^ m)\<close>
 | |
| 415 | using mask_eq_sum_exp [where ?'a = nat] by simp | |
| 416 | ||
| 71412 | 417 | context semiring_parity | 
| 418 | begin | |
| 419 | ||
| 74592 | 420 | lemma even_of_nat_iff [simp]: | 
| 421 | "even (of_nat n) \<longleftrightarrow> even n" | |
| 422 | by (induction n) simp_all | |
| 423 | ||
| 71412 | 424 | lemma even_sum_iff: | 
| 425 |   \<open>even (sum f A) \<longleftrightarrow> even (card {a\<in>A. odd (f a)})\<close> if \<open>finite A\<close>
 | |
| 426 | using that proof (induction A) | |
| 427 | case empty | |
| 428 | then show ?case | |
| 429 | by simp | |
| 430 | next | |
| 431 | case (insert a A) | |
| 432 |   moreover have \<open>{b \<in> insert a A. odd (f b)} = (if odd (f a) then {a} else {}) \<union> {b \<in> A. odd (f b)}\<close>
 | |
| 433 | by auto | |
| 434 | ultimately show ?case | |
| 435 | by simp | |
| 436 | qed | |
| 437 | ||
| 438 | lemma even_prod_iff: | |
| 439 | \<open>even (prod f A) \<longleftrightarrow> (\<exists>a\<in>A. even (f a))\<close> if \<open>finite A\<close> | |
| 440 | using that by (induction A) simp_all | |
| 441 | ||
| 442 | lemma even_mask_iff [simp]: | |
| 443 | \<open>even (2 ^ n - 1) \<longleftrightarrow> n = 0\<close> | |
| 444 | proof (cases \<open>n = 0\<close>) | |
| 445 | case True | |
| 446 | then show ?thesis | |
| 447 | by simp | |
| 448 | next | |
| 449 | case False | |
| 450 |   then have \<open>{a. a = 0 \<and> a < n} = {0}\<close>
 | |
| 451 | by auto | |
| 452 | then show ?thesis | |
| 453 | by (auto simp add: mask_eq_sum_exp even_sum_iff) | |
| 454 | qed | |
| 455 | ||
| 456 | end | |
| 457 | ||
| 71138 | 458 | |
| 60758 | 459 | subsection \<open>Parity and powers\<close> | 
| 58689 | 460 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
60867diff
changeset | 461 | context ring_1 | 
| 58689 | 462 | begin | 
| 463 | ||
| 63654 | 464 | lemma power_minus_even [simp]: "even n \<Longrightarrow> (- a) ^ n = a ^ n" | 
| 58690 | 465 | by (auto elim: evenE) | 
| 58689 | 466 | |
| 63654 | 467 | lemma power_minus_odd [simp]: "odd n \<Longrightarrow> (- a) ^ n = - (a ^ n)" | 
| 58690 | 468 | by (auto elim: oddE) | 
| 469 | ||
| 66815 | 470 | lemma uminus_power_if: | 
| 471 | "(- a) ^ n = (if even n then a ^ n else - (a ^ n))" | |
| 472 | by auto | |
| 473 | ||
| 63654 | 474 | lemma neg_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1" | 
| 58690 | 475 | by simp | 
| 58689 | 476 | |
| 63654 | 477 | lemma neg_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1" | 
| 58690 | 478 | by simp | 
| 58689 | 479 | |
| 66582 | 480 | lemma neg_one_power_add_eq_neg_one_power_diff: "k \<le> n \<Longrightarrow> (- 1) ^ (n + k) = (- 1) ^ (n - k)" | 
| 481 | by (cases "even (n + k)") auto | |
| 482 | ||
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 483 | 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 | 484 | by (induct n) auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
67083diff
changeset | 485 | |
| 63654 | 486 | end | 
| 58689 | 487 | |
| 488 | context linordered_idom | |
| 489 | begin | |
| 490 | ||
| 63654 | 491 | lemma zero_le_even_power: "even n \<Longrightarrow> 0 \<le> a ^ n" | 
| 58690 | 492 | by (auto elim: evenE) | 
| 58689 | 493 | |
| 63654 | 494 | lemma zero_le_odd_power: "odd n \<Longrightarrow> 0 \<le> a ^ n \<longleftrightarrow> 0 \<le> a" | 
| 58689 | 495 | by (auto simp add: power_even_eq zero_le_mult_iff elim: oddE) | 
| 496 | ||
| 63654 | 497 | lemma zero_le_power_eq: "0 \<le> a ^ n \<longleftrightarrow> even n \<or> odd n \<and> 0 \<le> a" | 
| 58787 | 498 | by (auto simp add: zero_le_even_power zero_le_odd_power) | 
| 63654 | 499 | |
| 500 | 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 | 501 | proof - | 
| 502 | have [simp]: "0 = a ^ n \<longleftrightarrow> a = 0 \<and> n > 0" | |
| 58787 | 503 | unfolding power_eq_0_iff [of a n, symmetric] by blast | 
| 58689 | 504 | show ?thesis | 
| 63654 | 505 | unfolding less_le zero_le_power_eq by auto | 
| 58689 | 506 | qed | 
| 507 | ||
| 63654 | 508 | lemma power_less_zero_eq [simp]: "a ^ n < 0 \<longleftrightarrow> odd n \<and> a < 0" | 
| 58689 | 509 | unfolding not_le [symmetric] zero_le_power_eq by auto | 
| 510 | ||
| 63654 | 511 | 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)" | 
| 512 | unfolding not_less [symmetric] zero_less_power_eq by auto | |
| 513 | ||
| 514 | lemma power_even_abs: "even n \<Longrightarrow> \<bar>a\<bar> ^ n = a ^ n" | |
| 58689 | 515 | using power_abs [of a n] by (simp add: zero_le_even_power) | 
| 516 | ||
| 517 | lemma power_mono_even: | |
| 518 | assumes "even n" and "\<bar>a\<bar> \<le> \<bar>b\<bar>" | |
| 519 | shows "a ^ n \<le> b ^ n" | |
| 520 | proof - | |
| 521 | have "0 \<le> \<bar>a\<bar>" by auto | |
| 63654 | 522 | with \<open>\<bar>a\<bar> \<le> \<bar>b\<bar>\<close> have "\<bar>a\<bar> ^ n \<le> \<bar>b\<bar> ^ n" | 
| 523 | by (rule power_mono) | |
| 524 | with \<open>even n\<close> show ?thesis | |
| 525 | by (simp add: power_even_abs) | |
| 58689 | 526 | qed | 
| 527 | ||
| 528 | lemma power_mono_odd: | |
| 529 | assumes "odd n" and "a \<le> b" | |
| 530 | shows "a ^ n \<le> b ^ n" | |
| 531 | proof (cases "b < 0") | |
| 63654 | 532 | case True | 
| 533 | with \<open>a \<le> b\<close> have "- b \<le> - a" and "0 \<le> - b" by auto | |
| 534 | then have "(- b) ^ n \<le> (- a) ^ n" by (rule power_mono) | |
| 60758 | 535 | with \<open>odd n\<close> show ?thesis by simp | 
| 58689 | 536 | next | 
| 63654 | 537 | case False | 
| 538 | then have "0 \<le> b" by auto | |
| 58689 | 539 | show ?thesis | 
| 540 | proof (cases "a < 0") | |
| 63654 | 541 | case True | 
| 542 | then have "n \<noteq> 0" and "a \<le> 0" using \<open>odd n\<close> [THEN odd_pos] by auto | |
| 60758 | 543 | then have "a ^ n \<le> 0" unfolding power_le_zero_eq using \<open>odd n\<close> by auto | 
| 63654 | 544 | moreover from \<open>0 \<le> b\<close> have "0 \<le> b ^ n" by auto | 
| 58689 | 545 | ultimately show ?thesis by auto | 
| 546 | next | |
| 63654 | 547 | case False | 
| 548 | then have "0 \<le> a" by auto | |
| 549 | with \<open>a \<le> b\<close> show ?thesis | |
| 550 | using power_mono by auto | |
| 58689 | 551 | qed | 
| 552 | qed | |
| 62083 | 553 | |
| 60758 | 554 | text \<open>Simplify, when the exponent is a numeral\<close> | 
| 58689 | 555 | |
| 556 | lemma zero_le_power_eq_numeral [simp]: | |
| 557 | "0 \<le> a ^ numeral w \<longleftrightarrow> even (numeral w :: nat) \<or> odd (numeral w :: nat) \<and> 0 \<le> a" | |
| 558 | by (fact zero_le_power_eq) | |
| 559 | ||
| 560 | lemma zero_less_power_eq_numeral [simp]: | |
| 63654 | 561 | "0 < a ^ numeral w \<longleftrightarrow> | 
| 562 | numeral w = (0 :: nat) \<or> | |
| 563 | even (numeral w :: nat) \<and> a \<noteq> 0 \<or> | |
| 564 | odd (numeral w :: nat) \<and> 0 < a" | |
| 58689 | 565 | by (fact zero_less_power_eq) | 
| 566 | ||
| 567 | lemma power_le_zero_eq_numeral [simp]: | |
| 63654 | 568 | "a ^ numeral w \<le> 0 \<longleftrightarrow> | 
| 569 | (0 :: nat) < numeral w \<and> | |
| 570 | (odd (numeral w :: nat) \<and> a \<le> 0 \<or> even (numeral w :: nat) \<and> a = 0)" | |
| 58689 | 571 | by (fact power_le_zero_eq) | 
| 572 | ||
| 573 | lemma power_less_zero_eq_numeral [simp]: | |
| 574 | "a ^ numeral w < 0 \<longleftrightarrow> odd (numeral w :: nat) \<and> a < 0" | |
| 575 | by (fact power_less_zero_eq) | |
| 576 | ||
| 577 | lemma power_even_abs_numeral [simp]: | |
| 578 | "even (numeral w :: nat) \<Longrightarrow> \<bar>a\<bar> ^ numeral w = a ^ numeral w" | |
| 579 | by (fact power_even_abs) | |
| 580 | ||
| 581 | end | |
| 582 | ||
| 71413 | 583 | context unique_euclidean_semiring_with_nat | 
| 584 | begin | |
| 585 | ||
| 586 | lemma even_mask_div_iff': | |
| 587 | \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> m \<le> n\<close> | |
| 588 | proof - | |
| 589 | have \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> even (of_nat ((2 ^ m - Suc 0) div 2 ^ n))\<close> | |
| 590 | by (simp only: of_nat_div) (simp add: of_nat_diff) | |
| 591 | also have \<open>\<dots> \<longleftrightarrow> even ((2 ^ m - Suc 0) div 2 ^ n)\<close> | |
| 592 | by simp | |
| 593 | also have \<open>\<dots> \<longleftrightarrow> m \<le> n\<close> | |
| 594 | proof (cases \<open>m \<le> n\<close>) | |
| 595 | case True | |
| 596 | then show ?thesis | |
| 597 | by (simp add: Suc_le_lessD) | |
| 598 | next | |
| 599 | case False | |
| 600 | then obtain r where r: \<open>m = n + Suc r\<close> | |
| 601 | using less_imp_Suc_add by fastforce | |
| 602 |     from r have \<open>{q. q < m} \<inter> {q. 2 ^ n dvd (2::nat) ^ q} = {q. n \<le> q \<and> q < m}\<close>
 | |
| 603 | by (auto simp add: dvd_power_iff_le) | |
| 604 |     moreover from r have \<open>{q. q < m} \<inter> {q. \<not> 2 ^ n dvd (2::nat) ^ q} = {q. q < n}\<close>
 | |
| 605 | by (auto simp add: dvd_power_iff_le) | |
| 606 |     moreover from False have \<open>{q. n \<le> q \<and> q < m \<and> q \<le> n} = {n}\<close>
 | |
| 607 | by auto | |
| 608 |     then have \<open>odd ((\<Sum>a\<in>{q. n \<le> q \<and> q < m}. 2 ^ a div (2::nat) ^ n) + sum ((^) 2) {q. q < n} div 2 ^ n)\<close>
 | |
| 609 | by (simp_all add: euclidean_semiring_cancel_class.power_diff_power_eq semiring_parity_class.even_sum_iff not_less mask_eq_sum_exp_nat [symmetric]) | |
| 610 |     ultimately have \<open>odd (sum ((^) (2::nat)) {q. q < m} div 2 ^ n)\<close>
 | |
| 611 | by (subst euclidean_semiring_cancel_class.sum_div_partition) simp_all | |
| 612 | with False show ?thesis | |
| 613 | by (simp add: mask_eq_sum_exp_nat) | |
| 614 | qed | |
| 615 | finally show ?thesis . | |
| 616 | qed | |
| 617 | ||
| 618 | end | |
| 619 | ||
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 620 | |
| 69593 | 621 | subsection \<open>Instance for \<^typ>\<open>int\<close>\<close> | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 622 | |
| 67816 | 623 | lemma even_diff_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 624 | "even (k - l) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 625 | by (fact even_diff) | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 626 | |
| 67816 | 627 | lemma even_abs_add_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 628 | "even (\<bar>k\<bar> + l) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 629 | by simp | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 630 | |
| 67816 | 631 | lemma even_add_abs_iff: | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 632 | "even (k + \<bar>l\<bar>) \<longleftrightarrow> even (k + l)" for k l :: int | 
| 67816 | 633 | by simp | 
| 66816 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 634 | |
| 
212a3334e7da
more fundamental definition of div and mod on int
 haftmann parents: 
66815diff
changeset | 635 | lemma even_nat_iff: "0 \<le> k \<Longrightarrow> even (nat k) \<longleftrightarrow> even k" | 
| 74592 | 636 | by (simp add: even_of_nat_iff [of "nat k", where ?'a = int, symmetric]) | 
| 71138 | 637 | |
| 71837 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 638 | context | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 639 |   assumes "SORT_CONSTRAINT('a::division_ring)"
 | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 640 | begin | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 641 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 642 | lemma power_int_minus_left: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 643 | "power_int (-a :: 'a) n = (if even n then power_int a n else -power_int a n)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 644 | by (auto simp: power_int_def minus_one_power_iff even_nat_iff) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 645 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 646 | lemma power_int_minus_left_even [simp]: "even n \<Longrightarrow> power_int (-a :: 'a) n = power_int a n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 647 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 648 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 649 | lemma power_int_minus_left_odd [simp]: "odd n \<Longrightarrow> power_int (-a :: 'a) n = -power_int a n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 650 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 651 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 652 | lemma power_int_minus_left_distrib: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 653 | "NO_MATCH (-1) x \<Longrightarrow> power_int (-a :: 'a) n = power_int (-1) n * power_int a n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 654 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 655 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 656 | lemma power_int_minus_one_minus: "power_int (-1 :: 'a) (-n) = power_int (-1) n" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 657 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 658 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 659 | lemma power_int_minus_one_diff_commute: "power_int (-1 :: 'a) (a - b) = power_int (-1) (b - a)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 660 | by (subst power_int_minus_one_minus [symmetric]) auto | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 661 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 662 | lemma power_int_minus_one_mult_self [simp]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 663 | "power_int (-1 :: 'a) m * power_int (-1) m = 1" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 664 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 665 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 666 | lemma power_int_minus_one_mult_self' [simp]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 667 | "power_int (-1 :: 'a) m * (power_int (-1) m * b) = b" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 668 | by (simp add: power_int_minus_left) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 669 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 670 | end | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71822diff
changeset | 671 | |
| 75937 | 672 | |
| 673 | subsection \<open>Computing congruences modulo \<open>2 ^ q\<close>\<close> | |
| 674 | ||
| 675 | context unique_euclidean_semiring_with_nat_division | |
| 676 | begin | |
| 677 | ||
| 678 | lemma cong_exp_iff_simps: | |
| 679 | "numeral n mod numeral Num.One = 0 | |
| 680 | \<longleftrightarrow> True" | |
| 681 | "numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) = 0 | |
| 682 | \<longleftrightarrow> numeral n mod numeral q = 0" | |
| 683 | "numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) = 0 | |
| 684 | \<longleftrightarrow> False" | |
| 685 | "numeral m mod numeral Num.One = (numeral n mod numeral Num.One) | |
| 686 | \<longleftrightarrow> True" | |
| 687 | "numeral Num.One mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) | |
| 688 | \<longleftrightarrow> True" | |
| 689 | "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) | |
| 690 | \<longleftrightarrow> False" | |
| 691 | "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) | |
| 692 | \<longleftrightarrow> (numeral n mod numeral q) = 0" | |
| 693 | "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) | |
| 694 | \<longleftrightarrow> False" | |
| 695 | "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) | |
| 696 | \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)" | |
| 697 | "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) | |
| 698 | \<longleftrightarrow> False" | |
| 699 | "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q)) | |
| 700 | \<longleftrightarrow> (numeral m mod numeral q) = 0" | |
| 701 | "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q)) | |
| 702 | \<longleftrightarrow> False" | |
| 703 | "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q)) | |
| 704 | \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)" | |
| 705 | by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even]) | |
| 706 | ||
| 707 | end | |
| 708 | ||
| 709 | ||
| 71853 | 710 | code_identifier | 
| 711 | code_module Parity \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith | |
| 712 | ||
| 74592 | 713 | lemmas even_of_nat = even_of_nat_iff | 
| 714 | ||
| 67816 | 715 | end |