| author | wenzelm | 
| Wed, 25 Jun 2008 22:11:17 +0200 | |
| changeset 27364 | a8672b0e2b15 | 
| parent 27105 | 5f139027c365 | 
| child 27435 | b3f8e9bdf9a7 | 
| permissions | -rw-r--r-- | 
| 20245 | 1 | (* Title : HOL/Hyperreal/HyperDef.thy | 
| 10751 | 2 | ID : $Id$ | 
| 3 | Author : Jacques D. Fleuriot | |
| 4 | Copyright : 1998 University of Cambridge | |
| 14468 | 5 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | 
| 13487 | 6 | *) | 
| 10751 | 7 | |
| 14468 | 8 | header{*Construction of Hyperreals Using Ultrafilters*}
 | 
| 9 | ||
| 15131 | 10 | theory HyperDef | 
| 21865 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 11 | imports HyperNat "../Real/Real" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 12 | uses ("hypreal_arith.ML")
 | 
| 15131 | 13 | begin | 
| 10751 | 14 | |
| 17298 | 15 | types hypreal = "real star" | 
| 14299 | 16 | |
| 19380 | 17 | abbreviation | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 18 | hypreal_of_real :: "real => real star" where | 
| 19380 | 19 | "hypreal_of_real == star_of" | 
| 10751 | 20 | |
| 21865 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 21 | abbreviation | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 22 | hypreal_of_hypnat :: "hypnat \<Rightarrow> hypreal" where | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 23 | "hypreal_of_hypnat \<equiv> of_hypnat" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 24 | |
| 19765 | 25 | definition | 
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 26 | omega :: hypreal where | 
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 27 |    -- {*an infinite number @{text "= [<1,2,3,...>]"} *}
 | 
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 28 | "omega = star_n (\<lambda>n. real (Suc n))" | 
| 13487 | 29 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 30 | definition | 
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 31 | epsilon :: hypreal where | 
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 32 |    -- {*an infinitesimal number @{text "= [<1,1/2,1/3,...>]"} *}
 | 
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 33 | "epsilon = star_n (\<lambda>n. inverse (real (Suc n)))" | 
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10834diff
changeset | 34 | |
| 21210 | 35 | notation (xsymbols) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 36 |   omega  ("\<omega>") and
 | 
| 19765 | 37 |   epsilon  ("\<epsilon>")
 | 
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10834diff
changeset | 38 | |
| 21210 | 39 | notation (HTML output) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 40 |   omega  ("\<omega>") and
 | 
| 19765 | 41 |   epsilon  ("\<epsilon>")
 | 
| 14565 | 42 | |
| 14329 | 43 | |
| 20555 | 44 | subsection {* Real vector class instances *}
 | 
| 45 | ||
| 26566 
36a93808642c
instantiation replacing primitive instance plus overloaded defs
 haftmann parents: 
24075diff
changeset | 46 | instantiation star :: (scaleR) scaleR | 
| 
36a93808642c
instantiation replacing primitive instance plus overloaded defs
 haftmann parents: 
24075diff
changeset | 47 | begin | 
| 20555 | 48 | |
| 26566 
36a93808642c
instantiation replacing primitive instance plus overloaded defs
 haftmann parents: 
24075diff
changeset | 49 | definition | 
| 20555 | 50 | star_scaleR_def [transfer_unfold]: "scaleR r \<equiv> *f* (scaleR r)" | 
| 51 | ||
| 26566 
36a93808642c
instantiation replacing primitive instance plus overloaded defs
 haftmann parents: 
24075diff
changeset | 52 | instance .. | 
| 
36a93808642c
instantiation replacing primitive instance plus overloaded defs
 haftmann parents: 
24075diff
changeset | 53 | |
| 
36a93808642c
instantiation replacing primitive instance plus overloaded defs
 haftmann parents: 
24075diff
changeset | 54 | end | 
| 
36a93808642c
instantiation replacing primitive instance plus overloaded defs
 haftmann parents: 
