| author | wenzelm | 
| Fri, 06 Mar 2015 23:53:52 +0100 | |
| changeset 59637 | f643308472ce | 
| parent 58889 | 5b7a9633cfa8 | 
| child 59816 | 034b13f4efae | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Parity.thy | 
| 2 | Author: Jeremy Avigad | |
| 3 | Author: Jacques D. Fleuriot | |
| 21256 | 4 | *) | 
| 5 | ||
| 58889 | 6 | section {* Parity in rings and semirings *}
 | 
| 21256 | 7 | |
| 8 | theory Parity | |
| 58778 
e29cae8eab1f
even further downshift of theory Parity in the hierarchy
 haftmann parents: 
58777diff
changeset | 9 | imports Nat_Transfer | 
| 21256 | 10 | begin | 
| 11 | ||
| 58787 | 12 | subsection {* Ring structures with parity and @{text even}/@{text odd} predicates *}
 | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 13 | |
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 14 | class semiring_parity = semiring_dvd + semiring_numeral + | 
| 58787 | 15 | assumes odd_one [simp]: "\<not> 2 dvd 1" | 
| 16 | assumes odd_even_add: "\<not> 2 dvd a \<Longrightarrow> \<not> 2 dvd b \<Longrightarrow> 2 dvd a + b" | |
| 17 | assumes even_multD: "2 dvd a * b \<Longrightarrow> 2 dvd a \<or> 2 dvd b" | |
| 18 | assumes odd_ex_decrement: "\<not> 2 dvd a \<Longrightarrow> \<exists>b. a = b + 1" | |
| 54227 
63b441f49645
moving generic lemmas out of theory parity, disregarding some unused auxiliary lemmas;
 haftmann parents: 
47225diff
changeset | 19 | begin | 
| 21256 | 20 | |
| 58740 | 21 | abbreviation even :: "'a \<Rightarrow> bool" | 
| 54228 | 22 | where | 
| 58740 | 23 | "even a \<equiv> 2 dvd a" | 
| 54228 | 24 | |
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 25 | abbreviation odd :: "'a \<Rightarrow> bool" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 26 | where | 
| 58740 | 27 | "odd a \<equiv> \<not> 2 dvd a" | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 28 | |
| 58787 | 29 | lemma even_zero [simp]: | 
| 30 | "even 0" | |
| 31 | by (fact dvd_0_right) | |
| 32 | ||
| 33 | lemma even_plus_one_iff [simp]: | |
| 34 | "even (a + 1) \<longleftrightarrow> odd a" | |
| 35 | by (auto simp add: dvd_add_right_iff intro: odd_even_add) | |
| 36 | ||
| 58690 | 37 | lemma evenE [elim?]: | 
| 38 | assumes "even a" | |
| 39 | obtains b where "a = 2 * b" | |
| 58740 | 40 | using assms by (rule dvdE) | 
| 58690 | 41 | |
| 58681 | 42 | lemma oddE [elim?]: | 
| 58680 | 43 | assumes "odd a" | 
| 44 | obtains b where "a = 2 * b + 1" | |
| 58787 | 45 | proof - | 
| 46 | from assms obtain b where *: "a = b + 1" | |
| 47 | by (blast dest: odd_ex_decrement) | |
| 48 | with assms have "even (b + 2)" by simp | |
| 49 | then have "even b" by simp | |
| 50 | then obtain c where "b = 2 * c" .. | |
| 51 | with * have "a = 2 * c + 1" by simp | |
| 52 | with that show thesis . | |
| 53 | qed | |
| 54 | ||
| 58770 | 55 | lemma even_times_iff [simp]: | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 56 | "even (a * b) \<longleftrightarrow> even a \<or> even b" | 
| 58787 | 57 | by (auto dest: even_multD) | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 58 | |
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 59 | lemma even_numeral [simp]: | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 60 | "even (numeral (Num.Bit0 n))" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 61 | proof - | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 62 | have "even (2 * numeral n)" | 
| 58740 | 63 | unfolding even_times_iff by simp | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 64 | then have "even (numeral n + numeral n)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 65 | unfolding mult_2 . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 66 | then show ?thesis | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 67 | unfolding numeral.simps . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 68 | qed | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 69 | |
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 70 | lemma odd_numeral [simp]: | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 71 | "odd (numeral (Num.Bit1 n))" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 72 | proof | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 73 | assume "even (numeral (num.Bit1 n))" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 74 | then have "even (numeral n + numeral n + 1)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 75 | unfolding numeral.simps . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 76 | then have "even (2 * numeral n + 1)" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 77 | unfolding mult_2 . | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 78 | then have "2 dvd numeral n * 2 + 1" | 
| 58740 | 79 | by (simp add: ac_simps) | 
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 80 | with dvd_add_times_triv_left_iff [of 2 "numeral n" 1] | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 81 | have "2 dvd 1" | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 82 | by simp | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 83 | then show False by simp | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 84 | qed | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 85 | |
| 58680 | 86 | lemma even_add [simp]: | 
| 87 | "even (a + b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)" | |
| 58787 | 88 | by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add) | 
| 58680 | 89 | |
| 90 | lemma odd_add [simp]: | |
| 91 | "odd (a + b) \<longleftrightarrow> (\<not> (odd a \<longleftrightarrow> odd b))" | |
| 92 | by simp | |
| 93 | ||
| 58770 | 94 | lemma even_power [simp]: | 
| 58771 
0997ea62e868
slight generalization and unification of simp rules for algebraic procedures
 haftmann parents: 
