| author | wenzelm | 
| Thu, 11 Apr 2024 12:05:01 +0200 | |
| changeset 80109 | dbcd6dc7f70f | 
| parent 79566 | f783490c6c99 | 
| child 80932 | 261cd8722677 | 
| 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  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Copyright 1997 University of Cambridge  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
4  | 
*)  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
5  | 
|
| 60758 | 6  | 
section \<open>Exponentiation\<close>  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
7  | 
|
| 15131 | 8  | 
theory Power  | 
| 63654 | 9  | 
imports Num  | 
| 15131 | 10  | 
begin  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
11  | 
|
| 60758 | 12  | 
subsection \<open>Powers for Arbitrary Monoids\<close>  | 
| 30960 | 13  | 
|
| 30996 | 14  | 
class power = one + times  | 
| 30960 | 15  | 
begin  | 
| 24996 | 16  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61944 
diff
changeset
 | 
17  | 
primrec power :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infixr "^" 80)  | 
| 63654 | 18  | 
where  | 
19  | 
power_0: "a ^ 0 = 1"  | 
|
20  | 
| power_Suc: "a ^ Suc n = a * a ^ n"  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
21  | 
|
| 30996 | 22  | 
notation (latex output)  | 
23  | 
  power ("(_\<^bsup>_\<^esup>)" [1000] 1000)
 | 
|
24  | 
||
| 60758 | 25  | 
text \<open>Special syntax for squares.\<close>  | 
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61944 
diff
changeset
 | 
26  | 
abbreviation power2 :: "'a \<Rightarrow> 'a"  ("(_\<^sup>2)" [1000] 999)
 | 
| 
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61944 
diff
changeset
 | 
27  | 
where "x\<^sup>2 \<equiv> x ^ 2"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
28  | 
|
| 30960 | 29  | 
end  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
30  | 
|
| 70928 | 31  | 
context  | 
32  | 
includes lifting_syntax  | 
|
33  | 
begin  | 
|
34  | 
||
35  | 
lemma power_transfer [transfer_rule]:  | 
|
36  | 
\<open>(R ===> (=) ===> R) (^) (^)\<close>  | 
|
37  | 
if [transfer_rule]: \<open>R 1 1\<close>  | 
|
38  | 
\<open>(R ===> R ===> R) (*) (*)\<close>  | 
|
39  | 
for R :: \<open>'a::power \<Rightarrow> 'b::power \<Rightarrow> bool\<close>  | 
|
40  | 
by (simp only: power_def [abs_def]) transfer_prover  | 
|
41  | 
||
42  | 
end  | 
|
43  | 
||
| 30996 | 44  | 
context monoid_mult  | 
45  | 
begin  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
46  | 
|
| 
39438
 
c5ece2a7a86e
Isar "default" step needs to fail for solved problems, for clear distinction of '.' and '..' for example -- amending lapse introduced in 9de4d64eee3b (April 2004);
 
wenzelm 
parents: 
36409 
diff
changeset
 | 
47  | 
subclass power .  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
48  | 
|
| 63654 | 49  | 
lemma power_one [simp]: "1 ^ n = 1"  | 
| 
30273
 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 
huffman 
parents: 
30242 
diff
changeset
 | 
50  | 
by (induct n) simp_all  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
51  | 
|
| 63654 | 52  | 
lemma power_one_right [simp]: "a ^ 1 = a"  | 
| 30996 | 53  | 
by simp  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
54  | 
|
| 63654 | 55  | 
lemma power_Suc0_right [simp]: "a ^ Suc 0 = a"  | 
| 
59741
 
5b762cd73a8e
Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
 
paulson <lp15@cam.ac.uk> 
parents: 
59009 
diff
changeset
 | 
56  | 
by simp  | 
| 
 
5b762cd73a8e
Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
 
paulson <lp15@cam.ac.uk> 
parents: 
59009 
diff
changeset
 | 
57  | 
|
| 63654 | 58  | 
lemma power_commutes: "a ^ n * a = a * a ^ n"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57418 
diff
changeset
 | 
59  | 
by (induct n) (simp_all add: mult.assoc)  | 
| 
21199
 
2d83f93c3580
* Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
 
krauss 
parents: 
17149 
diff
changeset
 | 
60  | 
|
| 63654 | 61  | 
lemma power_Suc2: "a ^ Suc n = a ^ n * a"  | 
| 
30273
 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 
huffman 
parents: 
30242 
diff
changeset
 | 
62  | 
by (simp add: power_commutes)  | 
| 
28131
 
3130d7b3149d
add lemma power_Suc2; generalize power_minus from class comm_ring_1 to ring_1
 
huffman 
parents: 
25874 
diff
changeset
 | 
63  | 
|
| 63654 | 64  | 
lemma power_add: "a ^ (m + n) = a ^ m * a ^ n"  | 
| 30996 | 65  | 
by (induct m) (simp_all add: algebra_simps)  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
66  | 
|
| 63654 | 67  | 
lemma power_mult: "a ^ (m * n) = (a ^ m) ^ n"  | 
| 
30273
 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 
huffman 
parents: 
30242 
diff
changeset
 | 
68  | 
by (induct n) (simp_all add: power_add)  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
69  | 
|
| 63654 | 70  | 
lemma power_even_eq: "a ^ (2 * n) = (a ^ n)\<^sup>2"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57418 
diff
changeset
 | 
71  | 
by (subst mult.commute) (simp add: power_mult)  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
72  | 
|
| 63654 | 73  | 
lemma power_odd_eq: "a ^ Suc (2*n) = a * (a ^ n)\<^sup>2"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
74  | 
by (simp add: power_even_eq)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
75  | 
|
| 63654 | 76  | 
lemma power_numeral_even: "z ^ numeral (Num.Bit0 w) = (let w = z ^ (numeral w) in w * w)"  | 
77  | 
by (simp only: numeral_Bit0 power_add Let_def)  | 
|
| 
47255
 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 
huffman 
parents: 
47241 
diff
changeset
 | 
78  | 
|
| 63654 | 79  | 
lemma power_numeral_odd: "z ^ numeral (Num.Bit1 w) = (let w = z ^ (numeral w) in z * w * w)"  | 
80  | 
by (simp only: numeral_Bit1 One_nat_def add_Suc_right add_0_right  | 
|
81  | 
power_Suc power_add Let_def mult.assoc)  | 
|
| 
47255
 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 
huffman 
parents: 
47241 
diff
changeset
 | 
82  | 
|
| 70724 | 83  | 
lemma power2_eq_square: "a\<^sup>2 = a * a"  | 
84  | 
by (simp add: numeral_2_eq_2)  | 
|
85  | 
||
86  | 
lemma power3_eq_cube: "a ^ 3 = a * a * a"  | 
|
87  | 
by (simp add: numeral_3_eq_3 mult.assoc)  | 
|
88  | 
||
89  | 
lemma power4_eq_xxxx: "x^4 = x * x * x * x"  | 
|
90  | 
by (simp add: mult.assoc power_numeral_even)  | 
|
91  | 
||
| 
77140
 
9a60c1759543
Lots more new material thanks to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
77138 
diff
changeset
 | 
92  | 
lemma power_numeral_reduce: "x ^ numeral n = x * x ^ pred_numeral n"  | 
| 
 
9a60c1759543
Lots more new material thanks to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
77138 
diff
changeset
 | 
93  | 
by (simp add: numeral_eq_Suc)  | 
| 
 
9a60c1759543
Lots more new material thanks to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
77138 
diff
changeset
 | 
94  | 
|
| 63654 | 95  | 
lemma funpow_times_power: "(times x ^^ f x) = times (x ^ f x)"  | 
| 49824 | 96  | 
proof (induct "f x" arbitrary: f)  | 
| 63654 | 97  | 
case 0  | 
98  | 
then show ?case by (simp add: fun_eq_iff)  | 
|
| 49824 | 99  | 
next  | 
100  | 
case (Suc n)  | 
|
| 63040 | 101  | 
define g where "g x = f x - 1" for x  | 
| 49824 | 102  | 
with Suc have "n = g x" by simp  | 
103  | 
with Suc have "times x ^^ g x = times (x ^ g x)" by simp  | 
|
104  | 
moreover from Suc g_def have "f x = g x + 1" by simp  | 
|
| 63654 | 105  | 
ultimately show ?case  | 
106  | 
by (simp add: power_add funpow_add fun_eq_iff mult.assoc)  | 
|
| 49824 | 107  | 
qed  | 
108  | 
||
| 58656 | 109  | 
lemma power_commuting_commutes:  | 
110  | 
assumes "x * y = y * x"  | 
|
111  | 
shows "x ^ n * y = y * x ^n"  | 
|
112  | 
proof (induct n)  | 
|
| 63654 | 113  | 
case 0  | 
114  | 
then show ?case by simp  | 
|
115  | 
next  | 
|
| 58656 | 116  | 
case (Suc n)  | 
117  | 
have "x ^ Suc n * y = x ^ n * y * x"  | 
|
118  | 
by (subst power_Suc2) (simp add: assms ac_simps)  | 
|
119  | 
also have "\<dots> = y * x ^ Suc n"  | 
|
| 63654 | 120  | 
by (simp only: Suc power_Suc2) (simp add: ac_simps)  | 
| 58656 | 121  | 
finally show ?case .  | 
| 63654 | 122  | 
qed  | 
| 58656 | 123  | 
|
| 63654 | 124  | 
lemma power_minus_mult: "0 < n \<Longrightarrow> a ^ (n - 1) * a = a ^ n"  | 
| 63648 | 125  | 
by (simp add: power_commutes split: nat_diff_split)  | 
| 62347 | 126  | 
|
| 
69791
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
127  | 
lemma left_right_inverse_power:  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
128  | 
assumes "x * y = 1"  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
129  | 
shows "x ^ n * y ^ n = 1"  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
130  | 
proof (induct n)  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
131  | 
case (Suc n)  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
132  | 
moreover have "x ^ Suc n * y ^ Suc n = x^n * (x * y) * y^n"  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
133  | 
by (simp add: power_Suc2[symmetric] mult.assoc[symmetric])  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
134  | 
ultimately show ?case by (simp add: assms)  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
135  | 
qed simp  | 
| 
 
195aeee8b30a
Formal Laurent series and overhaul of Formal power series (due to Jeremy Sylvestre)
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69700 
diff
changeset
 | 
136  | 
|
| 30996 | 137  | 
end  | 
138  | 
||
139  | 
context comm_monoid_mult  | 
|
140  | 
begin  | 
|
141  | 
||
| 
70817
 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 
haftmann 
parents: 
70724 
diff
changeset
 | 
