| author | wenzelm | 
| Tue, 14 Apr 2009 15:41:40 +0200 | |
| changeset 30890 | 0214d179c2be | 
| parent 30730 | 4d3565f2cb0e | 
| child 30960 | fec1a04b7220 | 
| permissions | -rw-r--r-- | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 1 | (* Title: HOL/Power.thy | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 4 | Copyright 1997 University of Cambridge | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 5 | |
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 6 | *) | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 7 | |
| 16733 
236dfafbeb63
linear arithmetic now takes "&" in assumptions apart.
 nipkow parents: 
15251diff
changeset | 8 | header{*Exponentiation*}
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 9 | |
| 15131 | 10 | theory Power | 
| 21413 | 11 | imports Nat | 
| 15131 | 12 | begin | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 13 | |
| 29608 | 14 | class power = | 
| 25062 | 15 | fixes power :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixr "^" 80) | 
| 24996 | 16 | |
| 21199 
2d83f93c3580
* Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
 krauss parents: 
17149diff
changeset | 17 | subsection{*Powers for Arbitrary Monoids*}
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 18 | |
| 22390 | 19 | class recpower = monoid_mult + power + | 
| 25062 | 20 | assumes power_0 [simp]: "a ^ 0 = 1" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 21 | assumes power_Suc [simp]: "a ^ Suc n = a * (a ^ n)" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 22 | |
| 21199 
2d83f93c3580
* Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
 krauss parents: 
17149diff
changeset | 23 | lemma power_0_Suc [simp]: "(0::'a::{recpower,semiring_0}) ^ (Suc n) = 0"
 | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 24 | by simp | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 25 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 26 | text{*It looks plausible as a simprule, but its effect can be strange.*}
 | 
| 21199 
2d83f93c3580
* Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
 krauss parents: 
17149diff
changeset | 27 | lemma power_0_left: "0^n = (if n=0 then 1 else (0::'a::{recpower,semiring_0}))"
 | 
| 23183 | 28 | by (induct n) simp_all | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 29 | |
| 15004 | 30 | lemma power_one [simp]: "1^n = (1::'a::recpower)" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 31 | by (induct n) simp_all | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 32 | |
| 15004 | 33 | lemma power_one_right [simp]: "(a::'a::recpower) ^ 1 = a" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 34 | unfolding One_nat_def by simp | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 35 | |
| 21199 
2d83f93c3580
* Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
 krauss parents: 
17149diff
changeset | 36 | lemma power_commutes: "(a::'a::recpower) ^ n * a = a * a ^ n" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 37 | by (induct n) (simp_all add: mult_assoc) | 
| 21199 
2d83f93c3580
* Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
 krauss parents: 
17149diff
changeset | 38 | |
| 28131 
3130d7b3149d
add lemma power_Suc2; generalize power_minus from class comm_ring_1 to ring_1
 huffman parents: 
25874diff
changeset | 39 | lemma power_Suc2: "(a::'a::recpower) ^ Suc n = a ^ n * a" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 40 | by (simp add: power_commutes) | 
| 28131 
3130d7b3149d
add lemma power_Suc2; generalize power_minus from class comm_ring_1 to ring_1
 huffman parents: 
25874diff
changeset | 41 | |
| 15004 | 42 | lemma power_add: "(a::'a::recpower) ^ (m+n) = (a^m) * (a^n)" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 43 | by (induct m) (simp_all add: mult_ac) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 44 | |
| 15004 | 45 | lemma power_mult: "(a::'a::recpower) ^ (m*n) = (a^m) ^ n" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 46 | by (induct n) (simp_all add: power_add) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 47 | |
| 21199 
2d83f93c3580
* Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
 krauss parents: 
17149diff
changeset | 48 | lemma power_mult_distrib: "((a::'a::{recpower,comm_monoid_mult}) * b) ^ n = (a^n) * (b^n)"
 | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 49 | by (induct n) (simp_all add: mult_ac) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 50 | |
