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