# HG changeset patch # User paulson # Date 1075391477 -3600 # Node ID b0064703967b2d837ee1133ef4dfcda5aec4825f # Parent c50188fe6366f6db017bd2f885d1f7197f980967 simplifications in the hyperreals diff -r c50188fe6366 -r b0064703967b src/HOL/Complex/CStar.ML --- a/src/HOL/Complex/CStar.ML Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Complex/CStar.ML Thu Jan 29 16:51:17 2004 +0100 @@ -95,7 +95,7 @@ Goal "*sc* (A - B) = *sc* A - *sc* B"; by (auto_tac (claset(),simpset() addsimps - [set_diff_iff2,STARC_Int,STARC_Compl])); + [Diff_eq,STARC_Int,STARC_Compl])); qed "STARC_diff"; Goalw [starsetC_n_def] diff -r c50188fe6366 -r b0064703967b src/HOL/Complex/NSComplex.thy --- a/src/HOL/Complex/NSComplex.thy Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Complex/NSComplex.thy Thu Jan 29 16:51:17 2004 +0100 @@ -6,17 +6,15 @@ theory NSComplex = NSInduct: -(* Move to one of the hyperreal theories *) +(* ???MOVE to one of the hyperreal theories: HRealAbs??? *) lemma hypreal_of_nat: "hypreal_of_nat m = Abs_hypreal(hyprel `` {%n. real m})" apply (induct_tac "m") apply (auto simp add: hypreal_zero_def hypreal_of_nat_Suc hypreal_zero_num hypreal_one_num hypreal_add real_of_nat_Suc) done -(* not proved already? strange! *) lemma hypreal_of_nat_le_iff: "(hypreal_of_nat n \ hypreal_of_nat m) = (n \ m)" -apply (unfold hypreal_le_def) -apply auto +apply (auto simp add: linorder_not_less [symmetric]) done declare hypreal_of_nat_le_iff [simp] @@ -1365,7 +1363,7 @@ lemma cos_harg_i_mult_zero [simp]: "y \ 0 ==> ( *f* cos) (harg(iii * hcomplex_of_hypreal y)) = 0" -apply (cut_tac x = "y" and y = "0" in hypreal_linear) +apply (cut_tac x = "y" and y = "0" in linorder_less_linear) apply (auto simp add: cos_harg_i_mult_zero_pos cos_harg_i_mult_zero_neg) done diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/HLog.ML --- a/src/HOL/Hyperreal/HLog.ML Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/HLog.ML Thu Jan 29 16:51:17 2004 +0100 @@ -114,7 +114,7 @@ Addsimps [powhr_less_cancel_iff]; Goal "1 < x ==> (x powhr a <= x powhr b) = (a <= b)"; -by (auto_tac (claset(),simpset() addsimps [hypreal_le_def])); +by (auto_tac (claset(),simpset() addsimps [linorder_not_less RS sym])); qed "powhr_le_cancel_iff"; Addsimps [powhr_le_cancel_iff]; @@ -246,7 +246,7 @@ Addsimps [hlog_less_cancel_iff]; Goal "[| 1 < a; 0 < x; 0 < y |] ==> (hlog a x <= hlog a y) = (x <= y)"; -by (auto_tac (claset(),simpset() addsimps [hypreal_le_def])); +by (auto_tac (claset(),simpset() addsimps [linorder_not_less RS sym])); qed "hlog_le_cancel_iff"; Addsimps [hlog_le_cancel_iff]; diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/HRealAbs.thy --- a/src/HOL/Hyperreal/HRealAbs.thy Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/HRealAbs.thy Thu Jan 29 16:51:17 2004 +0100 @@ -31,7 +31,7 @@ by (simp add: order_less_imp_le hrabs_eqI1) lemma hrabs_minus_eqI2: "x<(0::hypreal) ==> abs x = -x" -by (simp add: hypreal_le_def hrabs_def) +by (simp add: linorder_not_less [symmetric] hrabs_def) lemma hrabs_minus_eqI1: "x<=(0::hypreal) ==> abs x = -x" by (auto dest: order_antisym simp add: hrabs_def) diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/HTranscendental.ML --- a/src/HOL/Hyperreal/HTranscendental.ML Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/HTranscendental.ML Thu Jan 29 16:51:17 2004 +0100 @@ -63,7 +63,7 @@ Goal "[|0<=x; 0<=y |] ==> \ \ ( *f* sqrt)(x*y) = ( *f* sqrt)(x) * ( *f* sqrt)(y)"; by (auto_tac (claset() addIs [hypreal_sqrt_mult_distrib], - simpset() addsimps [hypreal_le_eq_less_or_eq])); + simpset() addsimps [order_le_less])); qed "hypreal_sqrt_mult_distrib2"; Goal "0 < x ==> (( *f* sqrt)(x) @= 0) = (x @= 0)"; @@ -102,7 +102,7 @@ Goal "0 <= x ==> 0 <= ( *f* sqrt)(x)"; by (auto_tac (claset() addIs [hypreal_sqrt_gt_zero], - simpset() addsimps [hypreal_le_eq_less_or_eq ])); + simpset() addsimps [order_le_less ])); qed "hypreal_sqrt_ge_zero"; Goal "( *f* sqrt)(x ^ 2) = abs(x)"; @@ -148,14 +148,14 @@ Addsimps [hypreal_sqrt_sum_squares_ge1]; Goal "[| 0 <= x; x : HFinite |] ==> ( *f* sqrt) x : HFinite"; -by (auto_tac (claset(),simpset() addsimps [hypreal_le_eq_less_or_eq])); +by (auto_tac (claset(),simpset() addsimps [order_le_less])); by (rtac (HFinite_square_iff RS iffD1) 1); by (dtac hypreal_sqrt_gt_zero_pow2 1); by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2])); qed "HFinite_hypreal_sqrt"; Goal "[| 0 <= x; ( *f* sqrt) x : HFinite |] ==> x : HFinite"; -by (auto_tac (claset(),simpset() addsimps [hypreal_le_eq_less_or_eq])); +by (auto_tac (claset(),simpset() addsimps [order_le_less])); by (dtac (HFinite_square_iff RS iffD2) 1); by (dtac hypreal_sqrt_gt_zero_pow2 1); by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2] delsimps [HFinite_square_iff])); @@ -175,14 +175,14 @@ Addsimps [HFinite_sqrt_sum_squares]; Goal "[| 0 <= x; x : Infinitesimal |] ==> ( *f* sqrt) x : Infinitesimal"; -by (auto_tac (claset(),simpset() addsimps [hypreal_le_eq_less_or_eq])); +by (auto_tac (claset(),simpset() addsimps [order_le_less])); by (rtac (Infinitesimal_square_iff RS iffD2) 1); by (dtac hypreal_sqrt_gt_zero_pow2 1); by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2])); qed "Infinitesimal_hypreal_sqrt"; Goal "[| 0 <= x; ( *f* sqrt) x : Infinitesimal |] ==> x : Infinitesimal"; -by (auto_tac (claset(),simpset() addsimps [hypreal_le_eq_less_or_eq])); +by (auto_tac (claset(),simpset() addsimps [order_le_less])); by (dtac (Infinitesimal_square_iff RS iffD1) 1); by (dtac hypreal_sqrt_gt_zero_pow2 1); by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2] @@ -203,14 +203,14 @@ Addsimps [Infinitesimal_sqrt_sum_squares]; Goal "[| 0 <= x; x : HInfinite |] ==> ( *f* sqrt) x : HInfinite"; -by (auto_tac (claset(),simpset() addsimps [hypreal_le_eq_less_or_eq])); +by (auto_tac (claset(),simpset() addsimps [order_le_less])); by (rtac (HInfinite_square_iff RS iffD1) 1); by (dtac hypreal_sqrt_gt_zero_pow2 1); by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2])); qed "HInfinite_hypreal_sqrt"; Goal "[| 0 <= x; ( *f* sqrt) x : HInfinite |] ==> x : HInfinite"; -by (auto_tac (claset(),simpset() addsimps [hypreal_le_eq_less_or_eq])); +by (auto_tac (claset(),simpset() addsimps [order_le_less])); by (dtac (HInfinite_square_iff RS iffD2) 1); by (dtac hypreal_sqrt_gt_zero_pow2 1); by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2] diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/HyperArith.thy --- a/src/HOL/Hyperreal/HyperArith.thy Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/HyperArith.thy Thu Jan 29 16:51:17 2004 +0100 @@ -143,6 +143,9 @@ lemma hypreal_add_minus_eq_minus: "x + y = (0::hypreal) ==> x = -y" by arith +lemma hypreal_le_add_order: "[| 0 \ x; 0 \ y |] ==> (0::hypreal) \ x + y" +by arith + subsubsection{*Division By @{term 1} and @{term "-1"}*} @@ -197,6 +200,7 @@ ML {* val hypreal_add_minus_eq_minus = thm "hypreal_add_minus_eq_minus"; +val hypreal_le_add_order = thm"hypreal_le_add_order"; *} end diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/HyperDef.thy --- a/src/HOL/Hyperreal/HyperDef.thy Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/HyperDef.thy Thu Jan 29 16:51:17 2004 +0100 @@ -67,7 +67,7 @@ epsilon :: hypreal ("\") -defs +defs (overloaded) hypreal_add_def: "P + Q == Abs_hypreal(\X \ Rep_hypreal(P). \Y \ Rep_hypreal(Q). @@ -77,12 +77,12 @@ "P * Q == Abs_hypreal(\X \ Rep_hypreal(P). \Y \ Rep_hypreal(Q). hyprel``{%n::nat. X n * Y n})" - hypreal_less_def: - "P < (Q::hypreal) == \X Y. X \ Rep_hypreal(P) & + hypreal_le_def: + "P \ (Q::hypreal) == \X Y. X \ Rep_hypreal(P) & Y \ Rep_hypreal(Q) & - {n::nat. X n < Y n} \ FreeUltrafilterNat" - hypreal_le_def: - "P \ (Q::hypreal) == ~(Q < P)" + {n::nat. X n \ Y n} \ FreeUltrafilterNat" + + hypreal_less_def: "(x < (y::hypreal)) == (x \ y & x \ y)" hrabs_def: "abs (r::hypreal) == (if 0 \ r then r else -r)" @@ -494,205 +494,88 @@ qed -subsection{*Theorems for Ordering*} - -text{*TODO: define @{text "\"} as the primitive concept and quickly -establish membership in class @{text linorder}. Then proofs could be -simplified, since properties of @{text "<"} would be generic.*} - -text{*TODO: The following theorem should be used througout the proofs - as it probably makes many of them more straightforward.*} -lemma hypreal_less: - "(Abs_hypreal(hyprel``{%n. X n}) < Abs_hypreal(hyprel``{%n. Y n})) = - ({n. X n < Y n} \ FreeUltrafilterNat)" -apply (unfold hypreal_less_def) -apply (auto intro!: lemma_hyprel_refl, ultra) -done - -lemma hypreal_less_not_refl: "~ (R::hypreal) < R" -apply (rule_tac z = R in eq_Abs_hypreal) -apply (auto simp add: hypreal_less_def, ultra) -done - -lemmas hypreal_less_irrefl = hypreal_less_not_refl [THEN notE, standard] -declare hypreal_less_irrefl [elim!] - -lemma hypreal_not_refl2: "!!(x::hypreal). x < y ==> x \ y" -by (auto simp add: hypreal_less_not_refl) - -lemma hypreal_less_trans: "!!(R1::hypreal). [| R1 < R2; R2 < R3 |] ==> R1 < R3" -apply (rule_tac z = R1 in eq_Abs_hypreal) -apply (rule_tac z = R2 in eq_Abs_hypreal) -apply (rule_tac z = R3 in eq_Abs_hypreal) -apply (auto intro!: exI simp add: hypreal_less_def, ultra) -done - -lemma hypreal_less_asym: "!! (R1::hypreal). [| R1 < R2; R2 < R1 |] ==> P" -apply (drule hypreal_less_trans, assumption) -apply (simp add: hypreal_less_not_refl) -done - - -subsection{*Trichotomy: the hyperreals are Linearly Ordered*} - -lemma lemma_hyprel_0_mem: "\x. x \ hyprel `` {%n. 0}" -apply (unfold hyprel_def) -apply (rule_tac x = "%n. 0" in exI, safe) -apply (auto intro!: FreeUltrafilterNat_Nat_set) -done - -lemma hypreal_trichotomy: "0 < x | x = 0 | x < (0::hypreal)" -apply (unfold hypreal_zero_def) -apply (rule_tac z = x in eq_Abs_hypreal) -apply (auto simp add: hypreal_less_def) -apply (cut_tac lemma_hyprel_0_mem, erule exE) -apply (drule_tac x = xa in spec) -apply (drule_tac x = x in spec) -apply (cut_tac x = x in lemma_hyprel_refl, auto) -apply (drule_tac x = x in spec) -apply (drule_tac x = xa in spec, auto, ultra) -done - -lemma hypreal_trichotomyE: - "[| (0::hypreal) < x ==> P; - x = 0 ==> P; - x < 0 ==> P |] ==> P" -apply (insert hypreal_trichotomy [of x], blast) -done - -lemma hypreal_less_minus_iff: "((x::hypreal) < y) = (0 < y + -x)" -apply (rule_tac z = x in eq_Abs_hypreal) -apply (rule_tac z = y in eq_Abs_hypreal) -apply (auto simp add: hypreal_add hypreal_zero_def hypreal_minus hypreal_less) -done - -lemma hypreal_less_minus_iff2: "((x::hypreal) < y) = (x + -y < 0)" -apply (rule_tac z = x in eq_Abs_hypreal) -apply (rule_tac z = y in eq_Abs_hypreal) -apply (auto simp add: hypreal_add hypreal_zero_def hypreal_minus hypreal_less) -done - -lemma hypreal_eq_minus_iff2: "((x::hypreal) = y) = (0 = y + - x)" -apply auto -apply (rule Ring_and_Field.add_right_cancel [of _ "-x", THEN iffD1], auto) -done - -lemma hypreal_linear: "(x::hypreal) < y | x = y | y < x" -apply (subst hypreal_eq_minus_iff2) -apply (rule_tac x1 = x in hypreal_less_minus_iff [THEN ssubst]) -apply (rule_tac x1 = y in hypreal_less_minus_iff2 [THEN ssubst]) -apply (rule hypreal_trichotomyE, auto) -done - -lemma hypreal_neq_iff: "((w::hypreal) \ z) = (w"} Relation*} lemma hypreal_le: "(Abs_hypreal(hyprel``{%n. X n}) \ Abs_hypreal(hyprel``{%n. Y n})) = ({n. X n \ Y n} \ FreeUltrafilterNat)" -apply (auto simp add: hypreal_less hypreal_le_def linorder_not_less[symmetric]) -apply (ultra+) -done - -lemma hypreal_le_imp_less_or_eq: "!!(x::hypreal). x \ y ==> x < y | x = y" apply (unfold hypreal_le_def) -apply (cut_tac hypreal_linear) -apply (fast elim: hypreal_less_irrefl hypreal_less_asym) +apply (auto intro!: lemma_hyprel_refl) +apply (ultra) done -lemma hypreal_less_or_eq_imp_le: "z z \(w::hypreal)" -apply (unfold hypreal_le_def) -apply (cut_tac hypreal_linear) -apply (fast elim: hypreal_less_irrefl hypreal_less_asym) -done - -lemma hypreal_le_eq_less_or_eq: "(x \ (y::hypreal)) = (x < y | x=y)" -by (blast intro!: hypreal_less_or_eq_imp_le dest: hypreal_le_imp_less_or_eq) - -lemmas hypreal_le_less = hypreal_le_eq_less_or_eq - lemma hypreal_le_refl: "w \ (w::hypreal)" -by (simp add: hypreal_le_eq_less_or_eq) - -(* Axiom 'linorder_linear' of class 'linorder': *) -lemma hypreal_le_linear: "(z::hypreal) \ w | w \ z" -apply (simp add: hypreal_le_less) -apply (cut_tac hypreal_linear, blast) +apply (rule eq_Abs_hypreal [of w]) +apply (simp add: hypreal_le) done lemma hypreal_le_trans: "[| i \ j; j \ k |] ==> i \ (k::hypreal)" -apply (drule hypreal_le_imp_less_or_eq) -apply (drule hypreal_le_imp_less_or_eq) -apply (rule hypreal_less_or_eq_imp_le) -apply (blast intro: hypreal_less_trans) +apply (rule eq_Abs_hypreal [of i]) +apply (rule eq_Abs_hypreal [of j]) +apply (rule eq_Abs_hypreal [of k]) +apply (simp add: hypreal_le) +apply ultra done lemma hypreal_le_anti_sym: "[| z \ w; w \ z |] ==> z = (w::hypreal)" -apply (drule hypreal_le_imp_less_or_eq) -apply (drule hypreal_le_imp_less_or_eq) -apply (fast elim: hypreal_less_irrefl hypreal_less_asym) +apply (rule eq_Abs_hypreal [of z]) +apply (rule eq_Abs_hypreal [of w]) +apply (simp add: hypreal_le) +apply ultra done (* Axiom 'order_less_le' of class 'order': *) lemma hypreal_less_le: "((w::hypreal) < z) = (w \ z & w \ z)" -apply (simp add: hypreal_le_def hypreal_neq_iff) -apply (blast intro: hypreal_less_asym) +apply (simp add: hypreal_less_def) done instance hypreal :: order - by (intro_classes, - (assumption | - rule hypreal_le_refl hypreal_le_trans hypreal_le_anti_sym - hypreal_less_le)+) +proof qed + (assumption | + rule hypreal_le_refl hypreal_le_trans hypreal_le_anti_sym hypreal_less_le)+ + + +(* Axiom 'linorder_linear' of class 'linorder': *) +lemma hypreal_le_linear: "(z::hypreal) \ w | w \ z" +apply (rule eq_Abs_hypreal [of z]) +apply (rule eq_Abs_hypreal [of w]) +apply (auto simp add: hypreal_le) +apply ultra +done instance hypreal :: linorder by (intro_classes, rule hypreal_le_linear) - -lemma hypreal_add_less_mono1: "(A::hypreal) < B ==> A + C < B + C" -apply (rule_tac z = A in eq_Abs_hypreal) -apply (rule_tac z = B in eq_Abs_hypreal) -apply (rule_tac z = C in eq_Abs_hypreal) -apply (auto intro!: exI simp add: hypreal_less_def hypreal_add, ultra) -done +lemma hypreal_not_refl2: "!!(x::hypreal). x < y ==> x \ y" +by (auto simp add: order_less_irrefl) -lemma hypreal_mult_order: "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y" -apply (unfold hypreal_zero_def) -apply (rule_tac z = x in eq_Abs_hypreal) -apply (rule_tac z = y in eq_Abs_hypreal) -apply (auto intro!: exI simp add: hypreal_less_def hypreal_mult, ultra) -apply (auto intro: real_mult_order) -done - -lemma hypreal_add_left_le_mono1: "(q1::hypreal) \ q2 ==> x + q1 \ x + q2" -apply (drule order_le_imp_less_or_eq) -apply (auto intro: order_less_imp_le hypreal_add_less_mono1 simp add: hypreal_add_commute) -done - -lemma hypreal_mult_less_mono1: "[| (0::hypreal) < z; x < y |] ==> x*z < y*z" -apply (rotate_tac 1) -apply (drule hypreal_less_minus_iff [THEN iffD1]) -apply (rule hypreal_less_minus_iff [THEN iffD2]) -apply (drule hypreal_mult_order, assumption) -apply (simp add: right_distrib hypreal_mult_commute) +lemma hypreal_add_left_mono: "x \ y ==> z + x \ z + (y::hypreal)" +apply (rule eq_Abs_hypreal [of x]) +apply (rule eq_Abs_hypreal [of y]) +apply (rule eq_Abs_hypreal [of z]) +apply (auto simp add: hypreal_le hypreal_add) done lemma hypreal_mult_less_mono2: "[| (0::hypreal) z*x y ==> z + x \ z + y" - by (rule hypreal_add_left_le_mono1) + by (rule hypreal_add_left_mono) show "x < y ==> 0 < z ==> z * x < z * y" by (simp add: hypreal_mult_less_mono2) show "\x\ = (if x < 0 then -x else x)" @@ -783,17 +666,17 @@ lemma hypreal_of_real_zero [simp]: "hypreal_of_real 0 = 0" by (unfold hypreal_of_real_def hypreal_zero_def, simp) -lemma hypreal_of_real_less_iff [simp]: - "(hypreal_of_real w < hypreal_of_real z) = (w < z)" -apply (unfold hypreal_less_def hypreal_of_real_def, auto) +lemma hypreal_of_real_le_iff [simp]: + "(hypreal_of_real w \ hypreal_of_real z) = (w \ z)" +apply (unfold hypreal_le_def hypreal_of_real_def, auto) apply (rule_tac [2] x = "%n. w" in exI, safe) apply (rule_tac [3] x = "%n. z" in exI, auto) apply (rule FreeUltrafilterNat_P, ultra) done -lemma hypreal_of_real_le_iff [simp]: - "(hypreal_of_real w \ hypreal_of_real z) = (w \ z)" -by (force simp add: hypreal_less hypreal_le_def linorder_not_less[symmetric]) +lemma hypreal_of_real_less_iff [simp]: + "(hypreal_of_real w < hypreal_of_real z) = (w < z)" +by (simp add: linorder_not_le [symmetric]) lemma hypreal_of_real_eq_iff [simp]: "(hypreal_of_real w = hypreal_of_real z) = (w = z)" @@ -821,14 +704,11 @@ lemma hypreal_of_real_minus [simp]: "hypreal_of_real (-r) = - hypreal_of_real r" -apply (unfold hypreal_of_real_def) -apply (auto simp add: hypreal_minus) -done +by (auto simp add: hypreal_of_real_def hypreal_minus) lemma hypreal_of_real_inverse [simp]: "hypreal_of_real (inverse r) = inverse (hypreal_of_real r)" -apply (case_tac "r=0") -apply (simp add: DIVISION_BY_ZERO INVERSE_ZERO HYPREAL_INVERSE_ZERO) +apply (case_tac "r=0", simp) apply (rule_tac c1 = "hypreal_of_real r" in hypreal_mult_left_cancel [THEN iffD1]) apply (auto simp add: hypreal_of_real_mult [symmetric]) done @@ -840,6 +720,13 @@ subsection{*Misc Others*} +lemma hypreal_less: + "(Abs_hypreal(hyprel``{%n. X n}) < Abs_hypreal(hyprel``{%n. Y n})) = + ({n. X n < Y n} \ FreeUltrafilterNat)" +apply (auto simp add: hypreal_le linorder_not_le [symmetric]) +apply ultra+ +done + lemma hypreal_zero_num: "0 = Abs_hypreal (hyprel `` {%n. 0})" by (simp add: hypreal_zero_def [THEN meta_eq_to_obj_eq, symmetric]) @@ -851,7 +738,6 @@ apply (auto simp add: hypreal_less hypreal_zero_num) done - lemma hypreal_hrabs: "abs (Abs_hypreal (hyprel `` {X})) = Abs_hypreal(hyprel `` {%n. abs (X n)})" @@ -859,6 +745,74 @@ apply (ultra, arith)+ done + + +lemma hypreal_add_zero_less_le_mono: "[|r < x; (0::hypreal) \ y|] ==> r < x+y" +by (auto dest: add_less_le_mono) + +text{*The precondition could be weakened to @{term "0\x"}*} +lemma hypreal_mult_less_mono: + "[| u u*x < v* y" + by (simp add: Ring_and_Field.mult_strict_mono order_less_imp_le) + + +subsection{*Existence of Infinite Hyperreal Number*} + +lemma Rep_hypreal_omega: "Rep_hypreal(omega) \ hypreal" +apply (unfold omega_def) +apply (rule Rep_hypreal) +done + +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 dest: inj_real_of_nat [THEN injD]) + +lemma lemma_finite_omega_set: "finite {n::nat. x = real n}" +by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto) + +lemma not_ex_hypreal_of_real_eq_omega: + "~ (\x. hypreal_of_real x = omega)" +apply (unfold omega_def hypreal_of_real_def) +apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric] + lemma_finite_omega_set [THEN FreeUltrafilterNat_finite]) +done + +lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \ omega" +by (cut_tac not_ex_hypreal_of_real_eq_omega, auto) + +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 simp add: inj_real_of_nat [THEN inj_eq]) + +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 not_ex_hypreal_of_real_eq_epsilon: + "~ (\x. hypreal_of_real x = epsilon)" +apply (unfold epsilon_def hypreal_of_real_def) +apply (auto simp add: lemma_finite_epsilon_set [THEN FreeUltrafilterNat_finite]) +done + +lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \ epsilon" +by (cut_tac not_ex_hypreal_of_real_eq_epsilon, auto) + +lemma hypreal_epsilon_not_zero: "epsilon \ 0" +by (unfold epsilon_def hypreal_zero_def, auto) + +lemma hypreal_epsilon_inverse_omega: "epsilon = inverse(omega)" +by (simp add: hypreal_inverse omega_def epsilon_def) + + ML {* val hrabs_def = thm "hrabs_def"; @@ -946,23 +900,12 @@ val hypreal_mult_not_0 = thm "hypreal_mult_not_0"; val hypreal_minus_inverse = thm "hypreal_minus_inverse"; val hypreal_inverse_distrib = thm "hypreal_inverse_distrib"; -val hypreal_less_not_refl = thm "hypreal_less_not_refl"; -val hypreal_less_irrefl = thm"hypreal_less_irrefl"; val hypreal_not_refl2 = thm "hypreal_not_refl2"; -val hypreal_less_trans = thm "hypreal_less_trans"; -val hypreal_less_asym = thm "hypreal_less_asym"; val hypreal_less = thm "hypreal_less"; -val hypreal_trichotomy = thm "hypreal_trichotomy"; -val hypreal_less_minus_iff = thm "hypreal_less_minus_iff"; -val hypreal_less_minus_iff2 = thm "hypreal_less_minus_iff2"; val hypreal_eq_minus_iff = thm "hypreal_eq_minus_iff"; -val hypreal_eq_minus_iff2 = thm "hypreal_eq_minus_iff2"; val hypreal_eq_minus_iff3 = thm "hypreal_eq_minus_iff3"; val hypreal_not_eq_minus_iff = thm "hypreal_not_eq_minus_iff"; -val hypreal_linear = thm "hypreal_linear"; val hypreal_le = thm "hypreal_le"; -val hypreal_le_imp_less_or_eq = thm "hypreal_le_imp_less_or_eq"; -val hypreal_le_eq_less_or_eq = thm "hypreal_le_eq_less_or_eq"; val hypreal_le_refl = thm "hypreal_le_refl"; val hypreal_le_linear = thm "hypreal_le_linear"; val hypreal_le_trans = thm "hypreal_le_trans"; @@ -984,6 +927,17 @@ val hypreal_zero_num = thm "hypreal_zero_num"; val hypreal_one_num = thm "hypreal_one_num"; val hypreal_omega_gt_zero = thm "hypreal_omega_gt_zero"; + +val hypreal_add_zero_less_le_mono = thm"hypreal_add_zero_less_le_mono"; +val Rep_hypreal_omega = thm"Rep_hypreal_omega"; +val lemma_omega_empty_singleton_disj = thm"lemma_omega_empty_singleton_disj"; +val lemma_finite_omega_set = thm"lemma_finite_omega_set"; +val not_ex_hypreal_of_real_eq_omega = thm"not_ex_hypreal_of_real_eq_omega"; +val hypreal_of_real_not_eq_omega = thm"hypreal_of_real_not_eq_omega"; +val not_ex_hypreal_of_real_eq_epsilon = thm"not_ex_hypreal_of_real_eq_epsilon"; +val hypreal_of_real_not_eq_epsilon = thm"hypreal_of_real_not_eq_epsilon"; +val hypreal_epsilon_not_zero = thm"hypreal_epsilon_not_zero"; +val hypreal_epsilon_inverse_omega = thm"hypreal_epsilon_inverse_omega"; *} end diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/HyperNat.ML --- a/src/HOL/Hyperreal/HyperNat.ML Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/HyperNat.ML Thu Jan 29 16:51:17 2004 +0100 @@ -5,6 +5,12 @@ ultrafilters *) +fun CLAIM_SIMP str thms = + prove_goal (the_context()) str + (fn prems => [cut_facts_tac prems 1, + auto_tac (claset(),simpset() addsimps thms)]); +fun CLAIM str = CLAIM_SIMP str []; + (* blast confuses different versions of < *) DelIffs [order_less_irrefl]; Addsimps [order_less_irrefl]; diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/HyperOrd.thy --- a/src/HOL/Hyperreal/HyperOrd.thy Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/HyperOrd.thy Thu Jan 29 16:51:17 2004 +0100 @@ -1,124 +1,9 @@ -(* Title: Real/Hyperreal/HyperOrd.thy - Author: Jacques D. Fleuriot - Copyright: 2000 University of Edinburgh - Description: Type "hypreal" is a linear order and also - satisfies plus_ac0: + is an AC-operator with zero -*) - theory HyperOrd = HyperDef: -(*** Monotonicity results ***) - -lemma hypreal_add_less_le_mono: "[|(i::hypreal)l |] ==> i + k < j + l" -by (auto dest!: order_le_imp_less_or_eq intro: hypreal_add_less_mono1 add_strict_mono) - -lemma hypreal_add_less_mono2: "!!(A::hypreal). A < B ==> C + A < C + B" -by (auto intro: hypreal_add_less_mono1 simp add: hypreal_add_commute) - -lemma hypreal_add_le_less_mono: "[|(i::hypreal)\j; k i + k < j + l" -apply (erule add_right_mono [THEN order_le_less_trans]) -apply (erule add_strict_left_mono) -done - -lemma hypreal_add_zero_less_le_mono: "[|r < x; (0::hypreal) \ y|] ==> r < x + y" -by (auto dest: hypreal_add_less_le_mono) - -lemma hypreal_add_order: "[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y" -apply (erule order_less_trans) -apply (drule hypreal_add_less_mono2, simp) -done - -lemma hypreal_le_add_order: "[| 0 \ x; 0 \ y |] ==> (0::hypreal) \ x + y" -apply (drule order_le_imp_less_or_eq)+ -apply (auto intro: hypreal_add_order order_less_imp_le) -done - -text{*The precondition could be weakened to @{term "0\x"}*} -lemma hypreal_mult_less_mono: - "[| u u*x < v* y" - by (simp add: Ring_and_Field.mult_strict_mono order_less_imp_le) - - -subsection{*Existence of Infinite Hyperreal Number*} - -lemma Rep_hypreal_omega: "Rep_hypreal(omega) \ hypreal" -apply (unfold omega_def) -apply (rule Rep_hypreal) -done - -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 dest: inj_real_of_nat [THEN injD]) - -lemma lemma_finite_omega_set: "finite {n::nat. x = real n}" -by (cut_tac x = x in lemma_omega_empty_singleton_disj, auto) - -lemma not_ex_hypreal_of_real_eq_omega: - "~ (\x. hypreal_of_real x = omega)" -apply (unfold omega_def hypreal_of_real_def) -apply (auto simp add: real_of_nat_Suc diff_eq_eq [symmetric] - lemma_finite_omega_set [THEN FreeUltrafilterNat_finite]) -done - -lemma hypreal_of_real_not_eq_omega: "hypreal_of_real x \ omega" -by (cut_tac not_ex_hypreal_of_real_eq_omega, auto) - -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})" -apply (auto simp add: inj_real_of_nat [THEN inj_eq]) -done - -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 not_ex_hypreal_of_real_eq_epsilon: - "~ (\x. hypreal_of_real x = epsilon)" -apply (unfold epsilon_def hypreal_of_real_def) -apply (auto simp add: lemma_finite_epsilon_set [THEN FreeUltrafilterNat_finite]) -done - -lemma hypreal_of_real_not_eq_epsilon: "hypreal_of_real x \ epsilon" -by (cut_tac not_ex_hypreal_of_real_eq_epsilon, auto) - -lemma hypreal_epsilon_not_zero: "epsilon \ 0" -by (unfold epsilon_def hypreal_zero_def, auto) - -lemma hypreal_epsilon_inverse_omega: "epsilon = inverse(omega)" -by (simp add: hypreal_inverse omega_def epsilon_def) - ML {* -val hypreal_add_less_mono1 = thm"hypreal_add_less_mono1"; -val hypreal_mult_order = thm"hypreal_mult_order"; -val hypreal_le_add_order = thm"hypreal_le_add_order"; -val hypreal_add_left_le_mono1 = thm"hypreal_add_left_le_mono1"; -val hypreal_add_less_le_mono = thm"hypreal_add_less_le_mono"; -val hypreal_add_le_less_mono = thm"hypreal_add_le_less_mono"; -val hypreal_add_zero_less_le_mono = thm"hypreal_add_zero_less_le_mono"; -val hypreal_mult_less_mono1 = thm"hypreal_mult_less_mono1"; -val hypreal_mult_less_mono2 = thm"hypreal_mult_less_mono2"; -val hypreal_mult_less_mono = thm"hypreal_mult_less_mono"; -val Rep_hypreal_omega = thm"Rep_hypreal_omega"; -val lemma_omega_empty_singleton_disj = thm"lemma_omega_empty_singleton_disj"; -val lemma_finite_omega_set = thm"lemma_finite_omega_set"; -val not_ex_hypreal_of_real_eq_omega = thm"not_ex_hypreal_of_real_eq_omega"; -val hypreal_of_real_not_eq_omega = thm"hypreal_of_real_not_eq_omega"; -val not_ex_hypreal_of_real_eq_epsilon = thm"not_ex_hypreal_of_real_eq_epsilon"; -val hypreal_of_real_not_eq_epsilon = thm"hypreal_of_real_not_eq_epsilon"; -val hypreal_epsilon_not_zero = thm"hypreal_epsilon_not_zero"; -val hypreal_epsilon_inverse_omega = thm"hypreal_epsilon_inverse_omega"; *} end diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/NSA.ML --- a/src/HOL/Hyperreal/NSA.ML Wed Jan 28 17:01:01 2004 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2384 +0,0 @@ -(* Title : NSA.ML - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Description : Infinite numbers, Infinitesimals, - infinitely close relation etc. -*) - -fun CLAIM_SIMP str thms = - prove_goal (the_context()) str - (fn prems => [cut_facts_tac prems 1, - auto_tac (claset(),simpset() addsimps thms)]); -fun CLAIM str = CLAIM_SIMP str []; - -(*-------------------------------------------------------------------- - Closure laws for members of (embedded) set standard real Reals - --------------------------------------------------------------------*) - -Goalw [SReal_def] "[| (x::hypreal): Reals; y: Reals |] ==> x + y: Reals"; -by (Step_tac 1); -by (res_inst_tac [("x","r + ra")] exI 1); -by (Simp_tac 1); -qed "SReal_add"; - -Goalw [SReal_def] "[| (x::hypreal): Reals; y: Reals |] ==> x * y: Reals"; -by (Step_tac 1); -by (res_inst_tac [("x","r * ra")] exI 1); -by (simp_tac (simpset() addsimps [hypreal_of_real_mult]) 1); -qed "SReal_mult"; - -Goalw [SReal_def] "(x::hypreal): Reals ==> inverse x : Reals"; -by (blast_tac (claset() addIs [hypreal_of_real_inverse RS sym]) 1); -qed "SReal_inverse"; - -Goal "[| (x::hypreal): Reals; y: Reals |] ==> x/y: Reals"; -by (asm_simp_tac (simpset() addsimps [SReal_mult,SReal_inverse, - hypreal_divide_def]) 1); -qed "SReal_divide"; - -Goalw [SReal_def] "(x::hypreal): Reals ==> -x : Reals"; -by (blast_tac (claset() addIs [hypreal_of_real_minus RS sym]) 1); -qed "SReal_minus"; - -Goal "(-x : Reals) = ((x::hypreal): Reals)"; -by Auto_tac; -by (etac SReal_minus 2); -by (dtac SReal_minus 1); -by Auto_tac; -qed "SReal_minus_iff"; -Addsimps [SReal_minus_iff]; - -Goal "[| (x::hypreal) + y : Reals; y: Reals |] ==> x: Reals"; -by (dres_inst_tac [("x","y")] SReal_minus 1); -by (dtac SReal_add 1); -by (assume_tac 1); -by Auto_tac; -qed "SReal_add_cancel"; - -Goalw [SReal_def] "(x::hypreal): Reals ==> abs x : Reals"; -by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_hrabs])); -qed "SReal_hrabs"; - -Goalw [SReal_def] "hypreal_of_real x: Reals"; -by (Blast_tac 1); -qed "SReal_hypreal_of_real"; -Addsimps [SReal_hypreal_of_real]; - -Goalw [hypreal_number_of_def] "(number_of w ::hypreal) : Reals"; -by (rtac SReal_hypreal_of_real 1); -qed "SReal_number_of"; -Addsimps [SReal_number_of]; - -(** As always with numerals, 0 and 1 are special cases **) - -Goal "(0::hypreal) : Reals"; -by (stac (hypreal_numeral_0_eq_0 RS sym) 1); -by (rtac SReal_number_of 1); -qed "Reals_0"; -Addsimps [Reals_0]; - -Goal "(1::hypreal) : Reals"; -by (stac (hypreal_numeral_1_eq_1 RS sym) 1); -by (rtac SReal_number_of 1); -qed "Reals_1"; -Addsimps [Reals_1]; - -Goalw [hypreal_divide_def] "r : Reals ==> r/(number_of w::hypreal) : Reals"; -by (blast_tac (claset() addSIs [SReal_number_of, SReal_mult, - SReal_inverse]) 1); -qed "SReal_divide_number_of"; - -(* Infinitesimal epsilon not in Reals *) - -Goalw [SReal_def] "epsilon ~: Reals"; -by (auto_tac (claset(), - simpset() addsimps [hypreal_of_real_not_eq_epsilon RS not_sym])); -qed "SReal_epsilon_not_mem"; - -Goalw [SReal_def] "omega ~: Reals"; -by (auto_tac (claset(), - simpset() addsimps [hypreal_of_real_not_eq_omega RS not_sym])); -qed "SReal_omega_not_mem"; - -Goalw [SReal_def] "{x. hypreal_of_real x : Reals} = (UNIV::real set)"; -by Auto_tac; -qed "SReal_UNIV_real"; - -Goalw [SReal_def] "(x: Reals) = (EX y. x = hypreal_of_real y)"; -by Auto_tac; -qed "SReal_iff"; - -Goalw [SReal_def] "hypreal_of_real `(UNIV::real set) = Reals"; -by Auto_tac; -qed "hypreal_of_real_image"; - -Goalw [SReal_def] "inv hypreal_of_real `Reals = (UNIV::real set)"; -by Auto_tac; -by (rtac (inj_hypreal_of_real RS inv_f_f RS subst) 1); -by (Blast_tac 1); -qed "inv_hypreal_of_real_image"; - -Goalw [SReal_def] - "[| EX x. x: P; P <= Reals |] ==> EX Q. P = hypreal_of_real ` Q"; -by (Best_tac 1); -qed "SReal_hypreal_of_real_image"; - -Goal "[| (x::hypreal): Reals; y: Reals; x EX r: Reals. x ((EX x:P. y < x) = \ -\ (EX X. hypreal_of_real X : P & y < hypreal_of_real X))"; -by (blast_tac (claset() addSDs [SReal_iff RS iffD1]) 1); -by (flexflex_tac ); -qed "SReal_sup_lemma"; - -Goal "[| P <= Reals; EX x. x: P; EX y : Reals. ALL x: P. x < y |] \ -\ ==> (EX X. X: {w. hypreal_of_real w : P}) & \ -\ (EX Y. ALL X: {w. hypreal_of_real w : P}. X < Y)"; -by (rtac conjI 1); -by (fast_tac (claset() addSDs [SReal_iff RS iffD1]) 1); -by (Auto_tac THEN ftac subsetD 1 THEN assume_tac 1); -by (dtac (SReal_iff RS iffD1) 1); -by (Auto_tac THEN res_inst_tac [("x","ya")] exI 1); -by Auto_tac; -qed "SReal_sup_lemma2"; - -(*------------------------------------------------------ - lifting of ub and property of lub - -------------------------------------------------------*) -Goalw [isUb_def,setle_def] - "(isUb (Reals) (hypreal_of_real ` Q) (hypreal_of_real Y)) = \ -\ (isUb (UNIV :: real set) Q Y)"; -by Auto_tac; -qed "hypreal_of_real_isUb_iff"; - -Goalw [isLub_def,leastP_def] - "isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y) \ -\ ==> isLub (UNIV :: real set) Q Y"; -by (auto_tac (claset() addIs [hypreal_of_real_isUb_iff RS iffD2], - simpset() addsimps [hypreal_of_real_isUb_iff, setge_def])); -qed "hypreal_of_real_isLub1"; - -Goalw [isLub_def,leastP_def] - "isLub (UNIV :: real set) Q Y \ -\ ==> isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y)"; -by (auto_tac (claset(), - simpset() addsimps [hypreal_of_real_isUb_iff, setge_def])); -by (forw_inst_tac [("x2","x")] (isUbD2a RS (SReal_iff RS iffD1) RS exE) 1); -by (assume_tac 2); -by (dres_inst_tac [("x","xa")] spec 1); -by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_isUb_iff])); -qed "hypreal_of_real_isLub2"; - -Goal "(isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y)) = \ -\ (isLub (UNIV :: real set) Q Y)"; -by (blast_tac (claset() addIs [hypreal_of_real_isLub1, - hypreal_of_real_isLub2]) 1); -qed "hypreal_of_real_isLub_iff"; - -(* lemmas *) -Goalw [isUb_def] - "isUb Reals P Y ==> EX Yo. isUb Reals P (hypreal_of_real Yo)"; -by (auto_tac (claset(), simpset() addsimps [SReal_iff])); -qed "lemma_isUb_hypreal_of_real"; - -Goalw [isLub_def,leastP_def,isUb_def] - "isLub Reals P Y ==> EX Yo. isLub Reals P (hypreal_of_real Yo)"; -by (auto_tac (claset(), simpset() addsimps [SReal_iff])); -qed "lemma_isLub_hypreal_of_real"; - -Goalw [isLub_def,leastP_def,isUb_def] - "EX Yo. isLub Reals P (hypreal_of_real Yo) ==> EX Y. isLub Reals P Y"; -by Auto_tac; -qed "lemma_isLub_hypreal_of_real2"; - -Goal "[| P <= Reals; EX x. x: P; EX Y. isUb Reals P Y |] \ -\ ==> EX t::hypreal. isLub Reals P t"; -by (ftac SReal_hypreal_of_real_image 1); -by (Auto_tac THEN dtac lemma_isUb_hypreal_of_real 1); -by (auto_tac (claset() addSIs [reals_complete, lemma_isLub_hypreal_of_real2], - simpset() addsimps [hypreal_of_real_isLub_iff,hypreal_of_real_isUb_iff])); -qed "SReal_complete"; - -(*-------------------------------------------------------------------- - Set of finite elements is a subring of the extended reals - --------------------------------------------------------------------*) -Goalw [HFinite_def] "[|x : HFinite; y : HFinite|] ==> (x+y) : HFinite"; -by (blast_tac (claset() addSIs [SReal_add,hrabs_add_less]) 1); -qed "HFinite_add"; - -Goalw [HFinite_def] "[|x : HFinite; y : HFinite|] ==> x*y : HFinite"; -by (Asm_full_simp_tac 1); -by (blast_tac (claset() addSIs [SReal_mult,abs_mult_less]) 1); -qed "HFinite_mult"; - -Goalw [HFinite_def] "(-x : HFinite) = (x : HFinite)"; -by (Simp_tac 1); -qed "HFinite_minus_iff"; - -Goalw [SReal_def,HFinite_def] "Reals <= HFinite"; -by Auto_tac; -by (res_inst_tac [("x","1 + abs(hypreal_of_real r)")] exI 1); -by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_hrabs])); -by (res_inst_tac [("x","1 + abs r")] exI 1); -by (Simp_tac 1); -qed "SReal_subset_HFinite"; - -Goal "hypreal_of_real x : HFinite"; -by (auto_tac (claset() addIs [(SReal_subset_HFinite RS subsetD)], - simpset())); -qed "HFinite_hypreal_of_real"; - -Addsimps [HFinite_hypreal_of_real]; - -Goalw [HFinite_def] "x : HFinite ==> EX t: Reals. abs x < t"; -by Auto_tac; -qed "HFiniteD"; - -Goalw [HFinite_def] "(abs x : HFinite) = (x : HFinite)"; -by Auto_tac; -qed "HFinite_hrabs_iff"; -AddIffs [HFinite_hrabs_iff]; - -Goal "number_of w : HFinite"; -by (rtac (SReal_number_of RS (SReal_subset_HFinite RS subsetD)) 1); -qed "HFinite_number_of"; -Addsimps [HFinite_number_of]; - -(** As always with numerals, 0 and 1 are special cases **) - -Goal "0 : HFinite"; -by (stac (hypreal_numeral_0_eq_0 RS sym) 1); -by (rtac HFinite_number_of 1); -qed "HFinite_0"; -Addsimps [HFinite_0]; - -Goal "1 : HFinite"; -by (stac (hypreal_numeral_1_eq_1 RS sym) 1); -by (rtac HFinite_number_of 1); -qed "HFinite_1"; -Addsimps [HFinite_1]; - -Goal "[|x : HFinite; y <= x; 0 <= y |] ==> y: HFinite"; -by (case_tac "x <= 0" 1); -by (dres_inst_tac [("y","x")] order_trans 1); -by (dtac hypreal_le_anti_sym 2); -by (auto_tac (claset(), simpset() addsimps [linorder_not_le])); -by (auto_tac (claset() addSIs [bexI] addIs [order_le_less_trans], - simpset() addsimps [hrabs_eqI1,hrabs_eqI2,hrabs_minus_eqI1,HFinite_def])); -qed "HFinite_bounded"; - -(*------------------------------------------------------------------ - Set of infinitesimals is a subring of the hyperreals - ------------------------------------------------------------------*) -Goalw [Infinitesimal_def] - "x : Infinitesimal ==> ALL r: Reals. 0 < r --> abs x < r"; -by Auto_tac; -qed "InfinitesimalD"; - -Goalw [Infinitesimal_def] "0 : Infinitesimal"; -by (Simp_tac 1); -qed "Infinitesimal_zero"; -AddIffs [Infinitesimal_zero]; - -Goal "x/(2::hypreal) + x/(2::hypreal) = x"; -by Auto_tac; -qed "hypreal_sum_of_halves"; - -Goal "0 < r ==> 0 < r/(2::hypreal)"; -by Auto_tac; -qed "hypreal_half_gt_zero"; - -Goalw [Infinitesimal_def] - "[| x : Infinitesimal; y : Infinitesimal |] ==> (x+y) : Infinitesimal"; -by Auto_tac; -by (rtac (hypreal_sum_of_halves RS subst) 1); -by (dtac hypreal_half_gt_zero 1); -by (blast_tac (claset() addIs [hrabs_add_less, hrabs_add_less, - SReal_divide_number_of]) 1); -qed "Infinitesimal_add"; - -Goalw [Infinitesimal_def] "(-x:Infinitesimal) = (x:Infinitesimal)"; -by (Full_simp_tac 1); -qed "Infinitesimal_minus_iff"; -Addsimps [Infinitesimal_minus_iff]; - -Goal "[| x : Infinitesimal; y : Infinitesimal |] ==> x-y : Infinitesimal"; -by (asm_simp_tac - (simpset() addsimps [hypreal_diff_def, Infinitesimal_add]) 1); -qed "Infinitesimal_diff"; - -Goalw [Infinitesimal_def] - "[| x : Infinitesimal; y : Infinitesimal |] ==> (x * y) : Infinitesimal"; -by Auto_tac; -by (case_tac "y=0" 1); -by (cut_inst_tac [("u","abs x"),("v","1"),("x","abs y"),("y","r")] - hypreal_mult_less_mono 2); -by Auto_tac; -qed "Infinitesimal_mult"; - -Goal "[| x : Infinitesimal; y : HFinite |] ==> (x * y) : Infinitesimal"; -by (auto_tac (claset() addSDs [HFiniteD], - simpset() addsimps [Infinitesimal_def])); -by (ftac hrabs_less_gt_zero 1); -by (dres_inst_tac [("x","r/t")] bspec 1); -by (blast_tac (claset() addIs [SReal_divide]) 1); -by (asm_full_simp_tac (simpset() addsimps [thm"Ring_and_Field.zero_less_divide_iff"]) 1); -by (case_tac "x=0 | y=0" 1); -by (cut_inst_tac [("u","abs x"),("v","r/t"),("x","abs y")] - hypreal_mult_less_mono 2); -by (auto_tac (claset(), simpset() addsimps [thm"Ring_and_Field.zero_less_divide_iff"])); -qed "Infinitesimal_HFinite_mult"; - -Goal "[| x : Infinitesimal; y : HFinite |] ==> (y * x) : Infinitesimal"; -by (auto_tac (claset() addDs [Infinitesimal_HFinite_mult], - simpset() addsimps [hypreal_mult_commute])); -qed "Infinitesimal_HFinite_mult2"; - -(*** rather long proof ***) -Goalw [HInfinite_def,Infinitesimal_def] - "x: HInfinite ==> inverse x: Infinitesimal"; -by Auto_tac; -by (eres_inst_tac [("x","inverse r")] ballE 1); -by (forw_inst_tac [("a1","r"),("z","abs x")] - (positive_imp_inverse_positive RS order_less_trans) 1); -by (assume_tac 1); -by (dtac ((inverse_inverse_eq RS sym) RS subst) 1); -by (rtac (inverse_less_iff_less RS iffD1) 1); -by (auto_tac (claset(), simpset() addsimps [SReal_inverse])); -qed "HInfinite_inverse_Infinitesimal"; - - - -Goalw [HInfinite_def] "[|x: HInfinite;y: HInfinite|] ==> (x*y): HInfinite"; -by Auto_tac; -by (eres_inst_tac [("x","1")] ballE 1); -by (eres_inst_tac [("x","r")] ballE 1); -by (case_tac "y=0" 1); -by (cut_inst_tac [("x","1"),("y","abs x"), - ("u","r"),("v","abs y")] hypreal_mult_less_mono 2); -by (auto_tac (claset(), simpset() addsimps mult_ac)); -qed "HInfinite_mult"; - -Goalw [HInfinite_def] - "[|x: HInfinite; 0 <= y; 0 <= x|] ==> (x + y): HInfinite"; -by (auto_tac (claset() addSIs [hypreal_add_zero_less_le_mono], - simpset() addsimps [hrabs_eqI1, hypreal_add_commute, - hypreal_le_add_order])); -qed "HInfinite_add_ge_zero"; - -Goal "[|x: HInfinite; 0 <= y; 0 <= x|] ==> (y + x): HInfinite"; -by (auto_tac (claset() addSIs [HInfinite_add_ge_zero], - simpset() addsimps [hypreal_add_commute])); -qed "HInfinite_add_ge_zero2"; - -Goal "[|x: HInfinite; 0 < y; 0 < x|] ==> (x + y): HInfinite"; -by (blast_tac (claset() addIs [HInfinite_add_ge_zero, - order_less_imp_le]) 1); -qed "HInfinite_add_gt_zero"; - -Goalw [HInfinite_def] "(-x: HInfinite) = (x: HInfinite)"; -by Auto_tac; -qed "HInfinite_minus_iff"; - -Goal "[|x: HInfinite; y <= 0; x <= 0|] ==> (x + y): HInfinite"; -by (dtac (HInfinite_minus_iff RS iffD2) 1); -by (rtac (HInfinite_minus_iff RS iffD1) 1); -by (auto_tac (claset() addIs [HInfinite_add_ge_zero], - simpset())); -qed "HInfinite_add_le_zero"; - -Goal "[|x: HInfinite; y < 0; x < 0|] ==> (x + y): HInfinite"; -by (blast_tac (claset() addIs [HInfinite_add_le_zero, - order_less_imp_le]) 1); -qed "HInfinite_add_lt_zero"; - -Goal "[|a: HFinite; b: HFinite; c: HFinite|] \ -\ ==> a*a + b*b + c*c : HFinite"; -by (auto_tac (claset() addIs [HFinite_mult,HFinite_add], simpset())); -qed "HFinite_sum_squares"; - -Goal "x ~: Infinitesimal ==> x ~= 0"; -by Auto_tac; -qed "not_Infinitesimal_not_zero"; - -Goal "x: HFinite - Infinitesimal ==> x ~= 0"; -by Auto_tac; -qed "not_Infinitesimal_not_zero2"; - -Goal "(abs x : Infinitesimal) = (x : Infinitesimal)"; -by (auto_tac (claset(), simpset() addsimps [hrabs_def])); -qed "Infinitesimal_hrabs_iff"; -AddIffs [Infinitesimal_hrabs_iff]; - -Goal "x : HFinite - Infinitesimal ==> abs x : HFinite - Infinitesimal"; -by (Blast_tac 1); -qed "HFinite_diff_Infinitesimal_hrabs"; - -Goalw [Infinitesimal_def] - "[| e : Infinitesimal; abs x < e |] ==> x : Infinitesimal"; -by (auto_tac (claset(), simpset() addsimps [abs_less_iff])); -qed "hrabs_less_Infinitesimal"; - -Goal "[| e : Infinitesimal; abs x <= e |] ==> x : Infinitesimal"; -by (blast_tac (claset() addDs [order_le_imp_less_or_eq] - addIs [hrabs_less_Infinitesimal]) 1); -qed "hrabs_le_Infinitesimal"; - -Goalw [Infinitesimal_def] - "[| e : Infinitesimal; \ -\ e' : Infinitesimal; \ -\ e' < x ; x < e |] ==> x : Infinitesimal"; -by (auto_tac (claset(), simpset() addsimps [abs_less_iff])); -qed "Infinitesimal_interval"; - -Goal "[| e : Infinitesimal; e' : Infinitesimal; \ -\ e' <= x ; x <= e |] ==> x : Infinitesimal"; -by (auto_tac (claset() addIs [Infinitesimal_interval], - simpset() addsimps [hypreal_le_eq_less_or_eq])); -qed "Infinitesimal_interval2"; - -Goalw [Infinitesimal_def] - "[| x ~: Infinitesimal; y ~: Infinitesimal|] ==> (x*y) ~:Infinitesimal"; -by (Clarify_tac 1); -by (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1); -by (eres_inst_tac [("x","r*ra")] ballE 1); -by (fast_tac (claset() addIs [SReal_mult]) 2); -by (auto_tac (claset(), simpset() addsimps [zero_less_mult_iff])); -by (cut_inst_tac [("c","ra"),("d","abs y"), - ("a","r"),("b","abs x")] mult_mono 1); -by Auto_tac; -qed "not_Infinitesimal_mult"; - -Goal "x*y : Infinitesimal ==> x : Infinitesimal | y : Infinitesimal"; -by (rtac ccontr 1); -by (dtac (de_Morgan_disj RS iffD1) 1); -by (fast_tac (claset() addDs [not_Infinitesimal_mult]) 1); -qed "Infinitesimal_mult_disj"; - -Goal "x: HFinite-Infinitesimal ==> x ~= 0"; -by (Blast_tac 1); -qed "HFinite_Infinitesimal_not_zero"; - -Goal "[| x : HFinite - Infinitesimal; \ -\ y : HFinite - Infinitesimal \ -\ |] ==> (x*y) : HFinite - Infinitesimal"; -by (Clarify_tac 1); -by (blast_tac (claset() addDs [HFinite_mult,not_Infinitesimal_mult]) 1); -qed "HFinite_Infinitesimal_diff_mult"; - -Goalw [Infinitesimal_def,HFinite_def] - "Infinitesimal <= HFinite"; -by Auto_tac; -by (res_inst_tac [("x","1")] bexI 1); -by Auto_tac; -qed "Infinitesimal_subset_HFinite"; - -Goal "x: Infinitesimal ==> x * hypreal_of_real r : Infinitesimal"; -by (etac (HFinite_hypreal_of_real RSN - (2,Infinitesimal_HFinite_mult)) 1); -qed "Infinitesimal_hypreal_of_real_mult"; - -Goal "x: Infinitesimal ==> hypreal_of_real r * x: Infinitesimal"; -by (etac (HFinite_hypreal_of_real RSN - (2,Infinitesimal_HFinite_mult2)) 1); -qed "Infinitesimal_hypreal_of_real_mult2"; - -(*---------------------------------------------------------------------- - Infinitely close relation @= - ----------------------------------------------------------------------*) - -Goalw [Infinitesimal_def,approx_def] - "(x:Infinitesimal) = (x @= 0)"; -by (Simp_tac 1); -qed "mem_infmal_iff"; - -Goalw [approx_def]" (x @= y) = (x + -y @= 0)"; -by (Simp_tac 1); -qed "approx_minus_iff"; - -Goalw [approx_def]" (x @= y) = (-y + x @= 0)"; -by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1); -qed "approx_minus_iff2"; - -Goalw [approx_def,Infinitesimal_def] "x @= x"; -by (Simp_tac 1); -qed "approx_refl"; -AddIffs [approx_refl]; - -Goalw [approx_def] "x @= y ==> y @= x"; -by (rtac (hypreal_minus_distrib1 RS subst) 1); -by (etac (Infinitesimal_minus_iff RS iffD2) 1); -qed "approx_sym"; - -Goalw [approx_def] "[| x @= y; y @= z |] ==> x @= z"; -by (dtac Infinitesimal_add 1); -by (assume_tac 1); -by Auto_tac; -qed "approx_trans"; - -Goal "[| r @= x; s @= x |] ==> r @= s"; -by (blast_tac (claset() addIs [approx_sym, approx_trans]) 1); -qed "approx_trans2"; - -Goal "[| x @= r; x @= s|] ==> r @= s"; -by (blast_tac (claset() addIs [approx_sym, approx_trans]) 1); -qed "approx_trans3"; - -Goal "(number_of w @= x) = (x @= number_of w)"; -by (blast_tac (claset() addIs [approx_sym]) 1); -qed "number_of_approx_reorient"; - -Goal "(0 @= x) = (x @= 0)"; -by (blast_tac (claset() addIs [approx_sym]) 1); -qed "zero_approx_reorient"; - -Goal "(1 @= x) = (x @= 1)"; -by (blast_tac (claset() addIs [approx_sym]) 1); -qed "one_approx_reorient"; - - -(*** re-orientation, following HOL/Integ/Bin.ML - We re-orient x @=y where x is 0, 1 or a numeral, unless y is as well! - ***) - -(*reorientation simprules using ==, for the following simproc*) -val meta_zero_approx_reorient = zero_approx_reorient RS eq_reflection; -val meta_one_approx_reorient = one_approx_reorient RS eq_reflection; -val meta_number_of_approx_reorient = number_of_approx_reorient RS eq_reflection; - -(*reorientation simplification procedure: reorients (polymorphic) - 0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*) -fun reorient_proc sg _ (_ $ t $ u) = - case u of - Const("0", _) => None - | Const("1", _) => None - | Const("Numeral.number_of", _) $ _ => None - | _ => Some (case t of - Const("0", _) => meta_zero_approx_reorient - | Const("1", _) => meta_one_approx_reorient - | Const("Numeral.number_of", _) $ _ => - meta_number_of_approx_reorient); - -val approx_reorient_simproc = - Bin_Simprocs.prep_simproc - ("reorient_simproc", ["0@=x", "1@=x", "number_of w @= x"], reorient_proc); - -Addsimprocs [approx_reorient_simproc]; - - -Goal "(x-y : Infinitesimal) = (x @= y)"; -by (auto_tac (claset(), - simpset() addsimps [hypreal_diff_def, approx_minus_iff RS sym, - mem_infmal_iff])); -qed "Infinitesimal_approx_minus"; - -Goalw [monad_def] "(x @= y) = (monad(x)=monad(y))"; -by (auto_tac (claset() addDs [approx_sym] - addSEs [approx_trans,equalityCE], - simpset())); -qed "approx_monad_iff"; - -Goal "[| x: Infinitesimal; y: Infinitesimal |] ==> x @= y"; -by (asm_full_simp_tac (simpset() addsimps [mem_infmal_iff]) 1); -by (blast_tac (claset() addIs [approx_trans, approx_sym]) 1); -qed "Infinitesimal_approx"; - -val prem1::prem2::rest = -goalw thy [approx_def] "[| a @= b; c @= d |] ==> a+c @= b+d"; -by (stac minus_add_distrib 1); -by (stac hypreal_add_assoc 1); -by (res_inst_tac [("b1","c")] (add_left_commute RS subst) 1); -by (rtac (hypreal_add_assoc RS subst) 1); -by (rtac ([prem1,prem2] MRS Infinitesimal_add) 1); -qed "approx_add"; - -Goal "a @= b ==> -a @= -b"; -by (rtac ((approx_minus_iff RS iffD2) RS approx_sym) 1); -by (dtac (approx_minus_iff RS iffD1) 1); -by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1); -qed "approx_minus"; - -Goal "-a @= -b ==> a @= b"; -by (auto_tac (claset() addDs [approx_minus], simpset())); -qed "approx_minus2"; - -Goal "(-a @= -b) = (a @= b)"; -by (blast_tac (claset() addIs [approx_minus,approx_minus2]) 1); -qed "approx_minus_cancel"; - -Addsimps [approx_minus_cancel]; - -Goal "[| a @= b; c @= d |] ==> a + -c @= b + -d"; -by (blast_tac (claset() addSIs [approx_add,approx_minus]) 1); -qed "approx_add_minus"; - -Goalw [approx_def] "[| a @= b; c: HFinite|] ==> a*c @= b*c"; -by (asm_full_simp_tac (simpset() addsimps [Infinitesimal_HFinite_mult, - minus_mult_left,left_distrib RS sym] - delsimps [minus_mult_left RS sym]) 1); -qed "approx_mult1"; - -Goal "[|a @= b; c: HFinite|] ==> c*a @= c*b"; -by (asm_simp_tac (simpset() addsimps [approx_mult1,hypreal_mult_commute]) 1); -qed "approx_mult2"; - -Goal "[|u @= v*x; x @= y; v: HFinite|] ==> u @= v*y"; -by (fast_tac (claset() addIs [approx_mult2,approx_trans]) 1); -qed "approx_mult_subst"; - -Goal "[| u @= x*v; x @= y; v: HFinite |] ==> u @= y*v"; -by (fast_tac (claset() addIs [approx_mult1,approx_trans]) 1); -qed "approx_mult_subst2"; - -Goal "[| u @= x*hypreal_of_real v; x @= y |] ==> u @= y*hypreal_of_real v"; -by (auto_tac (claset() addIs [approx_mult_subst2], simpset())); -qed "approx_mult_subst_SReal"; - -Goalw [approx_def] "a = b ==> a @= b"; -by (Asm_simp_tac 1); -qed "approx_eq_imp"; - -Goal "x: Infinitesimal ==> -x @= x"; -by (fast_tac (HOL_cs addIs [Infinitesimal_minus_iff RS iffD2, - mem_infmal_iff RS iffD1,approx_trans2]) 1); -qed "Infinitesimal_minus_approx"; - -Goalw [approx_def] "(EX y: Infinitesimal. x + -z = y) = (x @= z)"; -by (Blast_tac 1); -qed "bex_Infinitesimal_iff"; - -Goal "(EX y: Infinitesimal. x = z + y) = (x @= z)"; -by (asm_full_simp_tac (simpset() addsimps [bex_Infinitesimal_iff RS sym]) 1); -by (Force_tac 1); -qed "bex_Infinitesimal_iff2"; - -Goal "[| y: Infinitesimal; x + y = z |] ==> x @= z"; -by (rtac (bex_Infinitesimal_iff RS iffD1) 1); -by (dtac (Infinitesimal_minus_iff RS iffD2) 1); -by (auto_tac (claset(), simpset() addsimps [minus_add_distrib, - hypreal_add_assoc RS sym])); -qed "Infinitesimal_add_approx"; - -Goal "y: Infinitesimal ==> x @= x + y"; -by (rtac (bex_Infinitesimal_iff RS iffD1) 1); -by (dtac (Infinitesimal_minus_iff RS iffD2) 1); -by (auto_tac (claset(), simpset() addsimps [minus_add_distrib, - hypreal_add_assoc RS sym])); -qed "Infinitesimal_add_approx_self"; - -Goal "y: Infinitesimal ==> x @= y + x"; -by (auto_tac (claset() addDs [Infinitesimal_add_approx_self], - simpset() addsimps [hypreal_add_commute])); -qed "Infinitesimal_add_approx_self2"; - -Goal "y: Infinitesimal ==> x @= x + -y"; -by (blast_tac (claset() addSIs [Infinitesimal_add_approx_self, - Infinitesimal_minus_iff RS iffD2]) 1); -qed "Infinitesimal_add_minus_approx_self"; - -Goal "[| y: Infinitesimal; x+y @= z|] ==> x @= z"; -by (dres_inst_tac [("x","x")] (Infinitesimal_add_approx_self RS approx_sym) 1); -by (etac (approx_trans3 RS approx_sym) 1); -by (assume_tac 1); -qed "Infinitesimal_add_cancel"; - -Goal "[| y: Infinitesimal; x @= z + y|] ==> x @= z"; -by (dres_inst_tac [("x","z")] (Infinitesimal_add_approx_self2 RS approx_sym) 1); -by (etac (approx_trans3 RS approx_sym) 1); -by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 1); -by (etac approx_sym 1); -qed "Infinitesimal_add_right_cancel"; - -Goal "d + b @= d + c ==> b @= c"; -by (dtac (approx_minus_iff RS iffD1) 1); -by (asm_full_simp_tac (simpset() addsimps - [minus_add_distrib,approx_minus_iff RS sym] - @ add_ac) 1); -qed "approx_add_left_cancel"; - -Goal "b + d @= c + d ==> b @= c"; -by (rtac approx_add_left_cancel 1); -by (asm_full_simp_tac (simpset() addsimps - [hypreal_add_commute]) 1); -qed "approx_add_right_cancel"; - -Goal "b @= c ==> d + b @= d + c"; -by (rtac (approx_minus_iff RS iffD2) 1); -by (asm_full_simp_tac (simpset() addsimps - [minus_add_distrib,approx_minus_iff RS sym] - @ add_ac) 1); -qed "approx_add_mono1"; - -Goal "b @= c ==> b + a @= c + a"; -by (asm_simp_tac (simpset() addsimps - [hypreal_add_commute,approx_add_mono1]) 1); -qed "approx_add_mono2"; - -Goal "(a + b @= a + c) = (b @= c)"; -by (fast_tac (claset() addEs [approx_add_left_cancel, - approx_add_mono1]) 1); -qed "approx_add_left_iff"; - -Addsimps [approx_add_left_iff]; - -Goal "(b + a @= c + a) = (b @= c)"; -by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1); -qed "approx_add_right_iff"; - -Addsimps [approx_add_right_iff]; - -Goal "[| x: HFinite; x @= y |] ==> y: HFinite"; -by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1); -by (Step_tac 1); -by (dtac (Infinitesimal_subset_HFinite RS subsetD - RS (HFinite_minus_iff RS iffD2)) 1); -by (dtac HFinite_add 1); -by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc])); -qed "approx_HFinite"; - -Goal "x @= hypreal_of_real D ==> x: HFinite"; -by (rtac (approx_sym RSN (2,approx_HFinite)) 1); -by Auto_tac; -qed "approx_hypreal_of_real_HFinite"; - -Goal "[|a @= b; c @= d; b: HFinite; d: HFinite|] ==> a*c @= b*d"; -by (rtac approx_trans 1); -by (rtac approx_mult2 2); -by (rtac approx_mult1 1); -by (blast_tac (claset() addIs [approx_HFinite, approx_sym]) 2); -by Auto_tac; -qed "approx_mult_HFinite"; - -Goal "[|a @= hypreal_of_real b; c @= hypreal_of_real d |] \ -\ ==> a*c @= hypreal_of_real b*hypreal_of_real d"; -by (blast_tac (claset() addSIs [approx_mult_HFinite, - approx_hypreal_of_real_HFinite,HFinite_hypreal_of_real]) 1); -qed "approx_mult_hypreal_of_real"; - -Goal "[| a: Reals; a ~= 0; a*x @= 0 |] ==> x @= 0"; -by (dtac (SReal_inverse RS (SReal_subset_HFinite RS subsetD)) 1); -by (auto_tac (claset() addDs [approx_mult2], - simpset() addsimps [hypreal_mult_assoc RS sym])); -qed "approx_SReal_mult_cancel_zero"; - -(* REM comments: newly added *) -Goal "[| a: Reals; x @= 0 |] ==> x*a @= 0"; -by (auto_tac (claset() addDs [(SReal_subset_HFinite RS subsetD), - approx_mult1], simpset())); -qed "approx_mult_SReal1"; - -Goal "[| a: Reals; x @= 0 |] ==> a*x @= 0"; -by (auto_tac (claset() addDs [(SReal_subset_HFinite RS subsetD), - approx_mult2], simpset())); -qed "approx_mult_SReal2"; - -Goal "[|a : Reals; a ~= 0 |] ==> (a*x @= 0) = (x @= 0)"; -by (blast_tac (claset() addIs [approx_SReal_mult_cancel_zero, - approx_mult_SReal2]) 1); -qed "approx_mult_SReal_zero_cancel_iff"; -Addsimps [approx_mult_SReal_zero_cancel_iff]; - -Goal "[| a: Reals; a ~= 0; a* w @= a*z |] ==> w @= z"; -by (dtac (SReal_inverse RS (SReal_subset_HFinite RS subsetD)) 1); -by (auto_tac (claset() addDs [approx_mult2], - simpset() addsimps [hypreal_mult_assoc RS sym])); -qed "approx_SReal_mult_cancel"; - -Goal "[| a: Reals; a ~= 0|] ==> (a* w @= a*z) = (w @= z)"; -by (auto_tac (claset() addSIs [approx_mult2,SReal_subset_HFinite RS subsetD] - addIs [approx_SReal_mult_cancel], simpset())); -qed "approx_SReal_mult_cancel_iff1"; -Addsimps [approx_SReal_mult_cancel_iff1]; - -Goal "[| z <= f; f @= g; g <= z |] ==> f @= z"; -by (asm_full_simp_tac (simpset() addsimps [bex_Infinitesimal_iff2 RS sym]) 1); -by Auto_tac; -by (res_inst_tac [("x","g+y-z")] bexI 1); -by (Simp_tac 1); -by (rtac Infinitesimal_interval2 1); -by (rtac Infinitesimal_zero 2); -by Auto_tac; -qed "approx_le_bound"; - -(*----------------------------------------------------------------- - Zero is the only infinitesimal that is also a real - -----------------------------------------------------------------*) - -Goalw [Infinitesimal_def] - "[| x: Reals; y: Infinitesimal; 0 < x |] ==> y < x"; -by (rtac (abs_ge_self RS order_le_less_trans) 1); -by Auto_tac; -qed "Infinitesimal_less_SReal"; - -Goal "y: Infinitesimal ==> ALL r: Reals. 0 < r --> y < r"; -by (blast_tac (claset() addIs [Infinitesimal_less_SReal]) 1); -qed "Infinitesimal_less_SReal2"; - -Goalw [Infinitesimal_def] - "[| 0 < y; y: Reals|] ==> y ~: Infinitesimal"; -by (auto_tac (claset(), simpset() addsimps [hrabs_def])); -qed "SReal_not_Infinitesimal"; - -Goal "[| y < 0; y : Reals |] ==> y ~: Infinitesimal"; -by (stac (Infinitesimal_minus_iff RS sym) 1); -by (rtac SReal_not_Infinitesimal 1); -by Auto_tac; -qed "SReal_minus_not_Infinitesimal"; - -Goal "Reals Int Infinitesimal = {0}"; -by Auto_tac; -by (cut_inst_tac [("x","x"),("y","0")] hypreal_linear 1); -by (blast_tac (claset() addDs [SReal_not_Infinitesimal, - SReal_minus_not_Infinitesimal]) 1); -qed "SReal_Int_Infinitesimal_zero"; - -Goal "[| x: Reals; x: Infinitesimal|] ==> x = 0"; -by (cut_facts_tac [SReal_Int_Infinitesimal_zero] 1); -by (Blast_tac 1); -qed "SReal_Infinitesimal_zero"; - -Goal "[| x : Reals; x ~= 0 |] ==> x : HFinite - Infinitesimal"; -by (auto_tac (claset() addDs [SReal_Infinitesimal_zero, - SReal_subset_HFinite RS subsetD], - simpset())); -qed "SReal_HFinite_diff_Infinitesimal"; - -Goal "hypreal_of_real x ~= 0 ==> hypreal_of_real x : HFinite - Infinitesimal"; -by (rtac SReal_HFinite_diff_Infinitesimal 1); -by Auto_tac; -qed "hypreal_of_real_HFinite_diff_Infinitesimal"; - -Goal "(hypreal_of_real x : Infinitesimal) = (x=0)"; -by Auto_tac; -by (rtac ccontr 1); -by (rtac (hypreal_of_real_HFinite_diff_Infinitesimal RS DiffD2) 1); -by Auto_tac; -qed "hypreal_of_real_Infinitesimal_iff_0"; -AddIffs [hypreal_of_real_Infinitesimal_iff_0]; - -Goal "number_of w ~= (0::hypreal) ==> number_of w ~: Infinitesimal"; -by (fast_tac (claset() addDs [SReal_number_of RS SReal_Infinitesimal_zero]) 1); -qed "number_of_not_Infinitesimal"; -Addsimps [number_of_not_Infinitesimal]; - -(*again: 1 is a special case, but not 0 this time*) -Goal "1 ~: Infinitesimal"; -by (stac (hypreal_numeral_1_eq_1 RS sym) 1); -by (rtac number_of_not_Infinitesimal 1); -by (Simp_tac 1); -qed "one_not_Infinitesimal"; -Addsimps [one_not_Infinitesimal]; - -Goal "[| y: Reals; x @= y; y~= 0 |] ==> x ~= 0"; -by (cut_inst_tac [("x","y")] hypreal_trichotomy 1); -by (Asm_full_simp_tac 1); -by (blast_tac (claset() addDs - [approx_sym RS (mem_infmal_iff RS iffD2), - SReal_not_Infinitesimal, SReal_minus_not_Infinitesimal]) 1); -qed "approx_SReal_not_zero"; - -Goal "[| x @= y; y : HFinite - Infinitesimal |] \ -\ ==> x : HFinite - Infinitesimal"; -by (auto_tac (claset() addIs [approx_sym RSN (2,approx_HFinite)], - simpset() addsimps [mem_infmal_iff])); -by (dtac approx_trans3 1 THEN assume_tac 1); -by (blast_tac (claset() addDs [approx_sym]) 1); -qed "HFinite_diff_Infinitesimal_approx"; - -(*The premise y~=0 is essential; otherwise x/y =0 and we lose the - HFinite premise.*) -Goal "[| y ~= 0; y: Infinitesimal; x/y : HFinite |] ==> x : Infinitesimal"; -by (dtac Infinitesimal_HFinite_mult2 1); -by (assume_tac 1); -by (asm_full_simp_tac - (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 1); -qed "Infinitesimal_ratio"; - -(*------------------------------------------------------------------ - Standard Part Theorem: Every finite x: R* is infinitely - close to a unique real number (i.e a member of Reals) - ------------------------------------------------------------------*) -(*------------------------------------------------------------------ - Uniqueness: Two infinitely close reals are equal - ------------------------------------------------------------------*) - -Goal "[|x: Reals; y: Reals|] ==> (x @= y) = (x = y)"; -by Auto_tac; -by (rewtac approx_def); -by (dres_inst_tac [("x","y")] SReal_minus 1); -by (dtac SReal_add 1 THEN assume_tac 1); -by (dtac SReal_Infinitesimal_zero 1 THEN assume_tac 1); -by (dtac sym 1); -by (asm_full_simp_tac (simpset() addsimps [hypreal_eq_minus_iff RS sym]) 1); -qed "SReal_approx_iff"; - -Goal "(number_of v @= number_of w) = (number_of v = (number_of w :: hypreal))"; -by (rtac SReal_approx_iff 1); -by Auto_tac; -qed "number_of_approx_iff"; -Addsimps [number_of_approx_iff]; - -(*And also for 0 @= #nn and 1 @= #nn, #nn @= 0 and #nn @= 1.*) -Addsimps - (map (simplify (simpset())) - [inst "v" "bin.Pls" number_of_approx_iff, - inst "v" "bin.Pls BIT True" number_of_approx_iff, - inst "w" "bin.Pls" number_of_approx_iff, - inst "w" "bin.Pls BIT True" number_of_approx_iff]); - - -Goal "~ (0 @= 1)"; -by (stac SReal_approx_iff 1); -by Auto_tac; -qed "not_0_approx_1"; -Addsimps [not_0_approx_1]; - -Goal "~ (1 @= 0)"; -by (stac SReal_approx_iff 1); -by Auto_tac; -qed "not_1_approx_0"; -Addsimps [not_1_approx_0]; - -Goal "(hypreal_of_real k @= hypreal_of_real m) = (k = m)"; -by Auto_tac; -by (rtac (inj_hypreal_of_real RS injD) 1); -by (rtac (SReal_approx_iff RS iffD1) 1); -by Auto_tac; -qed "hypreal_of_real_approx_iff"; -Addsimps [hypreal_of_real_approx_iff]; - -Goal "(hypreal_of_real k @= number_of w) = (k = number_of w)"; -by (stac (hypreal_of_real_approx_iff RS sym) 1); -by Auto_tac; -qed "hypreal_of_real_approx_number_of_iff"; -Addsimps [hypreal_of_real_approx_number_of_iff]; - -(*And also for 0 and 1.*) -Addsimps - (map (simplify (simpset())) - [inst "w" "bin.Pls" hypreal_of_real_approx_number_of_iff, - inst "w" "bin.Pls BIT True" hypreal_of_real_approx_number_of_iff]); - -Goal "[| r: Reals; s: Reals; r @= x; s @= x|] ==> r = s"; -by (blast_tac (claset() addIs [(SReal_approx_iff RS iffD1), - approx_trans2]) 1); -qed "approx_unique_real"; - -(*------------------------------------------------------------------ - Existence of unique real infinitely close - ------------------------------------------------------------------*) -(* lemma about lubs *) -Goal "!!(x::hypreal). [| isLub R S x; isLub R S y |] ==> x = y"; -by (ftac isLub_isUb 1); -by (forw_inst_tac [("x","y")] isLub_isUb 1); -by (blast_tac (claset() addSIs [hypreal_le_anti_sym] - addSDs [isLub_le_isUb]) 1); -qed "hypreal_isLub_unique"; - -Goal "x: HFinite ==> EX u. isUb Reals {s. s: Reals & s < x} u"; -by (dtac HFiniteD 1 THEN Step_tac 1); -by (rtac exI 1 THEN rtac isUbI 1 THEN assume_tac 2); -by (auto_tac (claset() addIs [setleI,isUbI], simpset() addsimps [abs_less_iff])); -qed "lemma_st_part_ub"; - -Goal "x: HFinite ==> EX y. y: {s. s: Reals & s < x}"; -by (dtac HFiniteD 1 THEN Step_tac 1); -by (dtac SReal_minus 1); -by (res_inst_tac [("x","-t")] exI 1); -by (auto_tac (claset(), simpset() addsimps [abs_less_iff])); -qed "lemma_st_part_nonempty"; - -Goal "{s. s: Reals & s < x} <= Reals"; -by Auto_tac; -qed "lemma_st_part_subset"; - -Goal "x: HFinite ==> EX t. isLub Reals {s. s: Reals & s < x} t"; -by (blast_tac (claset() addSIs [SReal_complete,lemma_st_part_ub, - lemma_st_part_nonempty, lemma_st_part_subset]) 1); -qed "lemma_st_part_lub"; - -Goal "((t::hypreal) + r <= t) = (r <= 0)"; -by (Step_tac 1); -by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1); -by (dres_inst_tac [("x","t")] hypreal_add_left_le_mono1 2); -by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc RS sym])); -qed "lemma_hypreal_le_left_cancel"; - -Goal "[| x: HFinite; isLub Reals {s. s: Reals & s < x} t; \ -\ r: Reals; 0 < r |] ==> x <= t + r"; -by (ftac isLubD1a 1); -by (rtac ccontr 1 THEN dtac (linorder_not_le RS iffD2) 1); -by (dres_inst_tac [("x","t")] SReal_add 1 THEN assume_tac 1); -by (dres_inst_tac [("y","t + r")] (isLubD1 RS setleD) 1); -by Auto_tac; -qed "lemma_st_part_le1"; - -Goalw [setle_def] "!!x::hypreal. [| S *<= x; x < y |] ==> S *<= y"; -by (auto_tac (claset() addSDs [bspec,order_le_less_trans] - addIs [order_less_imp_le], - simpset())); -qed "hypreal_setle_less_trans"; - -Goalw [isUb_def] - "!!x::hypreal. [| isUb R S x; x < y; y: R |] ==> isUb R S y"; -by (blast_tac (claset() addIs [hypreal_setle_less_trans]) 1); -qed "hypreal_gt_isUb"; - -Goal "[| x: HFinite; x < y; y: Reals |] \ -\ ==> isUb Reals {s. s: Reals & s < x} y"; -by (auto_tac (claset() addDs [order_less_trans] - addIs [order_less_imp_le] addSIs [isUbI,setleI], simpset())); -qed "lemma_st_part_gt_ub"; - -Goal "t <= t + -r ==> r <= (0::hypreal)"; -by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1); -by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc RS sym])); -qed "lemma_minus_le_zero"; - -Goal "[| x: HFinite; \ -\ isLub Reals {s. s: Reals & s < x} t; \ -\ r: Reals; 0 < r |] \ -\ ==> t + -r <= x"; -by (ftac isLubD1a 1); -by (rtac ccontr 1 THEN dtac (linorder_not_le RS iffD1) 1); -by (dtac SReal_minus 1 THEN dres_inst_tac [("x","t")] - SReal_add 1 THEN assume_tac 1); -by (dtac lemma_st_part_gt_ub 1 THEN REPEAT(assume_tac 1)); -by (dtac isLub_le_isUb 1 THEN assume_tac 1); -by (dtac lemma_minus_le_zero 1); -by (auto_tac (claset() addDs [order_less_le_trans], simpset())); -qed "lemma_st_part_le2"; - -Goal "((x::hypreal) <= t + r) = (x + -t <= r)"; -by Auto_tac; -qed "lemma_hypreal_le_swap"; - -Goal "[| x: HFinite; \ -\ isLub Reals {s. s: Reals & s < x} t; \ -\ r: Reals; 0 < r |] \ -\ ==> x + -t <= r"; -by (blast_tac (claset() addSIs [lemma_hypreal_le_swap RS iffD1, - lemma_st_part_le1]) 1); -qed "lemma_st_part1a"; - -Goal "(t + -r <= x) = (-(x + -t) <= (r::hypreal))"; -by Auto_tac; -qed "lemma_hypreal_le_swap2"; - -Goal "[| x: HFinite; \ -\ isLub Reals {s. s: Reals & s < x} t; \ -\ r: Reals; 0 < r |] \ -\ ==> -(x + -t) <= r"; -by (blast_tac (claset() addSIs [lemma_hypreal_le_swap2 RS iffD1, - lemma_st_part_le2]) 1); -qed "lemma_st_part2a"; - -Goal "(x::hypreal): Reals ==> isUb Reals {s. s: Reals & s < x} x"; -by (auto_tac (claset() addIs [isUbI, setleI,order_less_imp_le], simpset())); -qed "lemma_SReal_ub"; - -Goal "(x::hypreal): Reals ==> isLub Reals {s. s: Reals & s < x} x"; -by (auto_tac (claset() addSIs [isLubI2,lemma_SReal_ub,setgeI], simpset())); -by (ftac isUbD2a 1); -by (res_inst_tac [("x","x"),("y","y")] linorder_cases 1); -by (auto_tac (claset() addSIs [order_less_imp_le], simpset())); -by (EVERY1[dtac SReal_dense, assume_tac, assume_tac, Step_tac]); -by (dres_inst_tac [("y","r")] isUbD 1); -by (auto_tac (claset() addDs [order_less_le_trans], simpset())); -qed "lemma_SReal_lub"; - -Goal "[| x: HFinite; \ -\ isLub Reals {s. s: Reals & s < x} t; \ -\ r: Reals; 0 < r |] \ -\ ==> x + -t ~= r"; -by Auto_tac; -by (forward_tac [isLubD1a RS SReal_minus] 1); -by (dtac SReal_add_cancel 1 THEN assume_tac 1); -by (dres_inst_tac [("x","x")] lemma_SReal_lub 1); -by (dtac hypreal_isLub_unique 1 THEN assume_tac 1); -by Auto_tac; -qed "lemma_st_part_not_eq1"; - -Goal "[| x: HFinite; \ -\ isLub Reals {s. s: Reals & s < x} t; \ -\ r: Reals; 0 < r |] \ -\ ==> -(x + -t) ~= r"; -by (auto_tac (claset(), simpset() addsimps [minus_add_distrib])); -by (ftac isLubD1a 1); -by (dtac SReal_add_cancel 1 THEN assume_tac 1); -by (dres_inst_tac [("x","-x")] SReal_minus 1); -by (Asm_full_simp_tac 1); -by (dres_inst_tac [("x","x")] lemma_SReal_lub 1); -by (dtac hypreal_isLub_unique 1 THEN assume_tac 1); -by Auto_tac; -qed "lemma_st_part_not_eq2"; - -Goal "[| x: HFinite; \ -\ isLub Reals {s. s: Reals & s < x} t; \ -\ r: Reals; 0 < r |] \ -\ ==> abs (x + -t) < r"; -by (ftac lemma_st_part1a 1); -by (ftac lemma_st_part2a 4); -by Auto_tac; -by (REPEAT(dtac order_le_imp_less_or_eq 1)); -by (auto_tac (claset() addDs [lemma_st_part_not_eq1, lemma_st_part_not_eq2], - simpset() addsimps [abs_less_iff])); -qed "lemma_st_part_major"; - -Goal "[| x: HFinite; \ -\ isLub Reals {s. s: Reals & s < x} t |] \ -\ ==> ALL r: Reals. 0 < r --> abs (x + -t) < r"; -by (blast_tac (claset() addSDs [lemma_st_part_major]) 1); -qed "lemma_st_part_major2"; - -(*---------------------------------------------- - Existence of real and Standard Part Theorem - ----------------------------------------------*) -Goal "x: HFinite ==> \ -\ EX t: Reals. ALL r: Reals. 0 < r --> abs (x + -t) < r"; -by (ftac lemma_st_part_lub 1 THEN Step_tac 1); -by (ftac isLubD1a 1); -by (blast_tac (claset() addDs [lemma_st_part_major2]) 1); -qed "lemma_st_part_Ex"; - -Goalw [approx_def,Infinitesimal_def] - "x:HFinite ==> EX t: Reals. x @= t"; -by (dtac lemma_st_part_Ex 1); -by Auto_tac; -qed "st_part_Ex"; - -(*-------------------------------- - Unique real infinitely close - -------------------------------*) -Goal "x:HFinite ==> EX! t. t: Reals & x @= t"; -by (dtac st_part_Ex 1 THEN Step_tac 1); -by (dtac approx_sym 2 THEN dtac approx_sym 2 - THEN dtac approx_sym 2); -by (auto_tac (claset() addSIs [approx_unique_real], simpset())); -qed "st_part_Ex1"; - -(*------------------------------------------------------------------ - Finite and Infinite --- more theorems - ------------------------------------------------------------------*) - -Goalw [HFinite_def,HInfinite_def] "HFinite Int HInfinite = {}"; -by (auto_tac (claset() addIs [hypreal_less_irrefl] addDs [order_less_trans], - simpset())); -qed "HFinite_Int_HInfinite_empty"; -Addsimps [HFinite_Int_HInfinite_empty]; - -Goal "x: HFinite ==> x ~: HInfinite"; -by (EVERY1[Step_tac, dtac IntI, assume_tac]); -by Auto_tac; -qed "HFinite_not_HInfinite"; - -Goalw [HInfinite_def, HFinite_def] "x~: HFinite ==> x: HInfinite"; -by Auto_tac; -by (dres_inst_tac [("x","r + 1")] bspec 1); -by (auto_tac (claset(), simpset() addsimps [SReal_add])); -qed "not_HFinite_HInfinite"; - -Goal "x : HInfinite | x : HFinite"; -by (blast_tac (claset() addIs [not_HFinite_HInfinite]) 1); -qed "HInfinite_HFinite_disj"; - -Goal "(x : HInfinite) = (x ~: HFinite)"; -by (blast_tac (claset() addDs [HFinite_not_HInfinite, - not_HFinite_HInfinite]) 1); -qed "HInfinite_HFinite_iff"; - -Goal "(x : HFinite) = (x ~: HInfinite)"; -by (simp_tac (simpset() addsimps [HInfinite_HFinite_iff]) 1); -qed "HFinite_HInfinite_iff"; - -(*------------------------------------------------------------------ - Finite, Infinite and Infinitesimal --- more theorems - ------------------------------------------------------------------*) - -Goal "x ~: Infinitesimal ==> x : HInfinite | x : HFinite - Infinitesimal"; -by (fast_tac (claset() addIs [not_HFinite_HInfinite]) 1); -qed "HInfinite_diff_HFinite_Infinitesimal_disj"; - -Goal "[| x : HFinite; x ~: Infinitesimal |] ==> inverse x : HFinite"; -by (cut_inst_tac [("x","inverse x")] HInfinite_HFinite_disj 1); -by (auto_tac (claset() addSDs [HInfinite_inverse_Infinitesimal], simpset())); -qed "HFinite_inverse"; - -Goal "x : HFinite - Infinitesimal ==> inverse x : HFinite"; -by (blast_tac (claset() addIs [HFinite_inverse]) 1); -qed "HFinite_inverse2"; - -(* stronger statement possible in fact *) -Goal "x ~: Infinitesimal ==> inverse(x) : HFinite"; -by (dtac HInfinite_diff_HFinite_Infinitesimal_disj 1); -by (blast_tac (claset() addIs [HFinite_inverse, - HInfinite_inverse_Infinitesimal, - Infinitesimal_subset_HFinite RS subsetD]) 1); -qed "Infinitesimal_inverse_HFinite"; - -Goal "x : HFinite - Infinitesimal ==> inverse x : HFinite - Infinitesimal"; -by (auto_tac (claset() addIs [Infinitesimal_inverse_HFinite], simpset())); -by (dtac Infinitesimal_HFinite_mult2 1); -by (assume_tac 1); -by (asm_full_simp_tac - (simpset() addsimps [not_Infinitesimal_not_zero, hypreal_mult_inverse]) 1); -qed "HFinite_not_Infinitesimal_inverse"; - -Goal "[| x @= y; y : HFinite - Infinitesimal |] \ -\ ==> inverse x @= inverse y"; -by (ftac HFinite_diff_Infinitesimal_approx 1); -by (assume_tac 1); -by (ftac not_Infinitesimal_not_zero2 1); -by (forw_inst_tac [("x","x")] not_Infinitesimal_not_zero2 1); -by (REPEAT(dtac HFinite_inverse2 1)); -by (dtac approx_mult2 1 THEN assume_tac 1); -by Auto_tac; -by (dres_inst_tac [("c","inverse x")] approx_mult1 1 - THEN assume_tac 1); -by (auto_tac (claset() addIs [approx_sym], - simpset() addsimps [hypreal_mult_assoc])); -qed "approx_inverse"; - -(*Used for NSLIM_inverse, NSLIMSEQ_inverse*) -bind_thm ("hypreal_of_real_approx_inverse", - hypreal_of_real_HFinite_diff_Infinitesimal RSN (2, approx_inverse)); - -Goal "[| x: HFinite - Infinitesimal; \ -\ h : Infinitesimal |] ==> inverse(x + h) @= inverse x"; -by (auto_tac (claset() addIs [approx_inverse, approx_sym, - Infinitesimal_add_approx_self], - simpset())); -qed "inverse_add_Infinitesimal_approx"; - -Goal "[| x: HFinite - Infinitesimal; \ -\ h : Infinitesimal |] ==> inverse(h + x) @= inverse x"; -by (rtac (hypreal_add_commute RS subst) 1); -by (blast_tac (claset() addIs [inverse_add_Infinitesimal_approx]) 1); -qed "inverse_add_Infinitesimal_approx2"; - -Goal "[| x : HFinite - Infinitesimal; \ -\ h : Infinitesimal |] ==> inverse(x + h) + -inverse x @= h"; -by (rtac approx_trans2 1); -by (auto_tac (claset() addIs [inverse_add_Infinitesimal_approx], - simpset() addsimps [mem_infmal_iff, - approx_minus_iff RS sym])); -qed "inverse_add_Infinitesimal_approx_Infinitesimal"; - -Goal "(x : Infinitesimal) = (x*x : Infinitesimal)"; -by (auto_tac (claset() addIs [Infinitesimal_mult], simpset())); -by (rtac ccontr 1 THEN ftac Infinitesimal_inverse_HFinite 1); -by (ftac not_Infinitesimal_not_zero 1); -by (auto_tac (claset() addDs [Infinitesimal_HFinite_mult], - simpset() addsimps [hypreal_mult_assoc])); -qed "Infinitesimal_square_iff"; -Addsimps [Infinitesimal_square_iff RS sym]; - -Goal "(x*x : HFinite) = (x : HFinite)"; -by (auto_tac (claset() addIs [HFinite_mult], simpset())); -by (auto_tac (claset() addDs [HInfinite_mult], - simpset() addsimps [HFinite_HInfinite_iff])); -qed "HFinite_square_iff"; -Addsimps [HFinite_square_iff]; - -Goal "(x*x : HInfinite) = (x : HInfinite)"; -by (auto_tac (claset(), simpset() addsimps - [HInfinite_HFinite_iff])); -qed "HInfinite_square_iff"; -Addsimps [HInfinite_square_iff]; - -Goal "[| a: HFinite-Infinitesimal; a* w @= a*z |] ==> w @= z"; -by (Step_tac 1); -by (ftac HFinite_inverse 1 THEN assume_tac 1); -by (dtac not_Infinitesimal_not_zero 1); -by (auto_tac (claset() addDs [approx_mult2], - simpset() addsimps [hypreal_mult_assoc RS sym])); -qed "approx_HFinite_mult_cancel"; - -Goal "a: HFinite-Infinitesimal ==> (a * w @= a * z) = (w @= z)"; -by (auto_tac (claset() addIs [approx_mult2, - approx_HFinite_mult_cancel], simpset())); -qed "approx_HFinite_mult_cancel_iff1"; - -(*------------------------------------------------------------------ - more about absolute value (hrabs) - ------------------------------------------------------------------*) - -Goal "abs x @= x | abs x @= -x"; -by (cut_inst_tac [("x","x")] hrabs_disj 1); -by Auto_tac; -qed "approx_hrabs_disj"; - -(*------------------------------------------------------------------ - Theorems about monads - ------------------------------------------------------------------*) - -Goal "monad (abs x) <= monad(x) Un monad(-x)"; -by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1); -by Auto_tac; -qed "monad_hrabs_Un_subset"; - -Goal "e : Infinitesimal ==> monad (x+e) = monad x"; -by (fast_tac (claset() addSIs [Infinitesimal_add_approx_self RS approx_sym, - approx_monad_iff RS iffD1]) 1); -qed "Infinitesimal_monad_eq"; - -Goalw [monad_def] "(u:monad x) = (-u:monad (-x))"; -by Auto_tac; -qed "mem_monad_iff"; - -Goalw [monad_def] "(x:Infinitesimal) = (x:monad 0)"; -by (auto_tac (claset() addIs [approx_sym], - simpset() addsimps [mem_infmal_iff])); -qed "Infinitesimal_monad_zero_iff"; - -Goal "(x:monad 0) = (-x:monad 0)"; -by (simp_tac (simpset() addsimps [Infinitesimal_monad_zero_iff RS sym]) 1); -qed "monad_zero_minus_iff"; - -Goal "(x:monad 0) = (abs x:monad 0)"; -by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1); -by (auto_tac (claset(), simpset() addsimps [monad_zero_minus_iff RS sym])); -qed "monad_zero_hrabs_iff"; - -Goalw [monad_def] "x:monad x"; -by (Simp_tac 1); -qed "mem_monad_self"; -Addsimps [mem_monad_self]; - -(*------------------------------------------------------------------ - Proof that x @= y ==> abs x @= abs y - ------------------------------------------------------------------*) -Goal "x @= y ==> {x,y}<=monad x"; -by (Simp_tac 1); -by (asm_full_simp_tac (simpset() addsimps - [approx_monad_iff]) 1); -qed "approx_subset_monad"; - -Goal "x @= y ==> {x,y}<=monad y"; -by (dtac approx_sym 1); -by (fast_tac (claset() addDs [approx_subset_monad]) 1); -qed "approx_subset_monad2"; - -Goalw [monad_def] "u:monad x ==> x @= u"; -by (Fast_tac 1); -qed "mem_monad_approx"; - -Goalw [monad_def] "x @= u ==> u:monad x"; -by (Fast_tac 1); -qed "approx_mem_monad"; - -Goalw [monad_def] "x @= u ==> x:monad u"; -by (blast_tac (claset() addSIs [approx_sym]) 1); -qed "approx_mem_monad2"; - -Goal "[| x @= y;x:monad 0 |] ==> y:monad 0"; -by (dtac mem_monad_approx 1); -by (fast_tac (claset() addIs [approx_mem_monad,approx_trans]) 1); -qed "approx_mem_monad_zero"; - -Goal "[| x @= y; x: Infinitesimal |] ==> abs x @= abs y"; -by (dtac (Infinitesimal_monad_zero_iff RS iffD1) 1); -by (blast_tac (claset() addIs [approx_mem_monad_zero, - monad_zero_hrabs_iff RS iffD1, mem_monad_approx, approx_trans3]) 1); -qed "Infinitesimal_approx_hrabs"; - -Goal "[| 0 < x; x ~:Infinitesimal; e :Infinitesimal |] ==> e < x"; -by (rtac ccontr 1); -by (auto_tac (claset() - addIs [Infinitesimal_zero RSN (2, Infinitesimal_interval)] - addSDs [order_le_imp_less_or_eq], - simpset() addsimps [linorder_not_less])); -qed "less_Infinitesimal_less"; - -Goal "[| 0 < x; x ~: Infinitesimal; u: monad x |] ==> 0 < u"; -by (dtac (mem_monad_approx RS approx_sym) 1); -by (etac (bex_Infinitesimal_iff2 RS iffD2 RS bexE) 1); -by (dres_inst_tac [("e","-xa")] less_Infinitesimal_less 1); -by Auto_tac; -qed "Ball_mem_monad_gt_zero"; - -Goal "[| x < 0; x ~: Infinitesimal; u: monad x |] ==> u < 0"; -by (dtac (mem_monad_approx RS approx_sym) 1); -by (etac (bex_Infinitesimal_iff RS iffD2 RS bexE) 1); -by (cut_inst_tac [("x","-x"),("e","xa")] less_Infinitesimal_less 1); -by Auto_tac; -qed "Ball_mem_monad_less_zero"; - -Goal "[|0 < x; x ~: Infinitesimal; x @= y|] ==> 0 < y"; -by (blast_tac (claset() addDs [Ball_mem_monad_gt_zero, - approx_subset_monad]) 1); -qed "lemma_approx_gt_zero"; - -Goal "[|x < 0; x ~: Infinitesimal; x @= y|] ==> y < 0"; -by (blast_tac (claset() addDs [Ball_mem_monad_less_zero, - approx_subset_monad]) 1); -qed "lemma_approx_less_zero"; - -Goal "[| x @= y; x < 0; x ~: Infinitesimal |] ==> abs x @= abs y"; -by (ftac lemma_approx_less_zero 1); -by (REPEAT(assume_tac 1)); -by (REPEAT(dtac hrabs_minus_eqI2 1)); -by Auto_tac; -qed "approx_hrabs1"; - -Goal "[| x @= y; 0 < x; x ~: Infinitesimal |] ==> abs x @= abs y"; -by (ftac lemma_approx_gt_zero 1); -by (REPEAT(assume_tac 1)); -by (REPEAT(dtac hrabs_eqI2 1)); -by Auto_tac; -qed "approx_hrabs2"; - -Goal "x @= y ==> abs x @= abs y"; -by (res_inst_tac [("Q","x:Infinitesimal")] (excluded_middle RS disjE) 1); -by (res_inst_tac [("x1","x"),("y1","0")] (hypreal_linear RS disjE) 1); -by (auto_tac (claset() addIs [approx_hrabs1,approx_hrabs2, - Infinitesimal_approx_hrabs], simpset())); -qed "approx_hrabs"; - -Goal "abs(x) @= 0 ==> x @= 0"; -by (cut_inst_tac [("x","x")] hrabs_disj 1); -by (auto_tac (claset() addDs [approx_minus], simpset())); -qed "approx_hrabs_zero_cancel"; - -Goal "e: Infinitesimal ==> abs x @= abs(x+e)"; -by (fast_tac (claset() addIs [approx_hrabs, - Infinitesimal_add_approx_self]) 1); -qed "approx_hrabs_add_Infinitesimal"; - -Goal "e: Infinitesimal ==> abs x @= abs(x + -e)"; -by (fast_tac (claset() addIs [approx_hrabs, - Infinitesimal_add_minus_approx_self]) 1); -qed "approx_hrabs_add_minus_Infinitesimal"; - -Goal "[| e: Infinitesimal; e': Infinitesimal; \ -\ abs(x+e) = abs(y+e')|] ==> abs x @= abs y"; -by (dres_inst_tac [("x","x")] approx_hrabs_add_Infinitesimal 1); -by (dres_inst_tac [("x","y")] approx_hrabs_add_Infinitesimal 1); -by (auto_tac (claset() addIs [approx_trans2], simpset())); -qed "hrabs_add_Infinitesimal_cancel"; - -Goal "[| e: Infinitesimal; e': Infinitesimal; \ -\ abs(x + -e) = abs(y + -e')|] ==> abs x @= abs y"; -by (dres_inst_tac [("x","x")] approx_hrabs_add_minus_Infinitesimal 1); -by (dres_inst_tac [("x","y")] approx_hrabs_add_minus_Infinitesimal 1); -by (auto_tac (claset() addIs [approx_trans2], simpset())); -qed "hrabs_add_minus_Infinitesimal_cancel"; - -(* interesting slightly counterintuitive theorem: necessary - for proving that an open interval is an NS open set -*) -Goalw [Infinitesimal_def] - "[| x < y; u: Infinitesimal |] \ -\ ==> hypreal_of_real x + u < hypreal_of_real y"; -by (dtac (hypreal_of_real_less_iff RS iffD2) 1); -by (dtac (hypreal_less_minus_iff RS iffD1) 1 THEN Step_tac 1); -by (rtac (hypreal_less_minus_iff RS iffD2) 1); -by (dres_inst_tac [("x","hypreal_of_real y + -hypreal_of_real x")] bspec 1); -by (auto_tac (claset(), - simpset() addsimps [hypreal_add_commute, - abs_less_iff, - SReal_add, SReal_minus])); -qed "Infinitesimal_add_hypreal_of_real_less"; - -Goal "[| x: Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |] \ -\ ==> abs (hypreal_of_real r + x) < hypreal_of_real y"; -by (dres_inst_tac [("x","hypreal_of_real r")] - approx_hrabs_add_Infinitesimal 1); -by (dtac (approx_sym RS (bex_Infinitesimal_iff2 RS iffD2)) 1); -by (auto_tac (claset() addSIs [Infinitesimal_add_hypreal_of_real_less], - simpset() addsimps [hypreal_of_real_hrabs])); -qed "Infinitesimal_add_hrabs_hypreal_of_real_less"; - -Goal "[| x: Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |] \ -\ ==> abs (x + hypreal_of_real r) < hypreal_of_real y"; -by (rtac (hypreal_add_commute RS subst) 1); -by (etac Infinitesimal_add_hrabs_hypreal_of_real_less 1); -by (assume_tac 1); -qed "Infinitesimal_add_hrabs_hypreal_of_real_less2"; - -Goalw [hypreal_le_def] - "[| u: Infinitesimal; hypreal_of_real x + u <= hypreal_of_real y |] \ -\ ==> hypreal_of_real x <= hypreal_of_real y"; -by (EVERY1 [rtac notI, rtac contrapos_np, assume_tac]); -by (res_inst_tac [("c1","-u")] (add_less_cancel_right RS iffD1) 1); -by (Asm_full_simp_tac 1); -by (dtac (Infinitesimal_minus_iff RS iffD2) 1); -by (dtac Infinitesimal_add_hypreal_of_real_less 1); -by (assume_tac 1); -by Auto_tac; -qed "Infinitesimal_add_cancel_hypreal_of_real_le"; - -Goal "[| u: Infinitesimal; hypreal_of_real x + u <= hypreal_of_real y |] \ -\ ==> x <= y"; -by (blast_tac (claset() addSIs [hypreal_of_real_le_iff RS iffD1, - Infinitesimal_add_cancel_hypreal_of_real_le]) 1); -qed "Infinitesimal_add_cancel_real_le"; - -Goal "[| u: Infinitesimal; v: Infinitesimal; \ -\ hypreal_of_real x + u <= hypreal_of_real y + v |] \ -\ ==> hypreal_of_real x <= hypreal_of_real y"; -by (asm_full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); -by Auto_tac; -by (dres_inst_tac [("u","v-u")] Infinitesimal_add_hypreal_of_real_less 1); -by (auto_tac (claset(), simpset() addsimps [Infinitesimal_diff])); -qed "hypreal_of_real_le_add_Infininitesimal_cancel"; - -Goal "[| u: Infinitesimal; v: Infinitesimal; \ -\ hypreal_of_real x + u <= hypreal_of_real y + v |] \ -\ ==> x <= y"; -by (blast_tac (claset() addSIs [hypreal_of_real_le_iff RS iffD1, - hypreal_of_real_le_add_Infininitesimal_cancel]) 1); -qed "hypreal_of_real_le_add_Infininitesimal_cancel2"; - -Goal "[| hypreal_of_real x < e; e: Infinitesimal |] ==> hypreal_of_real x <= 0"; -by (rtac (linorder_not_less RS iffD1) 1 THEN Safe_tac); -by (dtac Infinitesimal_interval 1); -by (dtac (SReal_hypreal_of_real RS SReal_Infinitesimal_zero) 4); -by Auto_tac; -qed "hypreal_of_real_less_Infinitesimal_le_zero"; - -(*used once, in Lim/NSDERIV_inverse*) -Goal "[| h: Infinitesimal; x ~= 0 |] ==> hypreal_of_real x + h ~= 0"; -by Auto_tac; -by (subgoal_tac "h = - hypreal_of_real x" 1); -by Auto_tac; -qed "Infinitesimal_add_not_zero"; - -Goal "x*x + y*y : Infinitesimal ==> x*x : Infinitesimal"; -by (rtac Infinitesimal_interval2 1); -by (rtac zero_le_square 3); -by (assume_tac 1); -by (auto_tac (claset(), simpset() addsimps [zero_le_square])); -qed "Infinitesimal_square_cancel"; -Addsimps [Infinitesimal_square_cancel]; - -Goal "x*x + y*y : HFinite ==> x*x : HFinite"; -by (rtac HFinite_bounded 1); -by (assume_tac 1); -by (auto_tac (claset(), simpset() addsimps [zero_le_square])); -qed "HFinite_square_cancel"; -Addsimps [HFinite_square_cancel]; - -Goal "x*x + y*y : Infinitesimal ==> y*y : Infinitesimal"; -by (rtac Infinitesimal_square_cancel 1); -by (rtac (hypreal_add_commute RS subst) 1); -by (Simp_tac 1); -qed "Infinitesimal_square_cancel2"; -Addsimps [Infinitesimal_square_cancel2]; - -Goal "x*x + y*y : HFinite ==> y*y : HFinite"; -by (rtac HFinite_square_cancel 1); -by (rtac (hypreal_add_commute RS subst) 1); -by (Simp_tac 1); -qed "HFinite_square_cancel2"; -Addsimps [HFinite_square_cancel2]; - -Goal "x*x + y*y + z*z : Infinitesimal ==> x*x : Infinitesimal"; -by (rtac Infinitesimal_interval2 1); -by (assume_tac 1); -by (rtac zero_le_square 2); -by (Asm_full_simp_tac 1); -by (cut_inst_tac [("a","y")] (thm"zero_le_square") 1); -by (cut_inst_tac [("a","z")] (thm"zero_le_square") 1); -by (asm_simp_tac (simpset() addsimps []) 1); -qed "Infinitesimal_sum_square_cancel"; -Addsimps [Infinitesimal_sum_square_cancel]; - -Goal "x*x + y*y + z*z : HFinite ==> x*x : HFinite"; -by (rtac HFinite_bounded 1); -by (assume_tac 1); -by (rtac zero_le_square 2); -by (cut_inst_tac [("a","y")] (thm"zero_le_square") 1); -by (cut_inst_tac [("a","z")] (thm"zero_le_square") 1); -by (asm_simp_tac (simpset() addsimps []) 1); -qed "HFinite_sum_square_cancel"; -Addsimps [HFinite_sum_square_cancel]; - -Goal "y*y + x*x + z*z : Infinitesimal ==> x*x : Infinitesimal"; -by (rtac Infinitesimal_sum_square_cancel 1); -by (asm_full_simp_tac (simpset() addsimps add_ac) 1); -qed "Infinitesimal_sum_square_cancel2"; -Addsimps [Infinitesimal_sum_square_cancel2]; - -Goal "y*y + x*x + z*z : HFinite ==> x*x : HFinite"; -by (rtac HFinite_sum_square_cancel 1); -by (asm_full_simp_tac (simpset() addsimps add_ac) 1); -qed "HFinite_sum_square_cancel2"; -Addsimps [HFinite_sum_square_cancel2]; - -Goal "z*z + y*y + x*x : Infinitesimal ==> x*x : Infinitesimal"; -by (rtac Infinitesimal_sum_square_cancel 1); -by (asm_full_simp_tac (simpset() addsimps add_ac) 1); -qed "Infinitesimal_sum_square_cancel3"; -Addsimps [Infinitesimal_sum_square_cancel3]; - -Goal "z*z + y*y + x*x : HFinite ==> x*x : HFinite"; -by (rtac HFinite_sum_square_cancel 1); -by (asm_full_simp_tac (simpset() addsimps add_ac) 1); -qed "HFinite_sum_square_cancel3"; -Addsimps [HFinite_sum_square_cancel3]; - -Goal "[| y: monad x; 0 < hypreal_of_real e |] \ -\ ==> abs (y + -x) < hypreal_of_real e"; -by (dtac (mem_monad_approx RS approx_sym) 1); -by (dtac (bex_Infinitesimal_iff RS iffD2) 1); -by (auto_tac (claset() addSDs [InfinitesimalD], simpset())); -qed "monad_hrabs_less"; - -Goal "x: monad (hypreal_of_real a) ==> x: HFinite"; -by (dtac (mem_monad_approx RS approx_sym) 1); -by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1); -by (step_tac (claset() addSDs - [Infinitesimal_subset_HFinite RS subsetD]) 1); -by (etac (SReal_hypreal_of_real RS (SReal_subset_HFinite - RS subsetD) RS HFinite_add) 1); -qed "mem_monad_SReal_HFinite"; - -(*------------------------------------------------------------------ - Theorems about standard part - ------------------------------------------------------------------*) - -Goalw [st_def] "x: HFinite ==> st x @= x"; -by (ftac st_part_Ex 1 THEN Step_tac 1); -by (rtac someI2 1); -by (auto_tac (claset() addIs [approx_sym], simpset())); -qed "st_approx_self"; - -Goalw [st_def] "x: HFinite ==> st x: Reals"; -by (ftac st_part_Ex 1 THEN Step_tac 1); -by (rtac someI2 1); -by (auto_tac (claset() addIs [approx_sym], simpset())); -qed "st_SReal"; - -Goal "x: HFinite ==> st x: HFinite"; -by (etac (st_SReal RS (SReal_subset_HFinite RS subsetD)) 1); -qed "st_HFinite"; - -Goalw [st_def] "x: Reals ==> st x = x"; -by (rtac some_equality 1); -by (fast_tac (claset() addIs [(SReal_subset_HFinite RS subsetD)]) 1); -by (blast_tac (claset() addDs [SReal_approx_iff RS iffD1]) 1); -qed "st_SReal_eq"; - -(* should be added to simpset *) -Goal "st (hypreal_of_real x) = hypreal_of_real x"; -by (rtac (SReal_hypreal_of_real RS st_SReal_eq) 1); -qed "st_hypreal_of_real"; - -Goal "[| x: HFinite; y: HFinite; st x = st y |] ==> x @= y"; -by (auto_tac (claset() addSDs [st_approx_self] - addSEs [approx_trans3], simpset())); -qed "st_eq_approx"; - -Goal "[| x: HFinite; y: HFinite; x @= y |] ==> st x = st y"; -by (EVERY1 [ftac st_approx_self , - forw_inst_tac [("x","y")] st_approx_self, - dtac st_SReal,dtac st_SReal]); -by (fast_tac (claset() addEs [approx_trans, - approx_trans2,SReal_approx_iff RS iffD1]) 1); -qed "approx_st_eq"; - -Goal "[| x: HFinite; y: HFinite|] \ -\ ==> (x @= y) = (st x = st y)"; -by (blast_tac (claset() addIs [approx_st_eq, - st_eq_approx]) 1); -qed "st_eq_approx_iff"; - -Goal "[| x: Reals; e: Infinitesimal |] ==> st(x + e) = x"; -by (forward_tac [st_SReal_eq RS subst] 1); -by (assume_tac 2); -by (forward_tac [SReal_subset_HFinite RS subsetD] 1); -by (forward_tac [Infinitesimal_subset_HFinite RS subsetD] 1); -by (dtac st_SReal_eq 1); -by (rtac approx_st_eq 1); -by (auto_tac (claset() addIs [HFinite_add], - simpset() addsimps [Infinitesimal_add_approx_self - RS approx_sym])); -qed "st_Infinitesimal_add_SReal"; - -Goal "[| x: Reals; e: Infinitesimal |] \ -\ ==> st(e + x) = x"; -by (rtac (hypreal_add_commute RS subst) 1); -by (blast_tac (claset() addSIs [st_Infinitesimal_add_SReal]) 1); -qed "st_Infinitesimal_add_SReal2"; - -Goal "x: HFinite ==> \ -\ EX e: Infinitesimal. x = st(x) + e"; -by (blast_tac (claset() addSDs [(st_approx_self RS - approx_sym),bex_Infinitesimal_iff2 RS iffD2]) 1); -qed "HFinite_st_Infinitesimal_add"; - -Goal "[| x: HFinite; y: HFinite |] ==> st (x + y) = st(x) + st(y)"; -by (ftac HFinite_st_Infinitesimal_add 1); -by (forw_inst_tac [("x","y")] HFinite_st_Infinitesimal_add 1); -by (Step_tac 1); -by (subgoal_tac "st (x + y) = st ((st x + e) + (st y + ea))" 1); -by (dtac sym 2 THEN dtac sym 2); -by (Asm_full_simp_tac 2); -by (asm_simp_tac (simpset() addsimps add_ac) 1); -by (REPEAT(dtac st_SReal 1)); -by (dtac SReal_add 1 THEN assume_tac 1); -by (dtac Infinitesimal_add 1 THEN assume_tac 1); -by (rtac (hypreal_add_assoc RS subst) 1); -by (blast_tac (claset() addSIs [st_Infinitesimal_add_SReal2]) 1); -qed "st_add"; - -Goal "st (number_of w) = number_of w"; -by (rtac (SReal_number_of RS st_SReal_eq) 1); -qed "st_number_of"; -Addsimps [st_number_of]; - -(*the theorem above for the special cases of zero and one*) -Addsimps - (map (simplify (simpset())) - [inst "w" "bin.Pls" st_number_of, inst "w" "bin.Pls BIT True" st_number_of]); - -Goal "y: HFinite ==> st(-y) = -st(y)"; -by (forward_tac [HFinite_minus_iff RS iffD2] 1); -by (rtac hypreal_add_minus_eq_minus 1); -by (dtac (st_add RS sym) 1 THEN assume_tac 1); -by Auto_tac; -qed "st_minus"; - -Goalw [hypreal_diff_def] - "[| x: HFinite; y: HFinite |] ==> st (x-y) = st(x) - st(y)"; -by (forw_inst_tac [("y1","y")] (st_minus RS sym) 1); -by (dres_inst_tac [("x1","y")] (HFinite_minus_iff RS iffD2) 1); -by (asm_simp_tac (simpset() addsimps [st_add]) 1); -qed "st_diff"; - -(* lemma *) -Goal "[| x: HFinite; y: HFinite; \ -\ e: Infinitesimal; \ -\ ea : Infinitesimal |] \ -\ ==> e*y + x*ea + e*ea: Infinitesimal"; -by (forw_inst_tac [("x","e"),("y","y")] Infinitesimal_HFinite_mult 1); -by (forw_inst_tac [("x","ea"),("y","x")] Infinitesimal_HFinite_mult 2); -by (dtac Infinitesimal_mult 3); -by (auto_tac (claset() addIs [Infinitesimal_add], - simpset() addsimps add_ac @ mult_ac)); -qed "lemma_st_mult"; - -Goal "[| x: HFinite; y: HFinite |] \ -\ ==> st (x * y) = st(x) * st(y)"; -by (ftac HFinite_st_Infinitesimal_add 1); -by (forw_inst_tac [("x","y")] HFinite_st_Infinitesimal_add 1); -by (Step_tac 1); -by (subgoal_tac "st (x * y) = st ((st x + e) * (st y + ea))" 1); -by (dtac sym 2 THEN dtac sym 2); -by (Asm_full_simp_tac 2); -by (thin_tac "x = st x + e" 1); -by (thin_tac "y = st y + ea" 1); -by (asm_full_simp_tac (simpset() addsimps [left_distrib,right_distrib]) 1); -by (REPEAT(dtac st_SReal 1)); -by (full_simp_tac (simpset() addsimps [hypreal_add_assoc]) 1); -by (rtac st_Infinitesimal_add_SReal 1); -by (blast_tac (claset() addSIs [SReal_mult]) 1); -by (REPEAT(dtac (SReal_subset_HFinite RS subsetD) 1)); -by (rtac (hypreal_add_assoc RS subst) 1); -by (blast_tac (claset() addSIs [lemma_st_mult]) 1); -qed "st_mult"; - -Goal "x: Infinitesimal ==> st x = 0"; -by (stac (hypreal_numeral_0_eq_0 RS sym) 1); -by (rtac (st_number_of RS subst) 1); -by (rtac approx_st_eq 1); -by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD], - simpset() addsimps [mem_infmal_iff RS sym])); -qed "st_Infinitesimal"; - -Goal "st(x) ~= 0 ==> x ~: Infinitesimal"; -by (fast_tac (claset() addIs [st_Infinitesimal]) 1); -qed "st_not_Infinitesimal"; - -Goal "[| x: HFinite; st x ~= 0 |] \ -\ ==> st(inverse x) = inverse (st x)"; -by (res_inst_tac [("c1","st x")] (hypreal_mult_left_cancel RS iffD1) 1); -by (auto_tac (claset(), - simpset() addsimps [st_mult RS sym, st_not_Infinitesimal, - HFinite_inverse])); -by (stac hypreal_mult_inverse 1); -by Auto_tac; -qed "st_inverse"; - -Goal "[| x: HFinite; y: HFinite; st y ~= 0 |] \ -\ ==> st(x/y) = (st x) / (st y)"; -by (auto_tac (claset(), - simpset() addsimps [hypreal_divide_def, st_mult, st_not_Infinitesimal, - HFinite_inverse, st_inverse])); -qed "st_divide"; -Addsimps [st_divide]; - -Goal "x: HFinite ==> st(st(x)) = st(x)"; -by (blast_tac (claset() addIs [st_HFinite, st_approx_self, - approx_st_eq]) 1); -qed "st_idempotent"; -Addsimps [st_idempotent]; - -(*** lemmas ***) -Goal "[| x: HFinite; y: HFinite; \ -\ u: Infinitesimal; st x < st y \ -\ |] ==> st x + u < st y"; -by (REPEAT(dtac st_SReal 1)); -by (auto_tac (claset() addSIs [Infinitesimal_add_hypreal_of_real_less], - simpset() addsimps [SReal_iff])); -qed "Infinitesimal_add_st_less"; - -Goalw [hypreal_le_def] - "[| x: HFinite; y: HFinite; \ -\ u: Infinitesimal; st x <= st y + u\ -\ |] ==> st x <= st y"; -by (auto_tac (claset() addDs [Infinitesimal_add_st_less], - simpset())); -qed "Infinitesimal_add_st_le_cancel"; - -Goal "[| x: HFinite; y: HFinite; x <= y |] ==> st(x) <= st(y)"; -by (ftac HFinite_st_Infinitesimal_add 1); -by (rotate_tac 1 1); -by (ftac HFinite_st_Infinitesimal_add 1); -by (Step_tac 1); -by (rtac Infinitesimal_add_st_le_cancel 1); -by (res_inst_tac [("x","ea"),("y","e")] - Infinitesimal_diff 3); -by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc RS sym])); -qed "st_le"; - -Goal "[| 0 <= x; x: HFinite |] ==> 0 <= st x"; -by (stac (hypreal_numeral_0_eq_0 RS sym) 1); -by (rtac (st_number_of RS subst) 1); -by (rtac st_le 1); -by Auto_tac; -qed "st_zero_le"; - -Goal "[| x <= 0; x: HFinite |] ==> st x <= 0"; -by (stac (hypreal_numeral_0_eq_0 RS sym) 1); -by (rtac (st_number_of RS subst) 1); -by (rtac st_le 1); -by Auto_tac; -qed "st_zero_ge"; - -Goal "x: HFinite ==> abs(st x) = st(abs x)"; -by (case_tac "0 <= x" 1); -by (auto_tac (claset() addSDs [order_less_imp_le], - simpset() addsimps [linorder_not_le,st_zero_le,hrabs_eqI1, hrabs_minus_eqI1, - st_zero_ge,st_minus])); -qed "st_hrabs"; - -(*-------------------------------------------------------------------- - Alternative definitions for HFinite using Free ultrafilter - --------------------------------------------------------------------*) - -Goal "[| X: Rep_hypreal x; Y: Rep_hypreal x |] \ -\ ==> {n. X n = Y n} : FreeUltrafilterNat"; -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by Auto_tac; -by (Ultra_tac 1); -qed "FreeUltrafilterNat_Rep_hypreal"; - -Goal "{n. Yb n < Y n} Int {n. -y = Yb n} <= {n. -y < Y n}"; -by Auto_tac; -qed "lemma_free1"; - -Goal "{n. Xa n < Yc n} Int {n. y = Yc n} <= {n. Xa n < y}"; -by Auto_tac; -qed "lemma_free2"; - -Goalw [HFinite_def] - "x : HFinite ==> EX X: Rep_hypreal x. \ -\ EX u. {n. abs (X n) < u}: FreeUltrafilterNat"; -by (auto_tac (claset(), simpset() addsimps [abs_less_iff, inst "a" "x" minus_less_iff])); -by (auto_tac (claset(), simpset() addsimps - [hypreal_less_def,SReal_iff,hypreal_minus, - hypreal_of_real_def])); -by (dtac FreeUltrafilterNat_Rep_hypreal 1 THEN assume_tac 1); -by (rename_tac "Z" 1); -by (res_inst_tac [("x","Z")] bexI 1 THEN assume_tac 2); -by (res_inst_tac [("x","y")] exI 1); -by (Ultra_tac 1); -qed "HFinite_FreeUltrafilterNat"; - -Goalw [HFinite_def] - "EX X: Rep_hypreal x. \ -\ EX u. {n. abs (X n) < u}: FreeUltrafilterNat\ -\ ==> x : HFinite"; -by (auto_tac (claset(), simpset() addsimps [abs_less_iff, inst "a" "x" minus_less_iff])); -by (res_inst_tac [("x","hypreal_of_real u")] bexI 1); -by (auto_tac (claset() addSIs [exI], - simpset() addsimps - [hypreal_less_def,SReal_iff,hypreal_minus, hypreal_of_real_def])); -by (ALLGOALS Ultra_tac); -qed "FreeUltrafilterNat_HFinite"; - -Goal "(x : HFinite) = (EX X: Rep_hypreal x. \ -\ EX u. {n. abs (X n) < u}: FreeUltrafilterNat)"; -by (blast_tac (claset() addSIs [HFinite_FreeUltrafilterNat, - FreeUltrafilterNat_HFinite]) 1); -qed "HFinite_FreeUltrafilterNat_iff"; - -(*-------------------------------------------------------------------- - Alternative definitions for HInfinite using Free ultrafilter - --------------------------------------------------------------------*) -Goal "- {n. (u::real) < abs (xa n)} = {n. abs (xa n) <= u}"; -by Auto_tac; -qed "lemma_Compl_eq"; - -Goal "- {n. abs (xa n) < (u::real)} = {n. u <= abs (xa n)}"; -by Auto_tac; -qed "lemma_Compl_eq2"; - -Goal "{n. abs (xa n) <= (u::real)} Int {n. u <= abs (xa n)} \ -\ = {n. abs(xa n) = u}"; -by Auto_tac; -qed "lemma_Int_eq1"; - -Goal "{n. abs (xa n) = u} <= {n. abs (xa n) < u + (1::real)}"; -by Auto_tac; -qed "lemma_FreeUltrafilterNat_one"; - -(*------------------------------------- - Exclude this type of sets from free - ultrafilter for Infinite numbers! - -------------------------------------*) -Goal "[| xa: Rep_hypreal x; \ -\ {n. abs (xa n) = u} : FreeUltrafilterNat \ -\ |] ==> x: HFinite"; -by (rtac FreeUltrafilterNat_HFinite 1); -by (res_inst_tac [("x","xa")] bexI 1); -by (res_inst_tac [("x","u + 1")] exI 1); -by (Ultra_tac 1 THEN assume_tac 1); -qed "FreeUltrafilterNat_const_Finite"; - -val [prem] = goal thy "x : HInfinite ==> EX X: Rep_hypreal x. \ -\ ALL u. {n. u < abs (X n)}: FreeUltrafilterNat"; -by (cut_facts_tac [(prem RS (HInfinite_HFinite_iff RS iffD1))] 1); -by (cut_inst_tac [("x","x")] Rep_hypreal_nonempty 1); -by (auto_tac (claset(), simpset() delsimps [Rep_hypreal_nonempty] - addsimps [HFinite_FreeUltrafilterNat_iff,Bex_def])); -by (REPEAT(dtac spec 1)); -by Auto_tac; -by (dres_inst_tac [("x","u")] spec 1 THEN - REPEAT(dtac FreeUltrafilterNat_Compl_mem 1)); -by (dtac FreeUltrafilterNat_Int 1 THEN assume_tac 1); -by (asm_full_simp_tac (simpset() addsimps [lemma_Compl_eq, - lemma_Compl_eq2,lemma_Int_eq1]) 1); -by (auto_tac (claset() addDs [FreeUltrafilterNat_const_Finite], - simpset() addsimps [(prem RS (HInfinite_HFinite_iff RS iffD1))])); -qed "HInfinite_FreeUltrafilterNat"; - -(* yet more lemmas! *) -Goal "{n. abs (Xa n) < u} Int {n. X n = Xa n} \ -\ <= {n. abs (X n) < (u::real)}"; -by Auto_tac; -qed "lemma_Int_HI"; - -Goal "{n. u < abs (X n)} Int {n. abs (X n) < (u::real)} = {}"; -by (auto_tac (claset() addIs [order_less_asym], simpset())); -qed "lemma_Int_HIa"; - -Goal "EX X: Rep_hypreal x. ALL u. \ -\ {n. u < abs (X n)}: FreeUltrafilterNat\ -\ ==> x : HInfinite"; -by (rtac (HInfinite_HFinite_iff RS iffD2) 1); -by (Step_tac 1 THEN dtac HFinite_FreeUltrafilterNat 1); -by Auto_tac; -by (dres_inst_tac [("x","u")] spec 1); -by (dtac FreeUltrafilterNat_Rep_hypreal 1 THEN assume_tac 1); -by (dres_inst_tac [("Y","{n. X n = Xa n}")] FreeUltrafilterNat_Int 1); -by (dtac (lemma_Int_HI RSN (2,FreeUltrafilterNat_subset)) 2); -by (dres_inst_tac [("Y","{n. abs (X n) < u}")] FreeUltrafilterNat_Int 2); -by (auto_tac (claset(), simpset() addsimps [lemma_Int_HIa, - FreeUltrafilterNat_empty])); -qed "FreeUltrafilterNat_HInfinite"; - -Goal "(x : HInfinite) = (EX X: Rep_hypreal x. \ -\ ALL u. {n. u < abs (X n)}: FreeUltrafilterNat)"; -by (blast_tac (claset() addSIs [HInfinite_FreeUltrafilterNat, - FreeUltrafilterNat_HInfinite]) 1); -qed "HInfinite_FreeUltrafilterNat_iff"; - -(*-------------------------------------------------------------------- - Alternative definitions for Infinitesimal using Free ultrafilter - --------------------------------------------------------------------*) - -Goal "{n. - u < Yd n} Int {n. xa n = Yd n} <= {n. -u < xa n}"; -by Auto_tac; -qed "lemma_free4"; - -Goal "{n. Yb n < u} Int {n. xa n = Yb n} <= {n. xa n < u}"; -by Auto_tac; -qed "lemma_free5"; - -Goalw [Infinitesimal_def] - "x : Infinitesimal ==> EX X: Rep_hypreal x. \ -\ ALL u. 0 < u --> {n. abs (X n) < u}: FreeUltrafilterNat"; -by (auto_tac (claset(), simpset() addsimps [abs_less_iff, inst "a" "x" minus_less_iff])); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (EVERY[Auto_tac, rtac bexI 1, rtac lemma_hyprel_refl 2, Step_tac 1]); -by (dtac (hypreal_of_real_less_iff RS iffD2) 1); -by (dres_inst_tac [("x","hypreal_of_real u")] bspec 1); -by Auto_tac; -by (auto_tac (claset(), - simpset() addsimps [hypreal_less_def, hypreal_minus, - hypreal_of_real_def])); -by (Ultra_tac 1); -qed "Infinitesimal_FreeUltrafilterNat"; - -Goalw [Infinitesimal_def] - "EX X: Rep_hypreal x. \ -\ ALL u. 0 < u --> {n. abs (X n) < u} : FreeUltrafilterNat \ -\ ==> x : Infinitesimal"; -by (auto_tac (claset(), - simpset() addsimps [abs_less_iff,abs_interval_iff, inst "a" "x" minus_less_iff])); -by (auto_tac (claset(), - simpset() addsimps [SReal_iff])); -by (auto_tac (claset() addSIs [exI] - addIs [FreeUltrafilterNat_subset], - simpset() addsimps [hypreal_less_def, hypreal_minus,hypreal_of_real_def])); -qed "FreeUltrafilterNat_Infinitesimal"; - -Goal "(x : Infinitesimal) = (EX X: Rep_hypreal x. \ -\ ALL u. 0 < u --> {n. abs (X n) < u}: FreeUltrafilterNat)"; -by (blast_tac (claset() addSIs [Infinitesimal_FreeUltrafilterNat, - FreeUltrafilterNat_Infinitesimal]) 1); -qed "Infinitesimal_FreeUltrafilterNat_iff"; - -(*------------------------------------------------------------------------ - Infinitesimals as smaller than 1/n for all n::nat (> 0) - ------------------------------------------------------------------------*) - -Goal "(ALL r. 0 < r --> x < r) = (ALL n. x < inverse(real (Suc n)))"; -by (auto_tac (claset(), simpset() addsimps [real_of_nat_Suc_gt_zero])); -by (blast_tac (claset() addSDs [reals_Archimedean] - addIs [order_less_trans]) 1); -qed "lemma_Infinitesimal"; - -Goal "(ALL r: Reals. 0 < r --> x < r) = \ -\ (ALL n. x < inverse(hypreal_of_nat (Suc n)))"; -by (Step_tac 1); -by (dres_inst_tac [("x","inverse (hypreal_of_real(real (Suc n)))")] - bspec 1); -by (full_simp_tac (simpset() addsimps [SReal_inverse]) 1); -by (rtac (real_of_nat_Suc_gt_zero RS positive_imp_inverse_positive RS - (hypreal_of_real_less_iff RS iffD2) RSN(2,impE)) 1); -by (assume_tac 2); -by (asm_full_simp_tac (simpset() addsimps - [real_of_nat_Suc_gt_zero, hypreal_of_nat_def]) 1); -by (auto_tac (claset() addSDs [reals_Archimedean], - simpset() addsimps [SReal_iff])); -by (dtac (hypreal_of_real_less_iff RS iffD2) 1); -by (asm_full_simp_tac (simpset() addsimps - [real_of_nat_Suc_gt_zero, hypreal_of_nat_def])1); -by (blast_tac (claset() addIs [order_less_trans]) 1); -qed "lemma_Infinitesimal2"; - -Goalw [Infinitesimal_def] - "Infinitesimal = {x. ALL n. abs x < inverse (hypreal_of_nat (Suc n))}"; -by (auto_tac (claset(), simpset() addsimps [lemma_Infinitesimal2])); -qed "Infinitesimal_hypreal_of_nat_iff"; - - -(*------------------------------------------------------------------------- - Proof that omega is an infinite number and - hence that epsilon is an infinitesimal number. - -------------------------------------------------------------------------*) -Goal "{n. n < Suc m} = {n. n < m} Un {n. n = m}"; -by (auto_tac (claset(), simpset() addsimps [less_Suc_eq])); -qed "Suc_Un_eq"; - -(*------------------------------------------- - Prove that any segment is finite and - hence cannot belong to FreeUltrafilterNat - -------------------------------------------*) -Goal "finite {n::nat. n < m}"; -by (induct_tac "m" 1); -by (auto_tac (claset(), simpset() addsimps [Suc_Un_eq])); -qed "finite_nat_segment"; - -Goal "finite {n::nat. real n < real (m::nat)}"; -by (auto_tac (claset() addIs [finite_nat_segment], simpset())); -qed "finite_real_of_nat_segment"; - -Goal "finite {n::nat. real n < u}"; -by (cut_inst_tac [("x","u")] reals_Archimedean2 1); -by (Step_tac 1); -by (rtac (finite_real_of_nat_segment RSN (2,finite_subset)) 1); -by (auto_tac (claset() addDs [order_less_trans], simpset())); -qed "finite_real_of_nat_less_real"; - -Goal "{n. f n <= u} = {n. f n < u} Un {n. u = (f n :: real)}"; -by (auto_tac (claset() addDs [order_le_imp_less_or_eq], - simpset() addsimps [order_less_imp_le])); -qed "lemma_real_le_Un_eq"; - -Goal "finite {n::nat. real n <= u}"; -by (auto_tac (claset(), - simpset() addsimps [lemma_real_le_Un_eq,lemma_finite_omega_set, - finite_real_of_nat_less_real])); -qed "finite_real_of_nat_le_real"; - -Goal "finite {n::nat. abs(real n) <= u}"; -by (simp_tac (simpset() addsimps [real_of_nat_Suc_gt_zero, - finite_real_of_nat_le_real]) 1); -qed "finite_rabs_real_of_nat_le_real"; - -Goal "{n. abs(real n) <= u} ~: FreeUltrafilterNat"; -by (blast_tac (claset() addSIs [FreeUltrafilterNat_finite, - finite_rabs_real_of_nat_le_real]) 1); -qed "rabs_real_of_nat_le_real_FreeUltrafilterNat"; - -Goal "{n. u < real n} : FreeUltrafilterNat"; -by (rtac ccontr 1 THEN dtac FreeUltrafilterNat_Compl_mem 1); -by (subgoal_tac "- {n::nat. u < real n} = {n. real n <= u}" 1); -by (Force_tac 2); -by (asm_full_simp_tac (simpset() addsimps - [finite_real_of_nat_le_real RS FreeUltrafilterNat_finite]) 1); -qed "FreeUltrafilterNat_nat_gt_real"; - -(*-------------------------------------------------------------- - The complement of {n. abs(real n) <= u} = - {n. u < abs (real n)} is in FreeUltrafilterNat - by property of (free) ultrafilters - --------------------------------------------------------------*) -Goal "- {n::nat. real n <= u} = {n. u < real n}"; -by (auto_tac (claset() addSDs [order_le_less_trans], - simpset() addsimps [linorder_not_le])); -val lemma = result(); - -(*----------------------------------------------- - Omega is a member of HInfinite - -----------------------------------------------*) - -Goal "hyprel``{%n::nat. real (Suc n)} : hypreal"; -by Auto_tac; -qed "hypreal_omega"; - -Goal "{n. u < real n} : FreeUltrafilterNat"; -by (cut_inst_tac [("u","u")] rabs_real_of_nat_le_real_FreeUltrafilterNat 1); -by (auto_tac (claset() addDs [FreeUltrafilterNat_Compl_mem], - simpset() addsimps [lemma])); -qed "FreeUltrafilterNat_omega"; - -Goalw [omega_def] "omega: HInfinite"; -by (auto_tac (claset() addSIs [FreeUltrafilterNat_HInfinite], - simpset())); -by (rtac bexI 1); -by (rtac lemma_hyprel_refl 2); -by Auto_tac; -by (simp_tac (simpset() addsimps [real_of_nat_Suc, diff_less_eq RS sym, - FreeUltrafilterNat_omega]) 1); -qed "HInfinite_omega"; -Addsimps [HInfinite_omega]; - -(*----------------------------------------------- - Epsilon is a member of Infinitesimal - -----------------------------------------------*) - -Goal "epsilon : Infinitesimal"; -by (auto_tac (claset() addSIs [HInfinite_inverse_Infinitesimal,HInfinite_omega], - simpset() addsimps [hypreal_epsilon_inverse_omega])); -qed "Infinitesimal_epsilon"; -Addsimps [Infinitesimal_epsilon]; - -Goal "epsilon : HFinite"; -by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD], - simpset())); -qed "HFinite_epsilon"; -Addsimps [HFinite_epsilon]; - -Goal "epsilon @= 0"; -by (simp_tac (simpset() addsimps [mem_infmal_iff RS sym]) 1); -qed "epsilon_approx_zero"; -Addsimps [epsilon_approx_zero]; - -(*------------------------------------------------------------------------ - Needed for proof that we define a hyperreal [ LIM. - -----------------------------------------------------------------------*) - -Goal "0 < u ==> \ -\ (u < inverse (real(Suc n))) = (real(Suc n) < inverse u)"; -by (asm_full_simp_tac (simpset() addsimps [inverse_eq_divide]) 1); -by (stac pos_less_divide_eq 1); -by (assume_tac 1); -by (stac pos_less_divide_eq 1); -by (simp_tac (simpset() addsimps [real_mult_commute]) 2); -by (simp_tac (simpset() addsimps [real_of_nat_Suc_gt_zero]) 1); -qed "real_of_nat_less_inverse_iff"; - -Goal "0 < u ==> finite {n. u < inverse(real(Suc n))}"; -by (asm_simp_tac (simpset() addsimps [real_of_nat_less_inverse_iff]) 1); -by (asm_simp_tac (simpset() addsimps [real_of_nat_Suc, - less_diff_eq RS sym]) 1); -by (rtac finite_real_of_nat_less_real 1); -qed "finite_inverse_real_of_posnat_gt_real"; - -Goal "{n. u <= inverse(real(Suc n))} = \ -\ {n. u < inverse(real(Suc n))} Un {n. u = inverse(real(Suc n))}"; -by (auto_tac (claset() addDs [order_le_imp_less_or_eq], - simpset() addsimps [order_less_imp_le])); -qed "lemma_real_le_Un_eq2"; - -Goal "(inverse (real(Suc n)) <= r) = (1 <= r * real(Suc n))"; -by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); -by (simp_tac (simpset() addsimps [inverse_eq_divide]) 1); -by (stac pos_less_divide_eq 1); -by (simp_tac (simpset() addsimps [real_of_nat_Suc_gt_zero]) 1); -by (simp_tac (simpset() addsimps [real_mult_commute]) 1); -qed "real_of_nat_inverse_le_iff"; - -Goal "(u = inverse (real(Suc n))) = (real(Suc n) = inverse u)"; -by (auto_tac (claset(), - simpset() addsimps [inverse_inverse_eq, real_of_nat_Suc_gt_zero, - real_not_refl2 RS not_sym])); -qed "real_of_nat_inverse_eq_iff"; - -Goal "finite {n::nat. u = inverse(real(Suc n))}"; -by (asm_simp_tac (simpset() addsimps [real_of_nat_inverse_eq_iff]) 1); -by (cut_inst_tac [("x","inverse u - 1")] lemma_finite_omega_set 1); -by (asm_full_simp_tac (simpset() addsimps [real_of_nat_Suc, - diff_eq_eq RS sym, eq_commute]) 1); -qed "lemma_finite_omega_set2"; - -Goal "0 < u ==> finite {n. u <= inverse(real(Suc n))}"; -by (auto_tac (claset(), - simpset() addsimps [lemma_real_le_Un_eq2,lemma_finite_omega_set2, - finite_inverse_real_of_posnat_gt_real])); -qed "finite_inverse_real_of_posnat_ge_real"; - -Goal "0 < u ==> \ -\ {n. u <= inverse(real(Suc n))} ~: FreeUltrafilterNat"; -by (blast_tac (claset() addSIs [FreeUltrafilterNat_finite, - finite_inverse_real_of_posnat_ge_real]) 1); -qed "inverse_real_of_posnat_ge_real_FreeUltrafilterNat"; - -(*-------------------------------------------------------------- - The complement of {n. u <= inverse(real(Suc n))} = - {n. inverse(real(Suc n)) < u} is in FreeUltrafilterNat - by property of (free) ultrafilters - --------------------------------------------------------------*) -Goal "- {n. u <= inverse(real(Suc n))} = \ -\ {n. inverse(real(Suc n)) < u}"; -by (auto_tac (claset() addSDs [order_le_less_trans], - simpset() addsimps [linorder_not_le])); -val lemma = result(); - -Goal "0 < u ==> \ -\ {n. inverse(real(Suc n)) < u} : FreeUltrafilterNat"; -by (cut_inst_tac [("u","u")] inverse_real_of_posnat_ge_real_FreeUltrafilterNat 1); -by (auto_tac (claset() addDs [FreeUltrafilterNat_Compl_mem], - simpset() addsimps [lemma])); -qed "FreeUltrafilterNat_inverse_real_of_posnat"; - -(*-------------------------------------------------------------- - 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 (the theorem was previously in REALTOPOS.thy). - -------------------------------------------------------------*) -(*----------------------------------------------------- - |X(n) - x| < 1/n ==> [] - hypreal_of_real x|: Infinitesimal - -----------------------------------------------------*) -Goal "ALL n. abs(X n + -x) < inverse(real(Suc n)) \ -\ ==> Abs_hypreal(hyprel``{X}) + -hypreal_of_real x : Infinitesimal"; -by (auto_tac (claset() addSIs [bexI] - addDs [FreeUltrafilterNat_inverse_real_of_posnat, - FreeUltrafilterNat_all,FreeUltrafilterNat_Int] - addIs [order_less_trans, FreeUltrafilterNat_subset], - simpset() addsimps [hypreal_minus, - hypreal_of_real_def,hypreal_add, - Infinitesimal_FreeUltrafilterNat_iff,hypreal_inverse])); -qed "real_seq_to_hypreal_Infinitesimal"; - -Goal "ALL n. abs(X n + -x) < inverse(real(Suc n)) \ -\ ==> Abs_hypreal(hyprel``{X}) @= hypreal_of_real x"; -by (stac approx_minus_iff 1); -by (rtac (mem_infmal_iff RS subst) 1); -by (etac real_seq_to_hypreal_Infinitesimal 1); -qed "real_seq_to_hypreal_approx"; - -Goal "ALL n. abs(x + -X n) < inverse(real(Suc n)) \ -\ ==> Abs_hypreal(hyprel``{X}) @= hypreal_of_real x"; -by (asm_full_simp_tac (simpset() addsimps [abs_minus_add_cancel, - real_seq_to_hypreal_approx]) 1); -qed "real_seq_to_hypreal_approx2"; - -Goal "ALL n. abs(X n + -Y n) < inverse(real(Suc n)) \ -\ ==> Abs_hypreal(hyprel``{X}) + \ -\ -Abs_hypreal(hyprel``{Y}) : Infinitesimal"; -by (auto_tac (claset() addSIs [bexI] - addDs [FreeUltrafilterNat_inverse_real_of_posnat, - FreeUltrafilterNat_all,FreeUltrafilterNat_Int] - addIs [order_less_trans, FreeUltrafilterNat_subset], - simpset() addsimps - [Infinitesimal_FreeUltrafilterNat_iff,hypreal_minus,hypreal_add, - hypreal_inverse])); -qed "real_seq_to_hypreal_Infinitesimal2"; - - - - - -(*MOVE UP*) -Goal "[| x + y : HInfinite; y: HFinite |] ==> x : HInfinite"; -by (rtac ccontr 1); -by (dtac (HFinite_HInfinite_iff RS iffD2) 1); -by (auto_tac (claset() addDs [HFinite_add],simpset() - addsimps [HInfinite_HFinite_iff])); -qed "HInfinite_HFinite_add_cancel"; - -Goal "[| x : HInfinite; y : HFinite |] ==> x + y : HInfinite"; -by (res_inst_tac [("y","-y")] HInfinite_HFinite_add_cancel 1); -by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc, - HFinite_minus_iff])); -qed "HInfinite_HFinite_add"; - -Goal "[| x : HInfinite; x <= y; 0 <= x |] ==> y : HInfinite"; -by (auto_tac (claset() addIs [HFinite_bounded],simpset() - addsimps [HInfinite_HFinite_iff])); -qed "HInfinite_ge_HInfinite"; - -Goal "[| x : Infinitesimal; x ~= 0 |] ==> inverse x : HInfinite"; -by (rtac ccontr 1 THEN dtac (HFinite_HInfinite_iff RS iffD2) 1); -by (auto_tac (claset() addDs [Infinitesimal_HFinite_mult2],simpset())); -qed "Infinitesimal_inverse_HInfinite"; - -Goal "[| x : HInfinite; y : HFinite - Infinitesimal |] \ -\ ==> x * y : HInfinite"; -by (rtac ccontr 1 THEN dtac (HFinite_HInfinite_iff RS iffD2) 1); -by (ftac HFinite_Infinitesimal_not_zero 1); -by (dtac HFinite_not_Infinitesimal_inverse 1); -by (Step_tac 1 THEN dtac HFinite_mult 1); -by (auto_tac (claset(),simpset() addsimps [hypreal_mult_assoc, - HFinite_HInfinite_iff])); -qed "HInfinite_HFinite_not_Infinitesimal_mult"; - -Goal "[| x : HInfinite; y : HFinite - Infinitesimal |] \ -\ ==> y * x : HInfinite"; -by (auto_tac (claset(),simpset() addsimps [hypreal_mult_commute, - HInfinite_HFinite_not_Infinitesimal_mult])); -qed "HInfinite_HFinite_not_Infinitesimal_mult2"; - -Goal "[| x : HInfinite; 0 < x; y : Reals |] ==> y < x"; -by (auto_tac (claset() addSDs [bspec],simpset() addsimps - [HInfinite_def,hrabs_def,order_less_imp_le])); -qed "HInfinite_gt_SReal"; - -Goal "[| x : HInfinite; 0 < x |] ==> 1 < x"; -by (auto_tac (claset() addIs [HInfinite_gt_SReal],simpset())); -qed "HInfinite_gt_zero_gt_one"; - - -Goal "1 ~: HInfinite"; -by (simp_tac (simpset() addsimps [HInfinite_HFinite_iff]) 1); -qed "not_HInfinite_one"; -Addsimps [not_HInfinite_one]; diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/NSA.thy --- a/src/HOL/Hyperreal/NSA.thy Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/NSA.thy Thu Jan 29 16:51:17 2004 +0100 @@ -1,46 +1,2346 @@ (* Title : NSA.thy Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge - Description : Infinite numbers, Infinitesimals, - infinitely close relation etc. -*) +*) -NSA = HRealAbs + RComplete + +header{*Infinite Numbers, Infinitesimals, Infinitely Close Relation*} + +theory NSA = HRealAbs + RComplete: constdefs Infinitesimal :: "hypreal set" - "Infinitesimal == {x. ALL r: Reals. 0 < r --> abs x < r}" + "Infinitesimal == {x. \r \ Reals. 0 < r --> abs x < r}" HFinite :: "hypreal set" - "HFinite == {x. EX r: Reals. abs x < r}" + "HFinite == {x. \r \ Reals. abs x < r}" HInfinite :: "hypreal set" - "HInfinite == {x. ALL r: Reals. r < abs x}" + "HInfinite == {x. \r \ Reals. r < abs x}" (* standard part map *) - st :: hypreal => hypreal - "st == (%x. @r. x : HFinite & r:Reals & r @= x)" + st :: "hypreal => hypreal" + "st == (%x. @r. x \ HFinite & r \ Reals & r @= x)" - monad :: hypreal => hypreal set + monad :: "hypreal => hypreal set" "monad x == {y. x @= y}" - galaxy :: hypreal => hypreal set - "galaxy x == {y. (x + -y) : HFinite}" - + galaxy :: "hypreal => hypreal set" + "galaxy x == {y. (x + -y) \ HFinite}" + (* infinitely close *) approx :: "[hypreal, hypreal] => bool" (infixl "@=" 50) - "x @= y == (x + -y) : Infinitesimal" + "x @= y == (x + -y) \ Infinitesimal" + + +defs + + (*standard real numbers as a subset of the hyperreals*) + SReal_def: "Reals == {x. \r. x = hypreal_of_real r}" + +syntax (xsymbols) + approx :: "[hypreal, hypreal] => bool" (infixl "\" 50) + + + +(*-------------------------------------------------------------------- + Closure laws for members of (embedded) set standard real Reals + --------------------------------------------------------------------*) + +lemma SReal_add: "[| (x::hypreal) \ Reals; y \ Reals |] ==> x + y \ Reals" +apply (auto simp add: SReal_def) +apply (rule_tac x = "r + ra" in exI, simp) +done + +lemma SReal_mult: "[| (x::hypreal) \ Reals; y \ Reals |] ==> x * y \ Reals" +apply (simp add: SReal_def, safe) +apply (rule_tac x = "r * ra" in exI) +apply (simp (no_asm) add: hypreal_of_real_mult) +done + +lemma SReal_inverse: "(x::hypreal) \ Reals ==> inverse x \ Reals" +apply (simp add: SReal_def) +apply (blast intro: hypreal_of_real_inverse [symmetric]) +done + +lemma SReal_divide: "[| (x::hypreal) \ Reals; y \ Reals |] ==> x/y \ Reals" +apply (simp (no_asm_simp) add: SReal_mult SReal_inverse hypreal_divide_def) +done + +lemma SReal_minus: "(x::hypreal) \ Reals ==> -x \ Reals" +apply (simp add: SReal_def) +apply (blast intro: hypreal_of_real_minus [symmetric]) +done + +lemma SReal_minus_iff: "(-x \ Reals) = ((x::hypreal) \ Reals)" +apply auto +apply (erule_tac [2] SReal_minus) +apply (drule SReal_minus, auto) +done +declare SReal_minus_iff [simp] + +lemma SReal_add_cancel: "[| (x::hypreal) + y \ Reals; y \ Reals |] ==> x \ Reals" +apply (drule_tac x = y in SReal_minus) +apply (drule SReal_add, assumption, auto) +done + +lemma SReal_hrabs: "(x::hypreal) \ Reals ==> abs x \ Reals" +apply (simp add: SReal_def) +apply (auto simp add: hypreal_of_real_hrabs) +done + +lemma SReal_hypreal_of_real: "hypreal_of_real x \ Reals" +by (simp add: SReal_def) +declare SReal_hypreal_of_real [simp] + +lemma SReal_number_of: "(number_of w ::hypreal) \ Reals" +apply (unfold hypreal_number_of_def) +apply (rule SReal_hypreal_of_real) +done +declare SReal_number_of [simp] + +(** As always with numerals, 0 and 1 are special cases **) + +lemma Reals_0: "(0::hypreal) \ Reals" +apply (subst hypreal_numeral_0_eq_0 [symmetric]) +apply (rule SReal_number_of) +done +declare Reals_0 [simp] + +lemma Reals_1: "(1::hypreal) \ Reals" +apply (subst hypreal_numeral_1_eq_1 [symmetric]) +apply (rule SReal_number_of) +done +declare Reals_1 [simp] + +lemma SReal_divide_number_of: "r \ Reals ==> r/(number_of w::hypreal) \ Reals" +apply (unfold hypreal_divide_def) +apply (blast intro!: SReal_number_of SReal_mult SReal_inverse) +done + +(* Infinitesimal epsilon not in Reals *) + +lemma SReal_epsilon_not_mem: "epsilon \ Reals" +apply (simp add: SReal_def) +apply (auto simp add: hypreal_of_real_not_eq_epsilon [THEN not_sym]) +done + +lemma SReal_omega_not_mem: "omega \ Reals" +apply (simp add: SReal_def) +apply (auto simp add: hypreal_of_real_not_eq_omega [THEN not_sym]) +done + +lemma SReal_UNIV_real: "{x. hypreal_of_real x \ Reals} = (UNIV::real set)" +by (simp add: SReal_def) + +lemma SReal_iff: "(x \ Reals) = (\y. x = hypreal_of_real y)" +by (simp add: SReal_def) + +lemma hypreal_of_real_image: "hypreal_of_real `(UNIV::real set) = Reals" +by (auto simp add: SReal_def) + +lemma inv_hypreal_of_real_image: "inv hypreal_of_real ` Reals = UNIV" +apply (auto simp add: SReal_def) +apply (rule inj_hypreal_of_real [THEN inv_f_f, THEN subst], blast) +done + +lemma SReal_hypreal_of_real_image: + "[| \x. x: P; P <= Reals |] ==> \Q. P = hypreal_of_real ` Q" +apply (simp add: SReal_def, blast) +done + +lemma SReal_dense: "[| (x::hypreal) \ Reals; y \ Reals; x \r \ Reals. x ((\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 <= Reals; \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 + +(*------------------------------------------------------ + lifting of ub and property of lub + -------------------------------------------------------*) +lemma hypreal_of_real_isUb_iff: + "(isUb (Reals) (hypreal_of_real ` Q) (hypreal_of_real Y)) = + (isUb (UNIV :: real set) Q Y)" +apply (simp add: isUb_def setle_def) +done + +lemma hypreal_of_real_isLub1: + "isLub Reals (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_isLub2: + "isLub (UNIV :: real set) Q Y + ==> isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y)" +apply (simp add: isLub_def leastP_def) +apply (auto simp add: hypreal_of_real_isUb_iff setge_def) +apply (frule_tac x2 = x in isUbD2a [THEN SReal_iff [THEN iffD1], THEN exE]) + prefer 2 apply assumption +apply (drule_tac x = xa in spec) +apply (auto simp add: hypreal_of_real_isUb_iff) +done + +lemma hypreal_of_real_isLub_iff: "(isLub Reals (hypreal_of_real ` Q) (hypreal_of_real Y)) = + (isLub (UNIV :: real set) Q Y)" +apply (blast intro: hypreal_of_real_isLub1 hypreal_of_real_isLub2) +done + +(* lemmas *) +lemma lemma_isUb_hypreal_of_real: + "isUb Reals P Y ==> \Yo. isUb Reals P (hypreal_of_real Yo)" +by (auto simp add: SReal_iff isUb_def) + +lemma lemma_isLub_hypreal_of_real: + "isLub Reals P Y ==> \Yo. isLub Reals 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 Reals P (hypreal_of_real Yo) ==> \Y. isLub Reals P Y" +by (auto simp add: isLub_def leastP_def isUb_def) + +lemma SReal_complete: "[| P <= Reals; \x. x \ P; \Y. isUb Reals P Y |] + ==> \t::hypreal. isLub Reals 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 + +(*-------------------------------------------------------------------- + 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!: SReal_add hrabs_add_less) +done + +lemma HFinite_mult: "[|x \ HFinite; y \ HFinite|] ==> x*y \ HFinite" +apply (simp add: HFinite_def) +apply (blast intro!: SReal_mult abs_mult_less) +done + +lemma HFinite_minus_iff: "(-x \ HFinite) = (x \ HFinite)" +by (simp add: HFinite_def) + +lemma SReal_subset_HFinite: "Reals <= HFinite" +apply (auto simp add: SReal_def HFinite_def) +apply (rule_tac x = "1 + abs (hypreal_of_real r) " in exI) +apply (auto simp add: hypreal_of_real_hrabs) +apply (rule_tac x = "1 + abs r" in exI, simp) +done + +lemma HFinite_hypreal_of_real [simp]: "hypreal_of_real x \ HFinite" +by (auto intro: SReal_subset_HFinite [THEN subsetD]) + +lemma HFiniteD: "x \ HFinite ==> \t \ Reals. abs x < t" +by (simp add: HFinite_def) + +lemma HFinite_hrabs_iff: "(abs x \ HFinite) = (x \ HFinite)" +by (simp add: HFinite_def) +declare HFinite_hrabs_iff [iff] + +lemma HFinite_number_of: "number_of w \ HFinite" +by (rule SReal_number_of [THEN SReal_subset_HFinite [THEN subsetD]]) +declare HFinite_number_of [simp] + +(** As always with numerals, 0 and 1 are special cases **) + +lemma HFinite_0: "0 \ HFinite" +apply (subst hypreal_numeral_0_eq_0 [symmetric]) +apply (rule HFinite_number_of) +done +declare HFinite_0 [simp] + +lemma HFinite_1: "1 \ HFinite" +apply (subst hypreal_numeral_1_eq_1 [symmetric]) +apply (rule HFinite_number_of) +done +declare HFinite_1 [simp] + +lemma HFinite_bounded: "[|x \ HFinite; y <= x; 0 <= y |] ==> y \ HFinite" +apply (case_tac "x <= 0") +apply (drule_tac y = x in order_trans) +apply (drule_tac [2] hypreal_le_anti_sym) +apply (auto simp add: linorder_not_le) +apply (auto intro: order_le_less_trans simp add: abs_if HFinite_def) +done + +(*------------------------------------------------------------------ + Set of infinitesimals is a subring of the hyperreals + ------------------------------------------------------------------*) +lemma InfinitesimalD: + "x \ Infinitesimal ==> \r \ Reals. 0 < r --> abs x < r" +apply (simp add: Infinitesimal_def) +done + +lemma Infinitesimal_zero: "0 \ Infinitesimal" +by (simp add: Infinitesimal_def) +declare Infinitesimal_zero [iff] + +lemma hypreal_sum_of_halves: "x/(2::hypreal) + x/(2::hypreal) = x" +by auto + +lemma hypreal_half_gt_zero: "0 < r ==> 0 < r/(2::hypreal)" +by auto + +lemma Infinitesimal_add: + "[| x \ Infinitesimal; y \ Infinitesimal |] ==> (x+y) \ Infinitesimal" +apply (auto simp add: Infinitesimal_def) +apply (rule hypreal_sum_of_halves [THEN subst]) +apply (drule hypreal_half_gt_zero) +apply (blast intro: hrabs_add_less hrabs_add_less SReal_divide_number_of) +done + +lemma Infinitesimal_minus_iff: "(-x:Infinitesimal) = (x:Infinitesimal)" +by (simp add: Infinitesimal_def) +declare Infinitesimal_minus_iff [simp] + +lemma Infinitesimal_diff: "[| x \ Infinitesimal; y \ Infinitesimal |] ==> x-y \ Infinitesimal" +by (simp add: hypreal_diff_def Infinitesimal_add) + +lemma Infinitesimal_mult: + "[| x \ Infinitesimal; y \ Infinitesimal |] ==> (x * y) \ Infinitesimal" +apply (auto simp add: Infinitesimal_def) +apply (case_tac "y=0") +apply (cut_tac [2] a = "abs x" and b = 1 and c = "abs y" and d = r in mult_strict_mono, auto) +done + +lemma Infinitesimal_HFinite_mult: "[| x \ Infinitesimal; y \ HFinite |] ==> (x * y) \ Infinitesimal" +apply (auto dest!: HFiniteD simp add: Infinitesimal_def) +apply (frule hrabs_less_gt_zero) +apply (drule_tac x = "r/t" in bspec) +apply (blast intro: SReal_divide) +apply (simp add: zero_less_divide_iff) +apply (case_tac "x=0 | y=0") +apply (cut_tac [2] a = "abs x" and b = "r/t" and c = "abs y" in mult_strict_mono) +apply (auto simp add: zero_less_divide_iff) +done + +lemma Infinitesimal_HFinite_mult2: "[| x \ Infinitesimal; y \ HFinite |] ==> (y * x) \ Infinitesimal" +by (auto dest: Infinitesimal_HFinite_mult simp add: hypreal_mult_commute) + +(*** rather long proof ***) +lemma HInfinite_inverse_Infinitesimal: + "x \ HInfinite ==> inverse x: Infinitesimal" +apply (auto simp add: HInfinite_def Infinitesimal_def) +apply (erule_tac x = "inverse r" in ballE) +apply (frule_tac a1 = r and z = "abs x" in positive_imp_inverse_positive [THEN order_less_trans], assumption) +apply (drule inverse_inverse_eq [symmetric, THEN subst]) +apply (rule inverse_less_iff_less [THEN iffD1]) +apply (auto simp add: SReal_inverse) +done + + + +lemma HInfinite_mult: "[|x \ HInfinite;y \ HInfinite|] ==> (x*y) \ HInfinite" +apply (simp add: HInfinite_def, auto) +apply (erule_tac x = 1 in ballE) +apply (erule_tac x = r in ballE) +apply (case_tac "y=0") +apply (cut_tac [2] c = 1 and d = "abs x" and a = r and b = "abs y" in mult_strict_mono) +apply (auto simp add: mult_ac) +done + +lemma HInfinite_add_ge_zero: + "[|x \ HInfinite; 0 <= y; 0 <= x|] ==> (x + y): HInfinite" +by (auto intro!: hypreal_add_zero_less_le_mono + simp add: abs_if hypreal_add_commute hypreal_le_add_order HInfinite_def) + +lemma HInfinite_add_ge_zero2: "[|x \ HInfinite; 0 <= y; 0 <= x|] ==> (y + x): HInfinite" +by (auto intro!: HInfinite_add_ge_zero simp add: hypreal_add_commute) + +lemma HInfinite_add_gt_zero: "[|x \ HInfinite; 0 < y; 0 < x|] ==> (x + y): HInfinite" +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_add_le_zero: "[|x \ HInfinite; y <= 0; x <= 0|] ==> (x + y): HInfinite" +apply (drule HInfinite_minus_iff [THEN iffD2]) +apply (rule HInfinite_minus_iff [THEN iffD1]) +apply (auto intro: HInfinite_add_ge_zero) +done + +lemma HInfinite_add_lt_zero: "[|x \ HInfinite; y < 0; x < 0|] ==> (x + y): HInfinite" +by (blast intro: HInfinite_add_le_zero order_less_imp_le) + +lemma HFinite_sum_squares: "[|a: HFinite; b: HFinite; c: HFinite|] + ==> a*a + b*b + c*c \ HFinite" +apply (auto intro: HFinite_mult HFinite_add) +done + +lemma not_Infinitesimal_not_zero: "x \ Infinitesimal ==> x \ 0" +by auto + +lemma not_Infinitesimal_not_zero2: "x \ HFinite - Infinitesimal ==> x \ 0" +by auto + +lemma Infinitesimal_hrabs_iff: "(abs x \ Infinitesimal) = (x \ Infinitesimal)" +by (auto simp add: hrabs_def) +declare Infinitesimal_hrabs_iff [iff] + +lemma HFinite_diff_Infinitesimal_hrabs: "x \ HFinite - Infinitesimal ==> abs x \ HFinite - Infinitesimal" +by blast + +lemma hrabs_less_Infinitesimal: + "[| e \ Infinitesimal; abs x < e |] ==> x \ Infinitesimal" +apply (auto simp add: Infinitesimal_def abs_less_iff) +done + +lemma hrabs_le_Infinitesimal: "[| e \ Infinitesimal; abs x <= e |] ==> x \ Infinitesimal" +by (blast dest: order_le_imp_less_or_eq intro: hrabs_less_Infinitesimal) + +lemma Infinitesimal_interval: + "[| e \ Infinitesimal; e' \ Infinitesimal; e' < x ; x < e |] + ==> x \ Infinitesimal" +apply (auto simp add: Infinitesimal_def abs_less_iff) +done + +lemma Infinitesimal_interval2: "[| e \ Infinitesimal; e' \ Infinitesimal; + e' <= x ; x <= e |] ==> x \ Infinitesimal" +apply (auto intro: Infinitesimal_interval simp add: order_le_less) +done + +lemma not_Infinitesimal_mult: + "[| x \ Infinitesimal; y \ Infinitesimal|] ==> (x*y) \Infinitesimal" +apply (unfold Infinitesimal_def, clarify) +apply (simp add: linorder_not_less) +apply (erule_tac x = "r*ra" in ballE) +prefer 2 apply (fast intro: SReal_mult) +apply (auto simp add: zero_less_mult_iff) +apply (cut_tac c = ra and d = "abs y" and a = r and b = "abs x" in mult_mono, auto) +done + +lemma Infinitesimal_mult_disj: "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 HFinite_Infinitesimal_not_zero: "x \ HFinite-Infinitesimal ==> x \ 0" +by blast + +lemma HFinite_Infinitesimal_diff_mult: "[| x \ HFinite - Infinitesimal; + y \ HFinite - Infinitesimal + |] ==> (x*y) \ HFinite - Infinitesimal" +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_hypreal_of_real_mult: "x \ Infinitesimal ==> x * hypreal_of_real r \ Infinitesimal" +by (erule HFinite_hypreal_of_real [THEN [2] Infinitesimal_HFinite_mult]) + +lemma Infinitesimal_hypreal_of_real_mult2: "x \ Infinitesimal ==> hypreal_of_real r * x \ Infinitesimal" +by (erule HFinite_hypreal_of_real [THEN [2] Infinitesimal_HFinite_mult2]) + +(*---------------------------------------------------------------------- + Infinitely close relation @= + ----------------------------------------------------------------------*) + +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_iff2: " (x @= y) = (-y + x @= 0)" +by (simp add: approx_def hypreal_add_commute) + +lemma approx_refl: "x @= x" +by (simp add: approx_def Infinitesimal_def) +declare approx_refl [iff] + +lemma approx_sym: "x @= y ==> y @= x" +apply (simp add: approx_def) +apply (rule hypreal_minus_distrib1 [THEN subst]) +apply (erule Infinitesimal_minus_iff [THEN iffD2]) +done + +lemma approx_trans: "[| x @= y; y @= z |] ==> x @= z" +apply (simp add: approx_def) +apply (drule Infinitesimal_add, assumption, auto) +done + +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 number_of_approx_reorient: "(number_of w @= x) = (x @= number_of w)" +by (blast intro: approx_sym) + +lemma zero_approx_reorient: "(0 @= x) = (x @= 0)" +by (blast intro: approx_sym) + +lemma one_approx_reorient: "(1 @= x) = (x @= 1)" +by (blast intro: approx_sym) -defs +ML +{* +val SReal_add = thm "SReal_add"; +val SReal_mult = thm "SReal_mult"; +val SReal_inverse = thm "SReal_inverse"; +val SReal_divide = thm "SReal_divide"; +val SReal_minus = thm "SReal_minus"; +val SReal_minus_iff = thm "SReal_minus_iff"; +val SReal_add_cancel = thm "SReal_add_cancel"; +val SReal_hrabs = thm "SReal_hrabs"; +val SReal_hypreal_of_real = thm "SReal_hypreal_of_real"; +val SReal_number_of = thm "SReal_number_of"; +val Reals_0 = thm "Reals_0"; +val Reals_1 = thm "Reals_1"; +val SReal_divide_number_of = thm "SReal_divide_number_of"; +val SReal_epsilon_not_mem = thm "SReal_epsilon_not_mem"; +val SReal_omega_not_mem = thm "SReal_omega_not_mem"; +val SReal_UNIV_real = thm "SReal_UNIV_real"; +val SReal_iff = thm "SReal_iff"; +val hypreal_of_real_image = thm "hypreal_of_real_image"; +val inv_hypreal_of_real_image = thm "inv_hypreal_of_real_image"; +val SReal_hypreal_of_real_image = thm "SReal_hypreal_of_real_image"; +val SReal_dense = thm "SReal_dense"; +val SReal_sup_lemma = thm "SReal_sup_lemma"; +val SReal_sup_lemma2 = thm "SReal_sup_lemma2"; +val hypreal_of_real_isUb_iff = thm "hypreal_of_real_isUb_iff"; +val hypreal_of_real_isLub1 = thm "hypreal_of_real_isLub1"; +val hypreal_of_real_isLub2 = thm "hypreal_of_real_isLub2"; +val hypreal_of_real_isLub_iff = thm "hypreal_of_real_isLub_iff"; +val lemma_isUb_hypreal_of_real = thm "lemma_isUb_hypreal_of_real"; +val lemma_isLub_hypreal_of_real = thm "lemma_isLub_hypreal_of_real"; +val lemma_isLub_hypreal_of_real2 = thm "lemma_isLub_hypreal_of_real2"; +val SReal_complete = thm "SReal_complete"; +val HFinite_add = thm "HFinite_add"; +val HFinite_mult = thm "HFinite_mult"; +val HFinite_minus_iff = thm "HFinite_minus_iff"; +val SReal_subset_HFinite = thm "SReal_subset_HFinite"; +val HFinite_hypreal_of_real = thm "HFinite_hypreal_of_real"; +val HFiniteD = thm "HFiniteD"; +val HFinite_hrabs_iff = thm "HFinite_hrabs_iff"; +val HFinite_number_of = thm "HFinite_number_of"; +val HFinite_0 = thm "HFinite_0"; +val HFinite_1 = thm "HFinite_1"; +val HFinite_bounded = thm "HFinite_bounded"; +val InfinitesimalD = thm "InfinitesimalD"; +val Infinitesimal_zero = thm "Infinitesimal_zero"; +val hypreal_sum_of_halves = thm "hypreal_sum_of_halves"; +val hypreal_half_gt_zero = thm "hypreal_half_gt_zero"; +val Infinitesimal_add = thm "Infinitesimal_add"; +val Infinitesimal_minus_iff = thm "Infinitesimal_minus_iff"; +val Infinitesimal_diff = thm "Infinitesimal_diff"; +val Infinitesimal_mult = thm "Infinitesimal_mult"; +val Infinitesimal_HFinite_mult = thm "Infinitesimal_HFinite_mult"; +val Infinitesimal_HFinite_mult2 = thm "Infinitesimal_HFinite_mult2"; +val HInfinite_inverse_Infinitesimal = thm "HInfinite_inverse_Infinitesimal"; +val HInfinite_mult = thm "HInfinite_mult"; +val HInfinite_add_ge_zero = thm "HInfinite_add_ge_zero"; +val HInfinite_add_ge_zero2 = thm "HInfinite_add_ge_zero2"; +val HInfinite_add_gt_zero = thm "HInfinite_add_gt_zero"; +val HInfinite_minus_iff = thm "HInfinite_minus_iff"; +val HInfinite_add_le_zero = thm "HInfinite_add_le_zero"; +val HInfinite_add_lt_zero = thm "HInfinite_add_lt_zero"; +val HFinite_sum_squares = thm "HFinite_sum_squares"; +val not_Infinitesimal_not_zero = thm "not_Infinitesimal_not_zero"; +val not_Infinitesimal_not_zero2 = thm "not_Infinitesimal_not_zero2"; +val Infinitesimal_hrabs_iff = thm "Infinitesimal_hrabs_iff"; +val HFinite_diff_Infinitesimal_hrabs = thm "HFinite_diff_Infinitesimal_hrabs"; +val hrabs_less_Infinitesimal = thm "hrabs_less_Infinitesimal"; +val hrabs_le_Infinitesimal = thm "hrabs_le_Infinitesimal"; +val Infinitesimal_interval = thm "Infinitesimal_interval"; +val Infinitesimal_interval2 = thm "Infinitesimal_interval2"; +val not_Infinitesimal_mult = thm "not_Infinitesimal_mult"; +val Infinitesimal_mult_disj = thm "Infinitesimal_mult_disj"; +val HFinite_Infinitesimal_not_zero = thm "HFinite_Infinitesimal_not_zero"; +val HFinite_Infinitesimal_diff_mult = thm "HFinite_Infinitesimal_diff_mult"; +val Infinitesimal_subset_HFinite = thm "Infinitesimal_subset_HFinite"; +val Infinitesimal_hypreal_of_real_mult = thm "Infinitesimal_hypreal_of_real_mult"; +val Infinitesimal_hypreal_of_real_mult2 = thm "Infinitesimal_hypreal_of_real_mult2"; +val mem_infmal_iff = thm "mem_infmal_iff"; +val approx_minus_iff = thm "approx_minus_iff"; +val approx_minus_iff2 = thm "approx_minus_iff2"; +val approx_refl = thm "approx_refl"; +val approx_sym = thm "approx_sym"; +val approx_trans = thm "approx_trans"; +val approx_trans2 = thm "approx_trans2"; +val approx_trans3 = thm "approx_trans3"; +val number_of_approx_reorient = thm "number_of_approx_reorient"; +val zero_approx_reorient = thm "zero_approx_reorient"; +val one_approx_reorient = thm "one_approx_reorient"; + +(*** re-orientation, following HOL/Integ/Bin.ML + We re-orient x @=y where x is 0, 1 or a numeral, unless y is as well! + ***) + +(*reorientation simprules using ==, for the following simproc*) +val meta_zero_approx_reorient = zero_approx_reorient RS eq_reflection; +val meta_one_approx_reorient = one_approx_reorient RS eq_reflection; +val meta_number_of_approx_reorient = number_of_approx_reorient RS eq_reflection + +(*reorientation simplification procedure: reorients (polymorphic) + 0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*) +fun reorient_proc sg _ (_ $ t $ u) = + case u of + Const("0", _) => None + | Const("1", _) => None + | Const("Numeral.number_of", _) $ _ => None + | _ => Some (case t of + Const("0", _) => meta_zero_approx_reorient + | Const("1", _) => meta_one_approx_reorient + | Const("Numeral.number_of", _) $ _ => + meta_number_of_approx_reorient); + +val approx_reorient_simproc = + Bin_Simprocs.prep_simproc + ("reorient_simproc", ["0@=x", "1@=x", "number_of w @= x"], reorient_proc); + +Addsimprocs [approx_reorient_simproc]; +*} + +lemma Infinitesimal_approx_minus: "(x-y \ Infinitesimal) = (x @= y)" +by (auto simp add: hypreal_diff_def 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 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" +proof (unfold approx_def) + assume inf: "a + - b \ Infinitesimal" "c + - d \ Infinitesimal" + have "a + c + - (b + d) = (a + - b) + (c + - d)" by arith + 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 (no_asm) add: hypreal_add_commute) +done + +lemma approx_minus2: "-a @= -b ==> a @= b" +by (auto dest: approx_minus) + +lemma approx_minus_cancel: "(-a @= -b) = (a @= b)" +by (blast intro: approx_minus approx_minus2) + +declare approx_minus_cancel [simp] + +lemma approx_add_minus: "[| a @= b; c @= d |] ==> a + -c @= b + -d" +by (blast intro!: approx_add approx_minus) + +lemma approx_mult1: "[| a @= b; c: HFinite|] ==> a*c @= b*c" +by (simp add: approx_def Infinitesimal_HFinite_mult minus_mult_left + left_distrib [symmetric] + del: minus_mult_left [symmetric]) + +lemma approx_mult2: "[|a @= b; c: HFinite|] ==> c*a @= c*b" +apply (simp (no_asm_simp) add: approx_mult1 hypreal_mult_commute) +done + +lemma approx_mult_subst: "[|u @= v*x; x @= y; v \ HFinite|] ==> u @= v*y" +by (blast intro: approx_mult2 approx_trans) + +lemma approx_mult_subst2: "[| u @= x*v; x @= y; v \ HFinite |] ==> u @= y*v" +by (blast intro: approx_mult1 approx_trans) + +lemma approx_mult_subst_SReal: "[| u @= x*hypreal_of_real v; x @= y |] ==> u @= y*hypreal_of_real v" +by (auto intro: approx_mult_subst2) + +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 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 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: minus_add_distrib hypreal_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: minus_add_distrib hypreal_add_assoc [symmetric]) +done + +lemma Infinitesimal_add_approx_self2: "y \ Infinitesimal ==> x @= y + x" +by (auto dest: Infinitesimal_add_approx_self simp add: hypreal_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_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: hypreal_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: minus_add_distrib approx_minus_iff [symmetric] add_ac) +done + +lemma approx_add_right_cancel: "b + d @= c + d ==> b @= c" +apply (rule approx_add_left_cancel) +apply (simp add: hypreal_add_commute) +done + +lemma approx_add_mono1: "b @= c ==> d + b @= d + c" +apply (rule approx_minus_iff [THEN iffD2]) +apply (simp add: minus_add_distrib approx_minus_iff [symmetric] add_ac) +done + +lemma approx_add_mono2: "b @= c ==> b + a @= c + a" +apply (simp (no_asm_simp) add: hypreal_add_commute approx_add_mono1) +done + +lemma approx_add_left_iff: "(a + b @= a + c) = (b @= c)" +by (fast elim: approx_add_left_cancel approx_add_mono1) + +declare approx_add_left_iff [simp] + +lemma approx_add_right_iff: "(b + a @= c + a) = (b @= c)" +apply (simp (no_asm) add: hypreal_add_commute) +done + +declare approx_add_right_iff [simp] + +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: hypreal_add_assoc) +done + +lemma approx_hypreal_of_real_HFinite: "x @= hypreal_of_real D ==> x \ HFinite" +by (rule approx_sym [THEN [2] approx_HFinite], auto) + +lemma approx_mult_HFinite: "[|a @= b; c @= d; b: HFinite; d: HFinite|] ==> a*c @= b*d" +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_hypreal_of_real: "[|a @= hypreal_of_real b; c @= hypreal_of_real d |] + ==> a*c @= hypreal_of_real b*hypreal_of_real d" +apply (blast intro!: approx_mult_HFinite approx_hypreal_of_real_HFinite HFinite_hypreal_of_real) +done + +lemma approx_SReal_mult_cancel_zero: "[| a \ Reals; a \ 0; a*x @= 0 |] ==> x @= 0" +apply (drule SReal_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) +apply (auto dest: approx_mult2 simp add: hypreal_mult_assoc [symmetric]) +done + +(* REM comments: newly added *) +lemma approx_mult_SReal1: "[| a \ Reals; x @= 0 |] ==> x*a @= 0" +by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult1) + +lemma approx_mult_SReal2: "[| a \ Reals; x @= 0 |] ==> a*x @= 0" +by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult2) + +lemma approx_mult_SReal_zero_cancel_iff: "[|a \ Reals; a \ 0 |] ==> (a*x @= 0) = (x @= 0)" +by (blast intro: approx_SReal_mult_cancel_zero approx_mult_SReal2) +declare approx_mult_SReal_zero_cancel_iff [simp] + +lemma approx_SReal_mult_cancel: "[| a \ Reals; a \ 0; a* w @= a*z |] ==> w @= z" +apply (drule SReal_inverse [THEN SReal_subset_HFinite [THEN subsetD]]) +apply (auto dest: approx_mult2 simp add: hypreal_mult_assoc [symmetric]) +done + +lemma approx_SReal_mult_cancel_iff1: "[| a \ Reals; a \ 0|] ==> (a* w @= a*z) = (w @= z)" +by (auto intro!: approx_mult2 SReal_subset_HFinite [THEN subsetD] intro: approx_SReal_mult_cancel) +declare approx_SReal_mult_cancel_iff1 [simp] + +lemma approx_le_bound: "[| z <= 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 + +(*----------------------------------------------------------------- + Zero is the only infinitesimal that is also a real + -----------------------------------------------------------------*) + +lemma Infinitesimal_less_SReal: + "[| x \ Reals; 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_SReal2: "y \ Infinitesimal ==> \r \ Reals. 0 < r --> y < r" +by (blast intro: Infinitesimal_less_SReal) + +lemma SReal_not_Infinitesimal: + "[| 0 < y; y \ Reals|] ==> y \ Infinitesimal" +apply (simp add: Infinitesimal_def) +apply (auto simp add: hrabs_def) +done + +lemma SReal_minus_not_Infinitesimal: "[| y < 0; y \ Reals |] ==> y \ Infinitesimal" +apply (subst Infinitesimal_minus_iff [symmetric]) +apply (rule SReal_not_Infinitesimal, auto) +done + +lemma SReal_Int_Infinitesimal_zero: "Reals Int Infinitesimal = {0}" +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 \ Reals; x \ Infinitesimal|] ==> x = 0" +by (cut_tac SReal_Int_Infinitesimal_zero, blast) + +lemma SReal_HFinite_diff_Infinitesimal: "[| x \ Reals; x \ 0 |] ==> x \ HFinite - Infinitesimal" +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) + +lemma hypreal_of_real_Infinitesimal_iff_0: "(hypreal_of_real x \ Infinitesimal) = (x=0)" +apply auto +apply (rule ccontr) +apply (rule hypreal_of_real_HFinite_diff_Infinitesimal [THEN DiffD2], auto) +done +declare hypreal_of_real_Infinitesimal_iff_0 [iff] + +lemma number_of_not_Infinitesimal: "number_of w \ (0::hypreal) ==> number_of w \ Infinitesimal" +by (fast dest: SReal_number_of [THEN SReal_Infinitesimal_zero]) +declare number_of_not_Infinitesimal [simp] + +(*again: 1 is a special case, but not 0 this time*) +lemma one_not_Infinitesimal: "1 \ Infinitesimal" +apply (subst hypreal_numeral_1_eq_1 [symmetric]) +apply (rule number_of_not_Infinitesimal) +apply (simp (no_asm)) +done +declare one_not_Infinitesimal [simp] + +lemma approx_SReal_not_zero: "[| y \ Reals; 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 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 + +(*The premise y\0 is essential; otherwise x/y =0 and we lose the + HFinite premise.*) +lemma Infinitesimal_ratio: "[| y \ 0; y \ Infinitesimal; x/y \ HFinite |] ==> x \ Infinitesimal" +apply (drule Infinitesimal_HFinite_mult2, assumption) +apply (simp add: hypreal_divide_def hypreal_mult_assoc) +done + +(*------------------------------------------------------------------ + Standard Part Theorem: Every finite x: R* is infinitely + close to a unique real number (i.e a member of Reals) + ------------------------------------------------------------------*) +(*------------------------------------------------------------------ + Uniqueness: Two infinitely close reals are equal + ------------------------------------------------------------------*) + +lemma SReal_approx_iff: "[|x \ Reals; y \ Reals|] ==> (x @= y) = (x = y)" +apply auto +apply (simp add: approx_def) +apply (drule_tac x = y in SReal_minus) +apply (drule SReal_add, assumption) +apply (drule SReal_Infinitesimal_zero, assumption) +apply (drule sym) +apply (simp add: hypreal_eq_minus_iff [symmetric]) +done + +lemma number_of_approx_iff: "(number_of v @= number_of w) = (number_of v = (number_of w :: hypreal))" +by (auto simp add: SReal_approx_iff) +declare number_of_approx_iff [simp] + +(*And also for 0 @= #nn and 1 @= #nn, #nn @= 0 and #nn @= 1.*) +lemma [simp]: "(0 @= number_of w) = ((number_of w :: hypreal) = 0)" + "(number_of w @= 0) = ((number_of w :: hypreal) = 0)" + "(1 @= number_of w) = ((number_of w :: hypreal) = 1)" + "(number_of w @= 1) = ((number_of w :: hypreal) = 1)" + "~ (0 @= 1)" "~ (1 @= 0)" +by (auto simp only: SReal_number_of SReal_approx_iff Reals_0 Reals_1) + +lemma hypreal_of_real_approx_iff: "(hypreal_of_real k @= hypreal_of_real m) = (k = m)" +apply auto +apply (rule inj_hypreal_of_real [THEN injD]) +apply (rule SReal_approx_iff [THEN iffD1], auto) +done +declare hypreal_of_real_approx_iff [simp] + +lemma hypreal_of_real_approx_number_of_iff: "(hypreal_of_real k @= number_of w) = (k = number_of w)" +by (subst hypreal_of_real_approx_iff [symmetric], auto) +declare hypreal_of_real_approx_number_of_iff [simp] + +(*And also for 0 and 1.*) +(*And also for 0 @= #nn and 1 @= #nn, #nn @= 0 and #nn @= 1.*) +lemma [simp]: "(hypreal_of_real k @= 0) = (k = 0)" + "(hypreal_of_real k @= 1) = (k = 1)" + by (simp_all add: hypreal_of_real_approx_iff [symmetric]) + +lemma approx_unique_real: "[| r \ Reals; s \ Reals; r @= x; s @= x|] ==> r = s" +by (blast intro: SReal_approx_iff [THEN iffD1] approx_trans2) + +(*------------------------------------------------------------------ + Existence of unique real infinitely close + ------------------------------------------------------------------*) +(* lemma about lubs *) +lemma hypreal_isLub_unique: + "[| isLub R S x; isLub R S y |] ==> x = (y::hypreal)" +apply (frule isLub_isUb) +apply (frule_tac x = y in isLub_isUb) +apply (blast intro!: hypreal_le_anti_sym dest!: isLub_le_isUb) +done + +lemma lemma_st_part_ub: "x \ HFinite ==> \u. isUb Reals {s. s \ Reals & 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 \ Reals & s < x}" +apply (drule HFiniteD, safe) +apply (drule SReal_minus) +apply (rule_tac x = "-t" in exI) +apply (auto simp add: abs_less_iff) +done + +lemma lemma_st_part_subset: "{s. s \ Reals & s < x} <= Reals" +by auto + +lemma lemma_st_part_lub: "x \ HFinite ==> \t. isLub Reals {s. s \ Reals & s < x} t" +by (blast intro!: SReal_complete lemma_st_part_ub lemma_st_part_nonempty lemma_st_part_subset) + +lemma lemma_hypreal_le_left_cancel: "((t::hypreal) + r <= t) = (r <= 0)" +apply safe +apply (drule_tac c = "-t" in add_left_mono) +apply (drule_tac [2] c = t in add_left_mono) +apply (auto simp add: hypreal_add_assoc [symmetric]) +done + +lemma lemma_st_part_le1: "[| x \ HFinite; isLub Reals {s. s \ Reals & s < x} t; + r \ Reals; 0 < r |] ==> x <= t + r" +apply (frule isLubD1a) +apply (rule ccontr, drule linorder_not_le [THEN iffD2]) +apply (drule_tac x = t in SReal_add, assumption) +apply (drule_tac y = "t + r" in isLubD1 [THEN setleD], auto) +done + +lemma hypreal_setle_less_trans: "!!x::hypreal. [| S *<= x; 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_gt_isUb: + "!!x::hypreal. [| isUb R S x; x < y; y \ R |] ==> isUb R S y" +apply (simp add: isUb_def) +apply (blast intro: hypreal_setle_less_trans) +done + +lemma lemma_st_part_gt_ub: "[| x \ HFinite; x < y; y \ Reals |] + ==> isUb Reals {s. s \ Reals & s < x} y" +apply (auto dest: order_less_trans intro: order_less_imp_le intro!: isUbI setleI) +done + +lemma lemma_minus_le_zero: "t <= t + -r ==> r <= (0::hypreal)" +apply (drule_tac c = "-t" in add_left_mono) +apply (auto simp add: hypreal_add_assoc [symmetric]) +done + +lemma lemma_st_part_le2: "[| x \ HFinite; + isLub Reals {s. s \ Reals & s < x} t; + r \ Reals; 0 < r |] + ==> t + -r <= x" +apply (frule isLubD1a) +apply (rule ccontr, drule linorder_not_le [THEN iffD1]) +apply (drule SReal_minus, drule_tac x = t in SReal_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_hypreal_le_swap: "((x::hypreal) <= t + r) = (x + -t <= r)" +by auto + +lemma lemma_st_part1a: "[| x \ HFinite; + isLub Reals {s. s \ Reals & s < x} t; + r \ Reals; 0 < r |] + ==> x + -t <= r" +apply (blast intro!: lemma_hypreal_le_swap [THEN iffD1] lemma_st_part_le1) +done + +lemma lemma_hypreal_le_swap2: "(t + -r <= x) = (-(x + -t) <= (r::hypreal))" +by auto + +lemma lemma_st_part2a: "[| x \ HFinite; + isLub Reals {s. s \ Reals & s < x} t; + r \ Reals; 0 < r |] + ==> -(x + -t) <= r" +apply (blast intro!: lemma_hypreal_le_swap2 [THEN iffD1] lemma_st_part_le2) +done + +lemma lemma_SReal_ub: "(x::hypreal) \ Reals ==> isUb Reals {s. s \ Reals & s < x} x" +by (auto intro: isUbI setleI order_less_imp_le) + +lemma lemma_SReal_lub: "(x::hypreal) \ Reals ==> isLub Reals {s. s \ Reals & 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_st_part_not_eq1: "[| x \ HFinite; + isLub Reals {s. s \ Reals & s < x} t; + r \ Reals; 0 < r |] + ==> x + -t \ r" +apply auto +apply (frule isLubD1a [THEN SReal_minus]) +apply (drule SReal_add_cancel, assumption) +apply (drule_tac x = x in lemma_SReal_lub) +apply (drule hypreal_isLub_unique, assumption, auto) +done + +lemma lemma_st_part_not_eq2: "[| x \ HFinite; + isLub Reals {s. s \ Reals & s < x} t; + r \ Reals; 0 < r |] + ==> -(x + -t) \ r" +apply (auto simp add: minus_add_distrib) +apply (frule isLubD1a) +apply (drule SReal_add_cancel, assumption) +apply (drule_tac x = "-x" in SReal_minus, simp) +apply (drule_tac x = x in lemma_SReal_lub) +apply (drule hypreal_isLub_unique, assumption, auto) +done + +lemma lemma_st_part_major: "[| x \ HFinite; + isLub Reals {s. s \ Reals & s < x} t; + r \ Reals; 0 < r |] + ==> abs (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 dest: lemma_st_part_not_eq1 lemma_st_part_not_eq2 simp add: abs_less_iff) +done + +lemma lemma_st_part_major2: "[| x \ HFinite; + isLub Reals {s. s \ Reals & s < x} t |] + ==> \r \ Reals. 0 < r --> abs (x + -t) < r" +apply (blast dest!: lemma_st_part_major) +done + +(*---------------------------------------------- + Existence of real and Standard Part Theorem + ----------------------------------------------*) +lemma lemma_st_part_Ex: "x \ HFinite ==> + \t \ Reals. \r \ Reals. 0 < r --> abs (x + -t) < r" +apply (frule lemma_st_part_lub, safe) +apply (frule isLubD1a) +apply (blast dest: lemma_st_part_major2) +done + +lemma st_part_Ex: + "x \ HFinite ==> \t \ Reals. x @= t" +apply (simp add: approx_def Infinitesimal_def) +apply (drule lemma_st_part_Ex, auto) +done + +(*-------------------------------- + Unique real infinitely close + -------------------------------*) +lemma st_part_Ex1: "x \ HFinite ==> EX! t. t \ Reals & 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 + +(*------------------------------------------------------------------ + Finite and Infinite --- more theorems + ------------------------------------------------------------------*) + +lemma HFinite_Int_HInfinite_empty: "HFinite Int HInfinite = {}" + +apply (simp add: HFinite_def HInfinite_def) +apply (auto dest: order_less_trans) +done +declare HFinite_Int_HInfinite_empty [simp] + +lemma HFinite_not_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 +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 simp add: SReal_add) +done + +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 HFinite_HInfinite_iff: "(x \ HFinite) = (x \ HInfinite)" +apply (simp (no_asm) add: HInfinite_HFinite_iff) +done + +(*------------------------------------------------------------------ + Finite, Infinite and Infinitesimal --- more theorems + ------------------------------------------------------------------*) + +lemma HInfinite_diff_HFinite_Infinitesimal_disj: "x \ Infinitesimal ==> x \ HInfinite | x \ HFinite - Infinitesimal" +by (fast intro: not_HFinite_HInfinite) + +lemma HFinite_inverse: "[| x \ HFinite; x \ Infinitesimal |] ==> inverse x \ HFinite" +apply (cut_tac x = "inverse x" in HInfinite_HFinite_disj) +apply (auto dest!: HInfinite_inverse_Infinitesimal) +done + +lemma HFinite_inverse2: "x \ HFinite - Infinitesimal ==> inverse x \ HFinite" +by (blast intro: HFinite_inverse) + +(* stronger statement possible in fact *) +lemma Infinitesimal_inverse_HFinite: "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 + +lemma HFinite_not_Infinitesimal_inverse: "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 hypreal_mult_inverse) +done + +lemma approx_inverse: "[| 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: hypreal_mult_assoc) +done + +(*Used for NSLIM_inverse, NSLIMSEQ_inverse*) +lemmas hypreal_of_real_approx_inverse = hypreal_of_real_HFinite_diff_Infinitesimal [THEN [2] approx_inverse] + +lemma inverse_add_Infinitesimal_approx: "[| x \ HFinite - Infinitesimal; + h \ Infinitesimal |] ==> inverse(x + h) @= inverse x" +apply (auto intro: approx_inverse approx_sym Infinitesimal_add_approx_self) +done + +lemma inverse_add_Infinitesimal_approx2: "[| x \ HFinite - Infinitesimal; + h \ Infinitesimal |] ==> inverse(h + x) @= inverse x" +apply (rule hypreal_add_commute [THEN subst]) +apply (blast intro: inverse_add_Infinitesimal_approx) +done + +lemma inverse_add_Infinitesimal_approx_Infinitesimal: "[| 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 + +lemma Infinitesimal_square_iff: "(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: hypreal_mult_assoc) +done +declare Infinitesimal_square_iff [symmetric, simp] + +lemma HFinite_square_iff: "(x*x \ HFinite) = (x \ HFinite)" +apply (auto intro: HFinite_mult) +apply (auto dest: HInfinite_mult simp add: HFinite_HInfinite_iff) +done +declare HFinite_square_iff [simp] + +lemma HInfinite_square_iff: "(x*x \ HInfinite) = (x \ HInfinite)" +by (auto simp add: HInfinite_HFinite_iff) +declare HInfinite_square_iff [simp] + +lemma approx_HFinite_mult_cancel: "[| 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: hypreal_mult_assoc [symmetric]) +done + +lemma approx_HFinite_mult_cancel_iff1: "a: HFinite-Infinitesimal ==> (a * w @= a * z) = (w @= z)" +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: "[| x \ HInfinite; y \ HFinite |] ==> x + y \ HInfinite" +apply (rule_tac y = "-y" in HInfinite_HFinite_add_cancel) +apply (auto simp add: hypreal_add_assoc HFinite_minus_iff) +done + +lemma HInfinite_ge_HInfinite: "[| x \ HInfinite; x <= y; 0 <= x |] ==> y \ HInfinite" +by (auto intro: HFinite_bounded simp add: HInfinite_HFinite_iff) + +lemma Infinitesimal_inverse_HInfinite: "[| x \ Infinitesimal; x \ 0 |] ==> inverse x \ HInfinite" +apply (rule ccontr, drule HFinite_HInfinite_iff [THEN iffD2]) +apply (auto dest: Infinitesimal_HFinite_mult2) +done + +lemma HInfinite_HFinite_not_Infinitesimal_mult: "[| 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: hypreal_mult_assoc HFinite_HInfinite_iff) +done + +lemma HInfinite_HFinite_not_Infinitesimal_mult2: "[| x \ HInfinite; y \ HFinite - Infinitesimal |] + ==> y * x \ HInfinite" +apply (auto simp add: hypreal_mult_commute HInfinite_HFinite_not_Infinitesimal_mult) +done + +lemma HInfinite_gt_SReal: "[| x \ HInfinite; 0 < x; y \ Reals |] ==> y < x" +by (auto dest!: bspec simp add: HInfinite_def hrabs_def order_less_imp_le) + +lemma HInfinite_gt_zero_gt_one: "[| x \ HInfinite; 0 < x |] ==> 1 < x" +by (auto intro: HInfinite_gt_SReal) + + +lemma not_HInfinite_one: "1 \ HInfinite" +apply (simp (no_asm) add: HInfinite_HFinite_iff) +done +declare not_HInfinite_one [simp] + +(*------------------------------------------------------------------ + more about absolute value (hrabs) + ------------------------------------------------------------------*) + +lemma approx_hrabs_disj: "abs x @= x | abs x @= -x" +by (cut_tac x = x in hrabs_disj, auto) + +(*------------------------------------------------------------------ + Theorems about monads + ------------------------------------------------------------------*) + +lemma monad_hrabs_Un_subset: "monad (abs x) <= monad(x) Un monad(-x)" +by (rule_tac x1 = x in hrabs_disj [THEN disjE], auto) + +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 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_hrabs_iff: "(x \ monad 0) = (abs x \ monad 0)" +apply (rule_tac x1 = x in hrabs_disj [THEN disjE]) +apply (auto simp add: monad_zero_minus_iff [symmetric]) +done + +lemma mem_monad_self: "x \ monad x" +by (simp add: monad_def) +declare mem_monad_self [simp] + +(*------------------------------------------------------------------ + Proof that x @= y ==> abs x @= abs 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_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 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_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 \ Infinitesimal |] ==> abs x @= abs 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 less_Infinitesimal_less: "[| 0 < x; x \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 Ball_mem_monad_gt_zero: "[| 0 < x; 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_less_zero: "[| x < 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 lemma_approx_gt_zero: "[|0 < x; x \ Infinitesimal; x @= y|] ==> 0 < y" +by (blast dest: Ball_mem_monad_gt_zero approx_subset_monad) + +lemma lemma_approx_less_zero: "[|x < 0; x \ Infinitesimal; x @= y|] ==> y < 0" +by (blast dest: Ball_mem_monad_less_zero approx_subset_monad) + +lemma approx_hrabs1: "[| x @= y; x < 0; x \ Infinitesimal |] ==> abs x @= abs y" +apply (frule lemma_approx_less_zero) +apply (assumption+) +apply (simp add: abs_if) +done + +lemma approx_hrabs2: "[| x @= y; 0 < x; x \ Infinitesimal |] ==> abs x @= abs y" +apply (frule lemma_approx_gt_zero) +apply (assumption+) +apply (simp add: abs_if) +done + +lemma approx_hrabs: "x @= y ==> abs x @= abs y" +apply (rule_tac Q = "x \ Infinitesimal" in excluded_middle [THEN disjE]) +apply (rule_tac x1 = x and y1 = 0 in linorder_less_linear [THEN disjE]) +apply (auto intro: approx_hrabs1 approx_hrabs2 Infinitesimal_approx_hrabs) +done + +lemma approx_hrabs_zero_cancel: "abs(x) @= 0 ==> x @= 0" +apply (cut_tac x = x in hrabs_disj) +apply (auto dest: approx_minus) +done + +lemma approx_hrabs_add_Infinitesimal: "e \ Infinitesimal ==> abs x @= abs(x+e)" +by (fast intro: approx_hrabs Infinitesimal_add_approx_self) + +lemma approx_hrabs_add_minus_Infinitesimal: "e \ Infinitesimal ==> abs x @= abs(x + -e)" +by (fast intro: approx_hrabs Infinitesimal_add_minus_approx_self) + +lemma hrabs_add_Infinitesimal_cancel: "[| e \ Infinitesimal; e' \ Infinitesimal; + abs(x+e) = abs(y+e')|] ==> abs x @= abs 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 + +lemma hrabs_add_minus_Infinitesimal_cancel: "[| e \ Infinitesimal; e' \ Infinitesimal; + abs(x + -e) = abs(y + -e')|] ==> abs x @= abs 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 + +lemma hypreal_less_minus_iff: "((x::hypreal) < y) = (0 < y + -x)" +by arith - (*standard real numbers as a subset of the hyperreals*) - SReal_def "Reals == {x. EX r. x = hypreal_of_real r}" +(* 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 hypreal_of_real_less_iff [THEN iffD2]) +apply (drule_tac x = "hypreal_of_real y + -hypreal_of_real x" in bspec) +apply (auto simp add: hypreal_add_commute abs_less_iff SReal_add SReal_minus) +done + +lemma Infinitesimal_add_hrabs_hypreal_of_real_less: "[| x \ Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |] + ==> abs (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 add: hypreal_of_real_hrabs) +done + +lemma Infinitesimal_add_hrabs_hypreal_of_real_less2: "[| x \ Infinitesimal; abs(hypreal_of_real r) < hypreal_of_real y |] + ==> abs (x + hypreal_of_real r) < hypreal_of_real y" +apply (rule hypreal_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 + +lemma hypreal_of_real_le_add_Infininitesimal_cancel2: "[| u \ Infinitesimal; v \ Infinitesimal; + hypreal_of_real x + u <= hypreal_of_real y + v |] + ==> x <= y" +apply (blast intro!: hypreal_of_real_le_iff [THEN iffD1] hypreal_of_real_le_add_Infininitesimal_cancel) +done + +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 + +(*used once, in Lim/NSDERIV_inverse*) +lemma Infinitesimal_add_not_zero: "[| h \ Infinitesimal; x \ 0 |] ==> hypreal_of_real x + h \ 0" +apply auto +apply (subgoal_tac "h = - hypreal_of_real x", auto) +done + +lemma Infinitesimal_square_cancel: "x*x + y*y \ Infinitesimal ==> x*x \ Infinitesimal" +apply (rule Infinitesimal_interval2) +apply (rule_tac [3] zero_le_square, assumption) +apply (auto simp add: zero_le_square) +done +declare Infinitesimal_square_cancel [simp] + +lemma HFinite_square_cancel: "x*x + y*y \ HFinite ==> x*x \ HFinite" +apply (rule HFinite_bounded, assumption) +apply (auto simp add: zero_le_square) +done +declare HFinite_square_cancel [simp] + +lemma Infinitesimal_square_cancel2: "x*x + y*y \ Infinitesimal ==> y*y \ Infinitesimal" +apply (rule Infinitesimal_square_cancel) +apply (rule hypreal_add_commute [THEN subst]) +apply (simp (no_asm)) +done +declare Infinitesimal_square_cancel2 [simp] + +lemma HFinite_square_cancel2: "x*x + y*y \ HFinite ==> y*y \ HFinite" +apply (rule HFinite_square_cancel) +apply (rule hypreal_add_commute [THEN subst]) +apply (simp (no_asm)) +done +declare HFinite_square_cancel2 [simp] + +lemma Infinitesimal_sum_square_cancel: "x*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) +done +declare Infinitesimal_sum_square_cancel [simp] + +lemma HFinite_sum_square_cancel: "x*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) +done +declare HFinite_sum_square_cancel [simp] + +lemma Infinitesimal_sum_square_cancel2: "y*y + x*x + z*z \ Infinitesimal ==> x*x \ Infinitesimal" +apply (rule Infinitesimal_sum_square_cancel) +apply (simp add: add_ac) +done +declare Infinitesimal_sum_square_cancel2 [simp] + +lemma HFinite_sum_square_cancel2: "y*y + x*x + z*z \ HFinite ==> x*x \ HFinite" +apply (rule HFinite_sum_square_cancel) +apply (simp add: add_ac) +done +declare HFinite_sum_square_cancel2 [simp] + +lemma Infinitesimal_sum_square_cancel3: "z*z + y*y + x*x \ Infinitesimal ==> x*x \ Infinitesimal" +apply (rule Infinitesimal_sum_square_cancel) +apply (simp add: add_ac) +done +declare Infinitesimal_sum_square_cancel3 [simp] + +lemma HFinite_sum_square_cancel3: "z*z + y*y + x*x \ HFinite ==> x*x \ HFinite" +apply (rule HFinite_sum_square_cancel) +apply (simp add: add_ac) +done +declare HFinite_sum_square_cancel3 [simp] + +lemma monad_hrabs_less: "[| y \ monad x; 0 < hypreal_of_real e |] + ==> abs (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 + +(*------------------------------------------------------------------ + 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_SReal: "x \ HFinite ==> st x \ Reals" +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_SReal_eq: "x \ Reals ==> st x = x" +apply (simp add: st_def) +apply (rule some_equality) +apply (fast intro: SReal_subset_HFinite [THEN subsetD]) +apply (blast dest: SReal_approx_iff [THEN iffD1]) +done + +(* ???should be added to simpset *) +lemma st_hypreal_of_real: "st (hypreal_of_real x) = hypreal_of_real x" +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 approx_st_eq: + assumes "x \ HFinite" and "y \ HFinite" and "x @= y" + shows "st x = st y" +proof - + have "st x @= x" "st y @= y" "st x \ Reals" "st y \ Reals" + by (simp_all add: st_approx_self st_SReal prems) + with prems show ?thesis + 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_Infinitesimal_add_SReal: "[| x \ Reals; e \ Infinitesimal |] ==> st(x + e) = x" +apply (frule st_SReal_eq [THEN subst]) +prefer 2 apply assumption +apply (frule SReal_subset_HFinite [THEN subsetD]) +apply (frule Infinitesimal_subset_HFinite [THEN subsetD]) +apply (drule st_SReal_eq) +apply (rule approx_st_eq) +apply (auto intro: HFinite_add simp add: Infinitesimal_add_approx_self [THEN approx_sym]) +done + +lemma st_Infinitesimal_add_SReal2: "[| x \ Reals; e \ Infinitesimal |] + ==> st(e + x) = x" +apply (rule hypreal_add_commute [THEN subst]) +apply (blast intro!: st_Infinitesimal_add_SReal) +done + +lemma HFinite_st_Infinitesimal_add: "x \ HFinite ==> + \e \ Infinitesimal. x = st(x) + e" +apply (blast dest!: st_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2]) +done + +lemma st_add: + assumes x: "x \ HFinite" and y: "y \ HFinite" + shows "st (x + y) = st(x) + st(y)" +proof - + from HFinite_st_Infinitesimal_add [OF x] + obtain ex where ex: "ex \ Infinitesimal" "st x + ex = x" + by (blast intro: sym) + from HFinite_st_Infinitesimal_add [OF y] + obtain ey where ey: "ey \ Infinitesimal" "st y + ey = y" + by (blast intro: sym) + have "st (x + y) = st ((st x + ex) + (st y + ey))" + by (simp add: ex ey) + also have "... = st ((ex + ey) + (st x + st y))" by (simp add: add_ac) + also have "... = st x + st y" + by (simp add: prems st_SReal SReal_add Infinitesimal_add + st_Infinitesimal_add_SReal2) + finally show ?thesis . +qed + +lemma st_number_of: "st (number_of w) = number_of w" +by (rule SReal_number_of [THEN st_SReal_eq]) +declare st_number_of [simp] + +(*the theorem above for the special cases of zero and one*) +lemma [simp]: "st 0 = 0" "st 1 = 1" +by (simp_all add: st_SReal_eq) + +lemma st_minus: assumes "y \ HFinite" shows "st(-y) = -st(y)" +proof - + have "st (- y) + st y = 0" + by (simp add: prems st_add [symmetric] HFinite_minus_iff) + thus ?thesis by arith +qed + +lemma st_diff: + "[| x \ HFinite; y \ HFinite |] ==> st (x-y) = st(x) - st(y)" +apply (simp add: hypreal_diff_def) +apply (frule_tac y1 = y in st_minus [symmetric]) +apply (drule_tac x1 = y in HFinite_minus_iff [THEN iffD2]) +apply (simp (no_asm_simp) add: st_add) +done + +(* lemma *) +lemma lemma_st_mult: "[| x \ HFinite; y \ HFinite; + e \ Infinitesimal; + ea \ Infinitesimal |] + ==> e*y + x*ea + e*ea \ Infinitesimal" +apply (frule_tac x = e and y = y in Infinitesimal_HFinite_mult) +apply (frule_tac [2] x = ea and y = x in Infinitesimal_HFinite_mult) +apply (drule_tac [3] Infinitesimal_mult) +apply (auto intro: Infinitesimal_add simp add: add_ac mult_ac) +done + +lemma st_mult: "[| x \ HFinite; y \ HFinite |] + ==> st (x * y) = st(x) * st(y)" +apply (frule HFinite_st_Infinitesimal_add) +apply (frule_tac x = y in HFinite_st_Infinitesimal_add, safe) +apply (subgoal_tac "st (x * y) = st ((st x + e) * (st y + ea))") +apply (drule_tac [2] sym, drule_tac [2] sym) + prefer 2 apply simp +apply (erule_tac V = "x = st x + e" in thin_rl) +apply (erule_tac V = "y = st y + ea" in thin_rl) +apply (simp add: left_distrib right_distrib) +apply (drule st_SReal)+ +apply (simp (no_asm_use) add: hypreal_add_assoc) +apply (rule st_Infinitesimal_add_SReal) +apply (blast intro!: SReal_mult) +apply (drule SReal_subset_HFinite [THEN subsetD])+ +apply (rule hypreal_add_assoc [THEN subst]) +apply (blast intro!: lemma_st_mult) +done + +lemma st_Infinitesimal: "x \ Infinitesimal ==> st x = 0" +apply (subst hypreal_numeral_0_eq_0 [symmetric]) +apply (rule st_number_of [THEN subst]) +apply (rule approx_st_eq) +apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] simp add: mem_infmal_iff [symmetric]) +done + +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 hypreal_mult_left_cancel [THEN iffD1]) +apply (auto simp add: st_mult [symmetric] st_not_Infinitesimal HFinite_inverse) +apply (subst hypreal_mult_inverse, auto) +done + +lemma st_divide: "[| x \ HFinite; y \ HFinite; st y \ 0 |] + ==> st(x/y) = (st x) / (st y)" +apply (auto simp add: hypreal_divide_def st_mult st_not_Infinitesimal HFinite_inverse st_inverse) +done +declare st_divide [simp] + +lemma st_idempotent: "x \ HFinite ==> st(st(x)) = st(x)" +by (blast intro: st_HFinite st_approx_self approx_st_eq) +declare st_idempotent [simp] + +(*** lemmas ***) +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 + +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 + +lemma st_le: "[| x \ HFinite; y \ HFinite; x <= y |] ==> st(x) <= st(y)" +apply (frule HFinite_st_Infinitesimal_add) +apply (rotate_tac 1) +apply (frule HFinite_st_Infinitesimal_add, safe) +apply (rule Infinitesimal_add_st_le_cancel) +apply (rule_tac [3] x = ea and y = e in Infinitesimal_diff) +apply (auto simp add: hypreal_add_assoc [symmetric]) +done + +lemma st_zero_le: "[| 0 <= x; x \ HFinite |] ==> 0 <= st x" +apply (subst hypreal_numeral_0_eq_0 [symmetric]) +apply (rule st_number_of [THEN subst]) +apply (rule st_le, auto) +done + +lemma st_zero_ge: "[| x <= 0; x \ HFinite |] ==> st x <= 0" +apply (subst hypreal_numeral_0_eq_0 [symmetric]) +apply (rule st_number_of [THEN subst]) +apply (rule st_le, auto) +done + +lemma st_hrabs: "x \ HFinite ==> abs(st x) = st(abs 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 + + + +(*-------------------------------------------------------------------- + Alternative definitions for HFinite using Free ultrafilter + --------------------------------------------------------------------*) + +lemma FreeUltrafilterNat_Rep_hypreal: "[| X \ Rep_hypreal x; Y \ Rep_hypreal x |] + ==> {n. X n = Y n} \ FreeUltrafilterNat" +apply (rule_tac z = x in eq_Abs_hypreal, auto, ultra) +done + +lemma HFinite_FreeUltrafilterNat: + "x \ HFinite + ==> \X \ Rep_hypreal x. \u. {n. abs (X n) < u} \ FreeUltrafilterNat" +apply (rule eq_Abs_hypreal [of x]) +apply (auto simp add: HFinite_def abs_less_iff minus_less_iff [of x] + hypreal_less SReal_iff hypreal_minus hypreal_of_real_def) +apply (rule_tac x=x in bexI) +apply (rule_tac x=y in exI, auto, ultra) +done + +lemma FreeUltrafilterNat_HFinite: + "\X \ Rep_hypreal x. + \u. {n. abs (X n) < u} \ FreeUltrafilterNat + ==> x \ HFinite" +apply (rule eq_Abs_hypreal [of x]) +apply (auto simp add: HFinite_def abs_less_iff minus_less_iff [of x]) +apply (rule_tac x = "hypreal_of_real u" in bexI) +apply (auto simp add: hypreal_less SReal_iff hypreal_minus hypreal_of_real_def, ultra+) +done + +lemma HFinite_FreeUltrafilterNat_iff: "(x \ HFinite) = (\X \ Rep_hypreal x. + \u. {n. abs (X n) < u} \ FreeUltrafilterNat)" +apply (blast intro!: HFinite_FreeUltrafilterNat FreeUltrafilterNat_HFinite) +done + +(*-------------------------------------------------------------------- + Alternative definitions for HInfinite using Free ultrafilter + --------------------------------------------------------------------*) +lemma lemma_Compl_eq: "- {n. (u::real) < abs (xa n)} = {n. abs (xa n) <= u}" +by auto + +lemma lemma_Compl_eq2: "- {n. abs (xa n) < (u::real)} = {n. u <= abs (xa n)}" +by auto + +lemma lemma_Int_eq1: "{n. abs (xa n) <= (u::real)} Int {n. u <= abs (xa n)} + = {n. abs(xa n) = u}" +apply auto +done + +lemma lemma_FreeUltrafilterNat_one: "{n. abs (xa n) = u} <= {n. abs (xa n) < u + (1::real)}" +by auto + +(*------------------------------------- + Exclude this type of sets from free + ultrafilter for Infinite numbers! + -------------------------------------*) +lemma FreeUltrafilterNat_const_Finite: "[| xa: Rep_hypreal x; + {n. abs (xa n) = u} \ FreeUltrafilterNat + |] ==> x \ HFinite" +apply (rule FreeUltrafilterNat_HFinite) +apply (rule_tac x = xa in bexI) +apply (rule_tac x = "u + 1" in exI) +apply (ultra, assumption) +done + +lemma HInfinite_FreeUltrafilterNat: + "x \ HInfinite ==> \X \ Rep_hypreal x. + \u. {n. u < abs (X n)} \ FreeUltrafilterNat" +apply (frule HInfinite_HFinite_iff [THEN iffD1]) +apply (cut_tac x = x in Rep_hypreal_nonempty) +apply (auto simp del: Rep_hypreal_nonempty simp add: HFinite_FreeUltrafilterNat_iff Bex_def) +apply (drule spec)+ +apply auto +apply (drule_tac x = u in spec) +apply (drule FreeUltrafilterNat_Compl_mem)+ +apply (drule FreeUltrafilterNat_Int, assumption) +apply (simp add: lemma_Compl_eq lemma_Compl_eq2 lemma_Int_eq1) +apply (auto dest: FreeUltrafilterNat_const_Finite simp + add: HInfinite_HFinite_iff [THEN iffD1]) +done + +(* yet more lemmas! *) +lemma lemma_Int_HI: "{n. abs (Xa n) < u} Int {n. X n = Xa n} + <= {n. abs (X n) < (u::real)}" +apply auto +done + +lemma lemma_Int_HIa: "{n. u < abs (X n)} Int {n. abs (X n) < (u::real)} = {}" +by (auto intro: order_less_asym) + +lemma FreeUltrafilterNat_HInfinite: "\X \ Rep_hypreal x. \u. + {n. u < abs (X n)} \ FreeUltrafilterNat + ==> x \ HInfinite" +apply (rule HInfinite_HFinite_iff [THEN iffD2]) +apply (safe, drule HFinite_FreeUltrafilterNat, auto) +apply (drule_tac x = u in spec) +apply (drule FreeUltrafilterNat_Rep_hypreal, assumption) +apply (drule_tac Y = "{n. X n = Xa n}" in FreeUltrafilterNat_Int, simp) +apply (drule lemma_Int_HI [THEN [2] FreeUltrafilterNat_subset]) +apply (drule_tac Y = "{n. abs (X n) < u}" in FreeUltrafilterNat_Int) +apply (auto simp add: lemma_Int_HIa FreeUltrafilterNat_empty) +done + +lemma HInfinite_FreeUltrafilterNat_iff: "(x \ HInfinite) = (\X \ Rep_hypreal x. + \u. {n. u < abs (X n)} \ FreeUltrafilterNat)" +apply (blast intro!: HInfinite_FreeUltrafilterNat FreeUltrafilterNat_HInfinite) +done + +(*-------------------------------------------------------------------- + Alternative definitions for Infinitesimal using Free ultrafilter + --------------------------------------------------------------------*) + -syntax (xsymbols) - approx :: "[hypreal, hypreal] => bool" (infixl "\\" 50) - +lemma Infinitesimal_FreeUltrafilterNat: + "x \ Infinitesimal ==> \X \ Rep_hypreal x. + \u. 0 < u --> {n. abs (X n) < u} \ FreeUltrafilterNat" +apply (simp add: Infinitesimal_def) +apply (auto simp add: abs_less_iff minus_less_iff [of x]) +apply (rule eq_Abs_hypreal [of x]) +apply (auto, rule bexI [OF _ lemma_hyprel_refl], safe) +apply (drule hypreal_of_real_less_iff [THEN iffD2]) +apply (drule_tac x = "hypreal_of_real u" in bspec, auto) +apply (auto simp add: hypreal_less hypreal_minus hypreal_of_real_def, ultra) +done + +lemma FreeUltrafilterNat_Infinitesimal: + "\X \ Rep_hypreal x. + \u. 0 < u --> {n. abs (X n) < u} \ FreeUltrafilterNat + ==> x \ Infinitesimal" +apply (simp add: Infinitesimal_def) +apply (rule eq_Abs_hypreal [of x]) +apply (auto simp add: abs_less_iff abs_interval_iff minus_less_iff [of x]) +apply (auto simp add: SReal_iff) +apply (drule_tac [!] x=y in spec) +apply (auto simp add: hypreal_less hypreal_minus hypreal_of_real_def, ultra+) +done + +lemma Infinitesimal_FreeUltrafilterNat_iff: "(x \ Infinitesimal) = (\X \ Rep_hypreal x. + \u. 0 < u --> {n. abs (X n) < u} \ FreeUltrafilterNat)" +apply (blast intro!: Infinitesimal_FreeUltrafilterNat FreeUltrafilterNat_Infinitesimal) +done + +(*------------------------------------------------------------------------ + 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 add: real_of_nat_Suc_gt_zero) +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 (no_asm_use) add: SReal_inverse) +apply (rule real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive, THEN hypreal_of_real_less_iff [THEN iffD2], THEN [2] impE]) +prefer 2 apply assumption +apply (simp add: real_of_nat_Suc_gt_zero hypreal_of_nat_def) +apply (auto dest!: reals_Archimedean simp add: SReal_iff) +apply (drule hypreal_of_real_less_iff [THEN iffD2]) +apply (simp add: real_of_nat_Suc_gt_zero hypreal_of_nat_def) +apply (blast intro: order_less_trans) +done + +lemma Infinitesimal_hypreal_of_nat_iff: + "Infinitesimal = {x. \n. abs x < inverse (hypreal_of_nat (Suc n))}" +apply (simp add: Infinitesimal_def) +apply (auto simp add: lemma_Infinitesimal2) +done + + +(*------------------------------------------------------------------------- + Proof that omega is an infinite number and + hence that epsilon 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) + +(*------------------------------------------- + Prove that any segment is finite and + hence cannot belong to FreeUltrafilterNat + -------------------------------------------*) +lemma finite_nat_segment: "finite {n::nat. n < m}" +apply (induct_tac "m") +apply (auto simp add: Suc_Un_eq) +done + +lemma finite_real_of_nat_segment: "finite {n::nat. real n < real (m::nat)}" +by (auto intro: finite_nat_segment) + +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 + +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 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) + +lemma finite_rabs_real_of_nat_le_real: "finite {n::nat. abs(real n) <= u}" +apply (simp (no_asm) add: real_of_nat_Suc_gt_zero finite_real_of_nat_le_real) +done + +lemma rabs_real_of_nat_le_real_FreeUltrafilterNat: "{n. abs(real n) <= u} \ FreeUltrafilterNat" +by (blast intro!: FreeUltrafilterNat_finite finite_rabs_real_of_nat_le_real) + +lemma FreeUltrafilterNat_nat_gt_real: "{n. u < real n} \ FreeUltrafilterNat" +apply (rule ccontr, drule FreeUltrafilterNat_Compl_mem) +apply (subgoal_tac "- {n::nat. u < real n} = {n. real n <= u}") +prefer 2 apply force +apply (simp add: finite_real_of_nat_le_real [THEN FreeUltrafilterNat_finite]) +done + +(*-------------------------------------------------------------- + The complement of {n. abs(real n) <= u} = + {n. u < abs (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) + +(*----------------------------------------------- + Omega is a member of HInfinite + -----------------------------------------------*) + +lemma hypreal_omega: "hyprel``{%n::nat. real (Suc n)} \ hypreal" +by auto + +lemma FreeUltrafilterNat_omega: "{n. u < real n} \ FreeUltrafilterNat" +apply (cut_tac u = u in rabs_real_of_nat_le_real_FreeUltrafilterNat) +apply (auto dest: FreeUltrafilterNat_Compl_mem simp add: Compl_real_le_eq) +done + +lemma HInfinite_omega: "omega: HInfinite" +apply (simp add: omega_def) +apply (auto intro!: FreeUltrafilterNat_HInfinite) +apply (rule bexI) +apply (rule_tac [2] lemma_hyprel_refl, auto) +apply (simp (no_asm) add: real_of_nat_Suc diff_less_eq [symmetric] FreeUltrafilterNat_omega) +done +declare HInfinite_omega [simp] + +(*----------------------------------------------- + Epsilon is a member of Infinitesimal + -----------------------------------------------*) + +lemma Infinitesimal_epsilon: "epsilon \ Infinitesimal" +by (auto intro!: HInfinite_inverse_Infinitesimal HInfinite_omega simp add: hypreal_epsilon_inverse_omega) +declare Infinitesimal_epsilon [simp] + +lemma HFinite_epsilon: "epsilon \ HFinite" +by (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]) +declare HFinite_epsilon [simp] + +lemma epsilon_approx_zero: "epsilon @= 0" +apply (simp (no_asm) add: mem_infmal_iff [symmetric]) +done +declare epsilon_approx_zero [simp] + +(*------------------------------------------------------------------------ + Needed for proof that we define a hyperreal [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 add: real_of_nat_Suc_gt_zero) +apply (simp add: real_mult_commute) +done + +lemma finite_inverse_real_of_posnat_gt_real: "0 < u ==> finite {n. u < inverse(real(Suc n))}" +apply (simp (no_asm_simp) add: real_of_nat_less_inverse_iff) +apply (simp (no_asm_simp) add: real_of_nat_Suc less_diff_eq [symmetric]) +apply (rule finite_real_of_nat_less_real) +done + +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))}" +apply (auto dest: order_le_imp_less_or_eq simp add: order_less_imp_le) +done + +lemma real_of_nat_inverse_le_iff: "(inverse (real(Suc n)) <= r) = (1 <= r * real(Suc n))" +apply (simp (no_asm) add: linorder_not_less [symmetric]) +apply (simp (no_asm) add: inverse_eq_divide) +apply (subst pos_less_divide_eq) +apply (simp (no_asm) add: real_of_nat_Suc_gt_zero) +apply (simp (no_asm) add: real_mult_commute) +done + +lemma real_of_nat_inverse_eq_iff: "(u = inverse (real(Suc n))) = (real(Suc n) = inverse u)" +by (auto simp add: inverse_inverse_eq real_of_nat_Suc_gt_zero real_not_refl2 [THEN not_sym]) + +lemma lemma_finite_omega_set2: "finite {n::nat. u = inverse(real(Suc n))}" +apply (simp (no_asm_simp) add: real_of_nat_inverse_eq_iff) +apply (cut_tac x = "inverse u - 1" in lemma_finite_omega_set) +apply (simp add: real_of_nat_Suc diff_eq_eq [symmetric] eq_commute) +done + +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_omega_set2 finite_inverse_real_of_posnat_gt_real) + +lemma inverse_real_of_posnat_ge_real_FreeUltrafilterNat: "0 < u ==> + {n. u <= inverse(real(Suc n))} \ FreeUltrafilterNat" +apply (blast intro!: FreeUltrafilterNat_finite finite_inverse_real_of_posnat_ge_real) +done + +(*-------------------------------------------------------------- + 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}" +apply (auto dest!: order_le_less_trans simp add: linorder_not_le) +done + +lemma FreeUltrafilterNat_inverse_real_of_posnat: "0 < u ==> + {n. inverse(real(Suc n)) < u} \ FreeUltrafilterNat" +apply (cut_tac u = u in inverse_real_of_posnat_ge_real_FreeUltrafilterNat) +apply (auto dest: FreeUltrafilterNat_Compl_mem simp add: Compl_le_inverse_eq) +done + +(*-------------------------------------------------------------- + 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 (the theorem was previously in REALTOPOS.thy). + -------------------------------------------------------------*) +(*----------------------------------------------------- + |X(n) - x| < 1/n ==> [] - hypreal_of_real x| \ Infinitesimal + -----------------------------------------------------*) +lemma real_seq_to_hypreal_Infinitesimal: "\n. abs(X n + -x) < inverse(real(Suc n)) + ==> Abs_hypreal(hyprel``{X}) + -hypreal_of_real x \ Infinitesimal" +apply (auto intro!: bexI dest: FreeUltrafilterNat_inverse_real_of_posnat FreeUltrafilterNat_all FreeUltrafilterNat_Int intro: order_less_trans FreeUltrafilterNat_subset simp add: hypreal_minus hypreal_of_real_def hypreal_add Infinitesimal_FreeUltrafilterNat_iff hypreal_inverse) +done + +lemma real_seq_to_hypreal_approx: "\n. abs(X n + -x) < inverse(real(Suc n)) + ==> Abs_hypreal(hyprel``{X}) @= hypreal_of_real x" +apply (subst approx_minus_iff) +apply (rule mem_infmal_iff [THEN subst]) +apply (erule real_seq_to_hypreal_Infinitesimal) +done + +lemma real_seq_to_hypreal_approx2: "\n. abs(x + -X n) < inverse(real(Suc n)) + ==> Abs_hypreal(hyprel``{X}) @= hypreal_of_real x" +apply (simp add: abs_minus_add_cancel real_seq_to_hypreal_approx) +done + +lemma real_seq_to_hypreal_Infinitesimal2: "\n. abs(X n + -Y n) < inverse(real(Suc n)) + ==> Abs_hypreal(hyprel``{X}) + + -Abs_hypreal(hyprel``{Y}) \ Infinitesimal" +by (auto intro!: bexI + dest: FreeUltrafilterNat_inverse_real_of_posnat + FreeUltrafilterNat_all FreeUltrafilterNat_Int + intro: order_less_trans FreeUltrafilterNat_subset + simp add: Infinitesimal_FreeUltrafilterNat_iff hypreal_minus + hypreal_add hypreal_inverse) + + +ML +{* +val Infinitesimal_def = thm"Infinitesimal_def"; +val HFinite_def = thm"HFinite_def"; +val HInfinite_def = thm"HInfinite_def"; +val st_def = thm"st_def"; +val monad_def = thm"monad_def"; +val galaxy_def = thm"galaxy_def"; +val approx_def = thm"approx_def"; +val SReal_def = thm"SReal_def"; + +val Infinitesimal_approx_minus = thm "Infinitesimal_approx_minus"; +val approx_monad_iff = thm "approx_monad_iff"; +val Infinitesimal_approx = thm "Infinitesimal_approx"; +val approx_add = thm "approx_add"; +val approx_minus = thm "approx_minus"; +val approx_minus2 = thm "approx_minus2"; +val approx_minus_cancel = thm "approx_minus_cancel"; +val approx_add_minus = thm "approx_add_minus"; +val approx_mult1 = thm "approx_mult1"; +val approx_mult2 = thm "approx_mult2"; +val approx_mult_subst = thm "approx_mult_subst"; +val approx_mult_subst2 = thm "approx_mult_subst2"; +val approx_mult_subst_SReal = thm "approx_mult_subst_SReal"; +val approx_eq_imp = thm "approx_eq_imp"; +val Infinitesimal_minus_approx = thm "Infinitesimal_minus_approx"; +val bex_Infinitesimal_iff = thm "bex_Infinitesimal_iff"; +val bex_Infinitesimal_iff2 = thm "bex_Infinitesimal_iff2"; +val Infinitesimal_add_approx = thm "Infinitesimal_add_approx"; +val Infinitesimal_add_approx_self = thm "Infinitesimal_add_approx_self"; +val Infinitesimal_add_approx_self2 = thm "Infinitesimal_add_approx_self2"; +val Infinitesimal_add_minus_approx_self = thm "Infinitesimal_add_minus_approx_self"; +val Infinitesimal_add_cancel = thm "Infinitesimal_add_cancel"; +val Infinitesimal_add_right_cancel = thm "Infinitesimal_add_right_cancel"; +val approx_add_left_cancel = thm "approx_add_left_cancel"; +val approx_add_right_cancel = thm "approx_add_right_cancel"; +val approx_add_mono1 = thm "approx_add_mono1"; +val approx_add_mono2 = thm "approx_add_mono2"; +val approx_add_left_iff = thm "approx_add_left_iff"; +val approx_add_right_iff = thm "approx_add_right_iff"; +val approx_HFinite = thm "approx_HFinite"; +val approx_hypreal_of_real_HFinite = thm "approx_hypreal_of_real_HFinite"; +val approx_mult_HFinite = thm "approx_mult_HFinite"; +val approx_mult_hypreal_of_real = thm "approx_mult_hypreal_of_real"; +val approx_SReal_mult_cancel_zero = thm "approx_SReal_mult_cancel_zero"; +val approx_mult_SReal1 = thm "approx_mult_SReal1"; +val approx_mult_SReal2 = thm "approx_mult_SReal2"; +val approx_mult_SReal_zero_cancel_iff = thm "approx_mult_SReal_zero_cancel_iff"; +val approx_SReal_mult_cancel = thm "approx_SReal_mult_cancel"; +val approx_SReal_mult_cancel_iff1 = thm "approx_SReal_mult_cancel_iff1"; +val approx_le_bound = thm "approx_le_bound"; +val Infinitesimal_less_SReal = thm "Infinitesimal_less_SReal"; +val Infinitesimal_less_SReal2 = thm "Infinitesimal_less_SReal2"; +val SReal_not_Infinitesimal = thm "SReal_not_Infinitesimal"; +val SReal_minus_not_Infinitesimal = thm "SReal_minus_not_Infinitesimal"; +val SReal_Int_Infinitesimal_zero = thm "SReal_Int_Infinitesimal_zero"; +val SReal_Infinitesimal_zero = thm "SReal_Infinitesimal_zero"; +val SReal_HFinite_diff_Infinitesimal = thm "SReal_HFinite_diff_Infinitesimal"; +val hypreal_of_real_HFinite_diff_Infinitesimal = thm "hypreal_of_real_HFinite_diff_Infinitesimal"; +val hypreal_of_real_Infinitesimal_iff_0 = thm "hypreal_of_real_Infinitesimal_iff_0"; +val number_of_not_Infinitesimal = thm "number_of_not_Infinitesimal"; +val one_not_Infinitesimal = thm "one_not_Infinitesimal"; +val approx_SReal_not_zero = thm "approx_SReal_not_zero"; +val HFinite_diff_Infinitesimal_approx = thm "HFinite_diff_Infinitesimal_approx"; +val Infinitesimal_ratio = thm "Infinitesimal_ratio"; +val SReal_approx_iff = thm "SReal_approx_iff"; +val number_of_approx_iff = thm "number_of_approx_iff"; +val hypreal_of_real_approx_iff = thm "hypreal_of_real_approx_iff"; +val hypreal_of_real_approx_number_of_iff = thm "hypreal_of_real_approx_number_of_iff"; +val approx_unique_real = thm "approx_unique_real"; +val hypreal_isLub_unique = thm "hypreal_isLub_unique"; +val hypreal_setle_less_trans = thm "hypreal_setle_less_trans"; +val hypreal_gt_isUb = thm "hypreal_gt_isUb"; +val st_part_Ex = thm "st_part_Ex"; +val st_part_Ex1 = thm "st_part_Ex1"; +val HFinite_Int_HInfinite_empty = thm "HFinite_Int_HInfinite_empty"; +val HFinite_not_HInfinite = thm "HFinite_not_HInfinite"; +val not_HFinite_HInfinite = thm "not_HFinite_HInfinite"; +val HInfinite_HFinite_disj = thm "HInfinite_HFinite_disj"; +val HInfinite_HFinite_iff = thm "HInfinite_HFinite_iff"; +val HFinite_HInfinite_iff = thm "HFinite_HInfinite_iff"; +val HInfinite_diff_HFinite_Infinitesimal_disj = thm "HInfinite_diff_HFinite_Infinitesimal_disj"; +val HFinite_inverse = thm "HFinite_inverse"; +val HFinite_inverse2 = thm "HFinite_inverse2"; +val Infinitesimal_inverse_HFinite = thm "Infinitesimal_inverse_HFinite"; +val HFinite_not_Infinitesimal_inverse = thm "HFinite_not_Infinitesimal_inverse"; +val approx_inverse = thm "approx_inverse"; +val hypreal_of_real_approx_inverse = thm "hypreal_of_real_approx_inverse"; +val inverse_add_Infinitesimal_approx = thm "inverse_add_Infinitesimal_approx"; +val inverse_add_Infinitesimal_approx2 = thm "inverse_add_Infinitesimal_approx2"; +val inverse_add_Infinitesimal_approx_Infinitesimal = thm "inverse_add_Infinitesimal_approx_Infinitesimal"; +val Infinitesimal_square_iff = thm "Infinitesimal_square_iff"; +val HFinite_square_iff = thm "HFinite_square_iff"; +val HInfinite_square_iff = thm "HInfinite_square_iff"; +val approx_HFinite_mult_cancel = thm "approx_HFinite_mult_cancel"; +val approx_HFinite_mult_cancel_iff1 = thm "approx_HFinite_mult_cancel_iff1"; +val approx_hrabs_disj = thm "approx_hrabs_disj"; +val monad_hrabs_Un_subset = thm "monad_hrabs_Un_subset"; +val Infinitesimal_monad_eq = thm "Infinitesimal_monad_eq"; +val mem_monad_iff = thm "mem_monad_iff"; +val Infinitesimal_monad_zero_iff = thm "Infinitesimal_monad_zero_iff"; +val monad_zero_minus_iff = thm "monad_zero_minus_iff"; +val monad_zero_hrabs_iff = thm "monad_zero_hrabs_iff"; +val mem_monad_self = thm "mem_monad_self"; +val approx_subset_monad = thm "approx_subset_monad"; +val approx_subset_monad2 = thm "approx_subset_monad2"; +val mem_monad_approx = thm "mem_monad_approx"; +val approx_mem_monad = thm "approx_mem_monad"; +val approx_mem_monad2 = thm "approx_mem_monad2"; +val approx_mem_monad_zero = thm "approx_mem_monad_zero"; +val Infinitesimal_approx_hrabs = thm "Infinitesimal_approx_hrabs"; +val less_Infinitesimal_less = thm "less_Infinitesimal_less"; +val Ball_mem_monad_gt_zero = thm "Ball_mem_monad_gt_zero"; +val Ball_mem_monad_less_zero = thm "Ball_mem_monad_less_zero"; +val approx_hrabs1 = thm "approx_hrabs1"; +val approx_hrabs2 = thm "approx_hrabs2"; +val approx_hrabs = thm "approx_hrabs"; +val approx_hrabs_zero_cancel = thm "approx_hrabs_zero_cancel"; +val approx_hrabs_add_Infinitesimal = thm "approx_hrabs_add_Infinitesimal"; +val approx_hrabs_add_minus_Infinitesimal = thm "approx_hrabs_add_minus_Infinitesimal"; +val hrabs_add_Infinitesimal_cancel = thm "hrabs_add_Infinitesimal_cancel"; +val hrabs_add_minus_Infinitesimal_cancel = thm "hrabs_add_minus_Infinitesimal_cancel"; +val hypreal_less_minus_iff = thm "hypreal_less_minus_iff"; +val Infinitesimal_add_hypreal_of_real_less = thm "Infinitesimal_add_hypreal_of_real_less"; +val Infinitesimal_add_hrabs_hypreal_of_real_less = thm "Infinitesimal_add_hrabs_hypreal_of_real_less"; +val Infinitesimal_add_hrabs_hypreal_of_real_less2 = thm "Infinitesimal_add_hrabs_hypreal_of_real_less2"; +val hypreal_of_real_le_add_Infininitesimal_cancel2 = thm"hypreal_of_real_le_add_Infininitesimal_cancel2"; +val hypreal_of_real_less_Infinitesimal_le_zero = thm "hypreal_of_real_less_Infinitesimal_le_zero"; +val Infinitesimal_add_not_zero = thm "Infinitesimal_add_not_zero"; +val Infinitesimal_square_cancel = thm "Infinitesimal_square_cancel"; +val HFinite_square_cancel = thm "HFinite_square_cancel"; +val Infinitesimal_square_cancel2 = thm "Infinitesimal_square_cancel2"; +val HFinite_square_cancel2 = thm "HFinite_square_cancel2"; +val Infinitesimal_sum_square_cancel = thm "Infinitesimal_sum_square_cancel"; +val HFinite_sum_square_cancel = thm "HFinite_sum_square_cancel"; +val Infinitesimal_sum_square_cancel2 = thm "Infinitesimal_sum_square_cancel2"; +val HFinite_sum_square_cancel2 = thm "HFinite_sum_square_cancel2"; +val Infinitesimal_sum_square_cancel3 = thm "Infinitesimal_sum_square_cancel3"; +val HFinite_sum_square_cancel3 = thm "HFinite_sum_square_cancel3"; +val monad_hrabs_less = thm "monad_hrabs_less"; +val mem_monad_SReal_HFinite = thm "mem_monad_SReal_HFinite"; +val st_approx_self = thm "st_approx_self"; +val st_SReal = thm "st_SReal"; +val st_HFinite = thm "st_HFinite"; +val st_SReal_eq = thm "st_SReal_eq"; +val st_hypreal_of_real = thm "st_hypreal_of_real"; +val st_eq_approx = thm "st_eq_approx"; +val approx_st_eq = thm "approx_st_eq"; +val st_eq_approx_iff = thm "st_eq_approx_iff"; +val st_Infinitesimal_add_SReal = thm "st_Infinitesimal_add_SReal"; +val st_Infinitesimal_add_SReal2 = thm "st_Infinitesimal_add_SReal2"; +val HFinite_st_Infinitesimal_add = thm "HFinite_st_Infinitesimal_add"; +val st_add = thm "st_add"; +val st_number_of = thm "st_number_of"; +val st_minus = thm "st_minus"; +val st_diff = thm "st_diff"; +val st_mult = thm "st_mult"; +val st_Infinitesimal = thm "st_Infinitesimal"; +val st_not_Infinitesimal = thm "st_not_Infinitesimal"; +val st_inverse = thm "st_inverse"; +val st_divide = thm "st_divide"; +val st_idempotent = thm "st_idempotent"; +val Infinitesimal_add_st_less = thm "Infinitesimal_add_st_less"; +val Infinitesimal_add_st_le_cancel = thm "Infinitesimal_add_st_le_cancel"; +val st_le = thm "st_le"; +val st_zero_le = thm "st_zero_le"; +val st_zero_ge = thm "st_zero_ge"; +val st_hrabs = thm "st_hrabs"; +val FreeUltrafilterNat_HFinite = thm "FreeUltrafilterNat_HFinite"; +val HFinite_FreeUltrafilterNat_iff = thm "HFinite_FreeUltrafilterNat_iff"; +val FreeUltrafilterNat_const_Finite = thm "FreeUltrafilterNat_const_Finite"; +val FreeUltrafilterNat_HInfinite = thm "FreeUltrafilterNat_HInfinite"; +val HInfinite_FreeUltrafilterNat_iff = thm "HInfinite_FreeUltrafilterNat_iff"; +val Infinitesimal_FreeUltrafilterNat = thm "Infinitesimal_FreeUltrafilterNat"; +val FreeUltrafilterNat_Infinitesimal = thm "FreeUltrafilterNat_Infinitesimal"; +val Infinitesimal_FreeUltrafilterNat_iff = thm "Infinitesimal_FreeUltrafilterNat_iff"; +val Infinitesimal_hypreal_of_nat_iff = thm "Infinitesimal_hypreal_of_nat_iff"; +val Suc_Un_eq = thm "Suc_Un_eq"; +val finite_nat_segment = thm "finite_nat_segment"; +val finite_real_of_nat_segment = thm "finite_real_of_nat_segment"; +val finite_real_of_nat_less_real = thm "finite_real_of_nat_less_real"; +val finite_real_of_nat_le_real = thm "finite_real_of_nat_le_real"; +val finite_rabs_real_of_nat_le_real = thm "finite_rabs_real_of_nat_le_real"; +val rabs_real_of_nat_le_real_FreeUltrafilterNat = thm "rabs_real_of_nat_le_real_FreeUltrafilterNat"; +val FreeUltrafilterNat_nat_gt_real = thm "FreeUltrafilterNat_nat_gt_real"; +val hypreal_omega = thm "hypreal_omega"; +val FreeUltrafilterNat_omega = thm "FreeUltrafilterNat_omega"; +val HInfinite_omega = thm "HInfinite_omega"; +val Infinitesimal_epsilon = thm "Infinitesimal_epsilon"; +val HFinite_epsilon = thm "HFinite_epsilon"; +val epsilon_approx_zero = thm "epsilon_approx_zero"; +val real_of_nat_less_inverse_iff = thm "real_of_nat_less_inverse_iff"; +val finite_inverse_real_of_posnat_gt_real = thm "finite_inverse_real_of_posnat_gt_real"; +val real_of_nat_inverse_le_iff = thm "real_of_nat_inverse_le_iff"; +val real_of_nat_inverse_eq_iff = thm "real_of_nat_inverse_eq_iff"; +val finite_inverse_real_of_posnat_ge_real = thm "finite_inverse_real_of_posnat_ge_real"; +val inverse_real_of_posnat_ge_real_FreeUltrafilterNat = thm "inverse_real_of_posnat_ge_real_FreeUltrafilterNat"; +val FreeUltrafilterNat_inverse_real_of_posnat = thm "FreeUltrafilterNat_inverse_real_of_posnat"; +val real_seq_to_hypreal_Infinitesimal = thm "real_seq_to_hypreal_Infinitesimal"; +val real_seq_to_hypreal_approx = thm "real_seq_to_hypreal_approx"; +val real_seq_to_hypreal_approx2 = thm "real_seq_to_hypreal_approx2"; +val real_seq_to_hypreal_Infinitesimal2 = thm "real_seq_to_hypreal_Infinitesimal2"; +val HInfinite_HFinite_add = thm "HInfinite_HFinite_add"; +val HInfinite_ge_HInfinite = thm "HInfinite_ge_HInfinite"; +val Infinitesimal_inverse_HInfinite = thm "Infinitesimal_inverse_HInfinite"; +val HInfinite_HFinite_not_Infinitesimal_mult = thm "HInfinite_HFinite_not_Infinitesimal_mult"; +val HInfinite_HFinite_not_Infinitesimal_mult2 = thm "HInfinite_HFinite_not_Infinitesimal_mult2"; +val HInfinite_gt_SReal = thm "HInfinite_gt_SReal"; +val HInfinite_gt_zero_gt_one = thm "HInfinite_gt_zero_gt_one"; +val not_HInfinite_one = thm "not_HInfinite_one"; +*} + end diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/Star.ML --- a/src/HOL/Hyperreal/Star.ML Wed Jan 28 17:01:01 2004 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,490 +0,0 @@ -(* Title : STAR.ML - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Description : *-transforms -*) - -(*-------------------------------------------------------- - Preamble - Pulling "EX" over "ALL" - ---------------------------------------------------------*) - -(* 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) -*) -Goal "ALL x. EX y. Q x y ==> EX (f :: nat => nat). ALL x. Q x (f x)"; -by (res_inst_tac [("x","%x. LEAST y. Q x y")] exI 1); -by (blast_tac (claset() addIs [LeastI]) 1); -qed "no_choice"; - -(*------------------------------------------------------------ - Properties of the *-transform applied to sets of reals - ------------------------------------------------------------*) - -Goalw [starset_def] "*s*(UNIV::real set) = (UNIV::hypreal set)"; -by (Auto_tac); -qed "STAR_real_set"; -Addsimps [STAR_real_set]; - -Goalw [starset_def] "*s* {} = {}"; -by (Step_tac 1); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (dres_inst_tac [("x","%n. xa n")] bspec 1); -by (Auto_tac); -qed "STAR_empty_set"; -Addsimps [STAR_empty_set]; - -Goalw [starset_def] "*s* (A Un B) = *s* A Un *s* B"; -by (Auto_tac); -by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2)); -by (dtac FreeUltrafilterNat_Compl_mem 1); -by (dtac bspec 1 THEN assume_tac 1); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (Auto_tac); -by (Fuf_tac 1); -qed "STAR_Un"; - -Goalw [starset_def] "*s* (A Int B) = *s* A Int *s* B"; -by (Auto_tac); -by (blast_tac (claset() addIs [FreeUltrafilterNat_Int, - FreeUltrafilterNat_subset]) 3); -by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1)); -qed "STAR_Int"; - -Goalw [starset_def] "*s* -A = -( *s* A)"; -by (Auto_tac); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 2); -by (REPEAT(Step_tac 1) THEN Auto_tac); -by (Fuf_empty_tac 1); -by (dtac FreeUltrafilterNat_Compl_mem 1); -by (Fuf_tac 1); -qed "STAR_Compl"; - -goal Set.thy "(A - B) = (A Int (- B))"; -by (Step_tac 1); -qed "set_diff_iff2"; - -Goal "x ~: *s* F ==> x : *s* (- F)"; -by (auto_tac (claset(),simpset() addsimps [STAR_Compl])); -qed "STAR_mem_Compl"; - -Goal "*s* (A - B) = *s* A - *s* B"; -by (auto_tac (claset(),simpset() addsimps - [set_diff_iff2,STAR_Int,STAR_Compl])); -qed "STAR_diff"; - -Goalw [starset_def] "A <= B ==> *s* A <= *s* B"; -by (REPEAT(blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1)); -qed "STAR_subset"; - -Goalw [starset_def,hypreal_of_real_def] - "a : A ==> hypreal_of_real a : *s* A"; -by (auto_tac (claset() addIs [FreeUltrafilterNat_subset],simpset())); -qed "STAR_mem"; - -Goalw [starset_def] "hypreal_of_real ` A <= *s* A"; -by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def])); -by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 1); -qed "STAR_hypreal_of_real_image_subset"; - -Goalw [starset_def] "*s* X Int Reals = hypreal_of_real ` X"; -by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def,SReal_def])); -by (fold_tac [hypreal_of_real_def]); -by (rtac imageI 1 THEN rtac ccontr 1); -by (dtac bspec 1); -by (rtac lemma_hyprel_refl 1); -by (blast_tac (claset() addIs [FreeUltrafilterNat_subset]) 2); -by (Auto_tac); -qed "STAR_hypreal_of_real_Int"; - -Goal "x ~: hypreal_of_real ` A ==> ALL y: A. x ~= hypreal_of_real y"; -by (Auto_tac); -qed "lemma_not_hyprealA"; - -Goal "- {n. X n = xa} = {n. X n ~= xa}"; -by (Auto_tac); -qed "lemma_Compl_eq"; - -Goalw [starset_def] - "ALL n. (X n) ~: M \ -\ ==> Abs_hypreal(hyprel``{X}) ~: *s* M"; -by (Auto_tac THEN rtac bexI 1 THEN rtac lemma_hyprel_refl 2); -by (Auto_tac); -qed "STAR_real_seq_to_hypreal"; - -Goalw [starset_def] "*s* {x} = {hypreal_of_real x}"; -by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def])); -by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); -by (auto_tac (claset() addIs [FreeUltrafilterNat_subset],simpset())); -qed "STAR_singleton"; -Addsimps [STAR_singleton]; - -Goal "x ~: F ==> hypreal_of_real x ~: *s* F"; -by (auto_tac (claset(),simpset() addsimps - [starset_def,hypreal_of_real_def])); -by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2); -by (Auto_tac); -qed "STAR_not_mem"; - -Goal "[| x : *s* A; A <= B |] ==> x : *s* B"; -by (blast_tac (claset() addDs [STAR_subset]) 1); -qed "STAR_subset_closed"; - -(*------------------------------------------------------------------ - Nonstandard extension of a set (defined using a constant - sequence) as a special case of an internal set - -----------------------------------------------------------------*) - -Goalw [starset_n_def,starset_def] - "ALL n. (As n = A) ==> *sn* As = *s* A"; -by (Auto_tac); -qed "starset_n_starset"; - - -(*----------------------------------------------------------------*) -(* Theorems about nonstandard extensions of functions *) -(*----------------------------------------------------------------*) - -(*----------------------------------------------------------------*) -(* Nonstandard extension of a function (defined using a *) -(* constant sequence) as a special case of an internal function *) -(*----------------------------------------------------------------*) - -Goalw [starfun_n_def,starfun_def] - "ALL n. (F n = f) ==> *fn* F = *f* f"; -by (Auto_tac); -qed "starfun_n_starfun"; - - -(* - Prove that hrabs is a nonstandard extension of rabs without - use of congruence property (proved after this for general - nonstandard extensions of real valued functions). This makes - proof much longer- see comments at end of HREALABS.thy where - we proved a congruence theorem for hrabs. - - NEW!!! No need to prove all the lemmas anymore. Use the ultrafilter - tactic! -*) - -Goalw [is_starext_def] "is_starext abs abs"; -by (Step_tac 1); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); -by Auto_tac; -by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2); -by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2); -by (auto_tac (claset() addSDs [spec], - simpset() addsimps [hypreal_minus,hrabs_def, hypreal_zero_def, - hypreal_le_def, hypreal_less_def])); -by (TRYALL(Ultra_tac)); -by (TRYALL(arith_tac)); -qed "hrabs_is_starext_rabs"; - -Goal "[| X: Rep_hypreal z; Y: Rep_hypreal z |] \ -\ ==> {n. X n = Y n} : FreeUltrafilterNat"; -by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); -by (Auto_tac THEN Fuf_tac 1); -qed "Rep_hypreal_FreeUltrafilterNat"; - -(*----------------------------------------------------------------------- - Nonstandard extension of functions- congruence - -----------------------------------------------------------------------*) - -Goalw [congruent_def] "congruent hyprel (%X. hyprel``{%n. f (X n)})"; -by Auto_tac; -by (ALLGOALS(Fuf_tac)); -qed "starfun_congruent"; - -Goalw [starfun_def] - "( *f* f) (Abs_hypreal(hyprel``{%n. X n})) = \ -\ Abs_hypreal(hyprel `` {%n. f (X n)})"; -by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1); -by (simp_tac (simpset() addsimps - [hyprel_in_hypreal RS Abs_hypreal_inverse,[equiv_hyprel, - starfun_congruent] MRS UN_equiv_class]) 1); -qed "starfun"; - -(*------------------------------------------- - multiplication: ( *f ) x ( *g ) = *(f x g) - ------------------------------------------*) -Goal "( *f* f) xa * ( *f* g) xa = ( *f* (%x. f x * g x)) xa"; -by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun,hypreal_mult])); -qed "starfun_mult"; -Addsimps [starfun_mult RS sym]; - -(*--------------------------------------- - addition: ( *f ) + ( *g ) = *(f + g) - ---------------------------------------*) -Goal "( *f* f) xa + ( *f* g) xa = ( *f* (%x. f x + g x)) xa"; -by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun,hypreal_add])); -qed "starfun_add"; -Addsimps [starfun_add RS sym]; - -(*-------------------------------------------- - subtraction: ( *f ) + -( *g ) = *(f + -g) - -------------------------------------------*) - -Goal "- ( *f* f) x = ( *f* (%x. - f x)) x"; -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun, hypreal_minus])); -qed "starfun_minus"; -Addsimps [starfun_minus RS sym]; - -(*FIXME: delete*) -Goal "( *f* f) xa + -( *f* g) xa = ( *f* (%x. f x + -g x)) xa"; -by (Simp_tac 1); -qed "starfun_add_minus"; -Addsimps [starfun_add_minus RS sym]; - -Goalw [hypreal_diff_def,real_diff_def] - "( *f* f) xa - ( *f* g) xa = ( *f* (%x. f x - g x)) xa"; -by (rtac starfun_add_minus 1); -qed "starfun_diff"; -Addsimps [starfun_diff RS sym]; - -(*-------------------------------------- - composition: ( *f ) o ( *g ) = *(f o g) - ---------------------------------------*) - -Goal "(%x. ( *f* f) (( *f* g) x)) = *f* (%x. f (g x))"; -by (rtac ext 1); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun])); -qed "starfun_o2"; - -Goalw [o_def] "( *f* f) o ( *f* g) = ( *f* (f o g))"; -by (simp_tac (simpset() addsimps [starfun_o2]) 1); -qed "starfun_o"; - -(*-------------------------------------- - NS extension of constant function - --------------------------------------*) -Goal "( *f* (%x. k)) xa = hypreal_of_real k"; -by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun, - hypreal_of_real_def])); -qed "starfun_const_fun"; - -Addsimps [starfun_const_fun]; - -(*---------------------------------------------------- - the NS extension of the identity function - ----------------------------------------------------*) - -Goal "x @= hypreal_of_real a ==> ( *f* (%x. x)) x @= hypreal_of_real a"; -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun])); -qed "starfun_Idfun_approx"; - -Goal "( *f* (%x. x)) x = x"; -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun])); -qed "starfun_Id"; -Addsimps [starfun_Id]; - -(*---------------------------------------------------------------------- - the *-function is a (nonstandard) extension of the function - ----------------------------------------------------------------------*) - -Goalw [is_starext_def] "is_starext ( *f* f) f"; -by (Auto_tac); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (res_inst_tac [("z","y")] eq_Abs_hypreal 1); -by (auto_tac (claset() addSIs [bexI] ,simpset() addsimps [starfun])); -qed "is_starext_starfun"; - -(*---------------------------------------------------------------------- - Any nonstandard extension is in fact the *-function - ----------------------------------------------------------------------*) - -Goalw [is_starext_def] "is_starext F f ==> F = *f* f"; -by (rtac ext 1); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (dres_inst_tac [("x","x")] spec 1); -by (dres_inst_tac [("x","( *f* f) x")] spec 1); -by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem], - simpset() addsimps [starfun])); -by (Fuf_empty_tac 1); -qed "is_starfun_starext"; - -Goal "(is_starext F f) = (F = *f* f)"; -by (blast_tac (claset() addIs [is_starfun_starext,is_starext_starfun]) 1); -qed "is_starext_starfun_iff"; - -(*-------------------------------------------------------- - extented function has same solution as its standard - version for real arguments. i.e they are the same - for all real arguments - -------------------------------------------------------*) -Goal "( *f* f) (hypreal_of_real a) = hypreal_of_real (f a)"; -by (auto_tac (claset(),simpset() addsimps - [starfun,hypreal_of_real_def])); -qed "starfun_eq"; - -Addsimps [starfun_eq]; - -Goal "( *f* f) (hypreal_of_real a) @= hypreal_of_real (f a)"; -by (Auto_tac); -qed "starfun_approx"; - -(* useful for NS definition of derivatives *) -Goal "( *f* (%h. f (x + h))) xa = ( *f* f) (hypreal_of_real x + xa)"; -by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun, - hypreal_of_real_def,hypreal_add])); -qed "starfun_lambda_cancel"; - -Goal "( *f* (%h. f(g(x + h)))) xa = ( *f* (f o g)) (hypreal_of_real x + xa)"; -by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun, - hypreal_of_real_def,hypreal_add])); -qed "starfun_lambda_cancel2"; - -Goal "[| ( *f* f) xa @= l; ( *f* g) xa @= m; \ -\ l: HFinite; m: HFinite \ -\ |] ==> ( *f* (%x. f x * g x)) xa @= l * m"; -by (dtac approx_mult_HFinite 1); -by (REPEAT(assume_tac 1)); -by (auto_tac (claset() addIs [approx_sym RSN (2,approx_HFinite)], - simpset())); -qed "starfun_mult_HFinite_approx"; - -Goal "[| ( *f* f) xa @= l; ( *f* g) xa @= m \ -\ |] ==> ( *f* (%x. f x + g x)) xa @= l + m"; -by (auto_tac (claset() addIs [approx_add], simpset())); -qed "starfun_add_approx"; - -(*---------------------------------------------------- - 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 *) - -Goal "*f* abs = abs"; -by (rtac (hrabs_is_starext_rabs RS - (is_starext_starfun_iff RS iffD1) RS sym) 1); -qed "starfun_rabs_hrabs"; - -Goal "( *f* inverse) x = inverse(x)"; -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (auto_tac (claset(), - simpset() addsimps [starfun, hypreal_inverse, hypreal_zero_def])); -qed "starfun_inverse_inverse"; -Addsimps [starfun_inverse_inverse]; - -Goal "inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"; -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (auto_tac (claset(), - simpset() addsimps [starfun, hypreal_inverse])); -qed "starfun_inverse"; -Addsimps [starfun_inverse RS sym]; - -Goalw [hypreal_divide_def,real_divide_def] - "( *f* f) xa / ( *f* g) xa = ( *f* (%x. f x / g x)) xa"; -by Auto_tac; -qed "starfun_divide"; -Addsimps [starfun_divide RS sym]; - -Goal "inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x"; -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (auto_tac (claset() addIs [FreeUltrafilterNat_subset] - addSDs [FreeUltrafilterNat_Compl_mem], - simpset() addsimps [starfun, hypreal_inverse, hypreal_zero_def])); -qed "starfun_inverse2"; - -(*------------------------------------------------------------- - General lemma/theorem needed for proofs in elementary - topology of the reals - ------------------------------------------------------------*) -Goalw [starset_def] - "( *f* f) x : *s* A ==> x : *s* {x. f x : A}"; -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun])); -by (rename_tac "X" 1); -by (dres_inst_tac [("x","%n. f (X n)")] bspec 1); -by (Auto_tac THEN Fuf_tac 1); -qed "starfun_mem_starset"; - -(*------------------------------------------------------------ - Alternative definition for hrabs with rabs function - applied entrywise to equivalence class representative. - This is easily proved using starfun and ns extension thm - ------------------------------------------------------------*) -Goal "abs (Abs_hypreal (hyprel `` {X})) = \ -\ Abs_hypreal(hyprel `` {%n. abs (X n)})"; -by (simp_tac (simpset() addsimps [starfun_rabs_hrabs RS sym,starfun]) 1); -qed "hypreal_hrabs"; - -(*---------------------------------------------------------------- - 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 ( *f* f). See second NS set extension below. - ----------------------------------------------------------------*) -Goalw [starset_def] - "*s* {x. abs (x + - y) < r} = \ -\ {x. abs(x + -hypreal_of_real y) < hypreal_of_real r}"; -by (Step_tac 1); -by (ALLGOALS(res_inst_tac [("z","x")] eq_Abs_hypreal)); -by (auto_tac (claset() addSIs [exI] addSDs [bspec], - simpset() addsimps [hypreal_minus, hypreal_of_real_def,hypreal_add, - hypreal_hrabs,hypreal_less_def])); -by (Fuf_tac 1); -qed "STAR_rabs_add_minus"; - -Goalw [starset_def] - "*s* {x. abs (f x + - y) < r} = \ -\ {x. abs(( *f* f) x + -hypreal_of_real y) < hypreal_of_real r}"; -by (Step_tac 1); -by (ALLGOALS(res_inst_tac [("z","x")] eq_Abs_hypreal)); -by (auto_tac (claset() addSIs [exI] addSDs [bspec], - simpset() addsimps [hypreal_minus, hypreal_of_real_def,hypreal_add, - hypreal_hrabs,hypreal_less_def,starfun])); -by (Fuf_tac 1); -qed "STAR_starfun_rabs_add_minus"; - -(*------------------------------------------------------------------- - Another characterization of Infinitesimal and one of @= relation. - In this theory since hypreal_hrabs proved here. (To Check:) Maybe - move both if possible? - -------------------------------------------------------------------*) -Goal "(x:Infinitesimal) = \ -\ (EX X:Rep_hypreal(x). \ -\ ALL m. {n. abs(X n) < inverse(real(Suc m))} \ -\ : FreeUltrafilterNat)"; -by (res_inst_tac [("z","x")] eq_Abs_hypreal 1); -by (auto_tac (claset() addSIs [bexI,lemma_hyprel_refl], - simpset() addsimps [Infinitesimal_hypreal_of_nat_iff, - hypreal_of_real_def,hypreal_inverse, - hypreal_hrabs,hypreal_less, hypreal_of_nat_def])); -by (asm_full_simp_tac (simpset() addsimps [real_of_nat_Suc_gt_zero, - real_not_refl2 RS not_sym]) 1) ; -by (dres_inst_tac [("x","n")] spec 1); -by (Fuf_tac 1); -qed "Infinitesimal_FreeUltrafilterNat_iff2"; - -Goal "(Abs_hypreal(hyprel``{X}) @= Abs_hypreal(hyprel``{Y})) = \ -\ (ALL m. {n. abs (X n + - Y n) < \ -\ inverse(real(Suc m))} : FreeUltrafilterNat)"; -by (stac approx_minus_iff 1); -by (rtac (mem_infmal_iff RS subst) 1); -by (auto_tac (claset(), - simpset() addsimps [hypreal_minus, hypreal_add, - Infinitesimal_FreeUltrafilterNat_iff2])); -by (dres_inst_tac [("x","m")] spec 1); -by (Fuf_tac 1); -qed "approx_FreeUltrafilterNat_iff"; - -Goal "inj starfun"; -by (rtac injI 1); -by (rtac ext 1 THEN rtac ccontr 1); -by (dres_inst_tac [("x","Abs_hypreal(hyprel ``{%n. xa})")] fun_cong 1); -by (auto_tac (claset(),simpset() addsimps [starfun])); -qed "inj_starfun"; diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/Star.thy --- a/src/HOL/Hyperreal/Star.thy Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/Star.thy Thu Jan 29 16:51:17 2004 +0100 @@ -1,39 +1,589 @@ (* Title : Star.thy Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge - Description : defining *-transforms in NSA which extends sets of reals, + Description : defining *-transforms in NSA which extends sets of reals, and real=>real functions -*) +*) -Star = NSA + +header{*Star-Transforms in Non-Standard Analysis*} + +theory Star = NSA: constdefs (* nonstandard extension of sets *) - starset :: real set => hypreal set ("*s* _" [80] 80) + starset :: "real set => hypreal set" ("*s* _" [80] 80) "*s* A == {x. ALL X: Rep_hypreal(x). {n::nat. X n : A}: FreeUltrafilterNat}" (* internal sets *) - starset_n :: (nat => real set) => hypreal set ("*sn* _" [80] 80) - "*sn* As == {x. ALL X: Rep_hypreal(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}" - + starset_n :: "(nat => real set) => hypreal set" ("*sn* _" [80] 80) + "*sn* As == {x. ALL X: Rep_hypreal(x). {n::nat. X n : (As n)}: FreeUltrafilterNat}" + InternalSets :: "hypreal set set" "InternalSets == {X. EX As. X = *sn* As}" (* nonstandard extension of function *) - is_starext :: [hypreal => hypreal, real => real] => bool + is_starext :: "[hypreal => hypreal, real => real] => bool" "is_starext F f == (ALL x y. EX X: Rep_hypreal(x). EX Y: Rep_hypreal(y). ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))" - - starfun :: (real => real) => hypreal => hypreal ("*f* _" [80] 80) - "*f* f == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel``{%n. f(X n)}))" + + starfun :: "(real => real) => hypreal => hypreal" ("*f* _" [80] 80) + "*f* f == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel``{%n. f(X n)}))" (* internal functions *) - starfun_n :: (nat => (real => real)) => hypreal => hypreal ("*fn* _" [80] 80) - "*fn* F == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel``{%n. (F n)(X n)}))" + starfun_n :: "(nat => (real => real)) => hypreal => hypreal" + ("*fn* _" [80] 80) + "*fn* F == (%x. Abs_hypreal(UN X: Rep_hypreal(x). hyprel``{%n. (F n)(X n)}))" - InternalFuns :: (hypreal => hypreal) set + InternalFuns :: "(hypreal => hypreal) set" "InternalFuns == {X. EX F. X = *fn* F}" -end +(*-------------------------------------------------------- + Preamble - Pulling "EX" over "ALL" + ---------------------------------------------------------*) + +(* 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: "ALL x. EX y. Q x y ==> EX (f :: nat => nat). ALL x. Q x (f x)" +apply (rule_tac x = "%x. LEAST y. Q x y" in exI) +apply (blast intro: LeastI) +done + +(*------------------------------------------------------------ + Properties of the *-transform applied to sets of reals + ------------------------------------------------------------*) + +lemma STAR_real_set: "*s*(UNIV::real set) = (UNIV::hypreal set)" + +apply (unfold starset_def) +apply auto +done +declare STAR_real_set [simp] + +lemma STAR_empty_set: "*s* {} = {}" +apply (unfold starset_def) +apply safe +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (drule_tac x = "%n. xa n" in bspec) +apply auto +done +declare STAR_empty_set [simp] + +lemma STAR_Un: "*s* (A Un B) = *s* A Un *s* B" +apply (unfold starset_def) +apply auto + prefer 3 apply (blast intro: FreeUltrafilterNat_subset) + prefer 2 apply (blast intro: FreeUltrafilterNat_subset) +apply (drule FreeUltrafilterNat_Compl_mem) +apply (drule bspec , assumption) +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply auto +apply ultra +done + +lemma STAR_Int: "*s* (A Int B) = *s* A Int *s* B" +apply (unfold starset_def) +apply auto +prefer 3 apply (blast intro: FreeUltrafilterNat_Int FreeUltrafilterNat_subset) +apply (blast intro: FreeUltrafilterNat_subset)+ +done + +lemma STAR_Compl: "*s* -A = -( *s* A)" +apply (auto simp add: starset_def) +apply (rule_tac [!] z = "x" in eq_Abs_hypreal) +apply (auto dest!: bspec); +apply ultra +apply (drule FreeUltrafilterNat_Compl_mem) +apply ultra +done + +lemma STAR_mem_Compl: "x \ *s* F ==> x : *s* (- F)" +apply (auto simp add: STAR_Compl) +done + +lemma STAR_diff: "*s* (A - B) = *s* A - *s* B" +apply (auto simp add: Diff_eq STAR_Int STAR_Compl) +done + +lemma STAR_subset: "A <= B ==> *s* A <= *s* B" +apply (unfold starset_def) +apply (blast intro: FreeUltrafilterNat_subset)+ +done + +lemma STAR_mem: "a : A ==> hypreal_of_real a : *s* A" +apply (unfold starset_def hypreal_of_real_def) +apply (auto intro: FreeUltrafilterNat_subset) +done + +lemma STAR_hypreal_of_real_image_subset: "hypreal_of_real ` A <= *s* A" +apply (unfold starset_def) +apply (auto simp add: hypreal_of_real_def) +apply (blast intro: FreeUltrafilterNat_subset) +done + +lemma STAR_hypreal_of_real_Int: "*s* X Int Reals = hypreal_of_real ` X" +apply (unfold starset_def) +apply (auto simp add: hypreal_of_real_def SReal_def) +apply (simp add: hypreal_of_real_def [symmetric]) +apply (rule imageI , rule ccontr) +apply (drule bspec) +apply (rule lemma_hyprel_refl) +prefer 2 apply (blast intro: FreeUltrafilterNat_subset) +apply auto +done + +lemma lemma_not_hyprealA: "x \ hypreal_of_real ` A ==> ALL y: A. x \ hypreal_of_real y" +apply auto +done + +lemma lemma_Compl_eq: "- {n. X n = xa} = {n. X n \ xa}" +apply auto +done + +lemma STAR_real_seq_to_hypreal: + "ALL n. (X n) \ M + ==> Abs_hypreal(hyprel``{X}) \ *s* M" +apply (unfold starset_def) +apply (auto , rule bexI , rule_tac [2] lemma_hyprel_refl) +apply auto +done + +lemma STAR_singleton: "*s* {x} = {hypreal_of_real x}" +apply (unfold starset_def) +apply (auto simp add: hypreal_of_real_def) +apply (rule_tac z = "xa" in eq_Abs_hypreal) +apply (auto intro: FreeUltrafilterNat_subset) +done +declare STAR_singleton [simp] + +lemma STAR_not_mem: "x \ F ==> hypreal_of_real x \ *s* F" +apply (auto simp add: starset_def hypreal_of_real_def) +apply (rule bexI , rule_tac [2] lemma_hyprel_refl) +apply auto +done + +lemma STAR_subset_closed: "[| x : *s* A; A <= B |] ==> x : *s* B" +apply (blast dest: STAR_subset) +done + +(*------------------------------------------------------------------ + Nonstandard extension of a set (defined using a constant + sequence) as a special case of an internal set + -----------------------------------------------------------------*) + +lemma starset_n_starset: + "ALL n. (As n = A) ==> *sn* As = *s* A" + +apply (unfold starset_n_def starset_def) +apply auto +done + + +(*----------------------------------------------------------------*) +(* Theorems about nonstandard extensions of functions *) +(*----------------------------------------------------------------*) + +(*----------------------------------------------------------------*) +(* Nonstandard extension of a function (defined using a *) +(* constant sequence) as a special case of an internal function *) +(*----------------------------------------------------------------*) + +lemma starfun_n_starfun: + "ALL n. (F n = f) ==> *fn* F = *f* f" + +apply (unfold starfun_n_def starfun_def) +apply auto +done + + +(* + Prove that hrabs is a nonstandard extension of rabs without + use of congruence property (proved after this for general + nonstandard extensions of real valued functions). This makes + proof much longer- see comments at end of HREALABS.thy where + we proved a congruence theorem for hrabs. + + NEW!!! No need to prove all the lemmas anymore. Use the ultrafilter + tactic! +*) + +lemma hrabs_is_starext_rabs: "is_starext abs abs" + +apply (unfold is_starext_def) +apply safe +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (rule_tac z = "y" in eq_Abs_hypreal) +apply auto +apply (rule bexI , rule_tac [2] lemma_hyprel_refl) +apply (rule bexI , rule_tac [2] lemma_hyprel_refl) +apply (auto dest!: spec simp add: hypreal_minus hrabs_def hypreal_zero_def hypreal_le_def hypreal_less_def) +apply (arith | ultra)+ +done + +lemma Rep_hypreal_FreeUltrafilterNat: "[| X: Rep_hypreal z; Y: Rep_hypreal z |] + ==> {n. X n = Y n} : FreeUltrafilterNat" +apply (rule_tac z = "z" in eq_Abs_hypreal) +apply (auto , ultra) +done + +(*----------------------------------------------------------------------- + Nonstandard extension of functions- congruence + -----------------------------------------------------------------------*) + +lemma starfun_congruent: "congruent hyprel (%X. hyprel``{%n. f (X n)})" +apply (unfold congruent_def) +apply auto +apply ultra +done + +lemma starfun: + "( *f* f) (Abs_hypreal(hyprel``{%n. X n})) = + Abs_hypreal(hyprel `` {%n. f (X n)})" +apply (unfold starfun_def) +apply (rule_tac f = "Abs_hypreal" in arg_cong) +apply (simp add: hyprel_in_hypreal [THEN Abs_hypreal_inverse] + UN_equiv_class [OF equiv_hyprel starfun_congruent]) +done + +(*------------------------------------------- + multiplication: ( *f ) x ( *g ) = *(f x g) + ------------------------------------------*) +lemma starfun_mult: "( *f* f) xa * ( *f* g) xa = ( *f* (%x. f x * g x)) xa" +apply (rule_tac z = "xa" in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_mult) +done +declare starfun_mult [symmetric, simp] + +(*--------------------------------------- + addition: ( *f ) + ( *g ) = *(f + g) + ---------------------------------------*) +lemma starfun_add: "( *f* f) xa + ( *f* g) xa = ( *f* (%x. f x + g x)) xa" +apply (rule_tac z = "xa" in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_add) +done +declare starfun_add [symmetric, simp] + +(*-------------------------------------------- + subtraction: ( *f ) + -( *g ) = *(f + -g) + -------------------------------------------*) + +lemma starfun_minus: "- ( *f* f) x = ( *f* (%x. - f x)) x" +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_minus) +done +declare starfun_minus [symmetric, simp] + +(*FIXME: delete*) +lemma starfun_add_minus: "( *f* f) xa + -( *f* g) xa = ( *f* (%x. f x + -g x)) xa" +apply (simp (no_asm)) +done +declare starfun_add_minus [symmetric, simp] + +lemma starfun_diff: + "( *f* f) xa - ( *f* g) xa = ( *f* (%x. f x - g x)) xa" +apply (unfold hypreal_diff_def real_diff_def) +apply (rule starfun_add_minus) +done +declare starfun_diff [symmetric, simp] + +(*-------------------------------------- + composition: ( *f ) o ( *g ) = *(f o g) + ---------------------------------------*) + +lemma starfun_o2: "(%x. ( *f* f) (( *f* g) x)) = *f* (%x. f (g x))" +apply (rule ext) +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (auto simp add: starfun) +done + +lemma starfun_o: "( *f* f) o ( *f* g) = ( *f* (f o g))" +apply (unfold o_def) +apply (simp (no_asm) add: starfun_o2) +done + +(*-------------------------------------- + NS extension of constant function + --------------------------------------*) +lemma starfun_const_fun: "( *f* (%x. k)) xa = hypreal_of_real k" +apply (rule_tac z = "xa" in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_of_real_def) +done + +declare starfun_const_fun [simp] + +(*---------------------------------------------------- + the NS extension of the identity function + ----------------------------------------------------*) + +lemma starfun_Idfun_approx: "x @= hypreal_of_real a ==> ( *f* (%x. x)) x @= hypreal_of_real a" +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (auto simp add: starfun) +done + +lemma starfun_Id: "( *f* (%x. x)) x = x" +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (auto simp add: starfun) +done +declare starfun_Id [simp] + +(*---------------------------------------------------------------------- + the *-function is a (nonstandard) extension of the function + ----------------------------------------------------------------------*) + +lemma is_starext_starfun: "is_starext ( *f* f) f" + +apply (unfold is_starext_def) +apply auto +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (rule_tac z = "y" in eq_Abs_hypreal) +apply (auto intro!: bexI simp add: starfun) +done + +(*---------------------------------------------------------------------- + Any nonstandard extension is in fact the *-function + ----------------------------------------------------------------------*) + +lemma is_starfun_starext: "is_starext F f ==> F = *f* f" + +apply (unfold is_starext_def) +apply (rule ext) +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (drule_tac x = "x" in spec) +apply (drule_tac x = "( *f* f) x" in spec) +apply (auto dest!: FreeUltrafilterNat_Compl_mem simp add: starfun) +apply ultra +done + +lemma is_starext_starfun_iff: "(is_starext F f) = (F = *f* f)" +apply (blast intro: is_starfun_starext is_starext_starfun) +done + +(*-------------------------------------------------------- + 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) (hypreal_of_real a) = hypreal_of_real (f a)" +apply (auto simp add: starfun hypreal_of_real_def) +done + +declare starfun_eq [simp] + +lemma starfun_approx: "( *f* f) (hypreal_of_real a) @= hypreal_of_real (f a)" +apply auto +done + +(* useful for NS definition of derivatives *) +lemma starfun_lambda_cancel: "( *f* (%h. f (x + h))) xa = ( *f* f) (hypreal_of_real x + xa)" +apply (rule_tac z = "xa" in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_of_real_def hypreal_add) +done + +lemma starfun_lambda_cancel2: "( *f* (%h. f(g(x + h)))) xa = ( *f* (f o g)) (hypreal_of_real x + xa)" +apply (rule_tac z = "xa" in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_of_real_def hypreal_add) +done + +lemma starfun_mult_HFinite_approx: "[| ( *f* f) xa @= l; ( *f* g) xa @= m; + l: HFinite; m: HFinite + |] ==> ( *f* (%x. f x * g x)) xa @= l * m" +apply (drule approx_mult_HFinite) +apply (assumption)+ +apply (auto intro: approx_HFinite [OF _ approx_sym]) +done + +lemma starfun_add_approx: "[| ( *f* f) xa @= l; ( *f* g) xa @= m + |] ==> ( *f* (%x. f x + g x)) xa @= l + m" +apply (auto intro: approx_add) +done + +(*---------------------------------------------------- + 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 *) + +lemma starfun_rabs_hrabs: "*f* abs = abs" +apply (rule hrabs_is_starext_rabs [THEN is_starext_starfun_iff [THEN iffD1], symmetric]) +done + +lemma starfun_inverse_inverse: "( *f* inverse) x = inverse(x)" +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_inverse hypreal_zero_def) +done +declare starfun_inverse_inverse [simp] + +lemma starfun_inverse: "inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x" +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_inverse) +done +declare starfun_inverse [symmetric, simp] + +lemma starfun_divide: + "( *f* f) xa / ( *f* g) xa = ( *f* (%x. f x / g x)) xa" +apply (unfold hypreal_divide_def real_divide_def) +apply auto +done +declare starfun_divide [symmetric, simp] + +lemma starfun_inverse2: "inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x" +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (auto intro: FreeUltrafilterNat_subset dest!: FreeUltrafilterNat_Compl_mem simp add: starfun hypreal_inverse hypreal_zero_def) +done + +(*------------------------------------------------------------- + General lemma/theorem needed for proofs in elementary + topology of the reals + ------------------------------------------------------------*) +lemma starfun_mem_starset: + "( *f* f) x : *s* A ==> x : *s* {x. f x : A}" +apply (unfold starset_def) +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (auto simp add: starfun) +apply (rename_tac "X") +apply (drule_tac x = "%n. f (X n) " in bspec) +apply (auto , ultra) +done + +(*------------------------------------------------------------ + 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: "abs (Abs_hypreal (hyprel `` {X})) = + Abs_hypreal(hyprel `` {%n. abs (X n)})" +apply (simp (no_asm) add: starfun_rabs_hrabs [symmetric] starfun) +done + +(*---------------------------------------------------------------- + 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 ( *f* f). See second NS set extension below. + ----------------------------------------------------------------*) +lemma STAR_rabs_add_minus: + "*s* {x. abs (x + - y) < r} = + {x. abs(x + -hypreal_of_real y) < hypreal_of_real r}" +apply (unfold starset_def) +apply safe +apply (rule_tac [!] z = "x" in eq_Abs_hypreal) +apply (auto intro!: exI dest!: bspec simp add: hypreal_minus hypreal_of_real_def hypreal_add hypreal_hrabs hypreal_less) +apply ultra +done + +lemma STAR_starfun_rabs_add_minus: + "*s* {x. abs (f x + - y) < r} = + {x. abs(( *f* f) x + -hypreal_of_real y) < hypreal_of_real r}" +apply (unfold starset_def) +apply safe +apply (rule_tac [!] z = "x" in eq_Abs_hypreal) +apply (auto intro!: exI dest!: bspec simp add: hypreal_minus hypreal_of_real_def hypreal_add hypreal_hrabs hypreal_less starfun) +apply ultra +done + +(*------------------------------------------------------------------- + Another characterization of Infinitesimal and one of @= relation. + In this theory since hypreal_hrabs proved here. (To Check:) Maybe + move both if possible? + -------------------------------------------------------------------*) +lemma Infinitesimal_FreeUltrafilterNat_iff2: "(x:Infinitesimal) = + (EX X:Rep_hypreal(x). + ALL m. {n. abs(X n) < inverse(real(Suc m))} + : FreeUltrafilterNat)" +apply (rule_tac z = "x" in eq_Abs_hypreal) +apply (auto intro!: bexI lemma_hyprel_refl + simp add: Infinitesimal_hypreal_of_nat_iff hypreal_of_real_def + hypreal_inverse hypreal_hrabs hypreal_less hypreal_of_nat_def) +apply (drule_tac x = "n" in spec) +apply ultra +done + +lemma approx_FreeUltrafilterNat_iff: "(Abs_hypreal(hyprel``{X}) @= Abs_hypreal(hyprel``{Y})) = + (ALL m. {n. abs (X n + - Y n) < + inverse(real(Suc m))} : FreeUltrafilterNat)" +apply (subst approx_minus_iff) +apply (rule mem_infmal_iff [THEN subst]) +apply (auto simp add: hypreal_minus hypreal_add Infinitesimal_FreeUltrafilterNat_iff2) +apply (drule_tac x = "m" in spec) +apply ultra +done + +lemma inj_starfun: "inj starfun" +apply (rule inj_onI) +apply (rule ext , rule ccontr) +apply (drule_tac x = "Abs_hypreal (hyprel ``{%n. xa}) " in fun_cong) +apply (auto simp add: starfun) +done + +ML +{* +val starset_def = thm"starset_def"; +val starset_n_def = thm"starset_n_def"; +val InternalSets_def = thm"InternalSets_def"; +val is_starext_def = thm"is_starext_def"; +val starfun_def = thm"starfun_def"; +val starfun_n_def = thm"starfun_n_def"; +val InternalFuns_def = thm"InternalFuns_def"; + +val no_choice = thm "no_choice"; +val STAR_real_set = thm "STAR_real_set"; +val STAR_empty_set = thm "STAR_empty_set"; +val STAR_Un = thm "STAR_Un"; +val STAR_Int = thm "STAR_Int"; +val STAR_Compl = thm "STAR_Compl"; +val STAR_mem_Compl = thm "STAR_mem_Compl"; +val STAR_diff = thm "STAR_diff"; +val STAR_subset = thm "STAR_subset"; +val STAR_mem = thm "STAR_mem"; +val STAR_hypreal_of_real_image_subset = thm "STAR_hypreal_of_real_image_subset"; +val STAR_hypreal_of_real_Int = thm "STAR_hypreal_of_real_Int"; +val STAR_real_seq_to_hypreal = thm "STAR_real_seq_to_hypreal"; +val STAR_singleton = thm "STAR_singleton"; +val STAR_not_mem = thm "STAR_not_mem"; +val STAR_subset_closed = thm "STAR_subset_closed"; +val starset_n_starset = thm "starset_n_starset"; +val starfun_n_starfun = thm "starfun_n_starfun"; +val hrabs_is_starext_rabs = thm "hrabs_is_starext_rabs"; +val Rep_hypreal_FreeUltrafilterNat = thm "Rep_hypreal_FreeUltrafilterNat"; +val starfun_congruent = thm "starfun_congruent"; +val starfun = thm "starfun"; +val starfun_mult = thm "starfun_mult"; +val starfun_add = thm "starfun_add"; +val starfun_minus = thm "starfun_minus"; +val starfun_add_minus = thm "starfun_add_minus"; +val starfun_diff = thm "starfun_diff"; +val starfun_o2 = thm "starfun_o2"; +val starfun_o = thm "starfun_o"; +val starfun_const_fun = thm "starfun_const_fun"; +val starfun_Idfun_approx = thm "starfun_Idfun_approx"; +val starfun_Id = thm "starfun_Id"; +val is_starext_starfun = thm "is_starext_starfun"; +val is_starfun_starext = thm "is_starfun_starext"; +val is_starext_starfun_iff = thm "is_starext_starfun_iff"; +val starfun_eq = thm "starfun_eq"; +val starfun_approx = thm "starfun_approx"; +val starfun_lambda_cancel = thm "starfun_lambda_cancel"; +val starfun_lambda_cancel2 = thm "starfun_lambda_cancel2"; +val starfun_mult_HFinite_approx = thm "starfun_mult_HFinite_approx"; +val starfun_add_approx = thm "starfun_add_approx"; +val starfun_rabs_hrabs = thm "starfun_rabs_hrabs"; +val starfun_inverse_inverse = thm "starfun_inverse_inverse"; +val starfun_inverse = thm "starfun_inverse"; +val starfun_divide = thm "starfun_divide"; +val starfun_inverse2 = thm "starfun_inverse2"; +val starfun_mem_starset = thm "starfun_mem_starset"; +val hypreal_hrabs = thm "hypreal_hrabs"; +val STAR_rabs_add_minus = thm "STAR_rabs_add_minus"; +val STAR_starfun_rabs_add_minus = thm "STAR_starfun_rabs_add_minus"; +val Infinitesimal_FreeUltrafilterNat_iff2 = thm "Infinitesimal_FreeUltrafilterNat_iff2"; +val approx_FreeUltrafilterNat_iff = thm "approx_FreeUltrafilterNat_iff"; +val inj_starfun = thm "inj_starfun"; +*} + +end diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/Transcendental.ML --- a/src/HOL/Hyperreal/Transcendental.ML Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/Transcendental.ML Thu Jan 29 16:51:17 2004 +0100 @@ -1216,7 +1216,6 @@ by Auto_tac; qed "exp_ln_eq"; -Addsimps [hypreal_less_not_refl]; (* ------------------------------------------------------------------------ *) (* Basic properties of the trig functions *) diff -r c50188fe6366 -r b0064703967b src/HOL/Hyperreal/hypreal_arith.ML --- a/src/HOL/Hyperreal/hypreal_arith.ML Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Hyperreal/hypreal_arith.ML Thu Jan 29 16:51:17 2004 +0100 @@ -9,6 +9,8 @@ *) (*FIXME DELETE*) +val hypreal_mult_less_mono2 = thm"hypreal_mult_less_mono2"; + val hypreal_mult_left_mono = read_instantiate_sg(sign_of (the_context())) [("a","?a::hypreal")] mult_left_mono; diff -r c50188fe6366 -r b0064703967b src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/IsaMakefile Thu Jan 29 16:51:17 2004 +0100 @@ -154,10 +154,10 @@ Hyperreal/Lim.ML Hyperreal/Lim.thy Hyperreal/Log.ML Hyperreal/Log.thy\ Hyperreal/MacLaurin.ML Hyperreal/MacLaurin.thy\ Hyperreal/NatStar.ML Hyperreal/NatStar.thy\ - Hyperreal/NSA.ML Hyperreal/NSA.thy Hyperreal/NthRoot.thy\ + Hyperreal/NSA.thy Hyperreal/NthRoot.thy\ Hyperreal/Poly.ML Hyperreal/Poly.thy\ Hyperreal/SEQ.ML Hyperreal/SEQ.thy Hyperreal/Series.ML Hyperreal/Series.thy\ - Hyperreal/Star.ML Hyperreal/Star.thy Hyperreal/Transcendental.ML\ + Hyperreal/Star.thy Hyperreal/Transcendental.ML\ Hyperreal/Transcendental.thy Hyperreal/fuf.ML Hyperreal/hypreal_arith.ML \ Complex/Complex_Main.thy\ Complex/CLim.ML Complex/CLim.thy\ diff -r c50188fe6366 -r b0064703967b src/HOL/Ring_and_Field.thy --- a/src/HOL/Ring_and_Field.thy Wed Jan 28 17:01:01 2004 +0100 +++ b/src/HOL/Ring_and_Field.thy Thu Jan 29 16:51:17 2004 +0100 @@ -249,7 +249,8 @@ apply (erule add_strict_left_mono) done -lemma add_less_le_mono: "[| ad |] ==> a + c < b + (d::'a::ordered_semiring)" +lemma add_less_le_mono: + "[| ad |] ==> a + c < b + (d::'a::ordered_semiring)" apply (erule add_strict_right_mono [THEN order_less_le_trans]) apply (erule add_left_mono) done