src/HOL/NSA/HyperDef.thy
author hoelzl
Sun, 12 Apr 2015 11:34:16 +0200
changeset 60041 6c86d58ab0ca
parent 59867 58043346ca64
child 60867 86e7560e07d0
permissions -rw-r--r--
replace Filters in NSA by HOL-Filters
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: 28562
diff changeset
     1
(*  Title       : HOL/NSA/HyperDef.thy
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     4
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     5
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     6
58878
f962e42e324d modernized header;
wenzelm
parents: 58410
diff changeset
     7
section{*Construction of Hyperreals Using Ultrafilters*}
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     8
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     9
theory HyperDef
51525
d3d170a2887f HOL-NSA should only import Complex_Main
hoelzl
parents: 49962
diff changeset
    10
imports Complex_Main HyperNat
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    11
begin
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    12
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 38715
diff changeset
    13
type_synonym hypreal = "real star"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    14
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    15
abbreviation
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    16
  hypreal_of_real :: "real => real star" where
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    17
  "hypreal_of_real == star_of"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    18
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    19
abbreviation
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    20
  hypreal_of_hypnat :: "hypnat \<Rightarrow> hypreal" where
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    21
  "hypreal_of_hypnat \<equiv> of_hypnat"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    22
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    23
definition
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    24
  omega :: hypreal where
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    25
   -- {*an infinite number @{text "= [<1,2,3,...>]"} *}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    26
  "omega = star_n (\<lambda>n. real (Suc n))"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    27
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    28
definition
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    29
  epsilon :: hypreal where
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    30
   -- {*an infinitesimal number @{text "= [<1,1/2,1/3,...>]"} *}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    31
  "epsilon = star_n (\<lambda>n. inverse (real (Suc n)))"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    32
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    33
notation (xsymbols)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    34
  omega  ("\<omega>") and
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    35
  epsilon  ("\<epsilon>")
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    36
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    37
notation (HTML output)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    38
  omega  ("\<omega>") and
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    39
  epsilon  ("\<epsilon>")
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    40
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    41
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    42
subsection {* Real vector class instances *}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    43
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    44
instantiation star :: (scaleR) scaleR
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    45
begin
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    46
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    47
definition
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 36409
diff changeset
    48
  star_scaleR_def [transfer_unfold]: "scaleR r \<equiv> *f* (scaleR r)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    49
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    50
instance ..
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    51
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    52
end
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    53
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    54
lemma Standard_scaleR [simp]: "x \<in> Standard \<Longrightarrow> scaleR r x \<in> Standard"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    55
by (simp add: star_scaleR_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    56
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    57
lemma star_of_scaleR [simp]: "star_of (scaleR r x) = scaleR r (star_of x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    58
by transfer (rule refl)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    59
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    60
instance star :: (real_vector) real_vector
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    61
proof
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    62
  fix a b :: real
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    63
  show "\<And>x y::'a star. scaleR a (x + y) = scaleR a x + scaleR a y"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    64
    by transfer (rule scaleR_right_distrib)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    65
  show "\<And>x::'a star. scaleR (a + b) x = scaleR a x + scaleR b x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    66
    by transfer (rule scaleR_left_distrib)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    67
  show "\<And>x::'a star. scaleR a (scaleR b x) = scaleR (a * b) x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    68
    by transfer (rule scaleR_scaleR)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    69
  show "\<And>x::'a star. scaleR 1 x = x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    70
    by transfer (rule scaleR_one)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    71
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    72
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    73
instance star :: (real_algebra) real_algebra
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    74
proof
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    75
  fix a :: real
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    76
  show "\<And>x y::'a star. scaleR a x * y = scaleR a (x * y)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    77
    by transfer (rule mult_scaleR_left)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    78
  show "\<And>x y::'a star. x * scaleR a y = scaleR a (x * y)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    79
    by transfer (rule mult_scaleR_right)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    80
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    81
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    82
instance star :: (real_algebra_1) real_algebra_1 ..
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    83
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    84
instance star :: (real_div_algebra) real_div_algebra ..
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    85
27553
d315a513a150 instance real_field < field_char_0;
huffman
parents: 27468
diff changeset
    86
instance star :: (field_char_0) field_char_0 ..
d315a513a150 instance real_field < field_char_0;
huffman
parents: 27468
diff changeset
    87
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    88
instance star :: (real_field) real_field ..
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    89
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    90
lemma star_of_real_def [transfer_unfold]: "of_real r = star_of (of_real r)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    91
by (unfold of_real_def, transfer, rule refl)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    92
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    93
lemma Standard_of_real [simp]: "of_real r \<in> Standard"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    94
by (simp add: star_of_real_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    95
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    96
lemma star_of_of_real [simp]: "star_of (of_real r) = of_real r"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    97
by transfer (rule refl)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    98
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    99
lemma of_real_eq_star_of [simp]: "of_real = star_of"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   100
proof
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   101
  fix r :: real
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   102
  show "of_real r = star_of r"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   103
    by transfer simp
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   104
qed
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   105
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   106
lemma Reals_eq_Standard: "(Reals :: hypreal set) = Standard"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   107
by (simp add: Reals_def Standard_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   108
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   109
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   110
subsection {* Injection from @{typ hypreal} *}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   111
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   112
definition
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   113
  of_hypreal :: "hypreal \<Rightarrow> 'a::real_algebra_1 star" where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 36409
diff changeset
   114
  [transfer_unfold]: "of_hypreal = *f* of_real"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   115
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   116
lemma Standard_of_hypreal [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   117
  "r \<in> Standard \<Longrightarrow> of_hypreal r \<in> Standard"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   118
by (simp add: of_hypreal_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   119
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   120
lemma of_hypreal_0 [simp]: "of_hypreal 0 = 0"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   121
by transfer (rule of_real_0)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   122
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   123
lemma of_hypreal_1 [simp]: "of_hypreal 1 = 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   124
by transfer (rule of_real_1)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   125
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   126
lemma of_hypreal_add [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   127
  "\<And>x y. of_hypreal (x + y) = of_hypreal x + of_hypreal y"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   128
by transfer (rule of_real_add)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   129
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   130
lemma of_hypreal_minus [simp]: "\<And>x. of_hypreal (- x) = - of_hypreal x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   131
by transfer (rule of_real_minus)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   132
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   133
lemma of_hypreal_diff [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   134
  "\<And>x y. of_hypreal (x - y) = of_hypreal x - of_hypreal y"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   135
by transfer (rule of_real_diff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   136
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   137
lemma of_hypreal_mult [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   138
  "\<And>x y. of_hypreal (x * y) = of_hypreal x * of_hypreal y"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   139
by transfer (rule of_real_mult)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   140
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   141
lemma of_hypreal_inverse [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   142
  "\<And>x. of_hypreal (inverse x) =
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59582
diff changeset
   143
   inverse (of_hypreal x :: 'a::{real_div_algebra, division_ring} star)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   144
by transfer (rule of_real_inverse)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   145
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   146
lemma of_hypreal_divide [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   147
  "\<And>x y. of_hypreal (x / y) =
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59582
diff changeset
   148
   (of_hypreal x / of_hypreal y :: 'a::{real_field, field} star)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   149
by transfer (rule of_real_divide)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   150
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   151
lemma of_hypreal_eq_iff [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   152
  "\<And>x y. (of_hypreal x = of_hypreal y) = (x = y)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   153
by transfer (rule of_real_eq_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   154
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   155
lemma of_hypreal_eq_0_iff [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   156
  "\<And>x. (of_hypreal x = 0) = (x = 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   157
by transfer (rule of_real_eq_0_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   158
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   159
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   160
subsection{*Properties of @{term starrel}*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   161
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   162
lemma lemma_starrel_refl [simp]: "x \<in> starrel `` {x}"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   163
by (simp add: starrel_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   164
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   165
lemma starrel_in_hypreal [simp]: "starrel``{x}:star"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   166
by (simp add: star_def starrel_def quotient_def, blast)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   167
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   168
declare Abs_star_inject [simp] Abs_star_inverse [simp]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   169
declare equiv_starrel [THEN eq_equiv_class_iff, simp]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   170
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55911
diff changeset
   171
subsection{*@{term hypreal_of_real}:
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   172
            the Injection from @{typ real} to @{typ hypreal}*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   173
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   174
lemma inj_star_of: "inj star_of"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   175
by (rule inj_onI, simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   176
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   177
lemma mem_Rep_star_iff: "(X \<in> Rep_star x) = (x = star_n X)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   178
by (cases x, simp add: star_n_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   179
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   180
lemma Rep_star_star_n_iff [simp]:
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 59867
diff changeset
   181
  "(X \<in> Rep_star (star_n Y)) = (eventually (\<lambda>n. Y n = X n) \<U>)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   182
by (simp add: star_n_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   183
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   184
lemma Rep_star_star_n: "X \<in> Rep_star (star_n X)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   185
by simp
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   186
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   187
subsection{* Properties of @{term star_n} *}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   188
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   189
lemma star_n_add:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   190
  "star_n X + star_n Y = star_n (%n. X n + Y n)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   191
by (simp only: star_add_def starfun2_star_n)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   192
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   193
lemma star_n_minus:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   194
   "- star_n X = star_n (%n. -(X n))"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   195
by (simp only: star_minus_def starfun_star_n)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   196
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   197
lemma star_n_diff:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   198
     "star_n X - star_n Y = star_n (%n. X n - Y n)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   199
by (simp only: star_diff_def starfun2_star_n)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   200
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   201
lemma star_n_mult:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   202
  "star_n X * star_n Y = star_n (%n. X n * Y n)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   203
by (simp only: star_mult_def starfun2_star_n)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   204
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   205
lemma star_n_inverse:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   206
      "inverse (star_n X) = star_n (%n. inverse(X n))"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   207
by (simp only: star_inverse_def starfun_star_n)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   208
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   209
lemma star_n_le:
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 59867
diff changeset
   210
      "star_n X \<le> star_n Y = (eventually (\<lambda>n. X n \<le> Y n) FreeUltrafilterNat)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   211
by (simp only: star_le_def starP2_star_n)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   212
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   213
lemma star_n_less:
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 59867
diff changeset
   214
      "star_n X < star_n Y = (eventually (\<lambda>n. X n < Y n) FreeUltrafilterNat)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   215
by (simp only: star_less_def starP2_star_n)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   216
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   217
lemma star_n_zero_num: "0 = star_n (%n. 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   218
by (simp only: star_zero_def star_of_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   219
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   220
lemma star_n_one_num: "1 = star_n (%n. 1)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   221
by (simp only: star_one_def star_of_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   222
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   223
lemma star_n_abs:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   224
     "abs (star_n X) = star_n (%n. abs (X n))"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   225
by (simp only: star_abs_def starfun_star_n)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   226
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   227
lemma hypreal_omega_gt_zero [simp]: "0 < omega"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   228
by (simp add: omega_def star_n_zero_num star_n_less)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   229
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   230
subsection{*Existence of Infinite Hyperreal Number*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   231
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   232
text{*Existence of infinite number not corresponding to any real number.
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   233
Use assumption that member @{term FreeUltrafilterNat} is not finite.*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   234
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   235
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   236
text{*A few lemmas first*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   237
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55911
diff changeset
   238
lemma lemma_omega_empty_singleton_disj: "{n::nat. x = real n} = {} |
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   239
      (\<exists>y. {n::nat. x = real n} = {y})"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   240
by force
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   241
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   242
lemma lemma_finite_omega_set: "finite {n::nat. x = real n}"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   243
by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   244
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55911
diff changeset
   245
lemma not_ex_hypreal_of_real_eq_omega:
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   246
      "~ (\<exists>x. hypreal_of_real x = omega)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   247
apply (simp add: omega_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   248
apply (simp add: star_of_def star_n_eq_iff)
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55911
diff changeset
   249
apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric]
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   250
            lemma_finite_omega_set [THEN FreeUltrafilterNat.finite])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   251
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   252
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   253
lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \<noteq> omega"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   254
by (insert not_ex_hypreal_of_real_eq_omega, auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   255
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   256
text{*Existence of infinitesimal number also not corresponding to any
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   257
 real number*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   258
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   259
lemma lemma_epsilon_empty_singleton_disj:
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55911
diff changeset
   260
     "{n::nat. x = inverse(real(Suc n))} = {} |
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   261
      (\<exists>y. {n::nat. x = inverse(real(Suc n))} = {y})"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   262
by auto
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   263
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   264
lemma lemma_finite_epsilon_set: "finite {n. x = inverse(real(Suc n))}"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   265
by (cut_tac x = x in lemma_epsilon_empty_singleton_disj, auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   266
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   267
lemma not_ex_hypreal_of_real_eq_epsilon: "~ (\<exists>x. hypreal_of_real x = epsilon)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   268
by (auto simp add: epsilon_def star_of_def star_n_eq_iff
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   269
                   lemma_finite_epsilon_set [THEN FreeUltrafilterNat.finite])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   270
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   271
lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \<noteq> epsilon"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   272
by (insert not_ex_hypreal_of_real_eq_epsilon, auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   273
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   274
lemma hypreal_epsilon_not_zero: "epsilon \<noteq> 0"
60041
6c86d58ab0ca replace Filters in NSA by HOL-Filters
hoelzl
parents: 59867
diff changeset
   275
by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff FreeUltrafilterNat.proper
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   276
         del: star_of_zero)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   277
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   278
lemma hypreal_epsilon_inverse_omega: "epsilon = inverse(omega)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   279
by (simp add: epsilon_def omega_def star_n_inverse)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   280
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   281
lemma hypreal_epsilon_gt_zero: "0 < epsilon"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   282
by (simp add: hypreal_epsilon_inverse_omega)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   283
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   284
subsection{*Absolute Value Function for the Hyperreals*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   285
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   286
lemma hrabs_add_less:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   287
     "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   288
by (simp add: abs_if split: split_if_asm)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   289
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   290
lemma hrabs_less_gt_zero: "abs x < r ==> (0::hypreal) < r"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   291
by (blast intro!: order_le_less_trans abs_ge_zero)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   292
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   293
lemma hrabs_disj: "abs x = (x::'a::abs_if) | abs x = -x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   294
by (simp add: abs_if)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   295
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   296
lemma hrabs_add_lemma_disj: "(y::hypreal) + - x + (y + - z) = abs (x + - z) ==> y = z | x = y"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   297
by (simp add: abs_if split add: split_if_asm)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   298
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   299
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   300
subsection{*Embedding the Naturals into the Hyperreals*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   301
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   302
abbreviation
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   303
  hypreal_of_nat :: "nat => hypreal" where
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   304
  "hypreal_of_nat == of_nat"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   305
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   306
lemma SNat_eq: "Nats = {n. \<exists>N. n = hypreal_of_nat N}"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   307
by (simp add: Nats_def image_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   308
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   309
(*------------------------------------------------------------*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   310
(* naturals embedded in hyperreals                            *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   311
(* is a hyperreal c.f. NS extension                           *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   312
(*------------------------------------------------------------*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   313
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   314
lemma hypreal_of_nat_eq:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   315
     "hypreal_of_nat (n::nat) = hypreal_of_real (real n)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   316
by (simp add: real_of_nat_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   317
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   318
lemma hypreal_of_nat:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   319
     "hypreal_of_nat m = star_n (%n. real m)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   320
apply (fold star_of_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   321
apply (simp add: real_of_nat_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   322
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   323
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   324
(*
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   325
FIXME: we should declare this, as for type int, but many proofs would break.
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   326
It replaces x+-y by x-y.
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   327
Addsimps [symmetric hypreal_diff_def]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   328
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   329
31100
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31017
diff changeset
   330
declaration {*
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31017
diff changeset
   331
  K (Lin_Arith.add_inj_thms [@{thm star_of_le} RS iffD2,
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31017
diff changeset
   332
    @{thm star_of_less} RS iffD2, @{thm star_of_eq} RS iffD2]
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31017
diff changeset
   333
  #> Lin_Arith.add_simps [@{thm star_of_zero}, @{thm star_of_one},
55911
d00023bd3554 remove simp rules made redundant by the replacement of neg_numeral with negated numerals
huffman
parents: 54489
diff changeset
   334
      @{thm star_of_numeral}, @{thm star_of_add},
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45605
diff changeset
   335
      @{thm star_of_minus}, @{thm star_of_diff}, @{thm star_of_mult}]
43595
7ae4a23b5be6 modernized some simproc setup;
wenzelm
parents: 42463
diff changeset
   336
  #> Lin_Arith.add_inj_const (@{const_name "StarDef.star_of"}, @{typ "real \<Rightarrow> hypreal"}))
31100
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31017
diff changeset
   337
*}
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   338
43595
7ae4a23b5be6 modernized some simproc setup;
wenzelm
parents: 42463
diff changeset
   339
simproc_setup fast_arith_hypreal ("(m::hypreal) < n" | "(m::hypreal) <= n" | "(m::hypreal) = n") =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 58878
diff changeset
   340
  {* fn _ => fn ss => fn ct => Lin_Arith.simproc ss (Thm.term_of ct) *}
43595
7ae4a23b5be6 modernized some simproc setup;
wenzelm
parents: 42463
diff changeset
   341
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   342
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   343
subsection {* Exponentials on the Hyperreals *}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   344
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   345
lemma hpowr_0 [simp]:   "r ^ 0       = (1::hypreal)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   346
by (rule power_0)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   347
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   348
lemma hpowr_Suc [simp]: "r ^ (Suc n) = (r::hypreal) * (r ^ n)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   349
by (rule power_Suc)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   350
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   351
lemma hrealpow_two: "(r::hypreal) ^ Suc (Suc 0) = r * r"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   352
by simp
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   353
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   354
lemma hrealpow_two_le [simp]: "(0::hypreal) \<le> r ^ Suc (Suc 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   355
by (auto simp add: zero_le_mult_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   356
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   357
lemma hrealpow_two_le_add_order [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   358
     "(0::hypreal) \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   359
by (simp only: hrealpow_two_le add_nonneg_nonneg)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   360
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   361
lemma hrealpow_two_le_add_order2 [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   362
     "(0::hypreal) \<le> u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   363
by (simp only: hrealpow_two_le add_nonneg_nonneg)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   364
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   365
lemma hypreal_add_nonneg_eq_0_iff:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   366
     "[| 0 \<le> x; 0 \<le> y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   367
by arith
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   368
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   369
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   370
text{*FIXME: DELETE THESE*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   371
lemma hypreal_three_squares_add_zero_iff:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   372
     "(x*x + y*y + z*z = 0) = (x = 0 & y = 0 & z = (0::hypreal))"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   373
apply (simp only: zero_le_square add_nonneg_nonneg hypreal_add_nonneg_eq_0_iff, auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   374
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   375
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   376
lemma hrealpow_three_squares_add_zero_iff [simp]:
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55911
diff changeset
   377
     "(x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + z ^ Suc (Suc 0) = (0::hypreal)) =
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   378
      (x = 0 & y = 0 & z = 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   379
by (simp only: hypreal_three_squares_add_zero_iff hrealpow_two)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   380
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   381
(*FIXME: This and RealPow.abs_realpow_two should be replaced by an abstract
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
   382
  result proved in Rings or Fields*)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   383
lemma hrabs_hrealpow_two [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   384
     "abs(x ^ Suc (Suc 0)) = (x::hypreal) ^ Suc (Suc 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   385
by (simp add: abs_mult)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   386
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   387
lemma two_hrealpow_ge_one [simp]: "(1::hypreal) \<le> 2 ^ n"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   388
by (insert power_increasing [of 0 n "2::hypreal"], simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   389
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   390
lemma two_hrealpow_gt [simp]: "hypreal_of_nat n < 2 ^ n"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   391
apply (induct n)
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47195
diff changeset
   392
apply (auto simp add: distrib_right)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   393
apply (cut_tac n = n in two_hrealpow_ge_one, arith)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   394
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   395
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   396
lemma hrealpow:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   397
    "star_n X ^ m = star_n (%n. (X n::real) ^ m)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   398
apply (induct_tac "m")
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   399
apply (auto simp add: star_n_one_num star_n_mult power_0)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   400
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   401
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   402
lemma hrealpow_sum_square_expand:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   403
     "(x + (y::hypreal)) ^ Suc (Suc 0) =
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   404
      x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + (hypreal_of_nat (Suc (Suc 0)))*x*y"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47195
diff changeset
   405
by (simp add: distrib_left distrib_right)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   406
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45605
diff changeset
   407
lemma power_hypreal_of_real_numeral:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45605
diff changeset
   408
     "(numeral v :: hypreal) ^ n = hypreal_of_real ((numeral v) ^ n)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   409
by simp
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45605
diff changeset
   410
declare power_hypreal_of_real_numeral [of _ "numeral w", simp] for w
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45605
diff changeset
   411
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45605
diff changeset
   412
lemma power_hypreal_of_real_neg_numeral:
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 51525
diff changeset
   413
     "(- numeral v :: hypreal) ^ n = hypreal_of_real ((- numeral v) ^ n)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45605
diff changeset
   414
by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 45605
diff changeset
   415
declare power_hypreal_of_real_neg_numeral [of _ "numeral w", simp] for w
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   416
(*
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   417
lemma hrealpow_HFinite:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 31001
diff changeset
   418
  fixes x :: "'a::{real_normed_algebra,power} star"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   419
  shows "x \<in> HFinite ==> x ^ n \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   420
apply (induct_tac "n")
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   421
apply (auto simp add: power_Suc intro: HFinite_mult)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   422
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   423
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   424
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   425
subsection{*Powers with Hypernatural Exponents*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   426
31001
7e6ffd8f51a9 cleaned up theory power further
haftmann
parents: 30968
diff changeset
   427
definition pow :: "['a::power star, nat star] \<Rightarrow> 'a star" (infixr "pow" 80) where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 36409
diff changeset
   428
  hyperpow_def [transfer_unfold]: "R pow N = ( *f2* op ^) R N"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   429
  (* hypernatural powers of hyperreals *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   430
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   431
lemma Standard_hyperpow [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   432
  "\<lbrakk>r \<in> Standard; n \<in> Standard\<rbrakk> \<Longrightarrow> r pow n \<in> Standard"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   433
unfolding hyperpow_def by simp
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   434
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   435
lemma hyperpow: "star_n X pow star_n Y = star_n (%n. X n ^ Y n)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   436
by (simp add: hyperpow_def starfun2_star_n)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   437
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   438
lemma hyperpow_zero [simp]:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 31001
diff changeset
   439
  "\<And>n. (0::'a::{power,semiring_0} star) pow (n + (1::hypnat)) = 0"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   440
by transfer simp
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   441
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   442
lemma hyperpow_not_zero:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 31001
diff changeset
   443
  "\<And>r n. r \<noteq> (0::'a::{field} star) ==> r pow n \<noteq> 0"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   444
by transfer (rule field_power_not_zero)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   445
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   446
lemma hyperpow_inverse:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59582
diff changeset
   447
  "\<And>r n. r \<noteq> (0::'a::field star)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   448
   \<Longrightarrow> inverse (r pow n) = (inverse r) pow n"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   449
by transfer (rule power_inverse)
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55911
diff changeset
   450
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   451
lemma hyperpow_hrabs:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 31101
diff changeset
   452
  "\<And>r n. abs (r::'a::{linordered_idom} star) pow n = abs (r pow n)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   453
by transfer (rule power_abs [symmetric])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   454
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   455
lemma hyperpow_add:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 31001
diff changeset
   456
  "\<And>r n m. (r::'a::monoid_mult star) pow (n + m) = (r pow n) * (r pow m)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   457
by transfer (rule power_add)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   458
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   459
lemma hyperpow_one [simp]:
31001
7e6ffd8f51a9 cleaned up theory power further
haftmann
parents: 30968
diff changeset
   460
  "\<And>r. (r::'a::monoid_mult star) pow (1::hypnat) = r"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   461
by transfer (rule power_one_right)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   462
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   463
lemma hyperpow_two:
45542
4849dbe6e310 HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
huffman
parents: 43595
diff changeset
   464
  "\<And>r. (r::'a::monoid_mult star) pow (2::hypnat) = r * r"
4849dbe6e310 HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
huffman
parents: 43595
diff changeset
   465
by transfer (rule power2_eq_square)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   466
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   467
lemma hyperpow_gt_zero:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 31101
diff changeset
   468
  "\<And>r n. (0::'a::{linordered_semidom} star) < r \<Longrightarrow> 0 < r pow n"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   469
by transfer (rule zero_less_power)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   470
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   471
lemma hyperpow_ge_zero:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 31101
diff changeset
   472
  "\<And>r n. (0::'a::{linordered_semidom} star) \<le> r \<Longrightarrow> 0 \<le> r pow n"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   473
by transfer (rule zero_le_power)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   474
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   475
lemma hyperpow_le:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 31101
diff changeset
   476
  "\<And>x y n. \<lbrakk>(0::'a::{linordered_semidom} star) < x; x \<le> y\<rbrakk>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   477
   \<Longrightarrow> x pow n \<le> y pow n"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   478
by transfer (rule power_mono [OF _ order_less_imp_le])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   479
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   480
lemma hyperpow_eq_one [simp]:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 31001
diff changeset
   481
  "\<And>n. 1 pow n = (1::'a::monoid_mult star)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   482
by transfer (rule power_one)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   483
55911
d00023bd3554 remove simp rules made redundant by the replacement of neg_numeral with negated numerals
huffman
parents: 54489
diff changeset
   484
lemma hrabs_hyperpow_minus [simp]:
d00023bd3554 remove simp rules made redundant by the replacement of neg_numeral with negated numerals
huffman
parents: 54489
diff changeset
   485
  "\<And>(a::'a::{linordered_idom} star) n. abs((-a) pow n) = abs (a pow n)"
d00023bd3554 remove simp rules made redundant by the replacement of neg_numeral with negated numerals
huffman
parents: 54489
diff changeset
   486
by transfer (rule abs_power_minus)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   487
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   488
lemma hyperpow_mult:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 31001
diff changeset
   489
  "\<And>r s n. (r * s::'a::{comm_monoid_mult} star) pow n
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   490
   = (r pow n) * (s pow n)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   491
by transfer (rule power_mult_distrib)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   492
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   493
lemma hyperpow_two_le [simp]:
45542
4849dbe6e310 HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
huffman
parents: 43595
diff changeset
   494
  "\<And>r. (0::'a::{monoid_mult,linordered_ring_strict} star) \<le> r pow 2"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   495
by (auto simp add: hyperpow_two zero_le_mult_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   496
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   497
lemma hrabs_hyperpow_two [simp]:
45542
4849dbe6e310 HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
huffman
parents: 43595
diff changeset
   498
  "abs(x pow 2) =
4849dbe6e310 HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
huffman
parents: 43595
diff changeset
   499
   (x::'a::{monoid_mult,linordered_ring_strict} star) pow 2"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   500
by (simp only: abs_of_nonneg hyperpow_two_le)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   501
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   502
lemma hyperpow_two_hrabs [simp]:
45542
4849dbe6e310 HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
huffman
parents: 43595
diff changeset
   503
  "abs(x::'a::{linordered_idom} star) pow 2 = x pow 2"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   504
by (simp add: hyperpow_hrabs)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   505
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   506
text{*The precondition could be weakened to @{term "0\<le>x"}*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   507
lemma hypreal_mult_less_mono:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   508
     "[| u<v;  x<y;  (0::hypreal) < v;  0 < x |] ==> u*x < v* y"
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
   509
 by (simp add: mult_strict_mono order_less_imp_le)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   510
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   511
lemma hyperpow_two_gt_one:
45542
4849dbe6e310 HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
huffman
parents: 43595
diff changeset
   512
  "\<And>r::'a::{linordered_semidom} star. 1 < r \<Longrightarrow> 1 < r pow 2"
4849dbe6e310 HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
huffman
parents: 43595
diff changeset
   513
by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   514
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   515
lemma hyperpow_two_ge_one:
45542
4849dbe6e310 HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
huffman
parents: 43595
diff changeset
   516
  "\<And>r::'a::{linordered_semidom} star. 1 \<le> r \<Longrightarrow> 1 \<le> r pow 2"
4849dbe6e310 HOL-NSA: add number_semiring instance, reformulate several lemmas using '2' instead of '1+1'
huffman
parents: 43595
diff changeset
   517
by transfer (rule one_le_power)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   518
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   519
lemma two_hyperpow_ge_one [simp]: "(1::hypreal) \<le> 2 pow n"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   520
apply (rule_tac y = "1 pow n" in order_trans)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   521
apply (rule_tac [2] hyperpow_le, auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   522
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   523
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   524
lemma hyperpow_minus_one2 [simp]:
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 56225
diff changeset
   525
     "\<And>n. (- 1) pow (2*n) = (1::hypreal)"
47195
836bf25fb70f remove duplicate lemmas power_m1_{even,odd} in favor of power_minus1_{even,odd}
huffman
parents: 47108
diff changeset
   526
by transfer (rule power_minus1_even)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   527
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   528
lemma hyperpow_less_le:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   529
     "!!r n N. [|(0::hypreal) \<le> r; r \<le> 1; n < N|] ==> r pow N \<le> r pow n"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   530
by transfer (rule power_decreasing [OF order_less_imp_le])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   531
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   532
lemma hyperpow_SHNat_le:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   533
     "[| 0 \<le> r;  r \<le> (1::hypreal);  N \<in> HNatInfinite |]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   534
      ==> ALL n: Nats. r pow N \<le> r pow n"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   535
by (auto intro!: hyperpow_less_le simp add: HNatInfinite_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   536
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   537
lemma hyperpow_realpow:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   538
      "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   539
by transfer (rule refl)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   540
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   541
lemma hyperpow_SReal [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   542
     "(hypreal_of_real r) pow (hypnat_of_nat n) \<in> Reals"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   543
by (simp add: Reals_eq_Standard)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   544
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   545
lemma hyperpow_zero_HNatInfinite [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   546
     "N \<in> HNatInfinite ==> (0::hypreal) pow N = 0"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   547
by (drule HNatInfinite_is_Suc, auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   548
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   549
lemma hyperpow_le_le:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   550
     "[| (0::hypreal) \<le> r; r \<le> 1; n \<le> N |] ==> r pow N \<le> r pow n"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   551
apply (drule order_le_less [of n, THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   552
apply (auto intro: hyperpow_less_le)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   553
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   554
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   555
lemma hyperpow_Suc_le_self2:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   556
     "[| (0::hypreal) \<le> r; r < 1 |] ==> r pow (n + (1::hypnat)) \<le> r"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   557
apply (drule_tac n = " (1::hypnat) " in hyperpow_le_le)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   558
apply auto
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   559
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   560
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   561
lemma hyperpow_hypnat_of_nat: "\<And>x. x pow hypnat_of_nat n = x ^ n"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   562
by transfer (rule refl)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   563
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   564
lemma of_hypreal_hyperpow:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   565
  "\<And>x n. of_hypreal (x pow n) =
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 31001
diff changeset
   566
   (of_hypreal x::'a::{real_algebra_1} star) pow n"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   567
by transfer (rule of_real_power)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   568
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   569
end