58770diff
changeset | 95 | "even (a ^ n) \<longleftrightarrow> even a \<and> n > 0" | 
| 58680 | 96 | by (induct n) auto | 
| 97 | ||
| 58678 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 98 | end | 
| 
398e05aa84d4
purely algebraic characterization of even and odd
 haftmann parents: 
58645diff
changeset | 99 | |
| 58787 | 100 | class ring_parity = comm_ring_1 + semiring_parity | 
| 58679 | 101 | begin | 
| 102 | ||
| 58770 | 103 | lemma even_minus [simp]: | 
| 58679 | 104 | "even (- a) \<longleftrightarrow> even a" | 
| 58740 | 105 | by (fact dvd_minus_iff) | 
| 58679 | 106 | |
| 58680 | 107 | lemma even_diff [simp]: | 
| 108 | "even (a - b) \<longleftrightarrow> even (a + b)" | |
| 109 | using even_add [of a "- b"] by simp | |
| 110 | ||
| 58679 | 111 | end | 
| 112 | ||
| 58710 
7216a10d69ba
augmented and tuned facts on even/odd and division
 haftmann parents: 
58709diff
changeset | 113 | |
| 58787 | 114 | subsection {* Instances for @{typ nat} and @{typ int} *}
 | 
| 115 | ||
| 116 | lemma even_Suc_Suc_iff [simp]: | |
| 117 | "even (Suc (Suc n)) \<longleftrightarrow> even n" | |
| 118 | using dvd_add_triv_right_iff [of 2 n] by simp | |
| 58687 | 119 | |
| 58770 | 120 | lemma even_Suc [simp]: | 
| 58787 | 121 | "even (Suc n) \<longleftrightarrow> odd n" | 
| 122 | by (induct n) auto | |
| 123 | ||
| 124 | lemma even_diff_nat [simp]: | |
| 125 | fixes m n :: nat | |
| 126 | shows "even (m - n) \<longleftrightarrow> m < n \<or> even (m + n)" | |
| 127 | proof (cases "n \<le> m") | |
| 128 | case True | |
| 129 | then have "m - n + n * 2 = m + n" by (simp add: mult_2_right) | |
| 130 | moreover have "even (m - n) \<longleftrightarrow> even (m - n + n * 2)" by simp | |
| 131 | ultimately have "even (m - n) \<longleftrightarrow> even (m + n)" by (simp only:) | |
| 132 | then show ?thesis by auto | |
| 133 | next | |
| 134 | case False | |
| 135 | then show ?thesis by simp | |
| 136 | qed | |
| 137 | ||
| 138 | lemma even_diff_iff [simp]: | |
| 139 | fixes k l :: int | |
| 140 | shows "even (k - l) \<longleftrightarrow> even (k + l)" | |
| 141 | using dvd_add_times_triv_right_iff [of 2 "k - l" l] by (simp add: mult_2_right) | |
| 142 | ||
| 143 | lemma even_abs_add_iff [simp]: | |
| 144 | fixes k l :: int | |
| 145 | shows "even (\<bar>k\<bar> + l) \<longleftrightarrow> even (k + l)" | |
| 146 | by (cases "k \<ge> 0") (simp_all add: ac_simps) | |
| 147 | ||
| 148 | lemma even_add_abs_iff [simp]: | |
| 149 | fixes k l :: int | |
| 150 | shows "even (k + \<bar>l\<bar>) \<longleftrightarrow> even (k + l)" | |
| 151 | using even_abs_add_iff [of l k] by (simp add: ac_simps) | |
| 152 | ||
| 153 | instance nat :: semiring_parity | |
| 154 | proof | |
| 155 | show "odd (1 :: nat)" | |
| 156 | by (rule notI, erule dvdE) simp | |
| 157 | next | |
| 158 | fix m n :: nat | |
| 159 | assume "odd m" | |
| 160 | moreover assume "odd n" | |
| 161 | ultimately have *: "even (Suc m) \<and> even (Suc n)" | |
| 162 | by simp | |
| 163 | then have "even (Suc m + Suc n)" | |
| 164 | by (blast intro: dvd_add) | |
| 165 | also have "Suc m + Suc n = m + n + 2" | |
| 166 | by simp | |
| 167 | finally show "even (m + n)" | |
| 168 | using dvd_add_triv_right_iff [of 2 "m + n"] by simp | |
| 169 | next | |
| 170 | fix m n :: nat | |
| 171 | assume *: "even (m * n)" | |
| 172 | show "even m \<or> even n" | |
| 173 | proof (rule disjCI) | |
| 174 | assume "odd n" | |
| 175 | then have "even (Suc n)" by simp | |
| 176 | then obtain r where "Suc n = 2 * r" .. | |
| 177 | moreover from * obtain s where "m * n = 2 * s" .. | |
| 178 | then have "2 * s + m = m * Suc n" by simp | |
| 179 | ultimately have " 2 * s + m = 2 * (m * r)" by (simp add: algebra_simps) | |
| 180 | then have "m = 2 * (m * r - s)" by simp | |
| 181 | then show "even m" .. | |
| 182 | qed | |
| 183 | next | |
| 184 | fix n :: nat | |
| 185 | assume "odd n" | |
| 186 | then show "\<exists>m. n = m + 1" | |
| 187 | by (cases n) simp_all | |
| 188 | qed | |
| 58687 | 189 | |
| 58689 | 190 | lemma odd_pos: | 
| 191 | "odd (n :: nat) \<Longrightarrow> 0 < n" | |
| 58690 | 192 | by (auto elim: oddE) | 
| 58689 | 193 | |
| 58787 | 194 | instance int :: ring_parity | 
| 195 | proof | |
| 196 | show "odd (1 :: int)" by (simp add: dvd_int_unfold_dvd_nat) | |
| 197 | fix k l :: int | |
| 198 | assume "odd k" | |
| 199 | moreover assume "odd l" | |
| 200 | ultimately have "even (nat \<bar>k\<bar> + nat \<bar>l\<bar>)" | |
| 201 | by (auto simp add: dvd_int_unfold_dvd_nat intro: odd_even_add) | |
| 202 | then have "even (\<bar>k\<bar> + \<bar>l\<bar>)" | |
| 203 | by (simp add: dvd_int_unfold_dvd_nat nat_add_distrib) | |
| 204 | then show "even (k + l)" | |
| 205 | by simp | |
| 206 | next | |
| 207 | fix k l :: int | |
| 208 | assume "even (k * l)" | |
| 209 | then show "even k \<or> even l" | |
| 210 | by (simp add: dvd_int_unfold_dvd_nat even_multD nat_abs_mult_distrib) | |
| 211 | next | |
| 212 | fix k :: int | |
| 213 | have "k = (k - 1) + 1" by simp | |
| 214 | then show "\<exists>l. k = l + 1" .. | |
| 215 | qed | |
| 58680 | 216 | |
| 58787 | 217 | lemma even_int_iff [simp]: | 
| 58679 | 218 | "even (int n) \<longleftrightarrow> even n" | 
| 58740 | 219 | by (simp add: dvd_int_iff) | 
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
31718diff
changeset | 220 | |
| 58687 | 221 | lemma even_nat_iff: | 
| 222 | "0 \<le> k \<Longrightarrow> even (nat k) \<longleftrightarrow> even k" | |
| 223 | by (simp add: even_int_iff [symmetric]) | |
| 224 | ||
| 225 | ||
| 58787 | 226 | subsection {* Parity and powers *}
 | 