24075diff
changeset | 55 | |
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 56 | lemma Standard_scaleR [simp]: "x \<in> Standard \<Longrightarrow> scaleR r x \<in> Standard" | 
| 20726 | 57 | by (simp add: star_scaleR_def) | 
| 58 | ||
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 59 | lemma star_of_scaleR [simp]: "star_of (scaleR r x) = scaleR r (star_of x)" | 
| 20726 | 60 | by transfer (rule refl) | 
| 61 | ||
| 20555 | 62 | instance star :: (real_vector) real_vector | 
| 63 | proof | |
| 64 | fix a b :: real | |
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 65 | show "\<And>x y::'a star. scaleR a (x + y) = scaleR a x + scaleR a y" | 
| 20555 | 66 | by transfer (rule scaleR_right_distrib) | 
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 67 | show "\<And>x::'a star. scaleR (a + b) x = scaleR a x + scaleR b x" | 
| 20555 | 68 | by transfer (rule scaleR_left_distrib) | 
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 69 | show "\<And>x::'a star. scaleR a (scaleR b x) = scaleR (a * b) x" | 
| 20765 | 70 | by transfer (rule scaleR_scaleR) | 
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 71 | show "\<And>x::'a star. scaleR 1 x = x" | 
| 20555 | 72 | by transfer (rule scaleR_one) | 
| 73 | qed | |
| 74 | ||
| 75 | instance star :: (real_algebra) real_algebra | |
| 76 | proof | |
| 77 | fix a :: real | |
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 78 | show "\<And>x y::'a star. scaleR a x * y = scaleR a (x * y)" | 
| 20555 | 79 | by transfer (rule mult_scaleR_left) | 
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 80 | show "\<And>x y::'a star. x * scaleR a y = scaleR a (x * y)" | 
| 20555 | 81 | by transfer (rule mult_scaleR_right) | 
| 82 | qed | |
| 83 | ||
| 84 | instance star :: (real_algebra_1) real_algebra_1 .. | |
| 85 | ||
| 20584 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20555diff
changeset | 86 | instance star :: (real_div_algebra) real_div_algebra .. | 
| 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20555diff
changeset | 87 | |
| 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20555diff
changeset | 88 | instance star :: (real_field) real_field .. | 
| 
60b1d52a455d
added classes real_div_algebra and real_field; added lemmas
 huffman parents: 
20555diff
changeset | 89 | |
| 20726 | 90 | lemma star_of_real_def [transfer_unfold]: "of_real r = star_of (of_real r)" | 
| 91 | by (unfold of_real_def, transfer, rule refl) | |
| 92 | ||
| 93 | lemma Standard_of_real [simp]: "of_real r \<in> Standard" | |
| 94 | by (simp add: star_of_real_def) | |
| 20555 | 95 | |
| 96 | lemma star_of_of_real [simp]: "star_of (of_real r) = of_real r" | |
| 97 | by transfer (rule refl) | |
| 98 | ||
| 20729 | 99 | lemma of_real_eq_star_of [simp]: "of_real = star_of" | 
| 100 | proof | |
| 101 | fix r :: real | |
| 102 | show "of_real r = star_of r" | |
| 103 | by transfer simp | |
| 104 | qed | |
| 105 | ||
| 106 | lemma Reals_eq_Standard: "(Reals :: hypreal set) = Standard" | |
| 20765 | 107 | by (simp add: Reals_def Standard_def) | 
| 20729 | 108 | |
| 20555 | 109 | |
| 22877 | 110 | subsection {* Injection from @{typ hypreal} *}
 | 
