| author | wenzelm | 
| Thu, 13 Nov 2008 21:29:19 +0100 | |
| changeset 28746 | c1b151a60a66 | 
| parent 26565 | 522f45a8604e | 
| child 28906 | 5f568bfc58d7 | 
| permissions | -rw-r--r-- | 
| 
9435
 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 
wenzelm 
parents: 
9013 
diff
changeset
 | 
1  | 
(* Title : HOL/Real/RealPow.thy  | 
| 7219 | 2  | 
ID : $Id$  | 
| 
7077
 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Author : Jacques D. Fleuriot  | 
| 
 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 
paulson 
parents:  
diff
changeset
 | 
4  | 
Copyright : 1998 University of Cambridge  | 
| 20634 | 5  | 
*)  | 
| 
7077
 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 
paulson 
parents:  
diff
changeset
 | 
6  | 
|
| 20634 | 7  | 
header {* Natural powers theory *}
 | 
| 
7077
 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 
paulson 
parents:  
diff
changeset
 | 
8  | 
|
| 15131 | 9  | 
theory RealPow  | 
| 15140 | 10  | 
imports RealDef  | 
| 15131 | 11  | 
begin  | 
| 
9435
 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 
wenzelm 
parents: 
9013 
diff
changeset
 | 
12  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
13  | 
declare abs_mult_self [simp]  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
14  | 
|
| 
26565
 
522f45a8604e
instantiation replacing primitive instance plus overloaded defs; realpow.simps now named power_real.simps
 
haftmann 
parents: 
25875 
diff
changeset
 | 
15  | 
instantiation real :: recpower  | 
| 
 
522f45a8604e
instantiation replacing primitive instance plus overloaded defs; realpow.simps now named power_real.simps
 
haftmann 
parents: 
25875 
diff
changeset
 | 
16  | 
begin  | 
| 
7077
 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 
paulson 
parents:  
diff
changeset
 | 
17  | 
|
| 
26565
 
522f45a8604e
instantiation replacing primitive instance plus overloaded defs; realpow.simps now named power_real.simps
 
haftmann 
parents: 
25875 
diff
changeset
 | 
18  | 
primrec power_real where  | 
| 
 
522f45a8604e
instantiation replacing primitive instance plus overloaded defs; realpow.simps now named power_real.simps
 
haftmann 
parents: 
25875 
diff
changeset
 | 
19  | 
realpow_0: "r ^ 0 = (1\<Colon>real)"  | 
| 
 
522f45a8604e
instantiation replacing primitive instance plus overloaded defs; realpow.simps now named power_real.simps
 
haftmann 
parents: 
25875 
diff
changeset
 | 
20  | 
| realpow_Suc: "r ^ Suc n = (r\<Colon>real) * r ^ n"  | 
| 
7077
 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 
paulson 
parents:  
diff
changeset
 | 
21  | 
|
| 
26565
 
522f45a8604e
instantiation replacing primitive instance plus overloaded defs; realpow.simps now named power_real.simps
 
haftmann 
parents: 
25875 
diff
changeset
 | 
22  | 
instance proof  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
23  | 
fix z :: real  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
24  | 
fix n :: nat  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
25  | 
show "z^0 = 1" by simp  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
26  | 
show "z^(Suc n) = z * (z^n)" by simp  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
27  | 
qed  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
28  | 
|
| 
26565
 
522f45a8604e
instantiation replacing primitive instance plus overloaded defs; realpow.simps now named power_real.simps
 
haftmann 
parents: 
25875 
diff
changeset
 | 
29  | 
end  | 
| 
 
522f45a8604e
instantiation replacing primitive instance plus overloaded defs; realpow.simps now named power_real.simps
 
haftmann 
parents: 
25875 
diff
changeset
 | 
30  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
31  | 
|
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
32  | 
lemma two_realpow_ge_one [simp]: "(1::real) \<le> 2 ^ n"  | 
| 25875 | 33  | 
by simp  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
34  | 
|
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
35  | 
lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n"  | 
| 15251 | 36  | 
apply (induct "n")  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
37  | 
apply (auto simp add: real_of_nat_Suc)  | 
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14352 
diff
changeset
 | 
