| author | paulson | 
| Thu, 19 Feb 2004 10:40:28 +0100 | |
| changeset 14395 | cc96cc06abf9 | 
| parent 14365 | 3d4df8c166ae | 
| child 14477 | cc61fd03e589 | 
| permissions | -rw-r--r-- | 
| 12196 | 1 | (* Title : NthRoot.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 4 | Description : Existence of nth root. Adapted from | |
| 5 | http://www.math.unl.edu/~webnotes | |
| 6 | *) | |
| 7 | ||
| 14324 | 8 | header{*Existence of Nth Root*}
 | 
| 9 | ||
| 10 | theory NthRoot = SEQ + HSeries: | |
| 11 | ||
| 12 | text{*Various lemmas needed for this result. We follow the proof
 | |
| 13 | given by John Lindsay Orr (jorr@math.unl.edu) in his Analysis | |
| 14 | Webnotes available on the www at http://www.math.unl.edu/~webnotes | |
| 15 | Lemmas about sequences of reals are used to reach the result.*} | |
| 16 | ||
| 17 | lemma lemma_nth_realpow_non_empty: | |
| 18 |      "[| (0::real) < a; 0 < n |] ==> \<exists>s. s : {x. x ^ n <= a & 0 < x}"
 | |
| 19 | apply (case_tac "1 <= a") | |
| 20 | apply (rule_tac x = "1" in exI) | |
| 14334 | 21 | apply (drule_tac [2] linorder_not_le [THEN iffD1]) | 
| 14324 | 22 | apply (drule_tac [2] less_not_refl2 [THEN not0_implies_Suc]) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 23 | apply (simp add: ); | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 24 | apply (force intro!: realpow_Suc_le_self simp del: realpow_Suc) | 
| 14324 | 25 | done | 
| 26 | ||
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 27 | text{*Used only just below*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 28 | lemma realpow_ge_self2: "[| (1::real) \<le> r; 0 < n |] ==> r \<le> r ^ n" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 29 | by (insert power_increasing [of 1 n r], simp) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 30 | |
| 14324 | 31 | lemma lemma_nth_realpow_isUb_ex: | 
| 32 | "[| (0::real) < a; 0 < n |] | |
| 33 |       ==> \<exists>u. isUb (UNIV::real set) {x. x ^ n <= a & 0 < x} u"
 | |
| 34 | apply (case_tac "1 <= a") | |
| 35 | apply (rule_tac x = "a" in exI) | |
| 14334 | 36 | apply (drule_tac [2] linorder_not_le [THEN iffD1]) | 
| 14324 | 37 | apply (rule_tac [2] x = "1" in exI) | 
| 38 | apply (rule_tac [!] setleI [THEN isUbI]) | |
| 39 | apply safe | |
| 40 | apply (simp_all (no_asm)) | |
| 41 | apply (rule_tac [!] ccontr) | |
| 14334 | 42 | apply (drule_tac [!] linorder_not_le [THEN iffD1]) | 
| 14324 | 43 | apply (drule realpow_ge_self2 , assumption) | 
| 44 | apply (drule_tac n = "n" in realpow_less) | |
| 45 | apply (assumption+) | |
| 46 | apply (drule real_le_trans , assumption) | |
| 47 | apply (drule_tac y = "y ^ n" in order_less_le_trans) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 48 | apply (assumption) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 49 | apply (simp); | 
| 14334 | 50 | apply (drule_tac n = "n" in zero_less_one [THEN realpow_less]) | 
| 14324 | 51 | apply auto | 
| 52 | done | |
| 53 | ||
| 54 | lemma nth_realpow_isLub_ex: | |
| 55 | "[| (0::real) < a; 0 < n |] | |
| 56 |       ==> \<exists>u. isLub (UNIV::real set) {x. x ^ n <= a & 0 < x} u"
 | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 57 | by (blast intro: lemma_nth_realpow_isUb_ex lemma_nth_realpow_non_empty reals_complete) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 58 | |
| 14324 | 59 | |
| 60 | subsection{*First Half -- Lemmas First*}
 | |
| 61 | ||
| 62 | lemma lemma_nth_realpow_seq: | |
| 63 |      "isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u  
 | |
| 64 |            ==> u + inverse(real (Suc k)) ~: {x. x ^ n <= a & 0 < x}"
 | |
| 65 | apply (safe , drule isLubD2 , blast) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 66 | apply (simp add: linorder_not_less [symmetric]) | 
| 14324 | 67 | done | 
| 68 | ||
| 69 | lemma lemma_nth_realpow_isLub_gt_zero: | |
| 70 |      "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u;  
 | |
| 71 | 0 < a; 0 < n |] ==> 0 < u" | |
| 72 | apply (drule lemma_nth_realpow_non_empty , auto) | |
| 73 | apply (drule_tac y = "s" in isLub_isUb [THEN isUbD]) | |
| 74 | apply (auto intro: order_less_le_trans) | |
| 75 | done | |
| 76 | ||
| 77 | lemma lemma_nth_realpow_isLub_ge: | |
| 78 |      "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u;  
 | |