| 25874 | 51 | lemma zero_less_power[simp]: | 
| 15004 | 52 |      "0 < (a::'a::{ordered_semidom,recpower}) ==> 0 < a^n"
 | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 53 | by (induct n) (simp_all add: mult_pos_pos) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 54 | |
| 25874 | 55 | lemma zero_le_power[simp]: | 
| 15004 | 56 |      "0 \<le> (a::'a::{ordered_semidom,recpower}) ==> 0 \<le> a^n"
 | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 57 | by (induct n) (simp_all add: mult_nonneg_nonneg) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 58 | |
| 25874 | 59 | lemma one_le_power[simp]: | 
| 15004 | 60 |      "1 \<le> (a::'a::{ordered_semidom,recpower}) ==> 1 \<le> a^n"
 | 
| 15251 | 61 | apply (induct "n") | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 62 | apply simp_all | 
| 14577 | 63 | apply (rule order_trans [OF _ mult_mono [of 1 _ 1]]) | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 64 | apply (simp_all add: order_trans [OF zero_le_one]) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 65 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 66 | |
| 14738 | 67 | lemma gt1_imp_ge0: "1 < a ==> 0 \<le> (a::'a::ordered_semidom)" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 68 | by (simp add: order_trans [OF zero_le_one order_less_imp_le]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 69 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 70 | lemma power_gt1_lemma: | 
| 15004 | 71 |   assumes gt1: "1 < (a::'a::{ordered_semidom,recpower})"
 | 
| 14577 | 72 | shows "1 < a * a^n" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 73 | proof - | 
| 14577 | 74 | have "1*1 < a*1" using gt1 by simp | 
| 75 | also have "\<dots> \<le> a * a^n" using gt1 | |
| 76 | by (simp only: mult_mono gt1_imp_ge0 one_le_power order_less_imp_le | |
| 77 | zero_le_one order_refl) | |
| 78 | finally show ?thesis by simp | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 79 | qed | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 80 | |
| 25874 | 81 | lemma one_less_power[simp]: | 
| 24376 | 82 |   "\<lbrakk>1 < (a::'a::{ordered_semidom,recpower}); 0 < n\<rbrakk> \<Longrightarrow> 1 < a ^ n"
 | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 83 | by (cases n, simp_all add: power_gt1_lemma) | 
| 24376 | 84 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 85 | lemma power_gt1: | 
| 15004 | 86 |      "1 < (a::'a::{ordered_semidom,recpower}) ==> 1 < a ^ (Suc n)"
 | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 87 | by (simp add: power_gt1_lemma) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 88 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 89 | lemma power_le_imp_le_exp: | 
| 15004 | 90 |   assumes gt1: "(1::'a::{recpower,ordered_semidom}) < a"
 | 
| 14577 | 91 | shows "!!n. a^m \<le> a^n ==> m \<le> n" | 
| 92 | proof (induct m) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 93 | case 0 | 
| 14577 | 94 | show ?case by simp | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 95 | next | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 96 | case (Suc m) | 
| 14577 | 97 | show ?case | 
| 98 | proof (cases n) | |
| 99 | case 0 | |
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 100 | from prems have "a * a^m \<le> 1" by simp | 
| 14577 | 101 | with gt1 show ?thesis | 
| 102 | by (force simp only: power_gt1_lemma | |
| 103 | linorder_not_less [symmetric]) | |
| 104 | next | |
| 105 | case (Suc n) | |
| 106 | from prems show ?thesis | |
| 107 | by (force dest: mult_left_le_imp_le | |
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 108 | simp add: order_less_trans [OF zero_less_one gt1]) | 
| 14577 | 109 | qed | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 110 | qed | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 111 | |
| 14577 | 112 | text{*Surely we can strengthen this? It holds for @{text "0<a<1"} too.*}
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 113 | lemma power_inject_exp [simp]: | 
| 15004 | 114 |      "1 < (a::'a::{ordered_semidom,recpower}) ==> (a^m = a^n) = (m=n)"
 | 
| 14577 | 115 | by (force simp add: order_antisym power_le_imp_le_exp) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 116 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 117 | text{*Can relax the first premise to @{term "0<a"} in the case of the
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 118 | natural numbers.*} | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 119 | lemma power_less_imp_less_exp: | 
| 15004 | 120 |      "[| (1::'a::{recpower,ordered_semidom}) < a; a^m < a^n |] ==> m < n"
 | 
| 14577 | 121 | by (simp add: order_less_le [of m n] order_less_le [of "a^m" "a^n"] | 
| 122 | power_le_imp_le_exp) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 123 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 124 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 125 | lemma power_mono: | 
| 15004 | 126 |      "[|a \<le> b; (0::'a::{recpower,ordered_semidom}) \<le> a|] ==> a^n \<le> b^n"
 | 
| 15251 | 127 | apply (induct "n") | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 128 | apply simp_all | 
| 25874 | 129 | apply (auto intro: mult_mono order_trans [of 0 a b]) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 130 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 131 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 132 | lemma power_strict_mono [rule_format]: | 
| 15004 | 133 |      "[|a < b; (0::'a::{recpower,ordered_semidom}) \<le> a|]
 | 
| 14577 | 134 | ==> 0 < n --> a^n < b^n" | 
| 15251 | 135 | apply (induct "n") | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 136 | apply (auto simp add: mult_strict_mono order_le_less_trans [of 0 a b]) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 137 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 138 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 139 | lemma power_eq_0_iff [simp]: | 
| 30056 | 140 | "(a^n = 0) \<longleftrightarrow> | 
| 141 |    (a = (0::'a::{mult_zero,zero_neq_one,no_zero_divisors,recpower}) & n\<noteq>0)"
 | |
| 15251 | 142 | apply (induct "n") | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 143 | apply (auto simp add: no_zero_divisors) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 144 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 145 | |
| 30056 | 146 | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 147 | lemma field_power_not_zero: | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 148 |   "a \<noteq> (0::'a::{ring_1_no_zero_divisors,recpower}) ==> a^n \<noteq> 0"
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 149 | by force | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 150 | |
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 151 | lemma nonzero_power_inverse: | 
| 22991 | 152 |   fixes a :: "'a::{division_ring,recpower}"
 | 
| 153 | shows "a \<noteq> 0 ==> inverse (a ^ n) = (inverse a) ^ n" | |
| 15251 | 154 | apply (induct "n") | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 155 | apply (auto simp add: nonzero_inverse_mult_distrib power_commutes) | 
| 22991 | 156 | done (* TODO: reorient or rename to nonzero_inverse_power *) | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 157 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 158 | text{*Perhaps these should be simprules.*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 159 | lemma power_inverse: | 
| 22991 | 160 |   fixes a :: "'a::{division_ring,division_by_zero,recpower}"
 | 
| 161 | shows "inverse (a ^ n) = (inverse a) ^ n" | |
| 162 | apply (cases "a = 0") | |
| 163 | apply (simp add: power_0_left) | |
| 164 | apply (simp add: nonzero_power_inverse) | |
| 165 | done (* TODO: reorient or rename to inverse_power *) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 166 | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16733diff
changeset | 167 | lemma power_one_over: "1 / (a::'a::{field,division_by_zero,recpower})^n = 
 | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16733diff
changeset | 168 | (1 / a)^n" | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16733diff
changeset | 169 | apply (simp add: divide_inverse) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16733diff
changeset | 170 | apply (rule power_inverse) | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16733diff
changeset | 171 | done | 
| 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
16733diff
changeset | 172 | |
| 14577 | 173 | lemma nonzero_power_divide: | 
| 15004 | 174 |     "b \<noteq> 0 ==> (a/b) ^ n = ((a::'a::{field,recpower}) ^ n) / (b ^ n)"
 | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 175 | by (simp add: divide_inverse power_mult_distrib nonzero_power_inverse) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 176 | |
| 14577 | 177 | lemma power_divide: | 
| 15004 | 178 |     "(a/b) ^ n = ((a::'a::{field,division_by_zero,recpower}) ^ n / b ^ n)"
 | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 179 | apply (case_tac "b=0", simp add: power_0_left) | 
| 14577 | 180 | apply (rule nonzero_power_divide) | 
| 181 | apply assumption | |
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 182 | done | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 183 | |
| 15004 | 184 | lemma power_abs: "abs(a ^ n) = abs(a::'a::{ordered_idom,recpower}) ^ n"
 | 
| 15251 | 185 | apply (induct "n") | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 186 | apply (auto simp add: abs_mult) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 187 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 188 | |
| 30730 | 189 | lemma abs_power_minus [simp]: | 
| 190 |   fixes a:: "'a::{ordered_idom,recpower}" shows "abs((-a) ^ n) = abs(a ^ n)"
 | |
| 191 | by (simp add: abs_minus_cancel power_abs) | |
| 192 | ||
| 24286 
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
 paulson parents: 
23544diff
changeset | 193 | lemma zero_less_power_abs_iff [simp,noatp]: | 
| 15004 | 194 |      "(0 < (abs a)^n) = (a \<noteq> (0::'a::{ordered_idom,recpower}) | n=0)"
 | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 195 | proof (induct "n") | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 196 | case 0 | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 197 | show ?case by simp | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 198 | next | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 199 | case (Suc n) | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 200 | show ?case by (auto simp add: prems zero_less_mult_iff) | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 201 | qed | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 202 | |
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 203 | lemma zero_le_power_abs [simp]: | 
| 15004 | 204 |      "(0::'a::{ordered_idom,recpower}) \<le> (abs a)^n"
 | 
| 22957 | 205 | by (rule zero_le_power [OF abs_ge_zero]) | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 206 | |
| 28131 
3130d7b3149d
add lemma power_Suc2; generalize power_minus from class comm_ring_1 to ring_1
 huffman parents: 
25874diff
changeset | 207 | lemma power_minus: "(-a) ^ n = (- 1)^n * (a::'a::{ring_1,recpower}) ^ n"
 | 
| 
3130d7b3149d
add lemma power_Suc2; generalize power_minus from class comm_ring_1 to ring_1
 huffman parents: 
25874diff
changeset | 208 | proof (induct n) | 
| 
3130d7b3149d
add lemma power_Suc2; generalize power_minus from class comm_ring_1 to ring_1
 huffman parents: 
25874diff
changeset | 209 | case 0 show ?case by simp | 
| 
3130d7b3149d
add lemma power_Suc2; generalize power_minus from class comm_ring_1 to ring_1
 huffman parents: 
25874diff
changeset | 210 | next | 
| 
3130d7b3149d
add lemma power_Suc2; generalize power_minus from class comm_ring_1 to ring_1
 huffman parents: 
25874diff
changeset | 211 | case (Suc n) then show ?case | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 212 | by (simp del: power_Suc add: power_Suc2 mult_assoc) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 213 | qed | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 214 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 215 | text{*Lemma for @{text power_strict_decreasing}*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 216 | lemma power_Suc_less: | 
| 15004 | 217 |      "[|(0::'a::{ordered_semidom,recpower}) < a; a < 1|]
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 218 | ==> a * a^n < a^n" | 
| 15251 | 219 | apply (induct n) | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 220 | apply (auto simp add: mult_strict_left_mono) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 221 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 222 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 223 | lemma power_strict_decreasing: | 
| 15004 | 224 |      "[|n < N; 0 < a; a < (1::'a::{ordered_semidom,recpower})|]
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 225 | ==> a^N < a^n" | 
| 14577 | 226 | apply (erule rev_mp) | 
| 15251 | 227 | apply (induct "N") | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 228 | apply (auto simp add: power_Suc_less less_Suc_eq) | 
| 14577 | 229 | apply (rename_tac m) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 230 | apply (subgoal_tac "a * a^m < 1 * a^n", simp) | 
| 14577 | 231 | apply (rule mult_strict_mono) | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 232 | apply (auto simp add: order_less_imp_le) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 233 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 234 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 235 | text{*Proof resembles that of @{text power_strict_decreasing}*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 236 | lemma power_decreasing: | 
| 15004 | 237 |      "[|n \<le> N; 0 \<le> a; a \<le> (1::'a::{ordered_semidom,recpower})|]
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 238 | ==> a^N \<le> a^n" | 
| 14577 | 239 | apply (erule rev_mp) | 
| 15251 | 240 | apply (induct "N") | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 241 | apply (auto simp add: le_Suc_eq) | 
| 14577 | 242 | apply (rename_tac m) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 243 | apply (subgoal_tac "a * a^m \<le> 1 * a^n", simp) | 
| 14577 | 244 | apply (rule mult_mono) | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 245 | apply auto | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 246 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 247 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 248 | lemma power_Suc_less_one: | 
| 15004 | 249 |      "[| 0 < a; a < (1::'a::{ordered_semidom,recpower}) |] ==> a ^ Suc n < 1"
 | 
| 14577 | 250 | apply (insert power_strict_decreasing [of 0 "Suc n" a], simp) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 251 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 252 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 253 | text{*Proof again resembles that of @{text power_strict_decreasing}*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 254 | lemma power_increasing: | 
| 15004 | 255 |      "[|n \<le> N; (1::'a::{ordered_semidom,recpower}) \<le> a|] ==> a^n \<le> a^N"
 | 
| 14577 | 256 | apply (erule rev_mp) | 
| 15251 | 257 | apply (induct "N") | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 258 | apply (auto simp add: le_Suc_eq) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 259 | apply (rename_tac m) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 260 | apply (subgoal_tac "1 * a^n \<le> a * a^m", simp) | 
| 14577 | 261 | apply (rule mult_mono) | 
| 25874 | 262 | apply (auto simp add: order_trans [OF zero_le_one]) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 263 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 264 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 265 | text{*Lemma for @{text power_strict_increasing}*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 266 | lemma power_less_power_Suc: | 
| 15004 | 267 |      "(1::'a::{ordered_semidom,recpower}) < a ==> a^n < a * a^n"
 | 
| 15251 | 268 | apply (induct n) | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 269 | apply (auto simp add: mult_strict_left_mono order_less_trans [OF zero_less_one]) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 270 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 271 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 272 | lemma power_strict_increasing: | 
| 15004 | 273 |      "[|n < N; (1::'a::{ordered_semidom,recpower}) < a|] ==> a^n < a^N"
 | 
| 14577 | 274 | apply (erule rev_mp) | 
| 15251 | 275 | apply (induct "N") | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 276 | apply (auto simp add: power_less_power_Suc less_Suc_eq) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 277 | apply (rename_tac m) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 278 | apply (subgoal_tac "1 * a^n < a * a^m", simp) | 
| 14577 | 279 | apply (rule mult_strict_mono) | 
| 25874 | 280 | apply (auto simp add: order_less_trans [OF zero_less_one] order_less_imp_le) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 281 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 282 | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 283 | lemma power_increasing_iff [simp]: | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 284 |   "1 < (b::'a::{ordered_semidom,recpower}) ==> (b ^ x \<le> b ^ y) = (x \<le> y)"
 | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 285 | by (blast intro: power_le_imp_le_exp power_increasing order_less_imp_le) | 
| 15066 | 286 | |
| 287 | lemma power_strict_increasing_iff [simp]: | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 288 |   "1 < (b::'a::{ordered_semidom,recpower}) ==> (b ^ x < b ^ y) = (x < y)"
 | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 289 | by (blast intro: power_less_imp_less_exp power_strict_increasing) | 
| 15066 | 290 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 291 | lemma power_le_imp_le_base: | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 292 | assumes le: "a ^ Suc n \<le> b ^ Suc n" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 293 |     and ynonneg: "(0::'a::{ordered_semidom,recpower}) \<le> b"
 | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 294 | shows "a \<le> b" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 295 | proof (rule ccontr) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 296 | assume "~ a \<le> b" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 297 | then have "b < a" by (simp only: linorder_not_le) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 298 | then have "b ^ Suc n < a ^ Suc n" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 299 | by (simp only: prems power_strict_mono) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 300 | from le and this show "False" | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 301 | by (simp add: linorder_not_less [symmetric]) | 
| 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25062diff
changeset | 302 | qed | 
| 14577 | 303 | |
| 22853 | 304 | lemma power_less_imp_less_base: | 
| 305 |   fixes a b :: "'a::{ordered_semidom,recpower}"
 | |
| 306 | assumes less: "a ^ n < b ^ n" | |
| 307 | assumes nonneg: "0 \<le> b" | |
| 308 | shows "a < b" | |
| 309 | proof (rule contrapos_pp [OF less]) | |
| 310 | assume "~ a < b" | |
| 311 | hence "b \<le> a" by (simp only: linorder_not_less) | |
| 312 | hence "b ^ n \<le> a ^ n" using nonneg by (rule power_mono) | |
| 313 | thus "~ a ^ n < b ^ n" by (simp only: linorder_not_less) | |
| 314 | qed | |
| 315 | ||
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 316 | lemma power_inject_base: | 
| 14577 | 317 | "[| a ^ Suc n = b ^ Suc n; 0 \<le> a; 0 \<le> b |] | 
| 15004 | 318 |       ==> a = (b::'a::{ordered_semidom,recpower})"
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 319 | by (blast intro: power_le_imp_le_base order_antisym order_eq_refl sym) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 320 | |
| 22955 | 321 | lemma power_eq_imp_eq_base: | 
| 322 |   fixes a b :: "'a::{ordered_semidom,recpower}"
 | |
| 323 | shows "\<lbrakk>a ^ n = b ^ n; 0 \<le> a; 0 \<le> b; 0 < n\<rbrakk> \<Longrightarrow> a = b" | |
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 324 | by (cases n, simp_all del: power_Suc, rule power_inject_base) | 
| 22955 | 325 | |
| 29978 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 326 | text {* The divides relation *}
 | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 327 | |
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 328 | lemma le_imp_power_dvd: | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 329 |   fixes a :: "'a::{comm_semiring_1,recpower}"
 | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 330 | assumes "m \<le> n" shows "a^m dvd a^n" | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 331 | proof | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 332 | have "a^n = a^(m + (n - m))" | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 333 | using `m \<le> n` by simp | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 334 | also have "\<dots> = a^m * a^(n - m)" | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 335 | by (rule power_add) | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 336 | finally show "a^n = a^m * a^(n - m)" . | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 337 | qed | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 338 | |
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 339 | lemma power_le_dvd: | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 340 |   fixes a b :: "'a::{comm_semiring_1,recpower}"
 | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 341 | shows "a^n dvd b \<Longrightarrow> m \<le> n \<Longrightarrow> a^m dvd b" | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 342 | by (rule dvd_trans [OF le_imp_power_dvd]) | 
| 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 huffman parents: 
29608diff
changeset | 343 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 344 | |
| 30313 | 345 | lemma dvd_power_same: | 
| 346 |   "(x::'a::{comm_semiring_1,recpower}) dvd y \<Longrightarrow> x^n dvd y^n"
 | |
| 347 | by (induct n) (auto simp add: mult_dvd_mono) | |
| 348 | ||
| 349 | lemma dvd_power_le: | |
| 350 |   "(x::'a::{comm_semiring_1,recpower}) dvd y \<Longrightarrow> m >= n \<Longrightarrow> x^n dvd y^m"
 | |
| 351 | by(rule power_le_dvd[OF dvd_power_same]) | |
| 352 | ||
| 353 | lemma dvd_power [simp]: | |
| 354 |   "n > 0 | (x::'a::{comm_semiring_1,recpower}) = 1 \<Longrightarrow> x dvd x^n"
 | |
| 355 | apply (erule disjE) | |
| 356 | apply (subgoal_tac "x ^ n = x^(Suc (n - 1))") | |
| 357 | apply (erule ssubst) | |
| 358 | apply (subst power_Suc) | |
| 359 | apply auto | |
| 360 | done | |
| 361 | ||
| 362 | ||
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 363 | subsection{*Exponentiation for the Natural Numbers*}
 | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 364 | |
| 25836 | 365 | instantiation nat :: recpower | 
| 366 | begin | |
| 21456 | 367 | |
| 25836 | 368 | primrec power_nat where | 
| 369 | "p ^ 0 = (1\<Colon>nat)" | |
| 370 | | "p ^ (Suc n) = (p\<Colon>nat) * (p ^ n)" | |
| 14577 | 371 | |
| 25836 | 372 | instance proof | 
| 14438 | 373 | fix z n :: nat | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 374 | show "z^0 = 1" by simp | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 375 | show "z^(Suc n) = z * (z^n)" by simp | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 376 | qed | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 377 | |
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 378 | declare power_nat.simps [simp del] | 
| 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 379 | |
| 25836 | 380 | end | 
| 381 | ||
| 23305 | 382 | lemma of_nat_power: | 
| 383 |   "of_nat (m ^ n) = (of_nat m::'a::{semiring_1,recpower}) ^ n"
 | |
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 384 | by (induct n, simp_all add: of_nat_mult) | 
| 23305 | 385 | |
| 30079 
293b896b9c25
make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
 huffman parents: 
30056diff
changeset | 386 | lemma nat_one_le_power [simp]: "Suc 0 \<le> i ==> Suc 0 \<le> i^n" | 
| 
293b896b9c25
make proofs work whether or not One_nat_def is a simp rule; replace 1 with Suc 0 in the rhs of some simp rules
 huffman parents: 
30056diff
changeset | 387 | by (rule one_le_power [of i n, unfolded One_nat_def]) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 388 | |
| 25162 | 389 | lemma nat_zero_less_power_iff [simp]: "(x^n > 0) = (x > (0::nat) | n=0)" | 
| 21413 | 390 | by (induct "n", auto) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 391 | |
| 30056 | 392 | lemma nat_power_eq_Suc_0_iff [simp]: | 
| 393 | "((x::nat)^m = Suc 0) = (m = 0 | x = Suc 0)" | |
| 394 | by (induct_tac m, auto) | |
| 395 | ||
| 396 | lemma power_Suc_0[simp]: "(Suc 0)^n = Suc 0" | |
| 397 | by simp | |
| 398 | ||
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 399 | text{*Valid for the naturals, but what if @{text"0<i<1"}?
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 400 | Premises cannot be weakened: consider the case where @{term "i=0"},
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 401 | @{term "m=1"} and @{term "n=0"}.*}
 | 
| 21413 | 402 | lemma nat_power_less_imp_less: | 
| 403 | assumes nonneg: "0 < (i\<Colon>nat)" | |
| 404 | assumes less: "i^m < i^n" | |
| 405 | shows "m < n" | |
| 406 | proof (cases "i = 1") | |
| 407 | case True with less power_one [where 'a = nat] show ?thesis by simp | |
| 408 | next | |
| 409 | case False with nonneg have "1 < i" by auto | |
| 410 | from power_strict_increasing_iff [OF this] less show ?thesis .. | |
| 411 | qed | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 412 | |
| 17149 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
16796diff
changeset | 413 | lemma power_diff: | 
| 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
16796diff
changeset | 414 | assumes nz: "a ~= 0" | 
| 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
16796diff
changeset | 415 |   shows "n <= m ==> (a::'a::{recpower, field}) ^ (m-n) = (a^m) / (a^n)"
 | 
| 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
16796diff
changeset | 416 | by (induct m n rule: diff_induct) | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 417 | (simp_all add: nonzero_mult_divide_cancel_left nz) | 
| 17149 
e2b19c92ef51
Lemmas on dvd, power and finite summation added or strengthened.
 ballarin parents: 
16796diff
changeset | 418 | |
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 419 | end |