38  | 
apply (subst mult_2)  | 
| 22962 | 39  | 
apply (rule add_less_le_mono)  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
40  | 
apply (auto simp add: two_realpow_ge_one)  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
41  | 
done  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
42  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
43  | 
lemma realpow_Suc_le_self: "[| 0 \<le> r; r \<le> (1::real) |] ==> r ^ Suc n \<le> r"  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
44  | 
by (insert power_decreasing [of 1 "Suc n" r], simp)  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
45  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
46  | 
lemma realpow_minus_mult [rule_format]:  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
47  | 
"0 < n --> (x::real) ^ (n - 1) * x = x ^ n"  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
48  | 
apply (simp split add: nat_diff_split)  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
49  | 
done  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
50  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
51  | 
lemma realpow_two_mult_inverse [simp]:  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
52  | 
"r \<noteq> 0 ==> r * inverse r ^Suc (Suc 0) = inverse (r::real)"  | 
| 23292 | 53  | 
by (simp add: real_mult_assoc [symmetric])  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
54  | 
|
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
55  | 
lemma realpow_two_minus [simp]: "(-x)^Suc (Suc 0) = (x::real)^Suc (Suc 0)"  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
56  | 
by simp  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
57  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
58  | 
lemma realpow_two_diff:  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
59  | 
"(x::real)^Suc (Suc 0) - y^Suc (Suc 0) = (x - y) * (x + y)"  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
60  | 
apply (unfold real_diff_def)  | 
| 
23477
 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 
nipkow 
parents: 
23293 
diff
changeset
 | 
61  | 
apply (simp add: ring_simps)  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
62  | 
done  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
63  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
64  | 
lemma realpow_two_disj:  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
65  | 
"((x::real)^Suc (Suc 0) = y^Suc (Suc 0)) = (x = y | x = -y)"  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
66  | 
apply (cut_tac x = x and y = y in realpow_two_diff)  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
67  | 
apply (auto simp del: realpow_Suc)  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
68  | 
done  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
69  | 
|
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
70  | 
lemma realpow_real_of_nat: "real (m::nat) ^ n = real (m ^ n)"  | 
| 15251 | 71  | 
apply (induct "n")  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
72  | 
apply (auto simp add: real_of_nat_one real_of_nat_mult)  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
73  | 
done  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
74  | 
|
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
75  | 
lemma realpow_real_of_nat_two_pos [simp] : "0 < real (Suc (Suc 0) ^ n)"  | 
| 15251 | 76  | 
apply (induct "n")  | 
| 14334 | 77  | 
apply (auto simp add: real_of_nat_mult zero_less_mult_iff)  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
78  | 
done  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
79  | 
|
| 22962 | 80  | 
(* used by AFP Integration theory *)  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
81  | 
lemma realpow_increasing:  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
82  | 
"[|(0::real) \<le> x; 0 \<le> y; x ^ Suc n \<le> y ^ Suc n|] ==> x \<le> y"  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
83  | 
by (rule power_le_imp_le_base)  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
84  | 
|
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
85  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
86  | 
subsection{*Literal Arithmetic Involving Powers, Type @{typ real}*}
 | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
87  | 
|
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
88  | 
lemma real_of_int_power: "real (x::int) ^ n = real (x ^ n)"  | 
| 15251 | 89  | 
apply (induct "n")  | 
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14352 
diff
changeset
 | 
90  | 
apply (simp_all add: nat_mult_distrib)  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
91  | 
done  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
92  | 
declare real_of_int_power [symmetric, simp]  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
93  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
94  | 
lemma power_real_number_of:  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
95  | 
"(number_of v :: real) ^ n = real ((number_of v :: int) ^ n)"  | 
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14352 
diff
changeset
 | 
96  | 
by (simp only: real_number_of [symmetric] real_of_int_power)  | 
| 
14265
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
97  | 
|
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
98  | 
declare power_real_number_of [of _ "number_of w", standard, simp]  | 
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
99  | 
|
| 
 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 
paulson 
parents: 
12018 
diff
changeset
 | 
