src/HOL/RealPow.thy
author huffman
Mon, 10 May 2010 12:12:58 -0700
changeset 36795 e05e1283c550
parent 36777 be5461582d0f
child 36821 9207505d1ee5
permissions -rw-r--r--
new construction of real numbers using Cauchy sequences
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
35578
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
     9
imports RealDef RComplete
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
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    27
(* TODO: no longer real-specific; rename and move elsewhere *)
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    28
lemma realpow_Suc_le_self:
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    29
  fixes r :: "'a::linordered_semidom"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    30
  shows "[| 0 \<le> r; r \<le> 1 |] ==> r ^ Suc n \<le> r"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    31
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
    32
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    33
(* TODO: no longer real-specific; rename and move elsewhere *)
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    34
lemma realpow_minus_mult:
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    35
  fixes x :: "'a::monoid_mult"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    36
  shows "0 < n \<Longrightarrow> x ^ (n - 1) * x = x ^ n"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    37
by (simp add: power_commutes 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
    38
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    39
(* TODO: no longer real-specific; rename and move elsewhere *)
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    40
lemma realpow_two_diff:
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    41
  fixes x :: "'a::comm_ring_1"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    42
  shows "x^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
    43
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
    44
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    45
(* TODO: move elsewhere *)
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    46
lemma add_eq_0_iff:
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    47
  fixes x y :: "'a::group_add"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    48
  shows "x + y = 0 \<longleftrightarrow> y = - x"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    49
by (auto dest: minus_unique)
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    50
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    51
(* TODO: no longer real-specific; rename and move elsewhere *)
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
    52
lemma realpow_two_disj:
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    53
  fixes x :: "'a::idom"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    54
  shows "(x^Suc (Suc 0) = y^Suc (Suc 0)) = (x = y | x = -y)"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    55
using realpow_two_diff [of x y]
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    56
by (auto simp add: add_eq_0_iff)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    57
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 12018
diff changeset
    58
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    59
subsection{* Squares of Reals *}
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    60
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    61
(* FIXME: declare this [simp] for all types, or not at all *)
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    62
lemma real_two_squares_add_zero_iff [simp]:
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    63
  "(x * x + y * y = 0) = ((x::real) = 0 \<and> y = 0)"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    64
by (rule sum_squares_eq_zero_iff)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    65
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    66
(* TODO: no longer real-specific; rename and move elsewhere *)
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    67
lemma real_squared_diff_one_factored:
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    68
  fixes x :: "'a::ring_1"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    69
  shows "x * x - 1 = (x + 1) * (x - 1)"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    70
by (simp add: algebra_simps)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    71
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    72
(* TODO: no longer real-specific; rename and move elsewhere *)
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    73
lemma real_mult_is_one [simp]:
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    74
  fixes x :: "'a::ring_1_no_zero_divisors"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    75
  shows "x * x = 1 \<longleftrightarrow> x = 1 \<or> x = - 1"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    76
proof -
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    77
  have "x * x = 1 \<longleftrightarrow> (x + 1) * (x - 1) = 0"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    78
    by (simp add: algebra_simps)
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    79
  also have "\<dots> \<longleftrightarrow> x = 1 \<or> x = - 1"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    80
    by (auto simp add: add_eq_0_iff minus_equation_iff [of _ 1])
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    81
  finally show ?thesis .
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    82
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
    83
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    84
(* FIXME: declare this [simp] for all types, or not at all *)
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    85
lemma realpow_two_sum_zero_iff [simp]:
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    86
     "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    87
by (rule sum_power2_eq_zero_iff)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    88
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    89
(* FIXME: declare this [simp] for all types, or not at all *)
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    90
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
    91
by (rule sum_power2_ge_zero)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    92
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
    93
(* FIXME: declare this [simp] for all types, or not at all *)
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    94
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
    95
by (intro add_nonneg_nonneg zero_le_power2)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    96
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    97
lemma real_minus_mult_self_le [simp]: "-(u * u) \<le> (x * (x::real))"
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 35632
diff changeset
    98
by (rule_tac y = 0 in order_trans, auto)
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
    99
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   100
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
   101
by (auto simp add: power2_eq_square)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   102
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   103
(* The following theorem is by Benjamin Porter *)
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
   104
(* TODO: no longer real-specific; rename and move elsewhere *)
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   105
lemma real_sq_order:
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
   106
  fixes x :: "'a::linordered_semidom"
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   107
  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
   108
  shows "x \<le> y"
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   109
proof -
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   110
  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
   111
    by (simp only: numeral_2_eq_2)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   112
  thus "x \<le> y" using ygt0
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   113
    by (rule power_le_imp_le_base)
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   114
qed
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   115
35578
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   116
subsection {*Floor*}
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   117
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   118
lemma floor_power:
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   119
  assumes "x = real (floor x)"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   120
  shows "floor (x ^ n) = floor x ^ n"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   121
proof -
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   122
  have *: "x ^ n = real (floor x ^ n)"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   123
    using assms by (induct n arbitrary: x) simp_all
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   124
  show ?thesis unfolding real_of_int_inject[symmetric]
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   125
    unfolding * floor_real_of_int ..
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   126
qed
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   127
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   128
lemma natfloor_power:
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   129
  assumes "x = real (natfloor x)"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   130
  shows "natfloor (x ^ n) = natfloor x ^ n"
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   131
proof -
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   132
  from assms have "0 \<le> floor x" by auto
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   133
  note assms[unfolded natfloor_def real_nat_eq_real[OF `0 \<le> floor x`]]
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   134
  from floor_power[OF this]
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   135
  show ?thesis unfolding natfloor_def nat_power_eq[OF `0 \<le> floor x`, symmetric]
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   136
    by simp
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   137
qed
22970
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   138
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   139
subsection {*Various Other Theorems*}
b5910e3dec4c move lemmas to RealPow.thy; tuned proofs
huffman
parents: 22967
diff changeset
   140
14304
cc0b4bbfbc43 minor tweaks
paulson
parents: 14288
diff changeset
   141
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
   142
by auto
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   143
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   144
lemma real_mult_inverse_cancel:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   145
     "[|(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
   146
      ==> inverse x * y < inverse x1 * u"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   147
apply (rule_tac c=x in mult_less_imp_less_left) 
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 35632
diff changeset
   148
apply (auto simp add: mult_assoc [symmetric])
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14304
diff changeset
   149
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
   150
apply (rule_tac c=x1 in mult_less_imp_less_right) 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14304
diff changeset
   151
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
   152
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   153
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
   154
lemma real_mult_inverse_cancel2:
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14334
diff changeset
   155
     "[|(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
   156
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
   157
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   158
35632
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
   159
(* TODO: no longer real-specific; rename and move elsewhere *)
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
   160
lemma realpow_num_eq_if:
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
   161
  fixes m :: "'a::power"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
   162
  shows "m ^ n = (if n=0 then 1 else m * m ^ (n - 1))"
61fd75e33137 generalize some lemmas, and remove a few unnecessary ones
huffman
parents: 35578
diff changeset
   163
by (cases n, auto)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14265
diff changeset
   164
35578
384ad08a1d1b Added natfloor and floor rules for multiplication and power.
hoelzl
parents: 35442
diff changeset
   165
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   166
end