# HG changeset patch # User paulson # Date 1556471207 -3600 # Node ID ab29bd01b8b2ca624e77fc749fe026112f39b995 # Parent 65b3bfc565b531610b42619143c82a8dc6590a66 further de-applying diff -r 65b3bfc565b5 -r ab29bd01b8b2 src/HOL/Nonstandard_Analysis/HTranscendental.thy --- a/src/HOL/Nonstandard_Analysis/HTranscendental.thy Sun Apr 28 16:50:19 2019 +0100 +++ b/src/HOL/Nonstandard_Analysis/HTranscendental.thy Sun Apr 28 18:06:47 2019 +0100 @@ -165,125 +165,144 @@ qed lemma coshr_zero [simp]: "coshr 0 = 1" -apply (simp add: coshr_def sumhr_split_add - [OF hypnat_one_less_hypnat_omega, symmetric]) -apply (rule st_unique, simp) -apply (rule subst [where P="\x. 1 \ x", OF _ approx_refl]) -apply (rule rev_mp [OF hypnat_one_less_hypnat_omega]) -apply (rule_tac x="whn" in spec) -apply (unfold sumhr_app, transfer, simp add: cos_coeff_def power_0_left) -done + proof - + have "\x>1. 1 = sumhr (0, 1, \n. cos_coeff n * 0 ^ n) + sumhr (1, x, \n. cos_coeff n * 0 ^ n)" + unfolding sumhr_app by transfer (simp add: power_0_left) + then have "sumhr (0, 1, \n. cos_coeff n * 0 ^ n) + sumhr (1, whn, \n. cos_coeff n * 0 ^ n) \ 1" + by auto + then show ?thesis + unfolding coshr_def + using sumhr_split_add [OF hypnat_one_less_hypnat_omega] st_unique by auto +qed lemma STAR_exp_zero_approx_one [simp]: "( *f* exp) (0::hypreal) \ 1" -apply (subgoal_tac "( *f* exp) (0::hypreal) = 1", simp) -apply (transfer, simp) -done + proof - + have "(*f* exp) (0::real star) = 1" + by transfer simp + then show ?thesis + by auto +qed -lemma STAR_exp_Infinitesimal: "x \ Infinitesimal \ ( *f* exp) (x::hypreal) \ 1" -apply (case_tac "x = 0") -apply (cut_tac [2] x = 0 in DERIV_exp) -apply (auto simp add: NSDERIV_DERIV_iff [symmetric] nsderiv_def) -apply (drule_tac x = x in bspec, auto) -apply (drule_tac c = x in approx_mult1) -apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] - simp add: mult.assoc) -apply (rule approx_add_right_cancel [where d="-1"]) -apply (rule approx_sym [THEN [2] approx_trans2]) -apply (auto simp add: mem_infmal_iff) -done +lemma STAR_exp_Infinitesimal: + assumes "x \ Infinitesimal" shows "( *f* exp) (x::hypreal) \ 1" +proof (cases "x = 0") + case False + have "NSDERIV exp 0 :> 1" + by (metis DERIV_exp NSDERIV_DERIV_iff exp_zero) + then have "((*f* exp) x - 1) / x \ 1" + using nsderiv_def False NSDERIVD2 assms by fastforce + then have "(*f* exp) x - 1 \ x" + using NSDERIVD4 \NSDERIV exp 0 :> 1\ assms by fastforce + then show ?thesis + by (meson Infinitesimal_approx approx_minus_iff approx_trans2 assms not_Infinitesimal_not_zero) +qed auto lemma STAR_exp_epsilon [simp]: "( *f* exp) \ \ 1" -by (auto intro: STAR_exp_Infinitesimal) + by (auto intro: STAR_exp_Infinitesimal) lemma STAR_exp_add: - "!!(x::'a:: {banach,real_normed_field} star) y. ( *f* exp)(x + y) = ( *f* exp) x * ( *f* exp) y" -by transfer (rule exp_add) + "\(x::'a:: {banach,real_normed_field} star) y. ( *f* exp)(x + y) = ( *f* exp) x * ( *f* exp) y" + by transfer (rule exp_add) lemma exphr_hypreal_of_real_exp_eq: "exphr x = hypreal_of_real (exp x)" -apply (simp add: exphr_def) -apply (rule st_unique, simp) -apply (subst starfunNat_sumr [symmetric]) -unfolding atLeast0LessThan -apply (rule NSLIMSEQ_D [THEN approx_sym]) -apply (rule LIMSEQ_NSLIMSEQ) -apply (subst sums_def [symmetric]) -apply (cut_tac exp_converges [where x=x], simp) -apply (rule HNatInfinite_whn) -done +proof - + have "(\n. inverse (fact n) * x ^ n) sums exp x" + using exp_converges [of x] by simp + then have "(\n. \n\<^sub>N\<^sub>S exp x" + using NSsums_def sums_NSsums_iff by blast + then have "hypreal_of_real (exp x) \ sumhr (0, whn, \n. inverse (fact n) * x ^ n)" + unfolding starfunNat_sumr [symmetric] atLeast0LessThan + using HNatInfinite_whn NSLIMSEQ_iff approx_sym by blast + then show ?thesis + unfolding exphr_def using st_eq_approx_iff by auto +qed lemma starfun_exp_ge_add_one_self [simp]: "\x::hypreal. 0 \ x \ (1 + x) \ ( *f* exp) x" -by transfer (rule exp_ge_add_one_self_aux) + by transfer (rule exp_ge_add_one_self_aux) -(* exp (oo) is infinite *) +text\exp maps infinities to infinities\ lemma starfun_exp_HInfinite: - "\x \ HInfinite; 0 \ x\ \ ( *f* exp) (x::hypreal) \ HInfinite" -apply (frule starfun_exp_ge_add_one_self) -apply (rule HInfinite_ge_HInfinite, assumption) -apply (rule order_trans [of _ "1+x"], auto) -done + fixes x :: hypreal + assumes "x \ HInfinite" "0 \ x" + shows "( *f* exp) x \ HInfinite" +proof - + have "x \ 1 + x" + by simp + also have "\ \ (*f* exp) x" + by (simp add: \0 \ x\) + finally show ?thesis + using HInfinite_ge_HInfinite assms by blast +qed lemma starfun_exp_minus: "\x::'a:: {banach,real_normed_field} star. ( *f* exp) (-x) = inverse(( *f* exp) x)" -by transfer (rule exp_minus) + by transfer (rule exp_minus) -(* exp (-oo) is infinitesimal *) +text\exp maps infinitesimals to infinitesimals\ lemma starfun_exp_Infinitesimal: - "\x \ HInfinite; x \ 0\ \ ( *f* exp) (x::hypreal) \ Infinitesimal" -apply (subgoal_tac "\y. x = - y") -apply (rule_tac [2] x = "- x" in exI) -apply (auto intro!: HInfinite_inverse_Infinitesimal starfun_exp_HInfinite - simp add: starfun_exp_minus HInfinite_minus_iff) -done + fixes x :: hypreal + assumes "x \ HInfinite" "x \ 0" + shows "( *f* exp) x \ Infinitesimal" +proof - + obtain y where "x = -y" "y \ 0" + by (metis abs_of_nonpos assms(2) eq_abs_iff') + then have "( *f* exp) y \ HInfinite" + using HInfinite_minus_iff assms(1) starfun_exp_HInfinite by blast + then show ?thesis + by (simp add: HInfinite_inverse_Infinitesimal \x = - y\ starfun_exp_minus) +qed lemma starfun_exp_gt_one [simp]: "\x::hypreal. 0 < x \ 1 < ( *f* exp) x" -by transfer (rule exp_gt_one) + by transfer (rule exp_gt_one) abbreviation real_ln :: "real \ real" where "real_ln \ ln" lemma starfun_ln_exp [simp]: "\x. ( *f* real_ln) (( *f* exp) x) = x" -by transfer (rule ln_exp) + by transfer (rule ln_exp) lemma starfun_exp_ln_iff [simp]: "\x. (( *f* exp)(( *f* real_ln) x) = x) = (0 < x)" -by transfer (rule exp_ln_iff) + by transfer (rule exp_ln_iff) -lemma starfun_exp_ln_eq: "!!u x. ( *f* exp) u = x \ ( *f* real_ln) x = u" -by transfer (rule ln_unique) +lemma starfun_exp_ln_eq: "\u x. ( *f* exp) u = x \ ( *f* real_ln) x = u" + by transfer (rule ln_unique) lemma starfun_ln_less_self [simp]: "\x. 0 < x \ ( *f* real_ln) x < x" -by transfer (rule ln_less_self) + by transfer (rule ln_less_self) lemma starfun_ln_ge_zero [simp]: "\x. 1 \ x \ 0 \ ( *f* real_ln) x" -by transfer (rule ln_ge_zero) + by transfer (rule ln_ge_zero) lemma starfun_ln_gt_zero [simp]: "\x .1 < x \ 0 < ( *f* real_ln) x" -by transfer (rule ln_gt_zero) + by transfer (rule ln_gt_zero) lemma starfun_ln_not_eq_zero [simp]: "\x. \0 < x; x \ 1\ \ ( *f* real_ln) x \ 0" -by transfer simp + by transfer simp lemma starfun_ln_HFinite: "\x \ HFinite; 1 \ x\ \ ( *f* real_ln) x \ HFinite" -apply (rule HFinite_bounded) -apply assumption -apply (simp_all add: starfun_ln_less_self order_less_imp_le) -done + by (metis HFinite_HInfinite_iff less_le_trans starfun_exp_HInfinite starfun_exp_ln_iff starfun_ln_ge_zero zero_less_one) lemma starfun_ln_inverse: "\x. 0 < x \ ( *f* real_ln) (inverse x) = -( *f* ln) x" -by transfer (rule ln_inverse) + by transfer (rule ln_inverse) lemma starfun_abs_exp_cancel: "\x. \( *f* exp) (x::hypreal)\ = ( *f* exp) x" -by transfer (rule abs_exp_cancel) + by transfer (rule abs_exp_cancel) lemma starfun_exp_less_mono: "\x y::hypreal. x < y \ ( *f* exp) x < ( *f* exp) y" -by transfer (rule exp_less_mono) + by transfer (rule exp_less_mono) -lemma starfun_exp_HFinite: "x \ HFinite \ ( *f* exp) (x::hypreal) \ HFinite" -apply (auto simp add: HFinite_def, rename_tac u) -apply (rule_tac x="( *f* exp) u" in rev_bexI) -apply (simp add: Reals_eq_Standard) -apply (simp add: starfun_abs_exp_cancel) -apply (simp add: starfun_exp_less_mono) -done +lemma starfun_exp_HFinite: + fixes x :: hypreal + assumes "x \ HFinite" + shows "( *f* exp) x \ HFinite" +proof - + obtain u where "u \ \" "\x\ < u" + using HFiniteD assms by force + with assms have "\(*f* exp) x\ < (*f* exp) u" + using starfun_abs_exp_cancel starfun_exp_less_mono by auto + with \u \ \\ show ?thesis + by (force simp: HFinite_def Reals_eq_Standard) +qed lemma starfun_exp_add_HFinite_Infinitesimal_approx: "\x \ Infinitesimal; z \ HFinite\ \ ( *f* exp) (z + x::hypreal) \ ( *f* exp) z" @@ -307,7 +326,7 @@ apply (auto intro: starfun_exp_HInfinite starfun_exp_Infinitesimal) done -(* check out this proof!!! *) +(* check out this proof\! *) lemma starfun_ln_HFinite_not_Infinitesimal: "\x \ HFinite - Infinitesimal; 0 < x\ \ ( *f* real_ln) x \ HFinite" apply (rule ccontr, drule HInfinite_HFinite_iff [THEN iffD2])