src/HOL/RealPow.thy
author haftmann
Tue, 02 Mar 2010 08:28:06 +0100
changeset 35437 fe196f61b970
parent 35348 c6331256b087
child 35442 992f9cb60b25
permissions -rw-r--r--
repaired definition (cf. d8d7d1b785af)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 28906
diff changeset
     1
(*  Title       : HOL/RealPow.thy
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot  
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
20634
45fe31e72391 add header
huffman
parents: 20517
diff changeset
     4
*)
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     5
20634
45fe31e72391 add header
huffman
parents: 20517
diff changeset
     6
header {* Natural powers theory *}
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15085
diff changeset
     8
theory RealPow
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
     9
imports RealDef
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15085
diff changeset
    10
begin
9435
c3a13a7d4424 lemmas [arith_split] = abs_split (*belongs to theory RealAbs*);
wenzelm
parents: 9013
diff changeset
    11
35347
be0c69c06176 remove redundant simp rules from RealPow.thy
huffman
parents: 35344
diff changeset
    12
(* FIXME: declare this in Rings.thy or not at all *)
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    13
declare abs_mult_self [simp]
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    14
35347
be0c69c06176 remove redundant simp rules from RealPow.thy
huffman
parents: 35344
diff changeset
    15
(* used by Import/HOL/real.imp *)
be0c69c06176 remove redundant simp rules from RealPow.thy
huffman
parents: 35344
diff changeset
    16
lemma two_realpow_ge_one: "(1::real) \<le> 2 ^ n"
25875
536dfdc25e0a added simp attributes/ proofs fixed
nipkow
parents: 23477
diff changeset
    17
by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    18
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    19
lemma two_realpow_gt [simp]: "real (n::nat) < 2 ^ n"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15229
diff changeset
    20
apply (induct "n")
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    21
apply (auto simp add: real_of_nat_Suc)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14352
diff changeset
    22
apply (subst mult_2)
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35123
diff changeset
    23
apply (erule add_less_le_mono)
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35123
diff changeset
    24
apply (rule two_realpow_ge_one)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    25
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    26
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    27
lemma realpow_Suc_le_self: "[| 0 \<le> r; r \<le> (1::real) |] ==> r ^ Suc n \<le> r"
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    28
by (insert power_decreasing [of 1 "Suc n" r], simp)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    29
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    30
lemma realpow_minus_mult [rule_format]:
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29667
diff changeset
    31
     "0 < n --> (x::real) ^ (n - 1) * x = x ^ n"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    32
apply (simp split add: nat_diff_split)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    33
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    34
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    35
lemma realpow_two_diff:
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    36
     "(x::real)^Suc (Suc 0) - y^Suc (Suc 0) = (x - y) * (x + y)"
35347
be0c69c06176 remove redundant simp rules from RealPow.thy
huffman
parents: 35344
diff changeset
    37
by (simp add: algebra_simps)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    38
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    39
lemma realpow_two_disj:
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    40
     "((x::real)^Suc (Suc 0) = y^Suc (Suc 0)) = (x = y | x = -y)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    41
apply (cut_tac x = x and y = y in realpow_two_diff)
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
    42
apply auto
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    43
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    44
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    45
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    46
subsection{* Squares of Reals *}
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    47
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    48
lemma real_two_squares_add_zero_iff [simp]:
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    49
  "(x * x + y * y = 0) = ((x::real) = 0 \<and> y = 0)"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    50
