author | wenzelm |
Tue, 03 Mar 2009 14:07:43 +0100 | |
changeset 30211 | 556d1810cdad |
parent 30082 | 43c5b7bfc791 |
child 30273 | ecd6f0ca62ea |
permissions | -rw-r--r-- |
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
28906
diff
changeset
|
1 |
(* Title : HOL/RealPow.thy |
7077
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
2 |
Author : Jacques D. Fleuriot |
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
3 |
Copyright : 1998 University of Cambridge |
20634 | 4 |
*) |
7077
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
5 |
|
20634 | 6 |
header {* Natural powers theory *} |
7077
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
7 |
|
15131 | 8 |
theory RealPow |
15140 | 9 |
imports RealDef |
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
28906
diff
changeset
|
10 |
uses ("Tools/float_syntax.ML") |
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]: |
30082
43c5b7bfc791
make more proofs work whether or not One_nat_def is a simp rule
huffman
parents:
29667
diff
changeset
|
47 |
"0 < n --> (x::real) ^ (n - 1) * x = x ^ n" |
43c5b7bfc791
make more proofs work whether or not One_nat_def is a simp rule
huffman
parents:
29667
diff
changeset
|
48 |
unfolding One_nat_def |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
49 |
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
|
50 |
done |
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
51 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
52 |
lemma realpow_two_mult_inverse [simp]: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
53 |
"r \<noteq> 0 ==> r * inverse r ^Suc (Suc 0) = inverse (r::real)" |
23292 | 54 |
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
|
55 |
|
14268
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
56 |
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
|
57 |
by simp |
14265
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
58 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
59 |
lemma realpow_two_diff: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
60 |
"(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
|
61 |
apply (unfold real_diff_def) |
29667 | 62 |
apply (simp add: algebra_simps) |
14265
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
63 |
done |
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
64 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
65 |
lemma realpow_two_disj: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
66 |
"((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
|
67 |
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
|
68 |
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
|
69 |
done |
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
70 |
|
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
71 |
lemma realpow_real_of_nat: "real (m::nat) ^ n = real (m ^ n)" |
15251 | 72 |
apply (induct "n") |
14265
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
73 |
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
|
74 |
done |
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
75 |
|
14268
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
76 |
lemma realpow_real_of_nat_two_pos [simp] : "0 < real (Suc (Suc 0) ^ n)" |
15251 | 77 |
apply (induct "n") |
14334 | 78 |
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
|
79 |
done |
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
80 |
|
22962 | 81 |
(* 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
|
82 |
lemma realpow_increasing: |
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
83 |
"[|(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
|
84 |
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
|
85 |
|
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
86 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
87 |
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
|
88 |
|
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
89 |
lemma real_of_int_power: "real (x::int) ^ n = real (x ^ n)" |
15251 | 90 |
apply (induct "n") |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14352
diff
changeset
|
91 |
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
|
92 |
done |
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
93 |
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
|
94 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
95 |
lemma power_real_number_of: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
96 |
"(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
|
97 |
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
|
98 |
|
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
99 |
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
|
100 |
|
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
12018
diff
changeset
|
101 |
|
22967 | 102 |
subsection {* Properties of Squares *} |
103 |
||
104 |
lemma sum_squares_ge_zero: |
|
105 |
fixes x y :: "'a::ordered_ring_strict" |
|
106 |
shows "0 \<le> x * x + y * y" |
|
107 |
by (intro add_nonneg_nonneg zero_le_square) |
|
108 |
||
109 |
lemma not_sum_squares_lt_zero: |
|
110 |
fixes x y :: "'a::ordered_ring_strict" |
|
111 |
shows "\<not> x * x + y * y < 0" |
|
112 |
by (simp add: linorder_not_less sum_squares_ge_zero) |
|
113 |
||
114 |
lemma sum_nonneg_eq_zero_iff: |
|
115 |
fixes x y :: "'a::pordered_ab_group_add" |
|
116 |
assumes x: "0 \<le> x" and y: "0 \<le> y" |
|
117 |
shows "(x + y = 0) = (x = 0 \<and> y = 0)" |
|
118 |
proof (auto) |
|
119 |
from y have "x + 0 \<le> x + y" by (rule add_left_mono) |
|
120 |
also assume "x + y = 0" |
|
121 |
finally have "x \<le> 0" by simp |
|
122 |
thus "x = 0" using x by (rule order_antisym) |
|
123 |
next |
|
124 |
from x have "0 + y \<le> x + y" by (rule add_right_mono) |
|
125 |
also assume "x + y = 0" |
|
126 |
finally have "y \<le> 0" by simp |
|
127 |
thus "y = 0" using y by (rule order_antisym) |
|
128 |
qed |
|
129 |
||
130 |
lemma sum_squares_eq_zero_iff: |
|
131 |
fixes x y :: "'a::ordered_ring_strict" |
|
132 |
shows "(x * x + y * y = 0) = (x = 0 \<and> y = 0)" |
|
23096 | 133 |
by (simp add: sum_nonneg_eq_zero_iff) |
22967 | 134 |
|
135 |
lemma sum_squares_le_zero_iff: |
|
136 |
fixes x y :: "'a::ordered_ring_strict" |
|
137 |
shows "(x * x + y * y \<le> 0) = (x = 0 \<and> y = 0)" |
|
138 |
by (simp add: order_le_less not_sum_squares_lt_zero sum_squares_eq_zero_iff) |
|
139 |
||
140 |
lemma sum_squares_gt_zero_iff: |
|
141 |
fixes x y :: "'a::ordered_ring_strict" |
|
142 |
shows "(0 < x * x + y * y) = (x \<noteq> 0 \<or> y \<noteq> 0)" |
|
143 |
by (simp add: order_less_le sum_squares_ge_zero sum_squares_eq_zero_iff) |
|
144 |
||
145 |
lemma sum_power2_ge_zero: |
|
146 |
fixes x y :: "'a::{ordered_idom,recpower}" |
|
147 |
shows "0 \<le> x\<twosuperior> + y\<twosuperior>" |
|
148 |
unfolding power2_eq_square by (rule sum_squares_ge_zero) |
|
149 |
||
150 |
lemma not_sum_power2_lt_zero: |
|
151 |
fixes x y :: "'a::{ordered_idom,recpower}" |
|
152 |
shows "\<not> x\<twosuperior> + y\<twosuperior> < 0" |
|
153 |
unfolding power2_eq_square by (rule not_sum_squares_lt_zero) |
|
154 |
||
155 |
lemma sum_power2_eq_zero_iff: |
|
156 |
fixes x y :: "'a::{ordered_idom,recpower}" |
|
157 |
shows "(x\<twosuperior> + y\<twosuperior> = 0) = (x = 0 \<and> y = 0)" |
|
158 |
unfolding power2_eq_square by (rule sum_squares_eq_zero_iff) |
|
159 |
||
160 |
lemma sum_power2_le_zero_iff: |
|
161 |
fixes x y :: "'a::{ordered_idom,recpower}" |
|
162 |
shows "(x\<twosuperior> + y\<twosuperior> \<le> 0) = (x = 0 \<and> y = 0)" |
|
163 |
unfolding power2_eq_square by (rule sum_squares_le_zero_iff) |
|
164 |
||
165 |
lemma sum_power2_gt_zero_iff: |
|
166 |
fixes x y :: "'a::{ordered_idom,recpower}" |
|
167 |
shows "(0 < x\<twosuperior> + y\<twosuperior>) = (x \<noteq> 0 \<or> y \<noteq> 0)" |
|
168 |
unfolding power2_eq_square by (rule sum_squares_gt_zero_iff) |
|
169 |
||
170 |
||
22970 | 171 |
subsection{* Squares of Reals *} |
172 |
||
173 |
lemma real_two_squares_add_zero_iff [simp]: |
|
174 |
"(x * x + y * y = 0) = ((x::real) = 0 \<and> y = 0)" |
|
175 |
by (rule sum_squares_eq_zero_iff) |
|
176 |
||
177 |
lemma real_sum_squares_cancel: "x * x + y * y = 0 ==> x = (0::real)" |
|
178 |
by simp |
|
179 |
||
180 |
lemma real_sum_squares_cancel2: "x * x + y * y = 0 ==> y = (0::real)" |
|
181 |
by simp |
|
182 |
||
183 |
lemma real_mult_self_sum_ge_zero: "(0::real) \<le> x*x + y*y" |
|
184 |
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
|
185 |
|
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
186 |
lemma real_sum_squares_cancel_a: "x * x = -(y * y) ==> x = (0::real) & y=0" |
22970 | 187 |
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
|
188 |
|
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
189 |
lemma real_squared_diff_one_factored: "x*x - (1::real) = (x + 1)*(x - 1)" |
22970 | 190 |
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
|
191 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
192 |
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
|
193 |
apply auto |
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
194 |
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
|
195 |
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
|
196 |
done |
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
197 |
|
22970 | 198 |
lemma real_sum_squares_not_zero: "x ~= 0 ==> x * x + y * y ~= (0::real)" |
199 |
by simp |
|
200 |
||
201 |
lemma real_sum_squares_not_zero2: "y ~= 0 ==> x * x + y * y ~= (0::real)" |
|
202 |
by simp |
|
203 |
||
204 |
lemma realpow_two_sum_zero_iff [simp]: |
|
205 |
"(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)" |
|
206 |
by (rule sum_power2_eq_zero_iff) |
|
207 |
||
208 |
lemma realpow_two_le_add_order [simp]: "(0::real) \<le> u ^ 2 + v ^ 2" |
|
209 |
by (rule sum_power2_ge_zero) |
|
210 |
||
211 |
lemma realpow_two_le_add_order2 [simp]: "(0::real) \<le> u ^ 2 + v ^ 2 + w ^ 2" |
|
212 |
by (intro add_nonneg_nonneg zero_le_power2) |
|
213 |
||
214 |
lemma real_sum_square_gt_zero: "x ~= 0 ==> (0::real) < x * x + y * y" |
|
215 |
by (simp add: sum_squares_gt_zero_iff) |
|
216 |
||
217 |
lemma real_sum_square_gt_zero2: "y ~= 0 ==> (0::real) < x * x + y * y" |
|
218 |
by (simp add: sum_squares_gt_zero_iff) |
|
219 |
||
220 |
lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))" |
|
221 |
by (rule_tac j = 0 in real_le_trans, auto) |
|
222 |
||
223 |
lemma realpow_square_minus_le [simp]: "-(u ^ 2) \<le> (x::real) ^ 2" |
|
224 |
by (auto simp add: power2_eq_square) |
|
225 |
||
226 |
(* The following theorem is by Benjamin Porter *) |
|
227 |
lemma real_sq_order: |
|
228 |
fixes x::real |
|
229 |
assumes xgt0: "0 \<le> x" and ygt0: "0 \<le> y" and sq: "x^2 \<le> y^2" |
|
230 |
shows "x \<le> y" |
|
231 |
proof - |
|
232 |
from sq have "x ^ Suc (Suc 0) \<le> y ^ Suc (Suc 0)" |
|
233 |
by (simp only: numeral_2_eq_2) |
|
234 |
thus "x \<le> y" using ygt0 |
|
235 |
by (rule power_le_imp_le_base) |
|
236 |
qed |
|
237 |
||
238 |
||
239 |
subsection {*Various Other Theorems*} |
|
240 |
||
14304 | 241 |
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
|
242 |
by auto |
14268
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
243 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
244 |
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
|
245 |
by auto |
14268
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
246 |
|
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
247 |
lemma real_mult_inverse_cancel: |
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
248 |
"[|(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
|
249 |
==> inverse x * y < inverse x1 * u" |
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
250 |
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
|
251 |
apply (auto simp add: real_mult_assoc [symmetric]) |
14334 | 252 |
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
|
253 |
apply (rule_tac c=x1 in mult_less_imp_less_right) |
14334 | 254 |
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
|
255 |
done |
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
256 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
257 |
lemma real_mult_inverse_cancel2: |
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
258 |
"[|(0::real) < x;0 < x1; x1 * y < x * u |] ==> y * inverse x < u * inverse x1" |
14334 | 259 |
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
|
260 |
done |
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
261 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
262 |
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
|
263 |
by simp |
14268
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
264 |
|
14348
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents:
14334
diff
changeset
|
265 |
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
|
266 |
by simp |
14268
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
267 |
|
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
268 |
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
|
269 |
by (case_tac "n", auto) |
14268
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents:
14265
diff
changeset
|
270 |
|
28906 | 271 |
subsection{* Float syntax *} |
272 |
||
273 |
syntax "_Float" :: "float_const \<Rightarrow> 'a" ("_") |
|
274 |
||
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
28906
diff
changeset
|
275 |
use "Tools/float_syntax.ML" |
28906 | 276 |
setup FloatSyntax.setup |
277 |
||
278 |
text{* Test: *} |
|
279 |
lemma "123.456 = -111.111 + 200 + 30 + 4 + 5/10 + 6/100 + (7/1000::real)" |
|
280 |
by simp |
|
281 |
||
7077
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff
changeset
|
282 |
end |