src/HOL/Real/RealPow.thy
author paulson
Fri, 21 Nov 2003 11:15:40 +0100
changeset 14265 95b42e69436c
parent 12018 ec054019c910
child 14268 5cf13e80be0e
permissions -rw-r--r--
HOL: installation of Ring_and_Field as the basis for Naturals and Reals
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9435
c3a13a7d4424 lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents: 9013
diff changeset
     1
(*  Title       : HOL/Real/RealPow.thy
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7077
diff changeset
     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
9435
c3a13a7d4424 lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents: 9013
diff changeset
     9
theory RealPow = RealAbs:
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    10
9435
c3a13a7d4424 lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents: 9013
diff changeset
    11
(*belongs to theory RealAbs*)
c3a13a7d4424 lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents: 9013
diff changeset
    12
lemmas [arith_split] = abs_split
c3a13a7d4424 lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents: 9013
diff changeset
    13
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 9435
diff changeset
    14
instance real :: power ..
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    15
8856
435187ffc64e fixed theory deps;
wenzelm
parents: 7219
diff changeset
    16
primrec (realpow)
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11701
diff changeset
    17
     realpow_0:   "r ^ 0       = 1"
9435
c3a13a7d4424 lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents: 9013
diff changeset
    18
     realpow_Suc: "r ^ (Suc n) = (r::real) * (r ^ n)"
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    19
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    20
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    21
(*FIXME: most of the theorems for Suc (Suc 0) are redundant!
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    22
*)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    23
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    24
lemma realpow_zero: "(0::real) ^ (Suc n) = 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    25
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    26
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    27
declare realpow_zero [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    28
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    29
lemma realpow_not_zero [rule_format (no_asm)]: "r \<noteq> (0::real) --> r ^ n \<noteq> 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    30
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    31
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    32
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    33
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    34
lemma realpow_zero_zero: "r ^ n = (0::real) ==> r = 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    35
apply (rule ccontr)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    36
apply (auto dest: realpow_not_zero)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    37
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    38
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    39
lemma realpow_inverse: "inverse ((r::real) ^ n) = (inverse r) ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    40
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    41
apply (auto simp add: real_inverse_distrib)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    42
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    43
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    44
lemma realpow_abs: "abs(r ^ n) = abs(r::real) ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    45
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    46
apply (auto simp add: abs_mult)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    47
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    48
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    49
lemma realpow_add: "(r::real) ^ (n + m) = (r ^ n) * (r ^ m)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    50
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    51
apply (auto simp add: real_mult_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    52
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    53
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    54
lemma realpow_one: "(r::real) ^ 1 = r"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    55
apply (simp (no_asm))
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    56
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    57
declare realpow_one [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    58
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    59
lemma realpow_two: "(r::real)^ (Suc (Suc 0)) = r * r"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    60
apply (simp (no_asm))
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    61
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    62
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    63
lemma realpow_gt_zero [rule_format (no_asm)]: "(0::real) < r --> 0 < r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    64
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    65
apply (auto intro: real_mult_order simp add: real_zero_less_one)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    66
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    67
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    68
lemma realpow_ge_zero [rule_format (no_asm)]: "(0::real) \<le> r --> 0 \<le> r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    69
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    70
apply (auto simp add: real_0_le_mult_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    71
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    72
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    73
lemma realpow_le [rule_format (no_asm)]: "(0::real) \<le> x & x \<le> y --> x ^ n \<le> y ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    74
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    75
apply (auto intro!: real_mult_le_mono)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    76
apply (simp (no_asm_simp) add: realpow_ge_zero)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    77
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    78
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    79
lemma realpow_0_left [rule_format, simp]:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    80
     "0 < n --> 0 ^ n = (0::real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    81
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    82
apply (auto ); 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    83
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    84
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    85
lemma realpow_less' [rule_format]:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    86
     "[|(0::real) \<le> x; x < y |] ==> 0 < n --> x ^ n < y ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    87
apply (induct n) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    88
apply (auto simp add: real_mult_less_mono' realpow_ge_zero); 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    89
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    90
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    91
text{*Legacy: weaker version of the theorem above*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    92
lemma realpow_less [rule_format]:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    93
     "[|(0::real) < x; x < y; 0 < n|] ==> x ^ n < y ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    94
apply (rule realpow_less') 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    95
apply (auto ); 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    96
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    97
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    98
lemma realpow_eq_one: "1 ^ n = (1::real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    99
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   100
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   101
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   102
declare realpow_eq_one [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   103
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   104
lemma abs_realpow_minus_one: "abs((-1) ^ n) = (1::real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   105
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   106
apply (auto simp add: abs_mult)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   107
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   108
declare abs_realpow_minus_one [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   109
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   110
lemma realpow_mult: "((r::real) * s) ^ n = (r ^ n) * (s ^ n)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   111
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   112
apply (auto simp add: real_mult_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   113
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   114
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   115
lemma realpow_two_le: "(0::real) \<le> r^ Suc (Suc 0)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   116
apply (simp (no_asm) add: real_le_square)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   117
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   118
declare realpow_two_le [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   119
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   120
lemma abs_realpow_two: "abs((x::real)^Suc (Suc 0)) = x^Suc (Suc 0)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   121
apply (simp (no_asm) add: abs_eqI1 real_le_square)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   122
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   123
declare abs_realpow_two [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   124
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   125
lemma realpow_two_abs: "abs(x::real)^Suc (Suc 0) = x^Suc (Suc 0)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   126
apply (simp (no_asm) add: realpow_abs [symmetric] abs_eqI1 del: realpow_Suc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   127
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   128
declare realpow_two_abs [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   129
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   130
lemma realpow_two_gt_one: "(1::real) < r ==> 1 < r^ (Suc (Suc 0))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   131
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   132
apply (cut_tac real_zero_less_one)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   133
apply (frule_tac x = "0" in order_less_trans)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   134
apply assumption
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   135
apply (drule_tac  z = "r" and x = "1" in real_mult_less_mono1)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   136
apply (auto intro: order_less_trans)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   137
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   138
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   139
lemma realpow_ge_one [rule_format (no_asm)]: "(1::real) < r --> 1 \<le> r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   140
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   141
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   142
apply (subgoal_tac "1*1 \<le> r * r^n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   143
apply (rule_tac [2] real_mult_le_mono)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   144
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   145
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   146
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   147
lemma realpow_ge_one2: "(1::real) \<le> r ==> 1 \<le> r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   148
apply (drule order_le_imp_less_or_eq)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   149
apply (auto dest: realpow_ge_one)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   150
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   151
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   152
lemma two_realpow_ge_one: "(1::real) \<le> 2 ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   153
apply (rule_tac y = "1 ^ n" in order_trans)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   154
apply (rule_tac [2] realpow_le)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   155
apply (auto intro: order_less_imp_le)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   156
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   157
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   158
lemma two_realpow_gt: "real (n::nat) < 2 ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   159
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   160
apply (auto simp add: real_of_nat_Suc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   161
apply (subst real_mult_2)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   162
apply (rule real_add_less_le_mono)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   163
apply (auto simp add: two_realpow_ge_one)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   164
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   165
declare two_realpow_gt [simp] two_realpow_ge_one [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   166
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   167
lemma realpow_minus_one: "(-1) ^ (2*n) = (1::real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   168
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   169
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   170
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   171
declare realpow_minus_one [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   172
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   173
lemma realpow_minus_one_odd: "(-1) ^ Suc (2*n) = -(1::real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   174
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   175
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   176
declare realpow_minus_one_odd [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   177
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   178
lemma realpow_minus_one_even: "(-1) ^ Suc (Suc (2*n)) = (1::real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   179
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   180
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   181
declare realpow_minus_one_even [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   182
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   183
lemma realpow_Suc_less [rule_format (no_asm)]: "(0::real) < r & r < (1::real) --> r ^ Suc n < r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   184
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   185
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   186
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   187
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   188
lemma realpow_Suc_le [rule_format (no_asm)]: "0 \<le> r & r < (1::real) --> r ^ Suc n \<le> r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   189
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   190
apply (auto intro: order_less_imp_le dest!: order_le_imp_less_or_eq)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   191
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   192
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   193
lemma realpow_zero_le: "(0::real) \<le> 0 ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   194
apply (case_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   195
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   196
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   197
declare realpow_zero_le [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   198
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   199
lemma realpow_Suc_le2 [rule_format (no_asm)]: "0 < r & r < (1::real) --> r ^ Suc n \<le> r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   200
apply (blast intro!: order_less_imp_le realpow_Suc_less)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   201
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   202
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   203
lemma realpow_Suc_le3: "[| 0 \<le> r; r < (1::real) |] ==> r ^ Suc n \<le> r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   204
apply (erule order_le_imp_less_or_eq [THEN disjE])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   205
apply (rule realpow_Suc_le2)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   206
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   207
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   208
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   209
lemma realpow_less_le [rule_format (no_asm)]: "0 \<le> r & r < (1::real) & n < N --> r ^ N \<le> r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   210
apply (induct_tac "N")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   211
apply (simp_all (no_asm_simp))
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   212
apply clarify
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   213
apply (subgoal_tac "r * r ^ na \<le> 1 * r ^ n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   214
apply simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   215
apply (rule real_mult_le_mono)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   216
apply (auto simp add: realpow_ge_zero less_Suc_eq)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   217
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   218
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   219
lemma realpow_le_le: "[| 0 \<le> r; r < (1::real); n \<le> N |] ==> r ^ N \<le> r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   220
apply (drule_tac n = "N" in le_imp_less_or_eq)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   221
apply (auto intro: realpow_less_le)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   222
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   223
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   224
lemma realpow_Suc_le_self: "[| 0 < r; r < (1::real) |] ==> r ^ Suc n \<le> r"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   225
apply (drule_tac n = "1" and N = "Suc n" in order_less_imp_le [THEN realpow_le_le])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   226
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   227
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   228
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   229
lemma realpow_Suc_less_one: "[| 0 < r; r < (1::real) |] ==> r ^ Suc n < 1"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   230
apply (blast intro: realpow_Suc_le_self order_le_less_trans)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   231
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   232
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   233
lemma realpow_le_Suc [rule_format (no_asm)]: "(1::real) \<le> r --> r ^ n \<le> r ^ Suc n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   234
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   235
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   236
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   237
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   238
lemma realpow_less_Suc [rule_format (no_asm)]: "(1::real) < r --> r ^ n < r ^ Suc n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   239
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   240
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   241
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   242
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   243
lemma realpow_le_Suc2 [rule_format (no_asm)]: "(1::real) < r --> r ^ n \<le> r ^ Suc n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   244
apply (blast intro!: order_less_imp_le realpow_less_Suc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   245
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   246
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   247
lemma realpow_gt_ge [rule_format (no_asm)]: "(1::real) < r & n < N --> r ^ n \<le> r ^ N"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   248
apply (induct_tac "N")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   249
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   250
apply (frule_tac [!] n = "na" in realpow_ge_one)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   251
apply (drule_tac [!] real_mult_self_le)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   252
apply assumption
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   253
prefer 2 apply (assumption)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   254
apply (auto intro: order_trans simp add: less_Suc_eq)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   255
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   256
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   257
lemma realpow_gt_ge2 [rule_format (no_asm)]: "(1::real) \<le> r & n < N --> r ^ n \<le> r ^ N"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   258
apply (induct_tac "N")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   259
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   260
apply (frule_tac [!] n = "na" in realpow_ge_one2)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   261
apply (drule_tac [!] real_mult_self_le2)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   262
apply assumption
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   263
prefer 2 apply (assumption)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   264
apply (auto intro: order_trans simp add: less_Suc_eq)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   265
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   266
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   267
lemma realpow_ge_ge: "[| (1::real) < r; n \<le> N |] ==> r ^ n \<le> r ^ N"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   268
apply (drule_tac n = "N" in le_imp_less_or_eq)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   269
apply (auto intro: realpow_gt_ge)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   270
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   271
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   272
lemma realpow_ge_ge2: "[| (1::real) \<le> r; n \<le> N |] ==> r ^ n \<le> r ^ N"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   273
apply (drule_tac n = "N" in le_imp_less_or_eq)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   274
apply (auto intro: realpow_gt_ge2)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   275
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   276
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   277
lemma realpow_Suc_ge_self [rule_format (no_asm)]: "(1::real) < r ==> r \<le> r ^ Suc n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   278
apply (drule_tac n = "1" and N = "Suc n" in realpow_ge_ge)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   279
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   280
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   281
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   282
lemma realpow_Suc_ge_self2 [rule_format (no_asm)]: "(1::real) \<le> r ==> r \<le> r ^ Suc n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   283
apply (drule_tac n = "1" and N = "Suc n" in realpow_ge_ge2)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   284
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   285
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   286
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   287
lemma realpow_ge_self: "[| (1::real) < r; 0 < n |] ==> r \<le> r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   288
apply (drule less_not_refl2 [THEN not0_implies_Suc])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   289
apply (auto intro!: realpow_Suc_ge_self)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   290
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   291
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   292
lemma realpow_ge_self2: "[| (1::real) \<le> r; 0 < n |] ==> r \<le> r ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   293
apply (drule less_not_refl2 [THEN not0_implies_Suc])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   294
apply (auto intro!: realpow_Suc_ge_self2)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   295
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   296
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   297
lemma realpow_minus_mult [rule_format (no_asm)]: "0 < n --> (x::real) ^ (n - 1) * x = x ^ n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   298
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   299
apply (auto simp add: real_mult_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   300
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   301
declare realpow_minus_mult [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   302
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   303
lemma realpow_two_mult_inverse: "r \<noteq> 0 ==> r * inverse r ^Suc (Suc 0) = inverse (r::real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   304
apply (simp (no_asm_simp) add: realpow_two real_mult_assoc [symmetric])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   305
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   306
declare realpow_two_mult_inverse [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   307
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   308
(* 05/00 *)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   309
lemma realpow_two_minus: "(-x)^Suc (Suc 0) = (x::real)^Suc (Suc 0)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   310
apply (simp (no_asm))
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   311
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   312
declare realpow_two_minus [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   313
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   314
lemma realpow_two_diff: "(x::real)^Suc (Suc 0) - y^Suc (Suc 0) = (x - y) * (x + y)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   315
apply (unfold real_diff_def)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   316
apply (simp add: real_add_mult_distrib2 real_add_mult_distrib real_mult_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   317
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   318
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   319
lemma realpow_two_disj: "((x::real)^Suc (Suc 0) = y^Suc (Suc 0)) = (x = y | x = -y)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   320
apply (cut_tac x = "x" and y = "y" in realpow_two_diff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   321
apply (auto simp del: realpow_Suc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   322
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   323
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   324
(* used in Transc *)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   325
lemma realpow_diff: "[|(x::real) \<noteq> 0; m \<le> n |] ==> x ^ (n - m) = x ^ n * inverse (x ^ m)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   326
apply (auto simp add: le_eq_less_or_eq less_iff_Suc_add realpow_add realpow_not_zero real_mult_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   327
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   328
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   329
lemma realpow_real_of_nat: "real (m::nat) ^ n = real (m ^ n)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   330
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   331
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: 12018
diff changeset
   332
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   333
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   334
lemma realpow_real_of_nat_two_pos: "0 < real (Suc (Suc 0) ^ n)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   335
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   336
apply (auto simp add: real_of_nat_mult real_0_less_mult_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   337
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   338
declare realpow_real_of_nat_two_pos [simp] 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   339
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   340
lemma realpow_increasing:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   341
  assumes xnonneg: "(0::real) \<le> x"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   342
      and ynonneg: "0 \<le> y"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   343
      and le: "x ^ Suc n \<le> y ^ Suc n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   344
  shows "x \<le> y"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   345
 proof (rule ccontr)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   346
   assume "~ x \<le> y"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   347
   then have "y<x" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   348
   then have "y ^ Suc n < x ^ Suc n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   349
     by (simp only: prems realpow_less') 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   350
   from le and this show "False"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   351
     by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   352
 qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   353
  
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   354
lemma realpow_Suc_cancel_eq: "[| (0::real) \<le> x; 0 \<le> y; x ^ Suc n = y ^ Suc n |] ==> x = y"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   355
apply (blast intro: realpow_increasing order_antisym order_eq_refl sym)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   356
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   357
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   358
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   359
(*** Logical equivalences for inequalities ***)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   360
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   361
lemma realpow_eq_0_iff: "(x^n = 0) = (x = (0::real) & 0<n)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   362
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   363
apply auto
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   364
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   365
declare realpow_eq_0_iff [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   366
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   367
lemma zero_less_realpow_abs_iff: "(0 < (abs x)^n) = (x \<noteq> (0::real) | n=0)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   368
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   369
apply (auto simp add: real_0_less_mult_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   370
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   371
declare zero_less_realpow_abs_iff [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   372
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   373
lemma zero_le_realpow_abs: "(0::real) \<le> (abs x)^n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   374
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   375
apply (auto simp add: real_0_le_mult_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   376
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   377
declare zero_le_realpow_abs [simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   378
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   379
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   380
(*** Literal arithmetic involving powers, type real ***)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   381
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   382
lemma real_of_int_power: "real (x::int) ^ n = real (x ^ n)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   383
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   384
apply (simp_all (no_asm_simp) add: nat_mult_distrib)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   385
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   386
declare real_of_int_power [symmetric, simp]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   387
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   388
lemma power_real_number_of: "(number_of v :: real) ^ n = real ((number_of v :: int) ^ n)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   389
apply (simp only: real_number_of_def real_of_int_power)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   390
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   391
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   392
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: 12018
diff changeset
   393
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   394
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   395
lemma real_power_two: "(r::real)\<twosuperior> = r * r"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   396
  by (simp add: numeral_2_eq_2)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   397
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   398
lemma real_sqr_ge_zero [iff]: "0 \<le> (r::real)\<twosuperior>"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   399
  by (simp add: real_power_two)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   400
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   401
lemma real_sqr_gt_zero: "(r::real) \<noteq> 0 ==> 0 < r\<twosuperior>"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   402
proof -
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   403
  assume "r \<noteq> 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   404
  hence "0 \<noteq> r\<twosuperior>" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   405
  also have "0 \<le> r\<twosuperior>" by (simp add: real_sqr_ge_zero)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   406
  finally show ?thesis .
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   407
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   408
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   409
lemma real_sqr_not_zero: "r \<noteq> 0 ==> (r::real)\<twosuperior> \<noteq> 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   410
  by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   411
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   412
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   413
ML
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   414
{*
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   415
val realpow_0 = thm "realpow_0";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   416
val realpow_Suc = thm "realpow_Suc";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   417
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   418
val realpow_zero = thm "realpow_zero";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   419
val realpow_not_zero = thm "realpow_not_zero";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   420
val realpow_zero_zero = thm "realpow_zero_zero";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   421
val realpow_inverse = thm "realpow_inverse";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   422
val realpow_abs = thm "realpow_abs";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   423
val realpow_add = thm "realpow_add";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   424
val realpow_one = thm "realpow_one";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   425
val realpow_two = thm "realpow_two";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   426
val realpow_gt_zero = thm "realpow_gt_zero";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   427
val realpow_ge_zero = thm "realpow_ge_zero";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   428
val realpow_le = thm "realpow_le";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   429
val realpow_0_left = thm "realpow_0_left";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   430
val realpow_less = thm "realpow_less";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   431
val realpow_eq_one = thm "realpow_eq_one";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   432
val abs_realpow_minus_one = thm "abs_realpow_minus_one";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   433
val realpow_mult = thm "realpow_mult";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   434
val realpow_two_le = thm "realpow_two_le";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   435
val abs_realpow_two = thm "abs_realpow_two";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   436
val realpow_two_abs = thm "realpow_two_abs";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   437
val realpow_two_gt_one = thm "realpow_two_gt_one";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   438
val realpow_ge_one = thm "realpow_ge_one";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   439
val realpow_ge_one2 = thm "realpow_ge_one2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   440
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: 12018
diff changeset
   441
val two_realpow_gt = thm "two_realpow_gt";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   442
val realpow_minus_one = thm "realpow_minus_one";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   443
val realpow_minus_one_odd = thm "realpow_minus_one_odd";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   444
val realpow_minus_one_even = thm "realpow_minus_one_even";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   445
val realpow_Suc_less = thm "realpow_Suc_less";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   446
val realpow_Suc_le = thm "realpow_Suc_le";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   447
val realpow_zero_le = thm "realpow_zero_le";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   448
val realpow_Suc_le2 = thm "realpow_Suc_le2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   449
val realpow_Suc_le3 = thm "realpow_Suc_le3";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   450
val realpow_less_le = thm "realpow_less_le";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   451
val realpow_le_le = thm "realpow_le_le";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   452
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: 12018
diff changeset
   453
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: 12018
diff changeset
   454
val realpow_le_Suc = thm "realpow_le_Suc";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   455
val realpow_less_Suc = thm "realpow_less_Suc";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   456
val realpow_le_Suc2 = thm "realpow_le_Suc2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   457
val realpow_gt_ge = thm "realpow_gt_ge";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   458
val realpow_gt_ge2 = thm "realpow_gt_ge2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   459
val realpow_ge_ge = thm "realpow_ge_ge";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   460
val realpow_ge_ge2 = thm "realpow_ge_ge2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   461
val realpow_Suc_ge_self = thm "realpow_Suc_ge_self";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   462
val realpow_Suc_ge_self2 = thm "realpow_Suc_ge_self2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   463
val realpow_ge_self = thm "realpow_ge_self";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   464
val realpow_ge_self2 = thm "realpow_ge_self2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   465
val realpow_minus_mult = thm "realpow_minus_mult";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   466
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: 12018
diff changeset
   467
val realpow_two_minus = thm "realpow_two_minus";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   468
val realpow_two_diff = thm "realpow_two_diff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   469
val realpow_two_disj = thm "realpow_two_disj";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   470
val realpow_diff = thm "realpow_diff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   471
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: 12018
diff changeset
   472
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: 12018
diff changeset
   473
val realpow_increasing = thm "realpow_increasing";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   474
val realpow_Suc_cancel_eq = thm "realpow_Suc_cancel_eq";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   475
val realpow_eq_0_iff = thm "realpow_eq_0_iff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   476
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: 12018
diff changeset
   477
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: 12018
diff changeset
   478
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: 12018
diff changeset
   479
val power_real_number_of = thm "power_real_number_of";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   480
val real_power_two = thm "real_power_two";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   481
val real_sqr_ge_zero = thm "real_sqr_ge_zero";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   482
val real_sqr_gt_zero = thm "real_sqr_gt_zero";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   483
val real_sqr_not_zero = thm "real_sqr_not_zero";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   484
*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   485
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
   486
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   487
end