# HG changeset patch # User paulson # Date 1556654661 -3600 # Node ID bde8ccb73fd2821c4ef18fb9e1df1d1f8bc4343b # Parent b21efbf6429208a8c6536dba13f92f19b3795ea5# Parent bca14283e1a81d6a2f561410b259977ae78a500e merged diff -r b21efbf64292 -r bde8ccb73fd2 src/HOL/Nonstandard_Analysis/NSA.thy --- a/src/HOL/Nonstandard_Analysis/NSA.thy Tue Apr 30 17:03:32 2019 +0100 +++ b/src/HOL/Nonstandard_Analysis/NSA.thy Tue Apr 30 21:04:21 2019 +0100 @@ -6,7 +6,7 @@ section \Infinite Numbers, Infinitesimals, Infinitely Close Relation\ theory NSA - imports HyperDef "HOL-Library.Lub_Glb" + imports HyperDef "HOL-Library.Lub_Glb" begin definition hnorm :: "'a::real_normed_vector star \ real star" @@ -162,37 +162,11 @@ by (simp add: Reals_eq_Standard Standard_def) lemma inv_hypreal_of_real_image: "inv hypreal_of_real ` \ = UNIV" - apply (auto simp add: SReal_def) - apply (rule inj_star_of [THEN inv_f_f, THEN subst], blast) - done - -lemma SReal_hypreal_of_real_image: "\x. x \ P \ P \ \ \ \Q. P = hypreal_of_real ` Q" - unfolding SReal_def image_def by blast + by (simp add: Reals_eq_Standard Standard_def inj_star_of) lemma SReal_dense: "x \ \ \ y \ \ \ x < y \ \r \ Reals. x < r \ r < y" for x y :: hypreal - apply (auto simp: SReal_def) - apply (drule dense) - apply auto - done - - -text \Completeness of Reals, but both lemmas are unused.\ - -lemma SReal_sup_lemma: - "P \ \ \ (\x \ P. y < x) = (\X. hypreal_of_real X \ P \ y < hypreal_of_real X)" - by (blast dest!: SReal_iff [THEN iffD1]) - -lemma SReal_sup_lemma2: - "P \ \ \ \x. x \ P \ \y \ Reals. \x \ P. x < y \ - (\X. X \ {w. hypreal_of_real w \ P}) \ - (\Y. \X \ {w. hypreal_of_real w \ P}. X < Y)" - apply (rule conjI) - apply (fast dest!: SReal_iff [THEN iffD1]) - apply (auto, frule subsetD, assumption) - apply (drule SReal_iff [THEN iffD1]) - apply (auto, rule_tac x = ya in exI, auto) - done + using dense by (fastforce simp add: SReal_def) subsection \Set of Finite Elements is a Subring of the Extended Reals\ @@ -211,11 +185,7 @@ by (simp add: HFinite_def) lemma HFinite_star_of [simp]: "star_of x \ HFinite" - apply (simp add: HFinite_def) - apply (rule_tac x="star_of (norm x) + 1" in bexI) - apply (transfer, simp) - apply (blast intro: Reals_add SReal_hypreal_of_real Reals_1) - done + by (simp add: HFinite_def) (metis SReal_hypreal_of_real gt_ex star_of_less star_of_norm) lemma SReal_subset_HFinite: "(\::hypreal set) \ HFinite" by (auto simp add: SReal_def) @@ -244,16 +214,22 @@ lemma hrealpow_HFinite: "x \ HFinite \ x ^ n \ HFinite" for x :: "'a::{real_normed_algebra,monoid_mult} star" - by (induct n) (auto simp add: power_Suc intro: HFinite_mult) + by (induct n) (auto intro: HFinite_mult) -lemma HFinite_bounded: "x \ HFinite \ y \ x \ 0 \ y \ y \ HFinite" - for x y :: hypreal - apply (cases "x \ 0") - apply (drule_tac y = x in order_trans) - apply (drule_tac [2] order_antisym) - apply (auto simp add: linorder_not_le) - apply (auto intro: order_le_less_trans simp add: abs_if HFinite_def) - done +lemma HFinite_bounded: + fixes x y :: hypreal + assumes "x \ HFinite" and y: "y \ x" "0 \ y" shows "y \ HFinite" +proof (cases "x \ 0") + case True + then have "y = 0" + using y by auto + then show ?thesis + by simp +next + case False + then show ?thesis + using assms le_less_trans by (auto simp: HFinite_def) +qed subsection \Set of Infinitesimals is a Subring of the Hyperreals\ @@ -273,12 +249,19 @@ lemma Infinitesimal_zero [iff]: "0 \ Infinitesimal" by (simp add: Infinitesimal_def) -lemma Infinitesimal_add: "x \ Infinitesimal \ y \ Infinitesimal \ x + y \ Infinitesimal" - apply (rule InfinitesimalI) - apply (rule field_sum_of_halves [THEN subst]) - apply (drule half_gt_zero) - apply (blast intro: hnorm_add_less SReal_divide_numeral dest: InfinitesimalD) - done +lemma Infinitesimal_add: + assumes "x \ Infinitesimal" "y \ Infinitesimal" + shows "x + y \ Infinitesimal" +proof (rule InfinitesimalI) + show "hnorm (x + y) < r" + if "r \ \" and "0 < r" for r :: "real star" + proof - + have "hnorm x < r/2" "hnorm y < r/2" + using InfinitesimalD SReal_divide_numeral assms half_gt_zero that by blast+ + then show ?thesis + using hnorm_add_less by fastforce + qed +qed lemma Infinitesimal_minus_iff [simp]: "- x \ Infinitesimal \ x \ Infinitesimal" by (simp add: Infinitesimal_def) @@ -297,83 +280,103 @@ lemma Infinitesimal_diff: "x \ Infinitesimal \ y \ Infinitesimal \ x - y \ Infinitesimal" using Infinitesimal_add [of x "- y"] by simp -lemma Infinitesimal_mult: "x \ Infinitesimal \ y \ Infinitesimal \ x * y \ Infinitesimal" - for x y :: "'a::real_normed_algebra star" - apply (rule InfinitesimalI) - apply (subgoal_tac "hnorm (x * y) < 1 * r") - apply (simp only: mult_1) - apply (rule hnorm_mult_less) - apply (simp_all add: InfinitesimalD) - done +lemma Infinitesimal_mult: + fixes x y :: "'a::real_normed_algebra star" + assumes "x \ Infinitesimal" "y \ Infinitesimal" + shows "x * y \ Infinitesimal" + proof (rule InfinitesimalI) + show "hnorm (x * y) < r" + if "r \ \" and "0 < r" for r :: "real star" + proof - + have "hnorm x < 1" "hnorm y < r" + using assms that by (auto simp add: InfinitesimalD) + then show ?thesis + using hnorm_mult_less by fastforce + qed +qed -lemma Infinitesimal_HFinite_mult: "x \ Infinitesimal \ y \ HFinite \ x * y \ Infinitesimal" - for x y :: "'a::real_normed_algebra star" - apply (rule InfinitesimalI) - apply (drule HFiniteD, clarify) - apply (subgoal_tac "0 < t") - apply (subgoal_tac "hnorm (x * y) < (r / t) * t", simp) - apply (subgoal_tac "0 < r / t") - apply (rule hnorm_mult_less) - apply (simp add: InfinitesimalD) - apply assumption - apply simp - apply (erule order_le_less_trans [OF hnorm_ge_zero]) - done +lemma Infinitesimal_HFinite_mult: + fixes x y :: "'a::real_normed_algebra star" + assumes "x \ Infinitesimal" "y \ HFinite" + shows "x * y \ Infinitesimal" +proof (rule InfinitesimalI) + obtain t where "hnorm y < t" "t \ Reals" + using HFiniteD \y \ HFinite\ by blast + then have "t > 0" + using hnorm_ge_zero le_less_trans by blast + show "hnorm (x * y) < r" + if "r \ \" and "0 < r" for r :: "real star" + proof - + have "hnorm x < r/t" + by (meson InfinitesimalD Reals_divide \hnorm y < t\ \t \ \\ assms(1) divide_pos_pos hnorm_ge_zero le_less_trans that) + then have "hnorm (x * y) < (r / t) * t" + using \hnorm y < t\ hnorm_mult_less by blast + then show ?thesis + using \0 < t\ by auto + qed +qed lemma Infinitesimal_HFinite_scaleHR: - "x \ Infinitesimal \ y \ HFinite \ scaleHR x y \ Infinitesimal" - apply (rule InfinitesimalI) - apply (drule HFiniteD, clarify) - apply (drule InfinitesimalD) - apply (simp add: hnorm_scaleHR) - apply (subgoal_tac "0 < t") - apply (subgoal_tac "\x\ * hnorm y < (r / t) * t", simp) - apply (subgoal_tac "0 < r / t") - apply (rule mult_strict_mono', simp_all) - apply (erule order_le_less_trans [OF hnorm_ge_zero]) - done + assumes "x \ Infinitesimal" "y \ HFinite" + shows "scaleHR x y \ Infinitesimal" +proof (rule InfinitesimalI) + obtain t where "hnorm y < t" "t \ Reals" + using HFiniteD \y \ HFinite\ by blast + then have "t > 0" + using hnorm_ge_zero le_less_trans by blast + show "hnorm (scaleHR x y) < r" + if "r \ \" and "0 < r" for r :: "real star" + proof - + have "\x\ * hnorm y < (r / t) * t" + by (metis InfinitesimalD Reals_divide \0 < t\ \hnorm y < t\ \t \ \\ assms(1) divide_pos_pos hnorm_ge_zero hypreal_hnorm_def mult_strict_mono' that) + then show ?thesis + by (simp add: \0 < t\ hnorm_scaleHR less_imp_not_eq2) + qed +qed lemma Infinitesimal_HFinite_mult2: - "x \ Infinitesimal \ y \ HFinite \ y * x \ Infinitesimal" - for x y :: "'a::real_normed_algebra star" - apply (rule InfinitesimalI) - apply (drule HFiniteD, clarify) - apply (subgoal_tac "0 < t") - apply (subgoal_tac "hnorm (y * x) < t * (r / t)", simp) - apply (subgoal_tac "0 < r / t") - apply (rule hnorm_mult_less) - apply assumption - apply (simp add: InfinitesimalD) - apply simp - apply (erule order_le_less_trans [OF hnorm_ge_zero]) - done + fixes x y :: "'a::real_normed_algebra star" + assumes "x \ Infinitesimal" "y \ HFinite" + shows "y * x \ Infinitesimal" +proof (rule InfinitesimalI) + obtain t where "hnorm y < t" "t \ Reals" + using HFiniteD \y \ HFinite\ by blast + then have "t > 0" + using hnorm_ge_zero le_less_trans by blast + show "hnorm (y * x) < r" + if "r \ \" and "0 < r" for r :: "real star" + proof - + have "hnorm x < r/t" + by (meson InfinitesimalD Reals_divide \hnorm y < t\ \t \ \\ assms(1) divide_pos_pos hnorm_ge_zero le_less_trans that) + then have "hnorm (y * x) < t * (r / t)" + using \hnorm y < t\ hnorm_mult_less by blast + then show ?thesis + using \0 < t\ by auto + qed +qed -lemma Infinitesimal_scaleR2: "x \ Infinitesimal \ a *\<^sub>R x \ Infinitesimal" - apply (case_tac "a = 0", simp) - apply (rule InfinitesimalI) - apply (drule InfinitesimalD) - apply (drule_tac x="r / \star_of a\" in bspec) - apply (simp add: Reals_eq_Standard) - apply simp - apply (simp add: hnorm_scaleR pos_less_divide_eq mult.commute) - done +lemma Infinitesimal_scaleR2: + assumes "x \ Infinitesimal" shows "a *\<^sub>R x \ Infinitesimal" + by (metis HFinite_star_of Infinitesimal_HFinite_mult2 Infinitesimal_hnorm_iff assms hnorm_scaleR hypreal_hnorm_def star_of_norm) lemma Compl_HFinite: "- HFinite = HInfinite" - apply (auto simp add: HInfinite_def HFinite_def linorder_not_less) - apply (rule_tac y="r + 1" in order_less_le_trans, simp) - apply simp - done + proof - + have "r < hnorm x" if *: "\s. s \ \ \ s \ hnorm x" and "r \ \" + for x :: "'a star" and r :: hypreal + using * [of "r+1"] \r \ \\ by auto + then show ?thesis + by (auto simp add: HInfinite_def HFinite_def linorder_not_less) +qed -lemma HInfinite_inverse_Infinitesimal: "x \ HInfinite \ inverse x \ Infinitesimal" +lemma HInfinite_inverse_Infinitesimal: + "x \ HInfinite \ inverse x \ Infinitesimal" for x :: "'a::real_normed_div_algebra star" - apply (rule InfinitesimalI) - apply (subgoal_tac "x \ 0") - apply (rule inverse_less_imp_less) - apply (simp add: nonzero_hnorm_inverse) - apply (simp add: HInfinite_def Reals_inverse) - apply assumption - apply (clarify, simp add: Compl_HFinite [symmetric]) - done + by (simp add: HInfinite_def InfinitesimalI hnorm_inverse inverse_less_imp_less) + +lemma inverse_Infinitesimal_iff_HInfinite: + "x \ 0 \ inverse x \ Infinitesimal \ x \ HInfinite" + for x :: "'a::real_normed_div_algebra star" + by (metis Compl_HFinite Compl_iff HInfinite_inverse_Infinitesimal InfinitesimalD Infinitesimal_HFinite_mult Reals_1 hnorm_one left_inverse less_irrefl zero_less_one) lemma HInfiniteI: "(\r. r \ \ \ r < hnorm x) \ x \ HInfinite" by (simp add: HInfinite_def) @@ -381,18 +384,26 @@ lemma HInfiniteD: "x \ HInfinite \ r \ \ \ r < hnorm x" by (simp add: HInfinite_def) -lemma HInfinite_mult: "x \ HInfinite \ y \ HInfinite \ x * y \ HInfinite" - for x y :: "'a::real_normed_div_algebra star" - apply (rule HInfiniteI, simp only: hnorm_mult) - apply (subgoal_tac "r * 1 < hnorm x * hnorm y", simp only: mult_1) - apply (case_tac "x = 0", simp add: HInfinite_def) - apply (rule mult_strict_mono) - apply (simp_all add: HInfiniteD) - done +lemma HInfinite_mult: + fixes x y :: "'a::real_normed_div_algebra star" + assumes "x \ HInfinite" "y \ HInfinite" shows "x * y \ HInfinite" +proof (rule HInfiniteI, simp only: hnorm_mult) + have "x \ 0" + using Compl_HFinite HFinite_0 assms by blast + show "r < hnorm x * hnorm y" + if "r \ \" for r :: "real star" + proof - + have "r = r * 1" + by simp + also have "\ < hnorm x * hnorm y" + by (meson HInfiniteD Reals_1 \x \ 0\ assms le_numeral_extra(1) mult_strict_mono that zero_less_hnorm_iff) + finally show ?thesis . + qed +qed lemma hypreal_add_zero_less_le_mono: "r < x \ 0 \ y \ r < x + y" for r x y :: hypreal - by (auto dest: add_less_le_mono) + by simp lemma HInfinite_add_ge_zero: "x \ HInfinite \ 0 \ y \ 0 \ x \ x + y \ HInfinite" for x y :: hypreal @@ -411,12 +422,7 @@ lemma HInfinite_add_le_zero: "x \ HInfinite \ y \ 0 \ x \ 0 \ x + y \ HInfinite" for x y :: hypreal - apply (drule HInfinite_minus_iff [THEN iffD2]) - apply (rule HInfinite_minus_iff [THEN iffD1]) - apply (simp only: minus_add add.commute) - apply (rule HInfinite_add_ge_zero) - apply simp_all - done + by (metis (no_types, lifting) HInfinite_add_ge_zero2 HInfinite_minus_iff add.inverse_distrib_swap neg_0_le_iff_le) lemma HInfinite_add_lt_zero: "x \ HInfinite \ y < 0 \ x < 0 \ x + y \ HInfinite" for x y :: hypreal @@ -465,17 +471,12 @@ lemma lemma_Infinitesimal_hyperpow: "x \ Infinitesimal \ 0 < N \ \x pow N\ \ \x\" for x :: hypreal - apply (unfold Infinitesimal_def) - apply (auto intro!: hyperpow_Suc_le_self2 - simp: hyperpow_hrabs [symmetric] hypnat_gt_zero_iff2 abs_ge_zero) - done + apply (clarsimp simp: Infinitesimal_def) + by (metis Reals_1 abs_ge_zero hyperpow_Suc_le_self2 hyperpow_hrabs hypnat_gt_zero_iff2 zero_less_one) lemma Infinitesimal_hyperpow: "x \ Infinitesimal \ 0 < N \ x pow N \ Infinitesimal" for x :: hypreal - apply (rule hrabs_le_Infinitesimal) - apply (rule_tac [2] lemma_Infinitesimal_hyperpow) - apply auto - done + using hrabs_le_Infinitesimal lemma_Infinitesimal_hyperpow by blast lemma hrealpow_hyperpow_Infinitesimal_iff: "(x ^ n \ Infinitesimal) \ x pow (hypnat_of_nat n) \ Infinitesimal" @@ -488,21 +489,11 @@ lemma not_Infinitesimal_mult: "x \ Infinitesimal \ y \ Infinitesimal \ x * y \ Infinitesimal" for x y :: "'a::real_normed_div_algebra star" - apply (unfold Infinitesimal_def, clarify, rename_tac r s) - apply (simp only: linorder_not_less hnorm_mult) - apply (drule_tac x = "r * s" in bspec) - apply (fast intro: Reals_mult) - apply simp - apply (drule_tac c = s and d = "hnorm y" and a = r and b = "hnorm x" in mult_mono) - apply simp_all - done + by (metis (no_types, lifting) inverse_Infinitesimal_iff_HInfinite ComplI Compl_HFinite Infinitesimal_HFinite_mult divide_inverse eq_divide_imp inverse_inverse_eq mult_zero_right) lemma Infinitesimal_mult_disj: "x * y \ Infinitesimal \ x \ Infinitesimal \ y \ Infinitesimal" for x y :: "'a::real_normed_div_algebra star" - apply (rule ccontr) - apply (drule de_Morgan_disj [THEN iffD1]) - apply (fast dest: not_Infinitesimal_mult) - done + using not_Infinitesimal_mult by blast lemma HFinite_Infinitesimal_not_zero: "x \ HFinite-Infinitesimal \ x \ 0" by blast @@ -510,16 +501,10 @@ lemma HFinite_Infinitesimal_diff_mult: "x \ HFinite - Infinitesimal \ y \ HFinite - Infinitesimal \ x * y \ HFinite - Infinitesimal" for x y :: "'a::real_normed_div_algebra star" - apply clarify - apply (blast dest: HFinite_mult not_Infinitesimal_mult) - done + by (simp add: HFinite_mult not_Infinitesimal_mult) lemma Infinitesimal_subset_HFinite: "Infinitesimal \ HFinite" - apply (simp add: Infinitesimal_def HFinite_def) - apply auto - apply (rule_tac x = 1 in bexI) - apply auto - done + using HFinite_def InfinitesimalD Reals_1 zero_less_one by blast lemma Infinitesimal_star_of_mult: "x \ Infinitesimal \ x * star_of r \ Infinitesimal" for x :: "'a::real_normed_algebra star" @@ -544,21 +529,19 @@ lemma approx_refl [iff]: "x \ x" by (simp add: approx_def Infinitesimal_def) -lemma hypreal_minus_distrib1: "- (y + - x) = x + -y" - for x y :: "'a::ab_group_add" - by (simp add: add.commute) +lemma approx_sym: "x \ y \ y \ x" + by (metis Infinitesimal_minus_iff approx_def minus_diff_eq) -lemma approx_sym: "x \ y \ y \ x" - apply (simp add: approx_def) - apply (drule Infinitesimal_minus_iff [THEN iffD2]) - apply simp - done - -lemma approx_trans: "x \ y \ y \ z \ x \ z" - apply (simp add: approx_def) - apply (drule (1) Infinitesimal_add) - apply simp - done +lemma approx_trans: + assumes "x \ y" "y \ z" shows "x \ z" +proof - + have "x - y \ Infinitesimal" "z - y \ Infinitesimal" + using assms approx_def approx_sym by auto + then have "x - z \ Infinitesimal" + using Infinitesimal_diff by force + then show ?thesis + by (simp add: approx_def) +qed lemma approx_trans2: "r \ x \ s \ x \ r \ s" by (blast intro: approx_sym approx_trans) @@ -587,12 +570,11 @@ by (simp add: approx_minus_iff [symmetric] mem_infmal_iff) lemma approx_monad_iff: "x \ y \ monad x = monad y" - by (auto simp add: monad_def dest: approx_sym elim!: approx_trans equalityCE) + apply (simp add: monad_def set_eq_iff) + using approx_reorient approx_trans by blast lemma Infinitesimal_approx: "x \ Infinitesimal \ y \ Infinitesimal \ x \ y" - apply (simp add: mem_infmal_iff) - apply (blast intro: approx_trans approx_sym) - done + by (simp add: Infinitesimal_diff approx_def) lemma approx_add: "a \ b \ c \ d \ a + c \ b + d" proof (unfold approx_def) @@ -604,10 +586,7 @@ qed lemma approx_minus: "a \ b \ - a \ - b" - apply (rule approx_minus_iff [THEN iffD2, THEN approx_sym]) - apply (drule approx_minus_iff [THEN iffD1]) - apply (simp add: add.commute) - done + by (metis approx_def approx_sym minus_diff_eq minus_diff_minus) lemma approx_minus2: "- a \ - b \ a \ b" by (auto dest: approx_minus) @@ -654,16 +633,10 @@ by (force simp add: bex_Infinitesimal_iff [symmetric]) lemma Infinitesimal_add_approx: "y \ Infinitesimal \ x + y = z \ x \ z" - apply (rule bex_Infinitesimal_iff [THEN iffD1]) - apply (drule Infinitesimal_minus_iff [THEN iffD2]) - apply (auto simp add: add.assoc [symmetric]) - done + using approx_sym bex_Infinitesimal_iff2 by blast lemma Infinitesimal_add_approx_self: "y \ Infinitesimal \ x \ x + y" - apply (rule bex_Infinitesimal_iff [THEN iffD1]) - apply (drule Infinitesimal_minus_iff [THEN iffD2]) - apply (auto simp add: add.assoc [symmetric]) - done + by (simp add: Infinitesimal_add_approx) lemma Infinitesimal_add_approx_self2: "y \ Infinitesimal \ x \ y + x" by (auto dest: Infinitesimal_add_approx_self simp add: add.commute) @@ -672,31 +645,19 @@ by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2]) lemma Infinitesimal_add_cancel: "y \ Infinitesimal \ x + y \ z \ x \ z" - apply (drule_tac x = x in Infinitesimal_add_approx_self [THEN approx_sym]) - apply (erule approx_trans3 [THEN approx_sym], assumption) - done + using Infinitesimal_add_approx approx_trans by blast lemma Infinitesimal_add_right_cancel: "y \ Infinitesimal \ x \ z + y \ x \ z" - apply (drule_tac x = z in Infinitesimal_add_approx_self2 [THEN approx_sym]) - apply (erule approx_trans3 [THEN approx_sym]) - apply (simp add: add.commute) - apply (erule approx_sym) - done + by (metis Infinitesimal_add_approx_self approx_monad_iff) -lemma approx_add_left_cancel: "d + b \ d + c \ b \ c" - apply (drule approx_minus_iff [THEN iffD1]) - apply (simp add: approx_minus_iff [symmetric] ac_simps) - done +lemma approx_add_left_cancel: "d + b \ d + c \ b \ c" + by (metis add_diff_cancel_left bex_Infinitesimal_iff) lemma approx_add_right_cancel: "b + d \ c + d \ b \ c" - apply (rule approx_add_left_cancel) - apply (simp add: add.commute) - done + by (simp add: approx_def) lemma approx_add_mono1: "b \ c \ d + b \ d + c" - apply (rule approx_minus_iff [THEN iffD2]) - apply (simp add: approx_minus_iff [symmetric] ac_simps) - done + by (simp add: approx_add) lemma approx_add_mono2: "b \ c \ b + a \ c + a" by (simp add: add.commute approx_add_mono1) @@ -708,52 +669,41 @@ by (simp add: add.commute) lemma approx_HFinite: "x \ HFinite \ x \ y \ y \ HFinite" - apply (drule bex_Infinitesimal_iff2 [THEN iffD2], safe) - apply (drule Infinitesimal_subset_HFinite [THEN subsetD, THEN HFinite_minus_iff [THEN iffD2]]) - apply (drule HFinite_add) - apply (auto simp add: add.assoc) - done + by (metis HFinite_add Infinitesimal_subset_HFinite approx_sym subsetD bex_Infinitesimal_iff2) lemma approx_star_of_HFinite: "x \ star_of D \ x \ HFinite" by (rule approx_sym [THEN [2] approx_HFinite], auto) lemma approx_mult_HFinite: "a \ b \ c \ d \ b \ HFinite \ d \ HFinite \ a * c \ b * d" for a b c d :: "'a::real_normed_algebra star" - apply (rule approx_trans) - apply (rule_tac [2] approx_mult2) - apply (rule approx_mult1) - prefer 2 apply (blast intro: approx_HFinite approx_sym, auto) - done + by (meson approx_HFinite approx_mult2 approx_mult_subst2 approx_sym) lemma scaleHR_left_diff_distrib: "\a b x. scaleHR (a - b) x = scaleHR a x - scaleHR b x" by transfer (rule scaleR_left_diff_distrib) lemma approx_scaleR1: "a \ star_of b \ c \ HFinite \ scaleHR a c \ b *\<^sub>R c" - apply (unfold approx_def) - apply (drule (1) Infinitesimal_HFinite_scaleHR) - apply (simp only: scaleHR_left_diff_distrib) - apply (simp add: scaleHR_def star_scaleR_def [symmetric]) - done + unfolding approx_def + by (metis Infinitesimal_HFinite_scaleHR scaleHR_def scaleHR_left_diff_distrib star_scaleR_def starfun2_star_of) lemma approx_scaleR2: "a \ b \ c *\<^sub>R a \ c *\<^sub>R b" by (simp add: approx_def Infinitesimal_scaleR2 scaleR_right_diff_distrib [symmetric]) lemma approx_scaleR_HFinite: "a \ star_of b \ c \ d \ d \ HFinite \ scaleHR a c \ b *\<^sub>R d" - apply (rule approx_trans) - apply (rule_tac [2] approx_scaleR2) - apply (rule approx_scaleR1) - prefer 2 apply (blast intro: approx_HFinite approx_sym, auto) - done + by (meson approx_HFinite approx_scaleR1 approx_scaleR2 approx_sym approx_trans) lemma approx_mult_star_of: "a \ star_of b \ c \ star_of d \ a * c \ star_of b * star_of d" for a c :: "'a::real_normed_algebra star" by (blast intro!: approx_mult_HFinite approx_star_of_HFinite HFinite_star_of) -lemma approx_SReal_mult_cancel_zero: "a \ \ \ a \ 0 \ a * x \ 0 \ x \ 0" - for a x :: hypreal - apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) - apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) - done +lemma approx_SReal_mult_cancel_zero: + fixes a x :: hypreal + assumes "a \ \" "a \ 0" "a * x \ 0" shows "x \ 0" +proof - + have "inverse a \ HFinite" + using Reals_inverse SReal_subset_HFinite assms(1) by blast + then show ?thesis + using assms by (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) +qed lemma approx_mult_SReal1: "a \ \ \ x \ 0 \ x * a \ 0" for a x :: hypreal @@ -767,25 +717,31 @@ for a x :: hypreal by (blast intro: approx_SReal_mult_cancel_zero approx_mult_SReal2) -lemma approx_SReal_mult_cancel: "a \ \ \ a \ 0 \ a * w \ a * z \ w \ z" - for a w z :: hypreal - apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) - apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) - done +lemma approx_SReal_mult_cancel: + fixes a w z :: hypreal + assumes "a \ \" "a \ 0" "a * w \ a * z" shows "w \ z" +proof - + have "inverse a \ HFinite" + using Reals_inverse SReal_subset_HFinite assms(1) by blast + then show ?thesis + using assms by (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) +qed lemma approx_SReal_mult_cancel_iff1 [simp]: "a \ \ \ a \ 0 \ a * w \ a * z \ w \ z" for a w z :: hypreal - by (auto intro!: approx_mult2 SReal_subset_HFinite [THEN subsetD] - intro: approx_SReal_mult_cancel) + by (meson SReal_subset_HFinite approx_SReal_mult_cancel approx_mult2 subsetD) -lemma approx_le_bound: "z \ f \ f \ g \ g \ z ==> f \ z" - for z :: hypreal - apply (simp add: bex_Infinitesimal_iff2 [symmetric], auto) - apply (rule_tac x = "g + y - z" in bexI) - apply simp - apply (rule Infinitesimal_interval2) - apply (rule_tac [2] Infinitesimal_zero, auto) - done +lemma approx_le_bound: + fixes z :: hypreal + assumes "z \ f" " f \ g" "g \ z" shows "f \ z" +proof - + obtain y where "z \ g + y" and "y \ Infinitesimal" "f = g + y" + using assms bex_Infinitesimal_iff2 by auto + then have "z - g \ Infinitesimal" + using assms(3) hrabs_le_Infinitesimal by auto + then show ?thesis + by (metis approx_def approx_trans2 assms(2)) +qed lemma approx_hnorm: "x \ y \ hnorm x \ hnorm y" for x y :: "'a::real_normed_vector star" @@ -810,9 +766,7 @@ lemma Infinitesimal_less_SReal: "x \ \ \ y \ Infinitesimal \ 0 < x \ y < x" for x y :: hypreal - apply (simp add: Infinitesimal_def) - apply (rule abs_ge_self [THEN order_le_less_trans], auto) - done + using InfinitesimalD by fastforce lemma Infinitesimal_less_SReal2: "y \ Infinitesimal \ \r \ Reals. 0 < r \ y < r" for y :: hypreal @@ -820,21 +774,19 @@ lemma SReal_not_Infinitesimal: "0 < y \ y \ \ ==> y \ Infinitesimal" for y :: hypreal - apply (simp add: Infinitesimal_def) - apply (auto simp add: abs_if) - done + by (auto simp add: Infinitesimal_def abs_if) lemma SReal_minus_not_Infinitesimal: "y < 0 \ y \ \ \ y \ Infinitesimal" for y :: hypreal - apply (subst Infinitesimal_minus_iff [symmetric]) - apply (rule SReal_not_Infinitesimal, auto) - done + using Infinitesimal_minus_iff Reals_minus SReal_not_Infinitesimal neg_0_less_iff_less by blast lemma SReal_Int_Infinitesimal_zero: "\ Int Infinitesimal = {0::hypreal}" - apply auto - apply (cut_tac x = x and y = 0 in linorder_less_linear) - apply (blast dest: SReal_not_Infinitesimal SReal_minus_not_Infinitesimal) - done + proof - + have "x = 0" if "x \ \" "x \ Infinitesimal" for x :: "real star" + using that SReal_minus_not_Infinitesimal SReal_not_Infinitesimal not_less_iff_gr_or_eq by blast + then show ?thesis + by auto +qed lemma SReal_Infinitesimal_zero: "x \ \ \ x \ Infinitesimal \ x = 0" for x :: hypreal @@ -849,12 +801,15 @@ by (rule SReal_HFinite_diff_Infinitesimal) auto lemma star_of_Infinitesimal_iff_0 [iff]: "star_of x \ Infinitesimal \ x = 0" - apply (auto simp add: Infinitesimal_def) - apply (drule_tac x="hnorm (star_of x)" in bspec) - apply (simp add: SReal_def) - apply (rule_tac x="norm x" in exI, simp) - apply simp - done +proof + show "x = 0" if "star_of x \ Infinitesimal" + proof - + have "hnorm (star_n (\n. x)) \ Standard" + by (metis Reals_eq_Standard SReal_iff star_of_def star_of_norm) + then show ?thesis + by (metis InfinitesimalD2 less_irrefl star_of_norm that zero_less_norm_iff) + qed +qed auto lemma star_of_HFinite_diff_Infinitesimal: "x \ 0 \ star_of x \ HFinite - Infinitesimal" by simp @@ -866,39 +821,26 @@ text \Again: \1\ is a special case, but not \0\ this time.\ lemma one_not_Infinitesimal [simp]: "(1::'a::{real_normed_vector,zero_neq_one} star) \ Infinitesimal" - apply (simp only: star_one_def star_of_Infinitesimal_iff_0) - apply simp - done + by (metis star_of_Infinitesimal_iff_0 star_one_def zero_neq_one) lemma approx_SReal_not_zero: "y \ \ \ x \ y \ y \ 0 \ x \ 0" for x y :: hypreal - apply (cut_tac x = 0 and y = y in linorder_less_linear, simp) - apply (blast dest: approx_sym [THEN mem_infmal_iff [THEN iffD2]] - SReal_not_Infinitesimal SReal_minus_not_Infinitesimal) - done + using SReal_Infinitesimal_zero approx_sym mem_infmal_iff by auto lemma HFinite_diff_Infinitesimal_approx: "x \ y \ y \ HFinite - Infinitesimal \ x \ HFinite - Infinitesimal" - apply (auto intro: approx_sym [THEN [2] approx_HFinite] simp: mem_infmal_iff) - apply (drule approx_trans3, assumption) - apply (blast dest: approx_sym) - done + by (meson Diff_iff approx_HFinite approx_sym approx_trans3 mem_infmal_iff) text \The premise \y \ 0\ is essential; otherwise \x / y = 0\ and we lose the \HFinite\ premise.\ lemma Infinitesimal_ratio: "y \ 0 \ y \ Infinitesimal \ x / y \ HFinite \ x \ Infinitesimal" for x y :: "'a::{real_normed_div_algebra,field} star" - apply (drule Infinitesimal_HFinite_mult2, assumption) - apply (simp add: divide_inverse mult.assoc) - done + using Infinitesimal_HFinite_mult by fastforce lemma Infinitesimal_SReal_divide: "x \ Infinitesimal \ y \ \ \ x / y \ Infinitesimal" for x y :: hypreal - apply (simp add: divide_inverse) - apply (auto intro!: Infinitesimal_HFinite_mult - dest!: Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) - done + by (metis HFinite_star_of Infinitesimal_HFinite_mult Reals_inverse SReal_iff divide_inverse) section \Standard Part Theorem\ @@ -912,28 +854,15 @@ subsection \Uniqueness: Two Infinitely Close Reals are Equal\ lemma star_of_approx_iff [simp]: "star_of x \ star_of y \ x = y" - apply safe - apply (simp add: approx_def) - apply (simp only: star_of_diff [symmetric]) - apply (simp only: star_of_Infinitesimal_iff_0) - apply simp - done + by (metis approx_def right_minus_eq star_of_Infinitesimal_iff_0 star_of_simps(2)) lemma SReal_approx_iff: "x \ \ \ y \ \ \ x \ y \ x = y" for x y :: hypreal - apply auto - apply (simp add: approx_def) - apply (drule (1) Reals_diff) - apply (drule (1) SReal_Infinitesimal_zero) - apply simp - done + by (meson Reals_diff SReal_Infinitesimal_zero approx_def approx_refl right_minus_eq) lemma numeral_approx_iff [simp]: - "(numeral v \ (numeral w :: 'a::{numeral,real_normed_vector} star)) = - (numeral v = (numeral w :: 'a))" - apply (unfold star_numeral_def) - apply (rule star_of_approx_iff) - done + "(numeral v \ (numeral w :: 'a::{numeral,real_normed_vector} star)) = (numeral v = (numeral w :: 'a))" + by (metis star_of_approx_iff star_of_numeral) text \And also for \0 \ #nn\ and \1 \ #nn\, \#nn \ 0\ and \#nn \ 1\.\ lemma [simp]: @@ -943,10 +872,8 @@ "((1::'b::{numeral,one,real_normed_vector} star) \ numeral w) = (numeral w = (1::'b))" "\ (0 \ (1::'c::{zero_neq_one,real_normed_vector} star))" "\ (1 \ (0::'c::{zero_neq_one,real_normed_vector} star))" - apply (unfold star_numeral_def star_zero_def star_one_def) - apply (unfold star_of_approx_iff) - apply (auto intro: sym) - done + unfolding star_numeral_def star_zero_def star_one_def star_of_approx_iff + by (auto intro: sym) lemma star_of_approx_numeral_iff [simp]: "star_of k \ numeral w \ k = numeral w" by (subst star_of_approx_iff [symmetric]) auto @@ -970,23 +897,19 @@ for Q :: "real set" and Y :: real by (simp add: isUb_def setle_def) -lemma hypreal_of_real_isLub1: "isLub \ (hypreal_of_real ` Q) (hypreal_of_real Y) \ isLub UNIV Q Y" - for Q :: "real set" and Y :: real - apply (simp add: isLub_def leastP_def) - apply (auto intro: hypreal_of_real_isUb_iff [THEN iffD2] - simp add: hypreal_of_real_isUb_iff setge_def) - done - -lemma hypreal_of_real_isLub2: "isLub UNIV Q Y \ isLub \ (hypreal_of_real ` Q) (hypreal_of_real Y)" +lemma hypreal_of_real_isLub_iff: + "isLub \ (hypreal_of_real ` Q) (hypreal_of_real Y) = isLub (UNIV :: real set) Q Y" (is "?lhs = ?rhs") for Q :: "real set" and Y :: real - apply (auto simp add: isLub_def leastP_def hypreal_of_real_isUb_iff setge_def) - apply (metis SReal_iff hypreal_of_real_isUb_iff isUbD2a star_of_le) - done - -lemma hypreal_of_real_isLub_iff: - "isLub \ (hypreal_of_real ` Q) (hypreal_of_real Y) = isLub (UNIV :: real set) Q Y" - for Q :: "real set" and Y :: real - by (blast intro: hypreal_of_real_isLub1 hypreal_of_real_isLub2) +proof + assume ?lhs + then show ?rhs + by (simp add: isLub_def leastP_def) (metis hypreal_of_real_isUb_iff mem_Collect_eq setge_def star_of_le) +next + assume ?rhs + then show ?lhs + apply (simp add: isLub_def leastP_def hypreal_of_real_isUb_iff setge_def) + by (metis SReal_iff hypreal_of_real_isUb_iff isUb_def star_of_le) +qed lemma lemma_isUb_hypreal_of_real: "isUb \ P Y \ \Yo. isUb \ P (hypreal_of_real Yo)" by (auto simp add: SReal_iff isUb_def) @@ -994,58 +917,47 @@ lemma lemma_isLub_hypreal_of_real: "isLub \ P Y \ \Yo. isLub \ P (hypreal_of_real Yo)" by (auto simp add: isLub_def leastP_def isUb_def SReal_iff) -lemma lemma_isLub_hypreal_of_real2: "\Yo. isLub \ P (hypreal_of_real Yo) \ \Y. isLub \ P Y" - by (auto simp add: isLub_def leastP_def isUb_def) - -lemma SReal_complete: "P \ \ \ \x. x \ P \ \Y. isUb \ P Y \ \t::hypreal. isLub \ P t" - apply (frule SReal_hypreal_of_real_image) - apply (auto, drule lemma_isUb_hypreal_of_real) - apply (auto intro!: reals_complete lemma_isLub_hypreal_of_real2 - simp add: hypreal_of_real_isLub_iff hypreal_of_real_isUb_iff) - done +lemma SReal_complete: + fixes P :: "hypreal set" + assumes "isUb \ P Y" "P \ \" "P \ {}" + shows "\t. isLub \ P t" +proof - + obtain Q where "P = hypreal_of_real ` Q" + by (metis \P \ \\ hypreal_of_real_image subset_imageE) + then show ?thesis + by (metis assms(1) \P \ {}\ equals0I hypreal_of_real_isLub_iff hypreal_of_real_isUb_iff image_empty lemma_isUb_hypreal_of_real reals_complete) +qed text \Lemmas about lubs.\ -lemma lemma_st_part_ub: "x \ HFinite \ \u. isUb \ {s. s \ \ \ s < x} u" - for x :: hypreal - apply (drule HFiniteD, safe) - apply (rule exI, rule isUbI) - apply (auto intro: setleI isUbI simp add: abs_less_iff) - done - -lemma lemma_st_part_nonempty: "x \ HFinite \ \y. y \ {s. s \ \ \ s < x}" - for x :: hypreal - apply (drule HFiniteD, safe) - apply (drule Reals_minus) - apply (rule_tac x = "-t" in exI) - apply (auto simp add: abs_less_iff) - done - -lemma lemma_st_part_lub: "x \ HFinite \ \t. isLub \ {s. s \ \ \ s < x} t" - for x :: hypreal - by (blast intro!: SReal_complete lemma_st_part_ub lemma_st_part_nonempty Collect_restrict) +lemma lemma_st_part_lub: + fixes x :: hypreal + assumes "x \ HFinite" + shows "\t. isLub \ {s. s \ \ \ s < x} t" +proof - + obtain t where t: "t \ \" "hnorm x < t" + using HFiniteD assms by blast + then have "isUb \ {s. s \ \ \ s < x} t" + by (simp add: abs_less_iff isUbI le_less_linear less_imp_not_less setleI) + moreover have "\y. y \ \ \ y < x" + using t by (rule_tac x = "-t" in exI) (auto simp add: abs_less_iff) + ultimately show ?thesis + using SReal_complete by fastforce +qed lemma lemma_st_part_le1: "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ r \ \ \ 0 < r \ x \ t + r" for x r t :: hypreal - apply (frule isLubD1a) - apply (rule ccontr, drule linorder_not_le [THEN iffD2]) - apply (drule (1) Reals_add) - apply (drule_tac y = "r + t" in isLubD1 [THEN setleD], auto) - done + by (metis (no_types, lifting) Reals_add add.commute isLubD1a isLubD2 less_add_same_cancel2 mem_Collect_eq not_le) lemma hypreal_setle_less_trans: "S *<= x \ x < y \ S *<= y" for x y :: hypreal - apply (simp add: setle_def) - apply (auto dest!: bspec order_le_less_trans intro: order_less_imp_le) - done + by (meson le_less_trans less_imp_le setle_def) lemma hypreal_gt_isUb: "isUb R S x \ x < y \ y \ R \ isUb R S y" for x y :: hypreal - apply (simp add: isUb_def) - apply (blast intro: hypreal_setle_less_trans) - done + using hypreal_setle_less_trans isUb_def by blast lemma lemma_st_part_gt_ub: "x \ HFinite \ x < y \ y \ \ \ isUb \ {s. s \ \ \ s < x} y" for x y :: hypreal @@ -1053,9 +965,7 @@ lemma lemma_minus_le_zero: "t \ t + -r \ r \ 0" for r t :: hypreal - apply (drule_tac c = "-t" in add_left_mono) - apply (auto simp add: add.assoc [symmetric]) - done + by simp lemma lemma_st_part_le2: "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ r \ \ \ 0 < r \ t + -r \ x" @@ -1072,9 +982,7 @@ lemma lemma_st_part1a: "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ r \ \ \ 0 < r \ x + -t \ r" for x r t :: hypreal - apply (subgoal_tac "x \ t + r") - apply (auto intro: lemma_st_part_le1) - done + using lemma_st_part_le1 by fastforce lemma lemma_st_part2a: "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ r \ \ \ 0 < r \ - (x + -t) \ r" @@ -1141,46 +1049,28 @@ lemma lemma_st_part_Ex: "x \ HFinite \ \t\Reals. \r \ Reals. 0 < r \ \x - t\ < r" for x :: hypreal - apply (frule lemma_st_part_lub, safe) - apply (frule isLubD1a) - apply (blast dest: lemma_st_part_major2) - done + by (meson isLubD1a lemma_st_part_lub lemma_st_part_major2) lemma st_part_Ex: "x \ HFinite \ \t\Reals. x \ t" for x :: hypreal - apply (simp add: approx_def Infinitesimal_def) - apply (drule lemma_st_part_Ex, auto) - done + by (metis InfinitesimalI approx_def hypreal_hnorm_def lemma_st_part_Ex) text \There is a unique real infinitely close.\ lemma st_part_Ex1: "x \ HFinite \ \!t::hypreal. t \ \ \ x \ t" - apply (drule st_part_Ex, safe) - apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym) - apply (auto intro!: approx_unique_real) - done + by (meson SReal_approx_iff approx_trans2 st_part_Ex) subsection \Finite, Infinite and Infinitesimal\ lemma HFinite_Int_HInfinite_empty [simp]: "HFinite Int HInfinite = {}" - apply (simp add: HFinite_def HInfinite_def) - apply (auto dest: order_less_trans) - done + using Compl_HFinite by blast lemma HFinite_not_HInfinite: - assumes x: "x \ HFinite" - shows "x \ HInfinite" -proof - assume x': "x \ HInfinite" - with x have "x \ HFinite \ HInfinite" by blast - then show False by auto -qed + assumes x: "x \ HFinite" shows "x \ HInfinite" + using Compl_HFinite x by blast lemma not_HFinite_HInfinite: "x \ HFinite \ x \ HInfinite" - apply (simp add: HInfinite_def HFinite_def, auto) - apply (drule_tac x = "r + 1" in bspec) - apply (auto) - done + using Compl_HFinite by blast lemma HInfinite_HFinite_disj: "x \ HInfinite \ x \ HFinite" by (blast intro: not_HFinite_HInfinite) @@ -1191,17 +1081,13 @@ lemma HFinite_HInfinite_iff: "x \ HFinite \ x \ HInfinite" by (simp add: HInfinite_HFinite_iff) - lemma HInfinite_diff_HFinite_Infinitesimal_disj: "x \ Infinitesimal \ x \ HInfinite \ x \ HFinite - Infinitesimal" by (fast intro: not_HFinite_HInfinite) lemma HFinite_inverse: "x \ HFinite \ x \ Infinitesimal \ inverse x \ HFinite" for x :: "'a::real_normed_div_algebra star" - apply (subgoal_tac "x \ 0") - apply (cut_tac x = "inverse x" in HInfinite_HFinite_disj) - apply (auto dest!: HInfinite_inverse_Infinitesimal simp: nonzero_inverse_inverse_eq) - done + using HInfinite_inverse_Infinitesimal not_HFinite_HInfinite by force lemma HFinite_inverse2: "x \ HFinite - Infinitesimal \ inverse x \ HFinite" for x :: "'a::real_normed_div_algebra star" @@ -1210,17 +1096,12 @@ text \Stronger statement possible in fact.\ lemma Infinitesimal_inverse_HFinite: "x \ Infinitesimal \ inverse x \ HFinite" for x :: "'a::real_normed_div_algebra star" - apply (drule HInfinite_diff_HFinite_Infinitesimal_disj) - apply (blast intro: HFinite_inverse HInfinite_inverse_Infinitesimal Infinitesimal_subset_HFinite [THEN subsetD]) - done + using HFinite_HInfinite_iff HInfinite_inverse_Infinitesimal by fastforce lemma HFinite_not_Infinitesimal_inverse: "x \ HFinite - Infinitesimal \ inverse x \ HFinite - Infinitesimal" for x :: "'a::real_normed_div_algebra star" - apply (auto intro: Infinitesimal_inverse_HFinite) - apply (drule Infinitesimal_HFinite_mult2, assumption) - apply (simp add: not_Infinitesimal_not_zero) - done + using HFinite_Infinitesimal_not_zero HFinite_inverse2 Infinitesimal_HFinite_mult2 by fastforce lemma approx_inverse: "x \ y \ y \ HFinite - Infinitesimal \ inverse x \ inverse y" for x y :: "'a::real_normed_div_algebra star" @@ -1245,32 +1126,21 @@ lemma inverse_add_Infinitesimal_approx2: "x \ HFinite - Infinitesimal \ h \ Infinitesimal \ inverse (h + x) \ inverse x" for x h :: "'a::real_normed_div_algebra star" - apply (rule add.commute [THEN subst]) - apply (blast intro: inverse_add_Infinitesimal_approx) - done + by (metis add.commute inverse_add_Infinitesimal_approx) lemma inverse_add_Infinitesimal_approx_Infinitesimal: "x \ HFinite - Infinitesimal \ h \ Infinitesimal \ inverse (x + h) - inverse x \ h" for x h :: "'a::real_normed_div_algebra star" - apply (rule approx_trans2) - apply (auto intro: inverse_add_Infinitesimal_approx - simp add: mem_infmal_iff approx_minus_iff [symmetric]) - done + by (meson Infinitesimal_approx bex_Infinitesimal_iff inverse_add_Infinitesimal_approx) lemma Infinitesimal_square_iff: "x \ Infinitesimal \ x * x \ Infinitesimal" for x :: "'a::real_normed_div_algebra star" - apply (auto intro: Infinitesimal_mult) - apply (rule ccontr, frule Infinitesimal_inverse_HFinite) - apply (frule not_Infinitesimal_not_zero) - apply (auto dest: Infinitesimal_HFinite_mult simp add: mult.assoc) - done + using Infinitesimal_mult Infinitesimal_mult_disj by auto declare Infinitesimal_square_iff [symmetric, simp] lemma HFinite_square_iff [simp]: "x * x \ HFinite \ x \ HFinite" for x :: "'a::real_normed_div_algebra star" - apply (auto intro: HFinite_mult) - apply (auto dest: HInfinite_mult simp add: HFinite_HInfinite_iff) - done + using HFinite_HInfinite_iff HFinite_mult HInfinite_mult by blast lemma HInfinite_square_iff [simp]: "x * x \ HInfinite \ x \ HInfinite" for x :: "'a::real_normed_div_algebra star" @@ -1278,26 +1148,17 @@ lemma approx_HFinite_mult_cancel: "a \ HFinite - Infinitesimal \ a * w \ a * z \ w \ z" for a w z :: "'a::real_normed_div_algebra star" - apply safe - apply (frule HFinite_inverse, assumption) - apply (drule not_Infinitesimal_not_zero) - apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) - done + by (metis DiffD2 Infinitesimal_mult_disj bex_Infinitesimal_iff right_diff_distrib) lemma approx_HFinite_mult_cancel_iff1: "a \ HFinite - Infinitesimal \ a * w \ a * z \ w \ z" for a w z :: "'a::real_normed_div_algebra star" by (auto intro: approx_mult2 approx_HFinite_mult_cancel) lemma HInfinite_HFinite_add_cancel: "x + y \ HInfinite \ y \ HFinite \ x \ HInfinite" - apply (rule ccontr) - apply (drule HFinite_HInfinite_iff [THEN iffD2]) - apply (auto dest: HFinite_add simp add: HInfinite_HFinite_iff) - done + using HFinite_add HInfinite_HFinite_iff by blast lemma HInfinite_HFinite_add: "x \ HInfinite \ y \ HFinite \ x + y \ HInfinite" - apply (rule_tac y = "-y" in HInfinite_HFinite_add_cancel) - apply (auto simp add: add.assoc HFinite_minus_iff) - done + by (metis (no_types, hide_lams) HFinite_HInfinite_iff HFinite_add HFinite_minus_iff add.commute add_minus_cancel) lemma HInfinite_ge_HInfinite: "x \ HInfinite \ x \ y \ 0 \ x \ y \ HInfinite" for x y :: hypreal @@ -1305,30 +1166,17 @@ lemma Infinitesimal_inverse_HInfinite: "x \ Infinitesimal \ x \ 0 \ inverse x \ HInfinite" for x :: "'a::real_normed_div_algebra star" - apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2]) - apply (auto dest: Infinitesimal_HFinite_mult2) - done + by (metis Infinitesimal_HFinite_mult not_HFinite_HInfinite one_not_Infinitesimal right_inverse) lemma HInfinite_HFinite_not_Infinitesimal_mult: "x \ HInfinite \ y \ HFinite - Infinitesimal \ x * y \ HInfinite" for x y :: "'a::real_normed_div_algebra star" - apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2]) - apply (frule HFinite_Infinitesimal_not_zero) - apply (drule HFinite_not_Infinitesimal_inverse) - apply (safe, drule HFinite_mult) - apply (auto simp add: mult.assoc HFinite_HInfinite_iff) - done + by (metis (no_types, hide_lams) HFinite_HInfinite_iff HFinite_Infinitesimal_not_zero HFinite_inverse2 HFinite_mult mult.assoc mult.right_neutral right_inverse) lemma HInfinite_HFinite_not_Infinitesimal_mult2: "x \ HInfinite \ y \ HFinite - Infinitesimal \ y * x \ HInfinite" for x y :: "'a::real_normed_div_algebra star" - apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2]) - apply (frule HFinite_Infinitesimal_not_zero) - apply (drule HFinite_not_Infinitesimal_inverse) - apply (safe, drule_tac x="inverse y" in HFinite_mult) - apply assumption - apply (auto simp add: mult.assoc [symmetric] HFinite_HInfinite_iff) - done + by (metis Diff_iff HInfinite_HFinite_iff HInfinite_inverse_Infinitesimal Infinitesimal_HFinite_mult2 divide_inverse mult_zero_right nonzero_eq_divide_eq) lemma HInfinite_gt_SReal: "x \ HInfinite \ 0 < x \ y \ \ \ y < x" for x y :: hypreal @@ -1338,7 +1186,6 @@ for x :: hypreal by (auto intro: HInfinite_gt_SReal) - lemma not_HInfinite_one [simp]: "1 \ HInfinite" by (simp add: HInfinite_HFinite_iff) @@ -1379,9 +1226,7 @@ by (simp (no_asm)) (simp add: approx_monad_iff) lemma approx_subset_monad2: "x \ y \ {x, y} \ monad y" - apply (drule approx_sym) - apply (fast dest: approx_subset_monad) - done + using approx_subset_monad approx_sym by auto lemma mem_monad_approx: "u \ monad x \ x \ u" by (simp add: monad_def) @@ -1390,28 +1235,18 @@ by (simp add: monad_def) lemma approx_mem_monad2: "x \ u \ x \ monad u" - apply (simp add: monad_def) - apply (blast intro!: approx_sym) - done + using approx_mem_monad approx_sym by blast lemma approx_mem_monad_zero: "x \ y \ x \ monad 0 \ y \ monad 0" - apply (drule mem_monad_approx) - apply (fast intro: approx_mem_monad approx_trans) - done + using approx_trans monad_def by blast lemma Infinitesimal_approx_hrabs: "x \ y \ x \ Infinitesimal \ \x\ \ \y\" for x y :: hypreal - apply (drule Infinitesimal_monad_zero_iff [THEN iffD1]) - apply (blast intro: approx_mem_monad_zero monad_zero_hrabs_iff [THEN iffD1] - mem_monad_approx approx_trans3) - done + using approx_hnorm by fastforce lemma less_Infinitesimal_less: "0 < x \ x \ Infinitesimal \ e \ Infinitesimal \ e < x" for x :: hypreal - apply (rule ccontr) - apply (auto intro: Infinitesimal_zero [THEN [2] Infinitesimal_interval] - dest!: order_le_imp_less_or_eq simp add: linorder_not_less) - done + using Infinitesimal_interval less_linear by blast lemma Ball_mem_monad_gt_zero: "0 < x \ x \ Infinitesimal \ u \ monad x \ 0 < u" for u x :: hypreal @@ -1454,18 +1289,12 @@ lemma hrabs_add_Infinitesimal_cancel: "e \ Infinitesimal \ e' \ Infinitesimal \ \x + e\ = \y + e'\ \ \x\ \ \y\" for e e' x y :: hypreal - apply (drule_tac x = x in approx_hrabs_add_Infinitesimal) - apply (drule_tac x = y in approx_hrabs_add_Infinitesimal) - apply (auto intro: approx_trans2) - done + by (metis approx_hrabs_add_Infinitesimal approx_trans2) lemma hrabs_add_minus_Infinitesimal_cancel: "e \ Infinitesimal \ e' \ Infinitesimal \ \x + -e\ = \y + -e'\ \ \x\ \ \y\" for e e' x y :: hypreal - apply (drule_tac x = x in approx_hrabs_add_minus_Infinitesimal) - apply (drule_tac x = y in approx_hrabs_add_minus_Infinitesimal) - apply (auto intro: approx_trans2) - done + by (meson Infinitesimal_minus_iff hrabs_add_Infinitesimal_cancel) subsection \More \<^term>\HFinite\ and \<^term>\Infinitesimal\ Theorems\ @@ -1493,9 +1322,7 @@ lemma Infinitesimal_add_hrabs_hypreal_of_real_less2: "x \ Infinitesimal \ \hypreal_of_real r\ < hypreal_of_real y \ \x + hypreal_of_real r\ < hypreal_of_real y" - apply (rule add.commute [THEN subst]) - apply (erule Infinitesimal_add_hrabs_hypreal_of_real_less, assumption) - done + using Infinitesimal_add_hrabs_hypreal_of_real_less by fastforce lemma hypreal_of_real_le_add_Infininitesimal_cancel: "u \ Infinitesimal \ v \ Infinitesimal \ @@ -1513,10 +1340,7 @@ lemma hypreal_of_real_less_Infinitesimal_le_zero: "hypreal_of_real x < e \ e \ Infinitesimal \ hypreal_of_real x \ 0" - apply (rule linorder_not_less [THEN iffD1], safe) - apply (drule Infinitesimal_interval) - apply (drule_tac [4] SReal_hypreal_of_real [THEN SReal_Infinitesimal_zero], auto) - done + by (metis Infinitesimal_interval eq_iff le_less_linear star_of_Infinitesimal_iff_0 star_of_eq_0) (*used once, in Lim/NSDERIV_inverse*) lemma Infinitesimal_add_not_zero: "h \ Infinitesimal \ x \ 0 \ star_of x + h \ 0" @@ -1527,16 +1351,11 @@ lemma Infinitesimal_square_cancel [simp]: "x * x + y * y \ Infinitesimal \ x * x \ Infinitesimal" for x y :: hypreal - apply (rule Infinitesimal_interval2) - apply (rule_tac [3] zero_le_square, assumption) - apply auto - done + by (meson Infinitesimal_interval2 le_add_same_cancel1 not_Infinitesimal_not_zero zero_le_square) lemma HFinite_square_cancel [simp]: "x * x + y * y \ HFinite \ x * x \ HFinite" for x y :: hypreal - apply (rule HFinite_bounded, assumption) - apply auto - done + using HFinite_bounded le_add_same_cancel1 zero_le_square by blast lemma Infinitesimal_square_cancel2 [simp]: "x * x + y * y \ Infinitesimal \ y * y \ Infinitesimal" for x y :: hypreal @@ -1596,27 +1415,16 @@ done lemma monad_hrabs_less: "y \ monad x \ 0 < hypreal_of_real e \ \y - x\ < hypreal_of_real e" - apply (drule mem_monad_approx [THEN approx_sym]) - apply (drule bex_Infinitesimal_iff [THEN iffD2]) - apply (auto dest!: InfinitesimalD) - done + by (simp add: Infinitesimal_approx_minus approx_sym less_Infinitesimal_less mem_monad_approx) lemma mem_monad_SReal_HFinite: "x \ monad (hypreal_of_real a) \ x \ HFinite" - apply (drule mem_monad_approx [THEN approx_sym]) - apply (drule bex_Infinitesimal_iff2 [THEN iffD2]) - apply (safe dest!: Infinitesimal_subset_HFinite [THEN subsetD]) - apply (erule SReal_hypreal_of_real [THEN SReal_subset_HFinite [THEN subsetD], THEN HFinite_add]) - done + using HFinite_star_of approx_HFinite mem_monad_approx by blast subsection \Theorems about Standard Part\ lemma st_approx_self: "x \ HFinite \ st x \ x" - apply (simp add: st_def) - apply (frule st_part_Ex, safe) - apply (rule someI2) - apply (auto intro: approx_sym) - done + by (metis (no_types, lifting) approx_refl approx_trans3 someI_ex st_def st_part_Ex st_part_Ex1) lemma st_SReal: "x \ HFinite \ st x \ \" apply (simp add: st_def) @@ -1813,9 +1621,6 @@ for u :: real by auto -lemma lemma_Int_HIa: "{n. u < norm (X n)} \ {n. norm (X n) < u} = {}" - by (auto intro: order_less_asym) - lemma FreeUltrafilterNat_HInfinite: "\u. eventually (\n. u < norm (X n)) \ \ star_n X \ HInfinite" apply (rule HInfinite_HFinite_iff [THEN iffD2]) @@ -1862,9 +1667,7 @@ text \Infinitesimals as smaller than \1/n\ for all \n::nat (> 0)\.\ lemma lemma_Infinitesimal: "(\r. 0 < r \ x < r) \ (\n. x < inverse (real (Suc n)))" - apply (auto simp del: of_nat_Suc) - apply (blast dest!: reals_Archimedean intro: order_less_trans) - done + by (meson inverse_positive_iff_positive less_trans of_nat_0_less_iff reals_Archimedean zero_less_Suc) lemma lemma_Infinitesimal2: "(\r \ Reals. 0 < r \ x < r) \ (\n. x < inverse(hypreal_of_nat (Suc n)))" @@ -1881,9 +1684,7 @@ lemma Infinitesimal_hypreal_of_nat_iff: "Infinitesimal = {x. \n. hnorm x < inverse (hypreal_of_nat (Suc n))}" - apply (simp add: Infinitesimal_def) - apply (auto simp add: lemma_Infinitesimal2) - done + using Infinitesimal_def lemma_Infinitesimal2 by auto subsection \Proof that \\\ is an infinite number\ @@ -1905,11 +1706,8 @@ apply (auto dest: order_less_trans) done -lemma lemma_real_le_Un_eq: "{n. f n \ u} = {n. f n < u} \ {n. u = (f n :: real)}" - by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le) - lemma finite_real_of_nat_le_real: "finite {n::nat. real n \ u}" - by (auto simp add: lemma_real_le_Un_eq lemma_finite_omega_set finite_real_of_nat_less_real) + by (metis infinite_nat_iff_unbounded leD le_nat_floor mem_Collect_eq) lemma finite_rabs_real_of_nat_le_real: "finite {n::nat. \real n\ \ u}" by (simp add: finite_real_of_nat_le_real) @@ -1927,9 +1725,6 @@ text \The complement of \{n. \real n\ \ u} = {n. u < \real n\}\ is in \\\ by property of (free) ultrafilters.\ -lemma Compl_real_le_eq: "- {n::nat. real n \ u} = {n. u < real n}" - by (auto dest!: order_le_less_trans simp add: linorder_not_le) - text \\<^term>\\\ is a member of \<^term>\HInfinite\.\ theorem HInfinite_omega [simp]: "\ \ HInfinite" apply (simp add: omega_def) @@ -1955,12 +1750,7 @@ text \Needed for proof that we define a hyperreal \[ hypreal_of_real a\ given that \\n. |X n - a| < 1/n\. Used in proof of \NSLIM \ LIM\.\ lemma real_of_nat_less_inverse_iff: "0 < u \ u < inverse (real(Suc n)) \ real(Suc n) < inverse u" - apply (simp add: inverse_eq_divide) - apply (subst pos_less_divide_eq, assumption) - apply (subst pos_less_divide_eq) - apply simp - apply (simp add: mult.commute) - done + using less_imp_inverse_less by force lemma finite_inverse_real_of_posnat_gt_real: "0 < u \ finite {n. u < inverse (real (Suc n))}" proof (simp only: real_of_nat_less_inverse_iff)