src/HOL/NSA/HTranscendental.thy
author huffman
Wed, 04 Mar 2009 17:12:23 -0800
changeset 30273 ecd6f0ca62ea
parent 27468 0783dd1dc13d
child 31271 0237e5e40b71
permissions -rw-r--r--
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     1
(*  Title       : HTranscendental.thy
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   : 2001 University of Edinburgh
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     4
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     5
Converted to Isar and polished by lcp
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     6
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     7
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     8
header{*Nonstandard Extensions of Transcendental Functions*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     9
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    10
theory HTranscendental
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    11
imports Transcendental HSeries HDeriv
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    12
begin
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    13
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    14
definition
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    15
  exphr :: "real => hypreal" where
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    16
    --{*define exponential function using standard part *}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    17
  "exphr x =  st(sumhr (0, whn, %n. inverse(real (fact n)) * (x ^ n)))"
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
definition
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    20
  sinhr :: "real => hypreal" where
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    21
  "sinhr x = st(sumhr (0, whn, %n. (if even(n) then 0 else
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    22
             ((-1) ^ ((n - 1) div 2))/(real (fact n))) * (x ^ n)))"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    23
  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    24
definition
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    25
  coshr :: "real => hypreal" where
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    26
  "coshr x = st(sumhr (0, whn, %n. (if even(n) then
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    27
            ((-1) ^ (n div 2))/(real (fact n)) else 0) * x ^ n))"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    28
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    29
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    30
subsection{*Nonstandard Extension of Square Root Function*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    31
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    32
lemma STAR_sqrt_zero [simp]: "( *f* sqrt) 0 = 0"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    33
by (simp add: starfun star_n_zero_num)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    34
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    35
lemma STAR_sqrt_one [simp]: "( *f* sqrt) 1 = 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    36
by (simp add: starfun star_n_one_num)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    37
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    38
lemma hypreal_sqrt_pow2_iff: "(( *f* sqrt)(x) ^ 2 = x) = (0 \<le> x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    39
apply (cases x)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    40
apply (auto simp add: star_n_le star_n_zero_num starfun hrealpow star_n_eq_iff
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 27468
diff changeset
    41
            simp del: hpowr_Suc power_Suc)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    42
done
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
lemma hypreal_sqrt_gt_zero_pow2: "!!x. 0 < x ==> ( *f* sqrt) (x) ^ 2 = x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    45
by (transfer, simp)
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
lemma hypreal_sqrt_pow2_gt_zero: "0 < x ==> 0 < ( *f* sqrt) (x) ^ 2"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    48
by (frule hypreal_sqrt_gt_zero_pow2, auto)
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
lemma hypreal_sqrt_not_zero: "0 < x ==> ( *f* sqrt) (x) \<noteq> 0"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    51
apply (frule hypreal_sqrt_pow2_gt_zero)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    52
apply (auto simp add: numeral_2_eq_2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    53
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    54
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    55
lemma hypreal_inverse_sqrt_pow2:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    56
     "0 < x ==> inverse (( *f* sqrt)(x)) ^ 2 = inverse x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    57
apply (cut_tac n = 2 and a = "( *f* sqrt) x" in power_inverse [symmetric])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    58
apply (auto dest: hypreal_sqrt_gt_zero_pow2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    59
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    60
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    61
lemma hypreal_sqrt_mult_distrib: 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    62
    "!!x y. [|0 < x; 0 <y |] ==>
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    63
      ( *f* sqrt)(x*y) = ( *f* sqrt)(x) * ( *f* sqrt)(y)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    64
apply transfer
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    65
apply (auto intro: real_sqrt_mult_distrib) 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    66
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    67
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    68
lemma hypreal_sqrt_mult_distrib2:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    69
     "[|0\<le>x; 0\<le>y |] ==>  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    70
     ( *f* sqrt)(x*y) =  ( *f* sqrt)(x) * ( *f* sqrt)(y)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    71
by (auto intro: hypreal_sqrt_mult_distrib simp add: order_le_less)
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
lemma hypreal_sqrt_approx_zero [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    74
     "0 < x ==> (( *f* sqrt)(x) @= 0) = (x @= 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    75
apply (auto simp add: mem_infmal_iff [symmetric])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    76
apply (rule hypreal_sqrt_gt_zero_pow2 [THEN subst])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    77
apply (auto intro: Infinitesimal_mult 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    78
            dest!: hypreal_sqrt_gt_zero_pow2 [THEN ssubst] 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    79
            simp add: numeral_2_eq_2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    80
done
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
lemma hypreal_sqrt_approx_zero2 [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    83
     "0 \<le> x ==> (( *f* sqrt)(x) @= 0) = (x @= 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    84
by (auto simp add: order_le_less)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    85
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    86
lemma hypreal_sqrt_sum_squares [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    87
     "(( *f* sqrt)(x*x + y*y + z*z) @= 0) = (x*x + y*y + z*z @= 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    88
apply (rule hypreal_sqrt_approx_zero2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    89
apply (rule add_nonneg_nonneg)+
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    90
apply (auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    91
done
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 hypreal_sqrt_sum_squares2 [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    94
     "(( *f* sqrt)(x*x + y*y) @= 0) = (x*x + y*y @= 0)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    95
apply (rule hypreal_sqrt_approx_zero2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    96
apply (rule add_nonneg_nonneg)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    97
apply (auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    98
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    99
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   100
lemma hypreal_sqrt_gt_zero: "!!x. 0 < x ==> 0 < ( *f* sqrt)(x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   101
apply transfer
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   102
apply (auto intro: real_sqrt_gt_zero)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   103
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   104
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   105
lemma hypreal_sqrt_ge_zero: "0 \<le> x ==> 0 \<le> ( *f* sqrt)(x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   106
by (auto intro: hypreal_sqrt_gt_zero simp add: order_le_less)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   107
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   108
lemma hypreal_sqrt_hrabs [simp]: "!!x. ( *f* sqrt)(x ^ 2) = abs(x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   109
by (transfer, simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   110
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   111
lemma hypreal_sqrt_hrabs2 [simp]: "!!x. ( *f* sqrt)(x*x) = abs(x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   112
by (transfer, simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   113
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   114
lemma hypreal_sqrt_hyperpow_hrabs [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   115
     "!!x. ( *f* sqrt)(x pow (hypnat_of_nat 2)) = abs(x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   116
by (transfer, simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   117
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   118
lemma star_sqrt_HFinite: "\<lbrakk>x \<in> HFinite; 0 \<le> x\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   119
apply (rule HFinite_square_iff [THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   120
apply (simp only: hypreal_sqrt_mult_distrib2 [symmetric], simp) 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   121
done
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 st_hypreal_sqrt:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   124
     "[| x \<in> HFinite; 0 \<le> x |] ==> st(( *f* sqrt) x) = ( *f* sqrt)(st x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   125
apply (rule power_inject_base [where n=1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   126
apply (auto intro!: st_zero_le hypreal_sqrt_ge_zero)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   127
apply (rule st_mult [THEN subst])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   128
apply (rule_tac [3] hypreal_sqrt_mult_distrib2 [THEN subst])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   129
apply (rule_tac [5] hypreal_sqrt_mult_distrib2 [THEN subst])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   130
apply (auto simp add: st_hrabs st_zero_le star_sqrt_HFinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   131
done
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 hypreal_sqrt_sum_squares_ge1 [simp]: "!!x y. x \<le> ( *f* sqrt)(x ^ 2 + y ^ 2)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   134
by transfer (rule real_sqrt_sum_squares_ge1)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   135
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   136
lemma HFinite_hypreal_sqrt:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   137
     "[| 0 \<le> x; x \<in> HFinite |] ==> ( *f* sqrt) x \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   138
apply (auto simp add: order_le_less)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   139
apply (rule HFinite_square_iff [THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   140
apply (drule hypreal_sqrt_gt_zero_pow2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   141
apply (simp add: numeral_2_eq_2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   142
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   143
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   144
lemma HFinite_hypreal_sqrt_imp_HFinite:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   145
     "[| 0 \<le> x; ( *f* sqrt) x \<in> HFinite |] ==> x \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   146
apply (auto simp add: order_le_less)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   147
apply (drule HFinite_square_iff [THEN iffD2])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   148
apply (drule hypreal_sqrt_gt_zero_pow2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   149
apply (simp add: numeral_2_eq_2 del: HFinite_square_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   150
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   151
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   152
lemma HFinite_hypreal_sqrt_iff [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   153
     "0 \<le> x ==> (( *f* sqrt) x \<in> HFinite) = (x \<in> HFinite)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   154
by (blast intro: HFinite_hypreal_sqrt HFinite_hypreal_sqrt_imp_HFinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   155
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   156
lemma HFinite_sqrt_sum_squares [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   157
     "(( *f* sqrt)(x*x + y*y) \<in> HFinite) = (x*x + y*y \<in> HFinite)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   158
apply (rule HFinite_hypreal_sqrt_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   159
apply (rule add_nonneg_nonneg)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   160
apply (auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   161
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   162
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   163
lemma Infinitesimal_hypreal_sqrt:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   164
     "[| 0 \<le> x; x \<in> Infinitesimal |] ==> ( *f* sqrt) x \<in> Infinitesimal"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   165
apply (auto simp add: order_le_less)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   166
apply (rule Infinitesimal_square_iff [THEN iffD2])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   167
apply (drule hypreal_sqrt_gt_zero_pow2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   168
apply (simp add: numeral_2_eq_2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   169
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   170
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   171
lemma Infinitesimal_hypreal_sqrt_imp_Infinitesimal:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   172
     "[| 0 \<le> x; ( *f* sqrt) x \<in> Infinitesimal |] ==> x \<in> Infinitesimal"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   173
apply (auto simp add: order_le_less)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   174
apply (drule Infinitesimal_square_iff [THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   175
apply (drule hypreal_sqrt_gt_zero_pow2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   176
apply (simp add: numeral_2_eq_2 del: Infinitesimal_square_iff [symmetric])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   177
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   178
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   179
lemma Infinitesimal_hypreal_sqrt_iff [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   180
     "0 \<le> x ==> (( *f* sqrt) x \<in> Infinitesimal) = (x \<in> Infinitesimal)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   181
by (blast intro: Infinitesimal_hypreal_sqrt_imp_Infinitesimal Infinitesimal_hypreal_sqrt)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   182
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   183
lemma Infinitesimal_sqrt_sum_squares [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   184
     "(( *f* sqrt)(x*x + y*y) \<in> Infinitesimal) = (x*x + y*y \<in> Infinitesimal)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   185
apply (rule Infinitesimal_hypreal_sqrt_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   186
apply (rule add_nonneg_nonneg)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   187
apply (auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   188
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   189
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   190
lemma HInfinite_hypreal_sqrt:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   191
     "[| 0 \<le> x; x \<in> HInfinite |] ==> ( *f* sqrt) x \<in> HInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   192
apply (auto simp add: order_le_less)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   193
apply (rule HInfinite_square_iff [THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   194
apply (drule hypreal_sqrt_gt_zero_pow2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   195
apply (simp add: numeral_2_eq_2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   196
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   197
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   198
lemma HInfinite_hypreal_sqrt_imp_HInfinite:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   199
     "[| 0 \<le> x; ( *f* sqrt) x \<in> HInfinite |] ==> x \<in> HInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   200
apply (auto simp add: order_le_less)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   201
apply (drule HInfinite_square_iff [THEN iffD2])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   202
apply (drule hypreal_sqrt_gt_zero_pow2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   203
apply (simp add: numeral_2_eq_2 del: HInfinite_square_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   204
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   205
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   206
lemma HInfinite_hypreal_sqrt_iff [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   207
     "0 \<le> x ==> (( *f* sqrt) x \<in> HInfinite) = (x \<in> HInfinite)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   208
by (blast intro: HInfinite_hypreal_sqrt HInfinite_hypreal_sqrt_imp_HInfinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   209
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   210
lemma HInfinite_sqrt_sum_squares [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   211
     "(( *f* sqrt)(x*x + y*y) \<in> HInfinite) = (x*x + y*y \<in> HInfinite)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   212
apply (rule HInfinite_hypreal_sqrt_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   213
apply (rule add_nonneg_nonneg)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   214
apply (auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   215
done
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 HFinite_exp [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   218
     "sumhr (0, whn, %n. inverse (real (fact n)) * x ^ n) \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   219
unfolding sumhr_app
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   220
apply (simp only: star_zero_def starfun2_star_of)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   221
apply (rule NSBseqD2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   222
apply (rule NSconvergent_NSBseq)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   223
apply (rule convergent_NSconvergent_iff [THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   224
apply (rule summable_convergent_sumr_iff [THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   225
apply (rule summable_exp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   226
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   227
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   228
lemma exphr_zero [simp]: "exphr 0 = 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   229
apply (simp add: exphr_def sumhr_split_add
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   230
                   [OF hypnat_one_less_hypnat_omega, symmetric])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   231
apply (rule st_unique, simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   232
apply (rule subst [where P="\<lambda>x. 1 \<approx> x", OF _ approx_refl])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   233
apply (rule rev_mp [OF hypnat_one_less_hypnat_omega])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   234
apply (rule_tac x="whn" in spec)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   235
apply (unfold sumhr_app, transfer, simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   236
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   237
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   238
lemma coshr_zero [simp]: "coshr 0 = 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   239
apply (simp add: coshr_def sumhr_split_add
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   240
                   [OF hypnat_one_less_hypnat_omega, symmetric]) 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   241
apply (rule st_unique, simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   242
apply (rule subst [where P="\<lambda>x. 1 \<approx> x", OF _ approx_refl])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   243
apply (rule rev_mp [OF hypnat_one_less_hypnat_omega])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   244
apply (rule_tac x="whn" in spec)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   245
apply (unfold sumhr_app, transfer, simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   246
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   247
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   248
lemma STAR_exp_zero_approx_one [simp]: "( *f* exp) (0::hypreal) @= 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   249
apply (subgoal_tac "( *f* exp) (0::hypreal) = 1", simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   250
apply (transfer, simp)
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 STAR_exp_Infinitesimal: "x \<in> Infinitesimal ==> ( *f* exp) (x::hypreal) @= 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   254
apply (case_tac "x = 0")
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   255
apply (cut_tac [2] x = 0 in DERIV_exp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   256
apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   257
apply (drule_tac x = x in bspec, auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   258
apply (drule_tac c = x in approx_mult1)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   259
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   260
            simp add: mult_assoc)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   261
apply (rule approx_add_right_cancel [where d="-1"])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   262
apply (rule approx_sym [THEN [2] approx_trans2])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   263
apply (auto simp add: diff_def mem_infmal_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   264
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   265
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   266
lemma STAR_exp_epsilon [simp]: "( *f* exp) epsilon @= 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   267
by (auto intro: STAR_exp_Infinitesimal)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   268
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   269
lemma STAR_exp_add: "!!x y. ( *f* exp)(x + y) = ( *f* exp) x * ( *f* exp) y"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   270
by transfer (rule exp_add)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   271
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   272
lemma exphr_hypreal_of_real_exp_eq: "exphr x = hypreal_of_real (exp x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   273
apply (simp add: exphr_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   274
apply (rule st_unique, simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   275
apply (subst starfunNat_sumr [symmetric])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   276
apply (rule NSLIMSEQ_D [THEN approx_sym])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   277
apply (rule LIMSEQ_NSLIMSEQ)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   278
apply (subst sums_def [symmetric])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   279
apply (cut_tac exp_converges [where x=x], simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   280
apply (rule HNatInfinite_whn)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   281
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   282
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   283
lemma starfun_exp_ge_add_one_self [simp]: "!!x::hypreal. 0 \<le> x ==> (1 + x) \<le> ( *f* exp) x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   284
by transfer (rule exp_ge_add_one_self_aux)
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
(* exp (oo) is infinite *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   287
lemma starfun_exp_HInfinite:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   288
     "[| x \<in> HInfinite; 0 \<le> x |] ==> ( *f* exp) (x::hypreal) \<in> HInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   289
apply (frule starfun_exp_ge_add_one_self)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   290
apply (rule HInfinite_ge_HInfinite, assumption)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   291
apply (rule order_trans [of _ "1+x"], auto) 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   292
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   293
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   294
lemma starfun_exp_minus: "!!x. ( *f* exp) (-x) = inverse(( *f* exp) x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   295
by transfer (rule exp_minus)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   296
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   297
(* exp (-oo) is infinitesimal *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   298
lemma starfun_exp_Infinitesimal:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   299
     "[| x \<in> HInfinite; x \<le> 0 |] ==> ( *f* exp) (x::hypreal) \<in> Infinitesimal"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   300
apply (subgoal_tac "\<exists>y. x = - y")
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   301
apply (rule_tac [2] x = "- x" in exI)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   302
apply (auto intro!: HInfinite_inverse_Infinitesimal starfun_exp_HInfinite
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   303
            simp add: starfun_exp_minus HInfinite_minus_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   304
done
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 starfun_exp_gt_one [simp]: "!!x::hypreal. 0 < x ==> 1 < ( *f* exp) x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   307
by transfer (rule exp_gt_one)
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
lemma starfun_ln_exp [simp]: "!!x. ( *f* ln) (( *f* exp) x) = x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   310
by transfer (rule ln_exp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   311
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   312
lemma starfun_exp_ln_iff [simp]: "!!x. (( *f* exp)(( *f* ln) x) = x) = (0 < x)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   313
by transfer (rule exp_ln_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   314
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   315
lemma starfun_exp_ln_eq: "!!u x. ( *f* exp) u = x ==> ( *f* ln) x = u"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   316
by transfer (rule exp_ln_eq)
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 starfun_ln_less_self [simp]: "!!x. 0 < x ==> ( *f* ln) x < x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   319
by transfer (rule ln_less_self)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   320
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   321
lemma starfun_ln_ge_zero [simp]: "!!x. 1 \<le> x ==> 0 \<le> ( *f* ln) x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   322
by transfer (rule ln_ge_zero)
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
lemma starfun_ln_gt_zero [simp]: "!!x .1 < x ==> 0 < ( *f* ln) x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   325
by transfer (rule ln_gt_zero)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   326
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   327
lemma starfun_ln_not_eq_zero [simp]: "!!x. [| 0 < x; x \<noteq> 1 |] ==> ( *f* ln) x \<noteq> 0"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   328
by transfer simp
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   329
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   330
lemma starfun_ln_HFinite: "[| x \<in> HFinite; 1 \<le> x |] ==> ( *f* ln) x \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   331
apply (rule HFinite_bounded)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   332
apply assumption 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   333
apply (simp_all add: starfun_ln_less_self order_less_imp_le)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   334
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   335
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   336
lemma starfun_ln_inverse: "!!x. 0 < x ==> ( *f* ln) (inverse x) = -( *f* ln) x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   337
by transfer (rule ln_inverse)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   338
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   339
lemma starfun_abs_exp_cancel: "\<And>x. \<bar>( *f* exp) (x::hypreal)\<bar> = ( *f* exp) x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   340
by transfer (rule abs_exp_cancel)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   341
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   342
lemma starfun_exp_less_mono: "\<And>x y::hypreal. x < y \<Longrightarrow> ( *f* exp) x < ( *f* exp) y"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   343
by transfer (rule exp_less_mono)
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 starfun_exp_HFinite: "x \<in> HFinite ==> ( *f* exp) (x::hypreal) \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   346
apply (auto simp add: HFinite_def, rename_tac u)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   347
apply (rule_tac x="( *f* exp) u" in rev_bexI)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   348
apply (simp add: Reals_eq_Standard)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   349
apply (simp add: starfun_abs_exp_cancel)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   350
apply (simp add: starfun_exp_less_mono)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   351
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   352
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   353
lemma starfun_exp_add_HFinite_Infinitesimal_approx:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   354
     "[|x \<in> Infinitesimal; z \<in> HFinite |] ==> ( *f* exp) (z + x::hypreal) @= ( *f* exp) z"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   355
apply (simp add: STAR_exp_add)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   356
apply (frule STAR_exp_Infinitesimal)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   357
apply (drule approx_mult2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   358
apply (auto intro: starfun_exp_HFinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   359
done
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
(* using previous result to get to result *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   362
lemma starfun_ln_HInfinite:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   363
     "[| x \<in> HInfinite; 0 < x |] ==> ( *f* ln) x \<in> HInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   364
apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   365
apply (drule starfun_exp_HFinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   366
apply (simp add: starfun_exp_ln_iff [THEN iffD2] HFinite_HInfinite_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   367
done
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
lemma starfun_exp_HInfinite_Infinitesimal_disj:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   370
 "x \<in> HInfinite ==> ( *f* exp) x \<in> HInfinite | ( *f* exp) (x::hypreal) \<in> Infinitesimal"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   371
apply (insert linorder_linear [of x 0]) 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   372
apply (auto intro: starfun_exp_HInfinite starfun_exp_Infinitesimal)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   373
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   374
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   375
(* check out this proof!!! *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   376
lemma starfun_ln_HFinite_not_Infinitesimal:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   377
     "[| x \<in> HFinite - Infinitesimal; 0 < x |] ==> ( *f* ln) x \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   378
apply (rule ccontr, drule HInfinite_HFinite_iff [THEN iffD2])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   379
apply (drule starfun_exp_HInfinite_Infinitesimal_disj)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   380
apply (simp add: starfun_exp_ln_iff [symmetric] HInfinite_HFinite_iff
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   381
            del: starfun_exp_ln_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   382
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   383
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   384
(* we do proof by considering ln of 1/x *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   385
lemma starfun_ln_Infinitesimal_HInfinite:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   386
     "[| x \<in> Infinitesimal; 0 < x |] ==> ( *f* ln) x \<in> HInfinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   387
apply (drule Infinitesimal_inverse_HInfinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   388
apply (frule positive_imp_inverse_positive)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   389
apply (drule_tac [2] starfun_ln_HInfinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   390
apply (auto simp add: starfun_ln_inverse HInfinite_minus_iff)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   391
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   392
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   393
lemma starfun_ln_less_zero: "!!x. [| 0 < x; x < 1 |] ==> ( *f* ln) x < 0"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   394
by transfer (rule ln_less_zero)
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 starfun_ln_Infinitesimal_less_zero:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   397
     "[| x \<in> Infinitesimal; 0 < x |] ==> ( *f* ln) x < 0"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   398
by (auto intro!: starfun_ln_less_zero simp add: Infinitesimal_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   399
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   400
lemma starfun_ln_HInfinite_gt_zero:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   401
     "[| x \<in> HInfinite; 0 < x |] ==> 0 < ( *f* ln) x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   402
by (auto intro!: starfun_ln_gt_zero simp add: HInfinite_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   403
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   404
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   405
(*
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   406
Goalw [NSLIM_def] "(%h. ((x powr h) - 1) / h) -- 0 --NS> ln x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   407
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   408
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   409
lemma HFinite_sin [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   410
     "sumhr (0, whn, %n. (if even(n) then 0 else  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   411
              (-1 ^ ((n - 1) div 2))/(real (fact n))) * x ^ n)  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   412
              \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   413
unfolding sumhr_app
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   414
apply (simp only: star_zero_def starfun2_star_of)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   415
apply (rule NSBseqD2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   416
apply (rule NSconvergent_NSBseq)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   417
apply (rule convergent_NSconvergent_iff [THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   418
apply (rule summable_convergent_sumr_iff [THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   419
apply (simp only: One_nat_def summable_sin)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   420
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   421
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   422
lemma STAR_sin_zero [simp]: "( *f* sin) 0 = 0"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   423
by transfer (rule sin_zero)
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
lemma STAR_sin_Infinitesimal [simp]: "x \<in> Infinitesimal ==> ( *f* sin) x @= x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   426
apply (case_tac "x = 0")
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   427
apply (cut_tac [2] x = 0 in DERIV_sin)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   428
apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   429
apply (drule bspec [where x = x], auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   430
apply (drule approx_mult1 [where c = x])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   431
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   432
           simp add: mult_assoc)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   433
done
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 HFinite_cos [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   436
     "sumhr (0, whn, %n. (if even(n) then  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   437
            (-1 ^ (n div 2))/(real (fact n)) else  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   438
            0) * x ^ n) \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   439
unfolding sumhr_app
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   440
apply (simp only: star_zero_def starfun2_star_of)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   441
apply (rule NSBseqD2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   442
apply (rule NSconvergent_NSBseq)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   443
apply (rule convergent_NSconvergent_iff [THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   444
apply (rule summable_convergent_sumr_iff [THEN iffD1])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   445
apply (rule summable_cos)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   446
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   447
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   448
lemma STAR_cos_zero [simp]: "( *f* cos) 0 = 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   449
by transfer (rule cos_zero)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   450
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   451
lemma STAR_cos_Infinitesimal [simp]: "x \<in> Infinitesimal ==> ( *f* cos) x @= 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   452
apply (case_tac "x = 0")
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   453
apply (cut_tac [2] x = 0 in DERIV_cos)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   454
apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   455
apply (drule bspec [where x = x])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   456
apply auto
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   457
apply (drule approx_mult1 [where c = x])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   458
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   459
            simp add: mult_assoc)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   460
apply (rule approx_add_right_cancel [where d = "-1"])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   461
apply (simp add: diff_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   462
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   463
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   464
lemma STAR_tan_zero [simp]: "( *f* tan) 0 = 0"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   465
by transfer (rule tan_zero)
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 STAR_tan_Infinitesimal: "x \<in> Infinitesimal ==> ( *f* tan) x @= x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   468
apply (case_tac "x = 0")
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   469
apply (cut_tac [2] x = 0 in DERIV_tan)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   470
apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   471
apply (drule bspec [where x = x], auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   472
apply (drule approx_mult1 [where c = x])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   473
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   474
             simp add: mult_assoc)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   475
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   476
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   477
lemma STAR_sin_cos_Infinitesimal_mult:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   478
     "x \<in> Infinitesimal ==> ( *f* sin) x * ( *f* cos) x @= x"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   479
apply (insert approx_mult_HFinite [of "( *f* sin) x" _ "( *f* cos) x" 1]) 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   480
apply (simp add: Infinitesimal_subset_HFinite [THEN subsetD])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   481
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   482
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   483
lemma HFinite_pi: "hypreal_of_real pi \<in> HFinite"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   484
by simp
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   485
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   486
(* lemmas *)
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 lemma_split_hypreal_of_real:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   489
     "N \<in> HNatInfinite  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   490
      ==> hypreal_of_real a =  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   491
          hypreal_of_hypnat N * (inverse(hypreal_of_hypnat N) * hypreal_of_real a)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   492
by (simp add: mult_assoc [symmetric] zero_less_HNatInfinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   493
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   494
lemma STAR_sin_Infinitesimal_divide:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   495
     "[|x \<in> Infinitesimal; x \<noteq> 0 |] ==> ( *f* sin) x/x @= 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   496
apply (cut_tac x = 0 in DERIV_sin)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   497
apply (simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   498
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   499
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   500
(*------------------------------------------------------------------------*) 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   501
(* sin* (1/n) * 1/(1/n) @= 1 for n = oo                                   *)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   502
(*------------------------------------------------------------------------*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   503
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   504
lemma lemma_sin_pi:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   505
     "n \<in> HNatInfinite  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   506
      ==> ( *f* sin) (inverse (hypreal_of_hypnat n))/(inverse (hypreal_of_hypnat n)) @= 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   507
apply (rule STAR_sin_Infinitesimal_divide)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   508
apply (auto simp add: zero_less_HNatInfinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   509
done
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 STAR_sin_inverse_HNatInfinite:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   512
     "n \<in> HNatInfinite  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   513
      ==> ( *f* sin) (inverse (hypreal_of_hypnat n)) * hypreal_of_hypnat n @= 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   514
apply (frule lemma_sin_pi)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   515
apply (simp add: divide_inverse)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   516
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   517
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   518
lemma Infinitesimal_pi_divide_HNatInfinite: 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   519
     "N \<in> HNatInfinite  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   520
      ==> hypreal_of_real pi/(hypreal_of_hypnat N) \<in> Infinitesimal"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   521
apply (simp add: divide_inverse)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   522
apply (auto intro: Infinitesimal_HFinite_mult2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   523
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   524
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   525
lemma pi_divide_HNatInfinite_not_zero [simp]:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   526
     "N \<in> HNatInfinite ==> hypreal_of_real pi/(hypreal_of_hypnat N) \<noteq> 0"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   527
by (simp add: zero_less_HNatInfinite)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   528
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   529
lemma STAR_sin_pi_divide_HNatInfinite_approx_pi:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   530
     "n \<in> HNatInfinite  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   531
      ==> ( *f* sin) (hypreal_of_real pi/(hypreal_of_hypnat n)) * hypreal_of_hypnat n  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   532
          @= hypreal_of_real pi"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   533
apply (frule STAR_sin_Infinitesimal_divide
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   534
               [OF Infinitesimal_pi_divide_HNatInfinite 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   535
                   pi_divide_HNatInfinite_not_zero])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   536
apply (auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   537
apply (rule approx_SReal_mult_cancel [of "inverse (hypreal_of_real pi)"])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   538
apply (auto intro: Reals_inverse simp add: divide_inverse mult_ac)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   539
done
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 STAR_sin_pi_divide_HNatInfinite_approx_pi2:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   542
     "n \<in> HNatInfinite  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   543
      ==> hypreal_of_hypnat n *  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   544
          ( *f* sin) (hypreal_of_real pi/(hypreal_of_hypnat n))  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   545
          @= hypreal_of_real pi"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   546
apply (rule mult_commute [THEN subst])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   547
apply (erule STAR_sin_pi_divide_HNatInfinite_approx_pi)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   548
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   549
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   550
lemma starfunNat_pi_divide_n_Infinitesimal: 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   551
     "N \<in> HNatInfinite ==> ( *f* (%x. pi / real x)) N \<in> Infinitesimal"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   552
by (auto intro!: Infinitesimal_HFinite_mult2 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   553
         simp add: starfun_mult [symmetric] divide_inverse
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   554
                   starfun_inverse [symmetric] starfunNat_real_of_nat)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   555
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   556
lemma STAR_sin_pi_divide_n_approx:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   557
     "N \<in> HNatInfinite ==>  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   558
      ( *f* sin) (( *f* (%x. pi / real x)) N) @=  
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   559
      hypreal_of_real pi/(hypreal_of_hypnat N)"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   560
apply (simp add: starfunNat_real_of_nat [symmetric])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   561
apply (rule STAR_sin_Infinitesimal)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   562
apply (simp add: divide_inverse)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   563
apply (rule Infinitesimal_HFinite_mult2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   564
apply (subst starfun_inverse)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   565
apply (erule starfunNat_inverse_real_of_nat_Infinitesimal)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   566
apply simp
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   567
done
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
lemma NSLIMSEQ_sin_pi: "(%n. real n * sin (pi / real n)) ----NS> pi"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   570
apply (auto simp add: NSLIMSEQ_def starfun_mult [symmetric] starfunNat_real_of_nat)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   571
apply (rule_tac f1 = sin in starfun_o2 [THEN subst])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   572
apply (auto simp add: starfun_mult [symmetric] starfunNat_real_of_nat divide_inverse)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   573
apply (rule_tac f1 = inverse in starfun_o2 [THEN subst])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   574
apply (auto dest: STAR_sin_pi_divide_HNatInfinite_approx_pi 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   575
            simp add: starfunNat_real_of_nat mult_commute divide_inverse)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   576
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   577
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   578
lemma NSLIMSEQ_cos_one: "(%n. cos (pi / real n))----NS> 1"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   579
apply (simp add: NSLIMSEQ_def, auto)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   580
apply (rule_tac f1 = cos in starfun_o2 [THEN subst])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   581
apply (rule STAR_cos_Infinitesimal)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   582
apply (auto intro!: Infinitesimal_HFinite_mult2 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   583
            simp add: starfun_mult [symmetric] divide_inverse
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   584
                      starfun_inverse [symmetric] starfunNat_real_of_nat)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   585
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   586
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   587
lemma NSLIMSEQ_sin_cos_pi:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   588
     "(%n. real n * sin (pi / real n) * cos (pi / real n)) ----NS> pi"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   589
by (insert NSLIMSEQ_mult [OF NSLIMSEQ_sin_pi NSLIMSEQ_cos_one], simp)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   590
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   591
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   592
text{*A familiar approximation to @{term "cos x"} when @{term x} is small*}
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   593
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   594
lemma STAR_cos_Infinitesimal_approx:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   595
     "x \<in> Infinitesimal ==> ( *f* cos) x @= 1 - x ^ 2"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   596
apply (rule STAR_cos_Infinitesimal [THEN approx_trans])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   597
apply (auto simp add: Infinitesimal_approx_minus [symmetric] 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   598
            diff_minus add_assoc [symmetric] numeral_2_eq_2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   599
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   600
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   601
lemma STAR_cos_Infinitesimal_approx2:
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   602
     "x \<in> Infinitesimal ==> ( *f* cos) x @= 1 - (x ^ 2)/2"
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   603
apply (rule STAR_cos_Infinitesimal [THEN approx_trans])
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   604
apply (auto intro: Infinitesimal_SReal_divide 
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   605
            simp add: Infinitesimal_approx_minus [symmetric] numeral_2_eq_2)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   606
done
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   607
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   608
end