src/HOL/Hyperreal/NthRoot.thy
author haftmann
Fri, 09 Mar 2007 08:45:50 +0100
changeset 22422 ee19cdb07528
parent 21865 55cc354fd2d9
child 22443 346729a55460
permissions -rw-r--r--
stepping towards uniform lattice theory development in HOL
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     1
(*  Title       : NthRoot.thy
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
     4
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     5
*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     6
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
     7
header{*Existence of Nth Root*}
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
     8
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15085
diff changeset
     9
theory NthRoot
21865
55cc354fd2d9 moved several theorems; rearranged theory dependencies
huffman
parents: 21404
diff changeset
    10
imports SEQ Parity
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15085
diff changeset
    11
begin
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    12
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    13
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20898
diff changeset
    14
  root :: "[nat, real] \<Rightarrow> real" where
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    15
  "root n x = (THE u. (0 < x \<longrightarrow> 0 < u) \<and> (u ^ n = x))"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    16
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20898
diff changeset
    17
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20898
diff changeset
    18
  sqrt :: "real \<Rightarrow> real" where
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    19
  "sqrt x = root 2 x"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    20
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    21
14767
d2b071e65e4c tuned document;
wenzelm
parents: 14477
diff changeset
    22
text {*
d2b071e65e4c tuned document;
wenzelm
parents: 14477
diff changeset
    23
  Various lemmas needed for this result. We follow the proof given by
d2b071e65e4c tuned document;
wenzelm
parents: 14477
diff changeset
    24
  John Lindsay Orr (\texttt{jorr@math.unl.edu}) in his Analysis
d2b071e65e4c tuned document;
wenzelm
parents: 14477
diff changeset
    25
  Webnotes available at \url{http://www.math.unl.edu/~webnotes}.
d2b071e65e4c tuned document;
wenzelm
parents: 14477
diff changeset
    26
d2b071e65e4c tuned document;
wenzelm
parents: 14477
diff changeset
    27
  Lemmas about sequences of reals are used to reach the result.
d2b071e65e4c tuned document;
wenzelm
parents: 14477
diff changeset
    28
*}
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    29
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    30
lemma lemma_nth_realpow_non_empty:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    31
     "[| (0::real) < a; 0 < n |] ==> \<exists>s. s : {x. x ^ n <= a & 0 < x}"
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    32
apply (case_tac "1 <= a")
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    33
apply (rule_tac x = 1 in exI)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14325
diff changeset
    34
apply (drule_tac [2] linorder_not_le [THEN iffD1])
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    35
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
    36
apply (force intro!: realpow_Suc_le_self simp del: realpow_Suc)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    37
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    38
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    39
text{*Used only just below*}
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    40
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
    41
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
    42
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    43
lemma lemma_nth_realpow_isUb_ex:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    44
     "[| (0::real) < a; 0 < n |]  
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    45
      ==> \<exists>u. isUb (UNIV::real set) {x. x ^ n <= a & 0 < x} u"
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    46
apply (case_tac "1 <= a")
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    47
apply (rule_tac x = a in exI)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14325
diff changeset
    48
apply (drule_tac [2] linorder_not_le [THEN iffD1])
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    49
apply (rule_tac [2] x = 1 in exI)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    50
apply (rule_tac [!] setleI [THEN isUbI], safe)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    51
apply (simp_all (no_asm))
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    52
apply (rule_tac [!] ccontr)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14325
diff changeset
    53
apply (drule_tac [!] linorder_not_le [THEN iffD1])
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    54
apply (drule realpow_ge_self2, assumption)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    55
apply (drule_tac n = n in realpow_less)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    56
apply (assumption+)
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    57
apply (drule real_le_trans, assumption)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    58
apply (drule_tac y = "y ^ n" in order_less_le_trans, assumption, simp) 
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    59
apply (drule_tac n = n in zero_less_one [THEN realpow_less], auto)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    60
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    61
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    62
lemma nth_realpow_isLub_ex:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    63
     "[| (0::real) < a; 0 < n |]  
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    64
      ==> \<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
    65
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
    66
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    67
 
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    68
subsection{*First Half -- Lemmas First*}
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    69
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    70
lemma lemma_nth_realpow_seq:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    71
     "isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u  
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    72
           ==> u + inverse(real (Suc k)) ~: {x. x ^ n <= a & 0 < x}"
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    73
apply (safe, drule isLubD2, blast)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14355
diff changeset
    74
apply (simp add: linorder_not_less [symmetric])
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    75
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    76
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    77
lemma lemma_nth_realpow_isLub_gt_zero:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    78
     "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u;  
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    79
         0 < a; 0 < n |] ==> 0 < u"
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    80
apply (drule lemma_nth_realpow_non_empty, auto)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    81
apply (drule_tac y = s in isLub_isUb [THEN isUbD])
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    82
apply (auto intro: order_less_le_trans)
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    83
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    84
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    85
lemma lemma_nth_realpow_isLub_ge:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    86
     "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u;  
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    87
         0 < a; 0 < n |] ==> ALL k. a <= (u + inverse(real (Suc k))) ^ n"
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    88
apply safe
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
    89
