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