| 58689 | 227 | |
| 228 | context comm_ring_1 | |
| 229 | begin | |
| 230 | ||
| 231 | lemma power_minus_even [simp]: | |
| 232 | "even n \<Longrightarrow> (- a) ^ n = a ^ n" | |
| 58690 | 233 | by (auto elim: evenE) | 
| 58689 | 234 | |
| 235 | lemma power_minus_odd [simp]: | |
| 236 | "odd n \<Longrightarrow> (- a) ^ n = - (a ^ n)" | |
| 58690 | 237 | by (auto elim: oddE) | 
| 238 | ||
| 58689 | 239 | lemma neg_one_even_power [simp]: | 
| 240 | "even n \<Longrightarrow> (- 1) ^ n = 1" | |
| 58690 | 241 | by simp | 
| 58689 | 242 | |
| 243 | lemma neg_one_odd_power [simp]: | |
| 244 | "odd n \<Longrightarrow> (- 1) ^ n = - 1" | |
| 58690 | 245 | by simp | 
| 58689 | 246 | |
| 247 | end | |
| 248 | ||
| 249 | context linordered_idom | |
| 250 | begin | |
| 251 | ||
| 252 | lemma zero_le_even_power: | |
| 253 | "even n \<Longrightarrow> 0 \<le> a ^ n" | |
| 58690 | 254 | by (auto elim: evenE) | 
| 58689 | 255 | |
| 256 | lemma zero_le_odd_power: | |
| 257 | "odd n \<Longrightarrow> 0 \<le> a ^ n \<longleftrightarrow> 0 \<le> a" | |
| 258 | by (auto simp add: power_even_eq zero_le_mult_iff elim: oddE) | |
| 259 | ||
| 58770 | 260 | lemma zero_le_power_eq: | 
| 58689 | 261 | "0 \<le> a ^ n \<longleftrightarrow> even n \<or> odd n \<and> 0 \<le> a" | 
| 58787 | 262 | by (auto simp add: zero_le_even_power zero_le_odd_power) | 
| 263 | ||
| 58770 | 264 | lemma zero_less_power_eq: | 
| 58689 | 265 | "0 < a ^ n \<longleftrightarrow> n = 0 \<or> even n \<and> a \<noteq> 0 \<or> odd n \<and> 0 < a" | 
| 266 | proof - | |
| 267 | have [simp]: "0 = a ^ n \<longleftrightarrow> a = 0 \<and> n > 0" | |
| 58787 | 268 | unfolding power_eq_0_iff [of a n, symmetric] by blast | 
| 58689 | 269 | show ?thesis | 
| 58710 
7216a10d69ba
augmented and tuned facts on even/odd and division
 haftmann parents: 