142  | 
lemma power_mult_distrib [algebra_simps, algebra_split_simps, field_simps, field_split_simps, divide_simps]:  | 
| 
 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 
haftmann 
parents: 
70724 
diff
changeset
 | 
143  | 
"(a * b) ^ n = (a ^ n) * (b ^ n)"  | 
| 
 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 
haftmann 
parents: 
70724 
diff
changeset
 | 
144  | 
by (induction n) (simp_all add: ac_simps)  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
145  | 
|
| 30996 | 146  | 
end  | 
147  | 
||
| 63654 | 148  | 
text \<open>Extract constant factors from powers.\<close>  | 
| 
59741
 
5b762cd73a8e
Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
 
paulson <lp15@cam.ac.uk> 
parents: 
59009 
diff
changeset
 | 
149  | 
declare power_mult_distrib [where a = "numeral w" for w, simp]  | 
| 
 
5b762cd73a8e
Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
 
paulson <lp15@cam.ac.uk> 
parents: 
59009 
diff
changeset
 | 
150  | 
declare power_mult_distrib [where b = "numeral w" for w, simp]  | 
| 
 
5b762cd73a8e
Lots of new material on complex-valued functions. Modified simplification of (x/n)^k
 
paulson <lp15@cam.ac.uk> 
parents: 
59009 
diff
changeset
 | 
151  | 
|
| 63654 | 152  | 
lemma power_add_numeral [simp]: "a^numeral m * a^numeral n = a^numeral (m + n)"  | 
153  | 
for a :: "'a::monoid_mult"  | 
|
| 
60155
 
91477b3a2d6b
Tidying. Improved simplification for numerals, esp in exponents.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
154  | 
by (simp add: power_add [symmetric])  | 
| 
 
91477b3a2d6b
Tidying. Improved simplification for numerals, esp in exponents.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
155  | 
|
| 63654 | 156  | 
lemma power_add_numeral2 [simp]: "a^numeral m * (a^numeral n * b) = a^numeral (m + n) * b"  | 
157  | 
for a :: "'a::monoid_mult"  | 
|
| 
60155
 
91477b3a2d6b
Tidying. Improved simplification for numerals, esp in exponents.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
158  | 
by (simp add: mult.assoc [symmetric])  | 
| 
 
91477b3a2d6b
Tidying. Improved simplification for numerals, esp in exponents.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
159  | 
|
| 63654 | 160  | 
lemma power_mult_numeral [simp]: "(a^numeral m)^numeral n = a^numeral (m * n)"  | 
161  | 
for a :: "'a::monoid_mult"  | 
|
| 
60155
 
91477b3a2d6b
Tidying. Improved simplification for numerals, esp in exponents.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
162  | 
by (simp only: numeral_mult power_mult)  | 
| 
 
91477b3a2d6b
Tidying. Improved simplification for numerals, esp in exponents.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
163  | 
|
| 47191 | 164  | 
context semiring_numeral  | 
165  | 
begin  | 
|
166  | 
||
167  | 
lemma numeral_sqr: "numeral (Num.sqr k) = numeral k * numeral k"  | 
|
168  | 
by (simp only: sqr_conv_mult numeral_mult)  | 
|
169  | 
||
170  | 
lemma numeral_pow: "numeral (Num.pow k l) = numeral k ^ numeral l"  | 
|
| 63654 | 171  | 
by (induct l)  | 
172  | 
(simp_all only: numeral_class.numeral.simps pow.simps  | 
|
173  | 
numeral_sqr numeral_mult power_add power_one_right)  | 
|
| 47191 | 174  | 
|
175  | 
lemma power_numeral [simp]: "numeral k ^ numeral l = numeral (Num.pow k l)"  | 
|
176  | 
by (rule numeral_pow [symmetric])  | 
|
177  | 
||
178  | 
end  | 
|
179  | 
||
| 30996 | 180  | 
context semiring_1  | 
181  | 
begin  | 
|
182  | 
||
| 63654 | 183  | 
lemma of_nat_power [simp]: "of_nat (m ^ n) = of_nat m ^ n"  | 
| 
63417
 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 
haftmann 
parents: 
63040 
diff
changeset
 | 
184  | 
by (induct n) simp_all  | 
| 30996 | 185  | 
|
| 63654 | 186  | 
lemma zero_power: "0 < n \<Longrightarrow> 0 ^ n = 0"  | 
| 
59009
 
348561aa3869
generalized lemmas (particularly concerning dvd) as far as appropriate
 
haftmann 
parents: 
58889 
diff
changeset
 | 
187  | 
by (cases n) simp_all  | 
| 
 
348561aa3869
generalized lemmas (particularly concerning dvd) as far as appropriate
 
haftmann 
parents: 
58889 
diff
changeset
 | 
188  | 
|
| 63654 | 189  | 
lemma power_zero_numeral [simp]: "0 ^ numeral k = 0"  | 
| 
47209
 
4893907fe872
add constant pred_numeral k = numeral k - (1::nat);
 
huffman 
parents: 
47192 
diff
changeset
 | 