apply (frule lemma_nth_realpow_seq, safe)
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 14767
diff changeset
    90
apply (auto elim: order_less_asym simp add: linorder_not_less [symmetric]
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 14767
diff changeset
    91
            iff: real_0_less_add_iff) --{*legacy iff rule!*}
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14355
diff changeset
    92
apply (simp add: linorder_not_less)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    93
apply (rule order_less_trans [of _ 0])
14325
94ac3895822f removing real_of_posnat
paulson
parents: 14324
diff changeset
    94
apply (auto intro: lemma_nth_realpow_isLub_gt_zero)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    95
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    96
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    97
text{*First result we want*}
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    98
lemma realpow_nth_ge:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    99
     "[| (0::real) < a; 0 < n;  
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   100
     isLub (UNIV::real set)  
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   101
     {x. x ^ n <= a & 0 < x} u |] ==> a <= u ^ n"
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   102
apply (frule lemma_nth_realpow_isLub_ge, safe)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   103
apply (rule LIMSEQ_inverse_real_of_nat_add [THEN LIMSEQ_pow, THEN LIMSEQ_le_const])
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14325
diff changeset
   104
apply (auto simp add: real_of_nat_def)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   105
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   106
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   107
subsection{*Second Half*}
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   108
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   109
lemma less_isLub_not_isUb:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   110
     "[| isLub (UNIV::real set) S u; x < u |]  
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   111
           ==> ~ isUb (UNIV::real set) S x"
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   112
apply safe
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   113
apply (drule isLub_le_isUb, assumption)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   114
apply (drule order_less_le_trans, auto)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   115
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   116
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   117
lemma not_isUb_less_ex:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   118
     "~ isUb (UNIV::real set) S u ==> \<exists>x \<in> S. u < x"
18585
5d379fe2eb74 replaced swap by contrapos_np;
wenzelm
parents: 16775
diff changeset
   119
apply (rule ccontr, erule contrapos_np)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   120
apply (rule setleI [THEN isUbI])
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14355
diff changeset
   121
apply (auto simp add: linorder_not_less [symmetric])
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   122
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   123
14325
94ac3895822f removing real_of_posnat
paulson
parents: 14324
diff changeset
   124
lemma real_mult_less_self: "0 < r ==> r * (1 + -inverse(real (Suc n))) < r"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14325
diff changeset
   125
apply (simp (no_asm) add: right_distrib)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14325
diff changeset
   126
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: 15140
diff changeset
   127
apply (auto intro: mult_pos_pos
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14325
diff changeset
   128
            simp add: add_assoc [symmetric] neg_less_0_iff_less)
14325
94ac3895822f removing real_of_posnat
paulson
parents: 14324
diff changeset
   129