58709diff
changeset | 270 | unfolding less_le zero_le_power_eq by auto | 
| 58689 | 271 | qed | 
| 272 | ||
| 58787 | 273 | lemma power_less_zero_eq [simp]: | 
| 58689 | 274 | "a ^ n < 0 \<longleftrightarrow> odd n \<and> a < 0" | 
| 275 | unfolding not_le [symmetric] zero_le_power_eq by auto | |
| 276 | ||
| 58770 | 277 | lemma power_le_zero_eq: | 
| 58689 | 278 | "a ^ n \<le> 0 \<longleftrightarrow> n > 0 \<and> (odd n \<and> a \<le> 0 \<or> even n \<and> a = 0)" | 
| 279 | unfolding not_less [symmetric] zero_less_power_eq by auto | |
| 280 | ||
| 281 | lemma power_even_abs: | |
| 282 | "even n \<Longrightarrow> \<bar>a\<bar> ^ n = a ^ n" | |
| 283 | using power_abs [of a n] by (simp add: zero_le_even_power) | |
| 284 | ||
| 285 | lemma power_mono_even: | |
| 286 | assumes "even n" and "\<bar>a\<bar> \<le> \<bar>b\<bar>" | |
| 287 | shows "a ^ n \<le> b ^ n" | |
| 288 | proof - | |
| 289 | have "0 \<le> \<bar>a\<bar>" by auto | |
| 290 | with `\<bar>a\<bar> \<le> \<bar>b\<bar>` | |
| 291 | have "\<bar>a\<bar> ^ n \<le> \<bar>b\<bar> ^ n" by (rule power_mono) | |
| 292 | with `even n` show ?thesis by (simp add: power_even_abs) | |
| 293 | qed | |
| 294 | ||
| 295 | lemma power_mono_odd: | |
| 296 | assumes "odd n" and "a \<le> b" | |
| 297 | shows "a ^ n \<le> b ^ n" | |
| 298 | proof (cases "b < 0") | |
| 299 | case True with `a \<le> b` have "- b \<le> - a" and "0 \<le> - b" by auto | |
| 300 | hence "(- b) ^ n \<le> (- a) ^ n" by (rule power_mono) | |
| 301 | with `odd n` show ?thesis by simp | |
| 302 | next | |
| 303 | case False then have "0 \<le> b" by auto | |
| 304 | show ?thesis | |
| 305 | proof (cases "a < 0") | |
| 306 | case True then have "n \<noteq> 0" and "a \<le> 0" using `odd n` [THEN odd_pos] by auto | |
| 307 | then have "a ^ n \<le> 0" unfolding power_le_zero_eq using `odd n` by auto | |
| 308 | moreover | |
| 309 | from `0 \<le> b` have "0 \<le> b ^ n" by auto | |
| 310 | ultimately show ?thesis by auto | |
| 311 | next | |
| 312 | case False then have "0 \<le> a" by auto | |
| 313 | with `a \<le> b` show ?thesis using power_mono by auto | |
| 314 | qed | |
| 315 | qed | |
| 316 | ||
| 317 | text {* Simplify, when the exponent is a numeral *}
 | |
