| author | paulson | 
| Mon, 23 Jan 2006 11:37:53 +0100 | |
| changeset 18751 | 38dc4ff2a32b | 
| parent 15251 | bb6f072c8d10 | 
| child 19279 | 48b527d0331b | 
| permissions | -rw-r--r-- | 
| 9435 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 wenzelm parents: 
9013diff
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 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 5 | Description : Natural powers theory | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 6 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 7 | *) | 
| 
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: 
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 | |
| 10309 | 15 | instance real :: power .. | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 16 | |
| 8856 | 17 | primrec (realpow) | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11701diff
changeset | 18 | realpow_0: "r ^ 0 = 1" | 
| 9435 
c3a13a7d4424
lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
 wenzelm parents: 
9013diff
changeset | 19 | realpow_Suc: "r ^ (Suc n) = (r::real) * (r ^ n)" | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 20 | |
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 21 | |
| 15003 | 22 | instance real :: recpower | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 23 | proof | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 24 | fix z :: real | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 25 | fix n :: nat | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 26 | show "z^0 = 1" by simp | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 27 | show "z^(Suc n) = z * (z^n)" by simp | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 28 | qed | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 29 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 30 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 31 | lemma realpow_not_zero: "r \<noteq> (0::real) ==> r ^ n \<noteq> 0" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 32 | by (rule field_power_not_zero) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 33 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 34 | lemma realpow_zero_zero: "r ^ n = (0::real) ==> r = 0" | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 35 | by simp | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 36 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 37 | lemma realpow_two: "(r::real)^ (Suc (Suc 0)) = r * r" | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 38 | by simp | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 39 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 40 | text{*Legacy: weaker version of the theorem @{text power_strict_mono},
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 41 | used 6 times in NthRoot and Transcendental*} | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 42 | lemma realpow_less: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 43 | "[|(0::real) < x; x < y; 0 < n|] ==> x ^ n < y ^ n" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 44 | apply (rule power_strict_mono, auto) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 45 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 46 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 47 | lemma realpow_two_le [simp]: "(0::real) \<le> r^ Suc (Suc 0)" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 48 | by (simp add: real_le_square) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 49 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 50 | lemma abs_realpow_two [simp]: "abs((x::real)^Suc (Suc 0)) = x^Suc (Suc 0)" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 51 | by (simp add: abs_mult) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 52 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 53 | lemma realpow_two_abs [simp]: "abs(x::real)^Suc (Suc 0) = x^Suc (Suc 0)" | 
| 15229 | 54 | by (simp add: power_abs [symmetric] del: realpow_Suc) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 55 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 56 | lemma two_realpow_ge_one [simp]: "(1::real) \<le> 2 ^ n" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 57 | by (insert power_increasing [of 0 n "2::real"], simp) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 58 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 59 | lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n" | 
| 15251 | 60 | apply (induct "n") | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 61 | apply (auto simp add: real_of_nat_Suc) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14352diff
changeset | 62 | apply (subst mult_2) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 63 | apply (rule real_add_less_le_mono) | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 64 | 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 | 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_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 | 68 | 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 | 69 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 70 | text{*Used ONCE in Transcendental*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 71 | lemma realpow_Suc_less_one: "[| 0 < r; r < (1::real) |] ==> r ^ Suc n < 1" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 72 | by (insert power_strict_decreasing [of 0 "Suc n" r], simp) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 73 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 74 | text{*Used ONCE in Lim.ML*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 75 | lemma realpow_minus_mult [rule_format]: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 76 | "0 < n --> (x::real) ^ (n - 1) * x = x ^ n" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 77 | 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 | 78 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 79 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 80 | lemma realpow_two_mult_inverse [simp]: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 81 | "r \<noteq> 0 ==> r * inverse r ^Suc (Suc 0) = inverse (r::real)" | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 82 | by (simp add: realpow_two real_mult_assoc [symmetric]) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 83 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 84 | 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 | 85 | by simp | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 86 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 87 | lemma realpow_two_diff: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 88 | "(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 | 89 | apply (unfold real_diff_def) | 
| 14334 | 90 | apply (simp add: right_distrib left_distrib mult_ac) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 91 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 92 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 93 | lemma realpow_two_disj: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 94 | "((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 | 95 | 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: 
12018diff
changeset | 96 | apply (auto simp del: realpow_Suc) | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 97 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 98 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 99 | lemma realpow_real_of_nat: "real (m::nat) ^ n = real (m ^ n)" | 
| 15251 | 100 | apply (induct "n") | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 101 | 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 | 102 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 103 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 104 | lemma realpow_real_of_nat_two_pos [simp] : "0 < real (Suc (Suc 0) ^ n)" | 
| 15251 | 105 | apply (induct "n") | 
| 14334 | 106 | 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 | 107 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 108 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 109 | lemma realpow_increasing: | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 110 | "[|(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 | 111 | 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 | 112 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 113 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 114 | lemma zero_less_realpow_abs_iff [simp]: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 115 | "(0 < (abs x)^n) = (x \<noteq> (0::real) | n=0)" | 
| 15251 | 116 | apply (induct "n") | 
| 14334 | 117 | apply (auto simp add: zero_less_mult_iff) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 118 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 119 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 120 | lemma zero_le_realpow_abs [simp]: "(0::real) \<le> (abs x)^n" | 
| 15251 | 121 | apply (induct "n") | 
| 14334 | 122 | apply (auto simp add: zero_le_mult_iff) | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 123 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 124 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 125 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 126 | 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 | 127 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 128 | lemma real_of_int_power: "real (x::int) ^ n = real (x ^ n)" | 
| 15251 | 129 | apply (induct "n") | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14352diff
changeset | 130 | 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 | 131 | done | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 132 | 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 | 133 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 134 | lemma power_real_number_of: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 135 | "(number_of v :: real) ^ n = real ((number_of v :: int) ^ n)" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14352diff
changeset | 136 | 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 | 137 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 138 | 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 | 139 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 140 | |
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 141 | subsection{*Various Other Theorems*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 142 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 143 | text{*Used several times in Hyperreal/Transcendental.ML*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 144 | lemma real_sum_squares_cancel_a: "x * x = -(y * y) ==> x = (0::real) & y=0" | 
| 15085 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15003diff
changeset | 145 | apply (auto dest: real_sum_squares_cancel simp add: real_add_eq_0_iff [symmetric]) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15003diff
changeset | 146 | apply (auto dest: real_sum_squares_cancel simp add: add_commute) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15003diff
changeset | 147 | done | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 148 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 149 | lemma real_squared_diff_one_factored: "x*x - (1::real) = (x + 1)*(x - 1)" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 150 | by (auto simp add: left_distrib right_distrib real_diff_def) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 151 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 152 | 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 | 153 | apply auto | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 154 | apply (drule right_minus_eq [THEN iffD2]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 155 | 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 | 156 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 157 | |
| 14304 | 158 | 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 | 159 | by auto | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 160 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 161 | 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 | 162 | by auto | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 163 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 164 | lemma real_mult_inverse_cancel: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 165 | "[|(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 | 166 | ==> inverse x * y < inverse x1 * u" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 167 | 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 | 168 | apply (auto simp add: real_mult_assoc [symmetric]) | 
| 14334 | 169 | 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 | 170 | apply (rule_tac c=x1 in mult_less_imp_less_right) | 
| 14334 | 171 | apply (auto simp add: mult_ac) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 172 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 173 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 174 | text{*Used once: in Hyperreal/Transcendental.ML*}
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 175 | lemma real_mult_inverse_cancel2: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 176 | "[|(0::real) < x;0 < x1; x1 * y < x * u |] ==> y * inverse x < u * inverse x1" | 
| 14334 | 177 | 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 | 178 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 179 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 180 | lemma inverse_real_of_nat_gt_zero [simp]: "0 < inverse (real (Suc n))" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 181 | by auto | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 182 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 183 | lemma inverse_real_of_nat_ge_zero [simp]: "0 \<le> inverse (real (Suc n))" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 184 | by auto | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 185 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 186 | lemma real_sum_squares_not_zero: "x ~= 0 ==> x * x + y * y ~= (0::real)" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 187 | by (blast dest!: real_sum_squares_cancel) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 188 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 189 | lemma real_sum_squares_not_zero2: "y ~= 0 ==> x * x + y * y ~= (0::real)" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 190 | by (blast dest!: real_sum_squares_cancel2) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 191 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 192 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 193 | subsection {*Various Other Theorems*}
 | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 194 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 195 | lemma realpow_divide: | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 196 | "(x/y) ^ n = ((x::real) ^ n/ y ^ n)" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 197 | apply (unfold real_divide_def) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 198 | apply (auto simp add: power_mult_distrib power_inverse) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 199 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 200 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 201 | lemma realpow_two_sum_zero_iff [simp]: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 202 | "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 203 | apply (auto intro: real_sum_squares_cancel real_sum_squares_cancel2 | 
| 14352 | 204 | simp add: power2_eq_square) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 205 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 206 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 207 | lemma realpow_two_le_add_order [simp]: "(0::real) \<le> u ^ 2 + v ^ 2" | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 208 | apply (rule real_le_add_order) | 
| 14352 | 209 | apply (auto simp add: power2_eq_square) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 210 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 211 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 212 | lemma realpow_two_le_add_order2 [simp]: "(0::real) \<le> u ^ 2 + v ^ 2 + w ^ 2" | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 213 | apply (rule real_le_add_order)+ | 
| 14352 | 214 | apply (auto simp add: power2_eq_square) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 215 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 216 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 217 | lemma real_sum_square_gt_zero: "x ~= 0 ==> (0::real) < x * x + y * y" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 218 | apply (cut_tac x = x and y = y in real_mult_self_sum_ge_zero) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 219 | apply (drule real_le_imp_less_or_eq) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 220 | apply (drule_tac y = y in real_sum_squares_not_zero, auto) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 221 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 222 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 223 | lemma real_sum_square_gt_zero2: "y ~= 0 ==> (0::real) < x * x + y * y" | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 224 | apply (rule real_add_commute [THEN subst]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 225 | apply (erule real_sum_square_gt_zero) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 226 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 227 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 228 | lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 229 | by (rule_tac j = 0 in real_le_trans, auto) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 230 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 231 | lemma realpow_square_minus_le [simp]: "-(u ^ 2) \<le> (x::real) ^ 2" | 
| 14352 | 232 | by (auto simp add: power2_eq_square) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 233 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 234 | 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 | 235 | by (case_tac "n", auto) | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 236 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 237 | lemma real_num_zero_less_two_pow [simp]: "0 < (2::real) ^ (4*d)" | 
| 15251 | 238 | apply (induct "d") | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 239 | apply (auto simp add: realpow_num_eq_if) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 240 | done | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 241 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 242 | lemma lemma_realpow_num_two_mono: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 243 | "x * (4::real) < y ==> x * (2 ^ 8) < y * (2 ^ 6)" | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 244 | apply (subgoal_tac " (2::real) ^ 8 = 4 * (2 ^ 6) ") | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 245 | apply (simp (no_asm_simp) add: real_mult_assoc [symmetric]) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 246 | apply (auto simp add: realpow_num_eq_if) | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 247 | done | 
| 
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 | |
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 250 | ML | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 251 | {*
 | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 252 | val realpow_0 = thm "realpow_0"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 253 | val realpow_Suc = thm "realpow_Suc"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 254 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 255 | val realpow_not_zero = thm "realpow_not_zero"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 256 | val realpow_zero_zero = thm "realpow_zero_zero"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 257 | val realpow_two = thm "realpow_two"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 258 | val realpow_less = thm "realpow_less"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 259 | val realpow_two_le = thm "realpow_two_le"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 260 | val abs_realpow_two = thm "abs_realpow_two"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 261 | val realpow_two_abs = thm "realpow_two_abs"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 262 | val two_realpow_ge_one = thm "two_realpow_ge_one"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 263 | val two_realpow_gt = thm "two_realpow_gt"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 264 | val realpow_Suc_le_self = thm "realpow_Suc_le_self"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 265 | val realpow_Suc_less_one = thm "realpow_Suc_less_one"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 266 | val realpow_minus_mult = thm "realpow_minus_mult"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 267 | val realpow_two_mult_inverse = thm "realpow_two_mult_inverse"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 268 | val realpow_two_minus = thm "realpow_two_minus"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 269 | val realpow_two_disj = thm "realpow_two_disj"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 270 | val realpow_real_of_nat = thm "realpow_real_of_nat"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 271 | val realpow_real_of_nat_two_pos = thm "realpow_real_of_nat_two_pos"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 272 | val realpow_increasing = thm "realpow_increasing"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 273 | val zero_less_realpow_abs_iff = thm "zero_less_realpow_abs_iff"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 274 | val zero_le_realpow_abs = thm "zero_le_realpow_abs"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 275 | val real_of_int_power = thm "real_of_int_power"; | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 276 | val power_real_number_of = thm "power_real_number_of"; | 
| 14268 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 277 | val real_sum_squares_cancel_a = thm "real_sum_squares_cancel_a"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 278 | val real_mult_inverse_cancel2 = thm "real_mult_inverse_cancel2"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 279 | val real_squared_diff_one_factored = thm "real_squared_diff_one_factored"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 280 | val real_mult_is_one = thm "real_mult_is_one"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 281 | val real_le_add_half_cancel = thm "real_le_add_half_cancel"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 282 | val real_minus_half_eq = thm "real_minus_half_eq"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 283 | val real_mult_inverse_cancel = thm "real_mult_inverse_cancel"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 284 | val real_mult_inverse_cancel2 = thm "real_mult_inverse_cancel2"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 285 | val inverse_real_of_nat_gt_zero = thm "inverse_real_of_nat_gt_zero"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 286 | val inverse_real_of_nat_ge_zero = thm "inverse_real_of_nat_ge_zero"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 287 | val real_sum_squares_not_zero = thm "real_sum_squares_not_zero"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 288 | val real_sum_squares_not_zero2 = thm "real_sum_squares_not_zero2"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 289 | |
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 290 | val realpow_divide = thm "realpow_divide"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 291 | val realpow_two_sum_zero_iff = thm "realpow_two_sum_zero_iff"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 292 | val realpow_two_le_add_order = thm "realpow_two_le_add_order"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 293 | val realpow_two_le_add_order2 = thm "realpow_two_le_add_order2"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 294 | val real_sum_square_gt_zero = thm "real_sum_square_gt_zero"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 295 | val real_sum_square_gt_zero2 = thm "real_sum_square_gt_zero2"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 296 | val real_minus_mult_self_le = thm "real_minus_mult_self_le"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 297 | val realpow_square_minus_le = thm "realpow_square_minus_le"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 298 | val realpow_num_eq_if = thm "realpow_num_eq_if"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 299 | val real_num_zero_less_two_pow = thm "real_num_zero_less_two_pow"; | 
| 
5cf13e80be0e
Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
 paulson parents: 
14265diff
changeset | 300 | val lemma_realpow_num_two_mono = thm "lemma_realpow_num_two_mono"; | 
| 14265 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 301 | *} | 
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 302 | |
| 
95b42e69436c
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
 paulson parents: 
12018diff
changeset | 303 | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 304 | end |