190  | 
by (simp add: numeral_eq_Suc)  | 
| 47191 | 191  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
192  | 
lemma zero_power2: "0\<^sup>2 = 0" (* delete? *)  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
193  | 
by (rule power_zero_numeral)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
194  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
195  | 
lemma one_power2: "1\<^sup>2 = 1" (* delete? *)  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
196  | 
by (rule power_one)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
197  | 
|
| 63654 | 198  | 
lemma power_0_Suc [simp]: "0 ^ Suc n = 0"  | 
| 60867 | 199  | 
by simp  | 
200  | 
||
| 63654 | 201  | 
text \<open>It looks plausible as a simprule, but its effect can be strange.\<close>  | 
202  | 
lemma power_0_left: "0 ^ n = (if n = 0 then 1 else 0)"  | 
|
| 60867 | 203  | 
by (cases n) simp_all  | 
204  | 
||
| 30996 | 205  | 
end  | 
206  | 
||
| 
66912
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
207  | 
context semiring_char_0 begin  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
208  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
209  | 
lemma numeral_power_eq_of_nat_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
210  | 
"numeral x ^ n = of_nat y \<longleftrightarrow> numeral x ^ n = y"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
211  | 
using of_nat_eq_iff by fastforce  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
212  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
213  | 
lemma real_of_nat_eq_numeral_power_cancel_iff [simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
214  | 
"of_nat y = numeral x ^ n \<longleftrightarrow> y = numeral x ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
215  | 
using numeral_power_eq_of_nat_cancel_iff [of x n y] by (metis (mono_tags))  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
216  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
217  | 
lemma of_nat_eq_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w = of_nat x \<longleftrightarrow> b ^ w = x"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
218  | 
by (metis of_nat_power of_nat_eq_iff)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
219  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
220  | 
lemma of_nat_power_eq_of_nat_cancel_iff[simp]: "of_nat x = (of_nat b) ^ w \<longleftrightarrow> x = b ^ w"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
221  | 
by (metis of_nat_eq_of_nat_power_cancel_iff)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
222  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
223  | 
end  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
224  | 
|
| 30996 | 225  | 
context comm_semiring_1  | 
226  | 
begin  | 
|
227  | 
||
| 63654 | 228  | 
text \<open>The divides relation.\<close>  | 
| 30996 | 229  | 
|
230  | 
lemma le_imp_power_dvd:  | 
|
| 63654 | 231  | 
assumes "m \<le> n"  | 
232  | 
shows "a ^ m dvd a ^ n"  | 
|
| 30996 | 233  | 
proof  | 
| 63654 | 234  | 
from assms have "a ^ n = a ^ (m + (n - m))" by simp  | 
235  | 
also have "\<dots> = a ^ m * a ^ (n - m)" by (rule power_add)  | 
|
| 30996 | 236  | 
finally show "a ^ n = a ^ m * a ^ (n - m)" .  | 
237  | 
qed  | 
|
238  | 
||
| 63654 | 239  | 
lemma power_le_dvd: "a ^ n dvd b \<Longrightarrow> m \<le> n \<Longrightarrow> a ^ m dvd b"  | 
| 30996 | 240  | 
by (rule dvd_trans [OF le_imp_power_dvd])  | 
241  | 
||
| 63654 | 242  | 
lemma dvd_power_same: "x dvd y \<Longrightarrow> x ^ n dvd y ^ n"  | 
| 30996 | 243  | 
by (induct n) (auto simp add: mult_dvd_mono)  | 
244  | 
||
| 63654 | 245  | 
lemma dvd_power_le: "x dvd y \<Longrightarrow> m \<ge> n \<Longrightarrow> x ^ n dvd y ^ m"  | 
| 30996 | 246  | 
by (rule power_le_dvd [OF dvd_power_same])  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
247  | 
|
| 30996 | 248  | 
lemma dvd_power [simp]:  | 
| 63654 | 249  | 
fixes n :: nat  | 
250  | 
assumes "n > 0 \<or> x = 1"  | 
|
| 30996 | 251  | 
shows "x dvd (x ^ n)"  | 
| 63654 | 252  | 
using assms  | 
253  | 
proof  | 
|
| 30996 | 254  | 
assume "0 < n"  | 
255  | 
then have "x ^ n = x ^ Suc (n - 1)" by simp  | 
|
256  | 
then show "x dvd (x ^ n)" by simp  | 
|
257  | 
next  | 
|
258  | 
assume "x = 1"  | 
|
259  | 
then show "x dvd (x ^ n)" by simp  | 
|
260  | 
qed  | 
|
261  | 
||
262  | 
end  | 
|
263  | 
||
| 
62481
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62366 
diff
changeset
 | 
264  | 
context semiring_1_no_zero_divisors  | 
| 60867 | 265  | 
begin  | 
266  | 
||
267  | 
subclass power .  | 
|
268  | 
||
| 63654 | 269  | 
lemma power_eq_0_iff [simp]: "a ^ n = 0 \<longleftrightarrow> a = 0 \<and> n > 0"  | 
| 60867 | 270  | 
by (induct n) auto  | 
271  | 
||
| 63654 | 272  | 
lemma power_not_zero: "a \<noteq> 0 \<Longrightarrow> a ^ n \<noteq> 0"  | 
| 60867 | 273  | 
by (induct n) auto  | 
274  | 
||
| 63654 | 275  | 
lemma zero_eq_power2 [simp]: "a\<^sup>2 = 0 \<longleftrightarrow> a = 0"  | 
| 60867 | 276  | 
unfolding power2_eq_square by simp  | 
277  | 
||
278  | 
end  | 
|
279  | 
||
| 30996 | 280  | 
context ring_1  | 
281  | 
begin  | 
|
282  | 
||
| 63654 | 283  | 
lemma power_minus: "(- a) ^ n = (- 1) ^ n * a ^ n"  | 
| 30996 | 284  | 
proof (induct n)  | 
| 63654 | 285  | 
case 0  | 
286  | 
show ?case by simp  | 
|
| 30996 | 287  | 
next  | 
| 63654 | 288  | 
case (Suc n)  | 
289  | 
then show ?case  | 
|
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57418 
diff
changeset
 | 
290  | 
by (simp del: power_Suc add: power_Suc2 mult.assoc)  | 
| 30996 | 291  | 
qed  | 
292  | 
||
| 
61531
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
293  | 
lemma power_minus': "NO_MATCH 1 x \<Longrightarrow> (-x) ^ n = (-1)^n * x ^ n"  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
294  | 
by (rule power_minus)  | 
| 
 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 
eberlm 
parents: 
61378 
diff
changeset
 | 
295  | 
|
| 63654 | 296  | 
lemma power_minus_Bit0: "(- x) ^ numeral (Num.Bit0 k) = x ^ numeral (Num.Bit0 k)"  | 
| 47191 | 297  | 
by (induct k, simp_all only: numeral_class.numeral.simps power_add  | 
298  | 
power_one_right mult_minus_left mult_minus_right minus_minus)  | 
|
299  | 
||
| 63654 | 300  | 
lemma power_minus_Bit1: "(- x) ^ numeral (Num.Bit1 k) = - (x ^ numeral (Num.Bit1 k))"  | 
| 
47220
 
52426c62b5d0
replace lemmas eval_nat_numeral with a simpler reformulation
 
huffman 
parents: 
47209 
diff
changeset
 | 
301  | 
by (simp only: eval_nat_numeral(3) power_Suc power_minus_Bit0 mult_minus_left)  | 
| 47191 | 302  | 
|
| 63654 | 303  | 
lemma power2_minus [simp]: "(- a)\<^sup>2 = a\<^sup>2"  | 
| 60867 | 304  | 
by (fact power_minus_Bit0)  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
305  | 
|
| 63654 | 306  | 
lemma power_minus1_even [simp]: "(- 1) ^ (2*n) = 1"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
307  | 
proof (induct n)  | 
| 63654 | 308  | 
case 0  | 
309  | 
show ?case by simp  | 
|
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
310  | 
next  | 
| 63654 | 311  | 
case (Suc n)  | 
312  | 
then show ?case by (simp add: power_add power2_eq_square)  | 
|
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
313  | 
qed  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
314  | 
|
| 63654 | 315  | 
lemma power_minus1_odd: "(- 1) ^ Suc (2*n) = -1"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
316  | 
by simp  | 
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61531 
diff
changeset
 | 
317  | 
|
| 63654 | 318  | 
lemma power_minus_even [simp]: "(-a) ^ (2*n) = a ^ (2*n)"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
319  | 
by (simp add: power_minus [of a])  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
320  | 
|
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
321  | 
end  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
322  | 
|
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
323  | 
context ring_1_no_zero_divisors  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
324  | 
begin  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
325  | 
|
| 63654 | 326  | 
lemma power2_eq_1_iff: "a\<^sup>2 = 1 \<longleftrightarrow> a = 1 \<or> a = - 1"  | 
| 60867 | 327  | 
using square_eq_1_iff [of a] by (simp add: power2_eq_square)  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
328  | 
|
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
329  | 
end  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
330  | 
|
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
331  | 
context idom  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
332  | 
begin  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
333  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
334  | 
lemma power2_eq_iff: "x\<^sup>2 = y\<^sup>2 \<longleftrightarrow> x = y \<or> x = - y"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
335  | 
unfolding power2_eq_square by (rule square_eq_iff)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
336  | 
|
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
337  | 
end  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
338  | 
|
| 66936 | 339  | 
context semidom_divide  | 
340  | 
begin  | 
|
341  | 
||
342  | 
lemma power_diff:  | 
|
343  | 
"a ^ (m - n) = (a ^ m) div (a ^ n)" if "a \<noteq> 0" and "n \<le> m"  | 
|
344  | 
proof -  | 
|
345  | 
define q where "q = m - n"  | 
|
346  | 
with \<open>n \<le> m\<close> have "m = q + n" by simp  | 
|
347  | 
with \<open>a \<noteq> 0\<close> q_def show ?thesis  | 
|
348  | 
by (simp add: power_add)  | 
|
349  | 
qed  | 
|
350  | 
||
351  | 
end  | 
|
352  | 
||
| 60867 | 353  | 
context algebraic_semidom  | 
354  | 
begin  | 
|
355  | 
||
| 63654 | 356  | 
lemma div_power: "b dvd a \<Longrightarrow> (a div b) ^ n = a ^ n div b ^ n"  | 
357  | 
by (induct n) (simp_all add: div_mult_div_if_dvd dvd_power_same)  | 
|
| 60867 | 358  | 
|
| 63654 | 359  | 
lemma is_unit_power_iff: "is_unit (a ^ n) \<longleftrightarrow> is_unit a \<or> n = 0"  | 
| 62366 | 360  | 
by (induct n) (auto simp add: is_unit_mult_iff)  | 
361  | 
||
| 63924 | 362  | 
lemma dvd_power_iff:  | 
363  | 
assumes "x \<noteq> 0"  | 
|
364  | 
shows "x ^ m dvd x ^ n \<longleftrightarrow> is_unit x \<or> m \<le> n"  | 
|
365  | 
proof  | 
|
366  | 
assume *: "x ^ m dvd x ^ n"  | 
|
367  | 
  {
 | 
|
368  | 
assume "m > n"  | 
|
369  | 
note *  | 
|
370  | 
also have "x ^ n = x ^ n * 1" by simp  | 
|
371  | 
also from \<open>m > n\<close> have "m = n + (m - n)" by simp  | 
|
372  | 
also have "x ^ \<dots> = x ^ n * x ^ (m - n)" by (rule power_add)  | 
|
373  | 
finally have "x ^ (m - n) dvd 1"  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
374  | 
using assms by (subst (asm) dvd_times_left_cancel_iff) simp_all  | 
| 63924 | 375  | 
with \<open>m > n\<close> have "is_unit x" by (simp add: is_unit_power_iff)  | 
376  | 
}  | 
|
377  | 
thus "is_unit x \<or> m \<le> n" by force  | 
|
378  | 
qed (auto intro: unit_imp_dvd simp: is_unit_power_iff le_imp_power_dvd)  | 
|
379  | 
||
380  | 
||
| 60867 | 381  | 
end  | 
382  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70928 
diff
changeset
 | 
383  | 
context normalization_semidom_multiplicative  | 
| 
60685
 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 
haftmann 
parents: 
60155 
diff
changeset
 | 
384  | 
begin  | 
| 
 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 
haftmann 
parents: 
60155 
diff
changeset
 | 
385  | 
|
| 63654 | 386  | 
lemma normalize_power: "normalize (a ^ n) = normalize a ^ n"  | 
| 
60685
 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 
haftmann 
parents: 
60155 
diff
changeset
 | 
387  | 
by (induct n) (simp_all add: normalize_mult)  | 
| 
 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 
haftmann 
parents: 
60155 
diff
changeset
 | 
388  | 
|
| 63654 | 389  | 
lemma unit_factor_power: "unit_factor (a ^ n) = unit_factor a ^ n"  | 
| 
60685
 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 
haftmann 
parents: 
60155 
diff
changeset
 | 
390  | 
by (induct n) (simp_all add: unit_factor_mult)  | 
| 
 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 
haftmann 
parents: 
60155 
diff
changeset
 | 
391  | 
|
| 
 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 
haftmann 
parents: 
60155 
diff
changeset
 | 
392  | 
end  | 
| 
 
cb21b7022b00
moved normalization and unit_factor into Main HOL corpus
 
haftmann 
parents: 
60155 
diff
changeset
 | 
393  | 
|
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
394  | 
context division_ring  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
395  | 
begin  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
396  | 
|
| 63654 | 397  | 
text \<open>Perhaps these should be simprules.\<close>  | 
| 
70817
 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 
haftmann 
parents: 
70724 
diff
changeset
 | 
398  | 
lemma power_inverse [field_simps, field_split_simps, divide_simps]: "inverse a ^ n = inverse (a ^ n)"  | 
| 60867 | 399  | 
proof (cases "a = 0")  | 
| 63654 | 400  | 
case True  | 
401  | 
then show ?thesis by (simp add: power_0_left)  | 
|
| 60867 | 402  | 
next  | 
| 63654 | 403  | 
case False  | 
404  | 
then have "inverse (a ^ n) = inverse a ^ n"  | 
|
| 60867 | 405  | 
by (induct n) (simp_all add: nonzero_inverse_mult_distrib power_commutes)  | 
406  | 
then show ?thesis by simp  | 
|
407  | 
qed  | 
|
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
408  | 
|
| 
70817
 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 
haftmann 
parents: 
70724 
diff
changeset
 | 
409  | 
lemma power_one_over [field_simps, field_split_simps, divide_simps]: "(1 / a) ^ n = 1 / a ^ n"  | 
| 60867 | 410  | 
using power_inverse [of a] by (simp add: divide_inverse)  | 
411  | 
||
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61531 
diff
changeset
 | 
412  | 
end  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
413  | 
|
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
414  | 
context field  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
415  | 
begin  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
416  | 
|
| 
70817
 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 
haftmann 
parents: 
70724 
diff
changeset
 | 
417  | 
lemma power_divide [field_simps, field_split_simps, divide_simps]: "(a / b) ^ n = a ^ n / b ^ n"  | 
| 60867 | 418  | 
by (induct n) simp_all  | 
419  | 
||
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
420  | 
end  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
421  | 
|
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
422  | 
|
| 60758 | 423  | 
subsection \<open>Exponentiation on ordered types\<close>  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
424  | 
|
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
33364 
diff
changeset
 | 
425  | 
context linordered_semidom  | 
| 30996 | 426  | 
begin  | 
427  | 
||
| 63654 | 428  | 
lemma zero_less_power [simp]: "0 < a \<Longrightarrow> 0 < a ^ n"  | 
| 56544 | 429  | 
by (induct n) simp_all  | 
| 30996 | 430  | 
|
| 63654 | 431  | 
lemma zero_le_power [simp]: "0 \<le> a \<Longrightarrow> 0 \<le> a ^ n"  | 
| 56536 | 432  | 
by (induct n) simp_all  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
433  | 
|
| 63654 | 434  | 
lemma power_mono: "a \<le> b \<Longrightarrow> 0 \<le> a \<Longrightarrow> a ^ n \<le> b ^ n"  | 
| 47241 | 435  | 
by (induct n) (auto intro: mult_mono order_trans [of 0 a b])  | 
436  | 
||
437  | 
lemma one_le_power [simp]: "1 \<le> a \<Longrightarrow> 1 \<le> a ^ n"  | 
|
438  | 
using power_mono [of 1 a n] by simp  | 
|
439  | 
||
| 63654 | 440  | 
lemma power_le_one: "0 \<le> a \<Longrightarrow> a \<le> 1 \<Longrightarrow> a ^ n \<le> 1"  | 
| 47241 | 441  | 
using power_mono [of a 1 n] by simp  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
442  | 
|
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
443  | 
lemma power_gt1_lemma:  | 
| 30996 | 444  | 
assumes gt1: "1 < a"  | 
445  | 
shows "1 < a * a ^ n"  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
446  | 
proof -  | 
| 30996 | 447  | 
from gt1 have "0 \<le> a"  | 
448  | 
by (fact order_trans [OF zero_le_one less_imp_le])  | 
|
| 63654 | 449  | 
from gt1 have "1 * 1 < a * 1" by simp  | 
450  | 
also from gt1 have "\<dots> \<le> a * a ^ n"  | 
|
451  | 
by (simp only: mult_mono \<open>0 \<le> a\<close> one_le_power order_less_imp_le zero_le_one order_refl)  | 
|
| 14577 | 452  | 
finally show ?thesis by simp  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
453  | 
qed  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
454  | 
|
| 63654 | 455  | 
lemma power_gt1: "1 < a \<Longrightarrow> 1 < a ^ Suc n"  | 
| 30996 | 456  | 
by (simp add: power_gt1_lemma)  | 
| 24376 | 457  | 
|
| 63654 | 458  | 
lemma one_less_power [simp]: "1 < a \<Longrightarrow> 0 < n \<Longrightarrow> 1 < a ^ n"  | 
| 30996 | 459  | 
by (cases n) (simp_all add: power_gt1_lemma)  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
460  | 
|
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
461  | 
lemma power_le_imp_le_exp:  | 
| 30996 | 462  | 
assumes gt1: "1 < a"  | 
463  | 
shows "a ^ m \<le> a ^ n \<Longrightarrow> m \<le> n"  | 
|
464  | 
proof (induct m arbitrary: n)  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
465  | 
case 0  | 
| 14577 | 466  | 
show ?case by simp  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
467  | 
next  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
468  | 
case (Suc m)  | 
| 14577 | 469  | 
show ?case  | 
470  | 
proof (cases n)  | 
|
471  | 
case 0  | 
|
| 63654 | 472  | 
with Suc have "a * a ^ m \<le> 1" by simp  | 
| 14577 | 473  | 
with gt1 show ?thesis  | 
| 63654 | 474  | 
by (force simp only: power_gt1_lemma not_less [symmetric])  | 
| 14577 | 475  | 
next  | 
476  | 
case (Suc n)  | 
|
| 30996 | 477  | 
with Suc.prems Suc.hyps show ?thesis  | 
| 63654 | 478  | 
by (force dest: mult_left_le_imp_le simp add: less_trans [OF zero_less_one gt1])  | 
| 14577 | 479  | 
qed  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
480  | 
qed  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
481  | 
|
| 63654 | 482  | 
lemma of_nat_zero_less_power_iff [simp]: "of_nat x ^ n > 0 \<longleftrightarrow> x > 0 \<or> n = 0"  | 
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61531 
diff
changeset
 | 
483  | 
by (induct n) auto  | 
| 
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61531 
diff
changeset
 | 
484  | 
|
| 63654 | 485  | 
text \<open>Surely we can strengthen this? It holds for \<open>0<a<1\<close> too.\<close>  | 
| 73411 | 486  | 
lemma power_inject_exp [simp]:  | 
487  | 
\<open>a ^ m = a ^ n \<longleftrightarrow> m = n\<close> if \<open>1 < a\<close>  | 
|
488  | 
using that by (force simp add: order_class.order.antisym power_le_imp_le_exp)  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
489  | 
|
| 63654 | 490  | 
text \<open>  | 
| 69593 | 491  | 
Can relax the first premise to \<^term>\<open>0<a\<close> in the case of the  | 
| 63654 | 492  | 
natural numbers.  | 
493  | 
\<close>  | 
|
494  | 
lemma power_less_imp_less_exp: "1 < a \<Longrightarrow> a ^ m < a ^ n \<Longrightarrow> m < n"  | 
|
495  | 
by (simp add: order_less_le [of m n] less_le [of "a^m" "a^n"] power_le_imp_le_exp)  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
496  | 
|
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
497  | 
lemma power_strict_mono: "a < b \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 < n \<Longrightarrow> a ^ n < b ^ n"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
498  | 
proof (induct n)  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
499  | 
case 0  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
500  | 
then show ?case by simp  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
501  | 
next  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
502  | 
case (Suc n)  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
503  | 
then show ?case  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
504  | 
by (cases "n = 0") (auto simp: mult_strict_mono le_less_trans [of 0 a b])  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
505  | 
qed  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
506  | 
|
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70331 
diff
changeset
 | 
507  | 
lemma power_mono_iff [simp]:  | 
| 
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70331 
diff
changeset
 | 
508  | 
shows "\<lbrakk>a \<ge> 0; b \<ge> 0; n>0\<rbrakk> \<Longrightarrow> a ^ n \<le> b ^ n \<longleftrightarrow> a \<le> b"  | 
| 
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70331 
diff
changeset
 | 
509  | 
using power_mono [of a b] power_strict_mono [of b a] not_le by auto  | 
| 
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70331 
diff
changeset
 | 
510  | 
|
| 61799 | 511  | 
text\<open>Lemma for \<open>power_strict_decreasing\<close>\<close>  | 
| 63654 | 512  | 
lemma power_Suc_less: "0 < a \<Longrightarrow> a < 1 \<Longrightarrow> a * a ^ n < a ^ n"  | 
513  | 
by (induct n) (auto simp: mult_strict_left_mono)  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
514  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
515  | 
lemma power_strict_decreasing: "n < N \<Longrightarrow> 0 < a \<Longrightarrow> a < 1 \<Longrightarrow> a ^ N < a ^ n"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
516  | 
proof (induction N)  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
517  | 
case 0  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
518  | 
then show ?case by simp  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
519  | 
next  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
520  | 
case (Suc N)  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
521  | 
then show ?case  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
522  | 
using mult_strict_mono[of a 1 "a ^ N" "a ^ n"]  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
523  | 
by (auto simp add: power_Suc_less less_Suc_eq)  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
524  | 
qed  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
525  | 
|
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
526  | 
text \<open>Proof resembles that of \<open>power_strict_decreasing\<close>.\<close>  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
527  | 
lemma power_decreasing: "n \<le> N \<Longrightarrow> 0 \<le> a \<Longrightarrow> a \<le> 1 \<Longrightarrow> a ^ N \<le> a ^ n"  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
528  | 
proof (induction N)  | 
| 63654 | 529  | 
case 0  | 
530  | 
then show ?case by simp  | 
|
| 30996 | 531  | 
next  | 
| 63654 | 532  | 
case (Suc N)  | 
533  | 
then show ?case  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
534  | 
using mult_mono[of a 1 "a^N" "a ^ n"]  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
535  | 
by (auto simp add: le_Suc_eq)  | 
| 30996 | 536  | 
qed  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
537  | 
|
| 
69700
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
538  | 
lemma power_decreasing_iff [simp]: "\<lbrakk>0 < b; b < 1\<rbrakk> \<Longrightarrow> b ^ m \<le> b ^ n \<longleftrightarrow> n \<le> m"  | 
| 
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
539  | 
using power_strict_decreasing [of m n b]  | 
| 
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
540  | 
by (auto intro: power_decreasing ccontr)  | 
| 
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
541  | 
|
| 
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
542  | 
lemma power_strict_decreasing_iff [simp]: "\<lbrakk>0 < b; b < 1\<rbrakk> \<Longrightarrow> b ^ m < b ^ n \<longleftrightarrow> n < m"  | 
| 
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
543  | 
using power_decreasing_iff [of b m n] unfolding le_less  | 
| 
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
544  | 
by (auto dest: power_strict_decreasing le_neq_implies_less)  | 
| 
 
7a92cbec7030
new material about summations and powers, along with some tweaks
 
paulson <lp15@cam.ac.uk> 
parents: 
69593 
diff
changeset
 | 
545  | 
|
| 63654 | 546  | 
lemma power_Suc_less_one: "0 < a \<Longrightarrow> a < 1 \<Longrightarrow> a ^ Suc n < 1"  | 
| 30996 | 547  | 
using power_strict_decreasing [of 0 "Suc n" a] by simp  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
548  | 
|
| 63654 | 549  | 
text \<open>Proof again resembles that of \<open>power_strict_decreasing\<close>.\<close>  | 
550  | 
lemma power_increasing: "n \<le> N \<Longrightarrow> 1 \<le> a \<Longrightarrow> a ^ n \<le> a ^ N"  | 
|
| 30996 | 551  | 
proof (induct N)  | 
| 63654 | 552  | 
case 0  | 
553  | 
then show ?case by simp  | 
|
| 30996 | 554  | 
next  | 
| 63654 | 555  | 
case (Suc N)  | 
556  | 
then show ?case  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
557  | 
using mult_mono[of 1 a "a ^ n" "a ^ N"]  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
558  | 
by (auto simp add: le_Suc_eq order_trans [OF zero_le_one])  | 
| 30996 | 559  | 
qed  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
560  | 
|
| 63654 | 561  | 
text \<open>Lemma for \<open>power_strict_increasing\<close>.\<close>  | 
562  | 
lemma power_less_power_Suc: "1 < a \<Longrightarrow> a ^ n < a * a ^ n"  | 
|
563  | 
by (induct n) (auto simp: mult_strict_left_mono less_trans [OF zero_less_one])  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
564  | 
|
| 63654 | 565  | 
lemma power_strict_increasing: "n < N \<Longrightarrow> 1 < a \<Longrightarrow> a ^ n < a ^ N"  | 
| 30996 | 566  | 
proof (induct N)  | 
| 63654 | 567  | 
case 0  | 
568  | 
then show ?case by simp  | 
|
| 30996 | 569  | 
next  | 
| 63654 | 570  | 
case (Suc N)  | 
571  | 
then show ?case  | 
|
| 
75669
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
572  | 
using mult_strict_mono[of 1 a "a^n" "a^N"]  | 
| 
 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 
Fabian Huch <huch@in.tum.de> 
parents: 
74438 
diff
changeset
 | 
573  | 
by (auto simp add: power_less_power_Suc less_Suc_eq less_trans [OF zero_less_one] less_imp_le)  | 
| 30996 | 574  | 
qed  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
575  | 
|
| 63654 | 576  | 
lemma power_increasing_iff [simp]: "1 < b \<Longrightarrow> b ^ x \<le> b ^ y \<longleftrightarrow> x \<le> y"  | 
| 30996 | 577  | 
by (blast intro: power_le_imp_le_exp power_increasing less_imp_le)  | 
| 15066 | 578  | 
|
| 63654 | 579  | 
lemma power_strict_increasing_iff [simp]: "1 < b \<Longrightarrow> b ^ x < b ^ y \<longleftrightarrow> x < y"  | 
580  | 
by (blast intro: power_less_imp_less_exp power_strict_increasing)  | 
|
| 15066 | 581  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
582  | 
lemma power_le_imp_le_base:  | 
| 30996 | 583  | 
assumes le: "a ^ Suc n \<le> b ^ Suc n"  | 
| 63654 | 584  | 
and "0 \<le> b"  | 
| 30996 | 585  | 
shows "a \<le> b"  | 
| 
25134
 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 
nipkow 
parents: 
25062 
diff
changeset
 | 
586  | 
proof (rule ccontr)  | 
| 63654 | 587  | 
assume "\<not> ?thesis"  | 
| 
25134
 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 
nipkow 
parents: 
25062 
diff
changeset
 | 
588  | 
then have "b < a" by (simp only: linorder_not_le)  | 
| 
 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 
nipkow 
parents: 
25062 
diff
changeset
 | 
589  | 
then have "b ^ Suc n < a ^ Suc n"  | 
| 63654 | 590  | 
by (simp only: assms(2) power_strict_mono)  | 
591  | 
with le show False  | 
|
| 
25134
 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 
nipkow 
parents: 
25062 
diff
changeset
 | 
592  | 
by (simp add: linorder_not_less [symmetric])  | 
| 
 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 
nipkow 
parents: 
25062 
diff
changeset
 | 
593  | 
qed  | 
| 14577 | 594  | 
|
| 22853 | 595  | 
lemma power_less_imp_less_base:  | 
596  | 
assumes less: "a ^ n < b ^ n"  | 
|
597  | 
assumes nonneg: "0 \<le> b"  | 
|
598  | 
shows "a < b"  | 
|
599  | 
proof (rule contrapos_pp [OF less])  | 
|
| 63654 | 600  | 
assume "\<not> ?thesis"  | 
601  | 
then have "b \<le> a" by (simp only: linorder_not_less)  | 
|
602  | 
from this nonneg have "b ^ n \<le> a ^ n" by (rule power_mono)  | 
|
603  | 
then show "\<not> a ^ n < b ^ n" by (simp only: linorder_not_less)  | 
|
| 22853 | 604  | 
qed  | 
605  | 
||
| 63654 | 606  | 
lemma power_inject_base: "a ^ Suc n = b ^ Suc n \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> a = b"  | 
| 73411 | 607  | 
by (blast intro: power_le_imp_le_base order.antisym eq_refl sym)  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
608  | 
|
| 63654 | 609  | 
lemma power_eq_imp_eq_base: "a ^ n = b ^ n \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 < n \<Longrightarrow> a = b"  | 
| 30996 | 610  | 
by (cases n) (simp_all del: power_Suc, rule power_inject_base)  | 
| 22955 | 611  | 
|
| 63654 | 612  | 
lemma power_eq_iff_eq_base: "0 < n \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> a ^ n = b ^ n \<longleftrightarrow> a = b"  | 
| 62347 | 613  | 
using power_eq_imp_eq_base [of a n b] by auto  | 
614  | 
||
| 63654 | 615  | 
lemma power2_le_imp_le: "x\<^sup>2 \<le> y\<^sup>2 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x \<le> y"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
616  | 
unfolding numeral_2_eq_2 by (rule power_le_imp_le_base)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
617  | 
|
| 63654 | 618  | 
lemma power2_less_imp_less: "x\<^sup>2 < y\<^sup>2 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x < y"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
619  | 
by (rule power_less_imp_less_base)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
620  | 
|
| 63654 | 621  | 
lemma power2_eq_imp_eq: "x\<^sup>2 = y\<^sup>2 \<Longrightarrow> 0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> x = y"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
622  | 
unfolding numeral_2_eq_2 by (erule (2) power_eq_imp_eq_base) simp  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
623  | 
|
| 63654 | 624  | 
lemma power_Suc_le_self: "0 \<le> a \<Longrightarrow> a \<le> 1 \<Longrightarrow> a ^ Suc n \<le> a"  | 
| 62347 | 625  | 
using power_decreasing [of 1 "Suc n" a] by simp  | 
626  | 
||
| 
65057
 
799bbbb3a395
Some new lemmas thanks to Lukas Bulwahn. Also, NEWS.
 
paulson <lp15@cam.ac.uk> 
parents: 
64964 
diff
changeset
 | 
627  | 
lemma power2_eq_iff_nonneg [simp]:  | 
| 
 
799bbbb3a395
Some new lemmas thanks to Lukas Bulwahn. Also, NEWS.
 
paulson <lp15@cam.ac.uk> 
parents: 
64964 
diff
changeset
 | 
628  | 
assumes "0 \<le> x" "0 \<le> y"  | 
| 
 
799bbbb3a395
Some new lemmas thanks to Lukas Bulwahn. Also, NEWS.
 
paulson <lp15@cam.ac.uk> 
parents: 
64964 
diff
changeset
 | 
629  | 
shows "(x ^ 2 = y ^ 2) \<longleftrightarrow> x = y"  | 
| 
 
799bbbb3a395
Some new lemmas thanks to Lukas Bulwahn. Also, NEWS.
 
paulson <lp15@cam.ac.uk> 
parents: 
64964 
diff
changeset
 | 
630  | 
using assms power2_eq_imp_eq by blast  | 
| 
 
799bbbb3a395
Some new lemmas thanks to Lukas Bulwahn. Also, NEWS.
 
paulson <lp15@cam.ac.uk> 
parents: 
64964 
diff
changeset
 | 
631  | 
|
| 
66912
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
632  | 
lemma of_nat_less_numeral_power_cancel_iff[simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
633  | 
"of_nat x < numeral i ^ n \<longleftrightarrow> x < numeral i ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
634  | 
using of_nat_less_iff[of x "numeral i ^ n", unfolded of_nat_numeral of_nat_power] .  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
635  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
636  | 
lemma of_nat_le_numeral_power_cancel_iff[simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
637  | 
"of_nat x \<le> numeral i ^ n \<longleftrightarrow> x \<le> numeral i ^ n"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
638  | 
using of_nat_le_iff[of x "numeral i ^ n", unfolded of_nat_numeral of_nat_power] .  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
639  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
640  | 
lemma numeral_power_less_of_nat_cancel_iff[simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
641  | 
"numeral i ^ n < of_nat x \<longleftrightarrow> numeral i ^ n < x"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
642  | 
using of_nat_less_iff[of "numeral i ^ n" x, unfolded of_nat_numeral of_nat_power] .  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
643  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
644  | 
lemma numeral_power_le_of_nat_cancel_iff[simp]:  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
645  | 
"numeral i ^ n \<le> of_nat x \<longleftrightarrow> numeral i ^ n \<le> x"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
646  | 
using of_nat_le_iff[of "numeral i ^ n" x, unfolded of_nat_numeral of_nat_power] .  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
647  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
648  | 
lemma of_nat_le_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w \<le> of_nat x \<longleftrightarrow> b ^ w \<le> x"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
649  | 
by (metis of_nat_le_iff of_nat_power)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
650  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
651  | 
lemma of_nat_power_le_of_nat_cancel_iff[simp]: "of_nat x \<le> (of_nat b) ^ w \<longleftrightarrow> x \<le> b ^ w"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
652  | 
by (metis of_nat_le_iff of_nat_power)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
653  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
654  | 
lemma of_nat_less_of_nat_power_cancel_iff[simp]: "(of_nat b) ^ w < of_nat x \<longleftrightarrow> b ^ w < x"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
655  | 
by (metis of_nat_less_iff of_nat_power)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
656  | 
|
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
657  | 
lemma of_nat_power_less_of_nat_cancel_iff[simp]: "of_nat x < (of_nat b) ^ w \<longleftrightarrow> x < b ^ w"  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
658  | 
by (metis of_nat_less_iff of_nat_power)  | 
| 
 
a99a7cbf0fb5
generalized lemmas cancelling real_of_int/real in (in)equalities with power; completed set of related simp rules; lemmas about floorlog/bitlen
 
immler 
parents: 
65057 
diff
changeset
 | 
659  | 
|
| 
77138
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
660  | 
lemma power2_nonneg_ge_1_iff:  | 
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
661  | 
assumes "x \<ge> 0"  | 
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
662  | 
shows "x ^ 2 \<ge> 1 \<longleftrightarrow> x \<ge> 1"  | 
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
663  | 
using assms by (auto intro: power2_le_imp_le)  | 
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
664  | 
|
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
665  | 
lemma power2_nonneg_gt_1_iff:  | 
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
666  | 
assumes "x \<ge> 0"  | 
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
667  | 
shows "x ^ 2 > 1 \<longleftrightarrow> x > 1"  | 
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
668  | 
using assms by (auto intro: power_less_imp_less_base)  | 
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
669  | 
|
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
670  | 
end  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
671  | 
|
| 70331 | 672  | 
text \<open>Some @{typ nat}-specific lemmas:\<close>
 | 
673  | 
||
674  | 
lemma mono_ge2_power_minus_self:  | 
|
675  | 
assumes "k \<ge> 2" shows "mono (\<lambda>m. k ^ m - m)"  | 
|
676  | 
unfolding mono_iff_le_Suc  | 
|
677  | 
proof  | 
|
678  | 
fix n  | 
|
679  | 
have "k ^ n < k ^ Suc n" using power_strict_increasing_iff[of k "n" "Suc n"] assms by linarith  | 
|
680  | 
thus "k ^ n - n \<le> k ^ Suc n - Suc n" by linarith  | 
|
681  | 
qed  | 
|
682  | 
||
683  | 
lemma self_le_ge2_pow[simp]:  | 
|
684  | 
assumes "k \<ge> 2" shows "m \<le> k ^ m"  | 
|
685  | 
proof (induction m)  | 
|
686  | 
case 0 show ?case by simp  | 
|
687  | 
next  | 
|
688  | 
case (Suc m)  | 
|
689  | 
hence "Suc m \<le> Suc (k ^ m)" by simp  | 
|
690  | 
also have "... \<le> k^m + k^m" using one_le_power[of k m] assms by linarith  | 
|
691  | 
also have "... \<le> k * k^m" by (metis mult_2 mult_le_mono1[OF assms])  | 
|
692  | 
finally show ?case by simp  | 
|
693  | 
qed  | 
|
694  | 
||
695  | 
lemma diff_le_diff_pow[simp]:  | 
|
696  | 
assumes "k \<ge> 2" shows "m - n \<le> k ^ m - k ^ n"  | 
|
697  | 
proof (cases "n \<le> m")  | 
|
698  | 
case True  | 
|
699  | 
thus ?thesis  | 
|
700  | 
using monoD[OF mono_ge2_power_minus_self[OF assms] True] self_le_ge2_pow[OF assms, of m]  | 
|
701  | 
by (simp add: le_diff_conv le_diff_conv2)  | 
|
702  | 
qed auto  | 
|
703  | 
||
704  | 
||
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
705  | 
context linordered_ring_strict  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
706  | 
begin  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
707  | 
|
| 63654 | 708  | 
lemma sum_squares_eq_zero_iff: "x * x + y * y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
709  | 
by (simp add: add_nonneg_eq_0_iff)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
710  | 
|
| 63654 | 711  | 
lemma sum_squares_le_zero_iff: "x * x + y * y \<le> 0 \<longleftrightarrow> x = 0 \<and> y = 0"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
712  | 
by (simp add: le_less not_sum_squares_lt_zero sum_squares_eq_zero_iff)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
713  | 
|
| 63654 | 714  | 
lemma sum_squares_gt_zero_iff: "0 < x * x + y * y \<longleftrightarrow> x \<noteq> 0 \<or> y \<noteq> 0"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
715  | 
by (simp add: not_le [symmetric] sum_squares_le_zero_iff)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
716  | 
|
| 30996 | 717  | 
end  | 
718  | 
||
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
33364 
diff
changeset
 | 
719  | 
context linordered_idom  | 
| 30996 | 720  | 
begin  | 
| 
29978
 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 
huffman 
parents: 
29608 
diff
changeset
 | 
721  | 
|
| 64715 | 722  | 
lemma zero_le_power2 [simp]: "0 \<le> a\<^sup>2"  | 
723  | 
by (simp add: power2_eq_square)  | 
|
724  | 
||
725  | 
lemma zero_less_power2 [simp]: "0 < a\<^sup>2 \<longleftrightarrow> a \<noteq> 0"  | 
|
726  | 
by (force simp add: power2_eq_square zero_less_mult_iff linorder_neq_iff)  | 
|
| 30996 | 727  | 
|
| 64715 | 728  | 
lemma power2_less_0 [simp]: "\<not> a\<^sup>2 < 0"  | 
729  | 
by (force simp add: power2_eq_square mult_less_0_iff)  | 
|
730  | 
||
| 67226 | 731  | 
lemma power_abs: "\<bar>a ^ n\<bar> = \<bar>a\<bar> ^ n" \<comment> \<open>FIXME simp?\<close>  | 
| 64715 | 732  | 
by (induct n) (simp_all add: abs_mult)  | 
733  | 
||
734  | 
lemma power_sgn [simp]: "sgn (a ^ n) = sgn a ^ n"  | 
|
735  | 
by (induct n) (simp_all add: sgn_mult)  | 
|
| 64964 | 736  | 
|
| 64715 | 737  | 
lemma abs_power_minus [simp]: "\<bar>(- a) ^ n\<bar> = \<bar>a ^ n\<bar>"  | 
| 35216 | 738  | 
by (simp add: power_abs)  | 
| 30996 | 739  | 
|
| 61944 | 740  | 
lemma zero_less_power_abs_iff [simp]: "0 < \<bar>a\<bar> ^ n \<longleftrightarrow> a \<noteq> 0 \<or> n = 0"  | 
| 30996 | 741  | 
proof (induct n)  | 
| 63654 | 742  | 
case 0  | 
743  | 
show ?case by simp  | 
|
| 30996 | 744  | 
next  | 
| 63654 | 745  | 
case Suc  | 
746  | 
then show ?case by (auto simp: zero_less_mult_iff)  | 
|
| 
29978
 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 
huffman 
parents: 
29608 
diff
changeset
 | 
747  | 
qed  | 
| 
 
33df3c4eb629
generalize le_imp_power_dvd and power_le_dvd; move from Divides to Power
 
huffman 
parents: 
29608 
diff
changeset
 | 
748  | 
|
| 61944 | 749  | 
lemma zero_le_power_abs [simp]: "0 \<le> \<bar>a\<bar> ^ n"  | 
| 30996 | 750  | 
by (rule zero_le_power [OF abs_ge_zero])  | 
751  | 
||
| 63654 | 752  | 
lemma power2_less_eq_zero_iff [simp]: "a\<^sup>2 \<le> 0 \<longleftrightarrow> a = 0"  | 
| 58787 | 753  | 
by (simp add: le_less)  | 
754  | 
||
| 61944 | 755  | 
lemma abs_power2 [simp]: "\<bar>a\<^sup>2\<bar> = a\<^sup>2"  | 
| 
63417
 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 
haftmann 
parents: 
63040 
diff
changeset
 | 
756  | 
by (simp add: power2_eq_square)  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
757  | 
|
| 61944 | 758  | 
lemma power2_abs [simp]: "\<bar>a\<bar>\<^sup>2 = a\<^sup>2"  | 
| 
63417
 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 
haftmann 
parents: 
63040 
diff
changeset
 | 
759  | 
by (simp add: power2_eq_square)  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
760  | 
|
| 64715 | 761  | 
lemma odd_power_less_zero: "a < 0 \<Longrightarrow> a ^ Suc (2 * n) < 0"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
762  | 
proof (induct n)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
763  | 
case 0  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
764  | 
then show ?case by simp  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
765  | 
next  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
766  | 
case (Suc n)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
767  | 
have "a ^ Suc (2 * Suc n) = (a*a) * a ^ Suc(2*n)"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
768  | 
by (simp add: ac_simps power_add power2_eq_square)  | 
| 63654 | 769  | 
then show ?case  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
770  | 
by (simp del: power_Suc add: Suc mult_less_0_iff mult_neg_neg)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
771  | 
qed  | 
| 30996 | 772  | 
|
| 64715 | 773  | 
lemma odd_0_le_power_imp_0_le: "0 \<le> a ^ Suc (2 * n) \<Longrightarrow> 0 \<le> a"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
774  | 
using odd_power_less_zero [of a n]  | 
| 63654 | 775  | 
by (force simp add: linorder_not_less [symmetric])  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
776  | 
|
| 64715 | 777  | 
lemma zero_le_even_power'[simp]: "0 \<le> a ^ (2 * n)"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
778  | 
proof (induct n)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
779  | 
case 0  | 
| 63654 | 780  | 
show ?case by simp  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
781  | 
next  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
782  | 
case (Suc n)  | 
| 63654 | 783  | 
have "a ^ (2 * Suc n) = (a*a) * a ^ (2*n)"  | 
784  | 
by (simp add: ac_simps power_add power2_eq_square)  | 
|
785  | 
then show ?case  | 
|
786  | 
by (simp add: Suc zero_le_mult_iff)  | 
|
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
787  | 
qed  | 
| 30996 | 788  | 
|
| 63654 | 789  | 
lemma sum_power2_ge_zero: "0 \<le> x\<^sup>2 + y\<^sup>2"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
790  | 
by (intro add_nonneg_nonneg zero_le_power2)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
791  | 
|
| 63654 | 792  | 
lemma not_sum_power2_lt_zero: "\<not> x\<^sup>2 + y\<^sup>2 < 0"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
793  | 
unfolding not_less by (rule sum_power2_ge_zero)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
794  | 
|
| 63654 | 795  | 
lemma sum_power2_eq_zero_iff: "x\<^sup>2 + y\<^sup>2 = 0 \<longleftrightarrow> x = 0 \<and> y = 0"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
796  | 
unfolding power2_eq_square by (simp add: add_nonneg_eq_0_iff)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
797  | 
|
| 63654 | 798  | 
lemma sum_power2_le_zero_iff: "x\<^sup>2 + y\<^sup>2 \<le> 0 \<longleftrightarrow> x = 0 \<and> y = 0"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
799  | 
by (simp add: le_less sum_power2_eq_zero_iff not_sum_power2_lt_zero)  | 
| 
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
800  | 
|
| 63654 | 801  | 
lemma sum_power2_gt_zero_iff: "0 < x\<^sup>2 + y\<^sup>2 \<longleftrightarrow> x \<noteq> 0 \<or> y \<noteq> 0"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
802  | 
unfolding not_le [symmetric] by (simp add: sum_power2_le_zero_iff)  | 
| 30996 | 803  | 
|
| 63654 | 804  | 
lemma abs_le_square_iff: "\<bar>x\<bar> \<le> \<bar>y\<bar> \<longleftrightarrow> x\<^sup>2 \<le> y\<^sup>2"  | 
805  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 
59865
 
8a20dd967385
rationalised and generalised some theorems concerning abs and x^2.
 
paulson <lp15@cam.ac.uk> 
parents: 
59741 
diff
changeset
 | 
806  | 
proof  | 
| 63654 | 807  | 
assume ?lhs  | 
808  | 
then have "\<bar>x\<bar>\<^sup>2 \<le> \<bar>y\<bar>\<^sup>2" by (rule power_mono) simp  | 
|
809  | 
then show ?rhs by simp  | 
|
| 
59865
 
8a20dd967385
rationalised and generalised some theorems concerning abs and x^2.
 
paulson <lp15@cam.ac.uk> 
parents: 
59741 
diff
changeset
 | 
810  | 
next  | 
| 63654 | 811  | 
assume ?rhs  | 
812  | 
then show ?lhs  | 
|
| 
59865
 
8a20dd967385
rationalised and generalised some theorems concerning abs and x^2.
 
paulson <lp15@cam.ac.uk> 
parents: 
59741 
diff
changeset
 | 
813  | 
by (auto intro!: power2_le_imp_le [OF _ abs_ge_zero])  | 
| 
 
8a20dd967385
rationalised and generalised some theorems concerning abs and x^2.
 
paulson <lp15@cam.ac.uk> 
parents: 
59741 
diff
changeset
 | 
814  | 
qed  | 
| 
 
8a20dd967385
rationalised and generalised some theorems concerning abs and x^2.
 
paulson <lp15@cam.ac.uk> 
parents: 
59741 
diff
changeset
 | 
815  | 
|
| 
74438
 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 
paulson <lp15@cam.ac.uk> 
parents: 
73869 
diff
changeset
 | 
816  | 
lemma power2_le_iff_abs_le:  | 
| 
 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 
paulson <lp15@cam.ac.uk> 
parents: 
73869 
diff
changeset
 | 
817  | 
"y \<ge> 0 \<Longrightarrow> x\<^sup>2 \<le> y\<^sup>2 \<longleftrightarrow> \<bar>x\<bar> \<le> y"  | 
| 
 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 
paulson <lp15@cam.ac.uk> 
parents: 
73869 
diff
changeset
 | 
818  | 
by (metis abs_le_square_iff abs_of_nonneg)  | 
| 
 
5827b91ef30e
new material from the Roth development, mostly about finite sets, disjoint famillies and partitions
 
paulson <lp15@cam.ac.uk> 
parents: 
73869 
diff
changeset
 | 
819  | 
|
| 61944 | 820  | 
lemma abs_square_le_1:"x\<^sup>2 \<le> 1 \<longleftrightarrow> \<bar>x\<bar> \<le> 1"  | 
| 63654 | 821  | 
using abs_le_square_iff [of x 1] by simp  | 
| 
59865
 
8a20dd967385
rationalised and generalised some theorems concerning abs and x^2.
 
paulson <lp15@cam.ac.uk> 
parents: 
59741 
diff
changeset
 | 
822  | 
|
| 61944 | 823  | 
lemma abs_square_eq_1: "x\<^sup>2 = 1 \<longleftrightarrow> \<bar>x\<bar> = 1"  | 
| 
59865
 
8a20dd967385
rationalised and generalised some theorems concerning abs and x^2.
 
paulson <lp15@cam.ac.uk> 
parents: 
59741 
diff
changeset
 | 
824  | 
by (auto simp add: abs_if power2_eq_1_iff)  | 
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61531 
diff
changeset
 | 
825  | 
|
| 61944 | 826  | 
lemma abs_square_less_1: "x\<^sup>2 < 1 \<longleftrightarrow> \<bar>x\<bar> < 1"  | 
| 63654 | 827  | 
using abs_square_eq_1 [of x] abs_square_le_1 [of x] by (auto simp add: le_less)  | 
| 
59865
 
8a20dd967385
rationalised and generalised some theorems concerning abs and x^2.
 
paulson <lp15@cam.ac.uk> 
parents: 
59741 
diff
changeset
 | 
828  | 
|
| 68611 | 829  | 
lemma square_le_1:  | 
830  | 
assumes "- 1 \<le> x" "x \<le> 1"  | 
|
831  | 
shows "x\<^sup>2 \<le> 1"  | 
|
832  | 
using assms  | 
|
833  | 
by (metis add.inverse_inverse linear mult_le_one neg_equal_0_iff_equal neg_le_iff_le power2_eq_square power_minus_Bit0)  | 
|
834  | 
||
| 30996 | 835  | 
end  | 
836  | 
||
| 60758 | 837  | 
subsection \<open>Miscellaneous rules\<close>  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
838  | 
|
| 79566 | 839  | 
context linordered_semidom  | 
840  | 
begin  | 
|
841  | 
||
842  | 
lemma self_le_power: "1 \<le> a \<Longrightarrow> 0 < n \<Longrightarrow> a \<le> a ^ n"  | 
|
| 60867 | 843  | 
using power_increasing [of 1 n a] power_one_right [of a] by auto  | 
| 
55718
 
34618f031ba9
A few lemmas about summations, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
55096 
diff
changeset
 | 
844  | 
|
| 79566 | 845  | 
lemma power_le_one_iff: "0 \<le> a \<Longrightarrow> a ^ n \<le> 1 \<longleftrightarrow> (n = 0 \<or> a \<le> 1)"  | 
846  | 
by (metis (mono_tags) gr0I nle_le one_le_power power_le_one self_le_power power_0)  | 
|
847  | 
||
848  | 
lemma power_less1_D: "a^n < 1 \<Longrightarrow> a < 1"  | 
|
849  | 
using not_le one_le_power by blast  | 
|
850  | 
||
851  | 
lemma power_less_one_iff: "0 \<le> a \<Longrightarrow> a ^ n < 1 \<longleftrightarrow> (n > 0 \<and> a < 1)"  | 
|
852  | 
by (metis (mono_tags) power_one power_strict_mono power_less1_D less_le_not_le neq0_conv power_0)  | 
|
853  | 
||
854  | 
end  | 
|
855  | 
||
| 
77138
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
856  | 
lemma power2_ge_1_iff: "x ^ 2 \<ge> 1 \<longleftrightarrow> x \<ge> 1 \<or> x \<le> (-1 :: 'a :: linordered_idom)"  | 
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
857  | 
using abs_le_square_iff[of 1 x] by (auto simp: abs_if split: if_splits)  | 
| 
 
c8597292cd41
Moved in a large number of highly useful library lemmas, mostly due to Manuel Eberl
 
paulson <lp15@cam.ac.uk> 
parents: 
75669 
diff
changeset
 | 
858  | 
|
| 63654 | 859  | 
lemma (in power) power_eq_if: "p ^ m = (if m=0 then 1 else p * (p ^ (m - 1)))"  | 
| 
47255
 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 
huffman 
parents: 
47241 
diff
changeset
 | 
860  | 
unfolding One_nat_def by (cases m) simp_all  | 
| 
 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 
huffman 
parents: 
47241 
diff
changeset
 | 
861  | 
|
| 63654 | 862  | 
lemma (in comm_semiring_1) power2_sum: "(x + y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 + 2 * x * y"  | 
| 
47192
 
0c0501cb6da6
move many lemmas from Nat_Numeral.thy to Power.thy or Num.thy
 
huffman 
parents: 
47191 
diff
changeset
 | 
863  | 
by (simp add: algebra_simps power2_eq_square mult_2_right)  | 
| 30996 | 864  | 
|
| 63654 | 865  | 
context comm_ring_1  | 
866  | 
begin  | 
|
867  | 
||
868  | 
lemma power2_diff: "(x - y)\<^sup>2 = x\<^sup>2 + y\<^sup>2 - 2 * x * y"  | 
|
| 58787 | 869  | 
by (simp add: algebra_simps power2_eq_square mult_2_right)  | 
| 30996 | 870  | 
|
| 63654 | 871  | 
lemma power2_commute: "(x - y)\<^sup>2 = (y - x)\<^sup>2"  | 
| 
60974
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60867 
diff
changeset
 | 
872  | 
by (simp add: algebra_simps power2_eq_square)  | 
| 
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60867 
diff
changeset
 | 
873  | 
|
| 63654 | 874  | 
lemma minus_power_mult_self: "(- a) ^ n * (- a) ^ n = a ^ (2 * n)"  | 
875  | 
by (simp add: power_mult_distrib [symmetric])  | 
|
876  | 
(simp add: power2_eq_square [symmetric] power_mult [symmetric])  | 
|
877  | 
||
878  | 
lemma minus_one_mult_self [simp]: "(- 1) ^ n * (- 1) ^ n = 1"  | 
|
| 
63417
 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 
haftmann 
parents: 
63040 
diff
changeset
 | 
879  | 
using minus_power_mult_self [of 1 n] by simp  | 
| 
 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 
haftmann 
parents: 
63040 
diff
changeset
 | 
880  | 
|
| 63654 | 881  | 
lemma left_minus_one_mult_self [simp]: "(- 1) ^ n * ((- 1) ^ n * a) = a"  | 
| 
63417
 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 
haftmann 
parents: 
63040 
diff
changeset
 | 
882  | 
by (simp add: mult.assoc [symmetric])  | 
| 
 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 
haftmann 
parents: 
63040 
diff
changeset
 | 
883  | 
|
| 63654 | 884  | 
end  | 
885  | 
||
| 60758 | 886  | 
text \<open>Simprules for comparisons where common factors can be cancelled.\<close>  | 
| 
47255
 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 
huffman 
parents: 
47241 
diff
changeset
 | 
887  | 
|
| 
 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 
huffman 
parents: 
47241 
diff
changeset
 | 
888  | 
lemmas zero_compare_simps =  | 
| 63654 | 889  | 
add_strict_increasing add_strict_increasing2 add_increasing  | 
890  | 
zero_le_mult_iff zero_le_divide_iff  | 
|
891  | 
zero_less_mult_iff zero_less_divide_iff  | 
|
892  | 
mult_le_0_iff divide_le_0_iff  | 
|
893  | 
mult_less_0_iff divide_less_0_iff  | 
|
894  | 
zero_le_power2 power2_less_0  | 
|
| 
47255
 
30a1692557b0
removed Nat_Numeral.thy, moving all theorems elsewhere
 
huffman 
parents: 
47241 
diff
changeset
 | 
895  | 
|
| 30313 | 896  | 
|
| 60758 | 897  | 
subsection \<open>Exponentiation for the Natural Numbers\<close>  | 
| 14577 | 898  | 
|
| 63654 | 899  | 
lemma nat_one_le_power [simp]: "Suc 0 \<le> i \<Longrightarrow> Suc 0 \<le> i ^ n"  | 
| 30996 | 900  | 
by (rule one_le_power [of i n, unfolded One_nat_def])  | 
| 23305 | 901  | 
|
| 63654 | 902  | 
lemma nat_zero_less_power_iff [simp]: "x ^ n > 0 \<longleftrightarrow> x > 0 \<or> n = 0"  | 
903  | 
for x :: nat  | 
|
| 30996 | 904  | 
by (induct n) auto  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
905  | 
|
| 63654 | 906  | 
lemma nat_power_eq_Suc_0_iff [simp]: "x ^ m = Suc 0 \<longleftrightarrow> m = 0 \<or> x = Suc 0"  | 
| 30996 | 907  | 
by (induct m) auto  | 
| 30056 | 908  | 
|
| 63654 | 909  | 
lemma power_Suc_0 [simp]: "Suc 0 ^ n = Suc 0"  | 
| 30996 | 910  | 
by simp  | 
| 30056 | 911  | 
|
| 63654 | 912  | 
text \<open>  | 
913  | 
Valid for the naturals, but what if \<open>0 < i < 1\<close>? Premises cannot be  | 
|
914  | 
weakened: consider the case where \<open>i = 0\<close>, \<open>m = 1\<close> and \<open>n = 0\<close>.  | 
|
915  | 
\<close>  | 
|
916  | 
||
| 21413 | 917  | 
lemma nat_power_less_imp_less:  | 
| 63654 | 918  | 
fixes i :: nat  | 
919  | 
assumes nonneg: "0 < i"  | 
|
| 30996 | 920  | 
assumes less: "i ^ m < i ^ n"  | 
| 21413 | 921  | 
shows "m < n"  | 
922  | 
proof (cases "i = 1")  | 
|
| 63654 | 923  | 
case True  | 
924  | 
with less power_one [where 'a = nat] show ?thesis by simp  | 
|
| 21413 | 925  | 
next  | 
| 63654 | 926  | 
case False  | 
927  | 
with nonneg have "1 < i" by auto  | 
|
| 21413 | 928  | 
from power_strict_increasing_iff [OF this] less show ?thesis ..  | 
929  | 
qed  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
8844 
diff
changeset
 | 
930  | 
|
| 
71435
 
d8fb621fea02
some lemmas about the lex ordering on lists, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
71398 
diff
changeset
 | 
931  | 
lemma power_gt_expt: "n > Suc 0 \<Longrightarrow> n^k > k"  | 
| 
 
d8fb621fea02
some lemmas about the lex ordering on lists, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
71398 
diff
changeset
 | 
932  | 
by (induction k) (auto simp: less_trans_Suc n_less_m_mult_n)  | 
| 
 
d8fb621fea02
some lemmas about the lex ordering on lists, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
71398 
diff
changeset
 | 
933  | 
|
| 73869 | 934  | 
lemma less_exp [simp]:  | 
| 72830 | 935  | 
\<open>n < 2 ^ n\<close>  | 
936  | 
by (simp add: power_gt_expt)  | 
|
937  | 
||
| 
71435
 
d8fb621fea02
some lemmas about the lex ordering on lists, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
71398 
diff
changeset
 | 
938  | 
lemma power_dvd_imp_le:  | 
| 
 
d8fb621fea02
some lemmas about the lex ordering on lists, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
71398 
diff
changeset
 | 
939  | 
fixes i :: nat  | 
| 
 
d8fb621fea02
some lemmas about the lex ordering on lists, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
71398 
diff
changeset
 | 
940  | 
assumes "i ^ m dvd i ^ n" "1 < i"  | 
| 
 
d8fb621fea02
some lemmas about the lex ordering on lists, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
71398 
diff
changeset
 | 
941  | 
shows "m \<le> n"  | 
| 
 
d8fb621fea02
some lemmas about the lex ordering on lists, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
71398 
diff
changeset
 | 
942  | 
using assms by (auto intro: power_le_imp_le_exp [OF \<open>1 < i\<close> dvd_imp_le])  | 
| 
33274
 
b6ff7db522b5
moved lemmas for dvd on nat to theories Nat and Power
 
haftmann 
parents: 
31998 
diff
changeset
 | 
943  | 
|
| 
70688
 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
944  | 
lemma dvd_power_iff_le:  | 
| 
 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
945  | 
fixes k::nat  | 
| 
 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
946  | 
shows "2 \<le> k \<Longrightarrow> ((k ^ m) dvd (k ^ n) \<longleftrightarrow> m \<le> n)"  | 
| 
 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
947  | 
using le_imp_power_dvd power_dvd_imp_le by force  | 
| 
 
3d894e1cfc75
new material on Analysis, plus some rearrangements
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
948  | 
|
| 63654 | 949  | 
lemma power2_nat_le_eq_le: "m\<^sup>2 \<le> n\<^sup>2 \<longleftrightarrow> m \<le> n"  | 
950  | 
for m n :: nat  | 
|
| 
51263
 
31e786e0e6a7
turned example into library for comparing growth of functions
 
haftmann 
parents: 
49824 
diff
changeset
 | 
951  | 
by (auto intro: power2_le_imp_le power_mono)  | 
| 
 
31e786e0e6a7
turned example into library for comparing growth of functions
 
haftmann 
parents: 
49824 
diff
changeset
 | 
952  | 
|
| 
 
31e786e0e6a7
turned example into library for comparing growth of functions
 
haftmann 
parents: 
49824 
diff
changeset
 | 
953  | 
lemma power2_nat_le_imp_le:  | 
| 
 
31e786e0e6a7
turned example into library for comparing growth of functions
 
haftmann 
parents: 
49824 
diff
changeset
 | 
954  | 
fixes m n :: nat  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
955  | 
assumes "m\<^sup>2 \<le> n"  | 
| 
51263
 
31e786e0e6a7
turned example into library for comparing growth of functions
 
haftmann 
parents: 
49824 
diff
changeset
 | 
956  | 
shows "m \<le> n"  | 
| 54249 | 957  | 
proof (cases m)  | 
| 63654 | 958  | 
case 0  | 
959  | 
then show ?thesis by simp  | 
|
| 54249 | 960  | 
next  | 
961  | 
case (Suc k)  | 
|
962  | 
show ?thesis  | 
|
963  | 
proof (rule ccontr)  | 
|
| 63654 | 964  | 
assume "\<not> ?thesis"  | 
| 54249 | 965  | 
then have "n < m" by simp  | 
966  | 
with assms Suc show False  | 
|
| 60867 | 967  | 
by (simp add: power2_eq_square)  | 
| 54249 | 968  | 
qed  | 
969  | 
qed  | 
|
| 
51263
 
31e786e0e6a7
turned example into library for comparing growth of functions
 
haftmann 
parents: 
49824 
diff
changeset
 | 
970  | 
|
| 64065 | 971  | 
lemma ex_power_ivl1: fixes b k :: nat assumes "b \<ge> 2"  | 
972  | 
shows "k \<ge> 1 \<Longrightarrow> \<exists>n. b^n \<le> k \<and> k < b^(n+1)" (is "_ \<Longrightarrow> \<exists>n. ?P k n")  | 
|
973  | 
proof(induction k)  | 
|
974  | 
case 0 thus ?case by simp  | 
|
975  | 
next  | 
|
976  | 
case (Suc k)  | 
|
977  | 
show ?case  | 
|
978  | 
proof cases  | 
|
979  | 
assume "k=0"  | 
|
980  | 
hence "?P (Suc k) 0" using assms by simp  | 
|
981  | 
thus ?case ..  | 
|
982  | 
next  | 
|
983  | 
assume "k\<noteq>0"  | 
|
984  | 
with Suc obtain n where IH: "?P k n" by auto  | 
|
985  | 
show ?case  | 
|
986  | 
proof (cases "k = b^(n+1) - 1")  | 
|
987  | 
case True  | 
|
988  | 
hence "?P (Suc k) (n+1)" using assms  | 
|
989  | 
by (simp add: power_less_power_Suc)  | 
|
990  | 
thus ?thesis ..  | 
|
991  | 
next  | 
|
992  | 
case False  | 
|
993  | 
hence "?P (Suc k) n" using IH by auto  | 
|
994  | 
thus ?thesis ..  | 
|
995  | 
qed  | 
|
996  | 
qed  | 
|
997  | 
qed  | 
|
998  | 
||
999  | 
lemma ex_power_ivl2: fixes b k :: nat assumes "b \<ge> 2" "k \<ge> 2"  | 
|
| 
71435
 
d8fb621fea02
some lemmas about the lex ordering on lists, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
71398 
diff
changeset
 | 
1000  | 
shows "\<exists>n. b^n < k \<and> k \<le> b^(n+1)"  | 
| 64065 | 1001  | 
proof -  | 
1002  | 
have "1 \<le> k - 1" using assms(2) by arith  | 
|
1003  | 
from ex_power_ivl1[OF assms(1) this]  | 
|
1004  | 
obtain n where "b ^ n \<le> k - 1 \<and> k - 1 < b ^ (n + 1)" ..  | 
|
1005  | 
hence "b^n < k \<and> k \<le> b^(n+1)" using assms by auto  | 
|
1006  | 
thus ?thesis ..  | 
|
1007  | 
qed  | 
|
1008  | 
||
| 63654 | 1009  | 
|
| 60758 | 1010  | 
subsubsection \<open>Cardinality of the Powerset\<close>  | 
| 55096 | 1011  | 
|
1012  | 
lemma card_UNIV_bool [simp]: "card (UNIV :: bool set) = 2"  | 
|
1013  | 
unfolding UNIV_bool by simp  | 
|
1014  | 
||
1015  | 
lemma card_Pow: "finite A \<Longrightarrow> card (Pow A) = 2 ^ card A"  | 
|
1016  | 
proof (induct rule: finite_induct)  | 
|
| 
61649
 
268d88ec9087
Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
 
paulson <lp15@cam.ac.uk> 
parents: 
61531 
diff
changeset
 | 
1017  | 
case empty  | 
| 64964 | 1018  | 
show ?case by simp  | 
| 55096 | 1019  | 
next  | 
1020  | 
case (insert x A)  | 
|
| 64964 | 1021  | 
  from \<open>x \<notin> A\<close> have disjoint: "Pow A \<inter> insert x ` Pow A = {}" by blast
 | 
1022  | 
from \<open>x \<notin> A\<close> have inj_on: "inj_on (insert x) (Pow A)"  | 
|
1023  | 
unfolding inj_on_def by auto  | 
|
1024  | 
||
1025  | 
have "card (Pow (insert x A)) = card (Pow A \<union> insert x ` Pow A)"  | 
|
1026  | 
by (simp only: Pow_insert)  | 
|
1027  | 
also have "\<dots> = card (Pow A) + card (insert x ` Pow A)"  | 
|
1028  | 
by (rule card_Un_disjoint) (use \<open>finite A\<close> disjoint in simp_all)  | 
|
1029  | 
also from inj_on have "card (insert x ` Pow A) = card (Pow A)"  | 
|
1030  | 
by (rule card_image)  | 
|
1031  | 
also have "\<dots> + \<dots> = 2 * \<dots>" by (simp add: mult_2)  | 
|
1032  | 
also from insert(3) have "\<dots> = 2 ^ Suc (card A)" by simp  | 
|
1033  | 
also from insert(1,2) have "Suc (card A) = card (insert x A)"  | 
|
1034  | 
by (rule card_insert_disjoint [symmetric])  | 
|
1035  | 
finally show ?case .  | 
|
| 55096 | 1036  | 
qed  | 
1037  | 
||
| 57418 | 1038  | 
|
| 60758 | 1039  | 
subsection \<open>Code generator tweak\<close>  | 
| 
31155
 
92d8ff6af82c
monomorphic code generation for power operations
 
haftmann 
parents: 
31021 
diff
changeset
 | 
1040  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51263 
diff
changeset
 | 
1041  | 
code_identifier  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
51263 
diff
changeset
 | 
1042  | 
code_module Power \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith  | 
| 33364 | 1043  | 
|
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
1044  | 
end  |