| 111 | ||
| 112 | definition | |
| 113 | of_hypreal :: "hypreal \<Rightarrow> 'a::real_algebra_1 star" where | |
| 114 | "of_hypreal = *f* of_real" | |
| 115 | ||
| 116 | declare of_hypreal_def [transfer_unfold] | |
| 117 | ||
| 118 | lemma Standard_of_hypreal [simp]: | |
| 119 | "r \<in> Standard \<Longrightarrow> of_hypreal r \<in> Standard" | |
| 120 | by (simp add: of_hypreal_def) | |
| 121 | ||
| 122 | lemma of_hypreal_0 [simp]: "of_hypreal 0 = 0" | |
| 123 | by transfer (rule of_real_0) | |
| 124 | ||
| 125 | lemma of_hypreal_1 [simp]: "of_hypreal 1 = 1" | |
| 126 | by transfer (rule of_real_1) | |
| 127 | ||
| 128 | lemma of_hypreal_add [simp]: | |
| 129 | "\<And>x y. of_hypreal (x + y) = of_hypreal x + of_hypreal y" | |
| 130 | by transfer (rule of_real_add) | |
| 131 | ||
| 132 | lemma of_hypreal_minus [simp]: "\<And>x. of_hypreal (- x) = - of_hypreal x" | |
| 133 | by transfer (rule of_real_minus) | |
| 134 | ||
| 135 | lemma of_hypreal_diff [simp]: | |
| 136 | "\<And>x y. of_hypreal (x - y) = of_hypreal x - of_hypreal y" | |
| 137 | by transfer (rule of_real_diff) | |
| 138 | ||
| 139 | lemma of_hypreal_mult [simp]: | |
| 140 | "\<And>x y. of_hypreal (x * y) = of_hypreal x * of_hypreal y" | |
| 141 | by transfer (rule of_real_mult) | |
| 142 | ||
| 143 | lemma of_hypreal_inverse [simp]: | |
| 144 | "\<And>x. of_hypreal (inverse x) = | |
| 145 |    inverse (of_hypreal x :: 'a::{real_div_algebra,division_by_zero} star)"
 | |
| 146 | by transfer (rule of_real_inverse) | |
| 147 | ||
| 148 | lemma of_hypreal_divide [simp]: | |
| 149 | "\<And>x y. of_hypreal (x / y) = | |
| 150 |    (of_hypreal x / of_hypreal y :: 'a::{real_field,division_by_zero} star)"
 | |
| 151 | by transfer (rule of_real_divide) | |
| 152 | ||
| 153 | lemma of_hypreal_eq_iff [simp]: | |
| 154 | "\<And>x y. (of_hypreal x = of_hypreal y) = (x = y)" | |
| 155 | by transfer (rule of_real_eq_iff) | |
| 156 | ||
| 157 | lemma of_hypreal_eq_0_iff [simp]: | |
| 158 | "\<And>x. (of_hypreal x = 0) = (x = 0)" | |
| 159 | by transfer (rule of_real_eq_0_iff) | |
| 160 | ||
| 161 | ||
| 17298 | 162 | subsection{*Properties of @{term starrel}*}
 | 
| 14329 | 163 | |
| 21787 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 164 | lemma lemma_starrel_refl [simp]: "x \<in> starrel `` {x}"
 | 
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 165 | by (simp add: starrel_def) | 
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 166 | |
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 167 | lemma starrel_in_hypreal [simp]: "starrel``{x}:star"
 | 
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 168 | by (simp add: star_def starrel_def quotient_def, blast) | 
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 169 | |
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 170 | declare Abs_star_inject [simp] Abs_star_inverse [simp] | 
| 
9edd495b6330
consistent naming for FreeUltrafilterNat lemmas; cleaned up
 huffman parents: 
21588diff
changeset | 171 | declare equiv_starrel [THEN eq_equiv_class_iff, simp] | 
| 14299 | 172 | |
| 14329 | 173 | subsection{*@{term hypreal_of_real}: 
 | 
| 174 |             the Injection from @{typ real} to @{typ hypreal}*}
 | |
| 14299 | 175 | |
| 21810 
b2d23672b003
generalized some lemmas; removed redundant lemmas; cleaned up some proofs
 huffman parents: 
21787diff
changeset | 176 | lemma inj_star_of: "inj star_of" | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 177 | by (rule inj_onI, simp) | 
| 14299 | 178 | |
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
20245diff
changeset | 179 | lemma mem_Rep_star_iff: "(X \<in> Rep_star x) = (x = star_n X)" | 
| 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
20245diff
changeset | 180 | by (cases x, simp add: star_n_def) | 
| 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
20245diff
changeset | 181 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 182 | lemma Rep_star_star_n_iff [simp]: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 183 |   "(X \<in> Rep_star (star_n Y)) = ({n. Y n = X n} \<in> \<U>)"
 | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 184 | by (simp add: star_n_def) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 185 | |
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 186 | lemma Rep_star_star_n: "X \<in> Rep_star (star_n X)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 187 | by simp | 
| 14329 | 188 | |
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 189 | subsection{* Properties of @{term star_n} *}
 | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 190 | |
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 191 | lemma star_n_add: | 
| 17298 | 192 | "star_n X + star_n Y = star_n (%n. X n + Y n)" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 193 | by (simp only: star_add_def starfun2_star_n) | 
| 14329 | 194 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 195 | lemma star_n_minus: | 
| 17298 | 196 | "- star_n X = star_n (%n. -(X n))" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 197 | by (simp only: star_minus_def starfun_star_n) | 
| 14299 | 198 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 199 | lemma star_n_diff: | 
| 17298 | 200 | "star_n X - star_n Y = star_n (%n. X n - Y n)" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 201 | by (simp only: star_diff_def starfun2_star_n) | 
| 14299 | 202 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 203 | lemma star_n_mult: | 
| 17298 | 204 | "star_n X * star_n Y = star_n (%n. X n * Y n)" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 205 | by (simp only: star_mult_def starfun2_star_n) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 206 | |
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 207 | lemma star_n_inverse: | 
| 17301 | 208 | "inverse (star_n X) = star_n (%n. inverse(X n))" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 209 | by (simp only: star_inverse_def starfun_star_n) | 
| 14299 | 210 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 211 | lemma star_n_le: | 
| 17298 | 212 | "star_n X \<le> star_n Y = | 
| 14365 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 paulson parents: 
14361diff
changeset | 213 |        ({n. X n \<le> Y n} \<in> FreeUltrafilterNat)"
 | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 214 | by (simp only: star_le_def starP2_star_n) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 215 | |
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 216 | lemma star_n_less: | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 217 |       "star_n X < star_n Y = ({n. X n < Y n} \<in> FreeUltrafilterNat)"
 | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 218 | by (simp only: star_less_def starP2_star_n) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 219 | |
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 220 | lemma star_n_zero_num: "0 = star_n (%n. 0)" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 221 | by (simp only: star_zero_def star_of_def) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 222 | |
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 223 | lemma star_n_one_num: "1 = star_n (%n. 1)" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 224 | by (simp only: star_one_def star_of_def) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 225 | |
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 226 | lemma star_n_abs: | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 227 | "abs (star_n X) = star_n (%n. abs (X n))" | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 228 | by (simp only: star_abs_def starfun_star_n) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 229 | |
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 230 | subsection{*Misc Others*}
 | 
| 14299 | 231 | |
| 14370 | 232 | lemma hypreal_not_refl2: "!!(x::hypreal). x < y ==> x \<noteq> y" | 
| 15539 | 233 | by (auto) | 
| 14329 | 234 | |
| 14331 | 235 | lemma hypreal_eq_minus_iff: "((x::hypreal) = y) = (x + - y = 0)" | 
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15229diff
changeset | 236 | by auto | 
| 14331 | 237 | |
| 14329 | 238 | lemma hypreal_mult_left_cancel: "(c::hypreal) \<noteq> 0 ==> (c*a=c*b) = (a=b)" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 239 | by auto | 
| 14329 | 240 | |
| 241 | lemma hypreal_mult_right_cancel: "(c::hypreal) \<noteq> 0 ==> (a*c=b*c) = (a=b)" | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 242 | by auto | 
| 14329 | 243 | |
| 14301 | 244 | lemma hypreal_omega_gt_zero [simp]: "0 < omega" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 245 | by (simp add: omega_def star_n_zero_num star_n_less) | 
| 14370 | 246 | |
| 247 | subsection{*Existence of Infinite Hyperreal Number*}
 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 248 | |
| 14370 | 249 | text{*Existence of infinite number not corresponding to any real number.
 | 
| 250 | Use assumption that member @{term FreeUltrafilterNat} is not finite.*}
 | |
| 251 | ||
| 252 | ||
| 253 | text{*A few lemmas first*}
 | |
| 254 | ||
| 255 | lemma lemma_omega_empty_singleton_disj: "{n::nat. x = real n} = {} |  
 | |
| 256 |       (\<exists>y. {n::nat. x = real n} = {y})"
 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 257 | by force | 
| 14370 | 258 | |
| 259 | lemma lemma_finite_omega_set: "finite {n::nat. x = real n}"
 | |
| 260 | by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto) | |
| 261 | ||
| 262 | lemma not_ex_hypreal_of_real_eq_omega: | |
| 263 | "~ (\<exists>x. hypreal_of_real x = omega)" | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 264 | apply (simp add: omega_def) | 
| 17298 | 265 | apply (simp add: star_of_def star_n_eq_iff) | 
| 14370 | 266 | apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric] | 
| 21855 
74909ecaf20a
remove ultra tactic and redundant FreeUltrafilterNat lemmas
 huffman parents: 