| 79 | 0 < a; 0 < n |] ==> ALL k. a <= (u + inverse(real (Suc k))) ^ n" | |
| 80 | apply (safe) | |
| 81 | apply (frule lemma_nth_realpow_seq , safe) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 82 | apply (auto elim: order_less_asym simp add: linorder_not_less [symmetric]) | 
| 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 83 | apply (simp add: linorder_not_less) | 
| 14324 | 84 | apply (rule order_less_trans [of _ 0]) | 
| 14325 | 85 | apply (auto intro: lemma_nth_realpow_isLub_gt_zero) | 
| 14324 | 86 | done | 
| 87 | ||
| 88 | text{*First result we want*}
 | |
| 89 | lemma realpow_nth_ge: | |
| 90 | "[| (0::real) < a; 0 < n; | |
| 91 | isLub (UNIV::real set) | |
| 92 |      {x. x ^ n <= a & 0 < x} u |] ==> a <= u ^ n"
 | |
| 93 | apply (frule lemma_nth_realpow_isLub_ge , safe) | |
| 94 | apply (rule LIMSEQ_inverse_real_of_nat_add [THEN LIMSEQ_pow, THEN LIMSEQ_le_const]) | |
| 14334 | 95 | apply (auto simp add: real_of_nat_def) | 
| 14324 | 96 | done | 
| 97 | ||
| 98 | subsection{*Second Half*}
 | |
| 99 | ||
| 100 | lemma less_isLub_not_isUb: | |
| 101 | "[| isLub (UNIV::real set) S u; x < u |] | |
| 102 | ==> ~ isUb (UNIV::real set) S x" | |
| 103 | apply (safe) | |
| 104 | apply (drule isLub_le_isUb) | |
| 105 | apply assumption | |
| 106 | apply (drule order_less_le_trans) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 107 | apply (auto) | 
| 14324 | 108 | done | 
| 109 | ||
| 110 | lemma not_isUb_less_ex: | |
| 111 | "~ isUb (UNIV::real set) S u ==> \<exists>x \<in> S. u < x" | |
| 112 | apply (rule ccontr , erule swap) | |
| 113 | apply (rule setleI [THEN isUbI]) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 114 | apply (auto simp add: linorder_not_less [symmetric]) | 
| 14324 | 115 | done | 
| 116 | ||
| 14325 | 117 | lemma real_mult_less_self: "0 < r ==> r * (1 + -inverse(real (Suc n))) < r" | 
| 14334 | 118 | apply (simp (no_asm) add: right_distrib) | 
| 119 | apply (rule add_less_cancel_left [of "-r", THEN iffD1]) | |
| 120 | apply (auto intro: mult_pos | |
| 121 | simp add: add_assoc [symmetric] neg_less_0_iff_less) | |
| 14325 | 122 | done | 
| 123 | ||
| 124 | lemma real_mult_add_one_minus_ge_zero: | |
| 125 | "0 < r ==> 0 <= r*(1 + -inverse(real (Suc n)))" | |
| 14355 
67e2e96bfe36
Told linear arithmetic package about injections "real" from nat/int into real.
 nipkow parents: 
14348diff
changeset | 126 | apply (simp add: zero_le_mult_iff real_of_nat_inverse_le_iff) | 
| 14325 | 127 | done | 
| 128 | ||
| 14324 | 129 | lemma lemma_nth_realpow_isLub_le: | 
| 130 |      "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u;  
 | |