100  | 
|
| 22967 | 101  | 
subsection {* Properties of Squares *}
 | 
102  | 
||
103  | 
lemma sum_squares_ge_zero:  | 
|
104  | 
fixes x y :: "'a::ordered_ring_strict"  | 
|
105  | 
shows "0 \<le> x * x + y * y"  | 
|
106  | 
by (intro add_nonneg_nonneg zero_le_square)  | 
|
107  | 
||
108  | 
lemma not_sum_squares_lt_zero:  | 
|
109  | 
fixes x y :: "'a::ordered_ring_strict"  | 
|
110  | 
shows "\<not> x * x + y * y < 0"  | 
|
111  | 
by (simp add: linorder_not_less sum_squares_ge_zero)  | 
|
112  | 
||
113  | 
lemma sum_nonneg_eq_zero_iff:  | 
|
114  | 
fixes x y :: "'a::pordered_ab_group_add"  | 
|
115  | 
assumes x: "0 \<le> x" and y: "0 \<le> y"  | 
|
116  | 
shows "(x + y = 0) = (x = 0 \<and> y = 0)"  | 
|
117  | 
proof (auto)  | 
|
118  | 
from y have "x + 0 \<le> x + y" by (rule add_left_mono)  | 
|
119  | 
also assume "x + y = 0"  | 
|
120  | 
finally have "x \<le> 0" by simp  | 
|
121  | 
thus "x = 0" using x by (rule order_antisym)  | 
|
122  | 
next  | 
|
123  | 
from x have "0 + y \<le> x + y" by (rule add_right_mono)  | 
|
124  | 
also assume "x + y = 0"  | 
|
125  | 
finally have "y \<le> 0" by simp  | 
|
126  | 
thus "y = 0" using y by (rule order_antisym)  | 
|
127  | 
qed  | 
|
128  | 
||
129  | 
lemma sum_squares_eq_zero_iff:  | 
|
130  | 
fixes x y :: "'a::ordered_ring_strict"  | 
|
131  | 
shows "(x * x + y * y = 0) = (x = 0 \<and> y = 0)"  | 
|
| 23096 | 132  | 
by (simp add: sum_nonneg_eq_zero_iff)  | 
| 22967 | 133  | 
|
134  | 
lemma sum_squares_le_zero_iff:  | 
|
135  | 
fixes x y :: "'a::ordered_ring_strict"  | 
|
136  | 
shows "(x * x + y * y \<le> 0) = (x = 0 \<and> y = 0)"  | 
|
137  | 
by (simp add: order_le_less not_sum_squares_lt_zero sum_squares_eq_zero_iff)  | 
|
138  | 
||
139  | 
lemma sum_squares_gt_zero_iff:  | 
|
140  | 
fixes x y :: "'a::ordered_ring_strict"  | 
|
141  | 
shows "(0 < x * x + y * y) = (x \<noteq> 0 \<or> y \<noteq> 0)"  | 
|
142  | 
by (simp add: order_less_le sum_squares_ge_zero sum_squares_eq_zero_iff)  | 
|
143  | 
||
144  | 
lemma sum_power2_ge_zero:  | 
|
145  | 
  fixes x y :: "'a::{ordered_idom,recpower}"
 | 
|
146  | 
shows "0 \<le> x\<twosuperior> + y\<twosuperior>"  | 
|
147  | 
unfolding power2_eq_square by (rule sum_squares_ge_zero)  | 
|
148  | 
||
149  | 
lemma not_sum_power2_lt_zero:  | 
|
150  | 
  fixes x y :: "'a::{ordered_idom,recpower}"
 | 
|
151  | 
shows "\<not> x\<twosuperior> + y\<twosuperior> < 0"  | 
|
152  | 
unfolding power2_eq_square by (rule not_sum_squares_lt_zero)  | 
|
153  | 
||
154  | 
lemma sum_power2_eq_zero_iff:  | 
|
155  | 
  fixes x y :: "'a::{ordered_idom,recpower}"
 | 
