| author | wenzelm | 
| Wed, 11 Dec 2024 10:40:57 +0100 | |
| changeset 81577 | a712bf5ccab0 | 
| parent 75866 | 9eeed5c424f9 | 
| permissions | -rw-r--r-- | 
| 62479 | 1 | (* Title: HOL/Nonstandard_Analysis/HTranscendental.thy | 
| 2 | Author: Jacques D. Fleuriot | |
| 3 | Copyright: 2001 University of Edinburgh | |
| 27468 | 4 | |
| 5 | Converted to Isar and polished by lcp | |
| 6 | *) | |
| 7 | ||
| 61975 | 8 | section\<open>Nonstandard Extensions of Transcendental Functions\<close> | 
| 27468 | 9 | |
| 10 | theory HTranscendental | |
| 70210 | 11 | imports Complex_Main HSeries HDeriv | 
| 27468 | 12 | begin | 
| 13 | ||
| 14 | definition | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 15 | exphr :: "real \<Rightarrow> hypreal" where | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67091diff
changeset | 16 | \<comment> \<open>define exponential function using standard part\<close> | 
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 17 | "exphr x \<equiv> st(sumhr (0, whn, \<lambda>n. inverse (fact n) * (x ^ n)))" | 
| 27468 | 18 | |
| 19 | definition | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 20 | sinhr :: "real \<Rightarrow> hypreal" where | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 21 | "sinhr x \<equiv> st(sumhr (0, whn, \<lambda>n. sin_coeff n * x ^ n))" | 
| 27468 | 22 | |
| 23 | definition | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 24 | coshr :: "real \<Rightarrow> hypreal" where | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 25 | "coshr x \<equiv> st(sumhr (0, whn, \<lambda>n. cos_coeff n * x ^ n))" | 
| 27468 | 26 | |
| 27 | ||
| 61975 | 28 | subsection\<open>Nonstandard Extension of Square Root Function\<close> | 
| 27468 | 29 | |
| 30 | lemma STAR_sqrt_zero [simp]: "( *f* sqrt) 0 = 0" | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 31 | by (simp add: starfun star_n_zero_num) | 
| 27468 | 32 | |
| 33 | lemma STAR_sqrt_one [simp]: "( *f* sqrt) 1 = 1" | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 34 | by (simp add: starfun star_n_one_num) | 
| 27468 | 35 | |
| 36 | lemma hypreal_sqrt_pow2_iff: "(( *f* sqrt)(x) ^ 2 = x) = (0 \<le> x)" | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 37 | proof (cases x) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 38 | case (star_n X) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 39 | then show ?thesis | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 40 | by (simp add: star_n_le star_n_zero_num starfun hrealpow star_n_eq_iff del: hpowr_Suc power_Suc) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 41 | qed | 
| 27468 | 42 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 43 | lemma hypreal_sqrt_gt_zero_pow2: "\<And>x. 0 < x \<Longrightarrow> ( *f* sqrt) (x) ^ 2 = x" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 44 | by transfer simp | 
| 27468 | 45 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 46 | lemma hypreal_sqrt_pow2_gt_zero: "0 < x \<Longrightarrow> 0 < ( *f* sqrt) (x) ^ 2" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 47 | by (frule hypreal_sqrt_gt_zero_pow2, auto) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 48 | |
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 49 | lemma hypreal_sqrt_not_zero: "0 < x \<Longrightarrow> ( *f* sqrt) (x) \<noteq> 0" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 50 | using hypreal_sqrt_gt_zero_pow2 by fastforce | 
| 27468 | 51 | |
| 52 | lemma hypreal_inverse_sqrt_pow2: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 53 | "0 < x \<Longrightarrow> inverse (( *f* sqrt)(x)) ^ 2 = inverse x" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 54 | by (simp add: hypreal_sqrt_gt_zero_pow2 power_inverse) | 
| 27468 | 55 | |
| 56 | lemma hypreal_sqrt_mult_distrib: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 57 | "\<And>x y. \<lbrakk>0 < x; 0 <y\<rbrakk> \<Longrightarrow> | 
| 27468 | 58 | ( *f* sqrt)(x*y) = ( *f* sqrt)(x) * ( *f* sqrt)(y)" | 
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 59 | by transfer (auto intro: real_sqrt_mult) | 
| 27468 | 60 | |
| 61 | lemma hypreal_sqrt_mult_distrib2: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 62 | "\<lbrakk>0\<le>x; 0\<le>y\<rbrakk> \<Longrightarrow> ( *f* sqrt)(x*y) = ( *f* sqrt)(x) * ( *f* sqrt)(y)" | 
| 27468 | 63 | by (auto intro: hypreal_sqrt_mult_distrib simp add: order_le_less) | 
| 64 | ||
| 65 | lemma hypreal_sqrt_approx_zero [simp]: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 66 | assumes "0 < x" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 67 | shows "(( *f* sqrt) x \<approx> 0) \<longleftrightarrow> (x \<approx> 0)" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 68 | proof - | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 69 | have "( *f* sqrt) x \<in> Infinitesimal \<longleftrightarrow> ((*f* sqrt) x)\<^sup>2 \<in> Infinitesimal" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 70 | by (metis Infinitesimal_hrealpow pos2 power2_eq_square Infinitesimal_square_iff) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 71 | also have "... \<longleftrightarrow> x \<in> Infinitesimal" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 72 | by (simp add: assms hypreal_sqrt_gt_zero_pow2) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 73 | finally show ?thesis | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 74 | using mem_infmal_iff by blast | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 75 | qed | 
| 27468 | 76 | |
| 77 | lemma hypreal_sqrt_approx_zero2 [simp]: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 78 | "0 \<le> x \<Longrightarrow> (( *f* sqrt)(x) \<approx> 0) = (x \<approx> 0)" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 79 | by (auto simp add: order_le_less) | 
| 27468 | 80 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 81 | lemma hypreal_sqrt_gt_zero: "\<And>x. 0 < x \<Longrightarrow> 0 < ( *f* sqrt)(x)" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 82 | by transfer (simp add: real_sqrt_gt_zero) | 
| 27468 | 83 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 84 | lemma hypreal_sqrt_ge_zero: "0 \<le> x \<Longrightarrow> 0 \<le> ( *f* sqrt)(x)" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 85 | by (auto intro: hypreal_sqrt_gt_zero simp add: order_le_less) | 
| 27468 | 86 | |
| 70216 | 87 | lemma hypreal_sqrt_lessI: | 
| 88 | "\<And>x u. \<lbrakk>0 < u; x < u\<^sup>2\<rbrakk> \<Longrightarrow> ( *f* sqrt) x < u" | |
| 89 | proof transfer | |
| 90 | show "\<And>x u. \<lbrakk>0 < u; x < u\<^sup>2\<rbrakk> \<Longrightarrow> sqrt x < u" | |
| 91 | by (metis less_le real_sqrt_less_iff real_sqrt_pow2 real_sqrt_power) | |
| 92 | qed | |
| 93 | ||
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 94 | lemma hypreal_sqrt_hrabs [simp]: "\<And>x. ( *f* sqrt)(x\<^sup>2) = \<bar>x\<bar>" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 95 | by transfer simp | 
| 27468 | 96 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 97 | lemma hypreal_sqrt_hrabs2 [simp]: "\<And>x. ( *f* sqrt)(x*x) = \<bar>x\<bar>" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 98 | by transfer simp | 
| 27468 | 99 | |
| 100 | lemma hypreal_sqrt_hyperpow_hrabs [simp]: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 101 | "\<And>x. ( *f* sqrt)(x pow (hypnat_of_nat 2)) = \<bar>x\<bar>" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 102 | by transfer simp | 
| 27468 | 103 | |
| 104 | lemma star_sqrt_HFinite: "\<lbrakk>x \<in> HFinite; 0 \<le> x\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> HFinite" | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 105 | by (metis HFinite_square_iff hypreal_sqrt_pow2_iff power2_eq_square) | 
| 27468 | 106 | |
| 107 | lemma st_hypreal_sqrt: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 108 | assumes "x \<in> HFinite" "0 \<le> x" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 109 | shows "st(( *f* sqrt) x) = ( *f* sqrt)(st x)" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 110 | proof (rule power_inject_base) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 111 | show "st ((*f* sqrt) x) ^ Suc 1 = (*f* sqrt) (st x) ^ Suc 1" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 112 | using assms hypreal_sqrt_pow2_iff [of x] | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 113 | by (metis HFinite_square_iff hypreal_sqrt_hrabs2 power2_eq_square st_hrabs st_mult) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 114 | show "0 \<le> st ((*f* sqrt) x)" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 115 | by (simp add: assms hypreal_sqrt_ge_zero st_zero_le star_sqrt_HFinite) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 116 | show "0 \<le> (*f* sqrt) (st x)" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 117 | by (simp add: assms hypreal_sqrt_ge_zero st_zero_le) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 118 | qed | 
| 27468 | 119 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 120 | lemma hypreal_sqrt_sum_squares_ge1 [simp]: "\<And>x y. x \<le> ( *f* sqrt)(x\<^sup>2 + y\<^sup>2)" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 121 | by transfer (rule real_sqrt_sum_squares_ge1) | 
| 27468 | 122 | |
| 123 | lemma HFinite_hypreal_sqrt_imp_HFinite: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 124 | "\<lbrakk>0 \<le> x; ( *f* sqrt) x \<in> HFinite\<rbrakk> \<Longrightarrow> x \<in> HFinite" | 
| 75866 | 125 | by (metis HFinite_mult hypreal_sqrt_pow2_iff power2_eq_square) | 
| 27468 | 126 | |
| 127 | lemma HFinite_hypreal_sqrt_iff [simp]: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 128 | "0 \<le> x \<Longrightarrow> (( *f* sqrt) x \<in> HFinite) = (x \<in> HFinite)" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 129 | by (blast intro: star_sqrt_HFinite HFinite_hypreal_sqrt_imp_HFinite) | 
| 27468 | 130 | |
| 131 | lemma Infinitesimal_hypreal_sqrt: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 132 | "\<lbrakk>0 \<le> x; x \<in> Infinitesimal\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> Infinitesimal" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 133 | by (simp add: mem_infmal_iff) | 
| 27468 | 134 | |
| 135 | lemma Infinitesimal_hypreal_sqrt_imp_Infinitesimal: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 136 | "\<lbrakk>0 \<le> x; ( *f* sqrt) x \<in> Infinitesimal\<rbrakk> \<Longrightarrow> x \<in> Infinitesimal" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 137 | using hypreal_sqrt_approx_zero2 mem_infmal_iff by blast | 
| 27468 | 138 | |
| 139 | lemma Infinitesimal_hypreal_sqrt_iff [simp]: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 140 | "0 \<le> x \<Longrightarrow> (( *f* sqrt) x \<in> Infinitesimal) = (x \<in> Infinitesimal)" | 
| 27468 | 141 | by (blast intro: Infinitesimal_hypreal_sqrt_imp_Infinitesimal Infinitesimal_hypreal_sqrt) | 
| 142 | ||
| 143 | lemma HInfinite_hypreal_sqrt: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 144 | "\<lbrakk>0 \<le> x; x \<in> HInfinite\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> HInfinite" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 145 | by (simp add: HInfinite_HFinite_iff) | 
| 27468 | 146 | |
| 147 | lemma HInfinite_hypreal_sqrt_imp_HInfinite: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 148 | "\<lbrakk>0 \<le> x; ( *f* sqrt) x \<in> HInfinite\<rbrakk> \<Longrightarrow> x \<in> HInfinite" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 149 | using HFinite_hypreal_sqrt_iff HInfinite_HFinite_iff by blast | 
| 27468 | 150 | |
| 151 | lemma HInfinite_hypreal_sqrt_iff [simp]: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 152 | "0 \<le> x \<Longrightarrow> (( *f* sqrt) x \<in> HInfinite) = (x \<in> HInfinite)" | 
| 27468 | 153 | by (blast intro: HInfinite_hypreal_sqrt HInfinite_hypreal_sqrt_imp_HInfinite) | 
| 154 | ||
| 155 | lemma HFinite_exp [simp]: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 156 | "sumhr (0, whn, \<lambda>n. inverse (fact n) * x ^ n) \<in> HFinite" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 157 | unfolding sumhr_app star_zero_def starfun2_star_of atLeast0LessThan | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 158 | by (metis NSBseqD2 NSconvergent_NSBseq convergent_NSconvergent_iff summable_iff_convergent summable_exp) | 
| 27468 | 159 | |
| 160 | lemma exphr_zero [simp]: "exphr 0 = 1" | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 161 | proof - | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 162 | have "\<forall>x>1. 1 = sumhr (0, 1, \<lambda>n. inverse (fact n) * 0 ^ n) + sumhr (1, x, \<lambda>n. inverse (fact n) * 0 ^ n)" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 163 | unfolding sumhr_app by transfer (simp add: power_0_left) | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 164 | then have "sumhr (0, 1, \<lambda>n. inverse (fact n) * 0 ^ n) + sumhr (1, whn, \<lambda>n. inverse (fact n) * 0 ^ n) \<approx> 1" | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 165 | by auto | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 166 | then show ?thesis | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 167 | unfolding exphr_def | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 168 | using sumhr_split_add [OF hypnat_one_less_hypnat_omega] st_unique by auto | 
| 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 169 | qed | 
| 27468 | 170 | |
| 171 | lemma coshr_zero [simp]: "coshr 0 = 1" | |
| 70209 | 172 | proof - | 
| 173 | have "\<forall>x>1. 1 = sumhr (0, 1, \<lambda>n. cos_coeff n * 0 ^ n) + sumhr (1, x, \<lambda>n. cos_coeff n * 0 ^ n)" | |
| 174 | unfolding sumhr_app by transfer (simp add: power_0_left) | |
| 175 | then have "sumhr (0, 1, \<lambda>n. cos_coeff n * 0 ^ n) + sumhr (1, whn, \<lambda>n. cos_coeff n * 0 ^ n) \<approx> 1" | |
| 176 | by auto | |
| 177 | then show ?thesis | |
| 178 | unfolding coshr_def | |
| 179 | using sumhr_split_add [OF hypnat_one_less_hypnat_omega] st_unique by auto | |
| 180 | qed | |
| 27468 | 181 | |
| 61982 | 182 | lemma STAR_exp_zero_approx_one [simp]: "( *f* exp) (0::hypreal) \<approx> 1" | 
| 70209 | 183 | proof - | 
| 184 | have "(*f* exp) (0::real star) = 1" | |
| 185 | by transfer simp | |
| 186 | then show ?thesis | |
| 187 | by auto | |
| 188 | qed | |
| 27468 | 189 | |
| 70209 | 190 | lemma STAR_exp_Infinitesimal: | 
| 191 | assumes "x \<in> Infinitesimal" shows "( *f* exp) (x::hypreal) \<approx> 1" | |
| 192 | proof (cases "x = 0") | |
| 193 | case False | |
| 194 | have "NSDERIV exp 0 :> 1" | |
| 195 | by (metis DERIV_exp NSDERIV_DERIV_iff exp_zero) | |
| 196 | then have "((*f* exp) x - 1) / x \<approx> 1" | |
| 197 | using nsderiv_def False NSDERIVD2 assms by fastforce | |
| 198 | then have "(*f* exp) x - 1 \<approx> x" | |
| 199 | using NSDERIVD4 \<open>NSDERIV exp 0 :> 1\<close> assms by fastforce | |
| 200 | then show ?thesis | |
| 201 | by (meson Infinitesimal_approx approx_minus_iff approx_trans2 assms not_Infinitesimal_not_zero) | |
| 202 | qed auto | |
| 27468 | 203 | |
| 61982 | 204 | lemma STAR_exp_epsilon [simp]: "( *f* exp) \<epsilon> \<approx> 1" | 
| 70209 | 205 | by (auto intro: STAR_exp_Infinitesimal) | 
| 27468 | 206 | |
| 58656 | 207 | lemma STAR_exp_add: | 
| 70209 | 208 |   "\<And>(x::'a:: {banach,real_normed_field} star) y. ( *f* exp)(x + y) = ( *f* exp) x * ( *f* exp) y"
 | 
