| author | wenzelm | 
| Tue, 08 Nov 2005 10:43:08 +0100 | |
| changeset 18117 | 61a430a67d7c | 
| parent 16775 | c1b87ef4a1c3 | 
| child 18585 | 5d379fe2eb74 | 
| permissions | -rw-r--r-- | 
| 12196 | 1 | (* Title : NthRoot.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 14477 | 4 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | 
| 12196 | 5 | *) | 
| 6 | ||
| 14324 | 7 | header{*Existence of Nth Root*}
 | 
| 8 | ||
| 15131 | 9 | theory NthRoot | 
| 15140 | 10 | imports SEQ HSeries | 
| 15131 | 11 | begin | 
| 14324 | 12 | |
| 14767 | 13 | text {*
 | 
| 14 | Various lemmas needed for this result. We follow the proof given by | |
| 15 |   John Lindsay Orr (\texttt{jorr@math.unl.edu}) in his Analysis
 | |
| 16 |   Webnotes available at \url{http://www.math.unl.edu/~webnotes}.
 | |
| 17 | ||
| 18 | Lemmas about sequences of reals are used to reach the result. | |
| 19 | *} | |
| 14324 | 20 | |
| 21 | lemma lemma_nth_realpow_non_empty: | |
| 22 |      "[| (0::real) < a; 0 < n |] ==> \<exists>s. s : {x. x ^ n <= a & 0 < x}"
 | |
| 23 | apply (case_tac "1 <= a") | |
| 14477 | 24 | apply (rule_tac x = 1 in exI) | 
| 14334 | 25 | apply (drule_tac [2] linorder_not_le [THEN iffD1]) | 
| 14477 | 26 | apply (drule_tac [2] less_not_refl2 [THEN not0_implies_Suc], simp) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 27 | apply (force intro!: realpow_Suc_le_self simp del: realpow_Suc) | 
| 14324 | 28 | done | 
| 29 | ||
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 30 | text{*Used only just below*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 31 | 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 | 32 | by (insert power_increasing [of 1 n r], simp) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 33 | |
| 14324 | 34 | lemma lemma_nth_realpow_isUb_ex: | 
| 35 | "[| (0::real) < a; 0 < n |] | |
| 36 |       ==> \<exists>u. isUb (UNIV::real set) {x. x ^ n <= a & 0 < x} u"
 | |
| 37 | apply (case_tac "1 <= a") | |
| 14477 | 38 | apply (rule_tac x = a in exI) | 
| 14334 | 39 | apply (drule_tac [2] linorder_not_le [THEN iffD1]) | 
| 14477 | 40 | apply (rule_tac [2] x = 1 in exI) | 
| 41 | apply (rule_tac [!] setleI [THEN isUbI], safe) | |
| 14324 | 42 | apply (simp_all (no_asm)) | 
| 43 | apply (rule_tac [!] ccontr) | |
| 14334 | 44 | apply (drule_tac [!] linorder_not_le [THEN iffD1]) | 
| 14477 | 45 | apply (drule realpow_ge_self2, assumption) | 
| 46 | apply (drule_tac n = n in realpow_less) | |
| 14324 | 47 | apply (assumption+) | 
| 14477 | 48 | apply (drule real_le_trans, assumption) | 
| 49 | apply (drule_tac y = "y ^ n" in order_less_le_trans, assumption, simp) | |
| 50 | apply (drule_tac n = n in zero_less_one [THEN realpow_less], auto) | |
| 14324 | 51 | done | 
| 52 | ||
| 53 | lemma nth_realpow_isLub_ex: | |
| 54 | "[| (0::real) < a; 0 < n |] | |
| 55 |       ==> \<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 | 56 | 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 | 57 | |
| 14324 | 58 | |
| 59 | subsection{*First Half -- Lemmas First*}
 | |
| 60 | ||
| 61 | lemma lemma_nth_realpow_seq: | |
| 62 |      "isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u  
 | |
| 63 |            ==> u + inverse(real (Suc k)) ~: {x. x ^ n <= a & 0 < x}"
 | |
| 14477 | 64 | apply (safe, drule isLubD2, blast) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 65 | apply (simp add: linorder_not_less [symmetric]) | 
| 14324 | 66 | done | 
| 67 | ||
| 68 | lemma lemma_nth_realpow_isLub_gt_zero: | |
| 69 |      "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u;  
 | |
| 70 | 0 < a; 0 < n |] ==> 0 < u" | |
| 14477 | 71 | apply (drule lemma_nth_realpow_non_empty, auto) | 
| 72 | apply (drule_tac y = s in isLub_isUb [THEN isUbD]) | |
| 14324 | 73 | apply (auto intro: order_less_le_trans) | 
| 74 | done | |
| 75 | ||
| 76 | lemma lemma_nth_realpow_isLub_ge: | |
| 77 |      "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u;  
 | |
| 78 | 0 < a; 0 < n |] ==> ALL k. a <= (u + inverse(real (Suc k))) ^ n" | |
| 14477 | 79 | apply safe | 
| 80 | apply (frule lemma_nth_realpow_seq, safe) | |
| 15085 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
14767diff
changeset | 81 | apply (auto elim: order_less_asym simp add: linorder_not_less [symmetric] | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
14767diff
changeset | 82 |             iff: real_0_less_add_iff) --{*legacy iff rule!*}
 | 
| 14365 
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"
 | |