|
156  | 
shows "(x\<twosuperior> + y\<twosuperior> = 0) = (x = 0 \<and> y = 0)"  | 
|
157  | 
unfolding power2_eq_square by (rule sum_squares_eq_zero_iff)  | 
|
158  | 
||
159  | 
lemma sum_power2_le_zero_iff:  | 
|
160  | 
  fixes x y :: "'a::{ordered_idom,recpower}"
 | 
|
161  | 
shows "(x\<twosuperior> + y\<twosuperior> \<le> 0) = (x = 0 \<and> y = 0)"  | 
|
162  | 
unfolding power2_eq_square by (rule sum_squares_le_zero_iff)  | 
|
163  | 
||
164  | 
lemma sum_power2_gt_zero_iff:  | 
|
165  | 
  fixes x y :: "'a::{ordered_idom,recpower}"
 | 
|
166  | 
shows "(0 < x\<twosuperior> + y\<twosuperior>) = (x \<noteq> 0 \<or> y \<noteq> 0)"  | 
|
167  | 
unfolding power2_eq_square by (rule sum_squares_gt_zero_iff)  | 
|
168  | 
||
169  | 
||
| 22970 | 170  | 
subsection{* Squares of Reals *}
 | 
171  | 
||
172  | 
lemma real_two_squares_add_zero_iff [simp]:  | 
|
173  | 
"(x * x + y * y = 0) = ((x::real) = 0 \<and> y = 0)"  | 
|
174  | 
by (rule sum_squares_eq_zero_iff)  | 
|
175  | 
||
176  | 
lemma real_sum_squares_cancel: "x * x + y * y = 0 ==> x = (0::real)"  | 
|
177  | 
by simp  | 
|
178  | 
||
179  | 
lemma real_sum_squares_cancel2: "x * x + y * y = 0 ==> y = (0::real)"  | 
|
180  | 
by simp  | 
|
181  | 
||
182  | 
lemma real_mult_self_sum_ge_zero: "(0::real) \<le> x*x + y*y"  | 
|
183  | 
by (rule sum_squares_ge_zero)  | 
|
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
184  | 
|
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
185  | 
lemma real_sum_squares_cancel_a: "x * x = -(y * y) ==> x = (0::real) & y=0"  | 
| 22970 | 186  | 
by (simp add: real_add_eq_0_iff [symmetric])  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
187  | 
|
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
188  | 
lemma real_squared_diff_one_factored: "x*x - (1::real) = (x + 1)*(x - 1)"  | 
| 22970 | 189  | 
by (simp add: left_distrib right_diff_distrib)  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
190  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
191  | 
lemma real_mult_is_one [simp]: "(x*x = (1::real)) = (x = 1 | x = - 1)"  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
192  | 
apply auto  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
193  | 
apply (drule right_minus_eq [THEN iffD2])  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
194  | 
apply (auto simp add: real_squared_diff_one_factored)  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
195  | 
done  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
196  | 
|
| 22970 | 197  | 
lemma real_sum_squares_not_zero: "x ~= 0 ==> x * x + y * y ~= (0::real)"  | 
198  | 
by simp  | 
|
199  | 
||
200  | 
lemma real_sum_squares_not_zero2: "y ~= 0 ==> x * x + y * y ~= (0::real)"  | 
|
201  | 
by simp  | 
|
202  | 
||
203  | 
lemma realpow_two_sum_zero_iff [simp]:  | 
|
204  | 
"(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)"  | 
|
205  | 
by (rule sum_power2_eq_zero_iff)  | 
|
206  | 
||
207  | 
lemma realpow_two_le_add_order [simp]: "(0::real) \<le> u ^ 2 + v ^ 2"  | 
|
208  | 
by (rule sum_power2_ge_zero)  | 
|
209  | 
||
210  | 
lemma realpow_two_le_add_order2 [simp]: "(0::real) \<le> u ^ 2 + v ^ 2 + w ^ 2"  | 
|
211  | 
by (intro add_nonneg_nonneg zero_le_power2)  | 
|
212  | 
||
213  | 
lemma real_sum_square_gt_zero: "x ~= 0 ==> (0::real) < x * x + y * y"  | 
|
214  | 
by (simp add: sum_squares_gt_zero_iff)  | 
|
215  | 
||
216  | 
lemma real_sum_square_gt_zero2: "y ~= 0 ==> (0::real) < x * x + y * y"  | 
|
217  | 
by (simp add: sum_squares_gt_zero_iff)  | 
|
218  | 
||
219  | 
lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))"  | 
|
220  | 
by (rule_tac j = 0 in real_le_trans, auto)  | 
|
221  | 
||
222  | 
lemma realpow_square_minus_le [simp]: "-(u ^ 2) \<le> (x::real) ^ 2"  | 
|
223  | 
by (auto simp add: power2_eq_square)  | 
|
224  | 
||
225  | 
(* The following theorem is by Benjamin Porter *)  | 
|
226  | 
lemma real_sq_order:  | 
|
227  | 
fixes x::real  | 
|
228  | 
assumes xgt0: "0 \<le> x" and ygt0: "0 \<le> y" and sq: "x^2 \<le> y^2"  | 
|
229  | 
shows "x \<le> y"  | 
|
230  | 
proof -  | 
|
231  | 
from sq have "x ^ Suc (Suc 0) \<le> y ^ Suc (Suc 0)"  | 
|
232  | 
by (simp only: numeral_2_eq_2)  | 
|
233  | 
thus "x \<le> y" using ygt0  | 
|
234  | 
by (rule power_le_imp_le_base)  | 
|
235  | 
qed  | 
|
236  | 
||
237  | 
||
238  | 
subsection {*Various Other Theorems*}
 | 