| 14325 | 131 | 0 < a; 0 < n |] ==> ALL k. (u*(1 + -inverse(real (Suc k)))) ^ n <= a" | 
| 14324 | 132 | apply (safe) | 
| 133 | apply (frule less_isLub_not_isUb [THEN not_isUb_less_ex]) | |
| 134 | apply (rule_tac n = "k" in real_mult_less_self) | |
| 135 | apply (blast intro: lemma_nth_realpow_isLub_gt_zero) | |
| 136 | apply (safe) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 137 | apply (drule_tac n = "k" in | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 138 | lemma_nth_realpow_isLub_gt_zero [THEN real_mult_add_one_minus_ge_zero]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 139 | apply assumption+ | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 140 | apply (blast intro: order_trans order_less_imp_le power_mono) | 
| 14324 | 141 | done | 
| 142 | ||
| 143 | text{*Second result we want*}
 | |
| 144 | lemma realpow_nth_le: | |
| 145 | "[| (0::real) < a; 0 < n; | |
| 146 | isLub (UNIV::real set) | |
| 147 |      {x. x ^ n <= a & 0 < x} u |] ==> u ^ n <= a"
 | |
| 148 | apply (frule lemma_nth_realpow_isLub_le , safe) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 149 | apply (rule LIMSEQ_inverse_real_of_nat_add_minus_mult | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 150 | [THEN LIMSEQ_pow, THEN LIMSEQ_le_const2]) | 
| 14334 | 151 | apply (auto simp add: real_of_nat_def) | 
| 14324 | 152 | done | 
| 153 | ||
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 154 | text{*The theorem at last!*}
 | 
| 14324 | 155 | lemma realpow_nth: "[| (0::real) < a; 0 < n |] ==> \<exists>r. r ^ n = a" | 
| 156 | apply (frule nth_realpow_isLub_ex , auto) | |
| 157 | apply (auto intro: realpow_nth_le realpow_nth_ge real_le_anti_sym) | |
| 158 | done | |
| 159 | ||
| 160 | (* positive only *) | |
| 161 | lemma realpow_pos_nth: "[| (0::real) < a; 0 < n |] ==> \<exists>r. 0 < r & r ^ n = a" | |
| 162 | apply (frule nth_realpow_isLub_ex , auto) | |
| 163 | apply (auto intro: realpow_nth_le realpow_nth_ge real_le_anti_sym lemma_nth_realpow_isLub_gt_zero) | |
| 164 | done | |
| 165 | ||
| 166 | lemma realpow_pos_nth2: "(0::real) < a ==> \<exists>r. 0 < r & r ^ Suc n = a" | |
| 167 | apply (blast intro: realpow_pos_nth) | |
| 168 | done | |
| 169 | ||
| 170 | (* uniqueness of nth positive root *) | |
| 171 | lemma realpow_pos_nth_unique: | |
| 172 | "[| (0::real) < a; 0 < n |] ==> EX! r. 0 < r & r ^ n = a" | |
| 173 | apply (auto intro!: realpow_pos_nth) | |
| 174 | apply (cut_tac x = "r" and y = "y" in linorder_less_linear) | |
| 175 | apply auto | |
| 176 | apply (drule_tac x = "r" in realpow_less) | |
| 177 | apply (drule_tac [4] x = "y" in realpow_less) | |
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 178 | apply (auto) | 
| 14324 | 179 | done | 
| 180 | ||
| 181 | ML | |
| 182 | {*
 | |
| 183 | val nth_realpow_isLub_ex = thm"nth_realpow_isLub_ex"; | |
| 184 | val realpow_nth_ge = thm"realpow_nth_ge"; | |
| 185 | val less_isLub_not_isUb = thm"less_isLub_not_isUb"; | |
| 186 | val not_isUb_less_ex = thm"not_isUb_less_ex"; | |
| 187 | val realpow_nth_le = thm"realpow_nth_le"; | |
| 188 | val realpow_nth = thm"realpow_nth"; | |
| 189 | val realpow_pos_nth = thm"realpow_pos_nth"; | |
| 190 | val realpow_pos_nth2 = thm"realpow_pos_nth2"; | |
| 191 | val realpow_pos_nth_unique = thm"realpow_pos_nth_unique"; | |
| 192 | *} | |
| 193 | ||
| 194 | end |