| 318 | ||
| 319 | lemma zero_le_power_eq_numeral [simp]: | |
| 320 | "0 \<le> a ^ numeral w \<longleftrightarrow> even (numeral w :: nat) \<or> odd (numeral w :: nat) \<and> 0 \<le> a" | |
| 321 | by (fact zero_le_power_eq) | |
| 322 | ||
| 323 | lemma zero_less_power_eq_numeral [simp]: | |
| 324 | "0 < a ^ numeral w \<longleftrightarrow> numeral w = (0 :: nat) | |
| 325 | \<or> even (numeral w :: nat) \<and> a \<noteq> 0 \<or> odd (numeral w :: nat) \<and> 0 < a" | |
| 326 | by (fact zero_less_power_eq) | |
| 327 | ||
| 328 | lemma power_le_zero_eq_numeral [simp]: | |
| 329 | "a ^ numeral w \<le> 0 \<longleftrightarrow> (0 :: nat) < numeral w | |
| 330 | \<and> (odd (numeral w :: nat) \<and> a \<le> 0 \<or> even (numeral w :: nat) \<and> a = 0)" | |
| 331 | by (fact power_le_zero_eq) | |
| 332 | ||
| 333 | lemma power_less_zero_eq_numeral [simp]: | |
| 334 | "a ^ numeral w < 0 \<longleftrightarrow> odd (numeral w :: nat) \<and> a < 0" | |
| 335 | by (fact power_less_zero_eq) | |
| 336 | ||
| 337 | lemma power_even_abs_numeral [simp]: | |
| 338 | "even (numeral w :: nat) \<Longrightarrow> \<bar>a\<bar> ^ numeral w = a ^ numeral w" | |
| 339 | by (fact power_even_abs) | |
| 340 | ||
| 341 | end | |
| 342 | ||
| 343 | ||
| 58687 | 344 | subsubsection {* Tools setup *}
 | 
| 345 | ||
| 58679 | 346 | declare transfer_morphism_int_nat [transfer add return: | 
| 347 | even_int_iff | |
| 33318 
ddd97d9dfbfb
moved Nat_Transfer before Divides; distributed Nat_Transfer setup accordingly
 haftmann parents: 
31718diff
changeset | 348 | ] | 
| 21256 | 349 | |
| 58770 | 350 | end | 
| 351 |