| 209 | by transfer (rule exp_add) | |
| 27468 | 210 | |
| 211 | lemma exphr_hypreal_of_real_exp_eq: "exphr x = hypreal_of_real (exp x)" | |
| 70209 | 212 | proof - | 
| 213 | have "(\<lambda>n. inverse (fact n) * x ^ n) sums exp x" | |
| 214 | using exp_converges [of x] by simp | |
| 215 | then have "(\<lambda>n. \<Sum>n<n. inverse (fact n) * x ^ n) \<longlonglongrightarrow>\<^sub>N\<^sub>S exp x" | |
| 216 | using NSsums_def sums_NSsums_iff by blast | |
| 217 | then have "hypreal_of_real (exp x) \<approx> sumhr (0, whn, \<lambda>n. inverse (fact n) * x ^ n)" | |
| 218 | unfolding starfunNat_sumr [symmetric] atLeast0LessThan | |
| 70228 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 paulson <lp15@cam.ac.uk> parents: 
70216diff
changeset | 219 | using HNatInfinite_whn NSLIMSEQ_def approx_sym by blast | 
| 70209 | 220 | then show ?thesis | 
| 221 | unfolding exphr_def using st_eq_approx_iff by auto | |
| 222 | qed | |
| 27468 | 223 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 224 | lemma starfun_exp_ge_add_one_self [simp]: "\<And>x::hypreal. 0 \<le> x \<Longrightarrow> (1 + x) \<le> ( *f* exp) x" | 
| 70209 | 225 | by transfer (rule exp_ge_add_one_self_aux) | 
| 27468 | 226 | |
| 70209 | 227 | text\<open>exp maps infinities to infinities\<close> | 
| 27468 | 228 | lemma starfun_exp_HInfinite: | 
| 70209 | 229 | fixes x :: hypreal | 
| 230 | assumes "x \<in> HInfinite" "0 \<le> x" | |
| 231 | shows "( *f* exp) x \<in> HInfinite" | |
| 232 | proof - | |
| 233 | have "x \<le> 1 + x" | |
| 234 | by simp | |
| 235 | also have "\<dots> \<le> (*f* exp) x" | |
| 236 | by (simp add: \<open>0 \<le> x\<close>) | |
| 237 | finally show ?thesis | |
| 238 | using HInfinite_ge_HInfinite assms by blast | |
| 239 | qed | |
| 27468 | 240 | |
| 58656 | 241 | lemma starfun_exp_minus: | 
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 242 |   "\<And>x::'a:: {banach,real_normed_field} star. ( *f* exp) (-x) = inverse(( *f* exp) x)"
 | 
