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