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