| 70209 | 243 | by transfer (rule exp_minus) | 
| 27468 | 244 | |
| 70209 | 245 | text\<open>exp maps infinitesimals to infinitesimals\<close> | 
| 27468 | 246 | lemma starfun_exp_Infinitesimal: | 
| 70209 | 247 | fixes x :: hypreal | 
| 248 | assumes "x \<in> HInfinite" "x \<le> 0" | |
| 249 | shows "( *f* exp) x \<in> Infinitesimal" | |
| 250 | proof - | |
| 251 | obtain y where "x = -y" "y \<ge> 0" | |
| 252 | by (metis abs_of_nonpos assms(2) eq_abs_iff') | |
| 253 | then have "( *f* exp) y \<in> HInfinite" | |
| 254 | using HInfinite_minus_iff assms(1) starfun_exp_HInfinite by blast | |
| 255 | then show ?thesis | |
| 256 | by (simp add: HInfinite_inverse_Infinitesimal \<open>x = - y\<close> starfun_exp_minus) | |
| 257 | qed | |
| 27468 | 258 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 259 | lemma starfun_exp_gt_one [simp]: "\<And>x::hypreal. 0 < x \<Longrightarrow> 1 < ( *f* exp) x" | 
| 70209 | 260 | by transfer (rule exp_gt_one) | 
| 27468 | 261 | |
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 262 | abbreviation real_ln :: "real \<Rightarrow> real" where | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 263 | "real_ln \<equiv> ln" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 264 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 265 | lemma starfun_ln_exp [simp]: "\<And>x. ( *f* real_ln) (( *f* exp) x) = x" | 
| 70209 | 266 | by transfer (rule ln_exp) | 
| 27468 | 267 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 268 | lemma starfun_exp_ln_iff [simp]: "\<And>x. (( *f* exp)(( *f* real_ln) x) = x) = (0 < x)" | 
| 70209 | 269 | by transfer (rule exp_ln_iff) | 
| 27468 | 270 | |
| 70209 | 271 | lemma starfun_exp_ln_eq: "\<And>u x. ( *f* exp) u = x \<Longrightarrow> ( *f* real_ln) x = u" | 
| 272 | by transfer (rule ln_unique) | |
| 27468 | 273 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 274 | lemma starfun_ln_less_self [simp]: "\<And>x. 0 < x \<Longrightarrow> ( *f* real_ln) x < x" | 
| 70209 | 275 | by transfer (rule ln_less_self) | 
| 27468 | 276 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 277 | lemma starfun_ln_ge_zero [simp]: "\<And>x. 1 \<le> x \<Longrightarrow> 0 \<le> ( *f* real_ln) x" | 
| 70209 | 278 | by transfer (rule ln_ge_zero) | 
| 27468 | 279 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 280 | lemma starfun_ln_gt_zero [simp]: "\<And>x .1 < x \<Longrightarrow> 0 < ( *f* real_ln) x" | 
| 70209 | 281 | by transfer (rule ln_gt_zero) | 
| 27468 | 282 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 283 | lemma starfun_ln_not_eq_zero [simp]: "\<And>x. \<lbrakk>0 < x; x \<noteq> 1\<rbrakk> \<Longrightarrow> ( *f* real_ln) x \<noteq> 0" | 
| 70209 | 284 | by transfer simp | 
| 27468 | 285 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 286 | lemma starfun_ln_HFinite: "\<lbrakk>x \<in> HFinite; 1 \<le> x\<rbrakk> \<Longrightarrow> ( *f* real_ln) x \<in> HFinite" | 
| 70209 | 287 | by (metis HFinite_HInfinite_iff less_le_trans starfun_exp_HInfinite starfun_exp_ln_iff starfun_ln_ge_zero zero_less_one) | 
| 27468 | 288 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 289 | lemma starfun_ln_inverse: "\<And>x. 0 < x \<Longrightarrow> ( *f* real_ln) (inverse x) = -( *f* ln) x" | 
| 70209 | 290 | by transfer (rule ln_inverse) | 
| 27468 | 291 | |
| 292 | lemma starfun_abs_exp_cancel: "\<And>x. \<bar>( *f* exp) (x::hypreal)\<bar> = ( *f* exp) x" | |
| 70209 | 293 | by transfer (rule abs_exp_cancel) | 
| 27468 | 294 | |
| 295 | lemma starfun_exp_less_mono: "\<And>x y::hypreal. x < y \<Longrightarrow> ( *f* exp) x < ( *f* exp) y" | |
| 70209 | 296 | by transfer (rule exp_less_mono) | 
| 27468 | 297 | |
| 70209 | 298 | lemma starfun_exp_HFinite: | 
| 299 | fixes x :: hypreal | |
| 300 | assumes "x \<in> HFinite" | |
| 301 | shows "( *f* exp) x \<in> HFinite" | |
| 302 | proof - | |
| 303 | obtain u where "u \<in> \<real>" "\<bar>x\<bar> < u" | |
| 304 | using HFiniteD assms by force | |
| 305 | with assms have "\<bar>(*f* exp) x\<bar> < (*f* exp) u" | |
| 306 | using starfun_abs_exp_cancel starfun_exp_less_mono by auto | |
| 307 | with \<open>u \<in> \<real>\<close> show ?thesis | |
| 308 | by (force simp: HFinite_def Reals_eq_Standard) | |
| 309 | qed | |
| 27468 | 310 | |
| 311 | lemma starfun_exp_add_HFinite_Infinitesimal_approx: | |
| 70210 | 312 | fixes x :: hypreal | 
| 313 | shows "\<lbrakk>x \<in> Infinitesimal; z \<in> HFinite\<rbrakk> \<Longrightarrow> ( *f* exp) (z + x::hypreal) \<approx> ( *f* exp) z" | |
| 314 | using STAR_exp_Infinitesimal approx_mult2 starfun_exp_HFinite by (fastforce simp add: STAR_exp_add) | |
| 27468 | 315 | |
| 316 | lemma starfun_ln_HInfinite: | |
| 70210 | 317 | "\<lbrakk>x \<in> HInfinite; 0 < x\<rbrakk> \<Longrightarrow> ( *f* real_ln) x \<in> HInfinite" | 
| 318 | by (metis HInfinite_HFinite_iff starfun_exp_HFinite starfun_exp_ln_iff) | |
| 27468 | 319 | |
| 320 | lemma starfun_exp_HInfinite_Infinitesimal_disj: | |
| 70210 | 321 | fixes x :: hypreal | 
| 322 | shows "x \<in> HInfinite \<Longrightarrow> ( *f* exp) x \<in> HInfinite \<or> ( *f* exp) (x::hypreal) \<in> Infinitesimal" | |
| 323 | by (meson linear starfun_exp_HInfinite starfun_exp_Infinitesimal) | |
| 27468 | 324 | |
| 325 | lemma starfun_ln_HFinite_not_Infinitesimal: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 326 | "\<lbrakk>x \<in> HFinite - Infinitesimal; 0 < x\<rbrakk> \<Longrightarrow> ( *f* real_ln) x \<in> HFinite" | 
| 70210 | 327 | by (metis DiffD1 DiffD2 HInfinite_HFinite_iff starfun_exp_HInfinite_Infinitesimal_disj starfun_exp_ln_iff) | 
| 27468 | 328 | |
| 329 | (* we do proof by considering ln of 1/x *) | |
| 330 | lemma starfun_ln_Infinitesimal_HInfinite: | |
| 70210 | 331 | assumes "x \<in> Infinitesimal" "0 < x" | 
| 332 | shows "( *f* real_ln) x \<in> HInfinite" | |
| 333 | proof - | |
| 334 | have "inverse x \<in> HInfinite" | |
| 335 | using Infinitesimal_inverse_HInfinite assms by blast | |
| 336 | then show ?thesis | |
| 337 | using HInfinite_minus_iff assms(2) starfun_ln_HInfinite starfun_ln_inverse by fastforce | |
| 338 | qed | |
| 27468 | 339 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 340 | lemma starfun_ln_less_zero: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> ( *f* real_ln) x < 0" | 
| 70210 | 341 | by transfer (rule ln_less_zero) | 
| 27468 | 342 | |
| 343 | lemma starfun_ln_Infinitesimal_less_zero: | |
| 70210 | 344 | "\<lbrakk>x \<in> Infinitesimal; 0 < x\<rbrakk> \<Longrightarrow> ( *f* real_ln) x < 0" | 
| 345 | by (auto intro!: starfun_ln_less_zero simp add: Infinitesimal_def) | |
| 27468 | 346 | |
| 347 | lemma starfun_ln_HInfinite_gt_zero: | |
| 70210 | 348 | "\<lbrakk>x \<in> HInfinite; 0 < x\<rbrakk> \<Longrightarrow> 0 < ( *f* real_ln) x" | 
| 349 | by (auto intro!: starfun_ln_gt_zero simp add: HInfinite_def) | |
| 27468 | 350 | |
| 351 | ||
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 352 | lemma HFinite_sin [simp]: "sumhr (0, whn, \<lambda>n. sin_coeff n * x ^ n) \<in> HFinite" | 
| 70210 | 353 | proof - | 
| 354 | have "summable (\<lambda>i. sin_coeff i * x ^ i)" | |
| 355 | using summable_norm_sin [of x] by (simp add: summable_rabs_cancel) | |
| 356 | then have "(*f* (\<lambda>n. \<Sum>n<n. sin_coeff n * x ^ n)) whn \<in> HFinite" | |
| 70228 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 paulson <lp15@cam.ac.uk> parents: 
70216diff
changeset | 357 | unfolding summable_sums_iff sums_NSsums_iff NSsums_def NSLIMSEQ_def | 
| 70210 | 358 | using HFinite_star_of HNatInfinite_whn approx_HFinite approx_sym by blast | 
| 359 | then show ?thesis | |
| 360 | unfolding sumhr_app | |
| 361 | by (simp only: star_zero_def starfun2_star_of atLeast0LessThan) | |
| 362 | qed | |
| 27468 | 363 | |
| 364 | lemma STAR_sin_zero [simp]: "( *f* sin) 0 = 0" | |
| 70210 | 365 | by transfer (rule sin_zero) | 
| 27468 | 366 | |
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 367 | lemma STAR_sin_Infinitesimal [simp]: | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 368 |   fixes x :: "'a::{real_normed_field,banach} star"
 | 
| 70210 | 369 | assumes "x \<in> Infinitesimal" | 
| 370 | shows "( *f* sin) x \<approx> x" | |
| 371 | proof (cases "x = 0") | |
| 372 | case False | |
| 373 | have "NSDERIV sin 0 :> 1" | |
| 374 | by (metis DERIV_sin NSDERIV_DERIV_iff cos_zero) | |
| 375 | then have "(*f* sin) x / x \<approx> 1" | |
| 376 | using False NSDERIVD2 assms by fastforce | |
| 377 | with assms show ?thesis | |
| 378 | unfolding star_one_def | |
| 379 | by (metis False Infinitesimal_approx Infinitesimal_ratio approx_star_of_HFinite) | |
| 380 | qed auto | |
| 27468 | 381 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 382 | lemma HFinite_cos [simp]: "sumhr (0, whn, \<lambda>n. cos_coeff n * x ^ n) \<in> HFinite" | 
| 70210 | 383 | proof - | 
| 384 | have "summable (\<lambda>i. cos_coeff i * x ^ i)" | |
| 385 | using summable_norm_cos [of x] by (simp add: summable_rabs_cancel) | |
| 386 | then have "(*f* (\<lambda>n. \<Sum>n<n. cos_coeff n * x ^ n)) whn \<in> HFinite" | |
| 70228 
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
 paulson <lp15@cam.ac.uk> parents: 
70216diff
changeset | 387 | unfolding summable_sums_iff sums_NSsums_iff NSsums_def NSLIMSEQ_def | 
| 70210 | 388 | using HFinite_star_of HNatInfinite_whn approx_HFinite approx_sym by blast | 
| 389 | then show ?thesis | |
| 390 | unfolding sumhr_app | |
| 391 | by (simp only: star_zero_def starfun2_star_of atLeast0LessThan) | |
| 392 | qed | |
| 27468 | 393 | |
| 394 | lemma STAR_cos_zero [simp]: "( *f* cos) 0 = 1" | |
| 70210 | 395 | by transfer (rule cos_zero) | 
| 27468 | 396 | |
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 397 | lemma STAR_cos_Infinitesimal [simp]: | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 398 |   fixes x :: "'a::{real_normed_field,banach} star"
 | 
| 70210 | 399 | assumes "x \<in> Infinitesimal" | 
| 400 | shows "( *f* cos) x \<approx> 1" | |
| 401 | proof (cases "x = 0") | |
| 402 | case False | |
| 403 | have "NSDERIV cos 0 :> 0" | |
| 404 | by (metis DERIV_cos NSDERIV_DERIV_iff add.inverse_neutral sin_zero) | |
| 405 | then have "(*f* cos) x - 1 \<approx> 0" | |
| 406 | using NSDERIV_approx assms by fastforce | |
| 407 | with assms show ?thesis | |
| 408 | using approx_minus_iff by blast | |
| 409 | qed auto | |
| 27468 | 410 | |
| 411 | lemma STAR_tan_zero [simp]: "( *f* tan) 0 = 0" | |
| 70210 | 412 | by transfer (rule tan_zero) | 
| 27468 | 413 | |
| 70210 | 414 | lemma STAR_tan_Infinitesimal [simp]: | 
| 415 | assumes "x \<in> Infinitesimal" | |
| 416 | shows "( *f* tan) x \<approx> x" | |
| 417 | proof (cases "x = 0") | |
| 418 | case False | |
| 419 | have "NSDERIV tan 0 :> 1" | |
| 420 | using DERIV_tan [of 0] by (simp add: NSDERIV_DERIV_iff) | |
| 421 | then have "(*f* tan) x / x \<approx> 1" | |
| 422 | using False NSDERIVD2 assms by fastforce | |
| 423 | with assms show ?thesis | |
| 424 | unfolding star_one_def | |
| 425 | by (metis False Infinitesimal_approx Infinitesimal_ratio approx_star_of_HFinite) | |
| 426 | qed auto | |
| 27468 | 427 | |
| 428 | lemma STAR_sin_cos_Infinitesimal_mult: | |
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 429 |   fixes x :: "'a::{real_normed_field,banach} star"
 | 
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 430 | shows "x \<in> Infinitesimal \<Longrightarrow> ( *f* sin) x * ( *f* cos) x \<approx> x" | 
| 70210 | 431 | using approx_mult_HFinite [of "( *f* sin) x" _ "( *f* cos) x" 1] | 
| 432 | by (simp add: Infinitesimal_subset_HFinite [THEN subsetD]) | |
| 27468 | 433 | |
| 434 | lemma HFinite_pi: "hypreal_of_real pi \<in> HFinite" | |
| 70210 | 435 | by simp | 
| 27468 | 436 | |
| 437 | ||
| 438 | lemma STAR_sin_Infinitesimal_divide: | |
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 439 |   fixes x :: "'a::{real_normed_field,banach} star"
 | 
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 440 | shows "\<lbrakk>x \<in> Infinitesimal; x \<noteq> 0\<rbrakk> \<Longrightarrow> ( *f* sin) x/x \<approx> 1" | 
| 70210 | 441 | using DERIV_sin [of "0::'a"] | 
| 442 | by (simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def) | |
| 27468 | 443 | |
| 70210 | 444 | subsection \<open>Proving $\sin* (1/n) \times 1/(1/n) \approx 1$ for $n = \infty$ \<close> | 
| 27468 | 445 | |
| 446 | lemma lemma_sin_pi: | |
| 70210 | 447 | "n \<in> HNatInfinite | 
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 448 | \<Longrightarrow> ( *f* sin) (inverse (hypreal_of_hypnat n))/(inverse (hypreal_of_hypnat n)) \<approx> 1" | 
| 70210 | 449 | by (simp add: STAR_sin_Infinitesimal_divide zero_less_HNatInfinite) | 
| 27468 | 450 | |
| 451 | lemma STAR_sin_inverse_HNatInfinite: | |
| 452 | "n \<in> HNatInfinite | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 453 | \<Longrightarrow> ( *f* sin) (inverse (hypreal_of_hypnat n)) * hypreal_of_hypnat n \<approx> 1" | 
| 70210 | 454 | by (metis field_class.field_divide_inverse inverse_inverse_eq lemma_sin_pi) | 
| 27468 | 455 | |
| 456 | lemma Infinitesimal_pi_divide_HNatInfinite: | |
| 457 | "N \<in> HNatInfinite | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 458 | \<Longrightarrow> hypreal_of_real pi/(hypreal_of_hypnat N) \<in> Infinitesimal" | 
| 70210 | 459 | by (simp add: Infinitesimal_HFinite_mult2 field_class.field_divide_inverse) | 
| 27468 | 460 | |
| 461 | lemma pi_divide_HNatInfinite_not_zero [simp]: | |
| 70210 | 462 | "N \<in> HNatInfinite \<Longrightarrow> hypreal_of_real pi/(hypreal_of_hypnat N) \<noteq> 0" | 
| 463 | by (simp add: zero_less_HNatInfinite) | |
| 27468 | 464 | |
| 465 | lemma STAR_sin_pi_divide_HNatInfinite_approx_pi: | |
| 70210 | 466 | assumes "n \<in> HNatInfinite" | 
| 467 | shows "(*f* sin) (hypreal_of_real pi / hypreal_of_hypnat n) * hypreal_of_hypnat n \<approx> | |
| 468 | hypreal_of_real pi" | |
| 469 | proof - | |
| 470 | have "(*f* sin) (hypreal_of_real pi / hypreal_of_hypnat n) / (hypreal_of_real pi / hypreal_of_hypnat n) \<approx> 1" | |
| 471 | using Infinitesimal_pi_divide_HNatInfinite STAR_sin_Infinitesimal_divide assms pi_divide_HNatInfinite_not_zero by blast | |
| 472 | then have "hypreal_of_hypnat n * star_of sin \<star> (hypreal_of_real pi / hypreal_of_hypnat n) / hypreal_of_real pi \<approx> 1" | |
| 473 | by (simp add: mult.commute starfun_def) | |
| 474 | then show ?thesis | |
| 475 | apply (simp add: starfun_def field_simps) | |
| 476 | by (metis (no_types, lifting) approx_mult_subst_star_of approx_refl mult_cancel_right1 nonzero_eq_divide_eq pi_neq_zero star_of_eq_0) | |
| 477 | qed | |
| 27468 | 478 | |
| 479 | lemma STAR_sin_pi_divide_HNatInfinite_approx_pi2: | |
| 480 | "n \<in> HNatInfinite | |
| 70210 | 481 | \<Longrightarrow> hypreal_of_hypnat n * ( *f* sin) (hypreal_of_real pi/(hypreal_of_hypnat n)) \<approx> hypreal_of_real pi" | 
| 482 | by (metis STAR_sin_pi_divide_HNatInfinite_approx_pi mult.commute) | |
| 27468 | 483 | |
| 484 | lemma starfunNat_pi_divide_n_Infinitesimal: | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 485 | "N \<in> HNatInfinite \<Longrightarrow> ( *f* (\<lambda>x. pi / real x)) N \<in> Infinitesimal" | 
| 70210 | 486 | by (simp add: Infinitesimal_HFinite_mult2 divide_inverse starfunNat_real_of_nat) | 
| 27468 | 487 | |
| 488 | lemma STAR_sin_pi_divide_n_approx: | |
| 70210 | 489 | assumes "N \<in> HNatInfinite" | 
| 490 | shows "( *f* sin) (( *f* (\<lambda>x. pi / real x)) N) \<approx> hypreal_of_real pi/(hypreal_of_hypnat N)" | |
| 491 | proof - | |
| 492 | have "\<exists>s. (*f* sin) ((*f* (\<lambda>n. pi / real n)) N) \<approx> s \<and> hypreal_of_real pi / hypreal_of_hypnat N \<approx> s" | |
| 493 | by (metis (lifting) Infinitesimal_approx Infinitesimal_pi_divide_HNatInfinite STAR_sin_Infinitesimal assms starfunNat_pi_divide_n_Infinitesimal) | |
| 494 | then show ?thesis | |
| 495 | by (meson approx_trans2) | |
| 496 | qed | |
| 27468 | 497 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 498 | lemma NSLIMSEQ_sin_pi: "(\<lambda>n. real n * sin (pi / real n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S pi" | 
| 70210 | 499 | proof - | 
| 500 | have *: "hypreal_of_hypnat N * (*f* sin) ((*f* (\<lambda>x. pi / real x)) N) \<approx> hypreal_of_real pi" | |
| 501 | if "N \<in> HNatInfinite" | |
| 502 | for N :: "nat star" | |
| 503 | using that | |
| 504 | by simp (metis STAR_sin_pi_divide_HNatInfinite_approx_pi2 starfunNat_real_of_nat) | |
| 505 | show ?thesis | |
| 506 | by (simp add: NSLIMSEQ_def starfunNat_real_of_nat) (metis * starfun_o2) | |
| 507 | qed | |
| 27468 | 508 | |
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 509 | lemma NSLIMSEQ_cos_one: "(\<lambda>n. cos (pi / real n))\<longlonglongrightarrow>\<^sub>N\<^sub>S 1" | 
| 70210 | 510 | proof - | 
| 511 | have "(*f* cos) ((*f* (\<lambda>x. pi / real x)) N) \<approx> 1" | |
| 512 | if "N \<in> HNatInfinite" for N | |
| 513 | using that STAR_cos_Infinitesimal starfunNat_pi_divide_n_Infinitesimal by blast | |
| 514 | then show ?thesis | |
| 515 | by (simp add: NSLIMSEQ_def) (metis STAR_cos_Infinitesimal starfunNat_pi_divide_n_Infinitesimal starfun_o2) | |
| 516 | qed | |
| 27468 | 517 | |
| 518 | lemma NSLIMSEQ_sin_cos_pi: | |
| 70210 | 519 | "(\<lambda>n. real n * sin (pi / real n) * cos (pi / real n)) \<longlonglongrightarrow>\<^sub>N\<^sub>S pi" | 
| 520 | using NSLIMSEQ_cos_one NSLIMSEQ_mult NSLIMSEQ_sin_pi by force | |
| 27468 | 521 | |
| 522 | ||
| 69597 | 523 | text\<open>A familiar approximation to \<^term>\<open>cos x\<close> when \<^term>\<open>x\<close> is small\<close> | 
| 27468 | 524 | |
| 525 | lemma STAR_cos_Infinitesimal_approx: | |
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 526 |   fixes x :: "'a::{real_normed_field,banach} star"
 | 
| 70208 
65b3bfc565b5
removal of ASCII connectives; some de-applying
 paulson <lp15@cam.ac.uk> parents: 
69597diff
changeset | 527 | shows "x \<in> Infinitesimal \<Longrightarrow> ( *f* cos) x \<approx> 1 - x\<^sup>2" | 
| 70210 | 528 | by (metis Infinitesimal_square_iff STAR_cos_Infinitesimal approx_diff approx_sym diff_zero mem_infmal_iff power2_eq_square) | 
| 27468 | 529 | |
| 530 | lemma STAR_cos_Infinitesimal_approx2: | |
| 70216 | 531 | fixes x :: hypreal | 
| 70210 | 532 | assumes "x \<in> Infinitesimal" | 
| 533 | shows "( *f* cos) x \<approx> 1 - (x\<^sup>2)/2" | |
| 534 | proof - | |
| 535 | have "1 \<approx> 1 - x\<^sup>2 / 2" | |
| 536 | using assms | |
| 537 | by (auto intro: Infinitesimal_SReal_divide simp add: Infinitesimal_approx_minus [symmetric] numeral_2_eq_2) | |
| 538 | then show ?thesis | |
| 539 | using STAR_cos_Infinitesimal approx_trans assms by blast | |
| 540 | qed | |
| 27468 | 541 | |
| 542 | end |