21810diff
changeset | 267 | lemma_finite_omega_set [THEN FreeUltrafilterNat.finite]) | 
| 14370 | 268 | done | 
| 269 | ||
| 270 | lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \<noteq> omega" | |
| 14705 | 271 | by (insert not_ex_hypreal_of_real_eq_omega, auto) | 
| 14370 | 272 | |
| 273 | text{*Existence of infinitesimal number also not corresponding to any
 | |
| 274 | real number*} | |
| 275 | ||
| 276 | lemma lemma_epsilon_empty_singleton_disj: | |
| 277 |      "{n::nat. x = inverse(real(Suc n))} = {} |  
 | |
| 278 |       (\<exists>y. {n::nat. x = inverse(real(Suc n))} = {y})"
 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 279 | by auto | 
| 14370 | 280 | |
| 281 | lemma lemma_finite_epsilon_set: "finite {n. x = inverse(real(Suc n))}"
 | |
| 282 | by (cut_tac x = x in lemma_epsilon_empty_singleton_disj, auto) | |
| 283 | ||
| 14705 | 284 | lemma not_ex_hypreal_of_real_eq_epsilon: "~ (\<exists>x. hypreal_of_real x = epsilon)" | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17301diff
changeset | 285 | by (auto simp add: epsilon_def star_of_def star_n_eq_iff | 
| 21855 
74909ecaf20a
remove ultra tactic and redundant FreeUltrafilterNat lemmas
 huffman parents: 
21810diff
changeset | 286 | lemma_finite_epsilon_set [THEN FreeUltrafilterNat.finite]) | 
| 14370 | 287 | |
| 288 | lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \<noteq> epsilon" | |
| 14705 | 289 | by (insert not_ex_hypreal_of_real_eq_epsilon, auto) | 
| 14370 | 290 | |
| 291 | lemma hypreal_epsilon_not_zero: "epsilon \<noteq> 0" | |
| 17298 | 292 | by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff | 
| 293 | del: star_of_zero) | |
| 14370 | 294 | |
| 295 | lemma hypreal_epsilon_inverse_omega: "epsilon = inverse(omega)" | |
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 296 | by (simp add: epsilon_def omega_def star_n_inverse) | 
| 14370 | 297 | |
| 20753 | 298 | lemma hypreal_epsilon_gt_zero: "0 < epsilon" | 
| 299 | by (simp add: hypreal_epsilon_inverse_omega) | |
| 300 | ||
| 21865 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 301 | subsection{*Absolute Value Function for the Hyperreals*}
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 302 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 303 | lemma hrabs_add_less: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 304 | "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 305 | by (simp add: abs_if split: split_if_asm) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 306 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 307 | lemma hrabs_less_gt_zero: "abs x < r ==> (0::hypreal) < r" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 308 | by (blast intro!: order_le_less_trans abs_ge_zero) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 309 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 310 | lemma hrabs_disj: "abs x = (x::'a::abs_if) | abs x = -x" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 311 | by (simp add: abs_if) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 312 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 313 | lemma hrabs_add_lemma_disj: "(y::hypreal) + - x + (y + - z) = abs (x + - z) ==> y = z | x = y" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 314 | by (simp add: abs_if split add: split_if_asm) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 315 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 316 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 317 | subsection{*Embedding the Naturals into the Hyperreals*}
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 318 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 319 | abbreviation | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 320 | hypreal_of_nat :: "nat => hypreal" where | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 321 | "hypreal_of_nat == of_nat" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 322 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 323 | lemma SNat_eq: "Nats = {n. \<exists>N. n = hypreal_of_nat N}"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 324 | by (simp add: Nats_def image_def) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 325 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 326 | (*------------------------------------------------------------*) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 327 | (* naturals embedded in hyperreals *) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 328 | (* is a hyperreal c.f. NS extension *) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 329 | (*------------------------------------------------------------*) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 330 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 331 | lemma hypreal_of_nat_eq: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 332 | "hypreal_of_nat (n::nat) = hypreal_of_real (real n)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 333 | by (simp add: real_of_nat_def) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 334 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 335 | lemma hypreal_of_nat: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 336 | "hypreal_of_nat m = star_n (%n. real m)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 337 | apply (fold star_of_def) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 338 | apply (simp add: real_of_nat_def) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 339 | done | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 340 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 341 | (* | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 342 | FIXME: we should declare this, as for type int, but many proofs would break. | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 343 | It replaces x+-y by x-y. | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 344 | Addsimps [symmetric hypreal_diff_def] | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 345 | *) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 346 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 347 | use "hypreal_arith.ML" | 
| 24075 | 348 | declaration {* K hypreal_arith_setup *}
 | 
| 21865 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 349 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 350 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 351 | subsection {* Exponentials on the Hyperreals *}
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 352 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 353 | lemma hpowr_0 [simp]: "r ^ 0 = (1::hypreal)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 354 | by (rule power_0) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 355 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 356 | lemma hpowr_Suc [simp]: "r ^ (Suc n) = (r::hypreal) * (r ^ n)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 357 | by (rule power_Suc) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 358 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 359 | lemma hrealpow_two: "(r::hypreal) ^ Suc (Suc 0) = r * r" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 360 | by simp | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 361 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 362 | lemma hrealpow_two_le [simp]: "(0::hypreal) \<le> r ^ Suc (Suc 0)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 363 | by (auto simp add: zero_le_mult_iff) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 364 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 365 | lemma hrealpow_two_le_add_order [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 366 | "(0::hypreal) \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 367 | by (simp only: hrealpow_two_le add_nonneg_nonneg) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 368 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 369 | lemma hrealpow_two_le_add_order2 [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 370 | "(0::hypreal) \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 371 | by (simp only: hrealpow_two_le add_nonneg_nonneg) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 372 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 373 | lemma hypreal_add_nonneg_eq_0_iff: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 374 | "[| 0 \<le> x; 0 \<le> y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 375 | by arith | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 376 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 377 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 378 | text{*FIXME: DELETE THESE*}
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 379 | lemma hypreal_three_squares_add_zero_iff: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 380 | "(x*x + y*y + z*z = 0) = (x = 0 & y = 0 & z = (0::hypreal))" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 381 | apply (simp only: zero_le_square add_nonneg_nonneg hypreal_add_nonneg_eq_0_iff, auto) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 382 | done | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 383 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 384 | lemma hrealpow_three_squares_add_zero_iff [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 385 | "(x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + z ^ Suc (Suc 0) = (0::hypreal)) = | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 386 | (x = 0 & y = 0 & z = 0)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 387 | by (simp only: hypreal_three_squares_add_zero_iff hrealpow_two) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 388 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 389 | (*FIXME: This and RealPow.abs_realpow_two should be replaced by an abstract | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 390 | result proved in Ring_and_Field*) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 391 | lemma hrabs_hrealpow_two [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 392 | "abs(x ^ Suc (Suc 0)) = (x::hypreal) ^ Suc (Suc 0)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 393 | by (simp add: abs_mult) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 394 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 395 | lemma two_hrealpow_ge_one [simp]: "(1::hypreal) \<le> 2 ^ n" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 396 | by (insert power_increasing [of 0 n "2::hypreal"], simp) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 397 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 398 | lemma two_hrealpow_gt [simp]: "hypreal_of_nat n < 2 ^ n" | 
| 27105 
5f139027c365
slightly tuning of some proofs involving case distinction and induction on natural numbers and similar
 haftmann parents: 
26566diff
changeset | 399 | apply (induct n) | 
| 21865 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 400 | apply (auto simp add: left_distrib) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 401 | apply (cut_tac n = n in two_hrealpow_ge_one, arith) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 402 | done | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 403 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 404 | lemma hrealpow: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 405 | "star_n X ^ m = star_n (%n. (X n::real) ^ m)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 406 | apply (induct_tac "m") | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 407 | apply (auto simp add: star_n_one_num star_n_mult power_0) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 408 | done | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 409 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 410 | lemma hrealpow_sum_square_expand: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 411 | "(x + (y::hypreal)) ^ Suc (Suc 0) = | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 412 | x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + (hypreal_of_nat (Suc (Suc 0)))*x*y" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 413 | by (simp add: right_distrib left_distrib) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 414 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 415 | lemma power_hypreal_of_real_number_of: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 416 | "(number_of v :: hypreal) ^ n = hypreal_of_real ((number_of v) ^ n)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 417 | by simp | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 418 | declare power_hypreal_of_real_number_of [of _ "number_of w", standard, simp] | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 419 | (* | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 420 | lemma hrealpow_HFinite: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 421 |   fixes x :: "'a::{real_normed_algebra,recpower} star"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 422 | shows "x \<in> HFinite ==> x ^ n \<in> HFinite" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 423 | apply (induct_tac "n") | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 424 | apply (auto simp add: power_Suc intro: HFinite_mult) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 425 | done | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 426 | *) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 427 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 428 | subsection{*Powers with Hypernatural Exponents*}
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 429 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 430 | definition | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 431 | (* hypernatural powers of hyperreals *) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 432 | pow :: "['a::power star, nat star] \<Rightarrow> 'a star" (infixr "pow" 80) where | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 433 | hyperpow_def [transfer_unfold]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 434 | "R pow N = ( *f2* op ^) R N" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 435 | |
| 22879 | 436 | lemma Standard_hyperpow [simp]: | 
| 437 | "\<lbrakk>r \<in> Standard; n \<in> Standard\<rbrakk> \<Longrightarrow> r pow n \<in> Standard" | |
| 438 | unfolding hyperpow_def by simp | |
| 439 | ||
| 21865 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 440 | lemma hyperpow: "star_n X pow star_n Y = star_n (%n. X n ^ Y n)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 441 | by (simp add: hyperpow_def starfun2_star_n) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 442 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 443 | lemma hyperpow_zero [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 444 |   "\<And>n. (0::'a::{recpower,semiring_0} star) pow (n + (1::hypnat)) = 0"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 445 | by transfer simp | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 446 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 447 | lemma hyperpow_not_zero: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 448 |   "\<And>r n. r \<noteq> (0::'a::{recpower,field} star) ==> r pow n \<noteq> 0"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 449 | by transfer (rule field_power_not_zero) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 450 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 451 | lemma hyperpow_inverse: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 452 |   "\<And>r n. r \<noteq> (0::'a::{recpower,division_by_zero,field} star)
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 453 | \<Longrightarrow> inverse (r pow n) = (inverse r) pow n" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 454 | by transfer (rule power_inverse) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 455 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 456 | lemma hyperpow_hrabs: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 457 |   "\<And>r n. abs (r::'a::{recpower,ordered_idom} star) pow n = abs (r pow n)"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 458 | by transfer (rule power_abs [symmetric]) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 459 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 460 | lemma hyperpow_add: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 461 | "\<And>r n m. (r::'a::recpower star) pow (n + m) = (r pow n) * (r pow m)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 462 | by transfer (rule power_add) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 463 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 464 | lemma hyperpow_one [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 465 | "\<And>r. (r::'a::recpower star) pow (1::hypnat) = r" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 466 | by transfer (rule power_one_right) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 467 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 468 | lemma hyperpow_two: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 469 | "\<And>r. (r::'a::recpower star) pow ((1::hypnat) + (1::hypnat)) = r * r" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 470 | by transfer (simp add: power_Suc) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 471 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 472 | lemma hyperpow_gt_zero: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 473 |   "\<And>r n. (0::'a::{recpower,ordered_semidom} star) < r \<Longrightarrow> 0 < r pow n"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 474 | by transfer (rule zero_less_power) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 475 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 476 | lemma hyperpow_ge_zero: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 477 |   "\<And>r n. (0::'a::{recpower,ordered_semidom} star) \<le> r \<Longrightarrow> 0 \<le> r pow n"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 478 | by transfer (rule zero_le_power) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 479 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 480 | lemma hyperpow_le: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 481 |   "\<And>x y n. \<lbrakk>(0::'a::{recpower,ordered_semidom} star) < x; x \<le> y\<rbrakk>
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 482 | \<Longrightarrow> x pow n \<le> y pow n" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 483 | by transfer (rule power_mono [OF _ order_less_imp_le]) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 484 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 485 | lemma hyperpow_eq_one [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 486 | "\<And>n. 1 pow n = (1::'a::recpower star)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 487 | by transfer (rule power_one) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 488 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 489 | lemma hrabs_hyperpow_minus_one [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 490 |   "\<And>n. abs(-1 pow n) = (1::'a::{number_ring,recpower,ordered_idom} star)"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 491 | by transfer (rule abs_power_minus_one) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 492 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 493 | lemma hyperpow_mult: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 494 |   "\<And>r s n. (r * s::'a::{comm_monoid_mult,recpower} star) pow n
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 495 | = (r pow n) * (s pow n)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 496 | by transfer (rule power_mult_distrib) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 497 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 498 | lemma hyperpow_two_le [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 499 |   "(0::'a::{recpower,ordered_ring_strict} star) \<le> r pow (1 + 1)"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 500 | by (auto simp add: hyperpow_two zero_le_mult_iff) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 501 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 502 | lemma hrabs_hyperpow_two [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 503 | "abs(x pow (1 + 1)) = | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 504 |    (x::'a::{recpower,ordered_ring_strict} star) pow (1 + 1)"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 505 | by (simp only: abs_of_nonneg hyperpow_two_le) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 506 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 507 | lemma hyperpow_two_hrabs [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 508 |   "abs(x::'a::{recpower,ordered_idom} star) pow (1 + 1)  = x pow (1 + 1)"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 509 | by (simp add: hyperpow_hrabs) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 510 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 511 | text{*The precondition could be weakened to @{term "0\<le>x"}*}
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 512 | lemma hypreal_mult_less_mono: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 513 | "[| u<v; x<y; (0::hypreal) < v; 0 < x |] ==> u*x < v* y" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 514 | by (simp add: Ring_and_Field.mult_strict_mono order_less_imp_le) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 515 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 516 | lemma hyperpow_two_gt_one: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 517 |   "\<And>r::'a::{recpower,ordered_semidom} star. 1 < r \<Longrightarrow> 1 < r pow (1 + 1)"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 518 | by transfer (simp add: power_gt1) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 519 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 520 | lemma hyperpow_two_ge_one: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 521 |   "\<And>r::'a::{recpower,ordered_semidom} star. 1 \<le> r \<Longrightarrow> 1 \<le> r pow (1 + 1)"
 | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 522 | by transfer (simp add: one_le_power) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 523 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 524 | lemma two_hyperpow_ge_one [simp]: "(1::hypreal) \<le> 2 pow n" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 525 | apply (rule_tac y = "1 pow n" in order_trans) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 526 | apply (rule_tac [2] hyperpow_le, auto) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 527 | done | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 528 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 529 | lemma hyperpow_minus_one2 [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 530 | "!!n. -1 pow ((1 + 1)*n) = (1::hypreal)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 531 | by transfer (subst power_mult, simp) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 532 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 533 | lemma hyperpow_less_le: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 534 | "!!r n N. [|(0::hypreal) \<le> r; r \<le> 1; n < N|] ==> r pow N \<le> r pow n" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 535 | by transfer (rule power_decreasing [OF order_less_imp_le]) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 536 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 537 | lemma hyperpow_SHNat_le: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 538 | "[| 0 \<le> r; r \<le> (1::hypreal); N \<in> HNatInfinite |] | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 539 | ==> ALL n: Nats. r pow N \<le> r pow n" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 540 | by (auto intro!: hyperpow_less_le simp add: HNatInfinite_iff) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 541 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 542 | lemma hyperpow_realpow: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 543 | "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 544 | by transfer (rule refl) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 545 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 546 | lemma hyperpow_SReal [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 547 | "(hypreal_of_real r) pow (hypnat_of_nat n) \<in> Reals" | 
| 22879 | 548 | by (simp add: Reals_eq_Standard) | 
| 21865 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 549 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 550 | lemma hyperpow_zero_HNatInfinite [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 551 | "N \<in> HNatInfinite ==> (0::hypreal) pow N = 0" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 552 | by (drule HNatInfinite_is_Suc, auto) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 553 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 554 | lemma hyperpow_le_le: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 555 | "[| (0::hypreal) \<le> r; r \<le> 1; n \<le> N |] ==> r pow N \<le> r pow n" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 556 | apply (drule order_le_less [of n, THEN iffD1]) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 557 | apply (auto intro: hyperpow_less_le) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 558 | done | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 559 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 560 | lemma hyperpow_Suc_le_self2: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 561 | "[| (0::hypreal) \<le> r; r < 1 |] ==> r pow (n + (1::hypnat)) \<le> r" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 562 | apply (drule_tac n = " (1::hypnat) " in hyperpow_le_le) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 563 | apply auto | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 564 | done | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 565 | |
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 566 | lemma hyperpow_hypnat_of_nat: "\<And>x. x pow hypnat_of_nat n = x ^ n" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 567 | by transfer (rule refl) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21856diff
changeset | 568 | |
| 22888 | 569 | lemma of_hypreal_hyperpow: | 
| 570 | "\<And>x n. of_hypreal (x pow n) = | |
| 571 |    (of_hypreal x::'a::{real_algebra_1,recpower} star) pow n"
 | |
| 572 | by transfer (rule of_real_power) | |
| 573 | ||
| 10751 | 574 | end |