by (rule sum_squares_eq_zero_iff)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    51
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    52
lemma real_sum_squares_cancel: "x * x + y * y = 0 ==> x = (0::real)"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    53
by simp
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    54
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    55
lemma real_sum_squares_cancel2: "x * x + y * y = 0 ==> y = (0::real)"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    56
by simp
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    57
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    58
lemma real_sum_squares_cancel_a: "x * x = -(y * y) ==> x = (0::real) & y=0"
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    59
by (simp add: real_add_eq_0_iff [symmetric])
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    60
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    61
lemma real_squared_diff_one_factored: "x*x - (1::real) = (x + 1)*(x - 1)"
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    62
by (simp add: left_distrib right_diff_distrib)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    63
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    64
lemma real_mult_is_one [simp]: "(x*x = (1::real)) = (x = 1 | x = - 1)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    65
apply auto
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    66
apply (drule right_minus_eq [THEN iffD2]) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    67
apply (auto simp add: real_squared_diff_one_factored)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    68
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    69
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    70
lemma real_sum_squares_not_zero: "x ~= 0 ==> x * x + y * y ~= (0::real)"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    71
by simp
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    72
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    73
lemma real_sum_squares_not_zero2: "y ~= 0 ==> x * x + y * y ~= (0::real)"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    74
by simp
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    75
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    76
lemma realpow_two_sum_zero_iff [simp]:
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    77
     "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    78
by (rule sum_power2_eq_zero_iff)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    79
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    80
lemma realpow_two_le_add_order [simp]: "(0::real) \<le> u ^ 2 + v ^ 2"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    81
by (rule sum_power2_ge_zero)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    82
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    83
lemma realpow_two_le_add_order2 [simp]: "(0::real) \<le> u ^ 2 + v ^ 2 + w ^ 2"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    84
by (intro add_nonneg_nonneg zero_le_power2)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    85
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    86
lemma real_sum_square_gt_zero: "x ~= 0 ==> (0::real) < x * x + y * y"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    87
by (simp add: sum_squares_gt_zero_iff)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    88
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    89
lemma real_sum_square_gt_zero2: "y ~= 0 ==> (0::real) < x * x + y * y"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    90
by (simp add: sum_squares_gt_zero_iff)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    91
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    92
lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    93
by (rule_tac j = 0 in real_le_trans, auto)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    94
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    95
lemma realpow_square_minus_le [simp]: "-(u ^ 2) \<le> (x::real) ^ 2"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    96
by (auto simp add: power2_eq_square)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    97
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    98
(* The following theorem is by Benjamin Porter *)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    99
lemma real_sq_order:
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   100
  fixes x::real
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   101
  assumes xgt0: "0 \<le> x" and ygt0: "0 \<le> y" and sq: "x^2 \<le> y^2"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   102
  shows "x \<le> y"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   103
proof -
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   104
  from sq have "x ^ Suc (Suc 0) \<le> y ^ Suc (Suc 0)"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   105
    by (simp only: numeral_2_eq_2)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   106
  thus "x \<le> y" using ygt0
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   107
    by (rule power_le_imp_le_base)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   108
qed
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   109
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   110
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   111
subsection {*Various Other Theorems*}
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   112
14304
cc0b4bbfbc43 minor tweaks
paulson
parents: 14288
diff changeset
   113
lemma real_le_add_half_cancel: "(x + y/2 \<le> (y::real)) = (x \<le> y /2)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
   114
by auto
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   115
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
   116
lemma real_minus_half_eq [simp]: "(x::real) - x/2 = x/2"
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
   117
by auto
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   118
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   119
lemma real_mult_inverse_cancel:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   120
     "[|(0::real) < x; 0 < x1; x1 * y < x * u |] 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   121
      ==> inverse x * y < inverse x1 * u"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   122
apply (rule_tac c=x in mult_less_imp_less_left) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   123
apply (auto simp add: real_mult_assoc [symmetric])
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14304
diff changeset
   124
apply (simp (no_asm) add: mult_ac)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   125
apply (rule_tac c=x1 in mult_less_imp_less_right) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14304
diff changeset
   126
apply (auto simp add: mult_ac)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   127
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   128
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
   129
lemma real_mult_inverse_cancel2:
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
   130
     "[|(0::real) < x;0 < x1; x1 * y < x * u |] ==> y * inverse x < u * inverse x1"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14304
diff changeset
   131
apply (auto dest: real_mult_inverse_cancel simp add: mult_ac)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   132
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   133
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   134
lemma realpow_num_eq_if: "(m::real) ^ n = (if n=0 then 1 else m * m ^ (n - 1))"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
   135
by (case_tac "n", auto)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   136
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   137
end