done
94ac3895822f removing real_of_posnat
paulson
parents: 14324
diff changeset
   130
94ac3895822f removing real_of_posnat
paulson
parents: 14324
diff changeset
   131
lemma real_mult_add_one_minus_ge_zero:
94ac3895822f removing real_of_posnat
paulson
parents: 14324
diff changeset
   132
     "0 < r ==>  0 <= r*(1 + -inverse(real (Suc n)))"
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 14767
diff changeset
   133
by (simp add: zero_le_mult_iff real_of_nat_inverse_le_iff real_0_le_add_iff)
14325
94ac3895822f removing real_of_posnat
paulson
parents: 14324
diff changeset
   134
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   135
lemma lemma_nth_realpow_isLub_le:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   136
     "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u;  
14325
94ac3895822f removing real_of_posnat
paulson
parents: 14324
diff changeset
   137
       0 < a; 0 < n |] ==> ALL k. (u*(1 + -inverse(real (Suc k)))) ^ n <= a"
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   138
apply safe
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   139
apply (frule less_isLub_not_isUb [THEN not_isUb_less_ex])
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   140
apply (rule_tac n = k in real_mult_less_self)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   141
apply (blast intro: lemma_nth_realpow_isLub_gt_zero, safe)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   142
apply (drule_tac n = k in
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   143
        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
   144
apply (blast intro: order_trans order_less_imp_le power_mono) 
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   145
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   146
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   147
text{*Second result we want*}
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   148
lemma realpow_nth_le:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   149
     "[| (0::real) < a; 0 < n;  
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   150
     isLub (UNIV::real set)  
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   151
     {x. x ^ n <= a & 0 < x} u |] ==> u ^ n <= a"
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   152
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
   153
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
   154
                [THEN LIMSEQ_pow, THEN LIMSEQ_le_const2])
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14325
diff changeset
   155
apply (auto simp add: real_of_nat_def)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   156
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   157
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
   158
text{*The theorem at last!*}
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   159
lemma realpow_nth: "[| (0::real) < a; 0 < n |] ==> \<exists>r. r ^ n = a"
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   160
apply (frule nth_realpow_isLub_ex, auto)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   161
apply (auto intro: realpow_nth_le realpow_nth_ge order_antisym)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   162
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   163
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   164
(* positive only *)
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   165
lemma realpow_pos_nth: "[| (0::real) < a; 0 < n |] ==> \<exists>r. 0 < r & r ^ n = a"
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   166
apply (frule nth_realpow_isLub_ex, auto)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   167
apply (auto intro: realpow_nth_le realpow_nth_ge order_antisym lemma_nth_realpow_isLub_gt_zero)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   168
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   169
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   170
lemma realpow_pos_nth2: "(0::real) < a  ==> \<exists>r. 0 < r & r ^ Suc n = a"
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   171
by (blast intro: realpow_pos_nth)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   172
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   173
(* uniqueness of nth positive root *)
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   174
lemma realpow_pos_nth_unique:
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   175
     "[| (0::real) < a; 0 < n |] ==> EX! r. 0 < r & r ^ n = a"
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   176
apply (auto intro!: realpow_pos_nth)
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   177
apply (cut_tac x = r and y = y in linorder_less_linear, auto)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   178
apply (drule_tac x = r in realpow_less)
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
   179
apply (drule_tac [4] x = y in realpow_less, auto)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   180
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   181
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   182
subsection {* Nth Root *}
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   183
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   184
lemma real_root_zero [simp]: "root (Suc n) 0 = 0"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   185
apply (simp add: root_def)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   186
apply (safe intro!: the_equality power_0_Suc elim!: realpow_zero_zero)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   187
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   188
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   189
lemma real_root_pow_pos: 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   190
     "0 < x ==> (root (Suc n) x) ^ (Suc n) = x"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   191
