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