| 14477 | 93 | apply (frule lemma_nth_realpow_isLub_ge, safe) | 
| 14324 | 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" | |
| 14477 | 103 | apply safe | 
| 104 | apply (drule isLub_le_isUb, assumption) | |
| 105 | apply (drule order_less_le_trans, auto) | |
| 14324 | 106 | done | 
| 107 | ||
| 108 | lemma not_isUb_less_ex: | |
| 109 | "~ isUb (UNIV::real set) S u ==> \<exists>x \<in> S. u < x" | |
| 14477 | 110 | apply (rule ccontr, erule swap) | 
| 14324 | 111 | apply (rule setleI [THEN isUbI]) | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14355diff
changeset | 112 | apply (auto simp add: linorder_not_less [symmetric]) | 
| 14324 | 113 | done | 
| 114 | ||
| 14325 | 115 | lemma real_mult_less_self: "0 < r ==> r * (1 + -inverse(real (Suc n))) < r" | 
| 14334 | 116 | apply (simp (no_asm) add: right_distrib) | 
| 117 | apply (rule add_less_cancel_left [of "-r", THEN iffD1]) | |
| 16775 
c1b87ef4a1c3
added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
 avigad parents: 
15140diff
changeset | 118 | apply (auto intro: mult_pos_pos | 
| 14334 | 119 | simp add: add_assoc [symmetric] neg_less_0_iff_less) | 
| 14325 | 120 | done | 
| 121 | ||
| 122 | lemma real_mult_add_one_minus_ge_zero: | |
| 123 | "0 < r ==> 0 <= r*(1 + -inverse(real (Suc n)))" | |
| 15085 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
14767diff
changeset | 124 | by (simp add: zero_le_mult_iff real_of_nat_inverse_le_iff real_0_le_add_iff) | 
| 14325 | 125 | |
| 14324 | 126 | lemma lemma_nth_realpow_isLub_le: | 
| 127 |      "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u;  
 | |
| 14325 | 128 | 0 < a; 0 < n |] ==> ALL k. (u*(1 + -inverse(real (Suc k)))) ^ n <= a" | 
| 14477 | 129 | apply safe | 
| 14324 | 130 | apply (frule less_isLub_not_isUb [THEN not_isUb_less_ex]) | 
| 14477 | 131 | apply (rule_tac n = k in real_mult_less_self) | 
| 132 | apply (blast intro: lemma_nth_realpow_isLub_gt_zero, safe) | |
| 133 | apply (drule_tac n = k in | |
| 134 | lemma_nth_realpow_isLub_gt_zero [THEN real_mult_add_one_minus_ge_zero], assumption+) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 135 | apply (blast intro: order_trans order_less_imp_le power_mono) | 
| 14324 | 136 | done | 
| 137 | ||
| 138 | text{*Second result we want*}
 | |
| 139 | lemma realpow_nth_le: | |
| 140 | "[| (0::real) < a; 0 < n; | |
| 141 | isLub (UNIV::real set) | |
| 142 |      {x. x ^ n <= a & 0 < x} u |] ==> u ^ n <= a"
 | |
| 14477 | 143 | apply (frule lemma_nth_realpow_isLub_le, safe) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 144 | 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 | 145 | [THEN LIMSEQ_pow, THEN LIMSEQ_le_const2]) | 
| 14334 | 146 | apply (auto simp add: real_of_nat_def) | 
| 14324 | 147 | done | 
| 148 | ||
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
14334diff
changeset | 149 | text{*The theorem at last!*}
 | 
| 14324 | 150 | lemma realpow_nth: "[| (0::real) < a; 0 < n |] ==> \<exists>r. r ^ n = a" | 
| 14477 | 151 | apply (frule nth_realpow_isLub_ex, auto) | 
| 152 | apply (auto intro: realpow_nth_le realpow_nth_ge order_antisym) | |
| 14324 | 153 | done | 
| 154 | ||
| 155 | (* positive only *) | |
| 156 | lemma realpow_pos_nth: "[| (0::real) < a; 0 < n |] ==> \<exists>r. 0 < r & r ^ n = a" | |
| 14477 | 157 | apply (frule nth_realpow_isLub_ex, auto) | 
| 158 | apply (auto intro: realpow_nth_le realpow_nth_ge order_antisym lemma_nth_realpow_isLub_gt_zero) | |
| 14324 | 159 | done | 
| 160 | ||
| 161 | lemma realpow_pos_nth2: "(0::real) < a ==> \<exists>r. 0 < r & r ^ Suc n = a" | |
| 14477 | 162 | by (blast intro: realpow_pos_nth) | 
| 14324 | 163 | |
| 164 | (* uniqueness of nth positive root *) | |
| 165 | lemma realpow_pos_nth_unique: | |
| 166 | "[| (0::real) < a; 0 < n |] ==> EX! r. 0 < r & r ^ n = a" | |
| 167 | apply (auto intro!: realpow_pos_nth) | |
| 14477 | 168 | apply (cut_tac x = r and y = y in linorder_less_linear, auto) | 
| 169 | apply (drule_tac x = r in realpow_less) | |
| 170 | apply (drule_tac [4] x = y in realpow_less, auto) | |
| 14324 | 171 | done | 
| 172 | ||
| 173 | ML | |
| 174 | {*
 | |
| 175 | val nth_realpow_isLub_ex = thm"nth_realpow_isLub_ex"; | |
| 176 | val realpow_nth_ge = thm"realpow_nth_ge"; | |
| 177 | val less_isLub_not_isUb = thm"less_isLub_not_isUb"; | |
| 178 | val not_isUb_less_ex = thm"not_isUb_less_ex"; | |
| 179 | val realpow_nth_le = thm"realpow_nth_le"; | |
| 180 | val realpow_nth = thm"realpow_nth"; | |
| 181 | val realpow_pos_nth = thm"realpow_pos_nth"; | |
| 182 | val realpow_pos_nth2 = thm"realpow_pos_nth2"; | |
| 183 | val realpow_pos_nth_unique = thm"realpow_pos_nth_unique"; | |
| 184 | *} | |
| 185 | ||
| 186 | end |