apply (simp add: root_def del: realpow_Suc)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   192
apply (drule_tac n="Suc n" in realpow_pos_nth_unique, simp)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   193
apply (erule theI' [THEN conjunct2])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   194
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   195
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   196
lemma real_root_pow_pos2: "0 \<le> x ==> (root (Suc n) x) ^ (Suc n) = x"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   197
by (auto dest!: real_le_imp_less_or_eq dest: real_root_pow_pos)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   198
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   199
lemma real_root_pos: 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   200
     "0 < x ==> root(Suc n) (x ^ (Suc n)) = x"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   201
apply (simp add: root_def)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   202
apply (rule the_equality)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   203
apply (frule_tac [2] n = n in zero_less_power)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   204
apply (auto simp add: zero_less_mult_iff)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   205
apply (rule_tac x = u and y = x in linorder_cases)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   206
apply (drule_tac n1 = n and x = u in zero_less_Suc [THEN [3] realpow_less])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   207
apply (drule_tac [4] n1 = n and x = x in zero_less_Suc [THEN [3] realpow_less])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   208
apply (auto)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   209
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   210
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   211
lemma real_root_pos2: "0 \<le> x ==> root(Suc n) (x ^ (Suc n)) = x"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   212
by (auto dest!: real_le_imp_less_or_eq real_root_pos)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   213
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   214
lemma real_root_gt_zero:
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   215
     "0 < x ==> 0 < root (Suc n) x"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   216
apply (simp add: root_def del: realpow_Suc)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   217
apply (drule_tac n="Suc n" in realpow_pos_nth_unique, simp)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   218
apply (erule theI' [THEN conjunct1])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   219
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   220
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   221
lemma real_root_pos_pos: 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   222
     "0 < x ==> 0 \<le> root(Suc n) x"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   223
by (rule real_root_gt_zero [THEN order_less_imp_le])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   224
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   225
lemma real_root_pos_pos_le: "0 \<le> x ==> 0 \<le> root(Suc n) x"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   226
by (auto simp add: order_le_less real_root_gt_zero)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   227
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   228
lemma real_root_one [simp]: "root (Suc n) 1 = 1"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   229
apply (simp add: root_def)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   230
apply (rule the_equality, auto)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   231
apply (rule ccontr)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   232
apply (rule_tac x = u and y = 1 in linorder_cases)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   233
apply (drule_tac n = n in realpow_Suc_less_one)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   234
apply (drule_tac [4] n = n in power_gt1_lemma)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   235
apply (auto)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   236
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   237
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   238
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   239
subsection{*Square Root*}
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   240
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   241
text{*needed because 2 is a binary numeral!*}
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   242
lemma root_2_eq [simp]: "root 2 = root (Suc (Suc 0))"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   243
by (simp del: nat_numeral_0_eq_0 nat_numeral_1_eq_1 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   244
         add: nat_numeral_0_eq_0 [symmetric])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   245
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   246
lemma real_sqrt_zero [simp]: "sqrt 0 = 0"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   247
by (simp add: sqrt_def)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   248
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   249
lemma real_sqrt_one [simp]: "sqrt 1 = 1"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   250
by (simp add: sqrt_def)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   251
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   252
lemma real_sqrt_pow2_iff [iff]: "((sqrt x)\<twosuperior> = x) = (0 \<le> x)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   253
apply (simp add: sqrt_def)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   254
apply (rule iffI) 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   255
 apply (cut_tac r = "root 2 x" in realpow_two_le)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   256
 apply (simp add: numeral_2_eq_2)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   257
apply (subst numeral_2_eq_2)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   258
apply (erule real_root_pow_pos2)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   259
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   260
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   261
lemma [simp]: "(sqrt(u2\<twosuperior> + v2\<twosuperior>))\<twosuperior> = u2\<twosuperior> + v2\<twosuperior>"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   262
by (rule realpow_two_le_add_order [THEN real_sqrt_pow2_iff [THEN iffD2]])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   263
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   264
lemma real_sqrt_pow2 [simp]: "0 \<le> x ==> (sqrt x)\<twosuperior> = x"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   265
by (simp)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   266
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   267
lemma real_sqrt_abs_abs [simp]: "sqrt\<bar>x\<bar> ^ 2 = \<bar>x\<bar>"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   268
by (rule real_sqrt_pow2_iff [THEN iffD2], arith)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   269
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   270
lemma real_pow_sqrt_eq_sqrt_pow: 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   271
      "0 \<le> x ==> (sqrt x)\<twosuperior> = sqrt(x\<twosuperior>)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   272
apply (simp add: sqrt_def)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   273
apply (simp only: numeral_2_eq_2 real_root_pow_pos2 real_root_pos2)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   274
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   275
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   276
lemma real_pow_sqrt_eq_sqrt_abs_pow2:
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   277
     "0 \<le> x ==> (sqrt x)\<twosuperior> = sqrt(\<bar>x\<bar> ^ 2)" 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   278
by (simp add: real_pow_sqrt_eq_sqrt_pow [symmetric])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   279
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   280
lemma real_sqrt_pow_abs: "0 \<le> x ==> (sqrt x)\<twosuperior> = \<bar>x\<bar>"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   281
apply (rule real_sqrt_abs_abs [THEN subst])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   282
apply (rule_tac x1 = x in real_pow_sqrt_eq_sqrt_abs_pow2 [THEN ssubst])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   283
apply (rule_tac [2] real_pow_sqrt_eq_sqrt_pow [symmetric])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   284
apply (assumption, arith)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   285
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   286
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   287
lemma not_real_square_gt_zero [simp]: "(~ (0::real) < x*x) = (x = 0)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   288
apply auto
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   289
apply (cut_tac x = x and y = 0 in linorder_less_linear)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   290
apply (simp add: zero_less_mult_iff)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   291
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   292
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   293
lemma real_sqrt_gt_zero: "0 < x ==> 0 < sqrt(x)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   294
by (simp add: sqrt_def real_root_gt_zero)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   295
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   296
lemma real_sqrt_ge_zero: "0 \<le> x ==> 0 \<le> sqrt(x)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   297
by (auto intro: real_sqrt_gt_zero simp add: order_le_less)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   298
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   299
lemma real_sqrt_mult_self_sum_ge_zero [simp]: "0 \<le> sqrt(x*x + y*y)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   300
by (rule real_sqrt_ge_zero [OF real_mult_self_sum_ge_zero]) 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   301
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   302
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   303
(*we need to prove something like this:
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   304
lemma "[|r ^ n = a; 0<n; 0 < a \<longrightarrow> 0 < r|] ==> root n a = r"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   305
apply (case_tac n, simp) 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   306
apply (simp add: root_def) 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   307
apply (rule someI2 [of _ r], safe)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   308
apply (auto simp del: realpow_Suc dest: power_inject_base)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   309
*)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   310
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   311
lemma sqrt_eqI: "[|r\<twosuperior> = a; 0 \<le> r|] ==> sqrt a = r"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   312
apply (erule subst)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   313
apply (simp add: sqrt_def numeral_2_eq_2 del: realpow_Suc)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   314
apply (erule real_root_pos2)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   315
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   316
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   317
lemma real_sqrt_mult_distrib: 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   318
     "[| 0 \<le> x; 0 \<le> y |] ==> sqrt(x*y) = sqrt(x) * sqrt(y)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   319
apply (rule sqrt_eqI)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   320
apply (simp add: power_mult_distrib)  
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   321
apply (simp add: zero_le_mult_iff real_sqrt_ge_zero) 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   322
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   323
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   324
lemma real_sqrt_mult_distrib2:
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   325
     "[|0\<le>x; 0\<le>y |] ==> sqrt(x*y) =  sqrt(x) * sqrt(y)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   326
by (auto intro: real_sqrt_mult_distrib simp add: order_le_less)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   327
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   328
lemma real_sqrt_sum_squares_ge_zero [simp]: "0 \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   329
by (auto intro!: real_sqrt_ge_zero)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   330
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   331
lemma real_sqrt_sum_squares_mult_ge_zero [simp]:
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   332
     "0 \<le> sqrt ((x\<twosuperior> + y\<twosuperior>)*(xa\<twosuperior> + ya\<twosuperior>))"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   333
by (auto intro!: real_sqrt_ge_zero simp add: zero_le_mult_iff)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   334
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   335
lemma real_sqrt_sum_squares_mult_squared_eq [simp]:
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   336
     "sqrt ((x\<twosuperior> + y\<twosuperior>) * (xa\<twosuperior> + ya\<twosuperior>)) ^ 2 = (x\<twosuperior> + y\<twosuperior>) * (xa\<twosuperior> + ya\<twosuperior>)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   337
by (auto simp add: zero_le_mult_iff simp del: realpow_Suc)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   338
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   339
lemma real_sqrt_abs [simp]: "sqrt(x\<twosuperior>) = \<bar>x\<bar>"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   340
apply (rule abs_realpow_two [THEN subst])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   341
apply (rule real_sqrt_abs_abs [THEN subst])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   342
apply (subst real_pow_sqrt_eq_sqrt_pow)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   343
apply (auto simp add: numeral_2_eq_2)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   344
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   345
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   346
lemma real_sqrt_abs2 [simp]: "sqrt(x*x) = \<bar>x\<bar>"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   347
apply (rule realpow_two [THEN subst])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   348
apply (subst numeral_2_eq_2 [symmetric])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   349
apply (rule real_sqrt_abs)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   350
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   351
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   352
lemma real_sqrt_pow2_gt_zero: "0 < x ==> 0 < (sqrt x)\<twosuperior>"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   353
by simp
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   354
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   355
lemma real_sqrt_not_eq_zero: "0 < x ==> sqrt x \<noteq> 0"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   356
apply (frule real_sqrt_pow2_gt_zero)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   357
apply (auto simp add: numeral_2_eq_2)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   358
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   359
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   360
lemma real_inv_sqrt_pow2: "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x"
20898
113c9516a2d7 attribute symmetric: zero_var_indexes;
wenzelm
parents: 20687
diff changeset
   361
by (cut_tac n = 2 and a = "sqrt x" in power_inverse [symmetric], auto)
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   362
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   363
lemma real_sqrt_eq_zero_cancel: "[| 0 \<le> x; sqrt(x) = 0|] ==> x = 0"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   364
apply (drule real_le_imp_less_or_eq)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   365
apply (auto dest: real_sqrt_not_eq_zero)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   366
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   367
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   368
lemma real_sqrt_eq_zero_cancel_iff [simp]: "0 \<le> x ==> ((sqrt x = 0) = (x=0))"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   369
by (auto simp add: real_sqrt_eq_zero_cancel)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   370
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   371
lemma real_sqrt_sum_squares_ge1 [simp]: "x \<le> sqrt(x\<twosuperior> + y\<twosuperior>)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   372
apply (subgoal_tac "x \<le> 0 | 0 \<le> x", safe)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   373
apply (rule real_le_trans)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   374
apply (auto simp del: realpow_Suc)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   375
apply (rule_tac n = 1 in realpow_increasing)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   376
apply (auto simp add: numeral_2_eq_2 [symmetric] simp del: realpow_Suc)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   377
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   378
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   379
lemma real_sqrt_sum_squares_ge2 [simp]: "y \<le> sqrt(z\<twosuperior> + y\<twosuperior>)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   380
apply (simp (no_asm) add: real_add_commute del: realpow_Suc)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   381
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   382
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   383
lemma real_sqrt_ge_one: "1 \<le> x ==> 1 \<le> sqrt x"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   384
apply (rule_tac n = 1 in realpow_increasing)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   385
apply (auto simp add: numeral_2_eq_2 [symmetric] real_sqrt_ge_zero simp 
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   386
            del: realpow_Suc)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   387
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   388
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   389
end