# HG changeset patch # User wenzelm # Date 1477957464 -3600 # Node ID c93b0e6131c3330d124f2826ed8a35db70f600b8 # Parent af5235830c1699b77ffbe36c1eb9a0019ccd5d24 misc tuning and modernization; diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/Free_Ultrafilter.thy --- a/src/HOL/Nonstandard_Analysis/Free_Ultrafilter.thy Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/Free_Ultrafilter.thy Tue Nov 01 00:44:24 2016 +0100 @@ -2,14 +2,15 @@ Author: Jacques D. Fleuriot, University of Cambridge Author: Lawrence C Paulson Author: Brian Huffman -*) +*) section \Filters and Ultrafilters\ theory Free_Ultrafilter -imports "~~/src/HOL/Library/Infinite_Set" + imports "~~/src/HOL/Library/Infinite_Set" begin + subsection \Definitions and basic properties\ subsubsection \Ultrafilters\ @@ -43,28 +44,29 @@ end + subsection \Maximal filter = Ultrafilter\ text \ - A filter F is an ultrafilter iff it is a maximal filter, - i.e. whenever G is a filter and @{term "F \ G"} then @{term "F = G"} + A filter \F\ is an ultrafilter iff it is a maximal filter, + i.e. whenever \G\ is a filter and @{prop "F \ G"} then @{prop "F = G"} \ + text \ - Lemmas that shows existence of an extension to what was assumed to + Lemma that shows existence of an extension to what was assumed to be a maximal filter. Will be used to derive contradiction in proof of property of ultrafilter. \ -lemma extend_filter: - "frequently P F \ inf F (principal {x. P x}) \ bot" - unfolding trivial_limit_def eventually_inf_principal by (simp add: not_eventually) +lemma extend_filter: "frequently P F \ inf F (principal {x. P x}) \ bot" + by (simp add: trivial_limit_def eventually_inf_principal not_eventually) lemma max_filter_ultrafilter: - assumes proper: "F \ bot" + assumes "F \ bot" assumes max: "\G. G \ bot \ G \ F \ F = G" shows "ultrafilter F" proof - fix P show "eventually P F \ (\\<^sub>Fx in F. \ P x)" + show "eventually P F \ (\\<^sub>Fx in F. \ P x)" for P proof (rule disjCI) assume "\ (\\<^sub>Fx in F. \ P x)" then have "inf F (principal {x. P x}) \ bot" @@ -84,7 +86,9 @@ done lemma (in ultrafilter) max_filter: - assumes G: "G \ bot" and sub: "G \ F" shows "F = G" + assumes G: "G \ bot" + and sub: "G \ F" + shows "F = G" proof (rule antisym) show "F \ G" using sub @@ -92,10 +96,9 @@ intro!: eventually_frequently G proper) qed fact + subsection \Ultrafilter Theorem\ -text "A local context makes proof of ultrafilter Theorem more modular" - lemma ex_max_ultrafilter: fixes F :: "'a filter" assumes F: "F \ bot" @@ -104,73 +107,77 @@ let ?X = "{G. G \ bot \ G \ F}" let ?R = "{(b, a). a \ bot \ a \ b \ b \ F}" - have bot_notin_R: "\c. c \ Chains ?R \ bot \ c" + have bot_notin_R: "c \ Chains ?R \ bot \ c" for c by (auto simp: Chains_def) have [simp]: "Field ?R = ?X" by (auto simp: Field_def bot_unique) - have "\m\Field ?R. \a\Field ?R. (m, a) \ ?R \ a = m" + have "\m\Field ?R. \a\Field ?R. (m, a) \ ?R \ a = m" (is "\m\?A. ?B m") proof (rule Zorns_po_lemma) show "Partial_order ?R" - unfolding partial_order_on_def preorder_on_def - by (auto simp: antisym_def refl_on_def trans_def Field_def bot_unique) + by (auto simp: partial_order_on_def preorder_on_def + antisym_def refl_on_def trans_def Field_def bot_unique) show "\C\Chains ?R. \u\Field ?R. \a\C. (a, u) \ ?R" proof (safe intro!: bexI del: notI) - fix c x assume c: "c \ Chains ?R" + fix c x + assume c: "c \ Chains ?R" - { assume "c \ {}" - with c have "Inf c = bot \ (\x\c. x = bot)" + have Inf_c: "Inf c \ bot" "Inf c \ F" if "c \ {}" + proof - + from c that have "Inf c = bot \ (\x\c. x = bot)" unfolding trivial_limit_def by (intro eventually_Inf_base) (auto simp: Chains_def) - with c have 1: "Inf c \ bot" + with c show "Inf c \ bot" by (simp add: bot_notin_R) - from \c \ {}\ obtain x where "x \ c" by auto - with c have 2: "Inf c \ F" + from that obtain x where "x \ c" by auto + with c show "Inf c \ F" by (auto intro!: Inf_lower2[of x] simp: Chains_def) - note 1 2 } - note Inf_c = this + qed then have [simp]: "inf F (Inf c) = (if c = {} then F else Inf c)" using c by (auto simp add: inf_absorb2) - show "inf F (Inf c) \ bot" - using c by (simp add: F Inf_c) + from c show "inf F (Inf c) \ bot" + by (simp add: F Inf_c) + from c show "inf F (Inf c) \ Field ?R" + by (simp add: Chains_def Inf_c F) - show "inf F (Inf c) \ Field ?R" - using c by (simp add: Chains_def Inf_c F) - - assume x: "x \ c" + assume "x \ c" with c show "inf F (Inf c) \ x" "x \ F" by (auto intro: Inf_lower simp: Chains_def) qed qed - then guess U .. - then show ?thesis - by (intro exI[of _ U] conjI max_filter_ultrafilter) auto + then obtain U where U: "U \ ?A" "?B U" .. + show ?thesis + proof + from U show "U \ F \ ultrafilter U" + by (auto intro!: max_filter_ultrafilter) + qed qed + subsubsection \Free Ultrafilters\ -text \There exists a free ultrafilter on any infinite set\ +text \There exists a free ultrafilter on any infinite set.\ locale freeultrafilter = ultrafilter + assumes infinite: "eventually P F \ infinite {x. P x}" begin lemma finite: "finite {x. P x} \ \ eventually P F" - by (erule contrapos_pn, erule infinite) + by (erule contrapos_pn) (erule infinite) lemma finite': "finite {x. \ P x} \ eventually P F" by (drule finite) (simp add: not_eventually frequently_eq_eventually) lemma le_cofinite: "F \ cofinite" by (intro filter_leI) - (auto simp add: eventually_cofinite not_eventually frequently_eq_eventually dest!: finite) + (auto simp add: eventually_cofinite not_eventually frequently_eq_eventually dest!: finite) lemma singleton: "\ eventually (\x. x = a) F" -by (rule finite, simp) + by (rule finite) simp lemma singleton': "\ eventually (op = a) F" -by (rule finite, simp) + by (rule finite) simp lemma ultrafilter: "ultrafilter F" .. @@ -186,7 +193,8 @@ interpret ultrafilter U by fact have "freeultrafilter U" proof - fix P assume "eventually P U" + fix P + assume "eventually P U" with proper have "frequently P U" by (rule eventually_frequently) then have "frequently P cofinite" diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/HDeriv.thy --- a/src/HOL/Nonstandard_Analysis/HDeriv.thy Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/HDeriv.thy Tue Nov 01 00:44:24 2016 +0100 @@ -4,284 +4,252 @@ Conversion to Isar and new proofs by Lawrence C Paulson, 2004 *) -section\Differentiation (Nonstandard)\ +section \Differentiation (Nonstandard)\ theory HDeriv -imports HLim + imports HLim begin -text\Nonstandard Definitions\ +text \Nonstandard Definitions.\ -definition - nsderiv :: "['a::real_normed_field \ 'a, 'a, 'a] \ bool" - ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where - "NSDERIV f x :> D = (\h \ Infinitesimal - {0}. - (( *f* f)(star_of x + h) - - star_of (f x))/h \ star_of D)" +definition nsderiv :: "['a::real_normed_field \ 'a, 'a, 'a] \ bool" + ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) + where "NSDERIV f x :> D \ + (\h \ Infinitesimal - {0}. (( *f* f)(star_of x + h) - star_of (f x)) / h \ star_of D)" -definition - NSdifferentiable :: "['a::real_normed_field \ 'a, 'a] \ bool" - (infixl "NSdifferentiable" 60) where - "f NSdifferentiable x = (\D. NSDERIV f x :> D)" +definition NSdifferentiable :: "['a::real_normed_field \ 'a, 'a] \ bool" + (infixl "NSdifferentiable" 60) + where "f NSdifferentiable x \ (\D. NSDERIV f x :> D)" -definition - increment :: "[real=>real,real,hypreal] => hypreal" where - "increment f x h = (@inc. f NSdifferentiable x & - inc = ( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x))" +definition increment :: "(real \ real) \ real \ hypreal \ hypreal" + where "increment f x h = + (SOME inc. f NSdifferentiable x \ inc = ( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x))" subsection \Derivatives\ -lemma DERIV_NS_iff: - "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) \0\\<^sub>N\<^sub>S D)" -by (simp add: DERIV_def LIM_NSLIM_iff) +lemma DERIV_NS_iff: "(DERIV f x :> D) \ (\h. (f (x + h) - f x) / h) \0\\<^sub>N\<^sub>S D" + by (simp add: DERIV_def LIM_NSLIM_iff) -lemma NS_DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) \0\\<^sub>N\<^sub>S D" -by (simp add: DERIV_def LIM_NSLIM_iff) +lemma NS_DERIV_D: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\\<^sub>N\<^sub>S D" + by (simp add: DERIV_def LIM_NSLIM_iff) -lemma hnorm_of_hypreal: - "\r. hnorm (( *f* of_real) r::'a::real_normed_div_algebra star) = \r\" -by transfer (rule norm_of_real) +lemma hnorm_of_hypreal: "\r. hnorm (( *f* of_real) r::'a::real_normed_div_algebra star) = \r\" + by transfer (rule norm_of_real) lemma Infinitesimal_of_hypreal: - "x \ Infinitesimal \ - (( *f* of_real) x::'a::real_normed_div_algebra star) \ Infinitesimal" -apply (rule InfinitesimalI2) -apply (drule (1) InfinitesimalD2) -apply (simp add: hnorm_of_hypreal) -done + "x \ Infinitesimal \ (( *f* of_real) x::'a::real_normed_div_algebra star) \ Infinitesimal" + apply (rule InfinitesimalI2) + apply (drule (1) InfinitesimalD2) + apply (simp add: hnorm_of_hypreal) + done -lemma of_hypreal_eq_0_iff: - "\x. (( *f* of_real) x = (0::'a::real_algebra_1 star)) = (x = 0)" -by transfer (rule of_real_eq_0_iff) +lemma of_hypreal_eq_0_iff: "\x. (( *f* of_real) x = (0::'a::real_algebra_1 star)) = (x = 0)" + by transfer (rule of_real_eq_0_iff) -lemma NSDeriv_unique: - "[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E" -apply (subgoal_tac "( *f* of_real) \ \ Infinitesimal - {0::'a star}") -apply (simp only: nsderiv_def) -apply (drule (1) bspec)+ -apply (drule (1) approx_trans3) -apply simp -apply (simp add: Infinitesimal_of_hypreal Infinitesimal_epsilon) -apply (simp add: of_hypreal_eq_0_iff hypreal_epsilon_not_zero) -done +lemma NSDeriv_unique: "NSDERIV f x :> D \ NSDERIV f x :> E \ D = E" + apply (subgoal_tac "( *f* of_real) \ \ Infinitesimal - {0::'a star}") + apply (simp only: nsderiv_def) + apply (drule (1) bspec)+ + apply (drule (1) approx_trans3) + apply simp + apply (simp add: Infinitesimal_of_hypreal Infinitesimal_epsilon) + apply (simp add: of_hypreal_eq_0_iff hypreal_epsilon_not_zero) + done -text \First NSDERIV in terms of NSLIM\ +text \First \NSDERIV\ in terms of \NSLIM\.\ -text\first equivalence\ -lemma NSDERIV_NSLIM_iff: - "(NSDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) \0\\<^sub>N\<^sub>S D)" -apply (simp add: nsderiv_def NSLIM_def, auto) -apply (drule_tac x = xa in bspec) -apply (rule_tac [3] ccontr) -apply (drule_tac [3] x = h in spec) -apply (auto simp add: mem_infmal_iff starfun_lambda_cancel) -done +text \First equivalence.\ +lemma NSDERIV_NSLIM_iff: "(NSDERIV f x :> D) \ (\h. (f (x + h) - f x) / h) \0\\<^sub>N\<^sub>S D" + apply (auto simp add: nsderiv_def NSLIM_def) + apply (drule_tac x = xa in bspec) + apply (rule_tac [3] ccontr) + apply (drule_tac [3] x = h in spec) + apply (auto simp add: mem_infmal_iff starfun_lambda_cancel) + done -text\second equivalence\ -lemma NSDERIV_NSLIM_iff2: - "(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) \x\\<^sub>N\<^sub>S D)" +text \Second equivalence.\ +lemma NSDERIV_NSLIM_iff2: "(NSDERIV f x :> D) \ (\z. (f z - f x) / (z - x)) \x\\<^sub>N\<^sub>S D" by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff LIM_NSLIM_iff [symmetric]) -(* while we're at it! *) - +text \While we're at it!\ lemma NSDERIV_iff2: - "(NSDERIV f x :> D) = - (\w. - w \ star_of x & w \ star_of x --> - ( *f* (%z. (f z - f x) / (z-x))) w \ star_of D)" -by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def) + "(NSDERIV f x :> D) \ + (\w. w \ star_of x & w \ star_of x \ ( *f* (%z. (f z - f x) / (z-x))) w \ star_of D)" + by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def) -(*FIXME DELETE*) -lemma hypreal_not_eq_minus_iff: - "(x \ a) = (x - a \ (0::'a::ab_group_add))" -by auto +(* FIXME delete *) +lemma hypreal_not_eq_minus_iff: "x \ a \ x - a \ (0::'a::ab_group_add)" + by auto lemma NSDERIVD5: - "(NSDERIV f x :> D) ==> - (\u. u \ hypreal_of_real x --> - ( *f* (%z. f z - f x)) u \ hypreal_of_real D * (u - hypreal_of_real x))" -apply (auto simp add: NSDERIV_iff2) -apply (case_tac "u = hypreal_of_real x", auto) -apply (drule_tac x = u in spec, auto) -apply (drule_tac c = "u - hypreal_of_real x" and b = "hypreal_of_real D" in approx_mult1) -apply (drule_tac [!] hypreal_not_eq_minus_iff [THEN iffD1]) -apply (subgoal_tac [2] "( *f* (%z. z-x)) u \ (0::hypreal) ") -apply (auto simp add: - approx_minus_iff [THEN iffD1, THEN mem_infmal_iff [THEN iffD2]] - Infinitesimal_subset_HFinite [THEN subsetD]) -done + "(NSDERIV f x :> D) \ + (\u. u \ hypreal_of_real x \ + ( *f* (\z. f z - f x)) u \ hypreal_of_real D * (u - hypreal_of_real x))" + apply (auto simp add: NSDERIV_iff2) + apply (case_tac "u = hypreal_of_real x", auto) + apply (drule_tac x = u in spec, auto) + apply (drule_tac c = "u - hypreal_of_real x" and b = "hypreal_of_real D" in approx_mult1) + apply (drule_tac [!] hypreal_not_eq_minus_iff [THEN iffD1]) + apply (subgoal_tac [2] "( *f* (%z. z-x)) u \ (0::hypreal) ") + apply (auto simp: approx_minus_iff [THEN iffD1, THEN mem_infmal_iff [THEN iffD2]] + Infinitesimal_subset_HFinite [THEN subsetD]) + done lemma NSDERIVD4: - "(NSDERIV f x :> D) ==> - (\h \ Infinitesimal. - (( *f* f)(hypreal_of_real x + h) - - hypreal_of_real (f x))\ (hypreal_of_real D) * h)" -apply (auto simp add: nsderiv_def) -apply (case_tac "h = (0::hypreal) ") -apply auto -apply (drule_tac x = h in bspec) -apply (drule_tac [2] c = h in approx_mult1) -apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]) -done + "(NSDERIV f x :> D) \ + (\h \ Infinitesimal. + ( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x) \ hypreal_of_real D * h)" + apply (auto simp add: nsderiv_def) + apply (case_tac "h = 0") + apply auto + apply (drule_tac x = h in bspec) + apply (drule_tac [2] c = h in approx_mult1) + apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]) + done lemma NSDERIVD3: - "(NSDERIV f x :> D) ==> - (\h \ Infinitesimal - {0}. - (( *f* f)(hypreal_of_real x + h) - - hypreal_of_real (f x))\ (hypreal_of_real D) * h)" -apply (auto simp add: nsderiv_def) -apply (rule ccontr, drule_tac x = h in bspec) -apply (drule_tac [2] c = h in approx_mult1) -apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult.assoc) -done + "(NSDERIV f x :> D) \ + \h \ Infinitesimal - {0}. + (( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)) \ hypreal_of_real D * h" + apply (auto simp add: nsderiv_def) + apply (rule ccontr, drule_tac x = h in bspec) + apply (drule_tac [2] c = h in approx_mult1) + apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult.assoc) + done -text\Differentiability implies continuity - nice and simple "algebraic" proof\ -lemma NSDERIV_isNSCont: "NSDERIV f x :> D ==> isNSCont f x" -apply (auto simp add: nsderiv_def isNSCont_NSLIM_iff NSLIM_def) -apply (drule approx_minus_iff [THEN iffD1]) -apply (drule hypreal_not_eq_minus_iff [THEN iffD1]) -apply (drule_tac x = "xa - star_of x" in bspec) - prefer 2 apply (simp add: add.assoc [symmetric]) -apply (auto simp add: mem_infmal_iff [symmetric] add.commute) -apply (drule_tac c = "xa - star_of x" in approx_mult1) -apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] - simp add: mult.assoc nonzero_mult_div_cancel_right) -apply (drule_tac x3=D in - HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult, - THEN mem_infmal_iff [THEN iffD1]]) -apply (auto simp add: mult.commute - intro: approx_trans approx_minus_iff [THEN iffD2]) -done +text \Differentiability implies continuity nice and simple "algebraic" proof.\ +lemma NSDERIV_isNSCont: "NSDERIV f x :> D \ isNSCont f x" + apply (auto simp add: nsderiv_def isNSCont_NSLIM_iff NSLIM_def) + apply (drule approx_minus_iff [THEN iffD1]) + apply (drule hypreal_not_eq_minus_iff [THEN iffD1]) + apply (drule_tac x = "xa - star_of x" in bspec) + prefer 2 apply (simp add: add.assoc [symmetric]) + apply (auto simp add: mem_infmal_iff [symmetric] add.commute) + apply (drule_tac c = "xa - star_of x" in approx_mult1) + apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mult.assoc) + apply (drule_tac x3=D in + HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult, THEN mem_infmal_iff [THEN iffD1]]) + apply (auto simp add: mult.commute intro: approx_trans approx_minus_iff [THEN iffD2]) + done -text\Differentiation rules for combinations of functions - follow from clear, straightforard, algebraic - manipulations\ -text\Constant function\ +text \Differentiation rules for combinations of functions + follow from clear, straightforard, algebraic manipulations.\ + +text \Constant function.\ (* use simple constant nslimit theorem *) -lemma NSDERIV_const [simp]: "(NSDERIV (%x. k) x :> 0)" -by (simp add: NSDERIV_NSLIM_iff) - -text\Sum of functions- proved easily\ +lemma NSDERIV_const [simp]: "NSDERIV (\x. k) x :> 0" + by (simp add: NSDERIV_NSLIM_iff) -lemma NSDERIV_add: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] - ==> NSDERIV (%x. f x + g x) x :> Da + Db" -apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) -apply (auto simp add: add_divide_distrib diff_divide_distrib dest!: spec) -apply (drule_tac b = "star_of Da" and d = "star_of Db" in approx_add) -apply (auto simp add: ac_simps algebra_simps) -done - -text\Product of functions - Proof is trivial but tedious - and long due to rearrangement of terms\ +text \Sum of functions- proved easily.\ -lemma lemma_nsderiv1: - fixes a b c d :: "'a::comm_ring star" - shows "(a*b) - (c*d) = (b*(a - c)) + (c*(b - d))" -by (simp add: right_diff_distrib ac_simps) +lemma NSDERIV_add: + "NSDERIV f x :> Da \ NSDERIV g x :> Db \ NSDERIV (\x. f x + g x) x :> Da + Db" + apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) + apply (auto simp add: add_divide_distrib diff_divide_distrib dest!: spec) + apply (drule_tac b = "star_of Da" and d = "star_of Db" in approx_add) + apply (auto simp add: ac_simps algebra_simps) + done -lemma lemma_nsderiv2: - fixes x y z :: "'a::real_normed_field star" - shows "[| (x - y) / z = star_of D + yb; z \ 0; - z \ Infinitesimal; yb \ Infinitesimal |] - ==> x - y \ 0" -apply (simp add: nonzero_divide_eq_eq) -apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add - simp add: mult.assoc mem_infmal_iff [symmetric]) -apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) -done +text \Product of functions - Proof is trivial but tedious + and long due to rearrangement of terms.\ + +lemma lemma_nsderiv1: "(a * b) - (c * d) = (b * (a - c)) + (c * (b - d))" + for a b c d :: "'a::comm_ring star" + by (simp add: right_diff_distrib ac_simps) -lemma NSDERIV_mult: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] - ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" -apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) -apply (auto dest!: spec - simp add: starfun_lambda_cancel lemma_nsderiv1) -apply (simp (no_asm) add: add_divide_distrib diff_divide_distrib) -apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+ -apply (auto simp add: times_divide_eq_right [symmetric] - simp del: times_divide_eq_right times_divide_eq_left) -apply (drule_tac D = Db in lemma_nsderiv2, assumption+) -apply (drule_tac - approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]]) -apply (auto intro!: approx_add_mono1 - simp add: distrib_right distrib_left mult.commute add.assoc) -apply (rule_tac b1 = "star_of Db * star_of (f x)" - in add.commute [THEN subst]) -apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym] - Infinitesimal_add Infinitesimal_mult - Infinitesimal_star_of_mult - Infinitesimal_star_of_mult2 - simp add: add.assoc [symmetric]) -done +lemma lemma_nsderiv2: "(x - y) / z = star_of D + yb \ z \ 0 \ + z \ Infinitesimal \ yb \ Infinitesimal \ x - y \ 0" + for x y z :: "'a::real_normed_field star" + apply (simp add: nonzero_divide_eq_eq) + apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add + simp add: mult.assoc mem_infmal_iff [symmetric]) + apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) + done -text\Multiplying by a constant\ -lemma NSDERIV_cmult: "NSDERIV f x :> D - ==> NSDERIV (%x. c * f x) x :> c*D" -apply (simp only: times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff - minus_mult_right right_diff_distrib [symmetric]) -apply (erule NSLIM_const [THEN NSLIM_mult]) -done +lemma NSDERIV_mult: + "NSDERIV f x :> Da \ NSDERIV g x :> Db \ + NSDERIV (\x. f x * g x) x :> (Da * g x) + (Db * f x)" + apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) + apply (auto dest!: spec simp add: starfun_lambda_cancel lemma_nsderiv1) + apply (simp (no_asm) add: add_divide_distrib diff_divide_distrib) + apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+ + apply (auto simp add: times_divide_eq_right [symmetric] + simp del: times_divide_eq_right times_divide_eq_left) + apply (drule_tac D = Db in lemma_nsderiv2, assumption+) + apply (drule_tac approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]]) + apply (auto intro!: approx_add_mono1 simp: distrib_right distrib_left mult.commute add.assoc) + apply (rule_tac b1 = "star_of Db * star_of (f x)" in add.commute [THEN subst]) + apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym] + Infinitesimal_add Infinitesimal_mult Infinitesimal_star_of_mult Infinitesimal_star_of_mult2 + simp add: add.assoc [symmetric]) + done -text\Negation of function\ -lemma NSDERIV_minus: "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D" +text \Multiplying by a constant.\ +lemma NSDERIV_cmult: "NSDERIV f x :> D \ NSDERIV (\x. c * f x) x :> c * D" + apply (simp only: times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff + minus_mult_right right_diff_distrib [symmetric]) + apply (erule NSLIM_const [THEN NSLIM_mult]) + done + +text \Negation of function.\ +lemma NSDERIV_minus: "NSDERIV f x :> D \ NSDERIV (\x. - f x) x :> - D" proof (simp add: NSDERIV_NSLIM_iff) assume "(\h. (f (x + h) - f x) / h) \0\\<^sub>N\<^sub>S D" - hence deriv: "(\h. - ((f(x+h) - f x) / h)) \0\\<^sub>N\<^sub>S - D" + then have deriv: "(\h. - ((f(x+h) - f x) / h)) \0\\<^sub>N\<^sub>S - D" by (rule NSLIM_minus) have "\h. - ((f (x + h) - f x) / h) = (- f (x + h) + f x) / h" by (simp add: minus_divide_left) - with deriv - have "(\h. (- f (x + h) + f x) / h) \0\\<^sub>N\<^sub>S - D" by simp - then show "(\h. (f (x + h) - f x) / h) \0\\<^sub>N\<^sub>S D \ - (\h. (f x - f (x + h)) / h) \0\\<^sub>N\<^sub>S - D" by simp + with deriv have "(\h. (- f (x + h) + f x) / h) \0\\<^sub>N\<^sub>S - D" + by simp + then show "(\h. (f (x + h) - f x) / h) \0\\<^sub>N\<^sub>S D \ (\h. (f x - f (x + h)) / h) \0\\<^sub>N\<^sub>S - D" + by simp qed -text\Subtraction\ -lemma NSDERIV_add_minus: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x + -g x) x :> Da + -Db" -by (blast dest: NSDERIV_add NSDERIV_minus) +text \Subtraction.\ +lemma NSDERIV_add_minus: + "NSDERIV f x :> Da \ NSDERIV g x :> Db \ NSDERIV (\x. f x + - g x) x :> Da + - Db" + by (blast dest: NSDERIV_add NSDERIV_minus) lemma NSDERIV_diff: - "NSDERIV f x :> Da \ NSDERIV g x :> Db \ NSDERIV (\x. f x - g x) x :> Da-Db" + "NSDERIV f x :> Da \ NSDERIV g x :> Db \ NSDERIV (\x. f x - g x) x :> Da - Db" using NSDERIV_add_minus [of f x Da g Db] by simp -text\Similarly to the above, the chain rule admits an entirely - straightforward derivation. Compare this with Harrison's - HOL proof of the chain rule, which proved to be trickier and - required an alternative characterisation of differentiability- - the so-called Carathedory derivative. Our main problem is - manipulation of terms.\ +text \Similarly to the above, the chain rule admits an entirely + straightforward derivation. Compare this with Harrison's + HOL proof of the chain rule, which proved to be trickier and + required an alternative characterisation of differentiability- + the so-called Carathedory derivative. Our main problem is + manipulation of terms.\ -(* lemmas *) + +subsection \Lemmas\ lemma NSDERIV_zero: - "[| NSDERIV g x :> D; - ( *f* g) (star_of x + xa) = star_of (g x); - xa \ Infinitesimal; - xa \ 0 - |] ==> D = 0" -apply (simp add: nsderiv_def) -apply (drule bspec, auto) -done + "NSDERIV g x :> D \ ( *f* g) (star_of x + xa) = star_of (g x) \ + xa \ Infinitesimal \ xa \ 0 \ D = 0" + apply (simp add: nsderiv_def) + apply (drule bspec) + apply auto + done -(* can be proved differently using NSLIM_isCont_iff *) +text \Can be proved differently using \NSLIM_isCont_iff\.\ lemma NSDERIV_approx: - "[| NSDERIV f x :> D; h \ Infinitesimal; h \ 0 |] - ==> ( *f* f) (star_of x + h) - star_of (f x) \ 0" -apply (simp add: nsderiv_def) -apply (simp add: mem_infmal_iff [symmetric]) -apply (rule Infinitesimal_ratio) -apply (rule_tac [3] approx_star_of_HFinite, auto) -done + "NSDERIV f x :> D \ h \ Infinitesimal \ h \ 0 \ + ( *f* f) (star_of x + h) - star_of (f x) \ 0" + apply (simp add: nsderiv_def) + apply (simp add: mem_infmal_iff [symmetric]) + apply (rule Infinitesimal_ratio) + apply (rule_tac [3] approx_star_of_HFinite, auto) + done -(*--------------------------------------------------------------- - from one version of differentiability +text \From one version of differentiability - f(x) - f(a) - --------------- \ Db - x - a - ---------------------------------------------------------------*) + \f x - f a\ + \-------------- \ Db\ + \x - a\ +\ lemma NSDERIVD1: "[| NSDERIV f (g x) :> Da; ( *f* g) (star_of(x) + xa) \ star_of (g x); @@ -290,186 +258,185 @@ - star_of (f (g x))) / (( *f* g) (star_of(x) + xa) - star_of (g x)) \ star_of(Da)" -by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def) + by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def) -(*-------------------------------------------------------------- - from other version of differentiability +text \From other version of differentiability - f(x + h) - f(x) - ----------------- \ Db - h - --------------------------------------------------------------*) + \f (x + h) - f x\ + \------------------ \ Db\ + \h\ +\ lemma NSDERIVD2: "[| NSDERIV g x :> Db; xa \ Infinitesimal; xa \ 0 |] ==> (( *f* g) (star_of(x) + xa) - star_of(g x)) / xa \ star_of(Db)" -by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff starfun_lambda_cancel) + by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff starfun_lambda_cancel) -lemma lemma_chain: "(z::'a::real_normed_field star) \ 0 ==> x*y = (x*inverse(z))*(z*y)" +lemma lemma_chain: "z \ 0 \ x * y = (x * inverse z) * (z * y)" + for x y z :: "'a::real_normed_field star" proof - assume z: "z \ 0" have "x * y = x * (inverse z * z) * y" by (simp add: z) - thus ?thesis by (simp add: mult.assoc) + then show ?thesis by (simp add: mult.assoc) qed -text\This proof uses both definitions of differentiability.\ -lemma NSDERIV_chain: "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] - ==> NSDERIV (f o g) x :> Da * Db" -apply (simp (no_asm_simp) add: NSDERIV_NSLIM_iff NSLIM_def - mem_infmal_iff [symmetric]) -apply clarify -apply (frule_tac f = g in NSDERIV_approx) -apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric]) -apply (case_tac "( *f* g) (star_of (x) + xa) = star_of (g x) ") -apply (drule_tac g = g in NSDERIV_zero) -apply (auto simp add: divide_inverse) -apply (rule_tac z1 = "( *f* g) (star_of (x) + xa) - star_of (g x) " and y1 = "inverse xa" in lemma_chain [THEN ssubst]) -apply (erule hypreal_not_eq_minus_iff [THEN iffD1]) -apply (rule approx_mult_star_of) -apply (simp_all add: divide_inverse [symmetric]) -apply (blast intro: NSDERIVD1 approx_minus_iff [THEN iffD2]) -apply (blast intro: NSDERIVD2) -done +text \This proof uses both definitions of differentiability.\ +lemma NSDERIV_chain: + "NSDERIV f (g x) :> Da \ NSDERIV g x :> Db \ NSDERIV (f \ g) x :> Da * Db" + apply (simp (no_asm_simp) add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff [symmetric]) + apply clarify + apply (frule_tac f = g in NSDERIV_approx) + apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric]) + apply (case_tac "( *f* g) (star_of (x) + xa) = star_of (g x) ") + apply (drule_tac g = g in NSDERIV_zero) + apply (auto simp add: divide_inverse) + apply (rule_tac z1 = "( *f* g) (star_of (x) + xa) - star_of (g x) " and y1 = "inverse xa" in lemma_chain [THEN ssubst]) + apply (erule hypreal_not_eq_minus_iff [THEN iffD1]) + apply (rule approx_mult_star_of) + apply (simp_all add: divide_inverse [symmetric]) + apply (blast intro: NSDERIVD1 approx_minus_iff [THEN iffD2]) + apply (blast intro: NSDERIVD2) + done -text\Differentiation of natural number powers\ -lemma NSDERIV_Id [simp]: "NSDERIV (%x. x) x :> 1" -by (simp add: NSDERIV_NSLIM_iff NSLIM_def del: divide_self_if) +text \Differentiation of natural number powers.\ +lemma NSDERIV_Id [simp]: "NSDERIV (\x. x) x :> 1" + by (simp add: NSDERIV_NSLIM_iff NSLIM_def del: divide_self_if) lemma NSDERIV_cmult_Id [simp]: "NSDERIV (op * c) x :> c" -by (cut_tac c = c and x = x in NSDERIV_Id [THEN NSDERIV_cmult], simp) + using NSDERIV_Id [THEN NSDERIV_cmult] by simp lemma NSDERIV_inverse: fixes x :: "'a::real_normed_field" assumes "x \ 0" \ \can't get rid of @{term "x \ 0"} because it isn't continuous at zero\ shows "NSDERIV (\x. inverse x) x :> - (inverse x ^ Suc (Suc 0))" proof - - { fix h :: "'a star" + { + fix h :: "'a star" assume h_Inf: "h \ Infinitesimal" - from this assms have not_0: "star_of x + h \ 0" by (rule Infinitesimal_add_not_zero) + from this assms have not_0: "star_of x + h \ 0" + by (rule Infinitesimal_add_not_zero) assume "h \ 0" - from h_Inf have "h * star_of x \ Infinitesimal" by (rule Infinitesimal_HFinite_mult) simp + from h_Inf have "h * star_of x \ Infinitesimal" + by (rule Infinitesimal_HFinite_mult) simp with assms have "inverse (- (h * star_of x) + - (star_of x * star_of x)) \ inverse (- (star_of x * star_of x))" - apply - apply (rule inverse_add_Infinitesimal_approx2) - apply (auto - dest!: hypreal_of_real_HFinite_diff_Infinitesimal + apply - + apply (rule inverse_add_Infinitesimal_approx2) + apply (auto dest!: hypreal_of_real_HFinite_diff_Infinitesimal simp add: inverse_minus_eq [symmetric] HFinite_minus_iff) done moreover from not_0 \h \ 0\ assms - have "inverse (- (h * star_of x) + - (star_of x * star_of x)) = - (inverse (star_of x + h) - inverse (star_of x)) / h" + have "inverse (- (h * star_of x) + - (star_of x * star_of x)) = + (inverse (star_of x + h) - inverse (star_of x)) / h" apply (simp add: division_ring_inverse_diff nonzero_inverse_mult_distrib [symmetric] - nonzero_inverse_minus_eq [symmetric] ac_simps ring_distribs) + nonzero_inverse_minus_eq [symmetric] ac_simps ring_distribs) apply (subst nonzero_inverse_minus_eq [symmetric]) using distrib_right [symmetric, of h "star_of x" "star_of x"] apply simp apply (simp add: field_simps) done ultimately have "(inverse (star_of x + h) - inverse (star_of x)) / h \ - (inverse (star_of x) * inverse (star_of x))" - using assms by simp - } then show ?thesis by (simp add: nsderiv_def) + using assms by simp + } + then show ?thesis by (simp add: nsderiv_def) qed + subsubsection \Equivalence of NS and Standard definitions\ lemma divideR_eq_divide: "x /\<^sub>R y = x / y" -by (simp add: divide_inverse mult.commute) + by (simp add: divide_inverse mult.commute) -text\Now equivalence between NSDERIV and DERIV\ -lemma NSDERIV_DERIV_iff: "(NSDERIV f x :> D) = (DERIV f x :> D)" -by (simp add: DERIV_def NSDERIV_NSLIM_iff LIM_NSLIM_iff) +text \Now equivalence between \NSDERIV\ and \DERIV\.\ +lemma NSDERIV_DERIV_iff: "NSDERIV f x :> D \ DERIV f x :> D" + by (simp add: DERIV_def NSDERIV_NSLIM_iff LIM_NSLIM_iff) -(* NS version *) -lemma NSDERIV_pow: "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" -by (simp add: NSDERIV_DERIV_iff DERIV_pow) +text \NS version.\ +lemma NSDERIV_pow: "NSDERIV (\x. x ^ n) x :> real n * (x ^ (n - Suc 0))" + by (simp add: NSDERIV_DERIV_iff DERIV_pow) -text\Derivative of inverse\ - +text \Derivative of inverse.\ lemma NSDERIV_inverse_fun: - fixes x :: "'a::{real_normed_field}" - shows "[| NSDERIV f x :> d; f(x) \ 0 |] - ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" -by (simp add: NSDERIV_DERIV_iff DERIV_inverse_fun del: power_Suc) + "NSDERIV f x :> d \ f x \ 0 \ + NSDERIV (\x. inverse (f x)) x :> (- (d * inverse (f x ^ Suc (Suc 0))))" + for x :: "'a::{real_normed_field}" + by (simp add: NSDERIV_DERIV_iff DERIV_inverse_fun del: power_Suc) -text\Derivative of quotient\ - +text \Derivative of quotient.\ lemma NSDERIV_quotient: - fixes x :: "'a::{real_normed_field}" - shows "[| NSDERIV f x :> d; NSDERIV g x :> e; g(x) \ 0 |] - ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) - - (e*f(x))) / (g(x) ^ Suc (Suc 0))" -by (simp add: NSDERIV_DERIV_iff DERIV_quotient del: power_Suc) + fixes x :: "'a::real_normed_field" + shows "NSDERIV f x :> d \ NSDERIV g x :> e \ g x \ 0 \ + NSDERIV (\y. f y / g y) x :> (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))" + by (simp add: NSDERIV_DERIV_iff DERIV_quotient del: power_Suc) -lemma CARAT_NSDERIV: "NSDERIV f x :> l ==> - \g. (\z. f z - f x = g z * (z-x)) & isNSCont g x & g x = l" -by (auto simp add: NSDERIV_DERIV_iff isNSCont_isCont_iff CARAT_DERIV - mult.commute) +lemma CARAT_NSDERIV: + "NSDERIV f x :> l \ \g. (\z. f z - f x = g z * (z - x)) \ isNSCont g x \ g x = l" + by (auto simp add: NSDERIV_DERIV_iff isNSCont_isCont_iff CARAT_DERIV mult.commute) -lemma hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))" -by auto +lemma hypreal_eq_minus_iff3: "x = y + z \ x + - z = y" + for x y z :: hypreal + by auto lemma CARAT_DERIVD: - assumes all: "\z. f z - f x = g z * (z-x)" - and nsc: "isNSCont g x" + assumes all: "\z. f z - f x = g z * (z - x)" + and nsc: "isNSCont g x" shows "NSDERIV f x :> g x" proof - - from nsc - have "\w. w \ star_of x \ w \ star_of x \ - ( *f* g) w * (w - star_of x) / (w - star_of x) \ - star_of (g x)" - by (simp add: isNSCont_def nonzero_mult_div_cancel_right) - thus ?thesis using all + from nsc have "\w. w \ star_of x \ w \ star_of x \ + ( *f* g) w * (w - star_of x) / (w - star_of x) \ star_of (g x)" + by (simp add: isNSCont_def) + with all show ?thesis by (simp add: NSDERIV_iff2 starfun_if_eq cong: if_cong) qed + subsubsection \Differentiability predicate\ -lemma NSdifferentiableD: "f NSdifferentiable x ==> \D. NSDERIV f x :> D" -by (simp add: NSdifferentiable_def) +lemma NSdifferentiableD: "f NSdifferentiable x \ \D. NSDERIV f x :> D" + by (simp add: NSdifferentiable_def) -lemma NSdifferentiableI: "NSDERIV f x :> D ==> f NSdifferentiable x" -by (force simp add: NSdifferentiable_def) +lemma NSdifferentiableI: "NSDERIV f x :> D \ f NSdifferentiable x" + by (force simp add: NSdifferentiable_def) subsection \(NS) Increment\ -lemma incrementI: - "f NSdifferentiable x ==> - increment f x h = ( *f* f) (hypreal_of_real(x) + h) - - hypreal_of_real (f x)" -by (simp add: increment_def) -lemma incrementI2: "NSDERIV f x :> D ==> - increment f x h = ( *f* f) (hypreal_of_real(x) + h) - - hypreal_of_real (f x)" -apply (erule NSdifferentiableI [THEN incrementI]) -done +lemma incrementI: + "f NSdifferentiable x \ + increment f x h = ( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)" + by (simp add: increment_def) + +lemma incrementI2: + "NSDERIV f x :> D \ + increment f x h = ( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)" + by (erule NSdifferentiableI [THEN incrementI]) -(* The Increment theorem -- Keisler p. 65 *) -lemma increment_thm: "[| NSDERIV f x :> D; h \ Infinitesimal; h \ 0 |] - ==> \e \ Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h" -apply (frule_tac h = h in incrementI2, simp add: nsderiv_def) -apply (drule bspec, auto) -apply (drule bex_Infinitesimal_iff2 [THEN iffD2], clarify) -apply (frule_tac b1 = "hypreal_of_real (D) + y" - in mult_right_cancel [THEN iffD2]) -apply (erule_tac [2] V = "(( *f* f) (hypreal_of_real (x) + h) - hypreal_of_real (f x)) / h = hypreal_of_real (D) + y" in thin_rl) -apply assumption -apply (simp add: times_divide_eq_right [symmetric]) -apply (auto simp add: distrib_right) -done +text \The Increment theorem -- Keisler p. 65.\ +lemma increment_thm: + "NSDERIV f x :> D \ h \ Infinitesimal \ h \ 0 \ + \e \ Infinitesimal. increment f x h = hypreal_of_real D * h + e * h" + apply (frule_tac h = h in incrementI2, simp add: nsderiv_def) + apply (drule bspec, auto) + apply (drule bex_Infinitesimal_iff2 [THEN iffD2], clarify) + apply (frule_tac b1 = "hypreal_of_real D + y" in mult_right_cancel [THEN iffD2]) + apply (erule_tac [2] + V = "(( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)) / h = hypreal_of_real (D) + y" + in thin_rl) + apply assumption + apply (simp add: times_divide_eq_right [symmetric]) + apply (auto simp add: distrib_right) + done lemma increment_thm2: - "[| NSDERIV f x :> D; h \ 0; h \ 0 |] - ==> \e \ Infinitesimal. increment f x h = - hypreal_of_real(D)*h + e*h" -by (blast dest!: mem_infmal_iff [THEN iffD2] intro!: increment_thm) - + "NSDERIV f x :> D \ h \ 0 \ h \ 0 \ + \e \ Infinitesimal. increment f x h = hypreal_of_real D * h + e * h" + by (blast dest!: mem_infmal_iff [THEN iffD2] intro!: increment_thm) -lemma increment_approx_zero: "[| NSDERIV f x :> D; h \ 0; h \ 0 |] - ==> increment f x h \ 0" -apply (drule increment_thm2, - auto intro!: Infinitesimal_HFinite_mult2 HFinite_add simp add: distrib_right [symmetric] mem_infmal_iff [symmetric]) -apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) -done +lemma increment_approx_zero: "NSDERIV f x :> D \ h \ 0 \ h \ 0 \ increment f x h \ 0" + apply (drule increment_thm2) + apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add + simp add: distrib_right [symmetric] mem_infmal_iff [symmetric]) + apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) + done end diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/HLim.thy --- a/src/HOL/Nonstandard_Analysis/HLim.thy Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/HLim.thy Tue Nov 01 00:44:24 2016 +0100 @@ -3,330 +3,308 @@ Author: Lawrence C Paulson *) -section\Limits and Continuity (Nonstandard)\ +section \Limits and Continuity (Nonstandard)\ theory HLim imports Star abbrevs "--->" = "\\\<^sub>N\<^sub>S" begin -text\Nonstandard Definitions\ +text \Nonstandard Definitions.\ -definition - NSLIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool" - ("((_)/ \(_)/\\<^sub>N\<^sub>S (_))" [60, 0, 60] 60) where - "f \a\\<^sub>N\<^sub>S L = - (\x. (x \ star_of a & x \ star_of a --> ( *f* f) x \ star_of L))" +definition NSLIM :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ 'b \ bool" + ("((_)/ \(_)/\\<^sub>N\<^sub>S (_))" [60, 0, 60] 60) + where "f \a\\<^sub>N\<^sub>S L \ (\x. x \ star_of a \ x \ star_of a \ ( *f* f) x \ star_of L)" -definition - isNSCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where - \\NS definition dispenses with limit notions\ - "isNSCont f a = (\y. y \ star_of a --> - ( *f* f) y \ star_of (f a))" +definition isNSCont :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ bool" + where \ \NS definition dispenses with limit notions\ + "isNSCont f a \ (\y. y \ star_of a \ ( *f* f) y \ star_of (f a))" -definition - isNSUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where - "isNSUCont f = (\x y. x \ y --> ( *f* f) x \ ( *f* f) y)" +definition isNSUCont :: "('a::real_normed_vector \ 'b::real_normed_vector) \ bool" + where "isNSUCont f \ (\x y. x \ y \ ( *f* f) x \ ( *f* f) y)" subsection \Limits of Functions\ -lemma NSLIM_I: - "(\x. \x \ star_of a; x \ star_of a\ \ starfun f x \ star_of L) - \ f \a\\<^sub>N\<^sub>S L" -by (simp add: NSLIM_def) +lemma NSLIM_I: "(\x. x \ star_of a \ x \ star_of a \ starfun f x \ star_of L) \ f \a\\<^sub>N\<^sub>S L" + by (simp add: NSLIM_def) + +lemma NSLIM_D: "f \a\\<^sub>N\<^sub>S L \ x \ star_of a \ x \ star_of a \ starfun f x \ star_of L" + by (simp add: NSLIM_def) -lemma NSLIM_D: - "\f \a\\<^sub>N\<^sub>S L; x \ star_of a; x \ star_of a\ - \ starfun f x \ star_of L" -by (simp add: NSLIM_def) +text \Proving properties of limits using nonstandard definition. + The properties hold for standard limits as well!\ -text\Proving properties of limits using nonstandard definition. - The properties hold for standard limits as well!\ - -lemma NSLIM_mult: - fixes l m :: "'a::real_normed_algebra" - shows "[| f \x\\<^sub>N\<^sub>S l; g \x\\<^sub>N\<^sub>S m |] - ==> (%x. f(x) * g(x)) \x\\<^sub>N\<^sub>S (l * m)" -by (auto simp add: NSLIM_def intro!: approx_mult_HFinite) +lemma NSLIM_mult: "f \x\\<^sub>N\<^sub>S l \ g \x\\<^sub>N\<^sub>S m \ (\x. f x * g x) \x\\<^sub>N\<^sub>S (l * m)" + for l m :: "'a::real_normed_algebra" + by (auto simp add: NSLIM_def intro!: approx_mult_HFinite) -lemma starfun_scaleR [simp]: - "starfun (\x. f x *\<^sub>R g x) = (\x. scaleHR (starfun f x) (starfun g x))" -by transfer (rule refl) +lemma starfun_scaleR [simp]: "starfun (\x. f x *\<^sub>R g x) = (\x. scaleHR (starfun f x) (starfun g x))" + by transfer (rule refl) -lemma NSLIM_scaleR: - "[| f \x\\<^sub>N\<^sub>S l; g \x\\<^sub>N\<^sub>S m |] - ==> (%x. f(x) *\<^sub>R g(x)) \x\\<^sub>N\<^sub>S (l *\<^sub>R m)" -by (auto simp add: NSLIM_def intro!: approx_scaleR_HFinite) +lemma NSLIM_scaleR: "f \x\\<^sub>N\<^sub>S l \ g \x\\<^sub>N\<^sub>S m \ (\x. f x *\<^sub>R g x) \x\\<^sub>N\<^sub>S (l *\<^sub>R m)" + by (auto simp add: NSLIM_def intro!: approx_scaleR_HFinite) -lemma NSLIM_add: - "[| f \x\\<^sub>N\<^sub>S l; g \x\\<^sub>N\<^sub>S m |] - ==> (%x. f(x) + g(x)) \x\\<^sub>N\<^sub>S (l + m)" -by (auto simp add: NSLIM_def intro!: approx_add) +lemma NSLIM_add: "f \x\\<^sub>N\<^sub>S l \ g \x\\<^sub>N\<^sub>S m \ (\x. f x + g x) \x\\<^sub>N\<^sub>S (l + m)" + by (auto simp add: NSLIM_def intro!: approx_add) -lemma NSLIM_const [simp]: "(%x. k) \x\\<^sub>N\<^sub>S k" -by (simp add: NSLIM_def) +lemma NSLIM_const [simp]: "(\x. k) \x\\<^sub>N\<^sub>S k" + by (simp add: NSLIM_def) -lemma NSLIM_minus: "f \a\\<^sub>N\<^sub>S L ==> (%x. -f(x)) \a\\<^sub>N\<^sub>S -L" -by (simp add: NSLIM_def) +lemma NSLIM_minus: "f \a\\<^sub>N\<^sub>S L \ (\x. - f x) \a\\<^sub>N\<^sub>S -L" + by (simp add: NSLIM_def) -lemma NSLIM_diff: - "\f \x\\<^sub>N\<^sub>S l; g \x\\<^sub>N\<^sub>S m\ \ (\x. f x - g x) \x\\<^sub>N\<^sub>S (l - m)" +lemma NSLIM_diff: "f \x\\<^sub>N\<^sub>S l \ g \x\\<^sub>N\<^sub>S m \ (\x. f x - g x) \x\\<^sub>N\<^sub>S (l - m)" by (simp only: NSLIM_add NSLIM_minus diff_conv_add_uminus) -lemma NSLIM_add_minus: "[| f \x\\<^sub>N\<^sub>S l; g \x\\<^sub>N\<^sub>S m |] ==> (%x. f(x) + -g(x)) \x\\<^sub>N\<^sub>S (l + -m)" -by (simp only: NSLIM_add NSLIM_minus) +lemma NSLIM_add_minus: "f \x\\<^sub>N\<^sub>S l \ g \x\\<^sub>N\<^sub>S m \ (\x. f x + - g x) \x\\<^sub>N\<^sub>S (l + -m)" + by (simp only: NSLIM_add NSLIM_minus) -lemma NSLIM_inverse: - fixes L :: "'a::real_normed_div_algebra" - shows "[| f \a\\<^sub>N\<^sub>S L; L \ 0 |] - ==> (%x. inverse(f(x))) \a\\<^sub>N\<^sub>S (inverse L)" -apply (simp add: NSLIM_def, clarify) -apply (drule spec) -apply (auto simp add: star_of_approx_inverse) -done +lemma NSLIM_inverse: "f \a\\<^sub>N\<^sub>S L \ L \ 0 \ (\x. inverse (f x)) \a\\<^sub>N\<^sub>S (inverse L)" + for L :: "'a::real_normed_div_algebra" + apply (simp add: NSLIM_def, clarify) + apply (drule spec) + apply (auto simp add: star_of_approx_inverse) + done lemma NSLIM_zero: - assumes f: "f \a\\<^sub>N\<^sub>S l" shows "(%x. f(x) - l) \a\\<^sub>N\<^sub>S 0" + assumes f: "f \a\\<^sub>N\<^sub>S l" + shows "(\x. f(x) - l) \a\\<^sub>N\<^sub>S 0" proof - have "(\x. f x - l) \a\\<^sub>N\<^sub>S l - l" by (rule NSLIM_diff [OF f NSLIM_const]) - thus ?thesis by simp + then show ?thesis by simp qed -lemma NSLIM_zero_cancel: "(%x. f(x) - l) \x\\<^sub>N\<^sub>S 0 ==> f \x\\<^sub>N\<^sub>S l" -apply (drule_tac g = "%x. l" and m = l in NSLIM_add) -apply (auto simp add: add.assoc) -done +lemma NSLIM_zero_cancel: "(\x. f x - l) \x\\<^sub>N\<^sub>S 0 \ f \x\\<^sub>N\<^sub>S l" + apply (drule_tac g = "%x. l" and m = l in NSLIM_add) + apply (auto simp add: add.assoc) + done -lemma NSLIM_const_not_eq: - fixes a :: "'a::real_normed_algebra_1" - shows "k \ L \ \ (\x. k) \a\\<^sub>N\<^sub>S L" -apply (simp add: NSLIM_def) -apply (rule_tac x="star_of a + of_hypreal \" in exI) -apply (simp add: hypreal_epsilon_not_zero approx_def) -done +lemma NSLIM_const_not_eq: "k \ L \ \ (\x. k) \a\\<^sub>N\<^sub>S L" + for a :: "'a::real_normed_algebra_1" + apply (simp add: NSLIM_def) + apply (rule_tac x="star_of a + of_hypreal \" in exI) + apply (simp add: hypreal_epsilon_not_zero approx_def) + done -lemma NSLIM_not_zero: - fixes a :: "'a::real_normed_algebra_1" - shows "k \ 0 \ \ (\x. k) \a\\<^sub>N\<^sub>S 0" -by (rule NSLIM_const_not_eq) +lemma NSLIM_not_zero: "k \ 0 \ \ (\x. k) \a\\<^sub>N\<^sub>S 0" + for a :: "'a::real_normed_algebra_1" + by (rule NSLIM_const_not_eq) -lemma NSLIM_const_eq: - fixes a :: "'a::real_normed_algebra_1" - shows "(\x. k) \a\\<^sub>N\<^sub>S L \ k = L" -apply (rule ccontr) -apply (blast dest: NSLIM_const_not_eq) -done +lemma NSLIM_const_eq: "(\x. k) \a\\<^sub>N\<^sub>S L \ k = L" + for a :: "'a::real_normed_algebra_1" + by (rule ccontr) (blast dest: NSLIM_const_not_eq) + +lemma NSLIM_unique: "f \a\\<^sub>N\<^sub>S L \ f \a\\<^sub>N\<^sub>S M \ L = M" + for a :: "'a::real_normed_algebra_1" + by (drule (1) NSLIM_diff) (auto dest!: NSLIM_const_eq) -lemma NSLIM_unique: - fixes a :: "'a::real_normed_algebra_1" - shows "\f \a\\<^sub>N\<^sub>S L; f \a\\<^sub>N\<^sub>S M\ \ L = M" -apply (drule (1) NSLIM_diff) -apply (auto dest!: NSLIM_const_eq) -done +lemma NSLIM_mult_zero: "f \x\\<^sub>N\<^sub>S 0 \ g \x\\<^sub>N\<^sub>S 0 \ (\x. f x * g x) \x\\<^sub>N\<^sub>S 0" + for f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" + by (drule NSLIM_mult) auto -lemma NSLIM_mult_zero: - fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" - shows "[| f \x\\<^sub>N\<^sub>S 0; g \x\\<^sub>N\<^sub>S 0 |] ==> (%x. f(x)*g(x)) \x\\<^sub>N\<^sub>S 0" -by (drule NSLIM_mult, auto) +lemma NSLIM_self: "(\x. x) \a\\<^sub>N\<^sub>S a" + by (simp add: NSLIM_def) -lemma NSLIM_self: "(%x. x) \a\\<^sub>N\<^sub>S a" -by (simp add: NSLIM_def) subsubsection \Equivalence of @{term filterlim} and @{term NSLIM}\ lemma LIM_NSLIM: - assumes f: "f \a\ L" shows "f \a\\<^sub>N\<^sub>S L" + assumes f: "f \a\ L" + shows "f \a\\<^sub>N\<^sub>S L" proof (rule NSLIM_I) fix x assume neq: "x \ star_of a" assume approx: "x \ star_of a" have "starfun f x - star_of L \ Infinitesimal" proof (rule InfinitesimalI2) - fix r::real assume r: "0 < r" - from LIM_D [OF f r] - obtain s where s: "0 < s" and - less_r: "\x. \x \ a; norm (x - a) < s\ \ norm (f x - L) < r" + fix r :: real + assume r: "0 < r" + from LIM_D [OF f r] obtain s + where s: "0 < s" and less_r: "\x. x \ a \ norm (x - a) < s \ norm (f x - L) < r" by fast from less_r have less_r': - "\x. \x \ star_of a; hnorm (x - star_of a) < star_of s\ - \ hnorm (starfun f x - star_of L) < star_of r" + "\x. x \ star_of a \ hnorm (x - star_of a) < star_of s \ + hnorm (starfun f x - star_of L) < star_of r" by transfer from approx have "x - star_of a \ Infinitesimal" - by (unfold approx_def) - hence "hnorm (x - star_of a) < star_of s" + by (simp only: approx_def) + then have "hnorm (x - star_of a) < star_of s" using s by (rule InfinitesimalD2) with neq show "hnorm (starfun f x - star_of L) < star_of r" by (rule less_r') qed - thus "starfun f x \ star_of L" + then show "starfun f x \ star_of L" by (unfold approx_def) qed lemma NSLIM_LIM: - assumes f: "f \a\\<^sub>N\<^sub>S L" shows "f \a\ L" + assumes f: "f \a\\<^sub>N\<^sub>S L" + shows "f \a\ L" proof (rule LIM_I) - fix r::real assume r: "0 < r" - have "\s>0. \x. x \ star_of a \ hnorm (x - star_of a) < s - \ hnorm (starfun f x - star_of L) < star_of r" + fix r :: real + assume r: "0 < r" + have "\s>0. \x. x \ star_of a \ hnorm (x - star_of a) < s \ + hnorm (starfun f x - star_of L) < star_of r" proof (rule exI, safe) - show "0 < \" by (rule hypreal_epsilon_gt_zero) + show "0 < \" + by (rule hypreal_epsilon_gt_zero) next - fix x assume neq: "x \ star_of a" + fix x + assume neq: "x \ star_of a" assume "hnorm (x - star_of a) < \" - with Infinitesimal_epsilon - have "x - star_of a \ Infinitesimal" + with Infinitesimal_epsilon have "x - star_of a \ Infinitesimal" by (rule hnorm_less_Infinitesimal) - hence "x \ star_of a" + then have "x \ star_of a" by (unfold approx_def) with f neq have "starfun f x \ star_of L" by (rule NSLIM_D) - hence "starfun f x - star_of L \ Infinitesimal" + then have "starfun f x - star_of L \ Infinitesimal" by (unfold approx_def) - thus "hnorm (starfun f x - star_of L) < star_of r" + then show "hnorm (starfun f x - star_of L) < star_of r" using r by (rule InfinitesimalD2) qed - thus "\s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - L) < r" + then show "\s>0. \x. x \ a \ norm (x - a) < s \ norm (f x - L) < r" by transfer qed -theorem LIM_NSLIM_iff: "(f \x\ L) = (f \x\\<^sub>N\<^sub>S L)" -by (blast intro: LIM_NSLIM NSLIM_LIM) +theorem LIM_NSLIM_iff: "f \x\ L \ f \x\\<^sub>N\<^sub>S L" + by (blast intro: LIM_NSLIM NSLIM_LIM) subsection \Continuity\ -lemma isNSContD: - "\isNSCont f a; y \ star_of a\ \ ( *f* f) y \ star_of (f a)" -by (simp add: isNSCont_def) +lemma isNSContD: "isNSCont f a \ y \ star_of a \ ( *f* f) y \ star_of (f a)" + by (simp add: isNSCont_def) + +lemma isNSCont_NSLIM: "isNSCont f a \ f \a\\<^sub>N\<^sub>S (f a)" + by (simp add: isNSCont_def NSLIM_def) -lemma isNSCont_NSLIM: "isNSCont f a ==> f \a\\<^sub>N\<^sub>S (f a) " -by (simp add: isNSCont_def NSLIM_def) +lemma NSLIM_isNSCont: "f \a\\<^sub>N\<^sub>S (f a) \ isNSCont f a" + apply (auto simp add: isNSCont_def NSLIM_def) + apply (case_tac "y = star_of a") + apply auto + done -lemma NSLIM_isNSCont: "f \a\\<^sub>N\<^sub>S (f a) ==> isNSCont f a" -apply (simp add: isNSCont_def NSLIM_def, auto) -apply (case_tac "y = star_of a", auto) -done +text \NS continuity can be defined using NS Limit in + similar fashion to standard definition of continuity.\ +lemma isNSCont_NSLIM_iff: "isNSCont f a \ f \a\\<^sub>N\<^sub>S (f a)" + by (blast intro: isNSCont_NSLIM NSLIM_isNSCont) -text\NS continuity can be defined using NS Limit in - similar fashion to standard definition of continuity\ -lemma isNSCont_NSLIM_iff: "(isNSCont f a) = (f \a\\<^sub>N\<^sub>S (f a))" -by (blast intro: isNSCont_NSLIM NSLIM_isNSCont) +text \Hence, NS continuity can be given in terms of standard limit.\ +lemma isNSCont_LIM_iff: "(isNSCont f a) = (f \a\ (f a))" + by (simp add: LIM_NSLIM_iff isNSCont_NSLIM_iff) -text\Hence, NS continuity can be given - in terms of standard limit\ -lemma isNSCont_LIM_iff: "(isNSCont f a) = (f \a\ (f a))" -by (simp add: LIM_NSLIM_iff isNSCont_NSLIM_iff) +text \Moreover, it's trivial now that NS continuity + is equivalent to standard continuity.\ +lemma isNSCont_isCont_iff: "isNSCont f a \ isCont f a" + by (simp add: isCont_def) (rule isNSCont_LIM_iff) -text\Moreover, it's trivial now that NS continuity - is equivalent to standard continuity\ -lemma isNSCont_isCont_iff: "(isNSCont f a) = (isCont f a)" -apply (simp add: isCont_def) -apply (rule isNSCont_LIM_iff) -done +text \Standard continuity \\\ NS continuity.\ +lemma isCont_isNSCont: "isCont f a \ isNSCont f a" + by (erule isNSCont_isCont_iff [THEN iffD2]) -text\Standard continuity ==> NS continuity\ -lemma isCont_isNSCont: "isCont f a ==> isNSCont f a" -by (erule isNSCont_isCont_iff [THEN iffD2]) +text \NS continuity \==>\ Standard continuity.\ +lemma isNSCont_isCont: "isNSCont f a ==> isCont f a" + by (erule isNSCont_isCont_iff [THEN iffD1]) -text\NS continuity ==> Standard continuity\ -lemma isNSCont_isCont: "isNSCont f a ==> isCont f a" -by (erule isNSCont_isCont_iff [THEN iffD1]) + +text \Alternative definition of continuity.\ -text\Alternative definition of continuity\ +text \Prove equivalence between NS limits -- + seems easier than using standard definition.\ +lemma NSLIM_h_iff: "f \a\\<^sub>N\<^sub>S L \ (\h. f (a + h)) \0\\<^sub>N\<^sub>S L" + apply (simp add: NSLIM_def, auto) + apply (drule_tac x = "star_of a + x" in spec) + apply (drule_tac [2] x = "- star_of a + x" in spec, safe, simp) + apply (erule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]]) + apply (erule_tac [3] approx_minus_iff2 [THEN iffD1]) + prefer 2 apply (simp add: add.commute) + apply (rule_tac x = x in star_cases) + apply (rule_tac [2] x = x in star_cases) + apply (auto simp add: starfun star_of_def star_n_minus star_n_add add.assoc star_n_zero_num) + done -(* Prove equivalence between NS limits - *) -(* seems easier than using standard definition *) -lemma NSLIM_h_iff: "(f \a\\<^sub>N\<^sub>S L) = ((%h. f(a + h)) \0\\<^sub>N\<^sub>S L)" -apply (simp add: NSLIM_def, auto) -apply (drule_tac x = "star_of a + x" in spec) -apply (drule_tac [2] x = "- star_of a + x" in spec, safe, simp) -apply (erule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]]) -apply (erule_tac [3] approx_minus_iff2 [THEN iffD1]) - prefer 2 apply (simp add: add.commute) -apply (rule_tac x = x in star_cases) -apply (rule_tac [2] x = x in star_cases) -apply (auto simp add: starfun star_of_def star_n_minus star_n_add add.assoc star_n_zero_num) -done - -lemma NSLIM_isCont_iff: "(f \a\\<^sub>N\<^sub>S f a) = ((%h. f(a + h)) \0\\<^sub>N\<^sub>S f a)" +lemma NSLIM_isCont_iff: "f \a\\<^sub>N\<^sub>S f a \ (\h. f (a + h)) \0\\<^sub>N\<^sub>S f a" by (fact NSLIM_h_iff) -lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%x. - f x) a" -by (simp add: isNSCont_def) +lemma isNSCont_minus: "isNSCont f a \ isNSCont (\x. - f x) a" + by (simp add: isNSCont_def) -lemma isNSCont_inverse: - fixes f :: "'a::real_normed_vector \ 'b::real_normed_div_algebra" - shows "[| isNSCont f x; f x \ 0 |] ==> isNSCont (%x. inverse (f x)) x" -by (auto intro: isCont_inverse simp add: isNSCont_isCont_iff) +lemma isNSCont_inverse: "isNSCont f x \ f x \ 0 \ isNSCont (\x. inverse (f x)) x" + for f :: "'a::real_normed_vector \ 'b::real_normed_div_algebra" + by (auto intro: isCont_inverse simp add: isNSCont_isCont_iff) -lemma isNSCont_const [simp]: "isNSCont (%x. k) a" -by (simp add: isNSCont_def) +lemma isNSCont_const [simp]: "isNSCont (\x. k) a" + by (simp add: isNSCont_def) -lemma isNSCont_abs [simp]: "isNSCont abs (a::real)" -apply (simp add: isNSCont_def) -apply (auto intro: approx_hrabs simp add: starfun_rabs_hrabs) -done +lemma isNSCont_abs [simp]: "isNSCont abs a" + for a :: real + by (auto simp: isNSCont_def intro: approx_hrabs simp: starfun_rabs_hrabs) subsection \Uniform Continuity\ -lemma isNSUContD: "[| isNSUCont f; x \ y|] ==> ( *f* f) x \ ( *f* f) y" -by (simp add: isNSUCont_def) +lemma isNSUContD: "isNSUCont f \ x \ y \ ( *f* f) x \ ( *f* f) y" + by (simp add: isNSUCont_def) lemma isUCont_isNSUCont: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" - assumes f: "isUCont f" shows "isNSUCont f" -proof (unfold isNSUCont_def, safe) + assumes f: "isUCont f" + shows "isNSUCont f" + unfolding isNSUCont_def +proof safe fix x y :: "'a star" assume approx: "x \ y" have "starfun f x - starfun f y \ Infinitesimal" proof (rule InfinitesimalI2) - fix r::real assume r: "0 < r" - with f obtain s where s: "0 < s" and - less_r: "\x y. norm (x - y) < s \ norm (f x - f y) < r" + fix r :: real + assume r: "0 < r" + with f obtain s where s: "0 < s" + and less_r: "\x y. norm (x - y) < s \ norm (f x - f y) < r" by (auto simp add: isUCont_def dist_norm) from less_r have less_r': - "\x y. hnorm (x - y) < star_of s - \ hnorm (starfun f x - starfun f y) < star_of r" + "\x y. hnorm (x - y) < star_of s \ hnorm (starfun f x - starfun f y) < star_of r" by transfer from approx have "x - y \ Infinitesimal" by (unfold approx_def) - hence "hnorm (x - y) < star_of s" + then have "hnorm (x - y) < star_of s" using s by (rule InfinitesimalD2) - thus "hnorm (starfun f x - starfun f y) < star_of r" + then show "hnorm (starfun f x - starfun f y) < star_of r" by (rule less_r') qed - thus "starfun f x \ starfun f y" + then show "starfun f x \ starfun f y" by (unfold approx_def) qed lemma isNSUCont_isUCont: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" - assumes f: "isNSUCont f" shows "isUCont f" -proof (unfold isUCont_def dist_norm, safe) - fix r::real assume r: "0 < r" - have "\s>0. \x y. hnorm (x - y) < s - \ hnorm (starfun f x - starfun f y) < star_of r" + assumes f: "isNSUCont f" + shows "isUCont f" + unfolding isUCont_def dist_norm +proof safe + fix r :: real + assume r: "0 < r" + have "\s>0. \x y. hnorm (x - y) < s \ hnorm (starfun f x - starfun f y) < star_of r" proof (rule exI, safe) - show "0 < \" by (rule hypreal_epsilon_gt_zero) + show "0 < \" + by (rule hypreal_epsilon_gt_zero) next fix x y :: "'a star" assume "hnorm (x - y) < \" - with Infinitesimal_epsilon - have "x - y \ Infinitesimal" + with Infinitesimal_epsilon have "x - y \ Infinitesimal" by (rule hnorm_less_Infinitesimal) - hence "x \ y" + then have "x \ y" by (unfold approx_def) with f have "starfun f x \ starfun f y" by (simp add: isNSUCont_def) - hence "starfun f x - starfun f y \ Infinitesimal" + then have "starfun f x - starfun f y \ Infinitesimal" by (unfold approx_def) - thus "hnorm (starfun f x - starfun f y) < star_of r" + then show "hnorm (starfun f x - starfun f y) < star_of r" using r by (rule InfinitesimalD2) qed - thus "\s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r" + then show "\s>0. \x y. norm (x - y) < s \ norm (f x - f y) < r" by transfer qed diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/HyperDef.thy --- a/src/HOL/Nonstandard_Analysis/HyperDef.thy Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/HyperDef.thy Tue Nov 01 00:44:24 2016 +0100 @@ -4,50 +4,42 @@ Conversion to Isar and new proofs by Lawrence C Paulson, 2004 *) -section\Construction of Hyperreals Using Ultrafilters\ +section \Construction of Hyperreals Using Ultrafilters\ theory HyperDef -imports Complex_Main HyperNat + imports Complex_Main HyperNat begin type_synonym hypreal = "real star" -abbreviation - hypreal_of_real :: "real => real star" where - "hypreal_of_real == star_of" +abbreviation hypreal_of_real :: "real \ real star" + where "hypreal_of_real \ star_of" -abbreviation - hypreal_of_hypnat :: "hypnat \ hypreal" where - "hypreal_of_hypnat \ of_hypnat" +abbreviation hypreal_of_hypnat :: "hypnat \ hypreal" + where "hypreal_of_hypnat \ of_hypnat" -definition - omega :: hypreal ("\") where - \ \an infinite number \= [<1,2,3,...>]\\ - "\ = star_n (\n. real (Suc n))" +definition omega :: hypreal ("\") + where "\ = star_n (\n. real (Suc n))" + \ \an infinite number \= [<1, 2, 3, \>]\\ -definition - epsilon :: hypreal ("\") where - \ \an infinitesimal number \= [<1,1/2,1/3,...>]\\ - "\ = star_n (\n. inverse (real (Suc n)))" +definition epsilon :: hypreal ("\") + where "\ = star_n (\n. inverse (real (Suc n)))" + \ \an infinitesimal number \= [<1, 1/2, 1/3, \>]\\ subsection \Real vector class instances\ instantiation star :: (scaleR) scaleR begin - -definition - star_scaleR_def [transfer_unfold]: "scaleR r \ *f* (scaleR r)" - -instance .. - + definition star_scaleR_def [transfer_unfold]: "scaleR r \ *f* (scaleR r)" + instance .. end lemma Standard_scaleR [simp]: "x \ Standard \ scaleR r x \ Standard" -by (simp add: star_scaleR_def) + by (simp add: star_scaleR_def) lemma star_of_scaleR [simp]: "star_of (scaleR r x) = scaleR r (star_of x)" -by transfer (rule refl) + by transfer (rule refl) instance star :: (real_vector) real_vector proof @@ -80,230 +72,213 @@ instance star :: (real_field) real_field .. lemma star_of_real_def [transfer_unfold]: "of_real r = star_of (of_real r)" -by (unfold of_real_def, transfer, rule refl) + by (unfold of_real_def, transfer, rule refl) lemma Standard_of_real [simp]: "of_real r \ Standard" -by (simp add: star_of_real_def) + by (simp add: star_of_real_def) lemma star_of_of_real [simp]: "star_of (of_real r) = of_real r" -by transfer (rule refl) + by transfer (rule refl) lemma of_real_eq_star_of [simp]: "of_real = star_of" proof - fix r :: real - show "of_real r = star_of r" + show "of_real r = star_of r" for r :: real by transfer simp qed lemma Reals_eq_Standard: "(\ :: hypreal set) = Standard" -by (simp add: Reals_def Standard_def) + by (simp add: Reals_def Standard_def) subsection \Injection from @{typ hypreal}\ -definition - of_hypreal :: "hypreal \ 'a::real_algebra_1 star" where - [transfer_unfold]: "of_hypreal = *f* of_real" +definition of_hypreal :: "hypreal \ 'a::real_algebra_1 star" + where [transfer_unfold]: "of_hypreal = *f* of_real" -lemma Standard_of_hypreal [simp]: - "r \ Standard \ of_hypreal r \ Standard" -by (simp add: of_hypreal_def) +lemma Standard_of_hypreal [simp]: "r \ Standard \ of_hypreal r \ Standard" + by (simp add: of_hypreal_def) lemma of_hypreal_0 [simp]: "of_hypreal 0 = 0" -by transfer (rule of_real_0) + by transfer (rule of_real_0) lemma of_hypreal_1 [simp]: "of_hypreal 1 = 1" -by transfer (rule of_real_1) + by transfer (rule of_real_1) -lemma of_hypreal_add [simp]: - "\x y. of_hypreal (x + y) = of_hypreal x + of_hypreal y" -by transfer (rule of_real_add) +lemma of_hypreal_add [simp]: "\x y. of_hypreal (x + y) = of_hypreal x + of_hypreal y" + by transfer (rule of_real_add) lemma of_hypreal_minus [simp]: "\x. of_hypreal (- x) = - of_hypreal x" -by transfer (rule of_real_minus) + by transfer (rule of_real_minus) -lemma of_hypreal_diff [simp]: - "\x y. of_hypreal (x - y) = of_hypreal x - of_hypreal y" -by transfer (rule of_real_diff) +lemma of_hypreal_diff [simp]: "\x y. of_hypreal (x - y) = of_hypreal x - of_hypreal y" + by transfer (rule of_real_diff) -lemma of_hypreal_mult [simp]: - "\x y. of_hypreal (x * y) = of_hypreal x * of_hypreal y" -by transfer (rule of_real_mult) +lemma of_hypreal_mult [simp]: "\x y. of_hypreal (x * y) = of_hypreal x * of_hypreal y" + by transfer (rule of_real_mult) lemma of_hypreal_inverse [simp]: "\x. of_hypreal (inverse x) = - inverse (of_hypreal x :: 'a::{real_div_algebra, division_ring} star)" -by transfer (rule of_real_inverse) + inverse (of_hypreal x :: 'a::{real_div_algebra, division_ring} star)" + by transfer (rule of_real_inverse) lemma of_hypreal_divide [simp]: "\x y. of_hypreal (x / y) = - (of_hypreal x / of_hypreal y :: 'a::{real_field, field} star)" -by transfer (rule of_real_divide) + (of_hypreal x / of_hypreal y :: 'a::{real_field, field} star)" + by transfer (rule of_real_divide) -lemma of_hypreal_eq_iff [simp]: - "\x y. (of_hypreal x = of_hypreal y) = (x = y)" -by transfer (rule of_real_eq_iff) +lemma of_hypreal_eq_iff [simp]: "\x y. (of_hypreal x = of_hypreal y) = (x = y)" + by transfer (rule of_real_eq_iff) -lemma of_hypreal_eq_0_iff [simp]: - "\x. (of_hypreal x = 0) = (x = 0)" -by transfer (rule of_real_eq_0_iff) +lemma of_hypreal_eq_0_iff [simp]: "\x. (of_hypreal x = 0) = (x = 0)" + by transfer (rule of_real_eq_0_iff) -subsection\Properties of @{term starrel}\ +subsection \Properties of @{term starrel}\ lemma lemma_starrel_refl [simp]: "x \ starrel `` {x}" -by (simp add: starrel_def) + by (simp add: starrel_def) lemma starrel_in_hypreal [simp]: "starrel``{x}:star" -by (simp add: star_def starrel_def quotient_def, blast) + by (simp add: star_def starrel_def quotient_def, blast) declare Abs_star_inject [simp] Abs_star_inverse [simp] declare equiv_starrel [THEN eq_equiv_class_iff, simp] -subsection\@{term hypreal_of_real}: - the Injection from @{typ real} to @{typ hypreal}\ + +subsection \@{term hypreal_of_real}: the Injection from @{typ real} to @{typ hypreal}\ lemma inj_star_of: "inj star_of" -by (rule inj_onI, simp) + by (rule inj_onI) simp -lemma mem_Rep_star_iff: "(X \ Rep_star x) = (x = star_n X)" -by (cases x, simp add: star_n_def) +lemma mem_Rep_star_iff: "X \ Rep_star x \ x = star_n X" + by (cases x) (simp add: star_n_def) -lemma Rep_star_star_n_iff [simp]: - "(X \ Rep_star (star_n Y)) = (eventually (\n. Y n = X n) \)" -by (simp add: star_n_def) +lemma Rep_star_star_n_iff [simp]: "X \ Rep_star (star_n Y) \ eventually (\n. Y n = X n) \" + by (simp add: star_n_def) lemma Rep_star_star_n: "X \ Rep_star (star_n X)" -by simp + by simp -subsection\Properties of @{term star_n}\ -lemma star_n_add: - "star_n X + star_n Y = star_n (%n. X n + Y n)" -by (simp only: star_add_def starfun2_star_n) +subsection \Properties of @{term star_n}\ + +lemma star_n_add: "star_n X + star_n Y = star_n (\n. X n + Y n)" + by (simp only: star_add_def starfun2_star_n) -lemma star_n_minus: - "- star_n X = star_n (%n. -(X n))" -by (simp only: star_minus_def starfun_star_n) +lemma star_n_minus: "- star_n X = star_n (\n. -(X n))" + by (simp only: star_minus_def starfun_star_n) -lemma star_n_diff: - "star_n X - star_n Y = star_n (%n. X n - Y n)" -by (simp only: star_diff_def starfun2_star_n) +lemma star_n_diff: "star_n X - star_n Y = star_n (\n. X n - Y n)" + by (simp only: star_diff_def starfun2_star_n) -lemma star_n_mult: - "star_n X * star_n Y = star_n (%n. X n * Y n)" -by (simp only: star_mult_def starfun2_star_n) +lemma star_n_mult: "star_n X * star_n Y = star_n (\n. X n * Y n)" + by (simp only: star_mult_def starfun2_star_n) -lemma star_n_inverse: - "inverse (star_n X) = star_n (%n. inverse(X n))" -by (simp only: star_inverse_def starfun_star_n) +lemma star_n_inverse: "inverse (star_n X) = star_n (\n. inverse (X n))" + by (simp only: star_inverse_def starfun_star_n) -lemma star_n_le: - "star_n X \ star_n Y = (eventually (\n. X n \ Y n) FreeUltrafilterNat)" -by (simp only: star_le_def starP2_star_n) +lemma star_n_le: "star_n X \ star_n Y = eventually (\n. X n \ Y n) FreeUltrafilterNat" + by (simp only: star_le_def starP2_star_n) + +lemma star_n_less: "star_n X < star_n Y = eventually (\n. X n < Y n) FreeUltrafilterNat" + by (simp only: star_less_def starP2_star_n) -lemma star_n_less: - "star_n X < star_n Y = (eventually (\n. X n < Y n) FreeUltrafilterNat)" -by (simp only: star_less_def starP2_star_n) +lemma star_n_zero_num: "0 = star_n (\n. 0)" + by (simp only: star_zero_def star_of_def) -lemma star_n_zero_num: "0 = star_n (%n. 0)" -by (simp only: star_zero_def star_of_def) +lemma star_n_one_num: "1 = star_n (\n. 1)" + by (simp only: star_one_def star_of_def) -lemma star_n_one_num: "1 = star_n (%n. 1)" -by (simp only: star_one_def star_of_def) - -lemma star_n_abs: "\star_n X\ = star_n (%n. \X n\)" -by (simp only: star_abs_def starfun_star_n) +lemma star_n_abs: "\star_n X\ = star_n (\n. \X n\)" + by (simp only: star_abs_def starfun_star_n) lemma hypreal_omega_gt_zero [simp]: "0 < \" -by (simp add: omega_def star_n_zero_num star_n_less) - -subsection\Existence of Infinite Hyperreal Number\ - -text\Existence of infinite number not corresponding to any real number. -Use assumption that member @{term FreeUltrafilterNat} is not finite.\ + by (simp add: omega_def star_n_zero_num star_n_less) -text\A few lemmas first\ +subsection \Existence of Infinite Hyperreal Number\ + +text \Existence of infinite number not corresponding to any real number. + Use assumption that member @{term FreeUltrafilterNat} is not finite.\ + +text \A few lemmas first.\ lemma lemma_omega_empty_singleton_disj: "{n::nat. x = real n} = {} \ (\y. {n::nat. x = real n} = {y})" -by force + by force lemma lemma_finite_omega_set: "finite {n::nat. x = real n}" using lemma_omega_empty_singleton_disj [of x] by auto -lemma not_ex_hypreal_of_real_eq_omega: - "~ (\x. hypreal_of_real x = \)" -apply (simp add: omega_def star_of_def star_n_eq_iff) -apply clarify -apply (rule_tac x2="x-1" in lemma_finite_omega_set [THEN FreeUltrafilterNat.finite, THEN notE]) -apply (erule eventually_mono) -apply auto -done +lemma not_ex_hypreal_of_real_eq_omega: "\x. hypreal_of_real x = \" + apply (simp add: omega_def star_of_def star_n_eq_iff) + apply clarify + apply (rule_tac x2="x-1" in lemma_finite_omega_set [THEN FreeUltrafilterNat.finite, THEN notE]) + apply (erule eventually_mono) + apply auto + done lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \ \" -by (insert not_ex_hypreal_of_real_eq_omega, auto) + using not_ex_hypreal_of_real_eq_omega by auto -text\Existence of infinitesimal number also not corresponding to any - real number\ +text \Existence of infinitesimal number also not corresponding to any real number.\ lemma lemma_epsilon_empty_singleton_disj: - "{n::nat. x = inverse(real(Suc n))} = {} | - (\y. {n::nat. x = inverse(real(Suc n))} = {y})" -by auto + "{n::nat. x = inverse(real(Suc n))} = {} \ (\y. {n::nat. x = inverse(real(Suc n))} = {y})" + by auto -lemma lemma_finite_epsilon_set: "finite {n. x = inverse(real(Suc n))}" -by (cut_tac x = x in lemma_epsilon_empty_singleton_disj, auto) +lemma lemma_finite_epsilon_set: "finite {n. x = inverse (real (Suc n))}" + using lemma_epsilon_empty_singleton_disj [of x] by auto -lemma not_ex_hypreal_of_real_eq_epsilon: "~ (\x. hypreal_of_real x = \)" -by (auto simp add: epsilon_def star_of_def star_n_eq_iff - lemma_finite_epsilon_set [THEN FreeUltrafilterNat.finite] simp del: of_nat_Suc) +lemma not_ex_hypreal_of_real_eq_epsilon: "\x. hypreal_of_real x = \" + by (auto simp: epsilon_def star_of_def star_n_eq_iff + lemma_finite_epsilon_set [THEN FreeUltrafilterNat.finite] simp del: of_nat_Suc) lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \ \" -by (insert not_ex_hypreal_of_real_eq_epsilon, auto) + using not_ex_hypreal_of_real_eq_epsilon by auto lemma hypreal_epsilon_not_zero: "\ \ 0" -by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff FreeUltrafilterNat.proper - del: star_of_zero) + by (simp add: epsilon_def star_zero_def star_of_def star_n_eq_iff FreeUltrafilterNat.proper + del: star_of_zero) lemma hypreal_epsilon_inverse_omega: "\ = inverse \" -by (simp add: epsilon_def omega_def star_n_inverse) + by (simp add: epsilon_def omega_def star_n_inverse) lemma hypreal_epsilon_gt_zero: "0 < \" -by (simp add: hypreal_epsilon_inverse_omega) - -subsection\Absolute Value Function for the Hyperreals\ - -lemma hrabs_add_less: "[| \x\ < r; \y\ < s |] ==> \x + y\ < r + (s::hypreal)" -by (simp add: abs_if split: if_split_asm) - -lemma hrabs_less_gt_zero: "\x\ < r ==> (0::hypreal) < r" -by (blast intro!: order_le_less_trans abs_ge_zero) - -lemma hrabs_disj: "\x\ = (x::'a::abs_if) \ \x\ = -x" -by (simp add: abs_if) - -lemma hrabs_add_lemma_disj: "(y::hypreal) + - x + (y + - z) = \x + - z\ ==> y = z | x = y" -by (simp add: abs_if split: if_split_asm) + by (simp add: hypreal_epsilon_inverse_omega) -subsection\Embedding the Naturals into the Hyperreals\ +subsection \Absolute Value Function for the Hyperreals\ + +lemma hrabs_add_less: "\x\ < r \ \y\ < s \ \x + y\ < r + s" + for x y r s :: hypreal + by (simp add: abs_if split: if_split_asm) + +lemma hrabs_less_gt_zero: "\x\ < r \ 0 < r" + for x r :: hypreal + by (blast intro!: order_le_less_trans abs_ge_zero) -abbreviation - hypreal_of_nat :: "nat => hypreal" where - "hypreal_of_nat == of_nat" +lemma hrabs_disj: "\x\ = x \ \x\ = -x" + for x :: "'a::abs_if" + by (simp add: abs_if) + +lemma hrabs_add_lemma_disj: "y + - x + (y + - z) = \x + - z\ \ y = z \ x = y" + for x y z :: hypreal + by (simp add: abs_if split: if_split_asm) + + +subsection \Embedding the Naturals into the Hyperreals\ + +abbreviation hypreal_of_nat :: "nat \ hypreal" + where "hypreal_of_nat \ of_nat" lemma SNat_eq: "Nats = {n. \N. n = hypreal_of_nat N}" -by (simp add: Nats_def image_def) + by (simp add: Nats_def image_def) -(*------------------------------------------------------------*) -(* naturals embedded in hyperreals *) -(* is a hyperreal c.f. NS extension *) -(*------------------------------------------------------------*) +text \Naturals embedded in hyperreals: is a hyperreal c.f. NS extension.\ -lemma hypreal_of_nat: "hypreal_of_nat m = star_n (%n. real m)" -by (simp add: star_of_def [symmetric]) +lemma hypreal_of_nat: "hypreal_of_nat m = star_n (\n. real m)" + by (simp add: star_of_def [symmetric]) declaration \ K (Lin_Arith.add_inj_thms [@{thm star_of_le} RS iffD2, @@ -314,75 +289,77 @@ #> Lin_Arith.add_inj_const (@{const_name "StarDef.star_of"}, @{typ "real \ hypreal"})) \ -simproc_setup fast_arith_hypreal ("(m::hypreal) < n" | "(m::hypreal) <= n" | "(m::hypreal) = n") = +simproc_setup fast_arith_hypreal ("(m::hypreal) < n" | "(m::hypreal) \ n" | "(m::hypreal) = n") = \K Lin_Arith.simproc\ subsection \Exponentials on the Hyperreals\ -lemma hpowr_0 [simp]: "r ^ 0 = (1::hypreal)" -by (rule power_0) - -lemma hpowr_Suc [simp]: "r ^ (Suc n) = (r::hypreal) * (r ^ n)" -by (rule power_Suc) +lemma hpowr_0 [simp]: "r ^ 0 = (1::hypreal)" + for r :: hypreal + by (rule power_0) -lemma hrealpow_two: "(r::hypreal) ^ Suc (Suc 0) = r * r" -by simp +lemma hpowr_Suc [simp]: "r ^ (Suc n) = r * (r ^ n)" + for r :: hypreal + by (rule power_Suc) -lemma hrealpow_two_le [simp]: "(0::hypreal) \ r ^ Suc (Suc 0)" -by (auto simp add: zero_le_mult_iff) +lemma hrealpow_two: "r ^ Suc (Suc 0) = r * r" + for r :: hypreal + by simp -lemma hrealpow_two_le_add_order [simp]: - "(0::hypreal) \ u ^ Suc (Suc 0) + v ^ Suc (Suc 0)" -by (simp only: hrealpow_two_le add_nonneg_nonneg) +lemma hrealpow_two_le [simp]: "0 \ r ^ Suc (Suc 0)" + for r :: hypreal + by (auto simp add: zero_le_mult_iff) + +lemma hrealpow_two_le_add_order [simp]: "0 \ u ^ Suc (Suc 0) + v ^ Suc (Suc 0)" + for u v :: hypreal + by (simp only: hrealpow_two_le add_nonneg_nonneg) -lemma hrealpow_two_le_add_order2 [simp]: - "(0::hypreal) \ u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)" -by (simp only: hrealpow_two_le add_nonneg_nonneg) +lemma hrealpow_two_le_add_order2 [simp]: "0 \ u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)" + for u v w :: hypreal + by (simp only: hrealpow_two_le add_nonneg_nonneg) -lemma hypreal_add_nonneg_eq_0_iff: - "[| 0 \ x; 0 \ y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))" -by arith +lemma hypreal_add_nonneg_eq_0_iff: "0 \ x \ 0 \ y \ x + y = 0 \ x = 0 \ y = 0" + for x y :: hypreal + by arith -text\FIXME: DELETE THESE\ -lemma hypreal_three_squares_add_zero_iff: - "(x*x + y*y + z*z = 0) = (x = 0 & y = 0 & z = (0::hypreal))" -apply (simp only: zero_le_square add_nonneg_nonneg hypreal_add_nonneg_eq_0_iff, auto) -done +(* FIXME: DELETE THESE *) +lemma hypreal_three_squares_add_zero_iff: "x * x + y * y + z * z = 0 \ x = 0 \ y = 0 \ z = 0" + for x y z :: hypreal + by (simp only: zero_le_square add_nonneg_nonneg hypreal_add_nonneg_eq_0_iff) auto lemma hrealpow_three_squares_add_zero_iff [simp]: - "(x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + z ^ Suc (Suc 0) = (0::hypreal)) = - (x = 0 & y = 0 & z = 0)" -by (simp only: hypreal_three_squares_add_zero_iff hrealpow_two) + "x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + z ^ Suc (Suc 0) = 0 \ x = 0 \ y = 0 \ z = 0" + for x y z :: hypreal + by (simp only: hypreal_three_squares_add_zero_iff hrealpow_two) (*FIXME: This and RealPow.abs_realpow_two should be replaced by an abstract result proved in Rings or Fields*) -lemma hrabs_hrealpow_two [simp]: "\x ^ Suc (Suc 0)\ = (x::hypreal) ^ Suc (Suc 0)" -by (simp add: abs_mult) +lemma hrabs_hrealpow_two [simp]: "\x ^ Suc (Suc 0)\ = x ^ Suc (Suc 0)" + for x :: hypreal + by (simp add: abs_mult) lemma two_hrealpow_ge_one [simp]: "(1::hypreal) \ 2 ^ n" -by (insert power_increasing [of 0 n "2::hypreal"], simp) + using power_increasing [of 0 n "2::hypreal"] by simp -lemma hrealpow: - "star_n X ^ m = star_n (%n. (X n::real) ^ m)" -apply (induct_tac "m") -apply (auto simp add: star_n_one_num star_n_mult power_0) -done +lemma hrealpow: "star_n X ^ m = star_n (\n. (X n::real) ^ m)" + by (induct m) (auto simp: star_n_one_num star_n_mult) lemma hrealpow_sum_square_expand: - "(x + (y::hypreal)) ^ Suc (Suc 0) = - x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + (hypreal_of_nat (Suc (Suc 0)))*x*y" -by (simp add: distrib_left distrib_right) + "(x + y) ^ Suc (Suc 0) = + x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + (hypreal_of_nat (Suc (Suc 0))) * x * y" + for x y :: hypreal + by (simp add: distrib_left distrib_right) lemma power_hypreal_of_real_numeral: - "(numeral v :: hypreal) ^ n = hypreal_of_real ((numeral v) ^ n)" -by simp + "(numeral v :: hypreal) ^ n = hypreal_of_real ((numeral v) ^ n)" + by simp declare power_hypreal_of_real_numeral [of _ "numeral w", simp] for w lemma power_hypreal_of_real_neg_numeral: - "(- numeral v :: hypreal) ^ n = hypreal_of_real ((- numeral v) ^ n)" -by simp + "(- numeral v :: hypreal) ^ n = hypreal_of_real ((- numeral v) ^ n)" + by simp declare power_hypreal_of_real_neg_numeral [of _ "numeral w", simp] for w (* lemma hrealpow_HFinite: @@ -393,148 +370,119 @@ done *) -subsection\Powers with Hypernatural Exponents\ -definition pow :: "['a::power star, nat star] \ 'a star" (infixr "pow" 80) where - hyperpow_def [transfer_unfold]: "R pow N = ( *f2* op ^) R N" - (* hypernatural powers of hyperreals *) +subsection \Powers with Hypernatural Exponents\ -lemma Standard_hyperpow [simp]: - "\r \ Standard; n \ Standard\ \ r pow n \ Standard" -unfolding hyperpow_def by simp +text \Hypernatural powers of hyperreals.\ +definition pow :: "'a::power star \ nat star \ 'a star" (infixr "pow" 80) + where hyperpow_def [transfer_unfold]: "R pow N = ( *f2* op ^) R N" -lemma hyperpow: "star_n X pow star_n Y = star_n (%n. X n ^ Y n)" -by (simp add: hyperpow_def starfun2_star_n) - -lemma hyperpow_zero [simp]: - "\n. (0::'a::{power,semiring_0} star) pow (n + (1::hypnat)) = 0" -by transfer simp +lemma Standard_hyperpow [simp]: "r \ Standard \ n \ Standard \ r pow n \ Standard" + by (simp add: hyperpow_def) -lemma hyperpow_not_zero: - "\r n. r \ (0::'a::{field} star) ==> r pow n \ 0" -by transfer (rule power_not_zero) +lemma hyperpow: "star_n X pow star_n Y = star_n (\n. X n ^ Y n)" + by (simp add: hyperpow_def starfun2_star_n) + +lemma hyperpow_zero [simp]: "\n. (0::'a::{power,semiring_0} star) pow (n + (1::hypnat)) = 0" + by transfer simp -lemma hyperpow_inverse: - "\r n. r \ (0::'a::field star) - \ inverse (r pow n) = (inverse r) pow n" -by transfer (rule power_inverse [symmetric]) +lemma hyperpow_not_zero: "\r n. r \ (0::'a::{field} star) \ r pow n \ 0" + by transfer (rule power_not_zero) -lemma hyperpow_hrabs: - "\r n. \r::'a::{linordered_idom} star\ pow n = \r pow n\" -by transfer (rule power_abs [symmetric]) +lemma hyperpow_inverse: "\r n. r \ (0::'a::field star) \ inverse (r pow n) = (inverse r) pow n" + by transfer (rule power_inverse [symmetric]) -lemma hyperpow_add: - "\r n m. (r::'a::monoid_mult star) pow (n + m) = (r pow n) * (r pow m)" -by transfer (rule power_add) +lemma hyperpow_hrabs: "\r n. \r::'a::{linordered_idom} star\ pow n = \r pow n\" + by transfer (rule power_abs [symmetric]) -lemma hyperpow_one [simp]: - "\r. (r::'a::monoid_mult star) pow (1::hypnat) = r" -by transfer (rule power_one_right) +lemma hyperpow_add: "\r n m. (r::'a::monoid_mult star) pow (n + m) = (r pow n) * (r pow m)" + by transfer (rule power_add) -lemma hyperpow_two: - "\r. (r::'a::monoid_mult star) pow (2::hypnat) = r * r" -by transfer (rule power2_eq_square) +lemma hyperpow_one [simp]: "\r. (r::'a::monoid_mult star) pow (1::hypnat) = r" + by transfer (rule power_one_right) -lemma hyperpow_gt_zero: - "\r n. (0::'a::{linordered_semidom} star) < r \ 0 < r pow n" -by transfer (rule zero_less_power) +lemma hyperpow_two: "\r. (r::'a::monoid_mult star) pow (2::hypnat) = r * r" + by transfer (rule power2_eq_square) -lemma hyperpow_ge_zero: - "\r n. (0::'a::{linordered_semidom} star) \ r \ 0 \ r pow n" -by transfer (rule zero_le_power) +lemma hyperpow_gt_zero: "\r n. (0::'a::{linordered_semidom} star) < r \ 0 < r pow n" + by transfer (rule zero_less_power) + +lemma hyperpow_ge_zero: "\r n. (0::'a::{linordered_semidom} star) \ r \ 0 \ r pow n" + by transfer (rule zero_le_power) -lemma hyperpow_le: - "\x y n. \(0::'a::{linordered_semidom} star) < x; x \ y\ - \ x pow n \ y pow n" -by transfer (rule power_mono [OF _ order_less_imp_le]) +lemma hyperpow_le: "\x y n. (0::'a::{linordered_semidom} star) < x \ x \ y \ x pow n \ y pow n" + by transfer (rule power_mono [OF _ order_less_imp_le]) -lemma hyperpow_eq_one [simp]: - "\n. 1 pow n = (1::'a::monoid_mult star)" -by transfer (rule power_one) +lemma hyperpow_eq_one [simp]: "\n. 1 pow n = (1::'a::monoid_mult star)" + by transfer (rule power_one) -lemma hrabs_hyperpow_minus [simp]: - "\(a::'a::{linordered_idom} star) n. \(-a) pow n\ = \a pow n\" -by transfer (rule abs_power_minus) +lemma hrabs_hyperpow_minus [simp]: "\(a::'a::linordered_idom star) n. \(-a) pow n\ = \a pow n\" + by transfer (rule abs_power_minus) -lemma hyperpow_mult: - "\r s n. (r * s::'a::{comm_monoid_mult} star) pow n - = (r pow n) * (s pow n)" -by transfer (rule power_mult_distrib) +lemma hyperpow_mult: "\r s n. (r * s::'a::comm_monoid_mult star) pow n = (r pow n) * (s pow n)" + by transfer (rule power_mult_distrib) -lemma hyperpow_two_le [simp]: - "\r. (0::'a::{monoid_mult,linordered_ring_strict} star) \ r pow 2" -by (auto simp add: hyperpow_two zero_le_mult_iff) +lemma hyperpow_two_le [simp]: "\r. (0::'a::{monoid_mult,linordered_ring_strict} star) \ r pow 2" + by (auto simp add: hyperpow_two zero_le_mult_iff) lemma hrabs_hyperpow_two [simp]: - "\x pow 2\ = - (x::'a::{monoid_mult,linordered_ring_strict} star) pow 2" -by (simp only: abs_of_nonneg hyperpow_two_le) + "\x pow 2\ = (x::'a::{monoid_mult,linordered_ring_strict} star) pow 2" + by (simp only: abs_of_nonneg hyperpow_two_le) -lemma hyperpow_two_hrabs [simp]: - "\x::'a::{linordered_idom} star\ pow 2 = x pow 2" -by (simp add: hyperpow_hrabs) +lemma hyperpow_two_hrabs [simp]: "\x::'a::linordered_idom star\ pow 2 = x pow 2" + by (simp add: hyperpow_hrabs) -text\The precondition could be weakened to @{term "0\x"}\ -lemma hypreal_mult_less_mono: - "[| u u*x < v* y" - by (simp add: mult_strict_mono order_less_imp_le) +text \The precondition could be weakened to @{term "0\x"}.\ +lemma hypreal_mult_less_mono: "u < v \ x < y \ 0 < v \ 0 < x \ u * x < v * y" + for u v x y :: hypreal + by (simp add: mult_strict_mono order_less_imp_le) -lemma hyperpow_two_gt_one: - "\r::'a::{linordered_semidom} star. 1 < r \ 1 < r pow 2" -by transfer simp +lemma hyperpow_two_gt_one: "\r::'a::linordered_semidom star. 1 < r \ 1 < r pow 2" + by transfer simp -lemma hyperpow_two_ge_one: - "\r::'a::{linordered_semidom} star. 1 \ r \ 1 \ r pow 2" -by transfer (rule one_le_power) +lemma hyperpow_two_ge_one: "\r::'a::linordered_semidom star. 1 \ r \ 1 \ r pow 2" + by transfer (rule one_le_power) lemma two_hyperpow_ge_one [simp]: "(1::hypreal) \ 2 pow n" -apply (rule_tac y = "1 pow n" in order_trans) -apply (rule_tac [2] hyperpow_le, auto) -done + apply (rule_tac y = "1 pow n" in order_trans) + apply (rule_tac [2] hyperpow_le) + apply auto + done -lemma hyperpow_minus_one2 [simp]: - "\n. (- 1) pow (2*n) = (1::hypreal)" -by transfer (rule power_minus1_even) +lemma hyperpow_minus_one2 [simp]: "\n. (- 1) pow (2 * n) = (1::hypreal)" + by transfer (rule power_minus1_even) -lemma hyperpow_less_le: - "!!r n N. [|(0::hypreal) \ r; r \ 1; n < N|] ==> r pow N \ r pow n" -by transfer (rule power_decreasing [OF order_less_imp_le]) +lemma hyperpow_less_le: "\r n N. (0::hypreal) \ r \ r \ 1 \ n < N \ r pow N \ r pow n" + by transfer (rule power_decreasing [OF order_less_imp_le]) lemma hyperpow_SHNat_le: - "[| 0 \ r; r \ (1::hypreal); N \ HNatInfinite |] - ==> ALL n: Nats. r pow N \ r pow n" -by (auto intro!: hyperpow_less_le simp add: HNatInfinite_iff) + "0 \ r \ r \ (1::hypreal) \ N \ HNatInfinite \ \n\Nats. r pow N \ r pow n" + by (auto intro!: hyperpow_less_le simp: HNatInfinite_iff) -lemma hyperpow_realpow: - "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)" -by transfer (rule refl) +lemma hyperpow_realpow: "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)" + by transfer (rule refl) -lemma hyperpow_SReal [simp]: - "(hypreal_of_real r) pow (hypnat_of_nat n) \ \" -by (simp add: Reals_eq_Standard) +lemma hyperpow_SReal [simp]: "(hypreal_of_real r) pow (hypnat_of_nat n) \ \" + by (simp add: Reals_eq_Standard) -lemma hyperpow_zero_HNatInfinite [simp]: - "N \ HNatInfinite ==> (0::hypreal) pow N = 0" -by (drule HNatInfinite_is_Suc, auto) +lemma hyperpow_zero_HNatInfinite [simp]: "N \ HNatInfinite \ (0::hypreal) pow N = 0" + by (drule HNatInfinite_is_Suc, auto) -lemma hyperpow_le_le: - "[| (0::hypreal) \ r; r \ 1; n \ N |] ==> r pow N \ r pow n" -apply (drule order_le_less [of n, THEN iffD1]) -apply (auto intro: hyperpow_less_le) -done +lemma hyperpow_le_le: "(0::hypreal) \ r \ r \ 1 \ n \ N \ r pow N \ r pow n" + apply (drule order_le_less [of n, THEN iffD1]) + apply (auto intro: hyperpow_less_le) + done -lemma hyperpow_Suc_le_self2: - "[| (0::hypreal) \ r; r < 1 |] ==> r pow (n + (1::hypnat)) \ r" -apply (drule_tac n = " (1::hypnat) " in hyperpow_le_le) -apply auto -done +lemma hyperpow_Suc_le_self2: "(0::hypreal) \ r \ r < 1 \ r pow (n + (1::hypnat)) \ r" + apply (drule_tac n = " (1::hypnat) " in hyperpow_le_le) + apply auto + done lemma hyperpow_hypnat_of_nat: "\x. x pow hypnat_of_nat n = x ^ n" -by transfer (rule refl) + by transfer (rule refl) lemma of_hypreal_hyperpow: - "\x n. of_hypreal (x pow n) = - (of_hypreal x::'a::{real_algebra_1} star) pow n" -by transfer (rule of_real_power) + "\x n. of_hypreal (x pow n) = (of_hypreal x::'a::{real_algebra_1} star) pow n" + by transfer (rule of_real_power) end diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/HyperNat.thy --- a/src/HOL/Nonstandard_Analysis/HyperNat.thy Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/HyperNat.thy Tue Nov 01 00:44:24 2016 +0100 @@ -5,410 +5,395 @@ Converted to Isar and polished by lcp *) -section\Hypernatural numbers\ +section \Hypernatural numbers\ theory HyperNat -imports StarDef + imports StarDef begin type_synonym hypnat = "nat star" -abbreviation - hypnat_of_nat :: "nat => nat star" where - "hypnat_of_nat == star_of" +abbreviation hypnat_of_nat :: "nat \ nat star" + where "hypnat_of_nat \ star_of" -definition - hSuc :: "hypnat => hypnat" where - hSuc_def [transfer_unfold]: "hSuc = *f* Suc" +definition hSuc :: "hypnat \ hypnat" + where hSuc_def [transfer_unfold]: "hSuc = *f* Suc" -subsection\Properties Transferred from Naturals\ + +subsection \Properties Transferred from Naturals\ lemma hSuc_not_zero [iff]: "\m. hSuc m \ 0" -by transfer (rule Suc_not_Zero) + by transfer (rule Suc_not_Zero) lemma zero_not_hSuc [iff]: "\m. 0 \ hSuc m" -by transfer (rule Zero_not_Suc) + by transfer (rule Zero_not_Suc) -lemma hSuc_hSuc_eq [iff]: "\m n. (hSuc m = hSuc n) = (m = n)" -by transfer (rule nat.inject) +lemma hSuc_hSuc_eq [iff]: "\m n. hSuc m = hSuc n \ m = n" + by transfer (rule nat.inject) lemma zero_less_hSuc [iff]: "\n. 0 < hSuc n" -by transfer (rule zero_less_Suc) + by transfer (rule zero_less_Suc) -lemma hypnat_minus_zero [simp]: "!!z. z - z = (0::hypnat)" -by transfer (rule diff_self_eq_0) +lemma hypnat_minus_zero [simp]: "\z::hypnat. z - z = 0" + by transfer (rule diff_self_eq_0) -lemma hypnat_diff_0_eq_0 [simp]: "!!n. (0::hypnat) - n = 0" -by transfer (rule diff_0_eq_0) +lemma hypnat_diff_0_eq_0 [simp]: "\n::hypnat. 0 - n = 0" + by transfer (rule diff_0_eq_0) -lemma hypnat_add_is_0 [iff]: "!!m n. (m+n = (0::hypnat)) = (m=0 & n=0)" -by transfer (rule add_is_0) +lemma hypnat_add_is_0 [iff]: "\m n::hypnat. m + n = 0 \ m = 0 \ n = 0" + by transfer (rule add_is_0) -lemma hypnat_diff_diff_left: "!!i j k. (i::hypnat) - j - k = i - (j+k)" -by transfer (rule diff_diff_left) +lemma hypnat_diff_diff_left: "\i j k::hypnat. i - j - k = i - (j + k)" + by transfer (rule diff_diff_left) -lemma hypnat_diff_commute: "!!i j k. (i::hypnat) - j - k = i-k-j" -by transfer (rule diff_commute) +lemma hypnat_diff_commute: "\i j k::hypnat. i - j - k = i - k - j" + by transfer (rule diff_commute) -lemma hypnat_diff_add_inverse [simp]: "!!m n. ((n::hypnat) + m) - n = m" -by transfer (rule diff_add_inverse) +lemma hypnat_diff_add_inverse [simp]: "\m n::hypnat. n + m - n = m" + by transfer (rule diff_add_inverse) -lemma hypnat_diff_add_inverse2 [simp]: "!!m n. ((m::hypnat) + n) - n = m" -by transfer (rule diff_add_inverse2) +lemma hypnat_diff_add_inverse2 [simp]: "\m n::hypnat. m + n - n = m" + by transfer (rule diff_add_inverse2) -lemma hypnat_diff_cancel [simp]: "!!k m n. ((k::hypnat) + m) - (k+n) = m - n" -by transfer (rule diff_cancel) +lemma hypnat_diff_cancel [simp]: "\k m n::hypnat. (k + m) - (k + n) = m - n" + by transfer (rule diff_cancel) -lemma hypnat_diff_cancel2 [simp]: "!!k m n. ((m::hypnat) + k) - (n+k) = m - n" -by transfer (rule diff_cancel2) +lemma hypnat_diff_cancel2 [simp]: "\k m n::hypnat. (m + k) - (n + k) = m - n" + by transfer (rule diff_cancel2) -lemma hypnat_diff_add_0 [simp]: "!!m n. (n::hypnat) - (n+m) = (0::hypnat)" -by transfer (rule diff_add_0) +lemma hypnat_diff_add_0 [simp]: "\m n::hypnat. n - (n + m) = 0" + by transfer (rule diff_add_0) -lemma hypnat_diff_mult_distrib: "!!k m n. ((m::hypnat) - n) * k = (m * k) - (n * k)" -by transfer (rule diff_mult_distrib) +lemma hypnat_diff_mult_distrib: "\k m n::hypnat. (m - n) * k = (m * k) - (n * k)" + by transfer (rule diff_mult_distrib) -lemma hypnat_diff_mult_distrib2: "!!k m n. (k::hypnat) * (m - n) = (k * m) - (k * n)" -by transfer (rule diff_mult_distrib2) +lemma hypnat_diff_mult_distrib2: "\k m n::hypnat. k * (m - n) = (k * m) - (k * n)" + by transfer (rule diff_mult_distrib2) -lemma hypnat_le_zero_cancel [iff]: "!!n. (n \ (0::hypnat)) = (n = 0)" -by transfer (rule le_0_eq) +lemma hypnat_le_zero_cancel [iff]: "\n::hypnat. n \ 0 \ n = 0" + by transfer (rule le_0_eq) -lemma hypnat_mult_is_0 [simp]: "!!m n. (m*n = (0::hypnat)) = (m=0 | n=0)" -by transfer (rule mult_is_0) +lemma hypnat_mult_is_0 [simp]: "\m n::hypnat. m * n = 0 \ m = 0 \ n = 0" + by transfer (rule mult_is_0) -lemma hypnat_diff_is_0_eq [simp]: "!!m n. ((m::hypnat) - n = 0) = (m \ n)" -by transfer (rule diff_is_0_eq) +lemma hypnat_diff_is_0_eq [simp]: "\m n::hypnat. m - n = 0 \ m \ n" + by transfer (rule diff_is_0_eq) -lemma hypnat_not_less0 [iff]: "!!n. ~ n < (0::hypnat)" -by transfer (rule not_less0) +lemma hypnat_not_less0 [iff]: "\n::hypnat. \ n < 0" + by transfer (rule not_less0) + +lemma hypnat_less_one [iff]: "\n::hypnat. n < 1 \ n = 0" + by transfer (rule less_one) -lemma hypnat_less_one [iff]: - "!!n. (n < (1::hypnat)) = (n=0)" -by transfer (rule less_one) +lemma hypnat_add_diff_inverse: "\m n::hypnat. \ m < n \ n + (m - n) = m" + by transfer (rule add_diff_inverse) -lemma hypnat_add_diff_inverse: "!!m n. ~ m n+(m-n) = (m::hypnat)" -by transfer (rule add_diff_inverse) +lemma hypnat_le_add_diff_inverse [simp]: "\m n::hypnat. n \ m \ n + (m - n) = m" + by transfer (rule le_add_diff_inverse) -lemma hypnat_le_add_diff_inverse [simp]: "!!m n. n \ m ==> n+(m-n) = (m::hypnat)" -by transfer (rule le_add_diff_inverse) - -lemma hypnat_le_add_diff_inverse2 [simp]: "!!m n. n\m ==> (m-n)+n = (m::hypnat)" -by transfer (rule le_add_diff_inverse2) +lemma hypnat_le_add_diff_inverse2 [simp]: "\m n::hypnat. n \ m \ (m - n) + n = m" + by transfer (rule le_add_diff_inverse2) declare hypnat_le_add_diff_inverse2 [OF order_less_imp_le] -lemma hypnat_le0 [iff]: "!!n. (0::hypnat) \ n" -by transfer (rule le0) +lemma hypnat_le0 [iff]: "\n::hypnat. 0 \ n" + by transfer (rule le0) -lemma hypnat_le_add1 [simp]: "!!x n. (x::hypnat) \ x + n" -by transfer (rule le_add1) +lemma hypnat_le_add1 [simp]: "\x n::hypnat. x \ x + n" + by transfer (rule le_add1) -lemma hypnat_add_self_le [simp]: "!!x n. (x::hypnat) \ n + x" -by transfer (rule le_add2) +lemma hypnat_add_self_le [simp]: "\x n::hypnat. x \ n + x" + by transfer (rule le_add2) -lemma hypnat_add_one_self_less [simp]: "(x::hypnat) < x + (1::hypnat)" +lemma hypnat_add_one_self_less [simp]: "x < x + 1" for x :: hypnat by (fact less_add_one) -lemma hypnat_neq0_conv [iff]: "!!n. (n \ 0) = (0 < (n::hypnat))" -by transfer (rule neq0_conv) +lemma hypnat_neq0_conv [iff]: "\n::hypnat. n \ 0 \ 0 < n" + by transfer (rule neq0_conv) -lemma hypnat_gt_zero_iff: "((0::hypnat) < n) = ((1::hypnat) \ n)" -by (auto simp add: linorder_not_less [symmetric]) +lemma hypnat_gt_zero_iff: "0 < n \ 1 \ n" for n :: hypnat + by (auto simp add: linorder_not_less [symmetric]) -lemma hypnat_gt_zero_iff2: "(0 < n) = (\m. n = m + (1::hypnat))" +lemma hypnat_gt_zero_iff2: "0 < n \ (\m. n = m + 1)" for n :: hypnat by (auto intro!: add_nonneg_pos exI[of _ "n - 1"] simp: hypnat_gt_zero_iff) -lemma hypnat_add_self_not_less: "~ (x + y < (x::hypnat))" -by (simp add: linorder_not_le [symmetric] add.commute [of x]) +lemma hypnat_add_self_not_less: "\ x + y < x" for x y :: hypnat + by (simp add: linorder_not_le [symmetric] add.commute [of x]) -lemma hypnat_diff_split: - "P(a - b::hypnat) = ((a P 0) & (ALL d. a = b + d --> P d))" - \ \elimination of \-\ on \hypnat\\ -proof (cases "a (a < b \ P 0) \ (\d. a = b + d \ P d)" + for a b :: hypnat + \ \elimination of \-\ on \hypnat\\ +proof (cases "a < b" rule: case_split) case True - thus ?thesis - by (auto simp add: hypnat_add_self_not_less order_less_imp_le - hypnat_diff_is_0_eq [THEN iffD2]) + then show ?thesis + by (auto simp add: hypnat_add_self_not_less order_less_imp_le hypnat_diff_is_0_eq [THEN iffD2]) next case False - thus ?thesis - by (auto simp add: linorder_not_less dest: order_le_less_trans) + then show ?thesis + by (auto simp add: linorder_not_less dest: order_le_less_trans) qed -subsection\Properties of the set of embedded natural numbers\ + +subsection \Properties of the set of embedded natural numbers\ lemma of_nat_eq_star_of [simp]: "of_nat = star_of" proof - fix n :: nat - show "of_nat n = star_of n" by transfer simp + show "of_nat n = star_of n" for n + by transfer simp qed lemma Nats_eq_Standard: "(Nats :: nat star set) = Standard" -by (auto simp add: Nats_def Standard_def) + by (auto simp: Nats_def Standard_def) lemma hypnat_of_nat_mem_Nats [simp]: "hypnat_of_nat n \ Nats" -by (simp add: Nats_eq_Standard) + by (simp add: Nats_eq_Standard) -lemma hypnat_of_nat_one [simp]: "hypnat_of_nat (Suc 0) = (1::hypnat)" -by transfer simp +lemma hypnat_of_nat_one [simp]: "hypnat_of_nat (Suc 0) = 1" + by transfer simp -lemma hypnat_of_nat_Suc [simp]: - "hypnat_of_nat (Suc n) = hypnat_of_nat n + (1::hypnat)" -by transfer simp +lemma hypnat_of_nat_Suc [simp]: "hypnat_of_nat (Suc n) = hypnat_of_nat n + 1" + by transfer simp -lemma of_nat_eq_add [rule_format]: - "\d::hypnat. of_nat m = of_nat n + d --> d \ range of_nat" -apply (induct n) -apply (auto simp add: add.assoc) -apply (case_tac x) -apply (auto simp add: add.commute [of 1]) -done +lemma of_nat_eq_add [rule_format]: "\d::hypnat. of_nat m = of_nat n + d --> d \ range of_nat" + apply (induct n) + apply (auto simp add: add.assoc) + apply (case_tac x) + apply (auto simp add: add.commute [of 1]) + done -lemma Nats_diff [simp]: "[|a \ Nats; b \ Nats|] ==> (a-b :: hypnat) \ Nats" -by (simp add: Nats_eq_Standard) +lemma Nats_diff [simp]: "a \ Nats \ b \ Nats \ a - b \ Nats" for a b :: hypnat + by (simp add: Nats_eq_Standard) -subsection\Infinite Hypernatural Numbers -- @{term HNatInfinite}\ +subsection \Infinite Hypernatural Numbers -- @{term HNatInfinite}\ + +text \The set of infinite hypernatural numbers.\ +definition HNatInfinite :: "hypnat set" + where "HNatInfinite = {n. n \ Nats}" -definition - (* the set of infinite hypernatural numbers *) - HNatInfinite :: "hypnat set" where - "HNatInfinite = {n. n \ Nats}" +lemma Nats_not_HNatInfinite_iff: "x \ Nats \ x \ HNatInfinite" + by (simp add: HNatInfinite_def) -lemma Nats_not_HNatInfinite_iff: "(x \ Nats) = (x \ HNatInfinite)" -by (simp add: HNatInfinite_def) - -lemma HNatInfinite_not_Nats_iff: "(x \ HNatInfinite) = (x \ Nats)" -by (simp add: HNatInfinite_def) +lemma HNatInfinite_not_Nats_iff: "x \ HNatInfinite \ x \ Nats" + by (simp add: HNatInfinite_def) lemma star_of_neq_HNatInfinite: "N \ HNatInfinite \ star_of n \ N" -by (auto simp add: HNatInfinite_def Nats_eq_Standard) + by (auto simp add: HNatInfinite_def Nats_eq_Standard) -lemma star_of_Suc_lessI: - "\N. \star_of n < N; star_of (Suc n) \ N\ \ star_of (Suc n) < N" -by transfer (rule Suc_lessI) +lemma star_of_Suc_lessI: "\N. star_of n < N \ star_of (Suc n) \ N \ star_of (Suc n) < N" + by transfer (rule Suc_lessI) lemma star_of_less_HNatInfinite: assumes N: "N \ HNatInfinite" shows "star_of n < N" proof (induct n) case 0 - from N have "star_of 0 \ N" by (rule star_of_neq_HNatInfinite) - thus "star_of 0 < N" by simp + from N have "star_of 0 \ N" + by (rule star_of_neq_HNatInfinite) + then show ?case by simp next case (Suc n) - from N have "star_of (Suc n) \ N" by (rule star_of_neq_HNatInfinite) - with Suc show "star_of (Suc n) < N" by (rule star_of_Suc_lessI) + from N have "star_of (Suc n) \ N" + by (rule star_of_neq_HNatInfinite) + with Suc show ?case + by (rule star_of_Suc_lessI) qed lemma star_of_le_HNatInfinite: "N \ HNatInfinite \ star_of n \ N" -by (rule star_of_less_HNatInfinite [THEN order_less_imp_le]) + by (rule star_of_less_HNatInfinite [THEN order_less_imp_le]) + subsubsection \Closure Rules\ -lemma Nats_less_HNatInfinite: "\x \ Nats; y \ HNatInfinite\ \ x < y" -by (auto simp add: Nats_def star_of_less_HNatInfinite) +lemma Nats_less_HNatInfinite: "x \ Nats \ y \ HNatInfinite \ x < y" + by (auto simp add: Nats_def star_of_less_HNatInfinite) -lemma Nats_le_HNatInfinite: "\x \ Nats; y \ HNatInfinite\ \ x \ y" -by (rule Nats_less_HNatInfinite [THEN order_less_imp_le]) +lemma Nats_le_HNatInfinite: "x \ Nats \ y \ HNatInfinite \ x \ y" + by (rule Nats_less_HNatInfinite [THEN order_less_imp_le]) lemma zero_less_HNatInfinite: "x \ HNatInfinite \ 0 < x" -by (simp add: Nats_less_HNatInfinite) + by (simp add: Nats_less_HNatInfinite) lemma one_less_HNatInfinite: "x \ HNatInfinite \ 1 < x" -by (simp add: Nats_less_HNatInfinite) + by (simp add: Nats_less_HNatInfinite) lemma one_le_HNatInfinite: "x \ HNatInfinite \ 1 \ x" -by (simp add: Nats_le_HNatInfinite) + by (simp add: Nats_le_HNatInfinite) lemma zero_not_mem_HNatInfinite [simp]: "0 \ HNatInfinite" -by (simp add: HNatInfinite_def) + by (simp add: HNatInfinite_def) -lemma Nats_downward_closed: - "\x \ Nats; (y::hypnat) \ x\ \ y \ Nats" -apply (simp only: linorder_not_less [symmetric]) -apply (erule contrapos_np) -apply (drule HNatInfinite_not_Nats_iff [THEN iffD2]) -apply (erule (1) Nats_less_HNatInfinite) -done +lemma Nats_downward_closed: "x \ Nats \ y \ x \ y \ Nats" for x y :: hypnat + apply (simp only: linorder_not_less [symmetric]) + apply (erule contrapos_np) + apply (drule HNatInfinite_not_Nats_iff [THEN iffD2]) + apply (erule (1) Nats_less_HNatInfinite) + done -lemma HNatInfinite_upward_closed: - "\x \ HNatInfinite; x \ y\ \ y \ HNatInfinite" -apply (simp only: HNatInfinite_not_Nats_iff) -apply (erule contrapos_nn) -apply (erule (1) Nats_downward_closed) -done +lemma HNatInfinite_upward_closed: "x \ HNatInfinite \ x \ y \ y \ HNatInfinite" + apply (simp only: HNatInfinite_not_Nats_iff) + apply (erule contrapos_nn) + apply (erule (1) Nats_downward_closed) + done lemma HNatInfinite_add: "x \ HNatInfinite \ x + y \ HNatInfinite" -apply (erule HNatInfinite_upward_closed) -apply (rule hypnat_le_add1) -done + apply (erule HNatInfinite_upward_closed) + apply (rule hypnat_le_add1) + done lemma HNatInfinite_add_one: "x \ HNatInfinite \ x + 1 \ HNatInfinite" -by (rule HNatInfinite_add) + by (rule HNatInfinite_add) -lemma HNatInfinite_diff: - "\x \ HNatInfinite; y \ Nats\ \ x - y \ HNatInfinite" -apply (frule (1) Nats_le_HNatInfinite) -apply (simp only: HNatInfinite_not_Nats_iff) -apply (erule contrapos_nn) -apply (drule (1) Nats_add, simp) -done +lemma HNatInfinite_diff: "x \ HNatInfinite \ y \ Nats \ x - y \ HNatInfinite" + apply (frule (1) Nats_le_HNatInfinite) + apply (simp only: HNatInfinite_not_Nats_iff) + apply (erule contrapos_nn) + apply (drule (1) Nats_add, simp) + done -lemma HNatInfinite_is_Suc: "x \ HNatInfinite ==> \y. x = y + (1::hypnat)" -apply (rule_tac x = "x - (1::hypnat) " in exI) -apply (simp add: Nats_le_HNatInfinite) -done +lemma HNatInfinite_is_Suc: "x \ HNatInfinite \ \y. x = y + 1" for x :: hypnat + apply (rule_tac x = "x - (1::hypnat) " in exI) + apply (simp add: Nats_le_HNatInfinite) + done -subsection\Existence of an infinite hypernatural number\ +subsection \Existence of an infinite hypernatural number\ -definition - (* \ is in fact an infinite hypernatural number = [<1,2,3,...>] *) - whn :: hypnat where - hypnat_omega_def: "whn = star_n (%n::nat. n)" +text \\\\ is in fact an infinite hypernatural number = \[<1, 2, 3, \>]\\ +definition whn :: hypnat + where hypnat_omega_def: "whn = star_n (\n::nat. n)" lemma hypnat_of_nat_neq_whn: "hypnat_of_nat n \ whn" -by (simp add: FreeUltrafilterNat.singleton' hypnat_omega_def star_of_def star_n_eq_iff) + by (simp add: FreeUltrafilterNat.singleton' hypnat_omega_def star_of_def star_n_eq_iff) lemma whn_neq_hypnat_of_nat: "whn \ hypnat_of_nat n" -by (simp add: FreeUltrafilterNat.singleton hypnat_omega_def star_of_def star_n_eq_iff) + by (simp add: FreeUltrafilterNat.singleton hypnat_omega_def star_of_def star_n_eq_iff) lemma whn_not_Nats [simp]: "whn \ Nats" -by (simp add: Nats_def image_def whn_neq_hypnat_of_nat) + by (simp add: Nats_def image_def whn_neq_hypnat_of_nat) lemma HNatInfinite_whn [simp]: "whn \ HNatInfinite" -by (simp add: HNatInfinite_def) + by (simp add: HNatInfinite_def) lemma lemma_unbounded_set [simp]: "eventually (\n::nat. m < n) \" by (rule filter_leD[OF FreeUltrafilterNat.le_cofinite]) (auto simp add: cofinite_eq_sequentially eventually_at_top_dense) -lemma hypnat_of_nat_eq: - "hypnat_of_nat m = star_n (%n::nat. m)" -by (simp add: star_of_def) +lemma hypnat_of_nat_eq: "hypnat_of_nat m = star_n (\n::nat. m)" + by (simp add: star_of_def) lemma SHNat_eq: "Nats = {n. \N. n = hypnat_of_nat N}" -by (simp add: Nats_def image_def) + by (simp add: Nats_def image_def) lemma Nats_less_whn: "n \ Nats \ n < whn" -by (simp add: Nats_less_HNatInfinite) + by (simp add: Nats_less_HNatInfinite) lemma Nats_le_whn: "n \ Nats \ n \ whn" -by (simp add: Nats_le_HNatInfinite) + by (simp add: Nats_le_HNatInfinite) lemma hypnat_of_nat_less_whn [simp]: "hypnat_of_nat n < whn" -by (simp add: Nats_less_whn) + by (simp add: Nats_less_whn) lemma hypnat_of_nat_le_whn [simp]: "hypnat_of_nat n \ whn" -by (simp add: Nats_le_whn) + by (simp add: Nats_le_whn) lemma hypnat_zero_less_hypnat_omega [simp]: "0 < whn" -by (simp add: Nats_less_whn) + by (simp add: Nats_less_whn) lemma hypnat_one_less_hypnat_omega [simp]: "1 < whn" -by (simp add: Nats_less_whn) + by (simp add: Nats_less_whn) -subsubsection\Alternative characterization of the set of infinite hypernaturals\ +subsubsection \Alternative characterization of the set of infinite hypernaturals\ -text\@{term "HNatInfinite = {N. \n \ Nats. n < N}"}\ +text \@{term "HNatInfinite = {N. \n \ Nats. n < N}"}\ (*??delete? similar reasoning in hypnat_omega_gt_SHNat above*) lemma HNatInfinite_FreeUltrafilterNat_lemma: assumes "\N::nat. eventually (\n. f n \ N) \" shows "eventually (\n. N < f n) \" -apply (induct N) -using assms -apply (drule_tac x = 0 in spec, simp) -using assms -apply (drule_tac x = "Suc N" in spec) -apply (auto elim: eventually_elim2) -done + apply (induct N) + using assms + apply (drule_tac x = 0 in spec, simp) + using assms + apply (drule_tac x = "Suc N" in spec) + apply (auto elim: eventually_elim2) + done lemma HNatInfinite_iff: "HNatInfinite = {N. \n \ Nats. n < N}" -apply (safe intro!: Nats_less_HNatInfinite) -apply (auto simp add: HNatInfinite_def) -done + apply (safe intro!: Nats_less_HNatInfinite) + apply (auto simp add: HNatInfinite_def) + done -subsubsection\Alternative Characterization of @{term HNatInfinite} using -Free Ultrafilter\ +subsubsection \Alternative Characterization of @{term HNatInfinite} using Free Ultrafilter\ lemma HNatInfinite_FreeUltrafilterNat: - "star_n X \ HNatInfinite ==> \u. eventually (\n. u < X n) FreeUltrafilterNat" -apply (auto simp add: HNatInfinite_iff SHNat_eq) -apply (drule_tac x="star_of u" in spec, simp) -apply (simp add: star_of_def star_less_def starP2_star_n) -done + "star_n X \ HNatInfinite \ \u. eventually (\n. u < X n) FreeUltrafilterNat" + apply (auto simp add: HNatInfinite_iff SHNat_eq) + apply (drule_tac x="star_of u" in spec, simp) + apply (simp add: star_of_def star_less_def starP2_star_n) + done lemma FreeUltrafilterNat_HNatInfinite: - "\u. eventually (\n. u < X n) FreeUltrafilterNat ==> star_n X \ HNatInfinite" -by (auto simp add: star_less_def starP2_star_n HNatInfinite_iff SHNat_eq hypnat_of_nat_eq) + "\u. eventually (\n. u < X n) FreeUltrafilterNat \ star_n X \ HNatInfinite" + by (auto simp add: star_less_def starP2_star_n HNatInfinite_iff SHNat_eq hypnat_of_nat_eq) lemma HNatInfinite_FreeUltrafilterNat_iff: - "(star_n X \ HNatInfinite) = (\u. eventually (\n. u < X n) FreeUltrafilterNat)" -by (rule iffI [OF HNatInfinite_FreeUltrafilterNat - FreeUltrafilterNat_HNatInfinite]) + "(star_n X \ HNatInfinite) = (\u. eventually (\n. u < X n) FreeUltrafilterNat)" + by (rule iffI [OF HNatInfinite_FreeUltrafilterNat FreeUltrafilterNat_HNatInfinite]) + subsection \Embedding of the Hypernaturals into other types\ -definition - of_hypnat :: "hypnat \ 'a::semiring_1_cancel star" where - of_hypnat_def [transfer_unfold]: "of_hypnat = *f* of_nat" +definition of_hypnat :: "hypnat \ 'a::semiring_1_cancel star" + where of_hypnat_def [transfer_unfold]: "of_hypnat = *f* of_nat" lemma of_hypnat_0 [simp]: "of_hypnat 0 = 0" -by transfer (rule of_nat_0) + by transfer (rule of_nat_0) lemma of_hypnat_1 [simp]: "of_hypnat 1 = 1" -by transfer (rule of_nat_1) + by transfer (rule of_nat_1) lemma of_hypnat_hSuc: "\m. of_hypnat (hSuc m) = 1 + of_hypnat m" -by transfer (rule of_nat_Suc) + by transfer (rule of_nat_Suc) -lemma of_hypnat_add [simp]: - "\m n. of_hypnat (m + n) = of_hypnat m + of_hypnat n" -by transfer (rule of_nat_add) +lemma of_hypnat_add [simp]: "\m n. of_hypnat (m + n) = of_hypnat m + of_hypnat n" + by transfer (rule of_nat_add) -lemma of_hypnat_mult [simp]: - "\m n. of_hypnat (m * n) = of_hypnat m * of_hypnat n" -by transfer (rule of_nat_mult) +lemma of_hypnat_mult [simp]: "\m n. of_hypnat (m * n) = of_hypnat m * of_hypnat n" + by transfer (rule of_nat_mult) lemma of_hypnat_less_iff [simp]: - "\m n. (of_hypnat m < (of_hypnat n::'a::linordered_semidom star)) = (m < n)" -by transfer (rule of_nat_less_iff) + "\m n. of_hypnat m < (of_hypnat n::'a::linordered_semidom star) \ m < n" + by transfer (rule of_nat_less_iff) lemma of_hypnat_0_less_iff [simp]: - "\n. (0 < (of_hypnat n::'a::linordered_semidom star)) = (0 < n)" -by transfer (rule of_nat_0_less_iff) + "\n. 0 < (of_hypnat n::'a::linordered_semidom star) \ 0 < n" + by transfer (rule of_nat_0_less_iff) -lemma of_hypnat_less_0_iff [simp]: - "\m. \ (of_hypnat m::'a::linordered_semidom star) < 0" -by transfer (rule of_nat_less_0_iff) +lemma of_hypnat_less_0_iff [simp]: "\m. \ (of_hypnat m::'a::linordered_semidom star) < 0" + by transfer (rule of_nat_less_0_iff) lemma of_hypnat_le_iff [simp]: - "\m n. (of_hypnat m \ (of_hypnat n::'a::linordered_semidom star)) = (m \ n)" -by transfer (rule of_nat_le_iff) + "\m n. of_hypnat m \ (of_hypnat n::'a::linordered_semidom star) \ m \ n" + by transfer (rule of_nat_le_iff) -lemma of_hypnat_0_le_iff [simp]: - "\n. 0 \ (of_hypnat n::'a::linordered_semidom star)" -by transfer (rule of_nat_0_le_iff) +lemma of_hypnat_0_le_iff [simp]: "\n. 0 \ (of_hypnat n::'a::linordered_semidom star)" + by transfer (rule of_nat_0_le_iff) -lemma of_hypnat_le_0_iff [simp]: - "\m. ((of_hypnat m::'a::linordered_semidom star) \ 0) = (m = 0)" -by transfer (rule of_nat_le_0_iff) +lemma of_hypnat_le_0_iff [simp]: "\m. (of_hypnat m::'a::linordered_semidom star) \ 0 \ m = 0" + by transfer (rule of_nat_le_0_iff) lemma of_hypnat_eq_iff [simp]: - "\m n. (of_hypnat m = (of_hypnat n::'a::linordered_semidom star)) = (m = n)" -by transfer (rule of_nat_eq_iff) + "\m n. of_hypnat m = (of_hypnat n::'a::linordered_semidom star) \ m = n" + by transfer (rule of_nat_eq_iff) -lemma of_hypnat_eq_0_iff [simp]: - "\m. ((of_hypnat m::'a::linordered_semidom star) = 0) = (m = 0)" -by transfer (rule of_nat_eq_0_iff) +lemma of_hypnat_eq_0_iff [simp]: "\m. (of_hypnat m::'a::linordered_semidom star) = 0 \ m = 0" + by transfer (rule of_nat_eq_0_iff) lemma HNatInfinite_of_hypnat_gt_zero: "N \ HNatInfinite \ (0::'a::linordered_semidom star) < of_hypnat N" -by (rule ccontr, simp add: linorder_not_less) + by (rule ccontr) (simp add: linorder_not_less) end diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/NSA.thy --- a/src/HOL/Nonstandard_Analysis/NSA.thy Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/NSA.thy Tue Nov 01 00:44:24 2016 +0100 @@ -3,649 +3,584 @@ Author: Lawrence C Paulson, University of Cambridge *) -section\Infinite Numbers, Infinitesimals, Infinitely Close Relation\ +section \Infinite Numbers, Infinitesimals, Infinitely Close Relation\ theory NSA -imports HyperDef "~~/src/HOL/Library/Lub_Glb" + imports HyperDef "~~/src/HOL/Library/Lub_Glb" begin -definition - hnorm :: "'a::real_normed_vector star \ real star" where - [transfer_unfold]: "hnorm = *f* norm" +definition hnorm :: "'a::real_normed_vector star \ real star" + where [transfer_unfold]: "hnorm = *f* norm" -definition - Infinitesimal :: "('a::real_normed_vector) star set" where - "Infinitesimal = {x. \r \ Reals. 0 < r --> hnorm x < r}" +definition Infinitesimal :: "('a::real_normed_vector) star set" + where "Infinitesimal = {x. \r \ Reals. 0 < r \ hnorm x < r}" -definition - HFinite :: "('a::real_normed_vector) star set" where - "HFinite = {x. \r \ Reals. hnorm x < r}" +definition HFinite :: "('a::real_normed_vector) star set" + where "HFinite = {x. \r \ Reals. hnorm x < r}" -definition - HInfinite :: "('a::real_normed_vector) star set" where - "HInfinite = {x. \r \ Reals. r < hnorm x}" +definition HInfinite :: "('a::real_normed_vector) star set" + where "HInfinite = {x. \r \ Reals. r < hnorm x}" -definition - approx :: "['a::real_normed_vector star, 'a star] => bool" (infixl "\" 50) where - \\the `infinitely close' relation\ - "(x \ y) = ((x - y) \ Infinitesimal)" +definition approx :: "'a::real_normed_vector star \ 'a star \ bool" (infixl "\" 50) + where "x \ y \ x - y \ Infinitesimal" + \ \the ``infinitely close'' relation\ -definition - st :: "hypreal => hypreal" where - \\the standard part of a hyperreal\ - "st = (%x. @r. x \ HFinite & r \ \ & r \ x)" +definition st :: "hypreal \ hypreal" + where "st = (\x. SOME r. x \ HFinite \ r \ \ \ r \ x)" + \ \the standard part of a hyperreal\ -definition - monad :: "'a::real_normed_vector star => 'a star set" where - "monad x = {y. x \ y}" +definition monad :: "'a::real_normed_vector star \ 'a star set" + where "monad x = {y. x \ y}" -definition - galaxy :: "'a::real_normed_vector star => 'a star set" where - "galaxy x = {y. (x + -y) \ HFinite}" +definition galaxy :: "'a::real_normed_vector star \ 'a star set" + where "galaxy x = {y. (x + -y) \ HFinite}" -lemma SReal_def: "\ == {x. \r. x = hypreal_of_real r}" -by (simp add: Reals_def image_def) +lemma SReal_def: "\ \ {x. \r. x = hypreal_of_real r}" + by (simp add: Reals_def image_def) + subsection \Nonstandard Extension of the Norm Function\ -definition - scaleHR :: "real star \ 'a star \ 'a::real_normed_vector star" where - [transfer_unfold]: "scaleHR = starfun2 scaleR" +definition scaleHR :: "real star \ 'a star \ 'a::real_normed_vector star" + where [transfer_unfold]: "scaleHR = starfun2 scaleR" lemma Standard_hnorm [simp]: "x \ Standard \ hnorm x \ Standard" -by (simp add: hnorm_def) + by (simp add: hnorm_def) lemma star_of_norm [simp]: "star_of (norm x) = hnorm (star_of x)" -by transfer (rule refl) + by transfer (rule refl) -lemma hnorm_ge_zero [simp]: - "\x::'a::real_normed_vector star. 0 \ hnorm x" -by transfer (rule norm_ge_zero) +lemma hnorm_ge_zero [simp]: "\x::'a::real_normed_vector star. 0 \ hnorm x" + by transfer (rule norm_ge_zero) -lemma hnorm_eq_zero [simp]: - "\x::'a::real_normed_vector star. (hnorm x = 0) = (x = 0)" -by transfer (rule norm_eq_zero) +lemma hnorm_eq_zero [simp]: "\x::'a::real_normed_vector star. hnorm x = 0 \ x = 0" + by transfer (rule norm_eq_zero) -lemma hnorm_triangle_ineq: - "\x y::'a::real_normed_vector star. hnorm (x + y) \ hnorm x + hnorm y" -by transfer (rule norm_triangle_ineq) - -lemma hnorm_triangle_ineq3: - "\x y::'a::real_normed_vector star. \hnorm x - hnorm y\ \ hnorm (x - y)" -by transfer (rule norm_triangle_ineq3) +lemma hnorm_triangle_ineq: "\x y::'a::real_normed_vector star. hnorm (x + y) \ hnorm x + hnorm y" + by transfer (rule norm_triangle_ineq) -lemma hnorm_scaleR: - "\x::'a::real_normed_vector star. - hnorm (a *\<^sub>R x) = \star_of a\ * hnorm x" -by transfer (rule norm_scaleR) +lemma hnorm_triangle_ineq3: "\x y::'a::real_normed_vector star. \hnorm x - hnorm y\ \ hnorm (x - y)" + by transfer (rule norm_triangle_ineq3) + +lemma hnorm_scaleR: "\x::'a::real_normed_vector star. hnorm (a *\<^sub>R x) = \star_of a\ * hnorm x" + by transfer (rule norm_scaleR) -lemma hnorm_scaleHR: - "\a (x::'a::real_normed_vector star). - hnorm (scaleHR a x) = \a\ * hnorm x" -by transfer (rule norm_scaleR) +lemma hnorm_scaleHR: "\a (x::'a::real_normed_vector star). hnorm (scaleHR a x) = \a\ * hnorm x" + by transfer (rule norm_scaleR) -lemma hnorm_mult_ineq: - "\x y::'a::real_normed_algebra star. hnorm (x * y) \ hnorm x * hnorm y" -by transfer (rule norm_mult_ineq) +lemma hnorm_mult_ineq: "\x y::'a::real_normed_algebra star. hnorm (x * y) \ hnorm x * hnorm y" + by transfer (rule norm_mult_ineq) -lemma hnorm_mult: - "\x y::'a::real_normed_div_algebra star. hnorm (x * y) = hnorm x * hnorm y" -by transfer (rule norm_mult) +lemma hnorm_mult: "\x y::'a::real_normed_div_algebra star. hnorm (x * y) = hnorm x * hnorm y" + by transfer (rule norm_mult) -lemma hnorm_hyperpow: - "\(x::'a::{real_normed_div_algebra} star) n. - hnorm (x pow n) = hnorm x pow n" -by transfer (rule norm_power) +lemma hnorm_hyperpow: "\(x::'a::{real_normed_div_algebra} star) n. hnorm (x pow n) = hnorm x pow n" + by transfer (rule norm_power) + +lemma hnorm_one [simp]: "hnorm (1::'a::real_normed_div_algebra star) = 1" + by transfer (rule norm_one) -lemma hnorm_one [simp]: - "hnorm (1::'a::real_normed_div_algebra star) = 1" -by transfer (rule norm_one) +lemma hnorm_zero [simp]: "hnorm (0::'a::real_normed_vector star) = 0" + by transfer (rule norm_zero) -lemma hnorm_zero [simp]: - "hnorm (0::'a::real_normed_vector star) = 0" -by transfer (rule norm_zero) +lemma zero_less_hnorm_iff [simp]: "\x::'a::real_normed_vector star. 0 < hnorm x \ x \ 0" + by transfer (rule zero_less_norm_iff) -lemma zero_less_hnorm_iff [simp]: - "\x::'a::real_normed_vector star. (0 < hnorm x) = (x \ 0)" -by transfer (rule zero_less_norm_iff) +lemma hnorm_minus_cancel [simp]: "\x::'a::real_normed_vector star. hnorm (- x) = hnorm x" + by transfer (rule norm_minus_cancel) -lemma hnorm_minus_cancel [simp]: - "\x::'a::real_normed_vector star. hnorm (- x) = hnorm x" -by transfer (rule norm_minus_cancel) +lemma hnorm_minus_commute: "\a b::'a::real_normed_vector star. hnorm (a - b) = hnorm (b - a)" + by transfer (rule norm_minus_commute) -lemma hnorm_minus_commute: - "\a b::'a::real_normed_vector star. hnorm (a - b) = hnorm (b - a)" -by transfer (rule norm_minus_commute) - -lemma hnorm_triangle_ineq2: - "\a b::'a::real_normed_vector star. hnorm a - hnorm b \ hnorm (a - b)" -by transfer (rule norm_triangle_ineq2) +lemma hnorm_triangle_ineq2: "\a b::'a::real_normed_vector star. hnorm a - hnorm b \ hnorm (a - b)" + by transfer (rule norm_triangle_ineq2) -lemma hnorm_triangle_ineq4: - "\a b::'a::real_normed_vector star. hnorm (a - b) \ hnorm a + hnorm b" -by transfer (rule norm_triangle_ineq4) +lemma hnorm_triangle_ineq4: "\a b::'a::real_normed_vector star. hnorm (a - b) \ hnorm a + hnorm b" + by transfer (rule norm_triangle_ineq4) -lemma abs_hnorm_cancel [simp]: - "\a::'a::real_normed_vector star. \hnorm a\ = hnorm a" -by transfer (rule abs_norm_cancel) +lemma abs_hnorm_cancel [simp]: "\a::'a::real_normed_vector star. \hnorm a\ = hnorm a" + by transfer (rule abs_norm_cancel) -lemma hnorm_of_hypreal [simp]: - "\r. hnorm (of_hypreal r::'a::real_normed_algebra_1 star) = \r\" -by transfer (rule norm_of_real) +lemma hnorm_of_hypreal [simp]: "\r. hnorm (of_hypreal r::'a::real_normed_algebra_1 star) = \r\" + by transfer (rule norm_of_real) lemma nonzero_hnorm_inverse: - "\a::'a::real_normed_div_algebra star. - a \ 0 \ hnorm (inverse a) = inverse (hnorm a)" -by transfer (rule nonzero_norm_inverse) + "\a::'a::real_normed_div_algebra star. a \ 0 \ hnorm (inverse a) = inverse (hnorm a)" + by transfer (rule nonzero_norm_inverse) lemma hnorm_inverse: - "\a::'a::{real_normed_div_algebra, division_ring} star. - hnorm (inverse a) = inverse (hnorm a)" -by transfer (rule norm_inverse) + "\a::'a::{real_normed_div_algebra, division_ring} star. hnorm (inverse a) = inverse (hnorm a)" + by transfer (rule norm_inverse) -lemma hnorm_divide: - "\a b::'a::{real_normed_field, field} star. - hnorm (a / b) = hnorm a / hnorm b" -by transfer (rule norm_divide) +lemma hnorm_divide: "\a b::'a::{real_normed_field, field} star. hnorm (a / b) = hnorm a / hnorm b" + by transfer (rule norm_divide) -lemma hypreal_hnorm_def [simp]: - "\r::hypreal. hnorm r = \r\" -by transfer (rule real_norm_def) +lemma hypreal_hnorm_def [simp]: "\r::hypreal. hnorm r = \r\" + by transfer (rule real_norm_def) lemma hnorm_add_less: - "\(x::'a::real_normed_vector star) y r s. - \hnorm x < r; hnorm y < s\ \ hnorm (x + y) < r + s" -by transfer (rule norm_add_less) + "\(x::'a::real_normed_vector star) y r s. hnorm x < r \ hnorm y < s \ hnorm (x + y) < r + s" + by transfer (rule norm_add_less) lemma hnorm_mult_less: - "\(x::'a::real_normed_algebra star) y r s. - \hnorm x < r; hnorm y < s\ \ hnorm (x * y) < r * s" -by transfer (rule norm_mult_less) + "\(x::'a::real_normed_algebra star) y r s. hnorm x < r \ hnorm y < s \ hnorm (x * y) < r * s" + by transfer (rule norm_mult_less) -lemma hnorm_scaleHR_less: - "\\x\ < r; hnorm y < s\ \ hnorm (scaleHR x y) < r * s" -apply (simp only: hnorm_scaleHR) -apply (simp add: mult_strict_mono') -done +lemma hnorm_scaleHR_less: "\x\ < r \ hnorm y < s \ hnorm (scaleHR x y) < r * s" + by (simp only: hnorm_scaleHR) (simp add: mult_strict_mono') + + +subsection \Closure Laws for the Standard Reals\ -subsection\Closure Laws for the Standard Reals\ +lemma Reals_minus_iff [simp]: "- x \ \ \ x \ \" + apply auto + apply (drule Reals_minus) + apply auto + done -lemma Reals_minus_iff [simp]: "(-x \ \) = (x \ \)" -apply auto -apply (drule Reals_minus, auto) -done +lemma Reals_add_cancel: "x + y \ \ \ y \ \ \ x \ \" + by (drule (1) Reals_diff) simp -lemma Reals_add_cancel: "\x + y \ \; y \ \\ \ x \ \" -by (drule (1) Reals_diff, simp) - -lemma SReal_hrabs: "(x::hypreal) \ \ ==> \x\ \ \" -by (simp add: Reals_eq_Standard) +lemma SReal_hrabs: "x \ \ \ \x\ \ \" + for x :: hypreal + by (simp add: Reals_eq_Standard) lemma SReal_hypreal_of_real [simp]: "hypreal_of_real x \ \" -by (simp add: Reals_eq_Standard) + by (simp add: Reals_eq_Standard) -lemma SReal_divide_numeral: "r \ \ ==> r/(numeral w::hypreal) \ \" -by simp +lemma SReal_divide_numeral: "r \ \ \ r / (numeral w::hypreal) \ \" + by simp text \\\\ is not in Reals because it is an infinitesimal\ lemma SReal_epsilon_not_mem: "\ \ \" -apply (simp add: SReal_def) -apply (auto simp add: hypreal_of_real_not_eq_epsilon [THEN not_sym]) -done + by (auto simp: SReal_def hypreal_of_real_not_eq_epsilon [symmetric]) lemma SReal_omega_not_mem: "\ \ \" -apply (simp add: SReal_def) -apply (auto simp add: hypreal_of_real_not_eq_omega [THEN not_sym]) -done + by (auto simp: SReal_def hypreal_of_real_not_eq_omega [symmetric]) lemma SReal_UNIV_real: "{x. hypreal_of_real x \ \} = (UNIV::real set)" -by simp + by simp -lemma SReal_iff: "(x \ \) = (\y. x = hypreal_of_real y)" -by (simp add: SReal_def) +lemma SReal_iff: "x \ \ \ (\y. x = hypreal_of_real y)" + by (simp add: SReal_def) lemma hypreal_of_real_image: "hypreal_of_real `(UNIV::real set) = \" -by (simp add: Reals_eq_Standard Standard_def) + 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 + 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" -by (simp add: SReal_def image_def, blast) +lemma SReal_hypreal_of_real_image: "\x. x \ P \ P \ \ \ \Q. P = hypreal_of_real ` Q" + unfolding SReal_def image_def by blast -lemma SReal_dense: - "[| (x::hypreal) \ \; y \ \; x \r \ Reals. 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.\ + +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]) + "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 + "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 -subsection\Set of Finite Elements is a Subring of the Extended Reals\ +subsection \Set of Finite Elements is a Subring of the Extended Reals\ -lemma HFinite_add: "[|x \ HFinite; y \ HFinite|] ==> (x+y) \ HFinite" -apply (simp add: HFinite_def) -apply (blast intro!: Reals_add hnorm_add_less) -done +lemma HFinite_add: "x \ HFinite \ y \ HFinite \ x + y \ HFinite" + unfolding HFinite_def by (blast intro!: Reals_add hnorm_add_less) -lemma HFinite_mult: - fixes x y :: "'a::real_normed_algebra star" - shows "[|x \ HFinite; y \ HFinite|] ==> x*y \ HFinite" -apply (simp add: HFinite_def) -apply (blast intro!: Reals_mult hnorm_mult_less) -done +lemma HFinite_mult: "x \ HFinite \ y \ HFinite \ x * y \ HFinite" + for x y :: "'a::real_normed_algebra star" + unfolding HFinite_def by (blast intro!: Reals_mult hnorm_mult_less) -lemma HFinite_scaleHR: - "[|x \ HFinite; y \ HFinite|] ==> scaleHR x y \ HFinite" -apply (simp add: HFinite_def) -apply (blast intro!: Reals_mult hnorm_scaleHR_less) -done +lemma HFinite_scaleHR: "x \ HFinite \ y \ HFinite \ scaleHR x y \ HFinite" + by (auto simp: HFinite_def intro!: Reals_mult hnorm_scaleHR_less) -lemma HFinite_minus_iff: "(-x \ HFinite) = (x \ HFinite)" -by (simp add: HFinite_def) +lemma HFinite_minus_iff: "- x \ HFinite \ x \ HFinite" + 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 + 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 lemma SReal_subset_HFinite: "(\::hypreal set) \ HFinite" -by (auto simp add: SReal_def) + by (auto simp add: SReal_def) -lemma HFiniteD: "x \ HFinite ==> \t \ Reals. hnorm x < t" -by (simp add: HFinite_def) +lemma HFiniteD: "x \ HFinite \ \t \ Reals. hnorm x < t" + by (simp add: HFinite_def) -lemma HFinite_hrabs_iff [iff]: "(\x::hypreal\ \ HFinite) = (x \ HFinite)" -by (simp add: HFinite_def) +lemma HFinite_hrabs_iff [iff]: "\x\ \ HFinite \ x \ HFinite" + for x :: hypreal + by (simp add: HFinite_def) -lemma HFinite_hnorm_iff [iff]: - "(hnorm (x::hypreal) \ HFinite) = (x \ HFinite)" -by (simp add: HFinite_def) +lemma HFinite_hnorm_iff [iff]: "hnorm x \ HFinite \ x \ HFinite" + for x :: hypreal + by (simp add: HFinite_def) lemma HFinite_numeral [simp]: "numeral w \ HFinite" -unfolding star_numeral_def by (rule HFinite_star_of) + unfolding star_numeral_def by (rule HFinite_star_of) -(** As always with numerals, 0 and 1 are special cases **) +text \As always with numerals, \0\ and \1\ are special cases.\ lemma HFinite_0 [simp]: "0 \ HFinite" -unfolding star_zero_def by (rule HFinite_star_of) + unfolding star_zero_def by (rule HFinite_star_of) lemma HFinite_1 [simp]: "1 \ HFinite" -unfolding star_one_def by (rule HFinite_star_of) + unfolding star_one_def by (rule HFinite_star_of) -lemma hrealpow_HFinite: - fixes x :: "'a::{real_normed_algebra,monoid_mult} star" - shows "x \ HFinite ==> x ^ n \ HFinite" -apply (induct n) -apply (auto simp add: power_Suc intro: HFinite_mult) -done +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) -lemma HFinite_bounded: - "[|(x::hypreal) \ HFinite; y \ x; 0 \ y |] ==> y \ HFinite" -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: "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 -subsection\Set of Infinitesimals is a Subring of the Hyperreals\ +subsection \Set of Infinitesimals is a Subring of the Hyperreals\ -lemma InfinitesimalI: - "(\r. \r \ \; 0 < r\ \ hnorm x < r) \ x \ Infinitesimal" -by (simp add: Infinitesimal_def) +lemma InfinitesimalI: "(\r. r \ \ \ 0 < r \ hnorm x < r) \ x \ Infinitesimal" + by (simp add: Infinitesimal_def) -lemma InfinitesimalD: - "x \ Infinitesimal ==> \r \ Reals. 0 < r --> hnorm x < r" -by (simp add: Infinitesimal_def) +lemma InfinitesimalD: "x \ Infinitesimal \ \r \ Reals. 0 < r \ hnorm x < r" + by (simp add: Infinitesimal_def) -lemma InfinitesimalI2: - "(\r. 0 < r \ hnorm x < star_of r) \ x \ Infinitesimal" -by (auto simp add: Infinitesimal_def SReal_def) +lemma InfinitesimalI2: "(\r. 0 < r \ hnorm x < star_of r) \ x \ Infinitesimal" + by (auto simp add: Infinitesimal_def SReal_def) -lemma InfinitesimalD2: - "\x \ Infinitesimal; 0 < r\ \ hnorm x < star_of r" -by (auto simp add: Infinitesimal_def SReal_def) +lemma InfinitesimalD2: "x \ Infinitesimal \ 0 < r \ hnorm x < star_of r" + by (auto simp add: Infinitesimal_def SReal_def) lemma Infinitesimal_zero [iff]: "0 \ Infinitesimal" -by (simp add: Infinitesimal_def) + by (simp add: Infinitesimal_def) -lemma hypreal_sum_of_halves: "x/(2::hypreal) + x/(2::hypreal) = x" -by auto +lemma hypreal_sum_of_halves: "x / 2 + x / 2 = x" + for x :: hypreal + by auto -lemma Infinitesimal_add: - "[| x \ Infinitesimal; y \ Infinitesimal |] ==> (x+y) \ Infinitesimal" -apply (rule InfinitesimalI) -apply (rule hypreal_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: "x \ Infinitesimal \ y \ Infinitesimal \ x + y \ Infinitesimal" + apply (rule InfinitesimalI) + apply (rule hypreal_sum_of_halves [THEN subst]) + apply (drule half_gt_zero) + apply (blast intro: hnorm_add_less SReal_divide_numeral dest: InfinitesimalD) + done -lemma Infinitesimal_minus_iff [simp]: "(-x:Infinitesimal) = (x:Infinitesimal)" -by (simp add: Infinitesimal_def) +lemma Infinitesimal_minus_iff [simp]: "- x \ Infinitesimal \ x \ Infinitesimal" + by (simp add: Infinitesimal_def) -lemma Infinitesimal_hnorm_iff: - "(hnorm x \ Infinitesimal) = (x \ Infinitesimal)" -by (simp add: Infinitesimal_def) +lemma Infinitesimal_hnorm_iff: "hnorm x \ Infinitesimal \ x \ Infinitesimal" + by (simp add: Infinitesimal_def) -lemma Infinitesimal_hrabs_iff [iff]: - "(\x::hypreal\ \ Infinitesimal) = (x \ Infinitesimal)" -by (simp add: abs_if) +lemma Infinitesimal_hrabs_iff [iff]: "\x\ \ Infinitesimal \ x \ Infinitesimal" + for x :: hypreal + by (simp add: abs_if) lemma Infinitesimal_of_hypreal_iff [simp]: - "((of_hypreal x::'a::real_normed_algebra_1 star) \ Infinitesimal) = - (x \ Infinitesimal)" -by (subst Infinitesimal_hnorm_iff [symmetric], simp) + "(of_hypreal x::'a::real_normed_algebra_1 star) \ Infinitesimal \ x \ Infinitesimal" + by (subst Infinitesimal_hnorm_iff [symmetric]) simp -lemma Infinitesimal_diff: - "[| x \ Infinitesimal; y \ Infinitesimal |] ==> x-y \ Infinitesimal" +lemma Infinitesimal_diff: "x \ Infinitesimal \ y \ Infinitesimal \ x - y \ Infinitesimal" using Infinitesimal_add [of x "- y"] by simp -lemma Infinitesimal_mult: - fixes x y :: "'a::real_normed_algebra star" - shows "[|x \ Infinitesimal; y \ Infinitesimal|] ==> (x * y) \ Infinitesimal" -apply (rule InfinitesimalI) -apply (subgoal_tac "hnorm (x * y) < 1 * r", simp only: mult_1) -apply (rule hnorm_mult_less) -apply (simp_all add: InfinitesimalD) -done +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_HFinite_mult: - fixes x y :: "'a::real_normed_algebra star" - shows "[| x \ Infinitesimal; y \ HFinite |] ==> (x * y) \ Infinitesimal" -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: "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_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 + "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 lemma Infinitesimal_HFinite_mult2: - fixes x y :: "'a::real_normed_algebra star" - shows "[| x \ Infinitesimal; y \ HFinite |] ==> (y * x) \ Infinitesimal" -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 + "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 -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: "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 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 + 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 -lemma HInfinite_inverse_Infinitesimal: - fixes x :: "'a::real_normed_div_algebra star" - shows "x \ HInfinite ==> inverse x \ Infinitesimal" -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 +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 lemma HInfiniteI: "(\r. r \ \ \ r < hnorm x) \ x \ HInfinite" -by (simp add: HInfinite_def) + by (simp add: HInfinite_def) -lemma HInfiniteD: "\x \ HInfinite; r \ \\ \ r < hnorm x" -by (simp add: HInfinite_def) +lemma HInfiniteD: "x \ HInfinite \ r \ \ \ r < hnorm x" + by (simp add: HInfinite_def) -lemma HInfinite_mult: - fixes x y :: "'a::real_normed_div_algebra star" - shows "[|x \ HInfinite; y \ HInfinite|] ==> (x*y) \ HInfinite" -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: "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 hypreal_add_zero_less_le_mono: "[|r < x; (0::hypreal) \ y|] ==> r < x+y" -by (auto dest: add_less_le_mono) +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) -lemma HInfinite_add_ge_zero: - "[|(x::hypreal) \ HInfinite; 0 \ y; 0 \ x|] ==> (x + y): HInfinite" -by (auto intro!: hypreal_add_zero_less_le_mono - simp add: abs_if add.commute add_nonneg_nonneg HInfinite_def) +lemma HInfinite_add_ge_zero: "x \ HInfinite \ 0 \ y \ 0 \ x \ x + y \ HInfinite" + for x y :: hypreal + by (auto simp: abs_if add.commute HInfinite_def) -lemma HInfinite_add_ge_zero2: - "[|(x::hypreal) \ HInfinite; 0 \ y; 0 \ x|] ==> (y + x): HInfinite" -by (auto intro!: HInfinite_add_ge_zero simp add: add.commute) +lemma HInfinite_add_ge_zero2: "x \ HInfinite \ 0 \ y \ 0 \ x \ y + x \ HInfinite" + for x y :: hypreal + by (auto intro!: HInfinite_add_ge_zero simp add: add.commute) -lemma HInfinite_add_gt_zero: - "[|(x::hypreal) \ HInfinite; 0 < y; 0 < x|] ==> (x + y): HInfinite" -by (blast intro: HInfinite_add_ge_zero order_less_imp_le) +lemma HInfinite_add_gt_zero: "x \ HInfinite \ 0 < y \ 0 < x \ x + y \ HInfinite" + for x y :: hypreal + by (blast intro: HInfinite_add_ge_zero order_less_imp_le) -lemma HInfinite_minus_iff: "(-x \ HInfinite) = (x \ HInfinite)" -by (simp add: HInfinite_def) +lemma HInfinite_minus_iff: "- x \ HInfinite \ x \ HInfinite" + by (simp add: HInfinite_def) -lemma HInfinite_add_le_zero: - "[|(x::hypreal) \ HInfinite; y \ 0; x \ 0|] ==> (x + y): HInfinite" -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 +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 -lemma HInfinite_add_lt_zero: - "[|(x::hypreal) \ HInfinite; y < 0; x < 0|] ==> (x + y): HInfinite" -by (blast intro: HInfinite_add_le_zero order_less_imp_le) +lemma HInfinite_add_lt_zero: "x \ HInfinite \ y < 0 \ x < 0 \ x + y \ HInfinite" + for x y :: hypreal + by (blast intro: HInfinite_add_le_zero order_less_imp_le) lemma HFinite_sum_squares: - fixes a b c :: "'a::real_normed_algebra star" - shows "[|a: HFinite; b: HFinite; c: HFinite|] - ==> a*a + b*b + c*c \ HFinite" -by (auto intro: HFinite_mult HFinite_add) + "a \ HFinite \ b \ HFinite \ c \ HFinite \ a * a + b * b + c * c \ HFinite" + for a b c :: "'a::real_normed_algebra star" + by (auto intro: HFinite_mult HFinite_add) -lemma not_Infinitesimal_not_zero: "x \ Infinitesimal ==> x \ 0" -by auto +lemma not_Infinitesimal_not_zero: "x \ Infinitesimal \ x \ 0" + by auto -lemma not_Infinitesimal_not_zero2: "x \ HFinite - Infinitesimal ==> x \ 0" -by auto +lemma not_Infinitesimal_not_zero2: "x \ HFinite - Infinitesimal \ x \ 0" + by auto lemma HFinite_diff_Infinitesimal_hrabs: - "(x::hypreal) \ HFinite - Infinitesimal ==> \x\ \ HFinite - Infinitesimal" -by blast + "x \ HFinite - Infinitesimal \ \x\ \ HFinite - Infinitesimal" + for x :: hypreal + by blast -lemma hnorm_le_Infinitesimal: - "\e \ Infinitesimal; hnorm x \ e\ \ x \ Infinitesimal" -by (auto simp add: Infinitesimal_def abs_less_iff) +lemma hnorm_le_Infinitesimal: "e \ Infinitesimal \ hnorm x \ e \ x \ Infinitesimal" + by (auto simp: Infinitesimal_def abs_less_iff) -lemma hnorm_less_Infinitesimal: - "\e \ Infinitesimal; hnorm x < e\ \ x \ Infinitesimal" -by (erule hnorm_le_Infinitesimal, erule order_less_imp_le) +lemma hnorm_less_Infinitesimal: "e \ Infinitesimal \ hnorm x < e \ x \ Infinitesimal" + by (erule hnorm_le_Infinitesimal, erule order_less_imp_le) -lemma hrabs_le_Infinitesimal: - "[| e \ Infinitesimal; \x::hypreal\ \ e |] ==> x \ Infinitesimal" -by (erule hnorm_le_Infinitesimal, simp) +lemma hrabs_le_Infinitesimal: "e \ Infinitesimal \ \x\ \ e \ x \ Infinitesimal" + for x :: hypreal + by (erule hnorm_le_Infinitesimal) simp -lemma hrabs_less_Infinitesimal: - "[| e \ Infinitesimal; \x::hypreal\ < e |] ==> x \ Infinitesimal" -by (erule hnorm_less_Infinitesimal, simp) +lemma hrabs_less_Infinitesimal: "e \ Infinitesimal \ \x\ < e \ x \ Infinitesimal" + for x :: hypreal + by (erule hnorm_less_Infinitesimal) simp lemma Infinitesimal_interval: - "[| e \ Infinitesimal; e' \ Infinitesimal; e' < x ; x < e |] - ==> (x::hypreal) \ Infinitesimal" -by (auto simp add: Infinitesimal_def abs_less_iff) + "e \ Infinitesimal \ e' \ Infinitesimal \ e' < x \ x < e \ x \ Infinitesimal" + for x :: hypreal + by (auto simp add: Infinitesimal_def abs_less_iff) lemma Infinitesimal_interval2: - "[| e \ Infinitesimal; e' \ Infinitesimal; - e' \ x ; x \ e |] ==> (x::hypreal) \ Infinitesimal" -by (auto intro: Infinitesimal_interval simp add: order_le_less) + "e \ Infinitesimal \ e' \ Infinitesimal \ e' \ x \ x \ e \ x \ Infinitesimal" + for x :: hypreal + by (auto intro: Infinitesimal_interval simp add: order_le_less) -lemma lemma_Infinitesimal_hyperpow: - "[| (x::hypreal) \ Infinitesimal; 0 < N |] ==> \x pow N\ \ \x\" -apply (unfold Infinitesimal_def) -apply (auto intro!: hyperpow_Suc_le_self2 - simp add: hyperpow_hrabs [symmetric] hypnat_gt_zero_iff2 abs_ge_zero) -done +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 -lemma Infinitesimal_hyperpow: - "[| (x::hypreal) \ Infinitesimal; 0 < N |] ==> x pow N \ Infinitesimal" -apply (rule hrabs_le_Infinitesimal) -apply (rule_tac [2] lemma_Infinitesimal_hyperpow, auto) -done +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 lemma hrealpow_hyperpow_Infinitesimal_iff: - "(x ^ n \ Infinitesimal) = (x pow (hypnat_of_nat n) \ Infinitesimal)" -by (simp only: hyperpow_hypnat_of_nat) + "(x ^ n \ Infinitesimal) \ x pow (hypnat_of_nat n) \ Infinitesimal" + by (simp only: hyperpow_hypnat_of_nat) -lemma Infinitesimal_hrealpow: - "[| (x::hypreal) \ Infinitesimal; 0 < n |] ==> x ^ n \ Infinitesimal" -by (simp add: hrealpow_hyperpow_Infinitesimal_iff Infinitesimal_hyperpow) +lemma Infinitesimal_hrealpow: "x \ Infinitesimal \ 0 < n \ x ^ n \ Infinitesimal" + for x :: hypreal + by (simp add: hrealpow_hyperpow_Infinitesimal_iff Infinitesimal_hyperpow) lemma not_Infinitesimal_mult: - fixes x y :: "'a::real_normed_div_algebra star" - shows "[| x \ Infinitesimal; y \ Infinitesimal|] ==> (x*y) \Infinitesimal" -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 (no_asm_simp)) -done + "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 -lemma Infinitesimal_mult_disj: - fixes x y :: "'a::real_normed_div_algebra star" - shows "x*y \ Infinitesimal ==> x \ Infinitesimal | y \ Infinitesimal" -apply (rule ccontr) -apply (drule de_Morgan_disj [THEN iffD1]) -apply (fast dest: not_Infinitesimal_mult) -done +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 -lemma HFinite_Infinitesimal_not_zero: "x \ HFinite-Infinitesimal ==> x \ 0" -by blast +lemma HFinite_Infinitesimal_not_zero: "x \ HFinite-Infinitesimal \ x \ 0" + by blast lemma HFinite_Infinitesimal_diff_mult: - fixes x y :: "'a::real_normed_div_algebra star" - shows "[| x \ HFinite - Infinitesimal; - y \ HFinite - Infinitesimal - |] ==> (x*y) \ HFinite - Infinitesimal" -apply clarify -apply (blast dest: HFinite_mult not_Infinitesimal_mult) -done + "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 -lemma Infinitesimal_subset_HFinite: - "Infinitesimal \ HFinite" -apply (simp add: Infinitesimal_def HFinite_def, auto) -apply (rule_tac x = 1 in bexI, auto) -done +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 -lemma Infinitesimal_star_of_mult: - fixes x :: "'a::real_normed_algebra star" - shows "x \ Infinitesimal ==> x * star_of r \ Infinitesimal" -by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult]) +lemma Infinitesimal_star_of_mult: "x \ Infinitesimal \ x * star_of r \ Infinitesimal" + for x :: "'a::real_normed_algebra star" + by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult]) -lemma Infinitesimal_star_of_mult2: - fixes x :: "'a::real_normed_algebra star" - shows "x \ Infinitesimal ==> star_of r * x \ Infinitesimal" -by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult2]) +lemma Infinitesimal_star_of_mult2: "x \ Infinitesimal \ star_of r * x \ Infinitesimal" + for x :: "'a::real_normed_algebra star" + by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult2]) -subsection\The Infinitely Close Relation\ +subsection \The Infinitely Close Relation\ -lemma mem_infmal_iff: "(x \ Infinitesimal) = (x \ 0)" -by (simp add: Infinitesimal_def approx_def) +lemma mem_infmal_iff: "x \ Infinitesimal \ x \ 0" + by (simp add: Infinitesimal_def approx_def) -lemma approx_minus_iff: " (x \ y) = (x - y \ 0)" -by (simp add: approx_def) +lemma approx_minus_iff: "x \ y \ x - y \ 0" + by (simp add: approx_def) -lemma approx_minus_iff2: " (x \ y) = (-y + x \ 0)" -by (simp add: approx_def add.commute) +lemma approx_minus_iff2: "x \ y \ - y + x \ 0" + by (simp add: approx_def add.commute) lemma approx_refl [iff]: "x \ x" -by (simp add: approx_def Infinitesimal_def) + by (simp add: approx_def Infinitesimal_def) -lemma hypreal_minus_distrib1: "-(y + -(x::'a::ab_group_add)) = x + -y" -by (simp add: add.commute) +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" -apply (simp add: approx_def) -apply (drule Infinitesimal_minus_iff [THEN iffD2]) -apply simp -done +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: "x \ y \ y \ z \ x \ z" + apply (simp add: approx_def) + apply (drule (1) Infinitesimal_add) + apply simp + done -lemma approx_trans2: "[| r \ x; s \ x |] ==> r \ s" -by (blast intro: approx_sym approx_trans) +lemma approx_trans2: "r \ x \ s \ x \ r \ s" + by (blast intro: approx_sym approx_trans) -lemma approx_trans3: "[| x \ r; x \ s|] ==> r \ s" -by (blast intro: approx_sym approx_trans) +lemma approx_trans3: "x \ r \ x \ s \ r \ s" + by (blast intro: approx_sym approx_trans) -lemma approx_reorient: "(x \ y) = (y \ x)" -by (blast intro: approx_sym) +lemma approx_reorient: "x \ y \ y \ x" + by (blast intro: approx_sym) -(*reorientation simplification procedure: reorients (polymorphic) - 0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*) +text \Reorientation simplification procedure: reorients (polymorphic) + \0 = x\, \1 = x\, \nnn = x\ provided \x\ isn't \0\, \1\ or a numeral.\ simproc_setup approx_reorient_simproc ("0 \ x" | "1 \ y" | "numeral w \ z" | "- 1 \ y" | "- numeral w \ r") = \ @@ -658,1133 +593,1067 @@ in proc end \ -lemma Infinitesimal_approx_minus: "(x-y \ Infinitesimal) = (x \ y)" -by (simp add: approx_minus_iff [symmetric] mem_infmal_iff) +lemma Infinitesimal_approx_minus: "x - y \ Infinitesimal \ x \ y" + by (simp add: approx_minus_iff [symmetric] mem_infmal_iff) -lemma approx_monad_iff: "(x \ y) = (monad(x)=monad(y))" -apply (simp add: monad_def) -apply (auto dest: approx_sym elim!: approx_trans equalityCE) -done +lemma approx_monad_iff: "x \ y \ monad x = monad y" + by (auto simp add: monad_def dest: approx_sym elim!: approx_trans equalityCE) -lemma Infinitesimal_approx: - "[| x \ Infinitesimal; y \ Infinitesimal |] ==> x \ y" -apply (simp add: mem_infmal_iff) -apply (blast intro: approx_trans approx_sym) -done +lemma Infinitesimal_approx: "x \ Infinitesimal \ y \ Infinitesimal \ x \ y" + apply (simp add: mem_infmal_iff) + apply (blast intro: approx_trans approx_sym) + done -lemma approx_add: "[| a \ b; c \ d |] ==> a+c \ b+d" +lemma approx_add: "a \ b \ c \ d \ a + c \ b + d" proof (unfold approx_def) assume inf: "a - b \ Infinitesimal" "c - d \ Infinitesimal" have "a + c - (b + d) = (a - b) + (c - d)" by simp - also have "... \ Infinitesimal" using inf by (rule Infinitesimal_add) + also have "... \ Infinitesimal" + using inf by (rule Infinitesimal_add) finally show "a + c - (b + d) \ Infinitesimal" . 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 +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 -lemma approx_minus2: "-a \ -b ==> a \ b" -by (auto dest: approx_minus) +lemma approx_minus2: "- a \ - b \ a \ b" + by (auto dest: approx_minus) -lemma approx_minus_cancel [simp]: "(-a \ -b) = (a \ b)" -by (blast intro: approx_minus approx_minus2) +lemma approx_minus_cancel [simp]: "- a \ - b \ a \ b" + by (blast intro: approx_minus approx_minus2) -lemma approx_add_minus: "[| a \ b; c \ d |] ==> a + -c \ b + -d" -by (blast intro!: approx_add approx_minus) +lemma approx_add_minus: "a \ b \ c \ d \ a + - c \ b + - d" + by (blast intro!: approx_add approx_minus) -lemma approx_diff: "[| a \ b; c \ d |] ==> a - c \ b - d" +lemma approx_diff: "a \ b \ c \ d \ a - c \ b - d" using approx_add [of a b "- c" "- d"] by simp -lemma approx_mult1: - fixes a b c :: "'a::real_normed_algebra star" - shows "[| a \ b; c: HFinite|] ==> a*c \ b*c" -by (simp add: approx_def Infinitesimal_HFinite_mult - left_diff_distrib [symmetric]) +lemma approx_mult1: "a \ b \ c \ HFinite \ a * c \ b * c" + for a b c :: "'a::real_normed_algebra star" + by (simp add: approx_def Infinitesimal_HFinite_mult left_diff_distrib [symmetric]) + +lemma approx_mult2: "a \ b \ c \ HFinite \ c * a \ c * b" + for a b c :: "'a::real_normed_algebra star" + by (simp add: approx_def Infinitesimal_HFinite_mult2 right_diff_distrib [symmetric]) -lemma approx_mult2: - fixes a b c :: "'a::real_normed_algebra star" - shows "[|a \ b; c: HFinite|] ==> c*a \ c*b" -by (simp add: approx_def Infinitesimal_HFinite_mult2 - right_diff_distrib [symmetric]) +lemma approx_mult_subst: "u \ v * x \ x \ y \ v \ HFinite \ u \ v * y" + for u v x y :: "'a::real_normed_algebra star" + by (blast intro: approx_mult2 approx_trans) -lemma approx_mult_subst: - fixes u v x y :: "'a::real_normed_algebra star" - shows "[|u \ v*x; x \ y; v \ HFinite|] ==> u \ v*y" -by (blast intro: approx_mult2 approx_trans) - -lemma approx_mult_subst2: - fixes u v x y :: "'a::real_normed_algebra star" - shows "[| u \ x*v; x \ y; v \ HFinite |] ==> u \ y*v" -by (blast intro: approx_mult1 approx_trans) +lemma approx_mult_subst2: "u \ x * v \ x \ y \ v \ HFinite \ u \ y * v" + for u v x y :: "'a::real_normed_algebra star" + by (blast intro: approx_mult1 approx_trans) -lemma approx_mult_subst_star_of: - fixes u x y :: "'a::real_normed_algebra star" - shows "[| u \ x*star_of v; x \ y |] ==> u \ y*star_of v" -by (auto intro: approx_mult_subst2) +lemma approx_mult_subst_star_of: "u \ x * star_of v \ x \ y \ u \ y * star_of v" + for u x y :: "'a::real_normed_algebra star" + by (auto intro: approx_mult_subst2) -lemma approx_eq_imp: "a = b ==> a \ b" -by (simp add: approx_def) +lemma approx_eq_imp: "a = b \ a \ b" + by (simp add: approx_def) -lemma Infinitesimal_minus_approx: "x \ Infinitesimal ==> -x \ x" -by (blast intro: Infinitesimal_minus_iff [THEN iffD2] - mem_infmal_iff [THEN iffD1] approx_trans2) +lemma Infinitesimal_minus_approx: "x \ Infinitesimal \ - x \ x" + by (blast intro: Infinitesimal_minus_iff [THEN iffD2] mem_infmal_iff [THEN iffD1] approx_trans2) -lemma bex_Infinitesimal_iff: "(\y \ Infinitesimal. x - z = y) = (x \ z)" -by (simp add: approx_def) +lemma bex_Infinitesimal_iff: "(\y \ Infinitesimal. x - z = y) \ x \ z" + by (simp add: approx_def) -lemma bex_Infinitesimal_iff2: "(\y \ Infinitesimal. x = z + y) = (x \ z)" -by (force simp add: bex_Infinitesimal_iff [symmetric]) +lemma bex_Infinitesimal_iff2: "(\y \ Infinitesimal. x = z + y) \ x \ z" + 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 +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 -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 +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 -lemma Infinitesimal_add_approx_self2: "y \ Infinitesimal ==> x \ y + x" -by (auto dest: Infinitesimal_add_approx_self simp add: add.commute) +lemma Infinitesimal_add_approx_self2: "y \ Infinitesimal \ x \ y + x" + by (auto dest: Infinitesimal_add_approx_self simp add: add.commute) -lemma Infinitesimal_add_minus_approx_self: "y \ Infinitesimal ==> x \ x + -y" -by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2]) +lemma Infinitesimal_add_minus_approx_self: "y \ Infinitesimal \ x \ x + - y" + 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 +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 -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 +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 -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" + apply (drule approx_minus_iff [THEN iffD1]) + apply (simp add: approx_minus_iff [symmetric] ac_simps) + done -lemma approx_add_right_cancel: "b + d \ c + d ==> b \ c" -apply (rule approx_add_left_cancel) -apply (simp add: add.commute) -done +lemma approx_add_right_cancel: "b + d \ c + d \ b \ c" + apply (rule approx_add_left_cancel) + apply (simp add: add.commute) + done -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 +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 -lemma approx_add_mono2: "b \ c ==> b + a \ c + a" -by (simp add: add.commute approx_add_mono1) +lemma approx_add_mono2: "b \ c \ b + a \ c + a" + by (simp add: add.commute approx_add_mono1) -lemma approx_add_left_iff [simp]: "(a + b \ a + c) = (b \ c)" -by (fast elim: approx_add_left_cancel approx_add_mono1) +lemma approx_add_left_iff [simp]: "a + b \ a + c \ b \ c" + by (fast elim: approx_add_left_cancel approx_add_mono1) -lemma approx_add_right_iff [simp]: "(b + a \ c + a) = (b \ c)" -by (simp add: add.commute) +lemma approx_add_right_iff [simp]: "b + a \ c + a \ b \ c" + 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 +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 -lemma approx_star_of_HFinite: "x \ star_of D ==> x \ HFinite" -by (rule approx_sym [THEN [2] approx_HFinite], auto) +lemma approx_star_of_HFinite: "x \ star_of D \ x \ HFinite" + by (rule approx_sym [THEN [2] approx_HFinite], auto) -lemma approx_mult_HFinite: - fixes a b c d :: "'a::real_normed_algebra star" - shows "[|a \ b; c \ d; b: HFinite; d: HFinite|] ==> a*c \ b*d" -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 +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 -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 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 +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 -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_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 -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 +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_mult_star_of: - fixes a c :: "'a::real_normed_algebra star" - shows "[|a \ star_of b; c \ star_of d |] - ==> a*c \ star_of b*star_of d" -by (blast intro!: approx_mult_HFinite approx_star_of_HFinite HFinite_star_of) +lemma approx_mult_SReal1: "a \ \ \ x \ 0 \ x * a \ 0" + for a x :: hypreal + by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult1) -lemma approx_SReal_mult_cancel_zero: - "[| (a::hypreal) \ \; a \ 0; a*x \ 0 |] ==> x \ 0" -apply (drule Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) -apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) -done +lemma approx_mult_SReal2: "a \ \ \ x \ 0 \ a * x \ 0" + for a x :: hypreal + by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult2) -lemma approx_mult_SReal1: "[| (a::hypreal) \ \; x \ 0 |] ==> x*a \ 0" -by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult1) - -lemma approx_mult_SReal2: "[| (a::hypreal) \ \; x \ 0 |] ==> a*x \ 0" -by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult2) +lemma approx_mult_SReal_zero_cancel_iff [simp]: "a \ \ \ a \ 0 \ a * x \ 0 \ x \ 0" + for a x :: hypreal + by (blast intro: approx_SReal_mult_cancel_zero approx_mult_SReal2) -lemma approx_mult_SReal_zero_cancel_iff [simp]: - "[|(a::hypreal) \ \; a \ 0 |] ==> (a*x \ 0) = (x \ 0)" -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: - "[| (a::hypreal) \ \; a \ 0; a* w \ a*z |] ==> w \ z" -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_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) -lemma approx_SReal_mult_cancel_iff1 [simp]: - "[| (a::hypreal) \ \; a \ 0|] ==> (a* w \ a*z) = (w \ z)" -by (auto intro!: approx_mult2 SReal_subset_HFinite [THEN subsetD] - intro: approx_SReal_mult_cancel) +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: "[| (z::hypreal) \ f; f \ g; g \ z |] ==> f \ z" -apply (simp add: bex_Infinitesimal_iff2 [symmetric], auto) -apply (rule_tac x = "g+y-z" in bexI) -apply (simp (no_asm)) -apply (rule Infinitesimal_interval2) -apply (rule_tac [2] Infinitesimal_zero, auto) -done - -lemma approx_hnorm: - fixes x y :: "'a::real_normed_vector star" - shows "x \ y \ hnorm x \ hnorm y" +lemma approx_hnorm: "x \ y \ hnorm x \ hnorm y" + for x y :: "'a::real_normed_vector star" proof (unfold approx_def) assume "x - y \ Infinitesimal" - hence 1: "hnorm (x - y) \ Infinitesimal" + then have "hnorm (x - y) \ Infinitesimal" by (simp only: Infinitesimal_hnorm_iff) - moreover have 2: "(0::real star) \ Infinitesimal" + moreover have "(0::real star) \ Infinitesimal" by (rule Infinitesimal_zero) - moreover have 3: "0 \ \hnorm x - hnorm y\" + moreover have "0 \ \hnorm x - hnorm y\" by (rule abs_ge_zero) - moreover have 4: "\hnorm x - hnorm y\ \ hnorm (x - y)" + moreover have "\hnorm x - hnorm y\ \ hnorm (x - y)" by (rule hnorm_triangle_ineq3) ultimately have "\hnorm x - hnorm y\ \ Infinitesimal" by (rule Infinitesimal_interval2) - thus "hnorm x - hnorm y \ Infinitesimal" + then show "hnorm x - hnorm y \ Infinitesimal" by (simp only: Infinitesimal_hrabs_iff) qed -subsection\Zero is the Only Infinitesimal that is also a Real\ +subsection \Zero is the Only Infinitesimal that is also a Real\ -lemma Infinitesimal_less_SReal: - "[| (x::hypreal) \ \; y \ Infinitesimal; 0 < x |] ==> y < x" -apply (simp add: Infinitesimal_def) -apply (rule abs_ge_self [THEN order_le_less_trans], auto) -done +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 -lemma Infinitesimal_less_SReal2: - "(y::hypreal) \ Infinitesimal ==> \r \ Reals. 0 < r --> y < r" -by (blast intro: Infinitesimal_less_SReal) +lemma Infinitesimal_less_SReal2: "y \ Infinitesimal \ \r \ Reals. 0 < r \ y < r" + for y :: hypreal + by (blast intro: Infinitesimal_less_SReal) -lemma SReal_not_Infinitesimal: - "[| 0 < y; (y::hypreal) \ \|] ==> y \ Infinitesimal" -apply (simp add: Infinitesimal_def) -apply (auto simp add: abs_if) -done +lemma SReal_not_Infinitesimal: "0 < y \ y \ \ ==> y \ Infinitesimal" + for y :: hypreal + apply (simp add: Infinitesimal_def) + apply (auto simp add: abs_if) + done -lemma SReal_minus_not_Infinitesimal: - "[| y < 0; (y::hypreal) \ \ |] ==> y \ Infinitesimal" -apply (subst Infinitesimal_minus_iff [symmetric]) -apply (rule SReal_not_Infinitesimal, auto) -done +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 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 + 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 -lemma SReal_Infinitesimal_zero: - "[| (x::hypreal) \ \; x \ Infinitesimal|] ==> x = 0" -by (cut_tac SReal_Int_Infinitesimal_zero, blast) +lemma SReal_Infinitesimal_zero: "x \ \ \ x \ Infinitesimal \ x = 0" + for x :: hypreal + using SReal_Int_Infinitesimal_zero by blast -lemma SReal_HFinite_diff_Infinitesimal: - "[| (x::hypreal) \ \; x \ 0 |] ==> x \ HFinite - Infinitesimal" -by (auto dest: SReal_Infinitesimal_zero SReal_subset_HFinite [THEN subsetD]) +lemma SReal_HFinite_diff_Infinitesimal: "x \ \ \ x \ 0 \ x \ HFinite - Infinitesimal" + for x :: hypreal + by (auto dest: SReal_Infinitesimal_zero SReal_subset_HFinite [THEN subsetD]) lemma hypreal_of_real_HFinite_diff_Infinitesimal: - "hypreal_of_real x \ 0 ==> hypreal_of_real x \ HFinite - Infinitesimal" -by (rule SReal_HFinite_diff_Infinitesimal, auto) + "hypreal_of_real x \ 0 \ hypreal_of_real x \ HFinite - Infinitesimal" + 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 +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 -lemma star_of_HFinite_diff_Infinitesimal: - "x \ 0 ==> star_of x \ HFinite - Infinitesimal" -by simp +lemma star_of_HFinite_diff_Infinitesimal: "x \ 0 \ star_of x \ HFinite - Infinitesimal" + by simp lemma numeral_not_Infinitesimal [simp]: - "numeral w \ (0::hypreal) ==> (numeral w :: hypreal) \ Infinitesimal" -by (fast dest: Reals_numeral [THEN SReal_Infinitesimal_zero]) + "numeral w \ (0::hypreal) \ (numeral w :: hypreal) \ Infinitesimal" + by (fast dest: Reals_numeral [THEN SReal_Infinitesimal_zero]) -(*again: 1 is a special case, but not 0 this time*) +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 + apply (simp only: star_one_def star_of_Infinitesimal_iff_0) + apply simp + done -lemma approx_SReal_not_zero: - "[| (y::hypreal) \ \; x \ y; y\ 0 |] ==> x \ 0" -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 +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 lemma HFinite_diff_Infinitesimal_approx: - "[| x \ y; y \ HFinite - Infinitesimal |] - ==> x \ HFinite - Infinitesimal" -apply (auto intro: approx_sym [THEN [2] approx_HFinite] - simp add: mem_infmal_iff) -apply (drule approx_trans3, assumption) -apply (blast dest: approx_sym) -done + "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 -(*The premise y\0 is essential; otherwise x/y =0 and we lose the - HFinite premise.*) +text \The premise \y \ 0\ is essential; otherwise \x / y = 0\ and we lose the + \HFinite\ premise.\ lemma Infinitesimal_ratio: - fixes x y :: "'a::{real_normed_div_algebra,field} star" - shows "[| y \ 0; y \ Infinitesimal; x/y \ HFinite |] - ==> x \ Infinitesimal" -apply (drule Infinitesimal_HFinite_mult2, assumption) -apply (simp add: divide_inverse mult.assoc) -done + "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 + +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 + + +section \Standard Part Theorem\ -lemma Infinitesimal_SReal_divide: - "[| (x::hypreal) \ Infinitesimal; y \ \ |] ==> x/y \ Infinitesimal" -apply (simp add: divide_inverse) -apply (auto intro!: Infinitesimal_HFinite_mult - dest!: Reals_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) -done +text \ + Every finite \x \ R*\ is infinitely close to a unique real number + (i.e. a member of \Reals\). +\ -(*------------------------------------------------------------------ - Standard Part Theorem: Every finite x: R* is infinitely - close to a unique real number (i.e a member of Reals) - ------------------------------------------------------------------*) -subsection\Uniqueness: Two Infinitely Close Reals are Equal\ +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 +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 -lemma SReal_approx_iff: - "[|(x::hypreal) \ \; y \ \|] ==> (x \ y) = (x = y)" -apply auto -apply (simp add: approx_def) -apply (drule (1) Reals_diff) -apply (drule (1) SReal_Infinitesimal_zero) -apply simp -done +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 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))" + apply (unfold star_numeral_def) + apply (rule star_of_approx_iff) + done -(*And also for 0 \ #nn and 1 \ #nn, #nn \ 0 and #nn \ 1.*) +text \And also for \0 \ #nn\ and \1 \ #nn\, \#nn \ 0\ and \#nn \ 1\.\ lemma [simp]: - "(numeral w \ (0::'a::{numeral,real_normed_vector} star)) = - (numeral w = (0::'a))" - "((0::'a::{numeral,real_normed_vector} star) \ numeral w) = - (numeral w = (0::'a))" - "(numeral w \ (1::'b::{numeral,one,real_normed_vector} star)) = - (numeral w = (1::'b))" - "((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) -by (auto intro: sym) + "(numeral w \ (0::'a::{numeral,real_normed_vector} star)) = (numeral w = (0::'a))" + "((0::'a::{numeral,real_normed_vector} star) \ numeral w) = (numeral w = (0::'a))" + "(numeral w \ (1::'b::{numeral,one,real_normed_vector} star)) = (numeral w = (1::'b))" + "((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 -lemma star_of_approx_numeral_iff [simp]: - "(star_of k \ numeral w) = (k = numeral w)" -by (subst star_of_approx_iff [symmetric], auto) +lemma star_of_approx_numeral_iff [simp]: "star_of k \ numeral w \ k = numeral w" + by (subst star_of_approx_iff [symmetric]) auto -lemma star_of_approx_zero_iff [simp]: "(star_of k \ 0) = (k = 0)" -by (simp_all add: star_of_approx_iff [symmetric]) +lemma star_of_approx_zero_iff [simp]: "star_of k \ 0 \ k = 0" + by (simp_all add: star_of_approx_iff [symmetric]) -lemma star_of_approx_one_iff [simp]: "(star_of k \ 1) = (k = 1)" -by (simp_all add: star_of_approx_iff [symmetric]) +lemma star_of_approx_one_iff [simp]: "star_of k \ 1 \ k = 1" + by (simp_all add: star_of_approx_iff [symmetric]) -lemma approx_unique_real: - "[| (r::hypreal) \ \; s \ \; r \ x; s \ x|] ==> r = s" -by (blast intro: SReal_approx_iff [THEN iffD1] approx_trans2) +lemma approx_unique_real: "r \ \ \ s \ \ \ r \ x \ s \ x \ r = s" + for r s :: hypreal + by (blast intro: SReal_approx_iff [THEN iffD1] approx_trans2) -subsection\Existence of Unique Real Infinitely Close\ +subsection \Existence of Unique Real Infinitely Close\ -subsubsection\Lifting of the Ub and Lub Properties\ +subsubsection \Lifting of the Ub and Lub Properties\ -lemma hypreal_of_real_isUb_iff: - "(isUb \ (hypreal_of_real ` Q) (hypreal_of_real Y)) = - (isUb (UNIV :: real set) Q Y)" -by (simp add: isUb_def setle_def) +lemma hypreal_of_real_isUb_iff: "isUb \ (hypreal_of_real ` Q) (hypreal_of_real Y) = isUb UNIV Q Y" + 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 :: real set) Q Y" -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_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 :: real set) Q Y - ==> isLub \ (hypreal_of_real ` Q) (hypreal_of_real Y)" -apply (auto simp add: isLub_def leastP_def hypreal_of_real_isUb_iff setge_def) -by (metis SReal_iff hypreal_of_real_isUb_iff isUbD2a star_of_le) +lemma hypreal_of_real_isLub2: "isLub UNIV Q Y \ isLub \ (hypreal_of_real ` Q) (hypreal_of_real Y)" + 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)" -by (blast intro: hypreal_of_real_isLub1 hypreal_of_real_isLub2) + "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) -lemma lemma_isUb_hypreal_of_real: - "isUb \ P Y ==> \Yo. isUb \ P (hypreal_of_real Yo)" -by (auto simp add: SReal_iff isUb_def) +lemma lemma_isUb_hypreal_of_real: "isUb \ P Y \ \Yo. isUb \ P (hypreal_of_real Yo)" + by (auto simp add: SReal_iff isUb_def) + +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_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 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: - "[| 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 +text \Lemmas about lubs.\ -(* lemma 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_ub: - "(x::hypreal) \ HFinite ==> \u. isUb \ {s. s \ \ & s < x} u" -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_nonempty: - "(x::hypreal) \ HFinite ==> \y. y \ {s. s \ \ & s < x}" -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::hypreal) \ HFinite ==> \t. isLub \ {s. s \ \ & s < x} t" -by (blast intro!: SReal_complete lemma_st_part_ub lemma_st_part_nonempty Collect_restrict) +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_le1: - "[| (x::hypreal) \ HFinite; isLub \ {s. s \ \ & s < x} t; - r \ \; 0 < r |] ==> x \ t + r" -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 + "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 -lemma hypreal_setle_less_trans: - "[| S *<= (x::hypreal); x < y |] ==> S *<= y" -apply (simp add: setle_def) -apply (auto dest!: bspec order_le_less_trans intro: order_less_imp_le) -done +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 -lemma hypreal_gt_isUb: - "[| isUb R S (x::hypreal); x < y; y \ R |] ==> isUb R S y" -apply (simp add: isUb_def) -apply (blast intro: hypreal_setle_less_trans) -done +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 -lemma lemma_st_part_gt_ub: - "[| (x::hypreal) \ HFinite; x < y; y \ \ |] - ==> isUb \ {s. s \ \ & s < x} y" -by (auto dest: order_less_trans intro: order_less_imp_le intro!: isUbI setleI) +lemma lemma_st_part_gt_ub: "x \ HFinite \ x < y \ y \ \ \ isUb \ {s. s \ \ \ s < x} y" + for x y :: hypreal + by (auto dest: order_less_trans intro: order_less_imp_le intro!: isUbI setleI) -lemma lemma_minus_le_zero: "t \ t + -r ==> r \ (0::hypreal)" -apply (drule_tac c = "-t" in add_left_mono) -apply (auto simp add: add.assoc [symmetric]) -done +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 lemma lemma_st_part_le2: - "[| (x::hypreal) \ HFinite; - isLub \ {s. s \ \ & s < x} t; - r \ \; 0 < r |] - ==> t + -r \ x" -apply (frule isLubD1a) -apply (rule ccontr, drule linorder_not_le [THEN iffD1]) -apply (drule Reals_minus, drule_tac a = t in Reals_add, assumption) -apply (drule lemma_st_part_gt_ub, assumption+) -apply (drule isLub_le_isUb, assumption) -apply (drule lemma_minus_le_zero) -apply (auto dest: order_less_le_trans) -done + "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ r \ \ \ 0 < r \ t + -r \ x" + for x r t :: hypreal + apply (frule isLubD1a) + apply (rule ccontr, drule linorder_not_le [THEN iffD1]) + apply (drule Reals_minus, drule_tac a = t in Reals_add, assumption) + apply (drule lemma_st_part_gt_ub, assumption+) + apply (drule isLub_le_isUb, assumption) + apply (drule lemma_minus_le_zero) + apply (auto dest: order_less_le_trans) + done lemma lemma_st_part1a: - "[| (x::hypreal) \ HFinite; - isLub \ {s. s \ \ & s < x} t; - r \ \; 0 < r |] - ==> x + -t \ r" -apply (subgoal_tac "x \ t+r") -apply (auto intro: lemma_st_part_le1) -done + "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 lemma lemma_st_part2a: - "[| (x::hypreal) \ HFinite; - isLub \ {s. s \ \ & s < x} t; - r \ \; 0 < r |] - ==> -(x + -t) \ r" -apply (subgoal_tac "(t + -r \ x)") -apply simp -apply (rule lemma_st_part_le2) -apply auto -done + "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ r \ \ \ 0 < r \ - (x + -t) \ r" + for x r t :: hypreal + apply (subgoal_tac "(t + -r \ x)") + apply simp + apply (rule lemma_st_part_le2) + apply auto + done -lemma lemma_SReal_ub: - "(x::hypreal) \ \ ==> isUb \ {s. s \ \ & s < x} x" -by (auto intro: isUbI setleI order_less_imp_le) +lemma lemma_SReal_ub: "x \ \ \ isUb \ {s. s \ \ \ s < x} x" + for x :: hypreal + by (auto intro: isUbI setleI order_less_imp_le) -lemma lemma_SReal_lub: - "(x::hypreal) \ \ ==> isLub \ {s. s \ \ & s < x} x" -apply (auto intro!: isLubI2 lemma_SReal_ub setgeI) -apply (frule isUbD2a) -apply (rule_tac x = x and y = y in linorder_cases) -apply (auto intro!: order_less_imp_le) -apply (drule SReal_dense, assumption, assumption, safe) -apply (drule_tac y = r in isUbD) -apply (auto dest: order_less_le_trans) -done +lemma lemma_SReal_lub: "x \ \ \ isLub \ {s. s \ \ \ s < x} x" + for x :: hypreal + apply (auto intro!: isLubI2 lemma_SReal_ub setgeI) + apply (frule isUbD2a) + apply (rule_tac x = x and y = y in linorder_cases) + apply (auto intro!: order_less_imp_le) + apply (drule SReal_dense, assumption, assumption, safe) + apply (drule_tac y = r in isUbD) + apply (auto dest: order_less_le_trans) + done lemma lemma_st_part_not_eq1: - "[| (x::hypreal) \ HFinite; - isLub \ {s. s \ \ & s < x} t; - r \ \; 0 < r |] - ==> x + -t \ r" -apply auto -apply (frule isLubD1a [THEN Reals_minus]) -using Reals_add_cancel [of x "- t"] apply simp -apply (drule_tac x = x in lemma_SReal_lub) -apply (drule isLub_unique, assumption, auto) -done + "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ r \ \ \ 0 < r \ x + - t \ r" + for x r t :: hypreal + apply auto + apply (frule isLubD1a [THEN Reals_minus]) + using Reals_add_cancel [of x "- t"] apply simp + apply (drule_tac x = x in lemma_SReal_lub) + apply (drule isLub_unique, assumption, auto) + done lemma lemma_st_part_not_eq2: - "[| (x::hypreal) \ HFinite; - isLub \ {s. s \ \ & s < x} t; - r \ \; 0 < r |] - ==> -(x + -t) \ r" -apply (auto) -apply (frule isLubD1a) -using Reals_add_cancel [of "- x" t] apply simp -apply (drule_tac x = x in lemma_SReal_lub) -apply (drule isLub_unique, assumption, auto) -done + "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ r \ \ \ 0 < r \ - (x + -t) \ r" + for x r t :: hypreal + apply (auto) + apply (frule isLubD1a) + using Reals_add_cancel [of "- x" t] apply simp + apply (drule_tac x = x in lemma_SReal_lub) + apply (drule isLub_unique, assumption, auto) + done lemma lemma_st_part_major: - "[| (x::hypreal) \ HFinite; - isLub \ {s. s \ \ & s < x} t; - r \ \; 0 < r |] - ==> \x - t\ < r" -apply (frule lemma_st_part1a) -apply (frule_tac [4] lemma_st_part2a, auto) -apply (drule order_le_imp_less_or_eq)+ -apply auto -using lemma_st_part_not_eq2 apply fastforce -using lemma_st_part_not_eq1 apply fastforce -done + "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ r \ \ \ 0 < r \ \x - t\ < r" + for x r t :: hypreal + apply (frule lemma_st_part1a) + apply (frule_tac [4] lemma_st_part2a, auto) + apply (drule order_le_imp_less_or_eq)+ + apply auto + using lemma_st_part_not_eq2 apply fastforce + using lemma_st_part_not_eq1 apply fastforce + done lemma lemma_st_part_major2: - "[| (x::hypreal) \ HFinite; isLub \ {s. s \ \ & s < x} t |] - ==> \r \ Reals. 0 < r --> \x - t\ < r" -by (blast dest!: lemma_st_part_major) - + "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ \r \ Reals. 0 < r \ \x - t\ < r" + for x t :: hypreal + by (blast dest!: lemma_st_part_major) -text\Existence of real and Standard Part Theorem\ -lemma lemma_st_part_Ex: - "(x::hypreal) \ HFinite - ==> \t \ Reals. \r \ Reals. 0 < r --> \x - t\ < r" -apply (frule lemma_st_part_lub, safe) -apply (frule isLubD1a) -apply (blast dest: lemma_st_part_major2) -done +text\Existence of real and Standard Part Theorem.\ + +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 -lemma st_part_Ex: - "(x::hypreal) \ HFinite ==> \t \ Reals. x \ t" -apply (simp add: approx_def Infinitesimal_def) -apply (drule lemma_st_part_Ex, auto) -done +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 -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 +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 -subsection\Finite, Infinite and Infinitesimal\ + +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 + apply (simp add: HFinite_def HInfinite_def) + apply (auto dest: order_less_trans) + done lemma HFinite_not_HInfinite: - assumes x: "x \ HFinite" shows "x \ HInfinite" + assumes x: "x \ HFinite" + shows "x \ HInfinite" proof assume x': "x \ HInfinite" with x have "x \ HFinite \ HInfinite" by blast - thus False by auto + then show False by auto qed -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 +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 -lemma HInfinite_HFinite_disj: "x \ HInfinite | x \ HFinite" -by (blast intro: not_HFinite_HInfinite) +lemma HInfinite_HFinite_disj: "x \ HInfinite \ x \ HFinite" + by (blast intro: not_HFinite_HInfinite) -lemma HInfinite_HFinite_iff: "(x \ HInfinite) = (x \ HFinite)" -by (blast dest: HFinite_not_HInfinite not_HFinite_HInfinite) +lemma HInfinite_HFinite_iff: "x \ HInfinite \ x \ HFinite" + by (blast dest: HFinite_not_HInfinite not_HFinite_HInfinite) -lemma HFinite_HInfinite_iff: "(x \ HFinite) = (x \ HInfinite)" -by (simp add: HInfinite_HFinite_iff) +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) + "x \ Infinitesimal \ x \ HInfinite \ x \ HFinite - Infinitesimal" + by (fast intro: not_HFinite_HInfinite) -lemma HFinite_inverse: - fixes x :: "'a::real_normed_div_algebra star" - shows "[| x \ HFinite; x \ Infinitesimal |] ==> inverse x \ HFinite" -apply (subgoal_tac "x \ 0") -apply (cut_tac x = "inverse x" in HInfinite_HFinite_disj) -apply (auto dest!: HInfinite_inverse_Infinitesimal - simp add: nonzero_inverse_inverse_eq) -done +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 -lemma HFinite_inverse2: - fixes x :: "'a::real_normed_div_algebra star" - shows "x \ HFinite - Infinitesimal ==> inverse x \ HFinite" -by (blast intro: HFinite_inverse) +lemma HFinite_inverse2: "x \ HFinite - Infinitesimal \ inverse x \ HFinite" + for x :: "'a::real_normed_div_algebra star" + by (blast intro: HFinite_inverse) -(* stronger statement possible in fact *) -lemma Infinitesimal_inverse_HFinite: - fixes x :: "'a::real_normed_div_algebra star" - shows "x \ Infinitesimal ==> inverse(x) \ HFinite" -apply (drule HInfinite_diff_HFinite_Infinitesimal_disj) -apply (blast intro: HFinite_inverse HInfinite_inverse_Infinitesimal Infinitesimal_subset_HFinite [THEN subsetD]) -done +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 lemma HFinite_not_Infinitesimal_inverse: - fixes x :: "'a::real_normed_div_algebra star" - shows "x \ HFinite - Infinitesimal ==> inverse x \ HFinite - Infinitesimal" -apply (auto intro: Infinitesimal_inverse_HFinite) -apply (drule Infinitesimal_HFinite_mult2, assumption) -apply (simp add: not_Infinitesimal_not_zero) -done + "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 -lemma approx_inverse: - fixes x y :: "'a::real_normed_div_algebra star" - shows - "[| x \ y; y \ HFinite - Infinitesimal |] - ==> inverse x \ inverse y" -apply (frule HFinite_diff_Infinitesimal_approx, assumption) -apply (frule not_Infinitesimal_not_zero2) -apply (frule_tac x = x in not_Infinitesimal_not_zero2) -apply (drule HFinite_inverse2)+ -apply (drule approx_mult2, assumption, auto) -apply (drule_tac c = "inverse x" in approx_mult1, assumption) -apply (auto intro: approx_sym simp add: mult.assoc) -done +lemma approx_inverse: "x \ y \ y \ HFinite - Infinitesimal \ inverse x \ inverse y" + for x y :: "'a::real_normed_div_algebra star" + apply (frule HFinite_diff_Infinitesimal_approx, assumption) + apply (frule not_Infinitesimal_not_zero2) + apply (frule_tac x = x in not_Infinitesimal_not_zero2) + apply (drule HFinite_inverse2)+ + apply (drule approx_mult2, assumption, auto) + apply (drule_tac c = "inverse x" in approx_mult1, assumption) + apply (auto intro: approx_sym simp add: mult.assoc) + done (*Used for NSLIM_inverse, NSLIMSEQ_inverse*) lemmas star_of_approx_inverse = star_of_HFinite_diff_Infinitesimal [THEN [2] approx_inverse] lemmas hypreal_of_real_approx_inverse = hypreal_of_real_HFinite_diff_Infinitesimal [THEN [2] approx_inverse] lemma inverse_add_Infinitesimal_approx: - fixes x h :: "'a::real_normed_div_algebra star" - shows - "[| x \ HFinite - Infinitesimal; - h \ Infinitesimal |] ==> inverse(x + h) \ inverse x" -apply (auto intro: approx_inverse approx_sym Infinitesimal_add_approx_self) -done + "x \ HFinite - Infinitesimal \ h \ Infinitesimal \ inverse (x + h) \ inverse x" + for x h :: "'a::real_normed_div_algebra star" + by (auto intro: approx_inverse approx_sym Infinitesimal_add_approx_self) lemma inverse_add_Infinitesimal_approx2: - fixes x h :: "'a::real_normed_div_algebra star" - shows - "[| x \ HFinite - Infinitesimal; - h \ Infinitesimal |] ==> inverse(h + x) \ inverse x" -apply (rule add.commute [THEN subst]) -apply (blast intro: inverse_add_Infinitesimal_approx) -done + "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 lemma inverse_add_Infinitesimal_approx_Infinitesimal: - fixes x h :: "'a::real_normed_div_algebra star" - shows - "[| x \ HFinite - Infinitesimal; - h \ Infinitesimal |] ==> inverse(x + h) - inverse x \ h" -apply (rule approx_trans2) -apply (auto intro: inverse_add_Infinitesimal_approx - simp add: mem_infmal_iff approx_minus_iff [symmetric]) -done + "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 -lemma Infinitesimal_square_iff: - fixes x :: "'a::real_normed_div_algebra star" - shows "(x \ Infinitesimal) = (x*x \ Infinitesimal)" -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 +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 declare Infinitesimal_square_iff [symmetric, simp] -lemma HFinite_square_iff [simp]: - fixes x :: "'a::real_normed_div_algebra star" - shows "(x*x \ HFinite) = (x \ HFinite)" -apply (auto intro: HFinite_mult) -apply (auto dest: HInfinite_mult simp add: HFinite_HInfinite_iff) -done +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 -lemma HInfinite_square_iff [simp]: - fixes x :: "'a::real_normed_div_algebra star" - shows "(x*x \ HInfinite) = (x \ HInfinite)" -by (auto simp add: HInfinite_HFinite_iff) +lemma HInfinite_square_iff [simp]: "x * x \ HInfinite \ x \ HInfinite" + for x :: "'a::real_normed_div_algebra star" + by (auto simp add: HInfinite_HFinite_iff) -lemma approx_HFinite_mult_cancel: - fixes a w z :: "'a::real_normed_div_algebra star" - shows "[| a: HFinite-Infinitesimal; a* w \ a*z |] ==> w \ z" -apply safe -apply (frule HFinite_inverse, assumption) -apply (drule not_Infinitesimal_not_zero) -apply (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) -done +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 -lemma approx_HFinite_mult_cancel_iff1: - fixes a w z :: "'a::real_normed_div_algebra star" - shows "a: HFinite-Infinitesimal ==> (a * w \ a * z) = (w \ z)" -by (auto intro: approx_mult2 approx_HFinite_mult_cancel) +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 +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 -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 +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 -lemma HInfinite_ge_HInfinite: - "[| (x::hypreal) \ HInfinite; x \ y; 0 \ x |] ==> y \ HInfinite" -by (auto intro: HFinite_bounded simp add: HInfinite_HFinite_iff) +lemma HInfinite_ge_HInfinite: "x \ HInfinite \ x \ y \ 0 \ x \ y \ HInfinite" + for x y :: hypreal + by (auto intro: HFinite_bounded simp add: HInfinite_HFinite_iff) -lemma Infinitesimal_inverse_HInfinite: - fixes x :: "'a::real_normed_div_algebra star" - shows "[| x \ Infinitesimal; x \ 0 |] ==> inverse x \ HInfinite" -apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2]) -apply (auto dest: Infinitesimal_HFinite_mult2) -done +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 lemma HInfinite_HFinite_not_Infinitesimal_mult: - fixes x y :: "'a::real_normed_div_algebra star" - shows "[| x \ HInfinite; y \ HFinite - Infinitesimal |] - ==> x * y \ HInfinite" -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 + "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 lemma HInfinite_HFinite_not_Infinitesimal_mult2: - fixes x y :: "'a::real_normed_div_algebra star" - shows "[| x \ HInfinite; y \ HFinite - Infinitesimal |] - ==> y * x \ HInfinite" -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 + "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 -lemma HInfinite_gt_SReal: - "[| (x::hypreal) \ HInfinite; 0 < x; y \ \ |] ==> y < x" -by (auto dest!: bspec simp add: HInfinite_def abs_if order_less_imp_le) +lemma HInfinite_gt_SReal: "x \ HInfinite \ 0 < x \ y \ \ \ y < x" + for x y :: hypreal + by (auto dest!: bspec simp add: HInfinite_def abs_if order_less_imp_le) -lemma HInfinite_gt_zero_gt_one: - "[| (x::hypreal) \ HInfinite; 0 < x |] ==> 1 < x" -by (auto intro: HInfinite_gt_SReal) +lemma HInfinite_gt_zero_gt_one: "x \ HInfinite \ 0 < x \ 1 < x" + for x :: hypreal + by (auto intro: HInfinite_gt_SReal) lemma not_HInfinite_one [simp]: "1 \ HInfinite" -apply (simp (no_asm) add: HInfinite_HFinite_iff) -done + by (simp add: HInfinite_HFinite_iff) -lemma approx_hrabs_disj: "\x::hypreal\ \ x \ \x\ \ -x" -by (cut_tac x = x in hrabs_disj, auto) +lemma approx_hrabs_disj: "\x\ \ x \ \x\ \ -x" + for x :: hypreal + using hrabs_disj [of x] by auto -subsection\Theorems about Monads\ - -lemma monad_hrabs_Un_subset: "monad \x\ \ monad(x::hypreal) Un monad(-x)" -by (rule_tac x1 = x in hrabs_disj [THEN disjE], auto) +subsection \Theorems about Monads\ -lemma Infinitesimal_monad_eq: "e \ Infinitesimal ==> monad (x+e) = monad x" -by (fast intro!: Infinitesimal_add_approx_self [THEN approx_sym] approx_monad_iff [THEN iffD1]) +lemma monad_hrabs_Un_subset: "monad \x\ \ monad x \ monad (- x)" + for x :: hypreal + by (rule hrabs_disj [of x, THEN disjE]) auto -lemma mem_monad_iff: "(u \ monad x) = (-u \ monad (-x))" -by (simp add: monad_def) +lemma Infinitesimal_monad_eq: "e \ Infinitesimal \ monad (x + e) = monad x" + by (fast intro!: Infinitesimal_add_approx_self [THEN approx_sym] approx_monad_iff [THEN iffD1]) -lemma Infinitesimal_monad_zero_iff: "(x \ Infinitesimal) = (x \ monad 0)" -by (auto intro: approx_sym simp add: monad_def mem_infmal_iff) +lemma mem_monad_iff: "u \ monad x \ - u \ monad (- x)" + by (simp add: monad_def) + +lemma Infinitesimal_monad_zero_iff: "x \ Infinitesimal \ x \ monad 0" + by (auto intro: approx_sym simp add: monad_def mem_infmal_iff) -lemma monad_zero_minus_iff: "(x \ monad 0) = (-x \ monad 0)" -apply (simp (no_asm) add: Infinitesimal_monad_zero_iff [symmetric]) -done +lemma monad_zero_minus_iff: "x \ monad 0 \ - x \ monad 0" + by (simp add: Infinitesimal_monad_zero_iff [symmetric]) -lemma monad_zero_hrabs_iff: "((x::hypreal) \ monad 0) = (\x\ \ monad 0)" -apply (rule_tac x1 = x in hrabs_disj [THEN disjE]) -apply (auto simp add: monad_zero_minus_iff [symmetric]) -done +lemma monad_zero_hrabs_iff: "x \ monad 0 \ \x\ \ monad 0" + for x :: hypreal + by (rule hrabs_disj [of x, THEN disjE]) (auto simp: monad_zero_minus_iff [symmetric]) lemma mem_monad_self [simp]: "x \ monad x" -by (simp add: monad_def) + by (simp add: monad_def) -subsection\Proof that @{term "x \ y"} implies @{term"\x\ \ \y\"}\ +subsection \Proof that @{term "x \ y"} implies @{term"\x\ \ \y\"}\ -lemma approx_subset_monad: "x \ y ==> {x,y} \ monad x" -apply (simp (no_asm)) -apply (simp add: approx_monad_iff) -done +lemma approx_subset_monad: "x \ y \ {x, y} \ monad x" + 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 +lemma approx_subset_monad2: "x \ y \ {x, y} \ monad y" + apply (drule approx_sym) + apply (fast dest: approx_subset_monad) + done -lemma mem_monad_approx: "u \ monad x ==> x \ u" -by (simp add: monad_def) +lemma mem_monad_approx: "u \ monad x \ x \ u" + by (simp add: monad_def) + +lemma approx_mem_monad: "x \ u \ u \ monad x" + by (simp add: monad_def) -lemma approx_mem_monad: "x \ u ==> u \ monad x" -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 +lemma approx_mem_monad2: "x \ u \ x \ monad u" + apply (simp add: monad_def) + apply (blast intro!: approx_sym) + done -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 +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 -lemma Infinitesimal_approx_hrabs: - "[| x \ y; (x::hypreal) \ Infinitesimal |] ==> \x\ \ \y\" -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 +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 -lemma less_Infinitesimal_less: - "[| 0 < x; (x::hypreal) \Infinitesimal; e :Infinitesimal |] ==> e < x" -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 +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 -lemma Ball_mem_monad_gt_zero: - "[| 0 < (x::hypreal); x \ Infinitesimal; u \ monad x |] ==> 0 < u" -apply (drule mem_monad_approx [THEN approx_sym]) -apply (erule bex_Infinitesimal_iff2 [THEN iffD2, THEN bexE]) -apply (drule_tac e = "-xa" in less_Infinitesimal_less, auto) -done +lemma Ball_mem_monad_gt_zero: "0 < x \ x \ Infinitesimal \ u \ monad x \ 0 < u" + for u x :: hypreal + apply (drule mem_monad_approx [THEN approx_sym]) + apply (erule bex_Infinitesimal_iff2 [THEN iffD2, THEN bexE]) + apply (drule_tac e = "-xa" in less_Infinitesimal_less, auto) + done -lemma Ball_mem_monad_less_zero: - "[| (x::hypreal) < 0; x \ Infinitesimal; u \ monad x |] ==> u < 0" -apply (drule mem_monad_approx [THEN approx_sym]) -apply (erule bex_Infinitesimal_iff [THEN iffD2, THEN bexE]) -apply (cut_tac x = "-x" and e = xa in less_Infinitesimal_less, auto) -done +lemma Ball_mem_monad_less_zero: "x < 0 \ x \ Infinitesimal \ u \ monad x \ u < 0" + for u x :: hypreal + apply (drule mem_monad_approx [THEN approx_sym]) + apply (erule bex_Infinitesimal_iff [THEN iffD2, THEN bexE]) + apply (cut_tac x = "-x" and e = xa in less_Infinitesimal_less, auto) + done -lemma lemma_approx_gt_zero: - "[|0 < (x::hypreal); x \ Infinitesimal; x \ y|] ==> 0 < y" -by (blast dest: Ball_mem_monad_gt_zero approx_subset_monad) +lemma lemma_approx_gt_zero: "0 < x \ x \ Infinitesimal \ x \ y \ 0 < y" + for x y :: hypreal + by (blast dest: Ball_mem_monad_gt_zero approx_subset_monad) -lemma lemma_approx_less_zero: - "[|(x::hypreal) < 0; x \ Infinitesimal; x \ y|] ==> y < 0" -by (blast dest: Ball_mem_monad_less_zero approx_subset_monad) +lemma lemma_approx_less_zero: "x < 0 \ x \ Infinitesimal \ x \ y \ y < 0" + for x y :: hypreal + by (blast dest: Ball_mem_monad_less_zero approx_subset_monad) -theorem approx_hrabs: "(x::hypreal) \ y ==> \x\ \ \y\" -by (drule approx_hnorm, simp) +lemma approx_hrabs: "x \ y \ \x\ \ \y\" + for x y :: hypreal + by (drule approx_hnorm) simp -lemma approx_hrabs_zero_cancel: "\x::hypreal\ \ 0 ==> x \ 0" -apply (cut_tac x = x in hrabs_disj) -apply (auto dest: approx_minus) -done +lemma approx_hrabs_zero_cancel: "\x\ \ 0 \ x \ 0" + for x :: hypreal + using hrabs_disj [of x] by (auto dest: approx_minus) -lemma approx_hrabs_add_Infinitesimal: - "(e::hypreal) \ Infinitesimal ==> \x\ \ \x + e\" -by (fast intro: approx_hrabs Infinitesimal_add_approx_self) +lemma approx_hrabs_add_Infinitesimal: "e \ Infinitesimal \ \x\ \ \x + e\" + for e x :: hypreal + by (fast intro: approx_hrabs Infinitesimal_add_approx_self) -lemma approx_hrabs_add_minus_Infinitesimal: - "(e::hypreal) \ Infinitesimal ==> \x\ \ \x + -e\" -by (fast intro: approx_hrabs Infinitesimal_add_minus_approx_self) +lemma approx_hrabs_add_minus_Infinitesimal: "e \ Infinitesimal ==> \x\ \ \x + -e\" + for e x :: hypreal + by (fast intro: approx_hrabs Infinitesimal_add_minus_approx_self) lemma hrabs_add_Infinitesimal_cancel: - "[| (e::hypreal) \ Infinitesimal; e' \ Infinitesimal; - \x + e\ = \y + e'\|] ==> \x\ \ \y\" -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 + "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 lemma hrabs_add_minus_Infinitesimal_cancel: - "[| (e::hypreal) \ Infinitesimal; e' \ Infinitesimal; - \x + -e\ = \y + -e'\|] ==> \x\ \ \y\" -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 + "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 + subsection \More @{term HFinite} and @{term Infinitesimal} Theorems\ -(* interesting slightly counterintuitive theorem: necessary - for proving that an open interval is an NS open set -*) +text \ + Interesting slightly counterintuitive theorem: necessary + for proving that an open interval is an NS open set. +\ lemma Infinitesimal_add_hypreal_of_real_less: - "[| x < y; u \ Infinitesimal |] - ==> hypreal_of_real x + u < hypreal_of_real y" -apply (simp add: Infinitesimal_def) -apply (drule_tac x = "hypreal_of_real y + -hypreal_of_real x" in bspec, simp) -apply (simp add: abs_less_iff) -done + "x < y \ u \ Infinitesimal \ hypreal_of_real x + u < hypreal_of_real y" + apply (simp add: Infinitesimal_def) + apply (drule_tac x = "hypreal_of_real y + -hypreal_of_real x" in bspec, simp) + apply (simp add: abs_less_iff) + done lemma Infinitesimal_add_hrabs_hypreal_of_real_less: - "[| x \ Infinitesimal; \hypreal_of_real r\ < hypreal_of_real y |] - ==> \hypreal_of_real r + x\ < hypreal_of_real y" -apply (drule_tac x = "hypreal_of_real r" in approx_hrabs_add_Infinitesimal) -apply (drule approx_sym [THEN bex_Infinitesimal_iff2 [THEN iffD2]]) -apply (auto intro!: Infinitesimal_add_hypreal_of_real_less - simp del: star_of_abs - simp add: star_of_abs [symmetric]) -done + "x \ Infinitesimal \ \hypreal_of_real r\ < hypreal_of_real y \ + \hypreal_of_real r + x\ < hypreal_of_real y" + apply (drule_tac x = "hypreal_of_real r" in approx_hrabs_add_Infinitesimal) + apply (drule approx_sym [THEN bex_Infinitesimal_iff2 [THEN iffD2]]) + apply (auto intro!: Infinitesimal_add_hypreal_of_real_less + simp del: star_of_abs simp add: star_of_abs [symmetric]) + done 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 + "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 lemma hypreal_of_real_le_add_Infininitesimal_cancel: - "[| u \ Infinitesimal; v \ Infinitesimal; - hypreal_of_real x + u \ hypreal_of_real y + v |] - ==> hypreal_of_real x \ hypreal_of_real y" -apply (simp add: linorder_not_less [symmetric], auto) -apply (drule_tac u = "v-u" in Infinitesimal_add_hypreal_of_real_less) -apply (auto simp add: Infinitesimal_diff) -done + "u \ Infinitesimal \ v \ Infinitesimal \ + hypreal_of_real x + u \ hypreal_of_real y + v \ + hypreal_of_real x \ hypreal_of_real y" + apply (simp add: linorder_not_less [symmetric], auto) + apply (drule_tac u = "v-u" in Infinitesimal_add_hypreal_of_real_less) + apply (auto simp add: Infinitesimal_diff) + done lemma hypreal_of_real_le_add_Infininitesimal_cancel2: - "[| u \ Infinitesimal; v \ Infinitesimal; - hypreal_of_real x + u \ hypreal_of_real y + v |] - ==> x \ y" -by (blast intro: star_of_le [THEN iffD1] - intro!: hypreal_of_real_le_add_Infininitesimal_cancel) + "u \ Infinitesimal \ v \ Infinitesimal \ + hypreal_of_real x + u \ hypreal_of_real y + v \ x \ y" + by (blast intro: star_of_le [THEN iffD1] intro!: hypreal_of_real_le_add_Infininitesimal_cancel) 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 + "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 (*used once, in Lim/NSDERIV_inverse*) -lemma Infinitesimal_add_not_zero: - "[| h \ Infinitesimal; x \ 0 |] ==> star_of x + h \ 0" -apply auto -apply (subgoal_tac "h = - star_of x", auto intro: minus_unique [symmetric]) -done +lemma Infinitesimal_add_not_zero: "h \ Infinitesimal \ x \ 0 \ star_of x + h \ 0" + apply auto + apply (subgoal_tac "h = - star_of x") + apply (auto intro: minus_unique [symmetric]) + done -lemma Infinitesimal_square_cancel [simp]: - "(x::hypreal)*x + y*y \ Infinitesimal ==> x*x \ Infinitesimal" -apply (rule Infinitesimal_interval2) -apply (rule_tac [3] zero_le_square, assumption) -apply (auto) -done +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 -lemma HFinite_square_cancel [simp]: - "(x::hypreal)*x + y*y \ HFinite ==> x*x \ HFinite" -apply (rule HFinite_bounded, assumption) -apply (auto) -done +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 -lemma Infinitesimal_square_cancel2 [simp]: - "(x::hypreal)*x + y*y \ Infinitesimal ==> y*y \ Infinitesimal" -apply (rule Infinitesimal_square_cancel) -apply (rule add.commute [THEN subst]) -apply (simp (no_asm)) -done +lemma Infinitesimal_square_cancel2 [simp]: "x * x + y * y \ Infinitesimal \ y * y \ Infinitesimal" + for x y :: hypreal + apply (rule Infinitesimal_square_cancel) + apply (rule add.commute [THEN subst]) + apply simp + done -lemma HFinite_square_cancel2 [simp]: - "(x::hypreal)*x + y*y \ HFinite ==> y*y \ HFinite" -apply (rule HFinite_square_cancel) -apply (rule add.commute [THEN subst]) -apply (simp (no_asm)) -done +lemma HFinite_square_cancel2 [simp]: "x * x + y * y \ HFinite \ y * y \ HFinite" + for x y :: hypreal + apply (rule HFinite_square_cancel) + apply (rule add.commute [THEN subst]) + apply simp + done lemma Infinitesimal_sum_square_cancel [simp]: - "(x::hypreal)*x + y*y + z*z \ Infinitesimal ==> x*x \ Infinitesimal" -apply (rule Infinitesimal_interval2, assumption) -apply (rule_tac [2] zero_le_square, simp) -apply (insert zero_le_square [of y]) -apply (insert zero_le_square [of z], simp del:zero_le_square) -done + "x * x + y * y + z * z \ Infinitesimal \ x * x \ Infinitesimal" + for x y z :: hypreal + apply (rule Infinitesimal_interval2, assumption) + apply (rule_tac [2] zero_le_square, simp) + apply (insert zero_le_square [of y]) + apply (insert zero_le_square [of z], simp del:zero_le_square) + done -lemma HFinite_sum_square_cancel [simp]: - "(x::hypreal)*x + y*y + z*z \ HFinite ==> x*x \ HFinite" -apply (rule HFinite_bounded, assumption) -apply (rule_tac [2] zero_le_square) -apply (insert zero_le_square [of y]) -apply (insert zero_le_square [of z], simp del:zero_le_square) -done +lemma HFinite_sum_square_cancel [simp]: "x * x + y * y + z * z \ HFinite \ x * x \ HFinite" + for x y z :: hypreal + apply (rule HFinite_bounded, assumption) + apply (rule_tac [2] zero_le_square) + apply (insert zero_le_square [of y]) + apply (insert zero_le_square [of z], simp del:zero_le_square) + done lemma Infinitesimal_sum_square_cancel2 [simp]: - "(y::hypreal)*y + x*x + z*z \ Infinitesimal ==> x*x \ Infinitesimal" -apply (rule Infinitesimal_sum_square_cancel) -apply (simp add: ac_simps) -done + "y * y + x * x + z * z \ Infinitesimal \ x * x \ Infinitesimal" + for x y z :: hypreal + apply (rule Infinitesimal_sum_square_cancel) + apply (simp add: ac_simps) + done -lemma HFinite_sum_square_cancel2 [simp]: - "(y::hypreal)*y + x*x + z*z \ HFinite ==> x*x \ HFinite" -apply (rule HFinite_sum_square_cancel) -apply (simp add: ac_simps) -done +lemma HFinite_sum_square_cancel2 [simp]: "y * y + x * x + z * z \ HFinite \ x * x \ HFinite" + for x y z :: hypreal + apply (rule HFinite_sum_square_cancel) + apply (simp add: ac_simps) + done lemma Infinitesimal_sum_square_cancel3 [simp]: - "(z::hypreal)*z + y*y + x*x \ Infinitesimal ==> x*x \ Infinitesimal" -apply (rule Infinitesimal_sum_square_cancel) -apply (simp add: ac_simps) -done + "z * z + y * y + x * x \ Infinitesimal \ x * x \ Infinitesimal" + for x y z :: hypreal + apply (rule Infinitesimal_sum_square_cancel) + apply (simp add: ac_simps) + done -lemma HFinite_sum_square_cancel3 [simp]: - "(z::hypreal)*z + y*y + x*x \ HFinite ==> x*x \ HFinite" -apply (rule HFinite_sum_square_cancel) -apply (simp add: ac_simps) -done +lemma HFinite_sum_square_cancel3 [simp]: "z * z + y * y + x * x \ HFinite \ x * x \ HFinite" + for x y z :: hypreal + apply (rule HFinite_sum_square_cancel) + apply (simp add: ac_simps) + 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 +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 -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 +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 -subsection\Theorems about Standard Part\ +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 +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 -lemma st_SReal: "x \ HFinite ==> st x \ \" -apply (simp add: st_def) -apply (frule st_part_Ex, safe) -apply (rule someI2) -apply (auto intro: approx_sym) -done +lemma st_SReal: "x \ HFinite \ st x \ \" + apply (simp add: st_def) + apply (frule st_part_Ex, safe) + apply (rule someI2) + apply (auto intro: approx_sym) + done -lemma st_HFinite: "x \ HFinite ==> st x \ HFinite" -by (erule st_SReal [THEN SReal_subset_HFinite [THEN subsetD]]) +lemma st_HFinite: "x \ HFinite \ st x \ HFinite" + by (erule st_SReal [THEN SReal_subset_HFinite [THEN subsetD]]) -lemma st_unique: "\r \ \; r \ x\ \ st x = r" -apply (frule SReal_subset_HFinite [THEN subsetD]) -apply (drule (1) approx_HFinite) -apply (unfold st_def) -apply (rule some_equality) -apply (auto intro: approx_unique_real) -done +lemma st_unique: "r \ \ \ r \ x \ st x = r" + apply (frule SReal_subset_HFinite [THEN subsetD]) + apply (drule (1) approx_HFinite) + apply (unfold st_def) + apply (rule some_equality) + apply (auto intro: approx_unique_real) + done -lemma st_SReal_eq: "x \ \ ==> st x = x" +lemma st_SReal_eq: "x \ \ \ st x = x" by (metis approx_refl st_unique) lemma st_hypreal_of_real [simp]: "st (hypreal_of_real x) = hypreal_of_real x" -by (rule SReal_hypreal_of_real [THEN st_SReal_eq]) + by (rule SReal_hypreal_of_real [THEN st_SReal_eq]) -lemma st_eq_approx: "[| x \ HFinite; y \ HFinite; st x = st y |] ==> x \ y" -by (auto dest!: st_approx_self elim!: approx_trans3) +lemma st_eq_approx: "x \ HFinite \ y \ HFinite \ st x = st y \ x \ y" + by (auto dest!: st_approx_self elim!: approx_trans3) lemma approx_st_eq: assumes x: "x \ HFinite" and y: "y \ HFinite" and xy: "x \ y" @@ -1796,32 +1665,27 @@ by (fast elim: approx_trans approx_trans2 SReal_approx_iff [THEN iffD1]) qed -lemma st_eq_approx_iff: - "[| x \ HFinite; y \ HFinite|] - ==> (x \ y) = (st x = st y)" -by (blast intro: approx_st_eq st_eq_approx) +lemma st_eq_approx_iff: "x \ HFinite \ y \ HFinite \ x \ y \ st x = st y" + by (blast intro: approx_st_eq st_eq_approx) -lemma st_Infinitesimal_add_SReal: - "[| x \ \; e \ Infinitesimal |] ==> st(x + e) = x" -apply (erule st_unique) -apply (erule Infinitesimal_add_approx_self) -done +lemma st_Infinitesimal_add_SReal: "x \ \ \ e \ Infinitesimal \ st (x + e) = x" + apply (erule st_unique) + apply (erule Infinitesimal_add_approx_self) + done -lemma st_Infinitesimal_add_SReal2: - "[| x \ \; e \ Infinitesimal |] ==> st(e + x) = x" -apply (erule st_unique) -apply (erule Infinitesimal_add_approx_self2) -done +lemma st_Infinitesimal_add_SReal2: "x \ \ \ e \ Infinitesimal \ st (e + x) = x" + apply (erule st_unique) + apply (erule Infinitesimal_add_approx_self2) + done -lemma HFinite_st_Infinitesimal_add: - "x \ HFinite ==> \e \ Infinitesimal. x = st(x) + e" -by (blast dest!: st_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2]) +lemma HFinite_st_Infinitesimal_add: "x \ HFinite \ \e \ Infinitesimal. x = st(x) + e" + by (blast dest!: st_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2]) -lemma st_add: "\x \ HFinite; y \ HFinite\ \ st (x + y) = st x + st y" -by (simp add: st_unique st_SReal st_approx_self approx_add) +lemma st_add: "x \ HFinite \ y \ HFinite \ st (x + y) = st x + st y" + by (simp add: st_unique st_SReal st_approx_self approx_add) lemma st_numeral [simp]: "st (numeral w) = numeral w" -by (rule Reals_numeral [THEN st_SReal_eq]) + by (rule Reals_numeral [THEN st_SReal_eq]) lemma st_neg_numeral [simp]: "st (- numeral w) = - numeral w" proof - @@ -1831,79 +1695,71 @@ qed lemma st_0 [simp]: "st 0 = 0" -by (simp add: st_SReal_eq) + by (simp add: st_SReal_eq) lemma st_1 [simp]: "st 1 = 1" -by (simp add: st_SReal_eq) + by (simp add: st_SReal_eq) lemma st_neg_1 [simp]: "st (- 1) = - 1" -by (simp add: st_SReal_eq) + by (simp add: st_SReal_eq) lemma st_minus: "x \ HFinite \ st (- x) = - st x" -by (simp add: st_unique st_SReal st_approx_self approx_minus) + by (simp add: st_unique st_SReal st_approx_self approx_minus) lemma st_diff: "\x \ HFinite; y \ HFinite\ \ st (x - y) = st x - st y" -by (simp add: st_unique st_SReal st_approx_self approx_diff) + by (simp add: st_unique st_SReal st_approx_self approx_diff) lemma st_mult: "\x \ HFinite; y \ HFinite\ \ st (x * y) = st x * st y" -by (simp add: st_unique st_SReal st_approx_self approx_mult_HFinite) + by (simp add: st_unique st_SReal st_approx_self approx_mult_HFinite) -lemma st_Infinitesimal: "x \ Infinitesimal ==> st x = 0" -by (simp add: st_unique mem_infmal_iff) +lemma st_Infinitesimal: "x \ Infinitesimal \ st x = 0" + by (simp add: st_unique mem_infmal_iff) -lemma st_not_Infinitesimal: "st(x) \ 0 ==> x \ Infinitesimal" +lemma st_not_Infinitesimal: "st(x) \ 0 \ x \ Infinitesimal" by (fast intro: st_Infinitesimal) -lemma st_inverse: - "[| x \ HFinite; st x \ 0 |] - ==> st(inverse x) = inverse (st x)" -apply (rule_tac c1 = "st x" in mult_left_cancel [THEN iffD1]) -apply (auto simp add: st_mult [symmetric] st_not_Infinitesimal HFinite_inverse) -apply (subst right_inverse, auto) -done +lemma st_inverse: "x \ HFinite \ st x \ 0 \ st (inverse x) = inverse (st x)" + apply (rule_tac c1 = "st x" in mult_left_cancel [THEN iffD1]) + apply (auto simp add: st_mult [symmetric] st_not_Infinitesimal HFinite_inverse) + apply (subst right_inverse, auto) + done -lemma st_divide [simp]: - "[| x \ HFinite; y \ HFinite; st y \ 0 |] - ==> st(x/y) = (st x) / (st y)" -by (simp add: divide_inverse st_mult st_not_Infinitesimal HFinite_inverse st_inverse) +lemma st_divide [simp]: "x \ HFinite \ y \ HFinite \ st y \ 0 \ st (x / y) = st x / st y" + by (simp add: divide_inverse st_mult st_not_Infinitesimal HFinite_inverse st_inverse) -lemma st_idempotent [simp]: "x \ HFinite ==> st(st(x)) = st(x)" -by (blast intro: st_HFinite st_approx_self approx_st_eq) +lemma st_idempotent [simp]: "x \ HFinite \ st (st x) = st x" + by (blast intro: st_HFinite st_approx_self approx_st_eq) lemma Infinitesimal_add_st_less: - "[| x \ HFinite; y \ HFinite; u \ Infinitesimal; st x < st y |] - ==> st x + u < st y" -apply (drule st_SReal)+ -apply (auto intro!: Infinitesimal_add_hypreal_of_real_less simp add: SReal_iff) -done + "x \ HFinite \ y \ HFinite \ u \ Infinitesimal \ st x < st y \ st x + u < st y" + apply (drule st_SReal)+ + apply (auto intro!: Infinitesimal_add_hypreal_of_real_less simp add: SReal_iff) + done lemma Infinitesimal_add_st_le_cancel: - "[| x \ HFinite; y \ HFinite; - u \ Infinitesimal; st x \ st y + u - |] ==> st x \ st y" -apply (simp add: linorder_not_less [symmetric]) -apply (auto dest: Infinitesimal_add_st_less) -done + "x \ HFinite \ y \ HFinite \ u \ Infinitesimal \ + st x \ st y + u \ st x \ st y" + apply (simp add: linorder_not_less [symmetric]) + apply (auto dest: Infinitesimal_add_st_less) + done -lemma st_le: "[| x \ HFinite; y \ HFinite; x \ y |] ==> st(x) \ st(y)" -by (metis approx_le_bound approx_sym linear st_SReal st_approx_self st_part_Ex1) +lemma st_le: "x \ HFinite \ y \ HFinite \ x \ y \ st x \ st y" + by (metis approx_le_bound approx_sym linear st_SReal st_approx_self st_part_Ex1) -lemma st_zero_le: "[| 0 \ x; x \ HFinite |] ==> 0 \ st x" -apply (subst st_0 [symmetric]) -apply (rule st_le, auto) -done +lemma st_zero_le: "0 \ x \ x \ HFinite \ 0 \ st x" + apply (subst st_0 [symmetric]) + apply (rule st_le, auto) + done -lemma st_zero_ge: "[| x \ 0; x \ HFinite |] ==> st x \ 0" -apply (subst st_0 [symmetric]) -apply (rule st_le, auto) -done +lemma st_zero_ge: "x \ 0 \ x \ HFinite \ st x \ 0" + apply (subst st_0 [symmetric]) + apply (rule st_le, auto) + done -lemma st_hrabs: "x \ HFinite ==> \st x\ = st \x\" -apply (simp add: linorder_not_le st_zero_le abs_if st_minus - linorder_not_less) -apply (auto dest!: st_zero_ge [OF order_less_imp_le]) -done - +lemma st_hrabs: "x \ HFinite \ \st x\ = st \x\" + apply (simp add: linorder_not_le st_zero_le abs_if st_minus linorder_not_less) + apply (auto dest!: st_zero_ge [OF order_less_imp_le]) + done subsection \Alternative Definitions using Free Ultrafilter\ @@ -1911,78 +1767,73 @@ subsubsection \@{term HFinite}\ lemma HFinite_FreeUltrafilterNat: - "star_n X \ HFinite - ==> \u. eventually (\n. norm (X n) < u) FreeUltrafilterNat" -apply (auto simp add: HFinite_def SReal_def) -apply (rule_tac x=r in exI) -apply (simp add: hnorm_def star_of_def starfun_star_n) -apply (simp add: star_less_def starP2_star_n) -done + "star_n X \ HFinite \ \u. eventually (\n. norm (X n) < u) FreeUltrafilterNat" + apply (auto simp add: HFinite_def SReal_def) + apply (rule_tac x=r in exI) + apply (simp add: hnorm_def star_of_def starfun_star_n) + apply (simp add: star_less_def starP2_star_n) + done lemma FreeUltrafilterNat_HFinite: - "\u. eventually (\n. norm (X n) < u) FreeUltrafilterNat - ==> star_n X \ HFinite" -apply (auto simp add: HFinite_def mem_Rep_star_iff) -apply (rule_tac x="star_of u" in bexI) -apply (simp add: hnorm_def starfun_star_n star_of_def) -apply (simp add: star_less_def starP2_star_n) -apply (simp add: SReal_def) -done + "\u. eventually (\n. norm (X n) < u) FreeUltrafilterNat \ star_n X \ HFinite" + apply (auto simp add: HFinite_def mem_Rep_star_iff) + apply (rule_tac x="star_of u" in bexI) + apply (simp add: hnorm_def starfun_star_n star_of_def) + apply (simp add: star_less_def starP2_star_n) + apply (simp add: SReal_def) + done lemma HFinite_FreeUltrafilterNat_iff: - "(star_n X \ HFinite) = (\u. eventually (\n. norm (X n) < u) FreeUltrafilterNat)" -by (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite) + "star_n X \ HFinite \ (\u. eventually (\n. norm (X n) < u) FreeUltrafilterNat)" + by (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite) + subsubsection \@{term HInfinite}\ lemma lemma_Compl_eq: "- {n. u < norm (f n)} = {n. norm (f n) \ u}" -by auto + by auto lemma lemma_Compl_eq2: "- {n. norm (f n) < u} = {n. u \ norm (f n)}" -by auto + by auto -lemma lemma_Int_eq1: - "{n. norm (f n) \ u} Int {n. u \ norm (f n)} = {n. norm(f n) = u}" -by auto +lemma lemma_Int_eq1: "{n. norm (f n) \ u} Int {n. u \ norm (f n)} = {n. norm(f n) = u}" + by auto -lemma lemma_FreeUltrafilterNat_one: - "{n. norm (f n) = u} \ {n. norm (f n) < u + (1::real)}" -by auto +lemma lemma_FreeUltrafilterNat_one: "{n. norm (f n) = u} \ {n. norm (f n) < u + (1::real)}" + by auto -(*------------------------------------- - Exclude this type of sets from free - ultrafilter for Infinite numbers! - -------------------------------------*) +text \Exclude this type of sets from free ultrafilter for Infinite numbers!\ lemma FreeUltrafilterNat_const_Finite: - "eventually (\n. norm (X n) = u) FreeUltrafilterNat ==> star_n X \ HFinite" -apply (rule FreeUltrafilterNat_HFinite) -apply (rule_tac x = "u + 1" in exI) -apply (auto elim: eventually_mono) -done + "eventually (\n. norm (X n) = u) FreeUltrafilterNat \ star_n X \ HFinite" + apply (rule FreeUltrafilterNat_HFinite) + apply (rule_tac x = "u + 1" in exI) + apply (auto elim: eventually_mono) + done lemma HInfinite_FreeUltrafilterNat: - "star_n X \ HInfinite ==> \u. eventually (\n. u < norm (X n)) FreeUltrafilterNat" -apply (drule HInfinite_HFinite_iff [THEN iffD1]) -apply (simp add: HFinite_FreeUltrafilterNat_iff) -apply (rule allI, drule_tac x="u + 1" in spec) -apply (simp add: FreeUltrafilterNat.eventually_not_iff[symmetric]) -apply (auto elim: eventually_mono) -done + "star_n X \ HInfinite \ \u. eventually (\n. u < norm (X n)) FreeUltrafilterNat" + apply (drule HInfinite_HFinite_iff [THEN iffD1]) + apply (simp add: HFinite_FreeUltrafilterNat_iff) + apply (rule allI, drule_tac x="u + 1" in spec) + apply (simp add: FreeUltrafilterNat.eventually_not_iff[symmetric]) + apply (auto elim: eventually_mono) + done -lemma lemma_Int_HI: - "{n. norm (Xa n) < u} Int {n. X n = Xa n} \ {n. norm (X n) < (u::real)}" -by auto +lemma lemma_Int_HI: "{n. norm (Xa n) < u} \ {n. X n = Xa n} \ {n. norm (X n) < u}" + for u :: real + by auto -lemma lemma_Int_HIa: "{n. u < norm (X n)} Int {n. norm (X n) < u} = {}" -by (auto intro: order_less_asym) +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)) FreeUltrafilterNat ==> star_n X \ HInfinite" -apply (rule HInfinite_HFinite_iff [THEN iffD2]) -apply (safe, drule HFinite_FreeUltrafilterNat, safe) -apply (drule_tac x = u in spec) + "\u. eventually (\n. u < norm (X n)) FreeUltrafilterNat \ star_n X \ HInfinite" + apply (rule HInfinite_HFinite_iff [THEN iffD2]) + apply (safe, drule HFinite_FreeUltrafilterNat, safe) + apply (drule_tac x = u in spec) proof - - fix u assume "\\<^sub>Fn in \. norm (X n) < u" "\\<^sub>Fn in \. u < norm (X n)" + fix u + assume "\\<^sub>Fn in \. norm (X n) < u" "\\<^sub>Fn in \. u < norm (X n)" then have "\\<^sub>F x in \. False" by eventually_elim auto then show False @@ -1990,230 +1841,204 @@ qed lemma HInfinite_FreeUltrafilterNat_iff: - "(star_n X \ HInfinite) = (\u. eventually (\n. u < norm (X n)) FreeUltrafilterNat)" -by (blast intro!: HInfinite_FreeUltrafilterNat FreeUltrafilterNat_HInfinite) + "star_n X \ HInfinite \ (\u. eventually (\n. u < norm (X n)) FreeUltrafilterNat)" + by (blast intro!: HInfinite_FreeUltrafilterNat FreeUltrafilterNat_HInfinite) + subsubsection \@{term Infinitesimal}\ -lemma ball_SReal_eq: "(\x::hypreal \ Reals. P x) = (\x::real. P (star_of x))" -by (unfold SReal_def, auto) +lemma ball_SReal_eq: "(\x::hypreal \ Reals. P x) \ (\x::real. P (star_of x))" + by (auto simp: SReal_def) lemma Infinitesimal_FreeUltrafilterNat: - "star_n X \ Infinitesimal ==> \u>0. eventually (\n. norm (X n) < u) \" -apply (simp add: Infinitesimal_def ball_SReal_eq) -apply (simp add: hnorm_def starfun_star_n star_of_def) -apply (simp add: star_less_def starP2_star_n) -done + "star_n X \ Infinitesimal \ \u>0. eventually (\n. norm (X n) < u) \" + apply (simp add: Infinitesimal_def ball_SReal_eq) + apply (simp add: hnorm_def starfun_star_n star_of_def) + apply (simp add: star_less_def starP2_star_n) + done lemma FreeUltrafilterNat_Infinitesimal: - "\u>0. eventually (\n. norm (X n) < u) \ ==> star_n X \ Infinitesimal" -apply (simp add: Infinitesimal_def ball_SReal_eq) -apply (simp add: hnorm_def starfun_star_n star_of_def) -apply (simp add: star_less_def starP2_star_n) -done + "\u>0. eventually (\n. norm (X n) < u) \ \ star_n X \ Infinitesimal" + apply (simp add: Infinitesimal_def ball_SReal_eq) + apply (simp add: hnorm_def starfun_star_n star_of_def) + apply (simp add: star_less_def starP2_star_n) + done lemma Infinitesimal_FreeUltrafilterNat_iff: - "(star_n X \ Infinitesimal) = (\u>0. eventually (\n. norm (X n) < u) \)" -by (blast intro!: Infinitesimal_FreeUltrafilterNat FreeUltrafilterNat_Infinitesimal) + "(star_n X \ Infinitesimal) = (\u>0. eventually (\n. norm (X n) < u) \)" + by (blast intro!: Infinitesimal_FreeUltrafilterNat FreeUltrafilterNat_Infinitesimal) + -(*------------------------------------------------------------------------ - Infinitesimals as smaller than 1/n for all n::nat (> 0) - ------------------------------------------------------------------------*) +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 +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 lemma lemma_Infinitesimal2: - "(\r \ Reals. 0 < r --> x < r) = - (\n. x < inverse(hypreal_of_nat (Suc n)))" -apply safe - apply (drule_tac x = "inverse (hypreal_of_real (real (Suc n))) " in bspec) - apply simp_all + "(\r \ Reals. 0 < r \ x < r) \ (\n. x < inverse(hypreal_of_nat (Suc n)))" + apply safe + apply (drule_tac x = "inverse (hypreal_of_real (real (Suc n))) " in bspec) + apply simp_all using less_imp_of_nat_less apply fastforce -apply (auto dest!: reals_Archimedean simp add: SReal_iff simp del: of_nat_Suc) -apply (drule star_of_less [THEN iffD2]) -apply simp -apply (blast intro: order_less_trans) -done + apply (auto dest!: reals_Archimedean simp add: SReal_iff simp del: of_nat_Suc) + apply (drule star_of_less [THEN iffD2]) + apply simp + apply (blast intro: order_less_trans) + done 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 + "Infinitesimal = {x. \n. hnorm x < inverse (hypreal_of_nat (Suc n))}" + apply (simp add: Infinitesimal_def) + apply (auto simp add: lemma_Infinitesimal2) + done -subsection\Proof that \\\ is an infinite number\ +subsection \Proof that \\\ is an infinite number\ -text\It will follow that \\\ is an infinitesimal number.\ +text \It will follow that \\\ is an infinitesimal number.\ lemma Suc_Un_eq: "{n. n < Suc m} = {n. n < m} Un {n. n = m}" -by (auto simp add: less_Suc_eq) + by (auto simp add: less_Suc_eq) -(*------------------------------------------- - Prove that any segment is finite and hence cannot belong to FreeUltrafilterNat - -------------------------------------------*) + +text \Prove that any segment is finite and hence cannot belong to \FreeUltrafilterNat\.\ lemma finite_real_of_nat_segment: "finite {n::nat. real n < real (m::nat)}" - by (auto intro: finite_Collect_less_nat) + by auto lemma finite_real_of_nat_less_real: "finite {n::nat. real n < u}" -apply (cut_tac x = u in reals_Archimedean2, safe) -apply (rule finite_real_of_nat_segment [THEN [2] finite_subset]) -apply (auto dest: order_less_trans) -done + apply (cut_tac x = u in reals_Archimedean2, safe) + apply (rule finite_real_of_nat_segment [THEN [2] finite_subset]) + apply (auto dest: order_less_trans) + done -lemma lemma_real_le_Un_eq: - "{n. f n \ u} = {n. f n < u} Un {n. u = (f n :: real)}" -by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le) +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 (auto simp add: lemma_real_le_Un_eq lemma_finite_omega_set finite_real_of_nat_less_real) lemma finite_rabs_real_of_nat_le_real: "finite {n::nat. \real n\ \ u}" -apply (simp (no_asm) add: finite_real_of_nat_le_real) -done + by (simp add: finite_real_of_nat_le_real) lemma rabs_real_of_nat_le_real_FreeUltrafilterNat: - "\ eventually (\n. \real n\ \ u) FreeUltrafilterNat" -by (blast intro!: FreeUltrafilterNat.finite finite_rabs_real_of_nat_le_real) + "\ eventually (\n. \real n\ \ u) FreeUltrafilterNat" + by (blast intro!: FreeUltrafilterNat.finite finite_rabs_real_of_nat_le_real) lemma FreeUltrafilterNat_nat_gt_real: "eventually (\n. u < real n) FreeUltrafilterNat" -apply (rule FreeUltrafilterNat.finite') -apply (subgoal_tac "{n::nat. \ u < real n} = {n. real n \ u}") -apply (auto simp add: finite_real_of_nat_le_real) -done + apply (rule FreeUltrafilterNat.finite') + apply (subgoal_tac "{n::nat. \ u < real n} = {n. real n \ u}") + apply (auto simp add: finite_real_of_nat_le_real) + done -(*-------------------------------------------------------------- - The complement of {n. \real n\ \ u} = - {n. u < \real n\} is in FreeUltrafilterNat - by property of (free) ultrafilters - --------------------------------------------------------------*) +text \The complement of \{n. \real n\ \ u} = {n. u < \real n\}\ is in + \FreeUltrafilterNat\ 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) + by (auto dest!: order_le_less_trans simp add: linorder_not_le) -text\@{term \} is a member of @{term HInfinite}\ - +text \@{term \} is a member of @{term HInfinite}.\ theorem HInfinite_omega [simp]: "\ \ HInfinite" -apply (simp add: omega_def) -apply (rule FreeUltrafilterNat_HInfinite) -apply clarify -apply (rule_tac u1 = "u-1" in eventually_mono [OF FreeUltrafilterNat_nat_gt_real]) -apply auto -done + apply (simp add: omega_def) + apply (rule FreeUltrafilterNat_HInfinite) + apply clarify + apply (rule_tac u1 = "u-1" in eventually_mono [OF FreeUltrafilterNat_nat_gt_real]) + apply auto + done -(*----------------------------------------------- - Epsilon is a member of Infinitesimal - -----------------------------------------------*) + +text \Epsilon is a member of Infinitesimal.\ lemma Infinitesimal_epsilon [simp]: "\ \ Infinitesimal" -by (auto intro!: HInfinite_inverse_Infinitesimal HInfinite_omega simp add: hypreal_epsilon_inverse_omega) + by (auto intro!: HInfinite_inverse_Infinitesimal HInfinite_omega + simp add: hypreal_epsilon_inverse_omega) lemma HFinite_epsilon [simp]: "\ \ HFinite" -by (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]) + by (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]) lemma epsilon_approx_zero [simp]: "\ \ 0" -apply (simp (no_asm) add: mem_infmal_iff [symmetric]) -done - -(*------------------------------------------------------------------------ - 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. - -----------------------------------------------------------------------*) + by (simp add: mem_infmal_iff [symmetric]) -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 +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 -lemma finite_inverse_real_of_posnat_gt_real: - "0 < u ==> finite {n. u < inverse(real(Suc n))}" +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) have "{n. 1 + real n < inverse u} = {n. real n < inverse u - 1}" by fastforce - thus "finite {n. real (Suc n) < inverse u}" - using finite_real_of_nat_less_real [of "inverse u - 1"] by auto + then show "finite {n. real (Suc n) < inverse u}" + using finite_real_of_nat_less_real [of "inverse u - 1"] + by auto qed lemma lemma_real_le_Un_eq2: - "{n. u \ inverse(real(Suc n))} = - {n. u < inverse(real(Suc n))} Un {n. u = inverse(real(Suc n))}" -by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le) + "{n. u \ inverse(real(Suc n))} = + {n. u < inverse(real(Suc n))} \ {n. u = inverse(real(Suc n))}" + by (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le) -lemma finite_inverse_real_of_posnat_ge_real: - "0 < u ==> finite {n. u \ inverse(real(Suc n))}" -by (auto simp add: lemma_real_le_Un_eq2 lemma_finite_epsilon_set finite_inverse_real_of_posnat_gt_real - simp del: of_nat_Suc) +lemma finite_inverse_real_of_posnat_ge_real: "0 < u \ finite {n. u \ inverse (real (Suc n))}" + by (auto simp add: lemma_real_le_Un_eq2 lemma_finite_epsilon_set finite_inverse_real_of_posnat_gt_real + simp del: of_nat_Suc) lemma inverse_real_of_posnat_ge_real_FreeUltrafilterNat: - "0 < u ==> \ eventually (\n. u \ inverse(real(Suc n))) FreeUltrafilterNat" -by (blast intro!: FreeUltrafilterNat.finite finite_inverse_real_of_posnat_ge_real) + "0 < u \ \ eventually (\n. u \ inverse(real(Suc n))) FreeUltrafilterNat" + by (blast intro!: FreeUltrafilterNat.finite finite_inverse_real_of_posnat_ge_real) -(*-------------------------------------------------------------- - The complement of {n. u \ inverse(real(Suc n))} = - {n. inverse(real(Suc n)) < u} is in FreeUltrafilterNat - by property of (free) ultrafilters - --------------------------------------------------------------*) -lemma Compl_le_inverse_eq: - "- {n. u \ inverse(real(Suc n))} = {n. inverse(real(Suc n)) < u}" -by (auto dest!: order_le_less_trans simp add: linorder_not_le) +text \The complement of \{n. u \ inverse(real(Suc n))} = {n. inverse (real (Suc n)) < u}\ + is in \FreeUltrafilterNat\ by property of (free) ultrafilters.\ +lemma Compl_le_inverse_eq: "- {n. u \ inverse(real(Suc n))} = {n. inverse(real(Suc n)) < u}" + by (auto dest!: order_le_less_trans simp add: linorder_not_le) lemma FreeUltrafilterNat_inverse_real_of_posnat: - "0 < u ==> eventually (\n. inverse(real(Suc n)) < u) FreeUltrafilterNat" -by (drule inverse_real_of_posnat_ge_real_FreeUltrafilterNat) - (simp add: FreeUltrafilterNat.eventually_not_iff not_le[symmetric]) + "0 < u \ eventually (\n. inverse(real(Suc n)) < u) FreeUltrafilterNat" + by (drule inverse_real_of_posnat_ge_real_FreeUltrafilterNat) + (simp add: FreeUltrafilterNat.eventually_not_iff not_le[symmetric]) -text\Example of an hypersequence (i.e. an extended standard sequence) - whose term with an hypernatural suffix is an infinitesimal i.e. - the whn'nth term of the hypersequence is a member of Infinitesimal\ +text \Example of an hypersequence (i.e. an extended standard sequence) + whose term with an hypernatural suffix is an infinitesimal i.e. + the whn'nth term of the hypersequence is a member of Infinitesimal\ -lemma SEQ_Infinitesimal: - "( *f* (%n::nat. inverse(real(Suc n)))) whn : Infinitesimal" -by (simp add: hypnat_omega_def starfun_star_n star_n_inverse Infinitesimal_FreeUltrafilterNat_iff - FreeUltrafilterNat_inverse_real_of_posnat del: of_nat_Suc) +lemma SEQ_Infinitesimal: "( *f* (\n::nat. inverse(real(Suc n)))) whn \ Infinitesimal" + by (simp add: hypnat_omega_def starfun_star_n star_n_inverse Infinitesimal_FreeUltrafilterNat_iff + FreeUltrafilterNat_inverse_real_of_posnat del: of_nat_Suc) -text\Example where we get a hyperreal from a real sequence - for which a particular property holds. The theorem is - used in proofs about equivalence of nonstandard and - standard neighbourhoods. Also used for equivalence of - nonstandard ans standard definitions of pointwise - limit.\ +text \Example where we get a hyperreal from a real sequence + for which a particular property holds. The theorem is + used in proofs about equivalence of nonstandard and + standard neighbourhoods. Also used for equivalence of + nonstandard ans standard definitions of pointwise + limit.\ -(*----------------------------------------------------- - |X(n) - x| < 1/n ==> [] - hypreal_of_real x| \ Infinitesimal - -----------------------------------------------------*) +text \\|X(n) - x| < 1/n \ [] - hypreal_of_real x| \ Infinitesimal\\ lemma real_seq_to_hypreal_Infinitesimal: - "\n. norm(X n - x) < inverse(real(Suc n)) - ==> star_n X - star_of x \ Infinitesimal" -unfolding star_n_diff star_of_def Infinitesimal_FreeUltrafilterNat_iff star_n_inverse -by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat - intro: order_less_trans elim!: eventually_mono) + "\n. norm (X n - x) < inverse (real (Suc n)) \ star_n X - star_of x \ Infinitesimal" + unfolding star_n_diff star_of_def Infinitesimal_FreeUltrafilterNat_iff star_n_inverse + by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat + intro: order_less_trans elim!: eventually_mono) lemma real_seq_to_hypreal_approx: - "\n. norm(X n - x) < inverse(real(Suc n)) - ==> star_n X \ star_of x" -by (metis bex_Infinitesimal_iff real_seq_to_hypreal_Infinitesimal) + "\n. norm (X n - x) < inverse (real (Suc n)) \ star_n X \ star_of x" + by (metis bex_Infinitesimal_iff real_seq_to_hypreal_Infinitesimal) lemma real_seq_to_hypreal_approx2: - "\n. norm(x - X n) < inverse(real(Suc n)) - ==> star_n X \ star_of x" -by (metis norm_minus_commute real_seq_to_hypreal_approx) + "\n. norm (x - X n) < inverse(real(Suc n)) \ star_n X \ star_of x" + by (metis norm_minus_commute real_seq_to_hypreal_approx) lemma real_seq_to_hypreal_Infinitesimal2: - "\n. norm(X n - Y n) < inverse(real(Suc n)) - ==> star_n X - star_n Y \ Infinitesimal" -unfolding Infinitesimal_FreeUltrafilterNat_iff star_n_diff -by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat - intro: order_less_trans elim!: eventually_mono) + "\n. norm(X n - Y n) < inverse(real(Suc n)) \ star_n X - star_n Y \ Infinitesimal" + unfolding Infinitesimal_FreeUltrafilterNat_iff star_n_diff + by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat + intro: order_less_trans elim!: eventually_mono) end diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/NSComplex.thy --- a/src/HOL/Nonstandard_Analysis/NSComplex.thy Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/NSComplex.thy Tue Nov 01 00:44:24 2016 +0100 @@ -3,654 +3,604 @@ Author: Lawrence C Paulson *) -section\Nonstandard Complex Numbers\ +section \Nonstandard Complex Numbers\ theory NSComplex -imports NSA + imports NSA begin type_synonym hcomplex = "complex star" -abbreviation - hcomplex_of_complex :: "complex => complex star" where - "hcomplex_of_complex == star_of" +abbreviation hcomplex_of_complex :: "complex \ complex star" + where "hcomplex_of_complex \ star_of" -abbreviation - hcmod :: "complex star => real star" where - "hcmod == hnorm" +abbreviation hcmod :: "complex star \ real star" + where "hcmod \ hnorm" - (*--- real and Imaginary parts ---*) +subsubsection \Real and Imaginary parts\ + +definition hRe :: "hcomplex \ hypreal" + where "hRe = *f* Re" + +definition hIm :: "hcomplex \ hypreal" + where "hIm = *f* Im" + -definition - hRe :: "hcomplex => hypreal" where - "hRe = *f* Re" +subsubsection \Imaginary unit\ + +definition iii :: hcomplex + where "iii = star_of \" -definition - hIm :: "hcomplex => hypreal" where - "hIm = *f* Im" + +subsubsection \Complex conjugate\ + +definition hcnj :: "hcomplex \ hcomplex" + where "hcnj = *f* cnj" - (*------ imaginary unit ----------*) +subsubsection \Argand\ -definition - iii :: hcomplex where - "iii = star_of \" - - (*------- complex conjugate ------*) +definition hsgn :: "hcomplex \ hcomplex" + where "hsgn = *f* sgn" -definition - hcnj :: "hcomplex => hcomplex" where - "hcnj = *f* cnj" - - (*------------ Argand -------------*) +definition harg :: "hcomplex \ hypreal" + where "harg = *f* arg" -definition - hsgn :: "hcomplex => hcomplex" where - "hsgn = *f* sgn" +definition \ \abbreviation for \cos a + i sin a\\ + hcis :: "hypreal \ hcomplex" + where "hcis = *f* cis" -definition - harg :: "hcomplex => hypreal" where - "harg = *f* arg" + +subsubsection \Injection from hyperreals\ -definition - (* abbreviation for (cos a + i sin a) *) - hcis :: "hypreal => hcomplex" where - "hcis = *f* cis" +abbreviation hcomplex_of_hypreal :: "hypreal \ hcomplex" + where "hcomplex_of_hypreal \ of_hypreal" - (*----- injection from hyperreals -----*) +definition \ \abbreviation for \r * (cos a + i sin a)\\ + hrcis :: "hypreal \ hypreal \ hcomplex" + where "hrcis = *f2* rcis" -abbreviation - hcomplex_of_hypreal :: "hypreal \ hcomplex" where - "hcomplex_of_hypreal \ of_hypreal" -definition - (* abbreviation for r*(cos a + i sin a) *) - hrcis :: "[hypreal, hypreal] => hcomplex" where - "hrcis = *f2* rcis" +subsubsection \\e ^ (x + iy)\\ - (*------------ e ^ (x + iy) ------------*) -definition - hExp :: "hcomplex => hcomplex" where - "hExp = *f* exp" +definition hExp :: "hcomplex \ hcomplex" + where "hExp = *f* exp" -definition - HComplex :: "[hypreal,hypreal] => hcomplex" where - "HComplex = *f2* Complex" +definition HComplex :: "hypreal \ hypreal \ hcomplex" + where "HComplex = *f2* Complex" lemmas hcomplex_defs [transfer_unfold] = hRe_def hIm_def iii_def hcnj_def hsgn_def harg_def hcis_def hrcis_def hExp_def HComplex_def lemma Standard_hRe [simp]: "x \ Standard \ hRe x \ Standard" -by (simp add: hcomplex_defs) + by (simp add: hcomplex_defs) lemma Standard_hIm [simp]: "x \ Standard \ hIm x \ Standard" -by (simp add: hcomplex_defs) + by (simp add: hcomplex_defs) lemma Standard_iii [simp]: "iii \ Standard" -by (simp add: hcomplex_defs) + by (simp add: hcomplex_defs) lemma Standard_hcnj [simp]: "x \ Standard \ hcnj x \ Standard" -by (simp add: hcomplex_defs) + by (simp add: hcomplex_defs) lemma Standard_hsgn [simp]: "x \ Standard \ hsgn x \ Standard" -by (simp add: hcomplex_defs) + by (simp add: hcomplex_defs) lemma Standard_harg [simp]: "x \ Standard \ harg x \ Standard" -by (simp add: hcomplex_defs) + by (simp add: hcomplex_defs) lemma Standard_hcis [simp]: "r \ Standard \ hcis r \ Standard" -by (simp add: hcomplex_defs) + by (simp add: hcomplex_defs) lemma Standard_hExp [simp]: "x \ Standard \ hExp x \ Standard" -by (simp add: hcomplex_defs) + by (simp add: hcomplex_defs) -lemma Standard_hrcis [simp]: - "\r \ Standard; s \ Standard\ \ hrcis r s \ Standard" -by (simp add: hcomplex_defs) +lemma Standard_hrcis [simp]: "r \ Standard \ s \ Standard \ hrcis r s \ Standard" + by (simp add: hcomplex_defs) -lemma Standard_HComplex [simp]: - "\r \ Standard; s \ Standard\ \ HComplex r s \ Standard" -by (simp add: hcomplex_defs) +lemma Standard_HComplex [simp]: "r \ Standard \ s \ Standard \ HComplex r s \ Standard" + by (simp add: hcomplex_defs) lemma hcmod_def: "hcmod = *f* cmod" -by (rule hnorm_def) + by (rule hnorm_def) -subsection\Properties of Nonstandard Real and Imaginary Parts\ +subsection \Properties of Nonstandard Real and Imaginary Parts\ -lemma hcomplex_hRe_hIm_cancel_iff: - "!!w z. (w=z) = (hRe(w) = hRe(z) & hIm(w) = hIm(z))" -by transfer (rule complex_Re_Im_cancel_iff) +lemma hcomplex_hRe_hIm_cancel_iff: "\w z. w = z \ hRe w = hRe z \ hIm w = hIm z" + by transfer (rule complex_Re_Im_cancel_iff) -lemma hcomplex_equality [intro?]: - "!!z w. hRe z = hRe w ==> hIm z = hIm w ==> z = w" -by transfer (rule complex_equality) +lemma hcomplex_equality [intro?]: "\z w. hRe z = hRe w \ hIm z = hIm w \ z = w" + by transfer (rule complex_equality) lemma hcomplex_hRe_zero [simp]: "hRe 0 = 0" -by transfer simp + by transfer simp lemma hcomplex_hIm_zero [simp]: "hIm 0 = 0" -by transfer simp + by transfer simp lemma hcomplex_hRe_one [simp]: "hRe 1 = 1" -by transfer simp + by transfer simp lemma hcomplex_hIm_one [simp]: "hIm 1 = 0" -by transfer simp + by transfer simp + + +subsection \Addition for Nonstandard Complex Numbers\ + +lemma hRe_add: "\x y. hRe (x + y) = hRe x + hRe y" + by transfer simp + +lemma hIm_add: "\x y. hIm (x + y) = hIm x + hIm y" + by transfer simp -subsection\Addition for Nonstandard Complex Numbers\ - -lemma hRe_add: "!!x y. hRe(x + y) = hRe(x) + hRe(y)" -by transfer simp +subsection \More Minus Laws\ -lemma hIm_add: "!!x y. hIm(x + y) = hIm(x) + hIm(y)" -by transfer simp +lemma hRe_minus: "\z. hRe (- z) = - hRe z" + by transfer (rule uminus_complex.sel) -subsection\More Minus Laws\ - -lemma hRe_minus: "!!z. hRe(-z) = - hRe(z)" -by transfer (rule uminus_complex.sel) +lemma hIm_minus: "\z. hIm (- z) = - hIm z" + by transfer (rule uminus_complex.sel) -lemma hIm_minus: "!!z. hIm(-z) = - hIm(z)" -by transfer (rule uminus_complex.sel) +lemma hcomplex_add_minus_eq_minus: "x + y = 0 \ x = - y" + for x y :: hcomplex + apply (drule minus_unique) + apply (simp add: minus_equation_iff [of x y]) + done -lemma hcomplex_add_minus_eq_minus: - "x + y = (0::hcomplex) ==> x = -y" -apply (drule minus_unique) -apply (simp add: minus_equation_iff [of x y]) -done +lemma hcomplex_i_mult_eq [simp]: "iii * iii = - 1" + by transfer (rule i_squared) -lemma hcomplex_i_mult_eq [simp]: "iii * iii = -1" -by transfer (rule i_squared) - -lemma hcomplex_i_mult_left [simp]: "!!z. iii * (iii * z) = -z" -by transfer (rule complex_i_mult_minus) +lemma hcomplex_i_mult_left [simp]: "\z. iii * (iii * z) = - z" + by transfer (rule complex_i_mult_minus) lemma hcomplex_i_not_zero [simp]: "iii \ 0" -by transfer (rule complex_i_not_zero) + by transfer (rule complex_i_not_zero) -subsection\More Multiplication Laws\ +subsection \More Multiplication Laws\ -lemma hcomplex_mult_minus_one: "- 1 * (z::hcomplex) = -z" -by simp - -lemma hcomplex_mult_minus_one_right: "(z::hcomplex) * - 1 = -z" -by simp +lemma hcomplex_mult_minus_one: "- 1 * z = - z" + for z :: hcomplex + by simp -lemma hcomplex_mult_left_cancel: - "(c::hcomplex) \ (0::hcomplex) ==> (c*a=c*b) = (a=b)" -by simp +lemma hcomplex_mult_minus_one_right: "z * - 1 = - z" + for z :: hcomplex + by simp -lemma hcomplex_mult_right_cancel: - "(c::hcomplex) \ (0::hcomplex) ==> (a*c=b*c) = (a=b)" -by simp +lemma hcomplex_mult_left_cancel: "c \ 0 \ c * a = c * b \ a = b" + for a b c :: hcomplex + by simp + +lemma hcomplex_mult_right_cancel: "c \ 0 \ a * c = b * c \ a = b" + for a b c :: hcomplex + by simp -subsection\Subraction and Division\ +subsection \Subtraction and Division\ -lemma hcomplex_diff_eq_eq [simp]: "((x::hcomplex) - y = z) = (x = z + y)" (* TODO: delete *) -by (rule diff_eq_eq) +lemma hcomplex_diff_eq_eq [simp]: "x - y = z \ x = z + y" + for x y z :: hcomplex + by (rule diff_eq_eq) -subsection\Embedding Properties for @{term hcomplex_of_hypreal} Map\ +subsection \Embedding Properties for @{term hcomplex_of_hypreal} Map\ + +lemma hRe_hcomplex_of_hypreal [simp]: "\z. hRe (hcomplex_of_hypreal z) = z" + by transfer (rule Re_complex_of_real) -lemma hRe_hcomplex_of_hypreal [simp]: "!!z. hRe(hcomplex_of_hypreal z) = z" -by transfer (rule Re_complex_of_real) +lemma hIm_hcomplex_of_hypreal [simp]: "\z. hIm (hcomplex_of_hypreal z) = 0" + by transfer (rule Im_complex_of_real) -lemma hIm_hcomplex_of_hypreal [simp]: "!!z. hIm(hcomplex_of_hypreal z) = 0" -by transfer (rule Im_complex_of_real) +lemma hcomplex_of_hypreal_epsilon_not_zero [simp]: "hcomplex_of_hypreal \ \ 0" + by (simp add: hypreal_epsilon_not_zero) -lemma hcomplex_of_hypreal_epsilon_not_zero [simp]: - "hcomplex_of_hypreal \ \ 0" -by (simp add: hypreal_epsilon_not_zero) + +subsection \\HComplex\ theorems\ -subsection\HComplex theorems\ +lemma hRe_HComplex [simp]: "\x y. hRe (HComplex x y) = x" + by transfer simp -lemma hRe_HComplex [simp]: "!!x y. hRe (HComplex x y) = x" -by transfer simp +lemma hIm_HComplex [simp]: "\x y. hIm (HComplex x y) = y" + by transfer simp -lemma hIm_HComplex [simp]: "!!x y. hIm (HComplex x y) = y" -by transfer simp - -lemma hcomplex_surj [simp]: "!!z. HComplex (hRe z) (hIm z) = z" -by transfer (rule complex_surj) +lemma hcomplex_surj [simp]: "\z. HComplex (hRe z) (hIm z) = z" + by transfer (rule complex_surj) lemma hcomplex_induct [case_names rect(*, induct type: hcomplex*)]: - "(\x y. P (HComplex x y)) ==> P z" -by (rule hcomplex_surj [THEN subst], blast) + "(\x y. P (HComplex x y)) \ P z" + by (rule hcomplex_surj [THEN subst]) blast -subsection\Modulus (Absolute Value) of Nonstandard Complex Number\ +subsection \Modulus (Absolute Value) of Nonstandard Complex Number\ lemma hcomplex_of_hypreal_abs: - "hcomplex_of_hypreal \x\ = - hcomplex_of_hypreal (hcmod (hcomplex_of_hypreal x))" -by simp + "hcomplex_of_hypreal \x\ = hcomplex_of_hypreal (hcmod (hcomplex_of_hypreal x))" + by simp -lemma HComplex_inject [simp]: - "!!x y x' y'. HComplex x y = HComplex x' y' = (x=x' & y=y')" -by transfer (rule complex.inject) +lemma HComplex_inject [simp]: "\x y x' y'. HComplex x y = HComplex x' y' \ x = x' \ y = y'" + by transfer (rule complex.inject) lemma HComplex_add [simp]: - "!!x1 y1 x2 y2. HComplex x1 y1 + HComplex x2 y2 = HComplex (x1+x2) (y1+y2)" -by transfer (rule complex_add) + "\x1 y1 x2 y2. HComplex x1 y1 + HComplex x2 y2 = HComplex (x1 + x2) (y1 + y2)" + by transfer (rule complex_add) -lemma HComplex_minus [simp]: "!!x y. - HComplex x y = HComplex (-x) (-y)" -by transfer (rule complex_minus) +lemma HComplex_minus [simp]: "\x y. - HComplex x y = HComplex (- x) (- y)" + by transfer (rule complex_minus) lemma HComplex_diff [simp]: - "!!x1 y1 x2 y2. HComplex x1 y1 - HComplex x2 y2 = HComplex (x1-x2) (y1-y2)" -by transfer (rule complex_diff) + "\x1 y1 x2 y2. HComplex x1 y1 - HComplex x2 y2 = HComplex (x1 - x2) (y1 - y2)" + by transfer (rule complex_diff) lemma HComplex_mult [simp]: - "!!x1 y1 x2 y2. HComplex x1 y1 * HComplex x2 y2 = - HComplex (x1*x2 - y1*y2) (x1*y2 + y1*x2)" -by transfer (rule complex_mult) + "\x1 y1 x2 y2. HComplex x1 y1 * HComplex x2 y2 = HComplex (x1*x2 - y1*y2) (x1*y2 + y1*x2)" + by transfer (rule complex_mult) -(*HComplex_inverse is proved below*) +text \\HComplex_inverse\ is proved below.\ -lemma hcomplex_of_hypreal_eq: "!!r. hcomplex_of_hypreal r = HComplex r 0" -by transfer (rule complex_of_real_def) +lemma hcomplex_of_hypreal_eq: "\r. hcomplex_of_hypreal r = HComplex r 0" + by transfer (rule complex_of_real_def) lemma HComplex_add_hcomplex_of_hypreal [simp]: - "!!x y r. HComplex x y + hcomplex_of_hypreal r = HComplex (x+r) y" -by transfer (rule Complex_add_complex_of_real) + "\x y r. HComplex x y + hcomplex_of_hypreal r = HComplex (x + r) y" + by transfer (rule Complex_add_complex_of_real) lemma hcomplex_of_hypreal_add_HComplex [simp]: - "!!r x y. hcomplex_of_hypreal r + HComplex x y = HComplex (r+x) y" -by transfer (rule complex_of_real_add_Complex) + "\r x y. hcomplex_of_hypreal r + HComplex x y = HComplex (r + x) y" + by transfer (rule complex_of_real_add_Complex) lemma HComplex_mult_hcomplex_of_hypreal: - "!!x y r. HComplex x y * hcomplex_of_hypreal r = HComplex (x*r) (y*r)" -by transfer (rule Complex_mult_complex_of_real) + "\x y r. HComplex x y * hcomplex_of_hypreal r = HComplex (x * r) (y * r)" + by transfer (rule Complex_mult_complex_of_real) lemma hcomplex_of_hypreal_mult_HComplex: - "!!r x y. hcomplex_of_hypreal r * HComplex x y = HComplex (r*x) (r*y)" -by transfer (rule complex_of_real_mult_Complex) + "\r x y. hcomplex_of_hypreal r * HComplex x y = HComplex (r * x) (r * y)" + by transfer (rule complex_of_real_mult_Complex) -lemma i_hcomplex_of_hypreal [simp]: - "!!r. iii * hcomplex_of_hypreal r = HComplex 0 r" -by transfer (rule i_complex_of_real) +lemma i_hcomplex_of_hypreal [simp]: "\r. iii * hcomplex_of_hypreal r = HComplex 0 r" + by transfer (rule i_complex_of_real) -lemma hcomplex_of_hypreal_i [simp]: - "!!r. hcomplex_of_hypreal r * iii = HComplex 0 r" -by transfer (rule complex_of_real_i) +lemma hcomplex_of_hypreal_i [simp]: "\r. hcomplex_of_hypreal r * iii = HComplex 0 r" + by transfer (rule complex_of_real_i) -subsection\Conjugation\ +subsection \Conjugation\ -lemma hcomplex_hcnj_cancel_iff [iff]: "!!x y. (hcnj x = hcnj y) = (x = y)" -by transfer (rule complex_cnj_cancel_iff) +lemma hcomplex_hcnj_cancel_iff [iff]: "\x y. hcnj x = hcnj y \ x = y" + by transfer (rule complex_cnj_cancel_iff) -lemma hcomplex_hcnj_hcnj [simp]: "!!z. hcnj (hcnj z) = z" -by transfer (rule complex_cnj_cnj) +lemma hcomplex_hcnj_hcnj [simp]: "\z. hcnj (hcnj z) = z" + by transfer (rule complex_cnj_cnj) lemma hcomplex_hcnj_hcomplex_of_hypreal [simp]: - "!!x. hcnj (hcomplex_of_hypreal x) = hcomplex_of_hypreal x" -by transfer (rule complex_cnj_complex_of_real) + "\x. hcnj (hcomplex_of_hypreal x) = hcomplex_of_hypreal x" + by transfer (rule complex_cnj_complex_of_real) -lemma hcomplex_hmod_hcnj [simp]: "!!z. hcmod (hcnj z) = hcmod z" -by transfer (rule complex_mod_cnj) +lemma hcomplex_hmod_hcnj [simp]: "\z. hcmod (hcnj z) = hcmod z" + by transfer (rule complex_mod_cnj) -lemma hcomplex_hcnj_minus: "!!z. hcnj (-z) = - hcnj z" -by transfer (rule complex_cnj_minus) +lemma hcomplex_hcnj_minus: "\z. hcnj (- z) = - hcnj z" + by transfer (rule complex_cnj_minus) -lemma hcomplex_hcnj_inverse: "!!z. hcnj(inverse z) = inverse(hcnj z)" -by transfer (rule complex_cnj_inverse) +lemma hcomplex_hcnj_inverse: "\z. hcnj (inverse z) = inverse (hcnj z)" + by transfer (rule complex_cnj_inverse) -lemma hcomplex_hcnj_add: "!!w z. hcnj(w + z) = hcnj(w) + hcnj(z)" -by transfer (rule complex_cnj_add) +lemma hcomplex_hcnj_add: "\w z. hcnj (w + z) = hcnj w + hcnj z" + by transfer (rule complex_cnj_add) -lemma hcomplex_hcnj_diff: "!!w z. hcnj(w - z) = hcnj(w) - hcnj(z)" -by transfer (rule complex_cnj_diff) +lemma hcomplex_hcnj_diff: "\w z. hcnj (w - z) = hcnj w - hcnj z" + by transfer (rule complex_cnj_diff) -lemma hcomplex_hcnj_mult: "!!w z. hcnj(w * z) = hcnj(w) * hcnj(z)" -by transfer (rule complex_cnj_mult) +lemma hcomplex_hcnj_mult: "\w z. hcnj (w * z) = hcnj w * hcnj z" + by transfer (rule complex_cnj_mult) -lemma hcomplex_hcnj_divide: "!!w z. hcnj(w / z) = (hcnj w)/(hcnj z)" -by transfer (rule complex_cnj_divide) +lemma hcomplex_hcnj_divide: "\w z. hcnj (w / z) = hcnj w / hcnj z" + by transfer (rule complex_cnj_divide) lemma hcnj_one [simp]: "hcnj 1 = 1" -by transfer (rule complex_cnj_one) + by transfer (rule complex_cnj_one) lemma hcomplex_hcnj_zero [simp]: "hcnj 0 = 0" -by transfer (rule complex_cnj_zero) - -lemma hcomplex_hcnj_zero_iff [iff]: "!!z. (hcnj z = 0) = (z = 0)" -by transfer (rule complex_cnj_zero_iff) - -lemma hcomplex_mult_hcnj: - "!!z. z * hcnj z = hcomplex_of_hypreal ((hRe z)\<^sup>2 + (hIm z)\<^sup>2)" -by transfer (rule complex_mult_cnj) - - -subsection\More Theorems about the Function @{term hcmod}\ + by transfer (rule complex_cnj_zero) -lemma hcmod_hcomplex_of_hypreal_of_nat [simp]: - "hcmod (hcomplex_of_hypreal(hypreal_of_nat n)) = hypreal_of_nat n" -by simp - -lemma hcmod_hcomplex_of_hypreal_of_hypnat [simp]: - "hcmod (hcomplex_of_hypreal(hypreal_of_hypnat n)) = hypreal_of_hypnat n" -by simp +lemma hcomplex_hcnj_zero_iff [iff]: "\z. hcnj z = 0 \ z = 0" + by transfer (rule complex_cnj_zero_iff) -lemma hcmod_mult_hcnj: "!!z. hcmod(z * hcnj(z)) = (hcmod z)\<^sup>2" -by transfer (rule complex_mod_mult_cnj) - -lemma hcmod_triangle_ineq2 [simp]: - "!!a b. hcmod(b + a) - hcmod b \ hcmod a" -by transfer (rule complex_mod_triangle_ineq2) - -lemma hcmod_diff_ineq [simp]: "!!a b. hcmod(a) - hcmod(b) \ hcmod(a + b)" -by transfer (rule norm_diff_ineq) +lemma hcomplex_mult_hcnj: "\z. z * hcnj z = hcomplex_of_hypreal ((hRe z)\<^sup>2 + (hIm z)\<^sup>2)" + by transfer (rule complex_mult_cnj) -subsection\Exponentiation\ +subsection \More Theorems about the Function @{term hcmod}\ + +lemma hcmod_hcomplex_of_hypreal_of_nat [simp]: + "hcmod (hcomplex_of_hypreal (hypreal_of_nat n)) = hypreal_of_nat n" + by simp + +lemma hcmod_hcomplex_of_hypreal_of_hypnat [simp]: + "hcmod (hcomplex_of_hypreal(hypreal_of_hypnat n)) = hypreal_of_hypnat n" + by simp + +lemma hcmod_mult_hcnj: "\z. hcmod (z * hcnj z) = (hcmod z)\<^sup>2" + by transfer (rule complex_mod_mult_cnj) -lemma hcomplexpow_0 [simp]: "z ^ 0 = (1::hcomplex)" -by (rule power_0) +lemma hcmod_triangle_ineq2 [simp]: "\a b. hcmod (b + a) - hcmod b \ hcmod a" + by transfer (rule complex_mod_triangle_ineq2) + +lemma hcmod_diff_ineq [simp]: "\a b. hcmod a - hcmod b \ hcmod (a + b)" + by transfer (rule norm_diff_ineq) + -lemma hcomplexpow_Suc [simp]: "z ^ (Suc n) = (z::hcomplex) * (z ^ n)" -by (rule power_Suc) +subsection \Exponentiation\ + +lemma hcomplexpow_0 [simp]: "z ^ 0 = 1" + for z :: hcomplex + by (rule power_0) + +lemma hcomplexpow_Suc [simp]: "z ^ (Suc n) = z * (z ^ n)" + for z :: hcomplex + by (rule power_Suc) lemma hcomplexpow_i_squared [simp]: "iii\<^sup>2 = -1" -by transfer (rule power2_i) + by transfer (rule power2_i) -lemma hcomplex_of_hypreal_pow: - "!!x. hcomplex_of_hypreal (x ^ n) = (hcomplex_of_hypreal x) ^ n" -by transfer (rule of_real_power) +lemma hcomplex_of_hypreal_pow: "\x. hcomplex_of_hypreal (x ^ n) = hcomplex_of_hypreal x ^ n" + by transfer (rule of_real_power) -lemma hcomplex_hcnj_pow: "!!z. hcnj(z ^ n) = hcnj(z) ^ n" -by transfer (rule complex_cnj_power) +lemma hcomplex_hcnj_pow: "\z. hcnj (z ^ n) = hcnj z ^ n" + by transfer (rule complex_cnj_power) -lemma hcmod_hcomplexpow: "!!x. hcmod(x ^ n) = hcmod(x) ^ n" -by transfer (rule norm_power) +lemma hcmod_hcomplexpow: "\x. hcmod (x ^ n) = hcmod x ^ n" + by transfer (rule norm_power) lemma hcpow_minus: - "!!x n. (-x::hcomplex) pow n = - (if ( *p* even) n then (x pow n) else -(x pow n))" -by transfer simp + "\x n. (- x :: hcomplex) pow n = (if ( *p* even) n then (x pow n) else - (x pow n))" + by transfer simp -lemma hcpow_mult: - "((r::hcomplex) * s) pow n = (r pow n) * (s pow n)" +lemma hcpow_mult: "(r * s) pow n = (r pow n) * (s pow n)" + for r s :: hcomplex by (fact hyperpow_mult) -lemma hcpow_zero2 [simp]: - "\n. 0 pow (hSuc n) = (0::'a::semiring_1 star)" +lemma hcpow_zero2 [simp]: "\n. 0 pow (hSuc n) = (0::'a::semiring_1 star)" by transfer (rule power_0_Suc) -lemma hcpow_not_zero [simp,intro]: - "!!r n. r \ 0 ==> r pow n \ (0::hcomplex)" +lemma hcpow_not_zero [simp,intro]: "\r n. r \ 0 \ r pow n \ (0::hcomplex)" by (fact hyperpow_not_zero) -lemma hcpow_zero_zero: - "r pow n = (0::hcomplex) ==> r = 0" +lemma hcpow_zero_zero: "r pow n = 0 \ r = 0" + for r :: hcomplex by (blast intro: ccontr dest: hcpow_not_zero) -subsection\The Function @{term hsgn}\ + +subsection \The Function @{term hsgn}\ lemma hsgn_zero [simp]: "hsgn 0 = 0" -by transfer (rule sgn_zero) + by transfer (rule sgn_zero) lemma hsgn_one [simp]: "hsgn 1 = 1" -by transfer (rule sgn_one) + by transfer (rule sgn_one) -lemma hsgn_minus: "!!z. hsgn (-z) = - hsgn(z)" -by transfer (rule sgn_minus) +lemma hsgn_minus: "\z. hsgn (- z) = - hsgn z" + by transfer (rule sgn_minus) -lemma hsgn_eq: "!!z. hsgn z = z / hcomplex_of_hypreal (hcmod z)" -by transfer (rule sgn_eq) +lemma hsgn_eq: "\z. hsgn z = z / hcomplex_of_hypreal (hcmod z)" + by transfer (rule sgn_eq) -lemma hcmod_i: "!!x y. hcmod (HComplex x y) = ( *f* sqrt) (x\<^sup>2 + y\<^sup>2)" -by transfer (rule complex_norm) +lemma hcmod_i: "\x y. hcmod (HComplex x y) = ( *f* sqrt) (x\<^sup>2 + y\<^sup>2)" + by transfer (rule complex_norm) lemma hcomplex_eq_cancel_iff1 [simp]: - "(hcomplex_of_hypreal xa = HComplex x y) = (xa = x & y = 0)" -by (simp add: hcomplex_of_hypreal_eq) + "hcomplex_of_hypreal xa = HComplex x y \ xa = x \ y = 0" + by (simp add: hcomplex_of_hypreal_eq) lemma hcomplex_eq_cancel_iff2 [simp]: - "(HComplex x y = hcomplex_of_hypreal xa) = (x = xa & y = 0)" -by (simp add: hcomplex_of_hypreal_eq) + "HComplex x y = hcomplex_of_hypreal xa \ x = xa \ y = 0" + by (simp add: hcomplex_of_hypreal_eq) -lemma HComplex_eq_0 [simp]: "!!x y. (HComplex x y = 0) = (x = 0 & y = 0)" -by transfer (rule Complex_eq_0) +lemma HComplex_eq_0 [simp]: "\x y. HComplex x y = 0 \ x = 0 \ y = 0" + by transfer (rule Complex_eq_0) -lemma HComplex_eq_1 [simp]: "!!x y. (HComplex x y = 1) = (x = 1 & y = 0)" -by transfer (rule Complex_eq_1) +lemma HComplex_eq_1 [simp]: "\x y. HComplex x y = 1 \ x = 1 \ y = 0" + by transfer (rule Complex_eq_1) lemma i_eq_HComplex_0_1: "iii = HComplex 0 1" -by transfer (simp add: complex_eq_iff) + by transfer (simp add: complex_eq_iff) -lemma HComplex_eq_i [simp]: "!!x y. (HComplex x y = iii) = (x = 0 & y = 1)" -by transfer (rule Complex_eq_i) - -lemma hRe_hsgn [simp]: "!!z. hRe(hsgn z) = hRe(z)/hcmod z" -by transfer (rule Re_sgn) +lemma HComplex_eq_i [simp]: "\x y. HComplex x y = iii \ x = 0 \ y = 1" + by transfer (rule Complex_eq_i) -lemma hIm_hsgn [simp]: "!!z. hIm(hsgn z) = hIm(z)/hcmod z" -by transfer (rule Im_sgn) +lemma hRe_hsgn [simp]: "\z. hRe (hsgn z) = hRe z / hcmod z" + by transfer (rule Re_sgn) -lemma HComplex_inverse: - "!!x y. inverse (HComplex x y) = HComplex (x/(x\<^sup>2 + y\<^sup>2)) (-y/(x\<^sup>2 + y\<^sup>2))" -by transfer (rule complex_inverse) - -lemma hRe_mult_i_eq[simp]: - "!!y. hRe (iii * hcomplex_of_hypreal y) = 0" -by transfer simp +lemma hIm_hsgn [simp]: "\z. hIm (hsgn z) = hIm z / hcmod z" + by transfer (rule Im_sgn) -lemma hIm_mult_i_eq [simp]: - "!!y. hIm (iii * hcomplex_of_hypreal y) = y" -by transfer simp +lemma HComplex_inverse: "\x y. inverse (HComplex x y) = HComplex (x / (x\<^sup>2 + y\<^sup>2)) (- y / (x\<^sup>2 + y\<^sup>2))" + by transfer (rule complex_inverse) -lemma hcmod_mult_i [simp]: "!!y. hcmod (iii * hcomplex_of_hypreal y) = \y\" -by transfer (simp add: norm_complex_def) - -lemma hcmod_mult_i2 [simp]: "!!y. hcmod (hcomplex_of_hypreal y * iii) = \y\" -by transfer (simp add: norm_complex_def) +lemma hRe_mult_i_eq[simp]: "\y. hRe (iii * hcomplex_of_hypreal y) = 0" + by transfer simp -(*---------------------------------------------------------------------------*) -(* harg *) -(*---------------------------------------------------------------------------*) +lemma hIm_mult_i_eq [simp]: "\y. hIm (iii * hcomplex_of_hypreal y) = y" + by transfer simp -lemma cos_harg_i_mult_zero [simp]: - "!!y. y \ 0 ==> ( *f* cos) (harg(HComplex 0 y)) = 0" -by transfer simp +lemma hcmod_mult_i [simp]: "\y. hcmod (iii * hcomplex_of_hypreal y) = \y\" + by transfer (simp add: norm_complex_def) -lemma hcomplex_of_hypreal_zero_iff [simp]: - "!!y. (hcomplex_of_hypreal y = 0) = (y = 0)" -by transfer (rule of_real_eq_0_iff) +lemma hcmod_mult_i2 [simp]: "\y. hcmod (hcomplex_of_hypreal y * iii) = \y\" + by transfer (simp add: norm_complex_def) -subsection\Polar Form for Nonstandard Complex Numbers\ +subsubsection \\harg\\ + +lemma cos_harg_i_mult_zero [simp]: "\y. y \ 0 \ ( *f* cos) (harg (HComplex 0 y)) = 0" + by transfer simp -lemma complex_split_polar2: - "\n. \r a. (z n) = complex_of_real r * (Complex (cos a) (sin a))" -by (auto intro: complex_split_polar) +lemma hcomplex_of_hypreal_zero_iff [simp]: "\y. hcomplex_of_hypreal y = 0 \ y = 0" + by transfer (rule of_real_eq_0_iff) + + +subsection \Polar Form for Nonstandard Complex Numbers\ + +lemma complex_split_polar2: "\n. \r a. (z n) = complex_of_real r * Complex (cos a) (sin a)" + by (auto intro: complex_split_polar) lemma hcomplex_split_polar: - "!!z. \r a. z = hcomplex_of_hypreal r * (HComplex(( *f* cos) a)(( *f* sin) a))" -by transfer (simp add: complex_split_polar) + "\z. \r a. z = hcomplex_of_hypreal r * (HComplex (( *f* cos) a) (( *f* sin) a))" + by transfer (simp add: complex_split_polar) lemma hcis_eq: - "!!a. hcis a = - (hcomplex_of_hypreal(( *f* cos) a) + - iii * hcomplex_of_hypreal(( *f* sin) a))" -by transfer (simp add: complex_eq_iff) + "\a. hcis a = hcomplex_of_hypreal (( *f* cos) a) + iii * hcomplex_of_hypreal (( *f* sin) a)" + by transfer (simp add: complex_eq_iff) -lemma hrcis_Ex: "!!z. \r a. z = hrcis r a" -by transfer (rule rcis_Ex) +lemma hrcis_Ex: "\z. \r a. z = hrcis r a" + by transfer (rule rcis_Ex) lemma hRe_hcomplex_polar [simp]: - "!!r a. hRe (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = - r * ( *f* cos) a" -by transfer simp + "\r a. hRe (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = r * ( *f* cos) a" + by transfer simp -lemma hRe_hrcis [simp]: "!!r a. hRe(hrcis r a) = r * ( *f* cos) a" -by transfer (rule Re_rcis) +lemma hRe_hrcis [simp]: "\r a. hRe (hrcis r a) = r * ( *f* cos) a" + by transfer (rule Re_rcis) lemma hIm_hcomplex_polar [simp]: - "!!r a. hIm (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = - r * ( *f* sin) a" -by transfer simp + "\r a. hIm (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = r * ( *f* sin) a" + by transfer simp -lemma hIm_hrcis [simp]: "!!r a. hIm(hrcis r a) = r * ( *f* sin) a" -by transfer (rule Im_rcis) +lemma hIm_hrcis [simp]: "\r a. hIm (hrcis r a) = r * ( *f* sin) a" + by transfer (rule Im_rcis) -lemma hcmod_unit_one [simp]: - "!!a. hcmod (HComplex (( *f* cos) a) (( *f* sin) a)) = 1" -by transfer (simp add: cmod_unit_one) +lemma hcmod_unit_one [simp]: "\a. hcmod (HComplex (( *f* cos) a) (( *f* sin) a)) = 1" + by transfer (simp add: cmod_unit_one) lemma hcmod_complex_polar [simp]: - "!!r a. hcmod (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = \r\" -by transfer (simp add: cmod_complex_polar) - -lemma hcmod_hrcis [simp]: "!!r a. hcmod(hrcis r a) = \r\" -by transfer (rule complex_mod_rcis) + "\r a. hcmod (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = \r\" + by transfer (simp add: cmod_complex_polar) -(*---------------------------------------------------------------------------*) -(* (r1 * hrcis a) * (r2 * hrcis b) = r1 * r2 * hrcis (a + b) *) -(*---------------------------------------------------------------------------*) +lemma hcmod_hrcis [simp]: "\r a. hcmod(hrcis r a) = \r\" + by transfer (rule complex_mod_rcis) -lemma hcis_hrcis_eq: "!!a. hcis a = hrcis 1 a" -by transfer (rule cis_rcis_eq) +text \\(r1 * hrcis a) * (r2 * hrcis b) = r1 * r2 * hrcis (a + b)\\ + +lemma hcis_hrcis_eq: "\a. hcis a = hrcis 1 a" + by transfer (rule cis_rcis_eq) declare hcis_hrcis_eq [symmetric, simp] -lemma hrcis_mult: - "!!a b r1 r2. hrcis r1 a * hrcis r2 b = hrcis (r1*r2) (a + b)" -by transfer (rule rcis_mult) +lemma hrcis_mult: "\a b r1 r2. hrcis r1 a * hrcis r2 b = hrcis (r1 * r2) (a + b)" + by transfer (rule rcis_mult) -lemma hcis_mult: "!!a b. hcis a * hcis b = hcis (a + b)" -by transfer (rule cis_mult) +lemma hcis_mult: "\a b. hcis a * hcis b = hcis (a + b)" + by transfer (rule cis_mult) lemma hcis_zero [simp]: "hcis 0 = 1" -by transfer (rule cis_zero) + by transfer (rule cis_zero) -lemma hrcis_zero_mod [simp]: "!!a. hrcis 0 a = 0" -by transfer (rule rcis_zero_mod) +lemma hrcis_zero_mod [simp]: "\a. hrcis 0 a = 0" + by transfer (rule rcis_zero_mod) -lemma hrcis_zero_arg [simp]: "!!r. hrcis r 0 = hcomplex_of_hypreal r" -by transfer (rule rcis_zero_arg) +lemma hrcis_zero_arg [simp]: "\r. hrcis r 0 = hcomplex_of_hypreal r" + by transfer (rule rcis_zero_arg) -lemma hcomplex_i_mult_minus [simp]: "!!x. iii * (iii * x) = - x" -by transfer (rule complex_i_mult_minus) +lemma hcomplex_i_mult_minus [simp]: "\x. iii * (iii * x) = - x" + by transfer (rule complex_i_mult_minus) lemma hcomplex_i_mult_minus2 [simp]: "iii * iii * x = - x" -by simp + by simp lemma hcis_hypreal_of_nat_Suc_mult: - "!!a. hcis (hypreal_of_nat (Suc n) * a) = - hcis a * hcis (hypreal_of_nat n * a)" -apply transfer -apply (simp add: distrib_right cis_mult) -done + "\a. hcis (hypreal_of_nat (Suc n) * a) = hcis a * hcis (hypreal_of_nat n * a)" + by transfer (simp add: distrib_right cis_mult) -lemma NSDeMoivre: "!!a. (hcis a) ^ n = hcis (hypreal_of_nat n * a)" -apply transfer -apply (rule DeMoivre) -done +lemma NSDeMoivre: "\a. (hcis a) ^ n = hcis (hypreal_of_nat n * a)" + by transfer (rule DeMoivre) lemma hcis_hypreal_of_hypnat_Suc_mult: - "!! a n. hcis (hypreal_of_hypnat (n + 1) * a) = - hcis a * hcis (hypreal_of_hypnat n * a)" -by transfer (simp add: distrib_right cis_mult) + "\a n. hcis (hypreal_of_hypnat (n + 1) * a) = hcis a * hcis (hypreal_of_hypnat n * a)" + by transfer (simp add: distrib_right cis_mult) -lemma NSDeMoivre_ext: - "!!a n. (hcis a) pow n = hcis (hypreal_of_hypnat n * a)" -by transfer (rule DeMoivre) +lemma NSDeMoivre_ext: "\a n. (hcis a) pow n = hcis (hypreal_of_hypnat n * a)" + by transfer (rule DeMoivre) + +lemma NSDeMoivre2: "\a r. (hrcis r a) ^ n = hrcis (r ^ n) (hypreal_of_nat n * a)" + by transfer (rule DeMoivre2) -lemma NSDeMoivre2: - "!!a r. (hrcis r a) ^ n = hrcis (r ^ n) (hypreal_of_nat n * a)" -by transfer (rule DeMoivre2) +lemma DeMoivre2_ext: "\a r n. (hrcis r a) pow n = hrcis (r pow n) (hypreal_of_hypnat n * a)" + by transfer (rule DeMoivre2) -lemma DeMoivre2_ext: - "!! a r n. (hrcis r a) pow n = hrcis (r pow n) (hypreal_of_hypnat n * a)" -by transfer (rule DeMoivre2) +lemma hcis_inverse [simp]: "\a. inverse (hcis a) = hcis (- a)" + by transfer (rule cis_inverse) -lemma hcis_inverse [simp]: "!!a. inverse(hcis a) = hcis (-a)" -by transfer (rule cis_inverse) +lemma hrcis_inverse: "\a r. inverse (hrcis r a) = hrcis (inverse r) (- a)" + by transfer (simp add: rcis_inverse inverse_eq_divide [symmetric]) -lemma hrcis_inverse: "!!a r. inverse(hrcis r a) = hrcis (inverse r) (-a)" -by transfer (simp add: rcis_inverse inverse_eq_divide [symmetric]) - -lemma hRe_hcis [simp]: "!!a. hRe(hcis a) = ( *f* cos) a" -by transfer simp +lemma hRe_hcis [simp]: "\a. hRe (hcis a) = ( *f* cos) a" + by transfer simp -lemma hIm_hcis [simp]: "!!a. hIm(hcis a) = ( *f* sin) a" -by transfer simp +lemma hIm_hcis [simp]: "\a. hIm (hcis a) = ( *f* sin) a" + by transfer simp -lemma cos_n_hRe_hcis_pow_n: "( *f* cos) (hypreal_of_nat n * a) = hRe(hcis a ^ n)" -by (simp add: NSDeMoivre) +lemma cos_n_hRe_hcis_pow_n: "( *f* cos) (hypreal_of_nat n * a) = hRe (hcis a ^ n)" + by (simp add: NSDeMoivre) -lemma sin_n_hIm_hcis_pow_n: "( *f* sin) (hypreal_of_nat n * a) = hIm(hcis a ^ n)" -by (simp add: NSDeMoivre) +lemma sin_n_hIm_hcis_pow_n: "( *f* sin) (hypreal_of_nat n * a) = hIm (hcis a ^ n)" + by (simp add: NSDeMoivre) -lemma cos_n_hRe_hcis_hcpow_n: "( *f* cos) (hypreal_of_hypnat n * a) = hRe(hcis a pow n)" -by (simp add: NSDeMoivre_ext) +lemma cos_n_hRe_hcis_hcpow_n: "( *f* cos) (hypreal_of_hypnat n * a) = hRe (hcis a pow n)" + by (simp add: NSDeMoivre_ext) -lemma sin_n_hIm_hcis_hcpow_n: "( *f* sin) (hypreal_of_hypnat n * a) = hIm(hcis a pow n)" -by (simp add: NSDeMoivre_ext) +lemma sin_n_hIm_hcis_hcpow_n: "( *f* sin) (hypreal_of_hypnat n * a) = hIm (hcis a pow n)" + by (simp add: NSDeMoivre_ext) -lemma hExp_add: "!!a b. hExp(a + b) = hExp(a) * hExp(b)" -by transfer (rule exp_add) +lemma hExp_add: "\a b. hExp (a + b) = hExp a * hExp b" + by transfer (rule exp_add) -subsection\@{term hcomplex_of_complex}: the Injection from - type @{typ complex} to to @{typ hcomplex}\ +subsection \@{term hcomplex_of_complex}: the Injection from type @{typ complex} to to @{typ hcomplex}\ lemma hcomplex_of_complex_i: "iii = hcomplex_of_complex \" -by (rule iii_def) + by (rule iii_def) -lemma hRe_hcomplex_of_complex: - "hRe (hcomplex_of_complex z) = hypreal_of_real (Re z)" -by transfer (rule refl) +lemma hRe_hcomplex_of_complex: "hRe (hcomplex_of_complex z) = hypreal_of_real (Re z)" + by transfer (rule refl) -lemma hIm_hcomplex_of_complex: - "hIm (hcomplex_of_complex z) = hypreal_of_real (Im z)" -by transfer (rule refl) +lemma hIm_hcomplex_of_complex: "hIm (hcomplex_of_complex z) = hypreal_of_real (Im z)" + by transfer (rule refl) -lemma hcmod_hcomplex_of_complex: - "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)" -by transfer (rule refl) +lemma hcmod_hcomplex_of_complex: "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)" + by transfer (rule refl) -subsection\Numerals and Arithmetic\ +subsection \Numerals and Arithmetic\ lemma hcomplex_of_hypreal_eq_hcomplex_of_complex: - "hcomplex_of_hypreal (hypreal_of_real x) = - hcomplex_of_complex (complex_of_real x)" -by transfer (rule refl) + "hcomplex_of_hypreal (hypreal_of_real x) = hcomplex_of_complex (complex_of_real x)" + by transfer (rule refl) lemma hcomplex_hypreal_numeral: "hcomplex_of_complex (numeral w) = hcomplex_of_hypreal(numeral w)" -by transfer (rule of_real_numeral [symmetric]) + by transfer (rule of_real_numeral [symmetric]) lemma hcomplex_hypreal_neg_numeral: "hcomplex_of_complex (- numeral w) = hcomplex_of_hypreal(- numeral w)" -by transfer (rule of_real_neg_numeral [symmetric]) + by transfer (rule of_real_neg_numeral [symmetric]) -lemma hcomplex_numeral_hcnj [simp]: - "hcnj (numeral v :: hcomplex) = numeral v" -by transfer (rule complex_cnj_numeral) +lemma hcomplex_numeral_hcnj [simp]: "hcnj (numeral v :: hcomplex) = numeral v" + by transfer (rule complex_cnj_numeral) -lemma hcomplex_numeral_hcmod [simp]: - "hcmod(numeral v :: hcomplex) = (numeral v :: hypreal)" -by transfer (rule norm_numeral) +lemma hcomplex_numeral_hcmod [simp]: "hcmod (numeral v :: hcomplex) = (numeral v :: hypreal)" + by transfer (rule norm_numeral) -lemma hcomplex_neg_numeral_hcmod [simp]: - "hcmod(- numeral v :: hcomplex) = (numeral v :: hypreal)" -by transfer (rule norm_neg_numeral) +lemma hcomplex_neg_numeral_hcmod [simp]: "hcmod (- numeral v :: hcomplex) = (numeral v :: hypreal)" + by transfer (rule norm_neg_numeral) -lemma hcomplex_numeral_hRe [simp]: - "hRe(numeral v :: hcomplex) = numeral v" -by transfer (rule complex_Re_numeral) +lemma hcomplex_numeral_hRe [simp]: "hRe (numeral v :: hcomplex) = numeral v" + by transfer (rule complex_Re_numeral) -lemma hcomplex_numeral_hIm [simp]: - "hIm(numeral v :: hcomplex) = 0" -by transfer (rule complex_Im_numeral) +lemma hcomplex_numeral_hIm [simp]: "hIm (numeral v :: hcomplex) = 0" + by transfer (rule complex_Im_numeral) end diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/NatStar.thy --- a/src/HOL/Nonstandard_Analysis/NatStar.thy Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/NatStar.thy Tue Nov 01 00:44:24 2016 +0100 @@ -5,237 +5,198 @@ Converted to Isar and polished by lcp *) -section\Star-transforms for the Hypernaturals\ +section \Star-transforms for the Hypernaturals\ theory NatStar -imports Star + imports Star begin lemma star_n_eq_starfun_whn: "star_n X = ( *f* X) whn" -by (simp add: hypnat_omega_def starfun_def star_of_def Ifun_star_n) + by (simp add: hypnat_omega_def starfun_def star_of_def Ifun_star_n) -lemma starset_n_Un: "*sn* (%n. (A n) Un (B n)) = *sn* A Un *sn* B" -apply (simp add: starset_n_def star_n_eq_starfun_whn Un_def) -apply (rule_tac x=whn in spec, transfer, simp) -done +lemma starset_n_Un: "*sn* (\n. (A n) \ (B n)) = *sn* A \ *sn* B" + apply (simp add: starset_n_def star_n_eq_starfun_whn Un_def) + apply (rule_tac x=whn in spec, transfer, simp) + done -lemma InternalSets_Un: - "[| X \ InternalSets; Y \ InternalSets |] - ==> (X Un Y) \ InternalSets" -by (auto simp add: InternalSets_def starset_n_Un [symmetric]) +lemma InternalSets_Un: "X \ InternalSets \ Y \ InternalSets \ X \ Y \ InternalSets" + by (auto simp add: InternalSets_def starset_n_Un [symmetric]) -lemma starset_n_Int: - "*sn* (%n. (A n) Int (B n)) = *sn* A Int *sn* B" -apply (simp add: starset_n_def star_n_eq_starfun_whn Int_def) -apply (rule_tac x=whn in spec, transfer, simp) -done +lemma starset_n_Int: "*sn* (\n. A n \ B n) = *sn* A \ *sn* B" + apply (simp add: starset_n_def star_n_eq_starfun_whn Int_def) + apply (rule_tac x=whn in spec, transfer, simp) + done -lemma InternalSets_Int: - "[| X \ InternalSets; Y \ InternalSets |] - ==> (X Int Y) \ InternalSets" -by (auto simp add: InternalSets_def starset_n_Int [symmetric]) +lemma InternalSets_Int: "X \ InternalSets \ Y \ InternalSets \ X \ Y \ InternalSets" + by (auto simp add: InternalSets_def starset_n_Int [symmetric]) -lemma starset_n_Compl: "*sn* ((%n. - A n)) = -( *sn* A)" -apply (simp add: starset_n_def star_n_eq_starfun_whn Compl_eq) -apply (rule_tac x=whn in spec, transfer, simp) -done +lemma starset_n_Compl: "*sn* ((\n. - A n)) = - ( *sn* A)" + apply (simp add: starset_n_def star_n_eq_starfun_whn Compl_eq) + apply (rule_tac x=whn in spec, transfer, simp) + done -lemma InternalSets_Compl: "X \ InternalSets ==> -X \ InternalSets" -by (auto simp add: InternalSets_def starset_n_Compl [symmetric]) +lemma InternalSets_Compl: "X \ InternalSets \ - X \ InternalSets" + by (auto simp add: InternalSets_def starset_n_Compl [symmetric]) -lemma starset_n_diff: "*sn* (%n. (A n) - (B n)) = *sn* A - *sn* B" -apply (simp add: starset_n_def star_n_eq_starfun_whn set_diff_eq) -apply (rule_tac x=whn in spec, transfer, simp) -done +lemma starset_n_diff: "*sn* (\n. (A n) - (B n)) = *sn* A - *sn* B" + apply (simp add: starset_n_def star_n_eq_starfun_whn set_diff_eq) + apply (rule_tac x=whn in spec, transfer, simp) + done -lemma InternalSets_diff: - "[| X \ InternalSets; Y \ InternalSets |] - ==> (X - Y) \ InternalSets" -by (auto simp add: InternalSets_def starset_n_diff [symmetric]) +lemma InternalSets_diff: "X \ InternalSets \ Y \ InternalSets \ X - Y \ InternalSets" + by (auto simp add: InternalSets_def starset_n_diff [symmetric]) lemma NatStar_SHNat_subset: "Nats \ *s* (UNIV:: nat set)" -by simp + by simp -lemma NatStar_hypreal_of_real_Int: - "*s* X Int Nats = hypnat_of_nat ` X" -by (auto simp add: SHNat_eq) +lemma NatStar_hypreal_of_real_Int: "*s* X Int Nats = hypnat_of_nat ` X" + by (auto simp add: SHNat_eq) -lemma starset_starset_n_eq: "*s* X = *sn* (%n. X)" -by (simp add: starset_n_starset) +lemma starset_starset_n_eq: "*s* X = *sn* (\n. X)" + by (simp add: starset_n_starset) lemma InternalSets_starset_n [simp]: "( *s* X) \ InternalSets" -by (auto simp add: InternalSets_def starset_starset_n_eq) + by (auto simp add: InternalSets_def starset_starset_n_eq) -lemma InternalSets_UNIV_diff: - "X \ InternalSets ==> UNIV - X \ InternalSets" -apply (subgoal_tac "UNIV - X = - X") -by (auto intro: InternalSets_Compl) +lemma InternalSets_UNIV_diff: "X \ InternalSets \ UNIV - X \ InternalSets" + apply (subgoal_tac "UNIV - X = - X") + apply (auto intro: InternalSets_Compl) + done -subsection\Nonstandard Extensions of Functions\ - -text\Example of transfer of a property from reals to hyperreals - --- used for limit comparison of sequences\ +subsection \Nonstandard Extensions of Functions\ -lemma starfun_le_mono: - "\n. N \ n --> f n \ g n - ==> \n. hypnat_of_nat N \ n --> ( *f* f) n \ ( *f* g) n" -by transfer +text \Example of transfer of a property from reals to hyperreals + --- used for limit comparison of sequences.\ + +lemma starfun_le_mono: "\n. N \ n \ f n \ g n \ + \n. hypnat_of_nat N \ n \ ( *f* f) n \ ( *f* g) n" + by transfer -(*****----- and another -----*****) +text \And another:\ lemma starfun_less_mono: - "\n. N \ n --> f n < g n - ==> \n. hypnat_of_nat N \ n --> ( *f* f) n < ( *f* g) n" -by transfer + "\n. N \ n \ f n < g n \ \n. hypnat_of_nat N \ n \ ( *f* f) n < ( *f* g) n" + by transfer -text\Nonstandard extension when we increment the argument by one\ +text \Nonstandard extension when we increment the argument by one.\ -lemma starfun_shift_one: - "!!N. ( *f* (%n. f (Suc n))) N = ( *f* f) (N + (1::hypnat))" -by (transfer, simp) +lemma starfun_shift_one: "\N. ( *f* (\n. f (Suc n))) N = ( *f* f) (N + (1::hypnat))" + by transfer simp -text\Nonstandard extension with absolute value\ - -lemma starfun_abs: "!!N. ( *f* (%n. \f n\)) N = \( *f* f) N\" -by (transfer, rule refl) +text \Nonstandard extension with absolute value.\ +lemma starfun_abs: "\N. ( *f* (\n. \f n\)) N = \( *f* f) N\" + by transfer (rule refl) -text\The hyperpow function as a nonstandard extension of realpow\ - -lemma starfun_pow: "!!N. ( *f* (%n. r ^ n)) N = (hypreal_of_real r) pow N" -by (transfer, rule refl) +text \The \hyperpow\ function as a nonstandard extension of \realpow\.\ +lemma starfun_pow: "\N. ( *f* (\n. r ^ n)) N = hypreal_of_real r pow N" + by transfer (rule refl) -lemma starfun_pow2: - "!!N. ( *f* (%n. (X n) ^ m)) N = ( *f* X) N pow hypnat_of_nat m" -by (transfer, rule refl) +lemma starfun_pow2: "\N. ( *f* (\n. X n ^ m)) N = ( *f* X) N pow hypnat_of_nat m" + by transfer (rule refl) -lemma starfun_pow3: "!!R. ( *f* (%r. r ^ n)) R = (R) pow hypnat_of_nat n" -by (transfer, rule refl) +lemma starfun_pow3: "\R. ( *f* (\r. r ^ n)) R = R pow hypnat_of_nat n" + by transfer (rule refl) -text\The @{term hypreal_of_hypnat} function as a nonstandard extension of - @{term real_of_nat}\ - +text \The @{term hypreal_of_hypnat} function as a nonstandard extension of + @{term real_of_nat}.\ lemma starfunNat_real_of_nat: "( *f* real) = hypreal_of_hypnat" -by transfer (simp add: fun_eq_iff) + by transfer (simp add: fun_eq_iff) lemma starfun_inverse_real_of_nat_eq: - "N \ HNatInfinite - ==> ( *f* (%x::nat. inverse(real x))) N = inverse(hypreal_of_hypnat N)" -apply (rule_tac f1 = inverse in starfun_o2 [THEN subst]) -apply (subgoal_tac "hypreal_of_hypnat N ~= 0") -apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat starfun_inverse_inverse) -done + "N \ HNatInfinite \ ( *f* (\x::nat. inverse (real x))) N = inverse (hypreal_of_hypnat N)" + apply (rule_tac f1 = inverse in starfun_o2 [THEN subst]) + apply (subgoal_tac "hypreal_of_hypnat N \ 0") + apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat) + done -text\Internal functions - some redundancy with *f* now\ +text \Internal functions -- some redundancy with \*f*\ now.\ -lemma starfun_n: "( *fn* f) (star_n X) = star_n (%n. f n (X n))" -by (simp add: starfun_n_def Ifun_star_n) - -text\Multiplication: \( *fn) x ( *gn) = *(fn x gn)\\ +lemma starfun_n: "( *fn* f) (star_n X) = star_n (\n. f n (X n))" + by (simp add: starfun_n_def Ifun_star_n) -lemma starfun_n_mult: - "( *fn* f) z * ( *fn* g) z = ( *fn* (% i x. f i x * g i x)) z" -apply (cases z) -apply (simp add: starfun_n star_n_mult) -done +text \Multiplication: \( *fn) x ( *gn) = *(fn x gn)\\ -text\Addition: \( *fn) + ( *gn) = *(fn + gn)\\ +lemma starfun_n_mult: "( *fn* f) z * ( *fn* g) z = ( *fn* (\i x. f i x * g i x)) z" + by (cases z) (simp add: starfun_n star_n_mult) -lemma starfun_n_add: - "( *fn* f) z + ( *fn* g) z = ( *fn* (%i x. f i x + g i x)) z" -apply (cases z) -apply (simp add: starfun_n star_n_add) -done +text \Addition: \( *fn) + ( *gn) = *(fn + gn)\\ +lemma starfun_n_add: "( *fn* f) z + ( *fn* g) z = ( *fn* (\i x. f i x + g i x)) z" + by (cases z) (simp add: starfun_n star_n_add) -text\Subtraction: \( *fn) - ( *gn) = *(fn + - gn)\\ - -lemma starfun_n_add_minus: - "( *fn* f) z + -( *fn* g) z = ( *fn* (%i x. f i x + -g i x)) z" -apply (cases z) -apply (simp add: starfun_n star_n_minus star_n_add) -done +text \Subtraction: \( *fn) - ( *gn) = *(fn + - gn)\\ +lemma starfun_n_add_minus: "( *fn* f) z + -( *fn* g) z = ( *fn* (\i x. f i x + -g i x)) z" + by (cases z) (simp add: starfun_n star_n_minus star_n_add) -text\Composition: \( *fn) o ( *gn) = *(fn o gn)\\ +text \Composition: \( *fn) \ ( *gn) = *(fn \ gn)\\ -lemma starfun_n_const_fun [simp]: - "( *fn* (%i x. k)) z = star_of k" -apply (cases z) -apply (simp add: starfun_n star_of_def) -done +lemma starfun_n_const_fun [simp]: "( *fn* (\i x. k)) z = star_of k" + by (cases z) (simp add: starfun_n star_of_def) -lemma starfun_n_minus: "- ( *fn* f) x = ( *fn* (%i x. - (f i) x)) x" -apply (cases x) -apply (simp add: starfun_n star_n_minus) -done +lemma starfun_n_minus: "- ( *fn* f) x = ( *fn* (\i x. - (f i) x)) x" + by (cases x) (simp add: starfun_n star_n_minus) -lemma starfun_n_eq [simp]: - "( *fn* f) (star_of n) = star_n (%i. f i n)" -by (simp add: starfun_n star_of_def) +lemma starfun_n_eq [simp]: "( *fn* f) (star_of n) = star_n (\i. f i n)" + by (simp add: starfun_n star_of_def) -lemma starfun_eq_iff: "(( *f* f) = ( *f* g)) = (f = g)" -by (transfer, rule refl) +lemma starfun_eq_iff: "(( *f* f) = ( *f* g)) \ f = g" + by transfer (rule refl) lemma starfunNat_inverse_real_of_nat_Infinitesimal [simp]: - "N \ HNatInfinite ==> ( *f* (%x. inverse (real x))) N \ Infinitesimal" -apply (rule_tac f1 = inverse in starfun_o2 [THEN subst]) -apply (subgoal_tac "hypreal_of_hypnat N ~= 0") -apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat) -done + "N \ HNatInfinite \ ( *f* (%x. inverse (real x))) N \ Infinitesimal" + apply (rule_tac f1 = inverse in starfun_o2 [THEN subst]) + apply (subgoal_tac "hypreal_of_hypnat N ~= 0") + apply (simp_all add: zero_less_HNatInfinite starfunNat_real_of_nat) + done -subsection\Nonstandard Characterization of Induction\ +subsection \Nonstandard Characterization of Induction\ lemma hypnat_induct_obj: - "!!n. (( *p* P) (0::hypnat) & - (\n. ( *p* P)(n) --> ( *p* P)(n + 1))) - --> ( *p* P)(n)" -by (transfer, induct_tac n, auto) + "\n. (( *p* P) (0::hypnat) \ (\n. ( *p* P) n \ ( *p* P) (n + 1))) \ ( *p* P) n" + by transfer (induct_tac n, auto) lemma hypnat_induct: - "!!n. [| ( *p* P) (0::hypnat); - !!n. ( *p* P)(n) ==> ( *p* P)(n + 1)|] - ==> ( *p* P)(n)" -by (transfer, induct_tac n, auto) + "\n. ( *p* P) (0::hypnat) \ (\n. ( *p* P) n \ ( *p* P) (n + 1)) \ ( *p* P) n" + by transfer (induct_tac n, auto) lemma starP2_eq_iff: "( *p2* (op =)) = (op =)" -by transfer (rule refl) + by transfer (rule refl) -lemma starP2_eq_iff2: "( *p2* (%x y. x = y)) X Y = (X = Y)" -by (simp add: starP2_eq_iff) +lemma starP2_eq_iff2: "( *p2* (\x y. x = y)) X Y \ X = Y" + by (simp add: starP2_eq_iff) -lemma nonempty_nat_set_Least_mem: - "c \ (S :: nat set) ==> (LEAST n. n \ S) \ S" -by (erule LeastI) +lemma nonempty_nat_set_Least_mem: "c \ S \ (LEAST n. n \ S) \ S" + for S :: "nat set" + by (erule LeastI) lemma nonempty_set_star_has_least: - "!!S::nat set star. Iset S \ {} ==> \n \ Iset S. \m \ Iset S. n \ m" -apply (transfer empty_def) -apply (rule_tac x="LEAST n. n \ S" in bexI) -apply (simp add: Least_le) -apply (rule LeastI_ex, auto) -done + "\S::nat set star. Iset S \ {} \ \n \ Iset S. \m \ Iset S. n \ m" + apply (transfer empty_def) + apply (rule_tac x="LEAST n. n \ S" in bexI) + apply (simp add: Least_le) + apply (rule LeastI_ex, auto) + done -lemma nonempty_InternalNatSet_has_least: - "[| (S::hypnat set) \ InternalSets; S \ {} |] ==> \n \ S. \m \ S. n \ m" -apply (clarsimp simp add: InternalSets_def starset_n_def) -apply (erule nonempty_set_star_has_least) -done +lemma nonempty_InternalNatSet_has_least: "S \ InternalSets \ S \ {} \ \n \ S. \m \ S. n \ m" + for S :: "hypnat set" + apply (clarsimp simp add: InternalSets_def starset_n_def) + apply (erule nonempty_set_star_has_least) + done -text\Goldblatt page 129 Thm 11.3.2\ +text \Goldblatt, page 129 Thm 11.3.2.\ lemma internal_induct_lemma: - "!!X::nat set star. [| (0::hypnat) \ Iset X; \n. n \ Iset X --> n + 1 \ Iset X |] - ==> Iset X = (UNIV:: hypnat set)" -apply (transfer UNIV_def) -apply (rule equalityI [OF subset_UNIV subsetI]) -apply (induct_tac x, auto) -done + "\X::nat set star. + (0::hypnat) \ Iset X \ \n. n \ Iset X \ n + 1 \ Iset X \ Iset X = (UNIV:: hypnat set)" + apply (transfer UNIV_def) + apply (rule equalityI [OF subset_UNIV subsetI]) + apply (induct_tac x, auto) + done lemma internal_induct: - "[| X \ InternalSets; (0::hypnat) \ X; \n. n \ X --> n + 1 \ X |] - ==> X = (UNIV:: hypnat set)" -apply (clarsimp simp add: InternalSets_def starset_n_def) -apply (erule (1) internal_induct_lemma) -done - + "X \ InternalSets \ (0::hypnat) \ X \ \n. n \ X \ n + 1 \ X \ X = (UNIV:: hypnat set)" + apply (clarsimp simp add: InternalSets_def starset_n_def) + apply (erule (1) internal_induct_lemma) + done end diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/Star.thy --- a/src/HOL/Nonstandard_Analysis/Star.thy Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/Star.thy Tue Nov 01 00:44:24 2016 +0100 @@ -4,340 +4,290 @@ Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 *) -section\Star-Transforms in Non-Standard Analysis\ +section \Star-Transforms in Non-Standard Analysis\ theory Star -imports NSA + imports NSA begin -definition - (* internal sets *) - starset_n :: "(nat => 'a set) => 'a star set" ("*sn* _" [80] 80) where - "*sn* As = Iset (star_n As)" +definition \ \internal sets\ + starset_n :: "(nat \ 'a set) \ 'a star set" ("*sn* _" [80] 80) + where "*sn* As = Iset (star_n As)" -definition - InternalSets :: "'a star set set" where - "InternalSets = {X. \As. X = *sn* As}" +definition InternalSets :: "'a star set set" + where "InternalSets = {X. \As. X = *sn* As}" -definition - (* nonstandard extension of function *) - is_starext :: "['a star => 'a star, 'a => 'a] => bool" where - "is_starext F f = - (\x y. \X \ Rep_star(x). \Y \ Rep_star(y). ((y = (F x)) = (eventually (\n. Y n = f(X n)) \)))" +definition \ \nonstandard extension of function\ + is_starext :: "('a star \ 'a star) \ ('a \ 'a) \ bool" + where "is_starext F f \ + (\x y. \X \ Rep_star x. \Y \ Rep_star y. y = F x \ eventually (\n. Y n = f(X n)) \)" -definition - (* internal functions *) - starfun_n :: "(nat => ('a => 'b)) => 'a star => 'b star" ("*fn* _" [80] 80) where - "*fn* F = Ifun (star_n F)" +definition \ \internal functions\ + starfun_n :: "(nat \ 'a \ 'b) \ 'a star \ 'b star" ("*fn* _" [80] 80) + where "*fn* F = Ifun (star_n F)" -definition - InternalFuns :: "('a star => 'b star) set" where - "InternalFuns = {X. \F. X = *fn* F}" +definition InternalFuns :: "('a star => 'b star) set" + where "InternalFuns = {X. \F. X = *fn* F}" -(*-------------------------------------------------------- - Preamble - Pulling "EX" over "ALL" - ---------------------------------------------------------*) +subsection \Preamble - Pulling \\\ over \\\\ -(* This proof does not need AC and was suggested by the - referee for the JCM Paper: let f(x) be least y such - that Q(x,y) -*) -lemma no_choice: "\x. \y. Q x y ==> \(f :: 'a => nat). \x. Q x (f x)" -apply (rule_tac x = "%x. LEAST y. Q x y" in exI) -apply (blast intro: LeastI) -done +text \This proof does not need AC and was suggested by the + referee for the JCM Paper: let \f x\ be least \y\ such + that \Q x y\.\ +lemma no_choice: "\x. \y. Q x y \ \f :: 'a \ nat. \x. Q x (f x)" + by (rule exI [where x = "\x. LEAST y. Q x y"]) (blast intro: LeastI) -subsection\Properties of the Star-transform Applied to Sets of Reals\ + +subsection \Properties of the Star-transform Applied to Sets of Reals\ -lemma STAR_star_of_image_subset: "star_of ` A <= *s* A" -by auto +lemma STAR_star_of_image_subset: "star_of ` A \ *s* A" + by auto -lemma STAR_hypreal_of_real_Int: "*s* X Int \ = hypreal_of_real ` X" -by (auto simp add: SReal_def) +lemma STAR_hypreal_of_real_Int: "*s* X \ \ = hypreal_of_real ` X" + by (auto simp add: SReal_def) -lemma STAR_star_of_Int: "*s* X Int Standard = star_of ` X" -by (auto simp add: Standard_def) +lemma STAR_star_of_Int: "*s* X \ Standard = star_of ` X" + by (auto simp add: Standard_def) -lemma lemma_not_hyprealA: "x \ hypreal_of_real ` A ==> \y \ A. x \ hypreal_of_real y" -by auto +lemma lemma_not_hyprealA: "x \ hypreal_of_real ` A \ \y \ A. x \ hypreal_of_real y" + by auto -lemma lemma_not_starA: "x \ star_of ` A ==> \y \ A. x \ star_of y" -by auto +lemma lemma_not_starA: "x \ star_of ` A \ \y \ A. x \ star_of y" + by auto lemma lemma_Compl_eq: "- {n. X n = xa} = {n. X n \ xa}" -by auto + by auto -lemma STAR_real_seq_to_hypreal: - "\n. (X n) \ M ==> star_n X \ *s* M" -apply (unfold starset_def star_of_def) -apply (simp add: Iset_star_n FreeUltrafilterNat.proper) -done +lemma STAR_real_seq_to_hypreal: "\n. (X n) \ M \ star_n X \ *s* M" + by (simp add: starset_def star_of_def Iset_star_n FreeUltrafilterNat.proper) lemma STAR_singleton: "*s* {x} = {star_of x}" -by simp + by simp -lemma STAR_not_mem: "x \ F ==> star_of x \ *s* F" -by transfer - -lemma STAR_subset_closed: "[| x : *s* A; A <= B |] ==> x : *s* B" -by (erule rev_subsetD, simp) +lemma STAR_not_mem: "x \ F \ star_of x \ *s* F" + by transfer -text\Nonstandard extension of a set (defined using a constant - sequence) as a special case of an internal set\ +lemma STAR_subset_closed: "x \ *s* A \ A \ B \ x \ *s* B" + by (erule rev_subsetD) simp -lemma starset_n_starset: "\n. (As n = A) ==> *sn* As = *s* A" -apply (drule fun_eq_iff [THEN iffD2]) -apply (simp add: starset_n_def starset_def star_of_def) -done +text \Nonstandard extension of a set (defined using a constant + sequence) as a special case of an internal set.\ +lemma starset_n_starset: "\n. As n = A \ *sn* As = *s* A" + by (drule fun_eq_iff [THEN iffD2]) (simp add: starset_n_def starset_def star_of_def) -(*----------------------------------------------------------------*) -(* Theorems about nonstandard extensions of functions *) -(*----------------------------------------------------------------*) +subsection \Theorems about nonstandard extensions of functions\ -(*----------------------------------------------------------------*) -(* Nonstandard extension of a function (defined using a *) -(* constant sequence) as a special case of an internal function *) -(*----------------------------------------------------------------*) +text \Nonstandard extension of a function (defined using a + constant sequence) as a special case of an internal function.\ -lemma starfun_n_starfun: "\n. (F n = f) ==> *fn* F = *f* f" -apply (drule fun_eq_iff [THEN iffD2]) -apply (simp add: starfun_n_def starfun_def star_of_def) -done - +lemma starfun_n_starfun: "\n. F n = f \ *fn* F = *f* f" + apply (drule fun_eq_iff [THEN iffD2]) + apply (simp add: starfun_n_def starfun_def star_of_def) + done -(* - Prove that abs for hypreal is a nonstandard extension of abs for real w/o - use of congruence property (proved after this for general - nonstandard extensions of real valued functions). +text \Prove that \abs\ for hypreal is a nonstandard extension of abs for real w/o + use of congruence property (proved after this for general + nonstandard extensions of real valued functions). - Proof now Uses the ultrafilter tactic! -*) + Proof now Uses the ultrafilter tactic!\ lemma hrabs_is_starext_rabs: "is_starext abs abs" -apply (simp add: is_starext_def, safe) -apply (rule_tac x=x in star_cases) -apply (rule_tac x=y in star_cases) -apply (unfold star_n_def, auto) -apply (rule bexI, rule_tac [2] lemma_starrel_refl) -apply (rule bexI, rule_tac [2] lemma_starrel_refl) -apply (fold star_n_def) -apply (unfold star_abs_def starfun_def star_of_def) -apply (simp add: Ifun_star_n star_n_eq_iff) -done + apply (simp add: is_starext_def, safe) + apply (rule_tac x=x in star_cases) + apply (rule_tac x=y in star_cases) + apply (unfold star_n_def, auto) + apply (rule bexI, rule_tac [2] lemma_starrel_refl) + apply (rule bexI, rule_tac [2] lemma_starrel_refl) + apply (fold star_n_def) + apply (unfold star_abs_def starfun_def star_of_def) + apply (simp add: Ifun_star_n star_n_eq_iff) + done -text\Nonstandard extension of functions\ -lemma starfun: - "( *f* f) (star_n X) = star_n (%n. f (X n))" -by (rule starfun_star_n) +text \Nonstandard extension of functions.\ + +lemma starfun: "( *f* f) (star_n X) = star_n (\n. f (X n))" + by (rule starfun_star_n) -lemma starfun_if_eq: - "!!w. w \ star_of x - ==> ( *f* (\z. if z = x then a else g z)) w = ( *f* g) w" -by (transfer, simp) +lemma starfun_if_eq: "\w. w \ star_of x \ ( *f* (\z. if z = x then a else g z)) w = ( *f* g) w" + by transfer simp -(*------------------------------------------- - multiplication: ( *f) x ( *g) = *(f x g) - ------------------------------------------*) -lemma starfun_mult: "!!x. ( *f* f) x * ( *f* g) x = ( *f* (%x. f x * g x)) x" -by (transfer, rule refl) +text \Multiplication: \( *f) x ( *g) = *(f x g)\\ +lemma starfun_mult: "\x. ( *f* f) x * ( *f* g) x = ( *f* (\x. f x * g x)) x" + by transfer (rule refl) declare starfun_mult [symmetric, simp] -(*--------------------------------------- - addition: ( *f) + ( *g) = *(f + g) - ---------------------------------------*) -lemma starfun_add: "!!x. ( *f* f) x + ( *f* g) x = ( *f* (%x. f x + g x)) x" -by (transfer, rule refl) +text \Addition: \( *f) + ( *g) = *(f + g)\\ +lemma starfun_add: "\x. ( *f* f) x + ( *f* g) x = ( *f* (\x. f x + g x)) x" + by transfer (rule refl) declare starfun_add [symmetric, simp] -(*-------------------------------------------- - subtraction: ( *f) + -( *g) = *(f + -g) - -------------------------------------------*) -lemma starfun_minus: "!!x. - ( *f* f) x = ( *f* (%x. - f x)) x" -by (transfer, rule refl) +text \Subtraction: \( *f) + -( *g) = *(f + -g)\\ +lemma starfun_minus: "\x. - ( *f* f) x = ( *f* (\x. - f x)) x" + by transfer (rule refl) declare starfun_minus [symmetric, simp] (*FIXME: delete*) -lemma starfun_add_minus: "!!x. ( *f* f) x + -( *f* g) x = ( *f* (%x. f x + -g x)) x" -by (transfer, rule refl) +lemma starfun_add_minus: "\x. ( *f* f) x + -( *f* g) x = ( *f* (\x. f x + -g x)) x" + by transfer (rule refl) declare starfun_add_minus [symmetric, simp] -lemma starfun_diff: "!!x. ( *f* f) x - ( *f* g) x = ( *f* (%x. f x - g x)) x" -by (transfer, rule refl) +lemma starfun_diff: "\x. ( *f* f) x - ( *f* g) x = ( *f* (\x. f x - g x)) x" + by transfer (rule refl) declare starfun_diff [symmetric, simp] -(*-------------------------------------- - composition: ( *f) o ( *g) = *(f o g) - ---------------------------------------*) +text \Composition: \( *f) \ ( *g) = *(f \ g)\\ +lemma starfun_o2: "(\x. ( *f* f) (( *f* g) x)) = *f* (\x. f (g x))" + by transfer (rule refl) -lemma starfun_o2: "(%x. ( *f* f) (( *f* g) x)) = *f* (%x. f (g x))" -by (transfer, rule refl) - -lemma starfun_o: "( *f* f) o ( *f* g) = ( *f* (f o g))" -by (transfer o_def, rule refl) +lemma starfun_o: "( *f* f) \ ( *f* g) = ( *f* (f \ g))" + by (transfer o_def) (rule refl) -text\NS extension of constant function\ -lemma starfun_const_fun [simp]: "!!x. ( *f* (%x. k)) x = star_of k" -by (transfer, rule refl) - -text\the NS extension of the identity function\ +text \NS extension of constant function.\ +lemma starfun_const_fun [simp]: "\x. ( *f* (\x. k)) x = star_of k" + by transfer (rule refl) -lemma starfun_Id [simp]: "!!x. ( *f* (%x. x)) x = x" -by (transfer, rule refl) +text \The NS extension of the identity function.\ +lemma starfun_Id [simp]: "\x. ( *f* (\x. x)) x = x" + by transfer (rule refl) -(* this is trivial, given starfun_Id *) -lemma starfun_Idfun_approx: - "x \ star_of a ==> ( *f* (%x. x)) x \ star_of a" -by (simp only: starfun_Id) +text \This is trivial, given \starfun_Id\.\ +lemma starfun_Idfun_approx: "x \ star_of a \ ( *f* (\x. x)) x \ star_of a" + by (simp only: starfun_Id) -text\The Star-function is a (nonstandard) extension of the function\ - +text \The Star-function is a (nonstandard) extension of the function.\ lemma is_starext_starfun: "is_starext ( *f* f) f" -apply (simp add: is_starext_def, auto) -apply (rule_tac x = x in star_cases) -apply (rule_tac x = y in star_cases) -apply (auto intro!: bexI [OF _ Rep_star_star_n] - simp add: starfun star_n_eq_iff) -done - -text\Any nonstandard extension is in fact the Star-function\ + apply (auto simp: is_starext_def) + apply (rule_tac x = x in star_cases) + apply (rule_tac x = y in star_cases) + apply (auto intro!: bexI [OF _ Rep_star_star_n] simp: starfun star_n_eq_iff) + done -lemma is_starfun_starext: "is_starext F f ==> F = *f* f" -apply (simp add: is_starext_def) -apply (rule ext) -apply (rule_tac x = x in star_cases) -apply (drule_tac x = x in spec) -apply (drule_tac x = "( *f* f) x" in spec) -apply (auto simp add: starfun_star_n) -apply (simp add: star_n_eq_iff [symmetric]) -apply (simp add: starfun_star_n [of f, symmetric]) -done +text \Any nonstandard extension is in fact the Star-function.\ +lemma is_starfun_starext: "is_starext F f \ F = *f* f" + apply (simp add: is_starext_def) + apply (rule ext) + apply (rule_tac x = x in star_cases) + apply (drule_tac x = x in spec) + apply (drule_tac x = "( *f* f) x" in spec) + apply (auto simp add: starfun_star_n) + apply (simp add: star_n_eq_iff [symmetric]) + apply (simp add: starfun_star_n [of f, symmetric]) + done -lemma is_starext_starfun_iff: "(is_starext F f) = (F = *f* f)" -by (blast intro: is_starfun_starext is_starext_starfun) +lemma is_starext_starfun_iff: "is_starext F f \ F = *f* f" + by (blast intro: is_starfun_starext is_starext_starfun) -text\extented function has same solution as its standard - version for real arguments. i.e they are the same - for all real arguments\ +text \Extented function has same solution as its standard version + for real arguments. i.e they are the same for all real arguments.\ lemma starfun_eq: "( *f* f) (star_of a) = star_of (f a)" -by (rule starfun_star_of) + by (rule starfun_star_of) lemma starfun_approx: "( *f* f) (star_of a) \ star_of (f a)" -by simp + by simp -(* useful for NS definition of derivatives *) -lemma starfun_lambda_cancel: - "!!x'. ( *f* (%h. f (x + h))) x' = ( *f* f) (star_of x + x')" -by (transfer, rule refl) +text \Useful for NS definition of derivatives.\ +lemma starfun_lambda_cancel: "\x'. ( *f* (\h. f (x + h))) x' = ( *f* f) (star_of x + x')" + by transfer (rule refl) -lemma starfun_lambda_cancel2: - "( *f* (%h. f(g(x + h)))) x' = ( *f* (f o g)) (star_of x + x')" -by (unfold o_def, rule starfun_lambda_cancel) +lemma starfun_lambda_cancel2: "( *f* (\h. f (g (x + h)))) x' = ( *f* (f \ g)) (star_of x + x')" + unfolding o_def by (rule starfun_lambda_cancel) lemma starfun_mult_HFinite_approx: - fixes l m :: "'a::real_normed_algebra star" - shows "[| ( *f* f) x \ l; ( *f* g) x \ m; - l: HFinite; m: HFinite - |] ==> ( *f* (%x. f x * g x)) x \ l * m" -apply (drule (3) approx_mult_HFinite) -apply (auto intro: approx_HFinite [OF _ approx_sym]) -done + "( *f* f) x \ l \ ( *f* g) x \ m \ l \ HFinite \ m \ HFinite \ + ( *f* (\x. f x * g x)) x \ l * m" + for l m :: "'a::real_normed_algebra star" + apply (drule (3) approx_mult_HFinite) + apply (auto intro: approx_HFinite [OF _ approx_sym]) + done -lemma starfun_add_approx: "[| ( *f* f) x \ l; ( *f* g) x \ m - |] ==> ( *f* (%x. f x + g x)) x \ l + m" -by (auto intro: approx_add) +lemma starfun_add_approx: "( *f* f) x \ l \ ( *f* g) x \ m \ ( *f* (%x. f x + g x)) x \ l + m" + by (auto intro: approx_add) -text\Examples: hrabs is nonstandard extension of rabs - inverse is nonstandard extension of inverse\ +text \Examples: \hrabs\ is nonstandard extension of \rabs\, + \inverse\ is nonstandard extension of \inverse\.\ -(* can be proved easily using theorem "starfun" and *) -(* properties of ultrafilter as for inverse below we *) -(* use the theorem we proved above instead *) +text \Can be proved easily using theorem \starfun\ and + properties of ultrafilter as for inverse below we + use the theorem we proved above instead.\ lemma starfun_rabs_hrabs: "*f* abs = abs" -by (simp only: star_abs_def) + by (simp only: star_abs_def) -lemma starfun_inverse_inverse [simp]: "( *f* inverse) x = inverse(x)" -by (simp only: star_inverse_def) +lemma starfun_inverse_inverse [simp]: "( *f* inverse) x = inverse x" + by (simp only: star_inverse_def) -lemma starfun_inverse: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x" -by (transfer, rule refl) +lemma starfun_inverse: "\x. inverse (( *f* f) x) = ( *f* (\x. inverse (f x))) x" + by transfer (rule refl) declare starfun_inverse [symmetric, simp] -lemma starfun_divide: "!!x. ( *f* f) x / ( *f* g) x = ( *f* (%x. f x / g x)) x" -by (transfer, rule refl) +lemma starfun_divide: "\x. ( *f* f) x / ( *f* g) x = ( *f* (\x. f x / g x)) x" + by transfer (rule refl) declare starfun_divide [symmetric, simp] -lemma starfun_inverse2: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x" -by (transfer, rule refl) +lemma starfun_inverse2: "\x. inverse (( *f* f) x) = ( *f* (\x. inverse (f x))) x" + by transfer (rule refl) -text\General lemma/theorem needed for proofs in elementary - topology of the reals\ -lemma starfun_mem_starset: - "!!x. ( *f* f) x : *s* A ==> x : *s* {x. f x \ A}" -by (transfer, simp) +text \General lemma/theorem needed for proofs in elementary topology of the reals.\ +lemma starfun_mem_starset: "\x. ( *f* f) x : *s* A \ x \ *s* {x. f x \ A}" + by transfer simp -text\Alternative definition for hrabs with rabs function - applied entrywise to equivalence class representative. - This is easily proved using starfun and ns extension thm\ -lemma hypreal_hrabs: "\star_n X\ = star_n (%n. \X n\)" -by (simp only: starfun_rabs_hrabs [symmetric] starfun) +text \Alternative definition for \hrabs\ with \rabs\ function applied + entrywise to equivalence class representative. + This is easily proved using @{thm [source] starfun} and ns extension thm.\ +lemma hypreal_hrabs: "\star_n X\ = star_n (\n. \X n\)" + by (simp only: starfun_rabs_hrabs [symmetric] starfun) -text\nonstandard extension of set through nonstandard extension - of rabs function i.e hrabs. A more general result should be - where we replace rabs by some arbitrary function f and hrabs +text \Nonstandard extension of set through nonstandard extension + of \rabs\ function i.e. \hrabs\. A more general result should be + where we replace \rabs\ by some arbitrary function \f\ and \hrabs\ by its NS extenson. See second NS set extension below.\ -lemma STAR_rabs_add_minus: - "*s* {x. \x + - y\ < r} = {x. \x + -star_of y\ < star_of r}" -by (transfer, rule refl) +lemma STAR_rabs_add_minus: "*s* {x. \x + - y\ < r} = {x. \x + -star_of y\ < star_of r}" + by transfer (rule refl) lemma STAR_starfun_rabs_add_minus: - "*s* {x. \f x + - y\ < r} = - {x. \( *f* f) x + -star_of y\ < star_of r}" -by (transfer, rule refl) + "*s* {x. \f x + - y\ < r} = {x. \( *f* f) x + -star_of y\ < star_of r}" + by transfer (rule refl) -text\Another characterization of Infinitesimal and one of \\\ relation. - In this theory since \hypreal_hrabs\ proved here. Maybe - move both theorems??\ +text \Another characterization of Infinitesimal and one of \\\ relation. + In this theory since \hypreal_hrabs\ proved here. Maybe move both theorems??\ lemma Infinitesimal_FreeUltrafilterNat_iff2: - "(star_n X \ Infinitesimal) = (\m. eventually (\n. norm(X n) < inverse(real(Suc m))) \)" -by (simp add: Infinitesimal_hypreal_of_nat_iff star_of_def - hnorm_def star_of_nat_def starfun_star_n - star_n_inverse star_n_less) + "star_n X \ Infinitesimal \ (\m. eventually (\n. norm (X n) < inverse (real (Suc m))) \)" + by (simp add: Infinitesimal_hypreal_of_nat_iff star_of_def hnorm_def + star_of_nat_def starfun_star_n star_n_inverse star_n_less) lemma HNatInfinite_inverse_Infinitesimal [simp]: - "n \ HNatInfinite ==> inverse (hypreal_of_hypnat n) \ Infinitesimal" -apply (cases n) -apply (auto simp add: of_hypnat_def starfun_star_n star_n_inverse real_norm_def - HNatInfinite_FreeUltrafilterNat_iff - Infinitesimal_FreeUltrafilterNat_iff2) -apply (drule_tac x="Suc m" in spec) -apply (auto elim!: eventually_mono) -done + "n \ HNatInfinite \ inverse (hypreal_of_hypnat n) \ Infinitesimal" + apply (cases n) + apply (auto simp: of_hypnat_def starfun_star_n star_n_inverse + HNatInfinite_FreeUltrafilterNat_iff Infinitesimal_FreeUltrafilterNat_iff2) + apply (drule_tac x = "Suc m" in spec) + apply (auto elim!: eventually_mono) + done -lemma approx_FreeUltrafilterNat_iff: "star_n X \ star_n Y = - (\r>0. eventually (\n. norm (X n - Y n) < r) \)" -apply (subst approx_minus_iff) -apply (rule mem_infmal_iff [THEN subst]) -apply (simp add: star_n_diff) -apply (simp add: Infinitesimal_FreeUltrafilterNat_iff) -done +lemma approx_FreeUltrafilterNat_iff: + "star_n X \ star_n Y \ (\r>0. eventually (\n. norm (X n - Y n) < r) \)" + apply (subst approx_minus_iff) + apply (rule mem_infmal_iff [THEN subst]) + apply (simp add: star_n_diff) + apply (simp add: Infinitesimal_FreeUltrafilterNat_iff) + done -lemma approx_FreeUltrafilterNat_iff2: "star_n X \ star_n Y = - (\m. eventually (\n. norm (X n - Y n) < inverse(real(Suc m))) \)" -apply (subst approx_minus_iff) -apply (rule mem_infmal_iff [THEN subst]) -apply (simp add: star_n_diff) -apply (simp add: Infinitesimal_FreeUltrafilterNat_iff2) -done +lemma approx_FreeUltrafilterNat_iff2: + "star_n X \ star_n Y \ (\m. eventually (\n. norm (X n - Y n) < inverse (real (Suc m))) \)" + apply (subst approx_minus_iff) + apply (rule mem_infmal_iff [THEN subst]) + apply (simp add: star_n_diff) + apply (simp add: Infinitesimal_FreeUltrafilterNat_iff2) + done lemma inj_starfun: "inj starfun" -apply (rule inj_onI) -apply (rule ext, rule ccontr) -apply (drule_tac x = "star_n (%n. xa)" in fun_cong) -apply (auto simp add: starfun star_n_eq_iff FreeUltrafilterNat.proper) -done + apply (rule inj_onI) + apply (rule ext, rule ccontr) + apply (drule_tac x = "star_n (\n. xa)" in fun_cong) + apply (auto simp add: starfun star_n_eq_iff FreeUltrafilterNat.proper) + done end diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/StarDef.thy --- a/src/HOL/Nonstandard_Analysis/StarDef.thy Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/StarDef.thy Tue Nov 01 00:44:24 2016 +0100 @@ -5,54 +5,58 @@ section \Construction of Star Types Using Ultrafilters\ theory StarDef -imports Free_Ultrafilter + imports Free_Ultrafilter begin subsection \A Free Ultrafilter over the Naturals\ -definition - FreeUltrafilterNat :: "nat filter" ("\") where - "\ = (SOME U. freeultrafilter U)" +definition FreeUltrafilterNat :: "nat filter" ("\") + where "\ = (SOME U. freeultrafilter U)" lemma freeultrafilter_FreeUltrafilterNat: "freeultrafilter \" -apply (unfold FreeUltrafilterNat_def) -apply (rule someI_ex) -apply (rule freeultrafilter_Ex) -apply (rule infinite_UNIV_nat) -done + apply (unfold FreeUltrafilterNat_def) + apply (rule someI_ex) + apply (rule freeultrafilter_Ex) + apply (rule infinite_UNIV_nat) + done interpretation FreeUltrafilterNat: freeultrafilter FreeUltrafilterNat -by (rule freeultrafilter_FreeUltrafilterNat) + by (rule freeultrafilter_FreeUltrafilterNat) + subsection \Definition of \star\ type constructor\ -definition - starrel :: "((nat \ 'a) \ (nat \ 'a)) set" where - "starrel = {(X,Y). eventually (\n. X n = Y n) \}" +definition starrel :: "((nat \ 'a) \ (nat \ 'a)) set" + where "starrel = {(X, Y). eventually (\n. X n = Y n) \}" definition "star = (UNIV :: (nat \ 'a) set) // starrel" typedef 'a star = "star :: (nat \ 'a) set set" - unfolding star_def by (auto intro: quotientI) + by (auto simp: star_def intro: quotientI) -definition - star_n :: "(nat \ 'a) \ 'a star" where - "star_n X = Abs_star (starrel `` {X})" +definition star_n :: "(nat \ 'a) \ 'a star" + where "star_n X = Abs_star (starrel `` {X})" theorem star_cases [case_names star_n, cases type: star]: - "(\X. x = star_n X \ P) \ P" -by (cases x, unfold star_n_def star_def, erule quotientE, fast) + obtains X where "x = star_n X" + by (cases x) (auto simp: star_n_def star_def elim: quotientE) -lemma all_star_eq: "(\x. P x) = (\X. P (star_n X))" -by (auto, rule_tac x=x in star_cases, simp) +lemma all_star_eq: "(\x. P x) \ (\X. P (star_n X))" + apply auto + apply (rule_tac x = x in star_cases) + apply simp + done -lemma ex_star_eq: "(\x. P x) = (\X. P (star_n X))" -by (auto, rule_tac x=x in star_cases, auto) +lemma ex_star_eq: "(\x. P x) \ (\X. P (star_n X))" + apply auto + apply (rule_tac x=x in star_cases) + apply auto + done -text \Proving that @{term starrel} is an equivalence relation\ +text \Proving that @{term starrel} is an equivalence relation.\ -lemma starrel_iff [iff]: "((X,Y) \ starrel) = (eventually (\n. X n = Y n) \)" -by (simp add: starrel_def) +lemma starrel_iff [iff]: "(X, Y) \ starrel \ eventually (\n. X n = Y n) \" + by (simp add: starrel_def) lemma equiv_starrel: "equiv UNIV starrel" proof (rule equivI) @@ -61,356 +65,322 @@ show "trans starrel" by (intro transI) (auto elim: eventually_elim2) qed -lemmas equiv_starrel_iff = - eq_equiv_class_iff [OF equiv_starrel UNIV_I UNIV_I] +lemmas equiv_starrel_iff = eq_equiv_class_iff [OF equiv_starrel UNIV_I UNIV_I] lemma starrel_in_star: "starrel``{x} \ star" -by (simp add: star_def quotientI) + by (simp add: star_def quotientI) -lemma star_n_eq_iff: "(star_n X = star_n Y) = (eventually (\n. X n = Y n) \)" -by (simp add: star_n_def Abs_star_inject starrel_in_star equiv_starrel_iff) +lemma star_n_eq_iff: "star_n X = star_n Y \ eventually (\n. X n = Y n) \" + by (simp add: star_n_def Abs_star_inject starrel_in_star equiv_starrel_iff) subsection \Transfer principle\ text \This introduction rule starts each transfer proof.\ -lemma transfer_start: - "P \ eventually (\n. Q) \ \ Trueprop P \ Trueprop Q" +lemma transfer_start: "P \ eventually (\n. Q) \ \ Trueprop P \ Trueprop Q" by (simp add: FreeUltrafilterNat.proper) text \Standard principles that play a central role in the transfer tactic.\ -definition - Ifun :: "('a \ 'b) star \ 'a star \ 'b star" ("_ \ _" [300,301] 300) where - "Ifun f \ \x. Abs_star - (\F\Rep_star f. \X\Rep_star x. starrel``{\n. F n (X n)})" +definition Ifun :: "('a \ 'b) star \ 'a star \ 'b star" ("_ \ _" [300,301] 300) + where "Ifun f \ + \x. Abs_star (\F\Rep_star f. \X\Rep_star x. starrel``{\n. F n (X n)})" -lemma Ifun_congruent2: - "congruent2 starrel starrel (\F X. starrel``{\n. F n (X n)})" -by (auto simp add: congruent2_def equiv_starrel_iff elim!: eventually_rev_mp) +lemma Ifun_congruent2: "congruent2 starrel starrel (\F X. starrel``{\n. F n (X n)})" + by (auto simp add: congruent2_def equiv_starrel_iff elim!: eventually_rev_mp) lemma Ifun_star_n: "star_n F \ star_n X = star_n (\n. F n (X n))" -by (simp add: Ifun_def star_n_def Abs_star_inverse starrel_in_star - UN_equiv_class2 [OF equiv_starrel equiv_starrel Ifun_congruent2]) + by (simp add: Ifun_def star_n_def Abs_star_inverse starrel_in_star + UN_equiv_class2 [OF equiv_starrel equiv_starrel Ifun_congruent2]) -lemma transfer_Ifun: - "\f \ star_n F; x \ star_n X\ \ f \ x \ star_n (\n. F n (X n))" -by (simp only: Ifun_star_n) +lemma transfer_Ifun: "f \ star_n F \ x \ star_n X \ f \ x \ star_n (\n. F n (X n))" + by (simp only: Ifun_star_n) -definition - star_of :: "'a \ 'a star" where - "star_of x == star_n (\n. x)" +definition star_of :: "'a \ 'a star" + where "star_of x \ star_n (\n. x)" text \Initialize transfer tactic.\ ML_file "transfer.ML" -method_setup transfer = \ - Attrib.thms >> (fn ths => fn ctxt => - SIMPLE_METHOD' (Transfer_Principle.transfer_tac ctxt ths)) -\ "transfer principle" +method_setup transfer = + \Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (Transfer_Principle.transfer_tac ctxt ths))\ + "transfer principle" text \Transfer introduction rules.\ lemma transfer_ex [transfer_intro]: - "\\X. p (star_n X) \ eventually (\n. P n (X n)) \\ - \ \x::'a star. p x \ eventually (\n. \x. P n x) \" -by (simp only: ex_star_eq eventually_ex) + "(\X. p (star_n X) \ eventually (\n. P n (X n)) \) \ + \x::'a star. p x \ eventually (\n. \x. P n x) \" + by (simp only: ex_star_eq eventually_ex) lemma transfer_all [transfer_intro]: - "\\X. p (star_n X) \ eventually (\n. P n (X n)) \\ - \ \x::'a star. p x \ eventually (\n. \x. P n x) \" -by (simp only: all_star_eq FreeUltrafilterNat.eventually_all_iff) + "(\X. p (star_n X) \ eventually (\n. P n (X n)) \) \ + \x::'a star. p x \ eventually (\n. \x. P n x) \" + by (simp only: all_star_eq FreeUltrafilterNat.eventually_all_iff) -lemma transfer_not [transfer_intro]: - "\p \ eventually P \\ \ \ p \ eventually (\n. \ P n) \" -by (simp only: FreeUltrafilterNat.eventually_not_iff) +lemma transfer_not [transfer_intro]: "p \ eventually P \ \ \ p \ eventually (\n. \ P n) \" + by (simp only: FreeUltrafilterNat.eventually_not_iff) lemma transfer_conj [transfer_intro]: - "\p \ eventually P \; q \ eventually Q \\ - \ p \ q \ eventually (\n. P n \ Q n) \" -by (simp only: eventually_conj_iff) + "p \ eventually P \ \ q \ eventually Q \ \ p \ q \ eventually (\n. P n \ Q n) \" + by (simp only: eventually_conj_iff) lemma transfer_disj [transfer_intro]: - "\p \ eventually P \; q \ eventually Q \\ - \ p \ q \ eventually (\n. P n \ Q n) \" -by (simp only: FreeUltrafilterNat.eventually_disj_iff) + "p \ eventually P \ \ q \ eventually Q \ \ p \ q \ eventually (\n. P n \ Q n) \" + by (simp only: FreeUltrafilterNat.eventually_disj_iff) lemma transfer_imp [transfer_intro]: - "\p \ eventually P \; q \ eventually Q \\ - \ p \ q \ eventually (\n. P n \ Q n) \" -by (simp only: FreeUltrafilterNat.eventually_imp_iff) + "p \ eventually P \ \ q \ eventually Q \ \ p \ q \ eventually (\n. P n \ Q n) \" + by (simp only: FreeUltrafilterNat.eventually_imp_iff) lemma transfer_iff [transfer_intro]: - "\p \ eventually P \; q \ eventually Q \\ - \ p = q \ eventually (\n. P n = Q n) \" -by (simp only: FreeUltrafilterNat.eventually_iff_iff) + "p \ eventually P \ \ q \ eventually Q \ \ p = q \ eventually (\n. P n = Q n) \" + by (simp only: FreeUltrafilterNat.eventually_iff_iff) lemma transfer_if_bool [transfer_intro]: - "\p \ eventually P \; x \ eventually X \; y \ eventually Y \\ - \ (if p then x else y) \ eventually (\n. if P n then X n else Y n) \" -by (simp only: if_bool_eq_conj transfer_conj transfer_imp transfer_not) + "p \ eventually P \ \ x \ eventually X \ \ y \ eventually Y \ \ + (if p then x else y) \ eventually (\n. if P n then X n else Y n) \" + by (simp only: if_bool_eq_conj transfer_conj transfer_imp transfer_not) lemma transfer_eq [transfer_intro]: - "\x \ star_n X; y \ star_n Y\ \ x = y \ eventually (\n. X n = Y n) \" -by (simp only: star_n_eq_iff) + "x \ star_n X \ y \ star_n Y \ x = y \ eventually (\n. X n = Y n) \" + by (simp only: star_n_eq_iff) lemma transfer_if [transfer_intro]: - "\p \ eventually (\n. P n) \; x \ star_n X; y \ star_n Y\ - \ (if p then x else y) \ star_n (\n. if P n then X n else Y n)" -apply (rule eq_reflection) -apply (auto simp add: star_n_eq_iff transfer_not elim!: eventually_mono) -done + "p \ eventually (\n. P n) \ \ x \ star_n X \ y \ star_n Y \ + (if p then x else y) \ star_n (\n. if P n then X n else Y n)" + by (rule eq_reflection) (auto simp: star_n_eq_iff transfer_not elim!: eventually_mono) lemma transfer_fun_eq [transfer_intro]: - "\\X. f (star_n X) = g (star_n X) - \ eventually (\n. F n (X n) = G n (X n)) \\ - \ f = g \ eventually (\n. F n = G n) \" -by (simp only: fun_eq_iff transfer_all) + "(\X. f (star_n X) = g (star_n X) \ eventually (\n. F n (X n) = G n (X n)) \) \ + f = g \ eventually (\n. F n = G n) \" + by (simp only: fun_eq_iff transfer_all) lemma transfer_star_n [transfer_intro]: "star_n X \ star_n (\n. X n)" -by (rule reflexive) + by (rule reflexive) lemma transfer_bool [transfer_intro]: "p \ eventually (\n. p) \" -by (simp add: FreeUltrafilterNat.proper) + by (simp add: FreeUltrafilterNat.proper) subsection \Standard elements\ -definition - Standard :: "'a star set" where - "Standard = range star_of" +definition Standard :: "'a star set" + where "Standard = range star_of" -text \Transfer tactic should remove occurrences of @{term star_of}\ +text \Transfer tactic should remove occurrences of @{term star_of}.\ setup \Transfer_Principle.add_const @{const_name star_of}\ -lemma star_of_inject: "(star_of x = star_of y) = (x = y)" -by (transfer, rule refl) +lemma star_of_inject: "star_of x = star_of y \ x = y" + by transfer (rule refl) lemma Standard_star_of [simp]: "star_of x \ Standard" -by (simp add: Standard_def) + by (simp add: Standard_def) + subsection \Internal functions\ -text \Transfer tactic should remove occurrences of @{term Ifun}\ +text \Transfer tactic should remove occurrences of @{term Ifun}.\ setup \Transfer_Principle.add_const @{const_name Ifun}\ lemma Ifun_star_of [simp]: "star_of f \ star_of x = star_of (f x)" -by (transfer, rule refl) + by transfer (rule refl) -lemma Standard_Ifun [simp]: - "\f \ Standard; x \ Standard\ \ f \ x \ Standard" -by (auto simp add: Standard_def) +lemma Standard_Ifun [simp]: "f \ Standard \ x \ Standard \ f \ x \ Standard" + by (auto simp add: Standard_def) -text \Nonstandard extensions of functions\ -definition - starfun :: "('a \ 'b) \ ('a star \ 'b star)" ("*f* _" [80] 80) where - "starfun f == \x. star_of f \ x" +text \Nonstandard extensions of functions.\ -definition - starfun2 :: "('a \ 'b \ 'c) \ ('a star \ 'b star \ 'c star)" - ("*f2* _" [80] 80) where - "starfun2 f == \x y. star_of f \ x \ y" +definition starfun :: "('a \ 'b) \ 'a star \ 'b star" ("*f* _" [80] 80) + where "starfun f \ \x. star_of f \ x" + +definition starfun2 :: "('a \ 'b \ 'c) \ 'a star \ 'b star \ 'c star" ("*f2* _" [80] 80) + where "starfun2 f \ \x y. star_of f \ x \ y" declare starfun_def [transfer_unfold] declare starfun2_def [transfer_unfold] lemma starfun_star_n: "( *f* f) (star_n X) = star_n (\n. f (X n))" -by (simp only: starfun_def star_of_def Ifun_star_n) + by (simp only: starfun_def star_of_def Ifun_star_n) -lemma starfun2_star_n: - "( *f2* f) (star_n X) (star_n Y) = star_n (\n. f (X n) (Y n))" -by (simp only: starfun2_def star_of_def Ifun_star_n) +lemma starfun2_star_n: "( *f2* f) (star_n X) (star_n Y) = star_n (\n. f (X n) (Y n))" + by (simp only: starfun2_def star_of_def Ifun_star_n) lemma starfun_star_of [simp]: "( *f* f) (star_of x) = star_of (f x)" -by (transfer, rule refl) + by transfer (rule refl) lemma starfun2_star_of [simp]: "( *f2* f) (star_of x) = *f* f x" -by (transfer, rule refl) + by transfer (rule refl) lemma Standard_starfun [simp]: "x \ Standard \ starfun f x \ Standard" -by (simp add: starfun_def) + by (simp add: starfun_def) -lemma Standard_starfun2 [simp]: - "\x \ Standard; y \ Standard\ \ starfun2 f x y \ Standard" -by (simp add: starfun2_def) +lemma Standard_starfun2 [simp]: "x \ Standard \ y \ Standard \ starfun2 f x y \ Standard" + by (simp add: starfun2_def) lemma Standard_starfun_iff: assumes inj: "\x y. f x = f y \ x = y" - shows "(starfun f x \ Standard) = (x \ Standard)" + shows "starfun f x \ Standard \ x \ Standard" proof assume "x \ Standard" - thus "starfun f x \ Standard" by simp + then show "starfun f x \ Standard" by simp next - have inj': "\x y. starfun f x = starfun f y \ x = y" - using inj by transfer + from inj have inj': "\x y. starfun f x = starfun f y \ x = y" + by transfer assume "starfun f x \ Standard" then obtain b where b: "starfun f x = star_of b" unfolding Standard_def .. - hence "\x. starfun f x = star_of b" .. - hence "\a. f a = b" by transfer + then have "\x. starfun f x = star_of b" .. + then have "\a. f a = b" by transfer then obtain a where "f a = b" .. - hence "starfun f (star_of a) = star_of b" by transfer + then have "starfun f (star_of a) = star_of b" by transfer with b have "starfun f x = starfun f (star_of a)" by simp - hence "x = star_of a" by (rule inj') - thus "x \ Standard" - unfolding Standard_def by auto + then have "x = star_of a" by (rule inj') + then show "x \ Standard" by (simp add: Standard_def) qed lemma Standard_starfun2_iff: assumes inj: "\a b a' b'. f a b = f a' b' \ a = a' \ b = b'" - shows "(starfun2 f x y \ Standard) = (x \ Standard \ y \ Standard)" + shows "starfun2 f x y \ Standard \ x \ Standard \ y \ Standard" proof assume "x \ Standard \ y \ Standard" - thus "starfun2 f x y \ Standard" by simp + then show "starfun2 f x y \ Standard" by simp next have inj': "\x y z w. starfun2 f x y = starfun2 f z w \ x = z \ y = w" using inj by transfer assume "starfun2 f x y \ Standard" then obtain c where c: "starfun2 f x y = star_of c" unfolding Standard_def .. - hence "\x y. starfun2 f x y = star_of c" by auto - hence "\a b. f a b = c" by transfer + then have "\x y. starfun2 f x y = star_of c" by auto + then have "\a b. f a b = c" by transfer then obtain a b where "f a b = c" by auto - hence "starfun2 f (star_of a) (star_of b) = star_of c" - by transfer - with c have "starfun2 f x y = starfun2 f (star_of a) (star_of b)" - by simp - hence "x = star_of a \ y = star_of b" - by (rule inj') - thus "x \ Standard \ y \ Standard" - unfolding Standard_def by auto + then have "starfun2 f (star_of a) (star_of b) = star_of c" by transfer + with c have "starfun2 f x y = starfun2 f (star_of a) (star_of b)" by simp + then have "x = star_of a \ y = star_of b" by (rule inj') + then show "x \ Standard \ y \ Standard" by (simp add: Standard_def) qed subsection \Internal predicates\ -definition unstar :: "bool star \ bool" where - "unstar b \ b = star_of True" +definition unstar :: "bool star \ bool" + where "unstar b \ b = star_of True" -lemma unstar_star_n: "unstar (star_n P) = (eventually P \)" -by (simp add: unstar_def star_of_def star_n_eq_iff) +lemma unstar_star_n: "unstar (star_n P) \ eventually P \" + by (simp add: unstar_def star_of_def star_n_eq_iff) lemma unstar_star_of [simp]: "unstar (star_of p) = p" -by (simp add: unstar_def star_of_inject) + by (simp add: unstar_def star_of_inject) -text \Transfer tactic should remove occurrences of @{term unstar}\ +text \Transfer tactic should remove occurrences of @{term unstar}.\ setup \Transfer_Principle.add_const @{const_name unstar}\ -lemma transfer_unstar [transfer_intro]: - "p \ star_n P \ unstar p \ eventually P \" -by (simp only: unstar_star_n) +lemma transfer_unstar [transfer_intro]: "p \ star_n P \ unstar p \ eventually P \" + by (simp only: unstar_star_n) -definition - starP :: "('a \ bool) \ 'a star \ bool" ("*p* _" [80] 80) where - "*p* P = (\x. unstar (star_of P \ x))" +definition starP :: "('a \ bool) \ 'a star \ bool" ("*p* _" [80] 80) + where "*p* P = (\x. unstar (star_of P \ x))" -definition - starP2 :: "('a \ 'b \ bool) \ 'a star \ 'b star \ bool" ("*p2* _" [80] 80) where - "*p2* P = (\x y. unstar (star_of P \ x \ y))" +definition starP2 :: "('a \ 'b \ bool) \ 'a star \ 'b star \ bool" ("*p2* _" [80] 80) + where "*p2* P = (\x y. unstar (star_of P \ x \ y))" declare starP_def [transfer_unfold] declare starP2_def [transfer_unfold] -lemma starP_star_n: "( *p* P) (star_n X) = (eventually (\n. P (X n)) \)" -by (simp only: starP_def star_of_def Ifun_star_n unstar_star_n) +lemma starP_star_n: "( *p* P) (star_n X) = eventually (\n. P (X n)) \" + by (simp only: starP_def star_of_def Ifun_star_n unstar_star_n) -lemma starP2_star_n: - "( *p2* P) (star_n X) (star_n Y) = (eventually (\n. P (X n) (Y n)) \)" -by (simp only: starP2_def star_of_def Ifun_star_n unstar_star_n) +lemma starP2_star_n: "( *p2* P) (star_n X) (star_n Y) = (eventually (\n. P (X n) (Y n)) \)" + by (simp only: starP2_def star_of_def Ifun_star_n unstar_star_n) lemma starP_star_of [simp]: "( *p* P) (star_of x) = P x" -by (transfer, rule refl) + by transfer (rule refl) lemma starP2_star_of [simp]: "( *p2* P) (star_of x) = *p* P x" -by (transfer, rule refl) + by transfer (rule refl) subsection \Internal sets\ -definition - Iset :: "'a set star \ 'a star set" where - "Iset A = {x. ( *p2* op \) x A}" +definition Iset :: "'a set star \ 'a star set" + where "Iset A = {x. ( *p2* op \) x A}" -lemma Iset_star_n: - "(star_n X \ Iset (star_n A)) = (eventually (\n. X n \ A n) \)" -by (simp add: Iset_def starP2_star_n) +lemma Iset_star_n: "(star_n X \ Iset (star_n A)) = (eventually (\n. X n \ A n) \)" + by (simp add: Iset_def starP2_star_n) -text \Transfer tactic should remove occurrences of @{term Iset}\ +text \Transfer tactic should remove occurrences of @{term Iset}.\ setup \Transfer_Principle.add_const @{const_name Iset}\ lemma transfer_mem [transfer_intro]: - "\x \ star_n X; a \ Iset (star_n A)\ - \ x \ a \ eventually (\n. X n \ A n) \" -by (simp only: Iset_star_n) + "x \ star_n X \ a \ Iset (star_n A) \ x \ a \ eventually (\n. X n \ A n) \" + by (simp only: Iset_star_n) lemma transfer_Collect [transfer_intro]: - "\\X. p (star_n X) \ eventually (\n. P n (X n)) \\ - \ Collect p \ Iset (star_n (\n. Collect (P n)))" -by (simp add: atomize_eq set_eq_iff all_star_eq Iset_star_n) + "(\X. p (star_n X) \ eventually (\n. P n (X n)) \) \ + Collect p \ Iset (star_n (\n. Collect (P n)))" + by (simp add: atomize_eq set_eq_iff all_star_eq Iset_star_n) lemma transfer_set_eq [transfer_intro]: - "\a \ Iset (star_n A); b \ Iset (star_n B)\ - \ a = b \ eventually (\n. A n = B n) \" -by (simp only: set_eq_iff transfer_all transfer_iff transfer_mem) + "a \ Iset (star_n A) \ b \ Iset (star_n B) \ a = b \ eventually (\n. A n = B n) \" + by (simp only: set_eq_iff transfer_all transfer_iff transfer_mem) lemma transfer_ball [transfer_intro]: - "\a \ Iset (star_n A); \X. p (star_n X) \ eventually (\n. P n (X n)) \\ - \ \x\a. p x \ eventually (\n. \x\A n. P n x) \" -by (simp only: Ball_def transfer_all transfer_imp transfer_mem) + "a \ Iset (star_n A) \ (\X. p (star_n X) \ eventually (\n. P n (X n)) \) \ + \x\a. p x \ eventually (\n. \x\A n. P n x) \" + by (simp only: Ball_def transfer_all transfer_imp transfer_mem) lemma transfer_bex [transfer_intro]: - "\a \ Iset (star_n A); \X. p (star_n X) \ eventually (\n. P n (X n)) \\ - \ \x\a. p x \ eventually (\n. \x\A n. P n x) \" -by (simp only: Bex_def transfer_ex transfer_conj transfer_mem) + "a \ Iset (star_n A) \ (\X. p (star_n X) \ eventually (\n. P n (X n)) \) \ + \x\a. p x \ eventually (\n. \x\A n. P n x) \" + by (simp only: Bex_def transfer_ex transfer_conj transfer_mem) -lemma transfer_Iset [transfer_intro]: - "\a \ star_n A\ \ Iset a \ Iset (star_n (\n. A n))" -by simp +lemma transfer_Iset [transfer_intro]: "a \ star_n A \ Iset a \ Iset (star_n (\n. A n))" + by simp + text \Nonstandard extensions of sets.\ -definition - starset :: "'a set \ 'a star set" ("*s* _" [80] 80) where - "starset A = Iset (star_of A)" +definition starset :: "'a set \ 'a star set" ("*s* _" [80] 80) + where "starset A = Iset (star_of A)" declare starset_def [transfer_unfold] -lemma starset_mem: "(star_of x \ *s* A) = (x \ A)" -by (transfer, rule refl) +lemma starset_mem: "star_of x \ *s* A \ x \ A" + by transfer (rule refl) lemma starset_UNIV: "*s* (UNIV::'a set) = (UNIV::'a star set)" -by (transfer UNIV_def, rule refl) + by (transfer UNIV_def) (rule refl) lemma starset_empty: "*s* {} = {}" -by (transfer empty_def, rule refl) + by (transfer empty_def) (rule refl) lemma starset_insert: "*s* (insert x A) = insert (star_of x) ( *s* A)" -by (transfer insert_def Un_def, rule refl) + by (transfer insert_def Un_def) (rule refl) lemma starset_Un: "*s* (A \ B) = *s* A \ *s* B" -by (transfer Un_def, rule refl) + by (transfer Un_def) (rule refl) lemma starset_Int: "*s* (A \ B) = *s* A \ *s* B" -by (transfer Int_def, rule refl) + by (transfer Int_def) (rule refl) lemma starset_Compl: "*s* -A = -( *s* A)" -by (transfer Compl_eq, rule refl) + by (transfer Compl_eq) (rule refl) lemma starset_diff: "*s* (A - B) = *s* A - *s* B" -by (transfer set_diff_eq, rule refl) + by (transfer set_diff_eq) (rule refl) lemma starset_image: "*s* (f ` A) = ( *f* f) ` ( *s* A)" -by (transfer image_def, rule refl) + by (transfer image_def) (rule refl) lemma starset_vimage: "*s* (f -` A) = ( *f* f) -` ( *s* A)" -by (transfer vimage_def, rule refl) + by (transfer vimage_def) (rule refl) -lemma starset_subset: "( *s* A \ *s* B) = (A \ B)" -by (transfer subset_eq, rule refl) +lemma starset_subset: "( *s* A \ *s* B) \ A \ B" + by (transfer subset_eq) (rule refl) -lemma starset_eq: "( *s* A = *s* B) = (A = B)" -by (transfer, rule refl) +lemma starset_eq: "( *s* A = *s* B) \ A = B" + by transfer (rule refl) lemmas starset_simps [simp] = starset_mem starset_UNIV @@ -425,127 +395,77 @@ instantiation star :: (zero) zero begin - -definition - star_zero_def: "0 \ star_of 0" - -instance .. - + definition star_zero_def: "0 \ star_of 0" + instance .. end instantiation star :: (one) one begin - -definition - star_one_def: "1 \ star_of 1" - -instance .. - + definition star_one_def: "1 \ star_of 1" + instance .. end instantiation star :: (plus) plus begin - -definition - star_add_def: "(op +) \ *f2* (op +)" - -instance .. - + definition star_add_def: "(op +) \ *f2* (op +)" + instance .. end instantiation star :: (times) times begin - -definition - star_mult_def: "(op *) \ *f2* (op *)" - -instance .. - + definition star_mult_def: "(op *) \ *f2* (op *)" + instance .. end instantiation star :: (uminus) uminus begin - -definition - star_minus_def: "uminus \ *f* uminus" - -instance .. - + definition star_minus_def: "uminus \ *f* uminus" + instance .. end instantiation star :: (minus) minus begin - -definition - star_diff_def: "(op -) \ *f2* (op -)" - -instance .. - + definition star_diff_def: "(op -) \ *f2* (op -)" + instance .. end instantiation star :: (abs) abs begin - -definition - star_abs_def: "abs \ *f* abs" - -instance .. - + definition star_abs_def: "abs \ *f* abs" + instance .. end instantiation star :: (sgn) sgn begin - -definition - star_sgn_def: "sgn \ *f* sgn" - -instance .. - + definition star_sgn_def: "sgn \ *f* sgn" + instance .. end instantiation star :: (divide) divide begin - -definition - star_divide_def: "divide \ *f2* divide" - -instance .. - + definition star_divide_def: "divide \ *f2* divide" + instance .. end instantiation star :: (inverse) inverse begin - -definition - star_inverse_def: "inverse \ *f* inverse" - -instance .. - + definition star_inverse_def: "inverse \ *f* inverse" + instance .. end instance star :: (Rings.dvd) Rings.dvd .. instantiation star :: (modulo) modulo begin - -definition - star_mod_def: "(op mod) \ *f2* (op mod)" - -instance .. - + definition star_mod_def: "(op mod) \ *f2* (op mod)" + instance .. end instantiation star :: (ord) ord begin - -definition - star_le_def: "(op \) \ *p2* (op \)" - -definition - star_less_def: "(op <) \ *p2* (op <)" - -instance .. - + definition star_le_def: "(op \) \ *p2* (op \)" + definition star_less_def: "(op <) \ *p2* (op <)" + instance .. end lemmas star_class_defs [transfer_unfold] = @@ -555,37 +475,38 @@ star_le_def star_less_def star_abs_def star_sgn_def star_mod_def -text \Class operations preserve standard elements\ + +text \Class operations preserve standard elements.\ lemma Standard_zero: "0 \ Standard" -by (simp add: star_zero_def) + by (simp add: star_zero_def) lemma Standard_one: "1 \ Standard" -by (simp add: star_one_def) + by (simp add: star_one_def) -lemma Standard_add: "\x \ Standard; y \ Standard\ \ x + y \ Standard" -by (simp add: star_add_def) +lemma Standard_add: "x \ Standard \ y \ Standard \ x + y \ Standard" + by (simp add: star_add_def) -lemma Standard_diff: "\x \ Standard; y \ Standard\ \ x - y \ Standard" -by (simp add: star_diff_def) +lemma Standard_diff: "x \ Standard \ y \ Standard \ x - y \ Standard" + by (simp add: star_diff_def) lemma Standard_minus: "x \ Standard \ - x \ Standard" -by (simp add: star_minus_def) + by (simp add: star_minus_def) -lemma Standard_mult: "\x \ Standard; y \ Standard\ \ x * y \ Standard" -by (simp add: star_mult_def) +lemma Standard_mult: "x \ Standard \ y \ Standard \ x * y \ Standard" + by (simp add: star_mult_def) -lemma Standard_divide: "\x \ Standard; y \ Standard\ \ x / y \ Standard" -by (simp add: star_divide_def) +lemma Standard_divide: "x \ Standard \ y \ Standard \ x / y \ Standard" + by (simp add: star_divide_def) lemma Standard_inverse: "x \ Standard \ inverse x \ Standard" -by (simp add: star_inverse_def) + by (simp add: star_inverse_def) lemma Standard_abs: "x \ Standard \ \x\ \ Standard" -by (simp add: star_abs_def) + by (simp add: star_abs_def) -lemma Standard_mod: "\x \ Standard; y \ Standard\ \ x mod y \ Standard" -by (simp add: star_mod_def) +lemma Standard_mod: "x \ Standard \ y \ Standard \ x mod y \ Standard" + by (simp add: star_mod_def) lemmas Standard_simps [simp] = Standard_zero Standard_one @@ -593,41 +514,44 @@ Standard_mult Standard_divide Standard_inverse Standard_abs Standard_mod -text \@{term star_of} preserves class operations\ + +text \@{term star_of} preserves class operations.\ lemma star_of_add: "star_of (x + y) = star_of x + star_of y" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_diff: "star_of (x - y) = star_of x - star_of y" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_minus: "star_of (-x) = - star_of x" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_mult: "star_of (x * y) = star_of x * star_of y" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_divide: "star_of (x / y) = star_of x / star_of y" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_inverse: "star_of (inverse x) = inverse (star_of x)" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_mod: "star_of (x mod y) = star_of x mod star_of y" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_abs: "star_of \x\ = \star_of x\" -by transfer (rule refl) + by transfer (rule refl) -text \@{term star_of} preserves numerals\ + +text \@{term star_of} preserves numerals.\ lemma star_of_zero: "star_of 0 = 0" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_one: "star_of 1 = 1" -by transfer (rule refl) + by transfer (rule refl) -text \@{term star_of} preserves orderings\ + +text \@{term star_of} preserves orderings.\ lemma star_of_less: "(star_of x < star_of y) = (x < y)" by transfer (rule refl) @@ -638,7 +562,8 @@ lemma star_of_eq: "(star_of x = star_of y) = (x = y)" by transfer (rule refl) -text\As above, for 0\ + +text \As above, for \0\.\ lemmas star_of_0_less = star_of_less [of 0, simplified star_of_zero] lemmas star_of_0_le = star_of_le [of 0, simplified star_of_zero] @@ -648,7 +573,8 @@ lemmas star_of_le_0 = star_of_le [of _ 0, simplified star_of_zero] lemmas star_of_eq_0 = star_of_eq [of _ 0, simplified star_of_zero] -text\As above, for 1\ + +text \As above, for \1\.\ lemmas star_of_1_less = star_of_less [of 1, simplified star_of_one] lemmas star_of_1_le = star_of_le [of 1, simplified star_of_one] @@ -669,36 +595,27 @@ star_of_1_less star_of_1_le star_of_1_eq star_of_less_1 star_of_le_1 star_of_eq_1 + subsection \Ordering and lattice classes\ instance star :: (order) order -apply (intro_classes) -apply (transfer, rule less_le_not_le) -apply (transfer, rule order_refl) -apply (transfer, erule (1) order_trans) -apply (transfer, erule (1) order_antisym) -done + apply intro_classes + apply (transfer, rule less_le_not_le) + apply (transfer, rule order_refl) + apply (transfer, erule (1) order_trans) + apply (transfer, erule (1) order_antisym) + done instantiation star :: (semilattice_inf) semilattice_inf begin - -definition - star_inf_def [transfer_unfold]: "inf \ *f2* inf" - -instance - by (standard; transfer) auto - + definition star_inf_def [transfer_unfold]: "inf \ *f2* inf" + instance by (standard; transfer) auto end instantiation star :: (semilattice_sup) semilattice_sup begin - -definition - star_sup_def [transfer_unfold]: "sup \ *f2* sup" - -instance - by (standard; transfer) auto - + definition star_sup_def [transfer_unfold]: "sup \ *f2* sup" + instance by (standard; transfer) auto end instance star :: (lattice) lattice .. @@ -706,102 +623,98 @@ instance star :: (distrib_lattice) distrib_lattice by (standard; transfer) (auto simp add: sup_inf_distrib1) -lemma Standard_inf [simp]: - "\x \ Standard; y \ Standard\ \ inf x y \ Standard" -by (simp add: star_inf_def) +lemma Standard_inf [simp]: "x \ Standard \ y \ Standard \ inf x y \ Standard" + by (simp add: star_inf_def) -lemma Standard_sup [simp]: - "\x \ Standard; y \ Standard\ \ sup x y \ Standard" -by (simp add: star_sup_def) +lemma Standard_sup [simp]: "x \ Standard \ y \ Standard \ sup x y \ Standard" + by (simp add: star_sup_def) lemma star_of_inf [simp]: "star_of (inf x y) = inf (star_of x) (star_of y)" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_sup [simp]: "star_of (sup x y) = sup (star_of x) (star_of y)" -by transfer (rule refl) + by transfer (rule refl) instance star :: (linorder) linorder -by (intro_classes, transfer, rule linorder_linear) + by (intro_classes, transfer, rule linorder_linear) lemma star_max_def [transfer_unfold]: "max = *f2* max" -apply (rule ext, rule ext) -apply (unfold max_def, transfer, fold max_def) -apply (rule refl) -done + apply (rule ext, rule ext) + apply (unfold max_def, transfer, fold max_def) + apply (rule refl) + done lemma star_min_def [transfer_unfold]: "min = *f2* min" -apply (rule ext, rule ext) -apply (unfold min_def, transfer, fold min_def) -apply (rule refl) -done + apply (rule ext, rule ext) + apply (unfold min_def, transfer, fold min_def) + apply (rule refl) + done -lemma Standard_max [simp]: - "\x \ Standard; y \ Standard\ \ max x y \ Standard" -by (simp add: star_max_def) +lemma Standard_max [simp]: "x \ Standard \ y \ Standard \ max x y \ Standard" + by (simp add: star_max_def) -lemma Standard_min [simp]: - "\x \ Standard; y \ Standard\ \ min x y \ Standard" -by (simp add: star_min_def) +lemma Standard_min [simp]: "x \ Standard \ y \ Standard \ min x y \ Standard" + by (simp add: star_min_def) lemma star_of_max [simp]: "star_of (max x y) = max (star_of x) (star_of y)" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_min [simp]: "star_of (min x y) = min (star_of x) (star_of y)" -by transfer (rule refl) + by transfer (rule refl) subsection \Ordered group classes\ instance star :: (semigroup_add) semigroup_add -by (intro_classes, transfer, rule add.assoc) + by (intro_classes, transfer, rule add.assoc) instance star :: (ab_semigroup_add) ab_semigroup_add -by (intro_classes, transfer, rule add.commute) + by (intro_classes, transfer, rule add.commute) instance star :: (semigroup_mult) semigroup_mult -by (intro_classes, transfer, rule mult.assoc) + by (intro_classes, transfer, rule mult.assoc) instance star :: (ab_semigroup_mult) ab_semigroup_mult -by (intro_classes, transfer, rule mult.commute) + by (intro_classes, transfer, rule mult.commute) instance star :: (comm_monoid_add) comm_monoid_add -by (intro_classes, transfer, rule comm_monoid_add_class.add_0) + by (intro_classes, transfer, rule comm_monoid_add_class.add_0) instance star :: (monoid_mult) monoid_mult -apply (intro_classes) -apply (transfer, rule mult_1_left) -apply (transfer, rule mult_1_right) -done + apply intro_classes + apply (transfer, rule mult_1_left) + apply (transfer, rule mult_1_right) + done instance star :: (power) power .. instance star :: (comm_monoid_mult) comm_monoid_mult -by (intro_classes, transfer, rule mult_1) + by (intro_classes, transfer, rule mult_1) instance star :: (cancel_semigroup_add) cancel_semigroup_add -apply (intro_classes) -apply (transfer, erule add_left_imp_eq) -apply (transfer, erule add_right_imp_eq) -done + apply intro_classes + apply (transfer, erule add_left_imp_eq) + apply (transfer, erule add_right_imp_eq) + done instance star :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add -by intro_classes (transfer, simp add: diff_diff_eq)+ + by intro_classes (transfer, simp add: diff_diff_eq)+ instance star :: (cancel_comm_monoid_add) cancel_comm_monoid_add .. instance star :: (ab_group_add) ab_group_add -apply (intro_classes) -apply (transfer, rule left_minus) -apply (transfer, rule diff_conv_add_uminus) -done + apply intro_classes + apply (transfer, rule left_minus) + apply (transfer, rule diff_conv_add_uminus) + done instance star :: (ordered_ab_semigroup_add) ordered_ab_semigroup_add -by (intro_classes, transfer, rule add_left_mono) + by (intro_classes, transfer, rule add_left_mono) instance star :: (ordered_cancel_ab_semigroup_add) ordered_cancel_ab_semigroup_add .. instance star :: (ordered_ab_semigroup_add_imp_le) ordered_ab_semigroup_add_imp_le -by (intro_classes, transfer, rule add_le_imp_le_left) + by (intro_classes, transfer, rule add_le_imp_le_left) instance star :: (ordered_comm_monoid_add) ordered_comm_monoid_add .. instance star :: (ordered_ab_semigroup_monoid_add_imp_le) ordered_ab_semigroup_monoid_add_imp_le .. @@ -809,8 +722,7 @@ instance star :: (ordered_ab_group_add) ordered_ab_group_add .. instance star :: (ordered_ab_group_add_abs) ordered_ab_group_add_abs - by intro_classes (transfer, - simp add: abs_ge_self abs_leI abs_triangle_ineq)+ + by intro_classes (transfer, simp add: abs_ge_self abs_leI abs_triangle_ineq)+ instance star :: (linordered_cancel_ab_semigroup_add) linordered_cancel_ab_semigroup_add .. @@ -909,14 +821,13 @@ instance star :: (linordered_field) linordered_field .. + subsection \Power\ -lemma star_power_def [transfer_unfold]: - "(op ^) \ \x n. ( *f* (\x. x ^ n)) x" +lemma star_power_def [transfer_unfold]: "(op ^) \ \x n. ( *f* (\x. x ^ n)) x" proof (rule eq_reflection, rule ext, rule ext) - fix n :: nat - show "\x::'a star. x ^ n = ( *f* (\x. x ^ n)) x" - proof (induct n) + show "x ^ n = ( *f* (\x. x ^ n)) x" for n :: nat and x :: "'a star" + proof (induct n arbitrary: x) case 0 have "\x::'a star. ( *f* (\x. 1)) x = 1" by transfer simp @@ -940,18 +851,17 @@ instance star :: (numeral) numeral .. -lemma star_numeral_def [transfer_unfold]: - "numeral k = star_of (numeral k)" -by (induct k, simp_all only: numeral.simps star_of_one star_of_add) +lemma star_numeral_def [transfer_unfold]: "numeral k = star_of (numeral k)" + by (induct k) (simp_all only: numeral.simps star_of_one star_of_add) lemma Standard_numeral [simp]: "numeral k \ Standard" -by (simp add: star_numeral_def) + by (simp add: star_numeral_def) lemma star_of_numeral [simp]: "star_of (numeral k) = numeral k" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_nat_def [transfer_unfold]: "of_nat n = star_of (of_nat n)" -by (induct n, simp_all) + by (induct n) simp_all lemmas star_of_compare_numeral [simp] = star_of_less [of "numeral k", simplified star_of_numeral] @@ -968,63 +878,66 @@ star_of_eq [of _ "- numeral k", simplified star_of_numeral] for k lemma Standard_of_nat [simp]: "of_nat n \ Standard" -by (simp add: star_of_nat_def) + by (simp add: star_of_nat_def) lemma star_of_of_nat [simp]: "star_of (of_nat n) = of_nat n" -by transfer (rule refl) + by transfer (rule refl) lemma star_of_int_def [transfer_unfold]: "of_int z = star_of (of_int z)" -by (rule_tac z=z in int_diff_cases, simp) + by (rule int_diff_cases [of z]) simp lemma Standard_of_int [simp]: "of_int z \ Standard" -by (simp add: star_of_int_def) + by (simp add: star_of_int_def) lemma star_of_of_int [simp]: "star_of (of_int z) = of_int z" -by transfer (rule refl) + by transfer (rule refl) instance star :: (semiring_char_0) semiring_char_0 proof - have "inj (star_of :: 'a \ 'a star)" by (rule injI) simp - then have "inj (star_of \ of_nat :: nat \ 'a star)" using inj_of_nat by (rule inj_comp) - then show "inj (of_nat :: nat \ 'a star)" by (simp add: comp_def) + have "inj (star_of :: 'a \ 'a star)" + by (rule injI) simp + then have "inj (star_of \ of_nat :: nat \ 'a star)" + using inj_of_nat by (rule inj_comp) + then show "inj (of_nat :: nat \ 'a star)" + by (simp add: comp_def) qed instance star :: (ring_char_0) ring_char_0 .. instance star :: (semiring_parity) semiring_parity -apply intro_classes -apply(transfer, rule odd_one) -apply(transfer, erule (1) odd_even_add) -apply(transfer, erule even_multD) -apply(transfer, erule odd_ex_decrement) -done + apply intro_classes + apply (transfer, rule odd_one) + apply (transfer, erule (1) odd_even_add) + apply (transfer, erule even_multD) + apply (transfer, erule odd_ex_decrement) + done instance star :: (semiring_div_parity) semiring_div_parity -apply intro_classes -apply(transfer, rule parity) -apply(transfer, rule one_mod_two_eq_one) -apply(transfer, rule zero_not_eq_two) -done + apply intro_classes + apply (transfer, rule parity) + apply (transfer, rule one_mod_two_eq_one) + apply (transfer, rule zero_not_eq_two) + done instantiation star :: (semiring_numeral_div) semiring_numeral_div begin definition divmod_star :: "num \ num \ 'a star \ 'a star" -where - divmod_star_def: "divmod_star m n = (numeral m div numeral n, numeral m mod numeral n)" + where divmod_star_def: "divmod_star m n = (numeral m div numeral n, numeral m mod numeral n)" definition divmod_step_star :: "num \ 'a star \ 'a star \ 'a star \ 'a star" -where - "divmod_step_star l qr = (let (q, r) = qr - in if r \ numeral l then (2 * q + 1, r - numeral l) - else (2 * q, r))" + where "divmod_step_star l qr = + (let (q, r) = qr + in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" -instance proof - show "divmod m n = (numeral m div numeral n :: 'a star, numeral m mod numeral n)" - for m n by (fact divmod_star_def) - show "divmod_step l qr = (let (q, r) = qr - in if r \ numeral l then (2 * q + 1, r - numeral l) - else (2 * q, r))" for l and qr :: "'a star \ 'a star" +instance +proof + show "divmod m n = (numeral m div numeral n :: 'a star, numeral m mod numeral n)" for m n + by (fact divmod_star_def) + show "divmod_step l qr = + (let (q, r) = qr + in if r \ numeral l then (2 * q + 1, r - numeral l) else (2 * q, r))" + for l and qr :: "'a star \ 'a star" by (fact divmod_step_star_def) qed (transfer, fact @@ -1046,13 +959,13 @@ subsection \Finite class\ lemma starset_finite: "finite A \ *s* A = star_of ` A" -by (erule finite_induct, simp_all) + by (erule finite_induct) simp_all instance star :: (finite) finite -apply (intro_classes) -apply (subst starset_UNIV [symmetric]) -apply (subst starset_finite [OF finite]) -apply (rule finite_imageI [OF finite]) -done + apply intro_classes + apply (subst starset_UNIV [symmetric]) + apply (subst starset_finite [OF finite]) + apply (rule finite_imageI [OF finite]) + done end diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/document/root.tex --- a/src/HOL/Nonstandard_Analysis/document/root.tex Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/document/root.tex Tue Nov 01 00:44:24 2016 +0100 @@ -1,6 +1,6 @@ \documentclass[11pt,a4paper]{article} \usepackage{graphicx,isabelle,isabellesym,latexsym,textcomp} -\usepackage{amsmath} +\usepackage{amsmath,amssymb} \usepackage{stmaryrd} \usepackage{pdfsetup} diff -r af5235830c16 -r c93b0e6131c3 src/HOL/Nonstandard_Analysis/transfer.ML --- a/src/HOL/Nonstandard_Analysis/transfer.ML Mon Oct 31 16:26:36 2016 +0100 +++ b/src/HOL/Nonstandard_Analysis/transfer.ML Tue Nov 01 00:44:24 2016 +0100 @@ -13,7 +13,7 @@ structure Transfer_Principle: TRANSFER_PRINCIPLE = struct -structure TransferData = Generic_Data +structure Data = Generic_Data ( type T = { intros: thm list, @@ -55,13 +55,13 @@ fun transfer_star_tac ctxt = let fun thm_of (Const (@{const_name Ifun}, _) $ t $ u) = @{thm transfer_Ifun} OF [thm_of t, thm_of u] - | thm_of (Const (@{const_name star_of}, _) $ _) = @{thm star_of_def} - | thm_of (Const (@{const_name star_n}, _) $ _) = @{thm Pure.reflexive} - | thm_of _ = raise MATCH; + | thm_of (Const (@{const_name star_of}, _) $ _) = @{thm star_of_def} + | thm_of (Const (@{const_name star_n}, _) $ _) = @{thm Pure.reflexive} + | thm_of _ = raise MATCH; fun thm_of_goal (Const (@{const_name Pure.eq}, _) $ t $ (Const (@{const_name star_n}, _) $ _)) = - thm_of t - | thm_of_goal _ = raise MATCH; + thm_of t + | thm_of_goal _ = raise MATCH; in SUBGOAL (fn (t, i) => resolve_tac ctxt [thm_of_goal (strip_all_body t |> Logic.strip_imp_concl)] i @@ -71,44 +71,43 @@ fun transfer_thm_of ctxt ths t = let - val {intros,unfolds,refolds,consts} = TransferData.get (Context.Proof ctxt); + val {intros,unfolds,refolds,consts} = Data.get (Context.Proof ctxt); val meta = Local_Defs.meta_rewrite_rule ctxt; val ths' = map meta ths; val unfolds' = map meta unfolds and refolds' = map meta refolds; val (_$_$t') = Thm.concl_of (Raw_Simplifier.rewrite ctxt true unfolds' (Thm.cterm_of ctxt t)) val u = unstar_term consts t' - val tac = + val tac = rewrite_goals_tac ctxt (ths' @ refolds' @ unfolds') THEN ALLGOALS (Object_Logic.full_atomize_tac ctxt) THEN match_tac ctxt [transitive_thm] 1 THEN resolve_tac ctxt [@{thm transfer_start}] 1 THEN REPEAT_ALL_NEW (resolve_tac ctxt intros ORELSE' transfer_star_tac ctxt) 1 THEN match_tac ctxt [reflexive_thm] 1 - in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end + in Goal.prove ctxt [] [] (Logic.mk_equals (t,u)) (K tac) end; -fun transfer_tac ctxt ths = - SUBGOAL (fn (t, _) => - (fn th => - let - val tr = transfer_thm_of ctxt ths t - val (_$l$r) = Thm.concl_of tr; - val trs = if l aconv r then [] else [tr]; - in rewrite_goals_tac ctxt trs th end)) +fun transfer_tac ctxt ths = SUBGOAL (fn (t, _) => (fn th => + let + val tr = transfer_thm_of ctxt ths t + val (_$ l $ r) = Thm.concl_of tr; + val trs = if l aconv r then [] else [tr]; + in rewrite_goals_tac ctxt trs th end)); local -fun map_intros f = TransferData.map +fun map_intros f = Data.map (fn {intros,unfolds,refolds,consts} => {intros=f intros, unfolds=unfolds, refolds=refolds, consts=consts}) -fun map_unfolds f = TransferData.map +fun map_unfolds f = Data.map (fn {intros,unfolds,refolds,consts} => {intros=intros, unfolds=f unfolds, refolds=refolds, consts=consts}) -fun map_refolds f = TransferData.map +fun map_refolds f = Data.map (fn {intros,unfolds,refolds,consts} => {intros=intros, unfolds=unfolds, refolds=f refolds, consts=consts}) in + val intro_add = Thm.declaration_attribute (map_intros o Thm.add_thm); val intro_del = Thm.declaration_attribute (map_intros o Thm.del_thm); @@ -117,9 +116,10 @@ val refold_add = Thm.declaration_attribute (map_refolds o Thm.add_thm); val refold_del = Thm.declaration_attribute (map_refolds o Thm.del_thm); + end -fun add_const c = Context.theory_map (TransferData.map +fun add_const c = Context.theory_map (Data.map (fn {intros,unfolds,refolds,consts} => {intros=intros, unfolds=unfolds, refolds=refolds, consts=c::consts}))