|
239  | 
||
| 14304 | 240  | 
lemma real_le_add_half_cancel: "(x + y/2 \<le> (y::real)) = (x \<le> y /2)"  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
241  | 
by auto  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
242  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
243  | 
lemma real_minus_half_eq [simp]: "(x::real) - x/2 = x/2"  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
244  | 
by auto  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
245  | 
|
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
246  | 
lemma real_mult_inverse_cancel:  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
247  | 
"[|(0::real) < x; 0 < x1; x1 * y < x * u |]  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
248  | 
==> inverse x * y < inverse x1 * u"  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
249  | 
apply (rule_tac c=x in mult_less_imp_less_left)  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
250  | 
apply (auto simp add: real_mult_assoc [symmetric])  | 
| 14334 | 251  | 
apply (simp (no_asm) add: mult_ac)  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
252  | 
apply (rule_tac c=x1 in mult_less_imp_less_right)  | 
| 14334 | 253  | 
apply (auto simp add: mult_ac)  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
254  | 
done  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
255  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
256  | 
lemma real_mult_inverse_cancel2:  | 
| 
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
257  | 
"[|(0::real) < x;0 < x1; x1 * y < x * u |] ==> y * inverse x < u * inverse x1"  | 
| 14334 | 258  | 
apply (auto dest: real_mult_inverse_cancel simp add: mult_ac)  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
259  | 
done  | 
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
260  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
261  | 
lemma inverse_real_of_nat_gt_zero [simp]: "0 < inverse (real (Suc n))"  | 
| 
20517
 
86343f2386a8
simplify some proofs, remove obsolete realpow_divide
 
huffman 
parents: 
19765 
diff
changeset
 | 
262  | 
by simp  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
263  | 
|
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
264  | 
lemma inverse_real_of_nat_ge_zero [simp]: "0 \<le> inverse (real (Suc n))"  | 
| 
20517
 
86343f2386a8
simplify some proofs, remove obsolete realpow_divide
 
huffman 
parents: 
19765 
diff
changeset
 | 
265  | 
by simp  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
266  | 
|
| 
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
267  | 
lemma realpow_num_eq_if: "(m::real) ^ n = (if n=0 then 1 else m * m ^ (n - 1))"  | 
| 
14348
 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 
paulson 
parents: 
14334 
diff
changeset
 | 
268  | 
by (case_tac "n", auto)  | 
| 
14268
 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 
paulson 
parents: 
14265 
diff
changeset
 | 
269  | 
|
| 
7077
 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 
paulson 
parents:  
diff
changeset
 | 
270  | 
end  |