# HG changeset patch # User paulson # Date 1079689863 -3600 # Node ID cc61fd03e589c9714c732ec86ef625389a702032 # Parent 758e7acdea2f3f0764ac5dae9b23e73aa501f09d conversion of Hyperreal/Lim to new-style diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/Hyperreal/HTranscendental.thy --- a/src/HOL/Hyperreal/HTranscendental.thy Fri Mar 19 10:50:06 2004 +0100 +++ b/src/HOL/Hyperreal/HTranscendental.thy Fri Mar 19 10:51:03 2004 +0100 @@ -573,7 +573,7 @@ apply (frule STAR_sin_Infinitesimal_divide [OF Infinitesimal_pi_divide_HNatInfinite pi_divide_HNatInfinite_not_zero]) -apply (auto simp add: hypreal_inverse_distrib) +apply (auto simp add: inverse_mult_distrib) apply (rule approx_SReal_mult_cancel [of "inverse (hypreal_of_real pi)"]) apply (auto intro: SReal_inverse simp add: divide_inverse mult_ac) done diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/Hyperreal/HyperDef.thy --- a/src/HOL/Hyperreal/HyperDef.thy Fri Mar 19 10:50:06 2004 +0100 +++ b/src/HOL/Hyperreal/HyperDef.thy Fri Mar 19 10:51:03 2004 +0100 @@ -543,64 +543,17 @@ by (auto dest: order_le_less_trans simp add: hrabs_def linorder_not_le) qed -lemma hypreal_mult_1_right: "z * (1::hypreal) = z" - by (rule Ring_and_Field.mult_1_right) - -lemma hypreal_mult_minus_1 [simp]: "(- (1::hypreal)) * z = -z" -by simp - -lemma hypreal_mult_minus_1_right [simp]: "z * (- (1::hypreal)) = -z" -by (subst hypreal_mult_commute, simp) - -(*Used ONCE: in NSA.ML*) -lemma hypreal_minus_distrib1: "-(y + -(x::hypreal)) = x + -y" -by (simp add: hypreal_add_commute) - -(*Used ONCE: in Lim.ML*) -lemma hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))" -by (auto simp add: hypreal_add_assoc) - lemma hypreal_eq_minus_iff: "((x::hypreal) = y) = (x + - y = 0)" apply auto apply (rule Ring_and_Field.add_right_cancel [of _ "-y", THEN iffD1], auto) done -(*Used 3 TIMES: in Lim.ML*) -lemma hypreal_not_eq_minus_iff: "(x \ a) = (x + -a \ (0::hypreal))" -by (auto dest: hypreal_eq_minus_iff [THEN iffD2]) - lemma hypreal_mult_left_cancel: "(c::hypreal) \ 0 ==> (c*a=c*b) = (a=b)" by auto lemma hypreal_mult_right_cancel: "(c::hypreal) \ 0 ==> (a*c=b*c) = (a=b)" by auto -lemma hypreal_mult_not_0: "[| x \ 0; y \ 0 |] ==> x * y \ (0::hypreal)" -by simp - -lemma hypreal_minus_inverse: "inverse(-x) = -inverse(x::hypreal)" - by (rule Ring_and_Field.inverse_minus_eq) - -lemma hypreal_inverse_distrib: "inverse(x*y) = inverse(x)*inverse(y::hypreal)" - by (rule Ring_and_Field.inverse_mult_distrib) - - -subsection{* Division lemmas *} - -lemma hypreal_divide_one: "x/(1::hypreal) = x" -by (simp add: hypreal_divide_def) - - -(** As with multiplication, pull minus signs OUT of the / operator **) - -lemma hypreal_add_divide_distrib: "(x+y)/(z::hypreal) = x/z + y/z" - by (rule Ring_and_Field.add_divide_distrib) - -lemma hypreal_inverse_add: - "[|(x::hypreal) \ 0; y \ 0 |] - ==> inverse(x) + inverse(y) = (x + y)*inverse(x*y)" -by (simp add: Ring_and_Field.inverse_add mult_assoc) - subsection{*The Embedding @{term hypreal_of_real} Preserves Field and Order Properties*} @@ -832,15 +785,11 @@ val hypreal_add_zero_right = thm "hypreal_add_zero_right"; val hypreal_add_minus = thm "hypreal_add_minus"; val hypreal_add_minus_left = thm "hypreal_add_minus_left"; -val hypreal_minus_distrib1 = thm "hypreal_minus_distrib1"; val hypreal_mult_congruent2 = thm "hypreal_mult_congruent2"; val hypreal_mult = thm "hypreal_mult"; val hypreal_mult_commute = thm "hypreal_mult_commute"; val hypreal_mult_assoc = thm "hypreal_mult_assoc"; val hypreal_mult_1 = thm "hypreal_mult_1"; -val hypreal_mult_1_right = thm "hypreal_mult_1_right"; -val hypreal_mult_minus_1 = thm "hypreal_mult_minus_1"; -val hypreal_mult_minus_1_right = thm "hypreal_mult_minus_1_right"; val hypreal_zero_not_eq_one = thm "hypreal_zero_not_eq_one"; val hypreal_inverse_congruent = thm "hypreal_inverse_congruent"; val hypreal_inverse = thm "hypreal_inverse"; @@ -848,14 +797,9 @@ val hypreal_mult_inverse_left = thm "hypreal_mult_inverse_left"; val hypreal_mult_left_cancel = thm "hypreal_mult_left_cancel"; val hypreal_mult_right_cancel = thm "hypreal_mult_right_cancel"; -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_not_refl2 = thm "hypreal_not_refl2"; val hypreal_less = thm "hypreal_less"; val hypreal_eq_minus_iff = thm "hypreal_eq_minus_iff"; -val hypreal_eq_minus_iff3 = thm "hypreal_eq_minus_iff3"; -val hypreal_not_eq_minus_iff = thm "hypreal_not_eq_minus_iff"; val hypreal_le = thm "hypreal_le"; val hypreal_le_refl = thm "hypreal_le_refl"; val hypreal_le_linear = thm "hypreal_le_linear"; @@ -872,9 +816,6 @@ val hypreal_of_real_zero = thm "hypreal_of_real_zero"; val hypreal_of_real_inverse = thm "hypreal_of_real_inverse"; val hypreal_of_real_divide = thm "hypreal_of_real_divide"; -val hypreal_divide_one = thm "hypreal_divide_one"; -val hypreal_add_divide_distrib = thm "hypreal_add_divide_distrib"; -val hypreal_inverse_add = thm "hypreal_inverse_add"; 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"; diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/Hyperreal/Integration.ML --- a/src/HOL/Hyperreal/Integration.ML Fri Mar 19 10:50:06 2004 +0100 +++ b/src/HOL/Hyperreal/Integration.ML Fri Mar 19 10:51:03 2004 +0100 @@ -508,7 +508,7 @@ by (Asm_full_simp_tac 2 THEN arith_tac 2); by (ALLGOALS (thin_tac "ALL xa. \ \ xa ~= x & abs (xa + - x) < s --> \ -\ abs ((f xa + - f x) / (xa + - x) + - f' x) * 2 < e")); +\ abs ((f xa - f x) / (xa - x) + - f' x) * 2 < e")); by (dres_inst_tac [("x","v")] spec 1 THEN Auto_tac); by (arith_tac 1); by (dres_inst_tac [("x","u")] spec 1 THEN Auto_tac); diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/Hyperreal/Lim.ML --- a/src/HOL/Hyperreal/Lim.ML Fri Mar 19 10:50:06 2004 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2307 +0,0 @@ -(* Title : Lim.ML - Author : Jacques D. Fleuriot - Copyright : 1998 University of Cambridge - Description : Theory of limits, continuity and - differentiation of real=>real functions -*) - -fun ARITH_PROVE str = prove_goal thy str - (fn prems => [cut_facts_tac prems 1,arith_tac 1]); - - -(*--------------------------------------------------------------- - Theory of limits, continuity and differentiation of - real=>real functions - ----------------------------------------------------------------*) - -Goalw [LIM_def] "(%x. k) -- x --> k"; -by Auto_tac; -qed "LIM_const"; -Addsimps [LIM_const]; - -(***-----------------------------------------------------------***) -(*** Some Purely Standard Proofs - Can be used for comparison ***) -(***-----------------------------------------------------------***) - -(*--------------- - LIM_add - ---------------*) -Goalw [LIM_def] - "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + g(x)) -- x --> (l + m)"; -by (Clarify_tac 1); -by (REPEAT(dres_inst_tac [("x","r/2")] spec 1)); -by (Asm_full_simp_tac 1); -by (Clarify_tac 1); -by (res_inst_tac [("x","s"),("y","sa")] linorder_cases 1); -by (res_inst_tac [("x","s")] exI 1); -by (res_inst_tac [("x","sa")] exI 2); -by (res_inst_tac [("x","sa")] exI 3); -by Safe_tac; -by (REPEAT(dres_inst_tac [("x","xa")] spec 1) - THEN step_tac (claset() addSEs [order_less_trans]) 1); -by (REPEAT(dres_inst_tac [("x","xa")] spec 2) - THEN step_tac (claset() addSEs [order_less_trans]) 2); -by (REPEAT(dres_inst_tac [("x","xa")] spec 3) - THEN step_tac (claset() addSEs [order_less_trans]) 3); -by (ALLGOALS(rtac (abs_sum_triangle_ineq RS order_le_less_trans))); -by (ALLGOALS(rtac (real_sum_of_halves RS subst))); -by (auto_tac (claset() addIs [add_strict_mono],simpset())); -qed "LIM_add"; - -Goalw [LIM_def] "f -- a --> L ==> (%x. -f(x)) -- a --> -L"; -by (subgoal_tac "ALL x. abs(- f x + L) = abs(f x + - L)" 1); -by (Asm_full_simp_tac 1); -by (asm_full_simp_tac (simpset() addsimps [real_abs_def]) 1); -qed "LIM_minus"; - -(*---------------------------------------------- - LIM_add_minus - ----------------------------------------------*) -Goal "[| f -- x --> l; g -- x --> m |] \ -\ ==> (%x. f(x) + -g(x)) -- x --> (l + -m)"; -by (blast_tac (claset() addDs [LIM_add,LIM_minus]) 1); -qed "LIM_add_minus"; - -(*---------------------------------------------- - LIM_zero - ----------------------------------------------*) -Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> 0"; -by (res_inst_tac [("a1","l")] ((right_minus RS subst)) 1); -by (rtac LIM_add_minus 1 THEN Auto_tac); -qed "LIM_zero"; - -(*-------------------------- - Limit not zero - --------------------------*) -Goalw [LIM_def] "k \\ 0 ==> ~ ((%x. k) -- x --> 0)"; -by (res_inst_tac [("x","k"),("y","0")] linorder_cases 1); -by (auto_tac (claset(), simpset() addsimps [real_abs_def])); -by (res_inst_tac [("x","-k")] exI 1); -by (res_inst_tac [("x","k")] exI 2); -by Auto_tac; -by (ALLGOALS(dres_inst_tac [("y","s")] real_dense)); -by Safe_tac; -by (ALLGOALS(res_inst_tac [("x","r + x")] exI)); -by Auto_tac; -qed "LIM_not_zero"; - -(* [| k \\ 0; (%x. k) -- x --> 0 |] ==> R *) -bind_thm("LIM_not_zeroE", LIM_not_zero RS notE); - -Goal "(%x. k) -- x --> L ==> k = L"; -by (rtac ccontr 1); -by (dtac LIM_zero 1); -by (rtac LIM_not_zeroE 1 THEN assume_tac 2); -by (arith_tac 1); -qed "LIM_const_eq"; - -(*------------------------ - Limit is Unique - ------------------------*) -Goal "[| f -- x --> L; f -- x --> M |] ==> L = M"; -by (dtac LIM_minus 1); -by (dtac LIM_add 1 THEN assume_tac 1); -by (auto_tac (claset() addSDs [LIM_const_eq RS sym], simpset())); -qed "LIM_unique"; - -(*------------- - LIM_mult_zero - -------------*) -Goalw [LIM_def] - "[| f -- x --> 0; g -- x --> 0 |] ==> (%x. f(x)*g(x)) -- x --> 0"; -by Safe_tac; -by (dres_inst_tac [("x","1")] spec 1); -by (dres_inst_tac [("x","r")] spec 1); -by (cut_facts_tac [real_zero_less_one] 1); -by (asm_full_simp_tac (simpset() addsimps [abs_mult]) 1); -by (Clarify_tac 1); -by (res_inst_tac [("x","s"),("y","sa")] - linorder_cases 1); -by (res_inst_tac [("x","s")] exI 1); -by (res_inst_tac [("x","sa")] exI 2); -by (res_inst_tac [("x","sa")] exI 3); -by Safe_tac; -by (REPEAT(dres_inst_tac [("x","xa")] spec 1) - THEN step_tac (claset() addSEs [order_less_trans]) 1); -by (REPEAT(dres_inst_tac [("x","xa")] spec 2) - THEN step_tac (claset() addSEs [order_less_trans]) 2); -by (REPEAT(dres_inst_tac [("x","xa")] spec 3) - THEN step_tac (claset() addSEs [order_less_trans]) 3); -by (ALLGOALS(res_inst_tac [("t","r")] (real_mult_1 RS subst))); -by (ALLGOALS(rtac abs_mult_less)); -by Auto_tac; -qed "LIM_mult_zero"; - -Goalw [LIM_def] "(%x. x) -- a --> a"; -by Auto_tac; -qed "LIM_self"; - -(*-------------------------------------------------------------- - Limits are equal for functions equal except at limit point - --------------------------------------------------------------*) -Goalw [LIM_def] - "[| \\x. x \\ a --> (f x = g x) |] \ -\ ==> (f -- a --> l) = (g -- a --> l)"; -by (auto_tac (claset(), simpset() addsimps [real_add_minus_iff])); -qed "LIM_equal"; - -Goal "[| (%x. f(x) + -g(x)) -- a --> 0; g -- a --> l |] \ -\ ==> f -- a --> l"; -by (dtac LIM_add 1 THEN assume_tac 1); -by (auto_tac (claset(), simpset() addsimps [real_add_assoc])); -qed "LIM_trans"; - -(***-------------------------------------------------------------***) -(*** End of Purely Standard Proofs ***) -(***-------------------------------------------------------------***) -(*-------------------------------------------------------------- - Standard and NS definitions of Limit - --------------------------------------------------------------*) -Goalw [LIM_def,NSLIM_def,approx_def] - "f -- x --> L ==> f -- x --NS> L"; -by (asm_full_simp_tac - (simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff]) 1); -by Safe_tac; -by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1); -by (auto_tac (claset(), - simpset() addsimps [real_add_minus_iff, starfun, hypreal_minus, - hypreal_of_real_def, hypreal_add])); -by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1); -by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1); -by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1); -by (subgoal_tac "\\n::nat. (xa n) \\ x & \ -\ abs ((xa n) + - x) < s --> abs (f (xa n) + - L) < u" 1); -by (Blast_tac 2); -by (dtac FreeUltrafilterNat_all 1); -by (Ultra_tac 1); -qed "LIM_NSLIM"; - -(*--------------------------------------------------------------------- - Limit: NS definition ==> standard definition - ---------------------------------------------------------------------*) - -Goal "\\s. 0 < s --> (\\xa. xa \\ x & \ -\ abs (xa + - x) < s & r \\ abs (f xa + -L)) \ -\ ==> \\n::nat. \\xa. xa \\ x & \ -\ abs(xa + -x) < inverse(real(Suc n)) & r \\ abs(f xa + -L)"; -by (Clarify_tac 1); -by (cut_inst_tac [("n1","n")] - (real_of_nat_Suc_gt_zero RS positive_imp_inverse_positive) 1); -by Auto_tac; -qed "lemma_LIM"; - -Goal "\\s. 0 < s --> (\\xa. xa \\ x & \ -\ abs (xa + - x) < s & r \\ abs (f xa + -L)) \ -\ ==> \\X. \\n::nat. X n \\ x & \ -\ abs(X n + -x) < inverse(real(Suc n)) & r \\ abs(f (X n) + -L)"; -by (dtac lemma_LIM 1); -by (dtac choice 1); -by (Blast_tac 1); -qed "lemma_skolemize_LIM2"; - -Goal "\\n. X n \\ x & \ -\ abs (X n + - x) < inverse (real(Suc n)) & \ -\ r \\ abs (f (X n) + - L) ==> \ -\ \\n. abs (X n + - x) < inverse (real(Suc n))"; -by (Auto_tac ); -qed "lemma_simp"; - -(*------------------- - NSLIM => LIM - -------------------*) - -Goalw [LIM_def,NSLIM_def,approx_def] - "f -- x --NS> L ==> f -- x --> L"; -by (asm_full_simp_tac - (simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff]) 1); -by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]); -by (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1); -by (dtac lemma_skolemize_LIM2 1); -by Safe_tac; -by (dres_inst_tac [("x","Abs_hypreal(hyprel``{X})")] spec 1); -by (auto_tac - (claset(), - simpset() addsimps [starfun, hypreal_minus, - hypreal_of_real_def,hypreal_add])); -by (dtac (lemma_simp RS real_seq_to_hypreal_Infinitesimal) 1); -by (asm_full_simp_tac - (simpset() addsimps - [Infinitesimal_FreeUltrafilterNat_iff,hypreal_of_real_def, - hypreal_minus, hypreal_add]) 1); -by (Blast_tac 1); -by (dtac spec 1 THEN dtac mp 1 THEN assume_tac 1); -by (dtac FreeUltrafilterNat_all 1); -by (Ultra_tac 1); -qed "NSLIM_LIM"; - - -(**** Key result ****) -Goal "(f -- x --> L) = (f -- x --NS> L)"; -by (blast_tac (claset() addIs [LIM_NSLIM,NSLIM_LIM]) 1); -qed "LIM_NSLIM_iff"; - -(*-------------------------------------------------------------------*) -(* Proving properties of limits using nonstandard definition and *) -(* hence, the properties hold for standard limits as well *) -(*-------------------------------------------------------------------*) -(*------------------------------------------------ - NSLIM_mult and hence (trivially) LIM_mult - ------------------------------------------------*) - -Goalw [NSLIM_def] - "[| f -- x --NS> l; g -- x --NS> m |] \ -\ ==> (%x. f(x) * g(x)) -- x --NS> (l * m)"; -by (auto_tac (claset() addSIs [approx_mult_HFinite], simpset())); -qed "NSLIM_mult"; - -Goal "[| f -- x --> l; g -- x --> m |] \ -\ ==> (%x. f(x) * g(x)) -- x --> (l * m)"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_mult]) 1); -qed "LIM_mult2"; - -(*---------------------------------------------- - NSLIM_add and hence (trivially) LIM_add - Note the much shorter proof - ----------------------------------------------*) -Goalw [NSLIM_def] - "[| f -- x --NS> l; g -- x --NS> m |] \ -\ ==> (%x. f(x) + g(x)) -- x --NS> (l + m)"; -by (auto_tac (claset() addSIs [approx_add], simpset())); -qed "NSLIM_add"; - -Goal "[| f -- x --> l; g -- x --> m |] \ -\ ==> (%x. f(x) + g(x)) -- x --> (l + m)"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_add]) 1); -qed "LIM_add2"; - -(*---------------------------------------------- - NSLIM_const - ----------------------------------------------*) -Goalw [NSLIM_def] "(%x. k) -- x --NS> k"; -by Auto_tac; -qed "NSLIM_const"; - -Addsimps [NSLIM_const]; - -Goal "(%x. k) -- x --> k"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1); -qed "LIM_const2"; - -(*---------------------------------------------- - NSLIM_minus - ----------------------------------------------*) -Goalw [NSLIM_def] - "f -- a --NS> L ==> (%x. -f(x)) -- a --NS> -L"; -by Auto_tac; -qed "NSLIM_minus"; - -Goal "f -- a --> L ==> (%x. -f(x)) -- a --> -L"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_minus]) 1); -qed "LIM_minus2"; - -(*---------------------------------------------- - NSLIM_add_minus - ----------------------------------------------*) -Goal "[| f -- x --NS> l; g -- x --NS> m |] \ -\ ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)"; -by (blast_tac (claset() addDs [NSLIM_add,NSLIM_minus]) 1); -qed "NSLIM_add_minus"; - -Goal "[| f -- x --> l; g -- x --> m |] \ -\ ==> (%x. f(x) + -g(x)) -- x --> (l + -m)"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, - NSLIM_add_minus]) 1); -qed "LIM_add_minus2"; - -(*----------------------------- - NSLIM_inverse - -----------------------------*) -Goalw [NSLIM_def] - "[| f -- a --NS> L; L \\ 0 |] \ -\ ==> (%x. inverse(f(x))) -- a --NS> (inverse L)"; -by (Clarify_tac 1); -by (dtac spec 1); -by (auto_tac (claset(), - simpset() addsimps [hypreal_of_real_approx_inverse])); -qed "NSLIM_inverse"; - -Goal "[| f -- a --> L; \ -\ L \\ 0 |] ==> (%x. inverse(f(x))) -- a --> (inverse L)"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_inverse]) 1); -qed "LIM_inverse"; - -(*------------------------------ - NSLIM_zero - ------------------------------*) -Goal "f -- a --NS> l ==> (%x. f(x) + -l) -- a --NS> 0"; -by (res_inst_tac [("a1","l")] ((right_minus RS subst)) 1); -by (rtac NSLIM_add_minus 1 THEN Auto_tac); -qed "NSLIM_zero"; - -Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> 0"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_zero]) 1); -qed "LIM_zero2"; - -Goal "(%x. f(x) - l) -- x --NS> 0 ==> f -- x --NS> l"; -by (dres_inst_tac [("g","%x. l"),("m","l")] NSLIM_add 1); -by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc])); -qed "NSLIM_zero_cancel"; - -Goal "(%x. f(x) - l) -- x --> 0 ==> f -- x --> l"; -by (dres_inst_tac [("g","%x. l"),("m","l")] LIM_add 1); -by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc])); -qed "LIM_zero_cancel"; - - -(*-------------------------- - NSLIM_not_zero - --------------------------*) -Goalw [NSLIM_def] "k \\ 0 ==> ~ ((%x. k) -- x --NS> 0)"; -by Auto_tac; -by (res_inst_tac [("x","hypreal_of_real x + epsilon")] exI 1); -by (auto_tac (claset() addIs [Infinitesimal_add_approx_self RS approx_sym], - simpset() addsimps [hypreal_epsilon_not_zero])); -qed "NSLIM_not_zero"; - -(* [| k \\ 0; (%x. k) -- x --NS> 0 |] ==> R *) -bind_thm("NSLIM_not_zeroE", NSLIM_not_zero RS notE); - -Goal "k \\ 0 ==> ~ ((%x. k) -- x --> 0)"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_not_zero]) 1); -qed "LIM_not_zero2"; - -(*------------------------------------- - NSLIM of constant function - -------------------------------------*) -Goal "(%x. k) -- x --NS> L ==> k = L"; -by (rtac ccontr 1); -by (dtac NSLIM_zero 1); -by (rtac NSLIM_not_zeroE 1 THEN assume_tac 2); -by (arith_tac 1); -qed "NSLIM_const_eq"; - -Goal "(%x. k) -- x --> L ==> k = L"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, - NSLIM_const_eq]) 1); -qed "LIM_const_eq2"; - -(*------------------------ - NS Limit is Unique - ------------------------*) -(* can actually be proved more easily by unfolding def! *) -Goal "[| f -- x --NS> L; f -- x --NS> M |] ==> L = M"; -by (dtac NSLIM_minus 1); -by (dtac NSLIM_add 1 THEN assume_tac 1); -by (auto_tac (claset() addSDs [NSLIM_const_eq RS sym], simpset())); -qed "NSLIM_unique"; - -Goal "[| f -- x --> L; f -- x --> M |] ==> L = M"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_unique]) 1); -qed "LIM_unique2"; - -(*-------------------- - NSLIM_mult_zero - --------------------*) -Goal "[| f -- x --NS> 0; g -- x --NS> 0 |] \ -\ ==> (%x. f(x)*g(x)) -- x --NS> 0"; -by (dtac NSLIM_mult 1 THEN Auto_tac); -qed "NSLIM_mult_zero"; - -(* we can use the corresponding thm LIM_mult2 *) -(* for standard definition of limit *) - -Goal "[| f -- x --> 0; g -- x --> 0 |] \ -\ ==> (%x. f(x)*g(x)) -- x --> 0"; -by (dtac LIM_mult2 1 THEN Auto_tac); -qed "LIM_mult_zero2"; - -(*---------------------------- - NSLIM_self - ----------------------------*) -Goalw [NSLIM_def] "(%x. x) -- a --NS> a"; -by (auto_tac (claset() addIs [starfun_Idfun_approx],simpset())); -qed "NSLIM_self"; - -Goal "(%x. x) -- a --> a"; -by (simp_tac (simpset() addsimps [LIM_NSLIM_iff,NSLIM_self]) 1); -qed "LIM_self2"; - -(*----------------------------------------------------------------------------- - Derivatives and Continuity - NS and Standard properties - -----------------------------------------------------------------------------*) -(*--------------- - Continuity - ---------------*) - -Goalw [isNSCont_def] - "[| isNSCont f a; y \\ hypreal_of_real a |] \ -\ ==> ( *f* f) y \\ hypreal_of_real (f a)"; -by (Blast_tac 1); -qed "isNSContD"; - -Goalw [isNSCont_def,NSLIM_def] - "isNSCont f a ==> f -- a --NS> (f a) "; -by (Blast_tac 1); -qed "isNSCont_NSLIM"; - -Goalw [isNSCont_def,NSLIM_def] - "f -- a --NS> (f a) ==> isNSCont f a"; -by Auto_tac; -by (res_inst_tac [("Q","y = hypreal_of_real a")] - (excluded_middle RS disjE) 1); -by Auto_tac; -qed "NSLIM_isNSCont"; - -(*----------------------------------------------------- - NS continuity can be defined using NS Limit in - similar fashion to standard def of continuity - -----------------------------------------------------*) -Goal "(isNSCont f a) = (f -- a --NS> (f a))"; -by (blast_tac (claset() addIs [isNSCont_NSLIM,NSLIM_isNSCont]) 1); -qed "isNSCont_NSLIM_iff"; - -(*---------------------------------------------- - Hence, NS continuity can be given - in terms of standard limit - ---------------------------------------------*) -Goal "(isNSCont f a) = (f -- a --> (f a))"; -by (asm_full_simp_tac (simpset() addsimps - [LIM_NSLIM_iff,isNSCont_NSLIM_iff]) 1); -qed "isNSCont_LIM_iff"; - -(*----------------------------------------------- - Moreover, it's trivial now that NS continuity - is equivalent to standard continuity - -----------------------------------------------*) -Goalw [isCont_def] "(isNSCont f a) = (isCont f a)"; -by (rtac isNSCont_LIM_iff 1); -qed "isNSCont_isCont_iff"; - -(*---------------------------------------- - Standard continuity ==> NS continuity - ----------------------------------------*) -Goal "isCont f a ==> isNSCont f a"; -by (etac (isNSCont_isCont_iff RS iffD2) 1); -qed "isCont_isNSCont"; - -(*---------------------------------------- - NS continuity ==> Standard continuity - ----------------------------------------*) -Goal "isNSCont f a ==> isCont f a"; -by (etac (isNSCont_isCont_iff RS iffD1) 1); -qed "isNSCont_isCont"; - -(*-------------------------------------------------------------------------- - Alternative definition of continuity - --------------------------------------------------------------------------*) -(* Prove equivalence between NS limits - *) -(* seems easier than using standard def *) -Goalw [NSLIM_def] "(f -- a --NS> L) = ((%h. f(a + h)) -- 0 --NS> L)"; -by Auto_tac; -by (dres_inst_tac [("x","hypreal_of_real a + x")] spec 1); -by (dres_inst_tac [("x","-hypreal_of_real a + x")] spec 2); -by Safe_tac; -by (Asm_full_simp_tac 1); -by (rtac ((mem_infmal_iff RS iffD2) RS - (Infinitesimal_add_approx_self RS approx_sym)) 1); -by (rtac (approx_minus_iff2 RS iffD1) 4); -by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 3); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 2); -by (res_inst_tac [("z","x")] eq_Abs_hypreal 4); -by (auto_tac (claset(), - simpset() addsimps [starfun, hypreal_of_real_def, hypreal_minus, - hypreal_add, real_add_assoc, approx_refl, hypreal_zero_def])); -qed "NSLIM_h_iff"; - -Goal "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)"; -by (rtac NSLIM_h_iff 1); -qed "NSLIM_isCont_iff"; - -Goal "(f -- a --> f a) = ((%h. f(a + h)) -- 0 --> f(a))"; -by (simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_isCont_iff]) 1); -qed "LIM_isCont_iff"; - -Goalw [isCont_def] "(isCont f x) = ((%h. f(x + h)) -- 0 --> f(x))"; -by (simp_tac (simpset() addsimps [LIM_isCont_iff]) 1); -qed "isCont_iff"; - -(*-------------------------------------------------------------------------- - Immediate application of nonstandard criterion for continuity can offer - very simple proofs of some standard property of continuous functions - --------------------------------------------------------------------------*) -(*------------------------ - sum continuous - ------------------------*) -Goal "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) + g(x)) a"; -by (auto_tac (claset() addIs [approx_add], - simpset() addsimps [isNSCont_isCont_iff RS sym, isNSCont_def])); -qed "isCont_add"; - -(*------------------------ - mult continuous - ------------------------*) -Goal "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) * g(x)) a"; -by (auto_tac (claset() addSIs [starfun_mult_HFinite_approx], - simpset() delsimps [starfun_mult RS sym] - addsimps [isNSCont_isCont_iff RS sym, isNSCont_def])); -qed "isCont_mult"; - -(*------------------------------------------- - composition of continuous functions - Note very short straightforard proof! - ------------------------------------------*) -Goal "[| isCont f a; isCont g (f a) |] \ -\ ==> isCont (g o f) a"; -by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym, - isNSCont_def,starfun_o RS sym])); -qed "isCont_o"; - -Goal "[| isCont f a; isCont g (f a) |] \ -\ ==> isCont (%x. g (f x)) a"; -by (auto_tac (claset() addDs [isCont_o],simpset() addsimps [o_def])); -qed "isCont_o2"; - -Goalw [isNSCont_def] "isNSCont f a ==> isNSCont (%x. - f x) a"; -by Auto_tac; -qed "isNSCont_minus"; - -Goal "isCont f a ==> isCont (%x. - f x) a"; -by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym, - isNSCont_minus])); -qed "isCont_minus"; - -Goalw [isCont_def] - "[| isCont f x; f x \\ 0 |] ==> isCont (%x. inverse (f x)) x"; -by (blast_tac (claset() addIs [LIM_inverse]) 1); -qed "isCont_inverse"; - -Goal "[| isNSCont f x; f x \\ 0 |] ==> isNSCont (%x. inverse (f x)) x"; -by (auto_tac (claset() addIs [isCont_inverse],simpset() addsimps - [isNSCont_isCont_iff])); -qed "isNSCont_inverse"; - -Goalw [real_diff_def] - "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) - g(x)) a"; -by (auto_tac (claset() addIs [isCont_add,isCont_minus],simpset())); -qed "isCont_diff"; - -Goalw [isCont_def] "isCont (%x. k) a"; -by (Simp_tac 1); -qed "isCont_const"; -Addsimps [isCont_const]; - -Goalw [isNSCont_def] "isNSCont (%x. k) a"; -by (Simp_tac 1); -qed "isNSCont_const"; -Addsimps [isNSCont_const]; - -Goalw [isNSCont_def] "isNSCont abs a"; -by (auto_tac (claset() addIs [approx_hrabs], - simpset() addsimps [hypreal_of_real_hrabs RS sym, - starfun_rabs_hrabs])); -qed "isNSCont_rabs"; -Addsimps [isNSCont_rabs]; - -Goal "isCont abs a"; -by (auto_tac (claset(), simpset() addsimps [isNSCont_isCont_iff RS sym])); -qed "isCont_rabs"; -Addsimps [isCont_rabs]; - -(**************************************************************** -(%* Leave as commented until I add topology theory or remove? *%) -(%*------------------------------------------------------------ - Elementary topology proof for a characterisation of - continuity now: a function f is continuous if and only - if the inverse image, {x. f(x) \\ A}, of any open set A - is always an open set - ------------------------------------------------------------*%) -Goal "[| isNSopen A; \\x. isNSCont f x |] \ -\ ==> isNSopen {x. f x \\ A}"; -by (auto_tac (claset(),simpset() addsimps [isNSopen_iff1])); -by (dtac (mem_monad_approx RS approx_sym) 1); -by (dres_inst_tac [("x","a")] spec 1); -by (dtac isNSContD 1 THEN assume_tac 1); -by (dtac bspec 1 THEN assume_tac 1); -by (dres_inst_tac [("x","( *f* f) x")] approx_mem_monad2 1); -by (blast_tac (claset() addIs [starfun_mem_starset]) 1); -qed "isNSCont_isNSopen"; - -Goalw [isNSCont_def] - "\\A. isNSopen A --> isNSopen {x. f x \\ A} \ -\ ==> isNSCont f x"; -by (auto_tac (claset() addSIs [(mem_infmal_iff RS iffD1) RS - (approx_minus_iff RS iffD2)],simpset() addsimps - [Infinitesimal_def,SReal_iff])); -by (dres_inst_tac [("x","{z. abs(z + -f(x)) < ya}")] spec 1); -by (etac (isNSopen_open_interval RSN (2,impE)) 1); -by (auto_tac (claset(),simpset() addsimps [isNSopen_def,isNSnbhd_def])); -by (dres_inst_tac [("x","x")] spec 1); -by (auto_tac (claset() addDs [approx_sym RS approx_mem_monad], - simpset() addsimps [hypreal_of_real_zero RS sym,STAR_starfun_rabs_add_minus])); -qed "isNSopen_isNSCont"; - -Goal "(\\x. isNSCont f x) = \ -\ (\\A. isNSopen A --> isNSopen {x. f(x) \\ A})"; -by (blast_tac (claset() addIs [isNSCont_isNSopen, - isNSopen_isNSCont]) 1); -qed "isNSCont_isNSopen_iff"; - -(%*------- Standard version of same theorem --------*%) -Goal "(\\x. isCont f x) = \ -\ (\\A. isopen A --> isopen {x. f(x) \\ A})"; -by (auto_tac (claset() addSIs [isNSCont_isNSopen_iff], - simpset() addsimps [isNSopen_isopen_iff RS sym, - isNSCont_isCont_iff RS sym])); -qed "isCont_isopen_iff"; -*******************************************************************) - -(*----------------------------------------------------------------- - Uniform continuity - ------------------------------------------------------------------*) -Goalw [isNSUCont_def] - "[| isNSUCont f; x \\ y|] ==> ( *f* f) x \\ ( *f* f) y"; -by (Blast_tac 1); -qed "isNSUContD"; - -Goalw [isUCont_def,isCont_def,LIM_def] - "isUCont f ==> isCont f x"; -by (Clarify_tac 1); -by (dtac spec 1); -by (Blast_tac 1); -qed "isUCont_isCont"; - -Goalw [isNSUCont_def,isUCont_def,approx_def] - "isUCont f ==> isNSUCont f"; -by (asm_full_simp_tac (simpset() addsimps - [Infinitesimal_FreeUltrafilterNat_iff]) 1); -by Safe_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(),simpset() addsimps [starfun, - hypreal_minus, hypreal_add])); -by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1); -by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1); -by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1); -by (subgoal_tac "\\n::nat. abs ((xa n) + - (xb n)) < s --> abs (f (xa n) + - f (xb n)) < u" 1); -by (Blast_tac 2); -by (thin_tac "\\x y. abs (x + - y) < s --> abs (f x + - f y) < u" 1); -by (dtac FreeUltrafilterNat_all 1); -by (Ultra_tac 1); -qed "isUCont_isNSUCont"; - -Goal "\\s. 0 < s --> (\\z y. abs (z + - y) < s & r \\ abs (f z + -f y)) \ -\ ==> \\n::nat. \\z y. \ -\ abs(z + -y) < inverse(real(Suc n)) & \ -\ r \\ abs(f z + -f y)"; -by (Clarify_tac 1); -by (cut_inst_tac [("n1","n")] - (real_of_nat_Suc_gt_zero RS positive_imp_inverse_positive) 1); -by Auto_tac; -qed "lemma_LIMu"; - -Goal "\\s. 0 < s --> (\\z y. abs (z + - y) < s & r \\ abs (f z + -f y)) \ -\ ==> \\X Y. \\n::nat. \ -\ abs(X n + -(Y n)) < inverse(real(Suc n)) & \ -\ r \\ abs(f (X n) + -f (Y n))"; -by (dtac lemma_LIMu 1); -by (dtac choice 1); -by Safe_tac; -by (dtac choice 1); -by (Blast_tac 1); -qed "lemma_skolemize_LIM2u"; - -Goal "\\n. abs (X n + -Y n) < inverse (real(Suc n)) & \ -\ r \\ abs (f (X n) + - f(Y n)) ==> \ -\ \\n. abs (X n + - Y n) < inverse (real(Suc n))"; -by (Auto_tac ); -qed "lemma_simpu"; - -Goalw [isNSUCont_def,isUCont_def,approx_def] - "isNSUCont f ==> isUCont f"; -by (asm_full_simp_tac (simpset() addsimps - [Infinitesimal_FreeUltrafilterNat_iff]) 1); -by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]); -by (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1); -by (dtac lemma_skolemize_LIM2u 1); -by Safe_tac; -by (dres_inst_tac [("x","Abs_hypreal(hyprel``{X})")] spec 1); -by (dres_inst_tac [("x","Abs_hypreal(hyprel``{Y})")] spec 1); -by (asm_full_simp_tac - (simpset() addsimps [starfun, hypreal_minus,hypreal_add]) 1); -by Auto_tac; -by (dtac (lemma_simpu RS real_seq_to_hypreal_Infinitesimal2) 1); -by (asm_full_simp_tac (simpset() addsimps - [Infinitesimal_FreeUltrafilterNat_iff, hypreal_minus,hypreal_add]) 1); -by (Blast_tac 1); -by (rotate_tac 2 1); -by (dres_inst_tac [("x","r")] spec 1); -by (Clarify_tac 1); -by (dtac FreeUltrafilterNat_all 1); -by (Ultra_tac 1); -qed "isNSUCont_isUCont"; - -(*------------------------------------------------------------------ - Derivatives - ------------------------------------------------------------------*) -Goalw [deriv_def] - "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --> D)"; -by (Blast_tac 1); -qed "DERIV_iff"; - -Goalw [deriv_def] - "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --NS> D)"; -by (simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1); -qed "DERIV_NS_iff"; - -Goalw [deriv_def] - "DERIV f x :> D \ -\ ==> (%h. (f(x + h) + - f(x))/h) -- 0 --> D"; -by (Blast_tac 1); -qed "DERIVD"; - -Goalw [deriv_def] "DERIV f x :> D ==> \ -\ (%h. (f(x + h) + - f(x))/h) -- 0 --NS> D"; -by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1); -qed "NS_DERIVD"; - -(* Uniqueness *) -Goalw [deriv_def] - "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E"; -by (blast_tac (claset() addIs [LIM_unique]) 1); -qed "DERIV_unique"; - -Goalw [nsderiv_def] - "[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E"; -by (cut_facts_tac [Infinitesimal_epsilon, hypreal_epsilon_not_zero] 1); -by (auto_tac (claset() addSDs [inst "x" "epsilon" bspec] - addSIs [inj_hypreal_of_real RS injD] - addDs [approx_trans3], - simpset())); -qed "NSDeriv_unique"; - -(*------------------------------------------------------------------------ - Differentiable - ------------------------------------------------------------------------*) - -Goalw [differentiable_def] - "f differentiable x ==> \\D. DERIV f x :> D"; -by (assume_tac 1); -qed "differentiableD"; - -Goalw [differentiable_def] - "DERIV f x :> D ==> f differentiable x"; -by (Blast_tac 1); -qed "differentiableI"; - -Goalw [NSdifferentiable_def] - "f NSdifferentiable x ==> \\D. NSDERIV f x :> D"; -by (assume_tac 1); -qed "NSdifferentiableD"; - -Goalw [NSdifferentiable_def] - "NSDERIV f x :> D ==> f NSdifferentiable x"; -by (Blast_tac 1); -qed "NSdifferentiableI"; - -(*-------------------------------------------------------- - Alternative definition for differentiability - -------------------------------------------------------*) - -Goalw [LIM_def] - "((%h. (f(a + h) + - f(a))/h) -- 0 --> D) = \ -\ ((%x. (f(x) + -f(a)) / (x + -a)) -- a --> D)"; -by Safe_tac; -by (ALLGOALS(dtac spec)); -by Safe_tac; -by (Blast_tac 1 THEN Blast_tac 2); -by (ALLGOALS(res_inst_tac [("x","s")] exI)); -by Safe_tac; -by (dres_inst_tac [("x","x + -a")] spec 1); -by (dres_inst_tac [("x","x + a")] spec 2); -by (auto_tac (claset(), simpset() addsimps add_ac)); -qed "DERIV_LIM_iff"; - -Goalw [deriv_def] "(DERIV f x :> D) = \ -\ ((%z. (f(z) + -f(x)) / (z + -x)) -- x --> D)"; -by (simp_tac (simpset() addsimps [DERIV_LIM_iff]) 1); -qed "DERIV_iff2"; - -(*-------------------------------------------------------- - Equivalence of NS and standard defs of differentiation - -------------------------------------------------------*) -(*------------------------------------------- - First NSDERIV in terms of NSLIM - -------------------------------------------*) - -(*--- first equivalence ---*) -Goalw [nsderiv_def,NSLIM_def] - "(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --NS> D)"; -by Auto_tac; -by (dres_inst_tac [("x","xa")] bspec 1); -by (rtac ccontr 3); -by (dres_inst_tac [("x","h")] spec 3); -by (auto_tac (claset(), - simpset() addsimps [mem_infmal_iff, starfun_lambda_cancel])); -qed "NSDERIV_NSLIM_iff"; - -(*--- second equivalence ---*) -Goal "(NSDERIV f x :> D) = \ -\ ((%z. (f(z) + -f(x)) / (z + -x)) -- x --NS> D)"; -by (full_simp_tac (simpset() addsimps - [NSDERIV_NSLIM_iff, DERIV_LIM_iff, LIM_NSLIM_iff RS sym]) 1); -qed "NSDERIV_NSLIM_iff2"; - -(* while we're at it! *) -Goalw [real_diff_def] - "(NSDERIV f x :> D) = \ -\ (\\xa. \ -\ xa \\ hypreal_of_real x & xa \\ hypreal_of_real x --> \ -\ ( *f* (%z. (f z - f x) / (z - x))) xa \\ hypreal_of_real D)"; -by (auto_tac (claset(), simpset() addsimps [NSDERIV_NSLIM_iff2, NSLIM_def])); -qed "NSDERIV_iff2"; - - -Goal "(NSDERIV f x :> D) ==> \ -\ (\\u. \ -\ u \\ hypreal_of_real x --> \ -\ ( *f* (%z. f z - f x)) u \\ hypreal_of_real D * (u - hypreal_of_real x))"; -by (auto_tac (claset(), simpset() addsimps [NSDERIV_iff2])); -by (case_tac "u = hypreal_of_real x" 1); -by (auto_tac (claset(), simpset() addsimps [hypreal_diff_def])); -by (dres_inst_tac [("x","u")] spec 1); -by Auto_tac; -by (dres_inst_tac [("c","u - hypreal_of_real x"),("b","hypreal_of_real D")] - approx_mult1 1); -by (ALLGOALS(dtac (hypreal_not_eq_minus_iff RS iffD1))); -by (subgoal_tac "( *f* (%z. z - x)) u \\ (0::hypreal)" 2); -by (auto_tac (claset(), - simpset() addsimps [real_diff_def, hypreal_diff_def, - (approx_minus_iff RS iffD1) RS (mem_infmal_iff RS iffD2), - Infinitesimal_subset_HFinite RS subsetD])); -qed "NSDERIVD5"; - -Goal "(NSDERIV f x :> D) ==> \ -\ (\\h \\ Infinitesimal. \ -\ (( *f* f)(hypreal_of_real x + h) - \ -\ hypreal_of_real (f x))\\ (hypreal_of_real D) * h)"; -by (auto_tac (claset(),simpset() addsimps [nsderiv_def])); -by (case_tac "h = (0::hypreal)" 1); -by (auto_tac (claset(),simpset() addsimps [hypreal_diff_def])); -by (dres_inst_tac [("x","h")] bspec 1); -by (dres_inst_tac [("c","h")] approx_mult1 2); -by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD], - simpset() addsimps [hypreal_diff_def])); -qed "NSDERIVD4"; - -Goal "(NSDERIV f x :> D) ==> \ -\ (\\h \\ Infinitesimal - {0}. \ -\ (( *f* f)(hypreal_of_real x + h) - \ -\ hypreal_of_real (f x))\\ (hypreal_of_real D) * h)"; -by (auto_tac (claset(),simpset() addsimps [nsderiv_def])); -by (rtac ccontr 1 THEN dres_inst_tac [("x","h")] bspec 1); -by (dres_inst_tac [("c","h")] approx_mult1 2); -by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD], - simpset() addsimps [hypreal_mult_assoc, hypreal_diff_def])); -qed "NSDERIVD3"; - -(*-------------------------------------------------------------- - Now equivalence between NSDERIV and DERIV - -------------------------------------------------------------*) -Goalw [deriv_def] "(NSDERIV f x :> D) = (DERIV f x :> D)"; -by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,LIM_NSLIM_iff]) 1); -qed "NSDERIV_DERIV_iff"; - -(*--------------------------------------------------- - Differentiability implies continuity - nice and simple "algebraic" proof - --------------------------------------------------*) -Goalw [nsderiv_def] - "NSDERIV f x :> D ==> isNSCont f x"; -by (auto_tac (claset(),simpset() addsimps - [isNSCont_NSLIM_iff,NSLIM_def])); -by (dtac (approx_minus_iff RS iffD1) 1); -by (dtac (hypreal_not_eq_minus_iff RS iffD1) 1); -by (dres_inst_tac [("x","-hypreal_of_real x + xa")] bspec 1); -by (asm_full_simp_tac (simpset() addsimps - [hypreal_add_assoc RS sym]) 2); -by (auto_tac (claset(),simpset() addsimps - [mem_infmal_iff RS sym,hypreal_add_commute])); -by (dres_inst_tac [("c","xa + -hypreal_of_real x")] approx_mult1 1); -by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite - RS subsetD],simpset() addsimps [hypreal_mult_assoc])); -by (dres_inst_tac [("x3","D")] (HFinite_hypreal_of_real RSN - (2,Infinitesimal_HFinite_mult) RS (mem_infmal_iff RS iffD1)) 1); -by (blast_tac (claset() addIs [approx_trans, - hypreal_mult_commute RS subst, - (approx_minus_iff RS iffD2)]) 1); -qed "NSDERIV_isNSCont"; - -(* Now Sandard proof *) -Goal "DERIV f x :> D ==> isCont f x"; -by (asm_full_simp_tac (simpset() addsimps - [NSDERIV_DERIV_iff RS sym, isNSCont_isCont_iff RS sym, - NSDERIV_isNSCont]) 1); -qed "DERIV_isCont"; - -(*---------------------------------------------------------------------------- - Differentiation rules for combinations of functions - follow from clear, straightforard, algebraic - manipulations - ----------------------------------------------------------------------------*) -(*------------------------- - Constant function - ------------------------*) - -(* use simple constant nslimit theorem *) -Goal "(NSDERIV (%x. k) x :> 0)"; -by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1); -qed "NSDERIV_const"; -Addsimps [NSDERIV_const]; - -Goal "(DERIV (%x. k) x :> 0)"; -by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1); -qed "DERIV_const"; -Addsimps [DERIV_const]; - -(*----------------------------------------------------- - Sum of functions- proved easily - ----------------------------------------------------*) - - -Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \ -\ ==> NSDERIV (%x. f x + g x) x :> Da + Db"; -by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff, - NSLIM_def]) 1 THEN REPEAT (Step_tac 1)); -by (auto_tac (claset(), - simpset() addsimps [add_divide_distrib])); -by (dres_inst_tac [("b","hypreal_of_real Da"), - ("d","hypreal_of_real Db")] approx_add 1); -by (auto_tac (claset(), simpset() addsimps add_ac)); -qed "NSDERIV_add"; - -(* Standard theorem *) -Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \ -\ ==> DERIV (%x. f x + g x) x :> Da + Db"; -by (asm_full_simp_tac (simpset() addsimps [NSDERIV_add, - NSDERIV_DERIV_iff RS sym]) 1); -qed "DERIV_add"; - -(*----------------------------------------------------- - Product of functions - Proof is trivial but tedious - and long due to rearrangement of terms - ----------------------------------------------------*) - -Goal "((a::hypreal)*b) + -(c*d) = (b*(a + -c)) + (c*(b + -d))"; -by (simp_tac (simpset() addsimps [right_distrib]) 1); -qed "lemma_nsderiv1"; - -Goal "[| (x + y) / z = hypreal_of_real D + yb; z \\ 0; \ -\ z \\ Infinitesimal; yb \\ Infinitesimal |] \ -\ ==> x + y \\ 0"; -by (forw_inst_tac [("c1","z")] (hypreal_mult_right_cancel RS iffD2) 1 - THEN assume_tac 1); -by (thin_tac "(x + y) / z = hypreal_of_real D + yb" 1); -by (auto_tac (claset() addSIs [Infinitesimal_HFinite_mult2, HFinite_add], - simpset() addsimps [hypreal_mult_assoc, mem_infmal_iff RS sym])); -by (etac (Infinitesimal_subset_HFinite RS subsetD) 1); -qed "lemma_nsderiv2"; - - -Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \ -\ ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))"; -by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def]) 1); -by (REPEAT (Step_tac 1)); -by (auto_tac (claset(), - simpset() addsimps [starfun_lambda_cancel, lemma_nsderiv1])); -by (simp_tac (simpset() addsimps [add_divide_distrib]) 1); -by (REPEAT(dtac (bex_Infinitesimal_iff2 RS iffD2) 1)); -by (auto_tac (claset(), - simpset() delsimps [times_divide_eq_right] - addsimps [times_divide_eq_right RS sym])); -by (dres_inst_tac [("D","Db")] lemma_nsderiv2 1); -by (dtac (approx_minus_iff RS iffD2 RS (bex_Infinitesimal_iff2 RS iffD2)) 4); -by (auto_tac (claset() addSIs [approx_add_mono1], - simpset() addsimps [left_distrib, right_distrib, - hypreal_mult_commute, hypreal_add_assoc])); -by (res_inst_tac [("w1","hypreal_of_real Db * hypreal_of_real (f x)")] - (hypreal_add_commute RS subst) 1); -by (auto_tac (claset() addSIs [Infinitesimal_add_approx_self2 RS approx_sym, - Infinitesimal_add, Infinitesimal_mult, - Infinitesimal_hypreal_of_real_mult, - Infinitesimal_hypreal_of_real_mult2], - simpset() addsimps [hypreal_add_assoc RS sym])); -qed "NSDERIV_mult"; - -Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \ -\ ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))"; -by (asm_full_simp_tac (simpset() addsimps [NSDERIV_mult, - NSDERIV_DERIV_iff RS sym]) 1); -qed "DERIV_mult"; - -(*---------------------------- - Multiplying by a constant - ---------------------------*) -Goal "NSDERIV f x :> D \ -\ ==> NSDERIV (%x. c * f x) x :> c*D"; -by (asm_full_simp_tac - (HOL_ss addsimps [times_divide_eq_right RS sym, NSDERIV_NSLIM_iff, - minus_mult_right, right_distrib RS sym]) 1); -by (etac (NSLIM_const RS NSLIM_mult) 1); -qed "NSDERIV_cmult"; - -(* let's do the standard proof though theorem *) -(* LIM_mult2 follows from a NS proof *) - -Goalw [deriv_def] - "DERIV f x :> D \ -\ ==> DERIV (%x. c * f x) x :> c*D"; -by (asm_full_simp_tac - (HOL_ss addsimps [times_divide_eq_right RS sym, NSDERIV_NSLIM_iff, - minus_mult_right, right_distrib RS sym]) 1); -by (etac (LIM_const RS LIM_mult2) 1); -qed "DERIV_cmult"; - -(*-------------------------------- - Negation of function - -------------------------------*) -Goal "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D"; -by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1); -by (dtac NSLIM_minus 1); -by (subgoal_tac "ALL a::real. ALL b. - a + b = - (a + - b)" 1); -by (asm_full_simp_tac (HOL_ss addsimps [thm"minus_divide_left" RS sym]) 1); -by (Asm_full_simp_tac 1); -qed "NSDERIV_minus"; - -Goal "DERIV f x :> D \ -\ ==> DERIV (%x. -(f x)) x :> -D"; -by (asm_full_simp_tac (simpset() addsimps - [NSDERIV_minus,NSDERIV_DERIV_iff RS sym]) 1); -qed "DERIV_minus"; - -(*------------------------------- - Subtraction - ------------------------------*) -Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \ -\ ==> NSDERIV (%x. f x + -g x) x :> Da + -Db"; -by (blast_tac (claset() addDs [NSDERIV_add,NSDERIV_minus]) 1); -qed "NSDERIV_add_minus"; - -Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \ -\ ==> DERIV (%x. f x + -g x) x :> Da + -Db"; -by (blast_tac (claset() addDs [DERIV_add,DERIV_minus]) 1); -qed "DERIV_add_minus"; - -Goalw [real_diff_def] - "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \ -\ ==> NSDERIV (%x. f x - g x) x :> Da - Db"; -by (blast_tac (claset() addIs [NSDERIV_add_minus]) 1); -qed "NSDERIV_diff"; - -Goalw [real_diff_def] - "[| DERIV f x :> Da; DERIV g x :> Db |] \ -\ ==> DERIV (%x. f x - g x) x :> Da - Db"; -by (blast_tac (claset() addIs [DERIV_add_minus]) 1); -qed "DERIV_diff"; - -(*--------------------------------------------------------------- - (NS) Increment - ---------------------------------------------------------------*) -Goalw [increment_def] - "f NSdifferentiable x ==> \ -\ increment f x h = ( *f* f) (hypreal_of_real(x) + h) + \ -\ -hypreal_of_real (f x)"; -by (Blast_tac 1); -qed "incrementI"; - -Goal "NSDERIV f x :> D ==> \ -\ increment f x h = ( *f* f) (hypreal_of_real(x) + h) + \ -\ -hypreal_of_real (f x)"; -by (etac (NSdifferentiableI RS incrementI) 1); -qed "incrementI2"; - -(* The Increment theorem -- Keisler p. 65 *) -Goal "[| NSDERIV f x :> D; h \\ Infinitesimal; h \\ 0 |] \ -\ ==> \\e \\ Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h"; -by (forw_inst_tac [("h","h")] incrementI2 1 THEN rewtac nsderiv_def); -by (dtac bspec 1 THEN Auto_tac); -by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1 THEN Step_tac 1); -by (forw_inst_tac [("b1","hypreal_of_real(D) + y")] - ((hypreal_mult_right_cancel RS iffD2)) 1); -by (thin_tac "(( *f* f) (hypreal_of_real(x) + h) + \ -\ - hypreal_of_real (f x)) / h = hypreal_of_real(D) + y" 2); -by (assume_tac 1); -by (asm_full_simp_tac (simpset() addsimps [times_divide_eq_right RS sym] - delsimps [times_divide_eq_right]) 1); -by (auto_tac (claset(), - simpset() addsimps [left_distrib])); -qed "increment_thm"; - -Goal "[| NSDERIV f x :> D; h \\ 0; h \\ 0 |] \ -\ ==> \\e \\ Infinitesimal. increment f x h = \ -\ hypreal_of_real(D)*h + e*h"; -by (blast_tac (claset() addSDs [mem_infmal_iff RS iffD2] - addSIs [increment_thm]) 1); -qed "increment_thm2"; - -Goal "[| NSDERIV f x :> D; h \\ 0; h \\ 0 |] \ -\ ==> increment f x h \\ 0"; -by (dtac increment_thm2 1 THEN auto_tac (claset() addSIs - [Infinitesimal_HFinite_mult2,HFinite_add],simpset() addsimps - [left_distrib RS sym,mem_infmal_iff RS sym])); -by (etac (Infinitesimal_subset_HFinite RS subsetD) 1); -qed "increment_approx_zero"; - -(*--------------------------------------------------------------- - Similarly to the above, the chain rule admits an entirely - straightforward derivation. Compare this with Harrison's - HOL proof of the chain rule, which proved to be trickier and - required an alternative characterisation of differentiability- - the so-called Carathedory derivative. Our main problem is - manipulation of terms. - --------------------------------------------------------------*) - -(* lemmas *) -Goalw [nsderiv_def] - "[| NSDERIV g x :> D; \ -\ ( *f* g) (hypreal_of_real(x) + xa) = hypreal_of_real(g x);\ -\ xa \\ Infinitesimal;\ -\ xa \\ 0 \ -\ |] ==> D = 0"; -by (dtac bspec 1); -by Auto_tac; -qed "NSDERIV_zero"; - -(* can be proved differently using NSLIM_isCont_iff *) -Goalw [nsderiv_def] - "[| NSDERIV f x :> D; h \\ Infinitesimal; h \\ 0 |] \ -\ ==> ( *f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) \\ 0"; -by (asm_full_simp_tac (simpset() addsimps - [mem_infmal_iff RS sym]) 1); -by (rtac Infinitesimal_ratio 1); -by (rtac approx_hypreal_of_real_HFinite 3); -by Auto_tac; -qed "NSDERIV_approx"; - -(*--------------------------------------------------------------- - from one version of differentiability - - f(x) - f(a) - --------------- \\ Db - x - a - ---------------------------------------------------------------*) -Goal "[| NSDERIV f (g x) :> Da; \ -\ ( *f* g) (hypreal_of_real(x) + xa) \\ hypreal_of_real (g x); \ -\ ( *f* g) (hypreal_of_real(x) + xa) \\ hypreal_of_real (g x) \ -\ |] ==> (( *f* f) (( *f* g) (hypreal_of_real(x) + xa)) \ -\ + - hypreal_of_real (f (g x))) \ -\ / (( *f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real (g x)) \ -\ \\ hypreal_of_real(Da)"; -by (auto_tac (claset(), - simpset() addsimps [NSDERIV_NSLIM_iff2, NSLIM_def])); -qed "NSDERIVD1"; - -(*-------------------------------------------------------------- - from other version of differentiability - - f(x + h) - f(x) - ----------------- \\ Db - h - --------------------------------------------------------------*) -Goal "[| NSDERIV g x :> Db; xa \\ Infinitesimal; xa \\ 0 |] \ -\ ==> (( *f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real(g x)) / xa \ -\ \\ hypreal_of_real(Db)"; -by (auto_tac (claset(), - simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def, - mem_infmal_iff, starfun_lambda_cancel])); -qed "NSDERIVD2"; - -Goal "(z::hypreal) \\ 0 ==> x*y = (x*inverse(z))*(z*y)"; -by Auto_tac; -qed "lemma_chain"; - -(*------------------------------------------------------ - This proof uses both definitions of differentiability. - ------------------------------------------------------*) -Goal "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] \ -\ ==> NSDERIV (f o g) x :> Da * Db"; -by (asm_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff, - NSLIM_def,mem_infmal_iff RS sym]) 1 THEN Step_tac 1); -by (forw_inst_tac [("f","g")] NSDERIV_approx 1); -by (auto_tac (claset(), - simpset() addsimps [starfun_lambda_cancel2, starfun_o RS sym])); -by (case_tac "( *f* g) (hypreal_of_real(x) + xa) = hypreal_of_real (g x)" 1); -by (dres_inst_tac [("g","g")] NSDERIV_zero 1); -by (auto_tac (claset(), simpset() addsimps [hypreal_divide_def])); -by (res_inst_tac [("z1","( *f* g) (hypreal_of_real(x) + xa) + -hypreal_of_real (g x)"), - ("y1","inverse xa")] (lemma_chain RS ssubst) 1); -by (etac (hypreal_not_eq_minus_iff RS iffD1) 1); -by (rtac approx_mult_hypreal_of_real 1); -by (fold_tac [hypreal_divide_def]); -by (blast_tac (claset() addIs [NSDERIVD1, - approx_minus_iff RS iffD2]) 1); -by (blast_tac (claset() addIs [NSDERIVD2]) 1); -qed "NSDERIV_chain"; - -(* standard version *) -Goal "[| DERIV f (g x) :> Da; \ -\ DERIV g x :> Db \ -\ |] ==> DERIV (f o g) x :> Da * Db"; -by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym, - NSDERIV_chain]) 1); -qed "DERIV_chain"; - -Goal "[| DERIV f (g x) :> Da; DERIV g x :> Db |] \ -\ ==> DERIV (%x. f (g x)) x :> Da * Db"; -by (auto_tac (claset() addDs [DERIV_chain], simpset() addsimps [o_def])); -qed "DERIV_chain2"; - -(*------------------------------------------------------------------ - Differentiation of natural number powers - ------------------------------------------------------------------*) -Goal "NSDERIV (%x. x) x :> 1"; -by (auto_tac (claset(), - simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def ,starfun_Id])); -qed "NSDERIV_Id"; -Addsimps [NSDERIV_Id]; - -(*derivative of the identity function*) -Goal "DERIV (%x. x) x :> 1"; -by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1); -qed "DERIV_Id"; -Addsimps [DERIV_Id]; - -bind_thm ("isCont_Id", DERIV_Id RS DERIV_isCont); - -(*derivative of linear multiplication*) -Goal "DERIV (op * c) x :> c"; -by (cut_inst_tac [("c","c"),("x","x")] (DERIV_Id RS DERIV_cmult) 1); -by (Asm_full_simp_tac 1); -qed "DERIV_cmult_Id"; -Addsimps [DERIV_cmult_Id]; - -Goal "NSDERIV (op * c) x :> c"; -by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff]) 1); -qed "NSDERIV_cmult_Id"; -Addsimps [NSDERIV_cmult_Id]; - -Goal "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))"; -by (induct_tac "n" 1); -by (dtac (DERIV_Id RS DERIV_mult) 2); -by (auto_tac (claset(), - simpset() addsimps [real_of_nat_Suc, left_distrib])); -by (case_tac "0 < n" 1); -by (dres_inst_tac [("x","x")] realpow_minus_mult 1); -by (auto_tac (claset(), - simpset() addsimps [real_mult_assoc, real_add_commute])); -qed "DERIV_pow"; - -(* NS version *) -Goal "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))"; -by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff, DERIV_pow]) 1); -qed "NSDERIV_pow"; - -(*--------------------------------------------------------------- - Power of -1 - ---------------------------------------------------------------*) - -(*Can't get rid of x \\ 0 because it isn't continuous at zero*) -Goalw [nsderiv_def] - "x \\ 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))"; -by (rtac ballI 1 THEN Asm_full_simp_tac 1 THEN Step_tac 1); -by (ftac Infinitesimal_add_not_zero 1); -by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 2); -by (auto_tac (claset(), - simpset() addsimps [starfun_inverse_inverse, realpow_two] - delsimps [minus_mult_left RS sym, - minus_mult_right RS sym])); -by (asm_full_simp_tac - (simpset() addsimps [hypreal_inverse_add, - hypreal_inverse_distrib RS sym, hypreal_minus_inverse RS sym] - @ add_ac @ mult_ac - delsimps [inverse_mult_distrib,inverse_minus_eq, - minus_mult_left RS sym, - minus_mult_right RS sym] ) 1); -by (asm_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym, - right_distrib] - delsimps [minus_mult_left RS sym, - minus_mult_right RS sym]) 1); -by (res_inst_tac [("y"," inverse(- hypreal_of_real x * hypreal_of_real x)")] - approx_trans 1); -by (rtac inverse_add_Infinitesimal_approx2 1); -by (auto_tac (claset() addSDs [hypreal_of_real_HFinite_diff_Infinitesimal], - simpset() addsimps [hypreal_minus_inverse RS sym, - HFinite_minus_iff])); -by (rtac Infinitesimal_HFinite_mult2 1); -by Auto_tac; -qed "NSDERIV_inverse"; - - -Goal "x \\ 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))"; -by (asm_simp_tac (simpset() addsimps [NSDERIV_inverse, - NSDERIV_DERIV_iff RS sym] delsimps [realpow_Suc]) 1); -qed "DERIV_inverse"; - -(*-------------------------------------------------------------- - Derivative of inverse - -------------------------------------------------------------*) -Goal "[| DERIV f x :> d; f(x) \\ 0 |] \ -\ ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))"; -by (rtac (real_mult_commute RS subst) 1); -by (asm_simp_tac (HOL_ss addsimps [minus_mult_left, power_inverse]) 1); -by (fold_goals_tac [o_def]); -by (blast_tac (claset() addSIs [DERIV_chain,DERIV_inverse]) 1); -qed "DERIV_inverse_fun"; - -Goal "[| NSDERIV f x :> d; f(x) \\ 0 |] \ -\ ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))"; -by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff, - DERIV_inverse_fun] delsimps [realpow_Suc]) 1); -qed "NSDERIV_inverse_fun"; - -(*-------------------------------------------------------------- - Derivative of quotient - -------------------------------------------------------------*) -Goal "[| DERIV f x :> d; DERIV g x :> e; g(x) \\ 0 |] \ -\ ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) + -(e*f(x))) / (g(x) ^ Suc (Suc 0))"; -by (dres_inst_tac [("f","g")] DERIV_inverse_fun 1); -by (dtac DERIV_mult 2); -by (REPEAT(assume_tac 1)); -by (asm_full_simp_tac - (simpset() addsimps [real_divide_def, right_distrib, - power_inverse,minus_mult_left] @ mult_ac - delsimps [realpow_Suc, minus_mult_right RS sym, minus_mult_left RS sym]) 1); -qed "DERIV_quotient"; - -Goal "[| NSDERIV f x :> d; DERIV g x :> e; g(x) \\ 0 |] \ -\ ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) \ -\ + -(e*f(x))) / (g(x) ^ Suc (Suc 0))"; -by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff, - DERIV_quotient] delsimps [realpow_Suc]) 1); -qed "NSDERIV_quotient"; - -(* ------------------------------------------------------------------------ *) -(* Caratheodory formulation of derivative at a point: standard proof *) -(* ------------------------------------------------------------------------ *) - -Goal "(DERIV f x :> l) = \ -\ (\\g. (\\z. f z - f x = g z * (z - x)) & isCont g x & g x = l)"; -by Safe_tac; -by (res_inst_tac - [("x","%z. if z = x then l else (f(z) - f(x)) / (z - x)")] exI 1); -by (auto_tac (claset(),simpset() addsimps [real_mult_assoc, - ARITH_PROVE "z \\ x ==> z - x \\ (0::real)"])); -by (auto_tac (claset(),simpset() addsimps [isCont_iff,DERIV_iff])); -by (ALLGOALS(rtac (LIM_equal RS iffD1))); -by (auto_tac (claset(),simpset() addsimps [real_diff_def,real_mult_assoc])); -qed "CARAT_DERIV"; - -Goal "NSDERIV f x :> l ==> \ -\ \\g. (\\z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l"; -by (auto_tac (claset(),simpset() addsimps [NSDERIV_DERIV_iff, - isNSCont_isCont_iff,CARAT_DERIV])); -qed "CARAT_NSDERIV"; - -(* How about a NS proof? *) -Goal "(\\z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l \ -\ ==> NSDERIV f x :> l"; -by (auto_tac (claset(), - simpset() delsimprocs field_cancel_factor - addsimps [NSDERIV_iff2])); -by (auto_tac (claset(), - simpset() addsimps [hypreal_mult_assoc])); -by (asm_full_simp_tac (simpset() addsimps [hypreal_eq_minus_iff3 RS sym, - hypreal_diff_def]) 1); -by (asm_full_simp_tac (simpset() addsimps [isNSCont_def]) 1); -qed "CARAT_DERIVD"; - - - -(*--------------------------------------------------------------------------*) -(* Lemmas about nested intervals and proof by bisection (cf.Harrison) *) -(* All considerably tidied by lcp *) -(*--------------------------------------------------------------------------*) - -Goal "(\\n. (f::nat=>real) n \\ f (Suc n)) --> f m \\ f(m + no)"; -by (induct_tac "no" 1); -by (auto_tac (claset() addIs [order_trans], simpset())); -qed_spec_mp "lemma_f_mono_add"; - -Goal "[| \\n. f(n) \\ f(Suc n); \ -\ \\n. g(Suc n) \\ g(n); \ -\ \\n. f(n) \\ g(n) |] \ -\ ==> Bseq f"; -by (res_inst_tac [("k","f 0"),("K","g 0")] BseqI2 1 THEN rtac allI 1); -by (induct_tac "n" 1); -by (auto_tac (claset() addIs [order_trans], simpset())); -by (res_inst_tac [("y","g(Suc na)")] order_trans 1); -by (induct_tac "na" 2); -by (auto_tac (claset() addIs [order_trans], simpset())); -qed "f_inc_g_dec_Beq_f"; - -Goal "[| \\n. f(n) \\ f(Suc n); \ -\ \\n. g(Suc n) \\ g(n); \ -\ \\n. f(n) \\ g(n) |] \ -\ ==> Bseq g"; -by (stac (Bseq_minus_iff RS sym) 1); -by (res_inst_tac [("g","%x. -(f x)")] f_inc_g_dec_Beq_f 1); -by Auto_tac; -qed "f_inc_g_dec_Beq_g"; - -Goal "[| \\n. f n \\ f (Suc n); convergent f |] ==> f n \\ lim f"; -by (rtac (linorder_not_less RS iffD1) 1); -by (auto_tac (claset(), - simpset() addsimps [convergent_LIMSEQ_iff, LIMSEQ_iff, monoseq_Suc])); -by (dtac real_less_sum_gt_zero 1); -by (dres_inst_tac [("x","f n + - lim f")] spec 1); -by Safe_tac; -by (dres_inst_tac [("P","%na. no\\na --> ?Q na"),("x","no + n")] spec 1); -by Auto_tac; -by (subgoal_tac "lim f \\ f(no + n)" 1); -by (induct_tac "no" 2); -by (auto_tac (claset() addIs [order_trans], - simpset() addsimps [real_diff_def, real_abs_def])); -by (dres_inst_tac [("x","f(no + n)"),("no1","no")] - (lemma_f_mono_add RSN (2,order_less_le_trans)) 1); -by (auto_tac (claset(), simpset() addsimps [add_commute])); -qed "f_inc_imp_le_lim"; - -Goal "convergent g ==> lim (%x. - g x) = - (lim g)"; -by (rtac (LIMSEQ_minus RS limI) 1); -by (asm_full_simp_tac (simpset() addsimps [convergent_LIMSEQ_iff]) 1); -qed "lim_uminus"; - -Goal "[| \\n. g(Suc n) \\ g(n); convergent g |] ==> lim g \\ g n"; -by (subgoal_tac "- (g n) \\ - (lim g)" 1); -by (cut_inst_tac [("f", "%x. - (g x)")] f_inc_imp_le_lim 2); -by (auto_tac (claset(), - simpset() addsimps [lim_uminus, convergent_minus_iff RS sym])); -qed "g_dec_imp_lim_le"; - -Goal "[| \\n. f(n) \\ f(Suc n); \ -\ \\n. g(Suc n) \\ g(n); \ -\ \\n. f(n) \\ g(n) |] \ -\ ==> \\l m. l \\ m & ((\\n. f(n) \\ l) & f ----> l) & \ -\ ((\\n. m \\ g(n)) & g ----> m)"; -by (subgoal_tac "monoseq f & monoseq g" 1); -by (force_tac (claset(), simpset() addsimps [LIMSEQ_iff,monoseq_Suc]) 2); -by (subgoal_tac "Bseq f & Bseq g" 1); -by (blast_tac (claset() addIs [f_inc_g_dec_Beq_f, f_inc_g_dec_Beq_g]) 2); -by (auto_tac (claset() addSDs [Bseq_monoseq_convergent], - simpset() addsimps [convergent_LIMSEQ_iff])); -by (res_inst_tac [("x","lim f")] exI 1); -by (res_inst_tac [("x","lim g")] exI 1); -by (auto_tac (claset() addIs [LIMSEQ_le], simpset())); -by (auto_tac (claset(), - simpset() addsimps [f_inc_imp_le_lim, g_dec_imp_lim_le, - convergent_LIMSEQ_iff])); -qed "lemma_nest"; - -Goal "[| \\n. f(n) \\ f(Suc n); \ -\ \\n. g(Suc n) \\ g(n); \ -\ \\n. f(n) \\ g(n); \ -\ (%n. f(n) - g(n)) ----> 0 |] \ -\ ==> \\l. ((\\n. f(n) \\ l) & f ----> l) & \ -\ ((\\n. l \\ g(n)) & g ----> l)"; -by (dtac lemma_nest 1 THEN Auto_tac); -by (subgoal_tac "l = m" 1); -by (dres_inst_tac [("X","f")] LIMSEQ_diff 2); -by (auto_tac (claset() addIs [LIMSEQ_unique], simpset())); -qed "lemma_nest_unique"; - - -Goal "a \\ b ==> \ -\ \\n. fst (Bolzano_bisect P a b n) \\ snd (Bolzano_bisect P a b n)"; -by (rtac allI 1); -by (induct_tac "n" 1); -by (auto_tac (claset(), simpset() addsimps [Let_def, split_def])); -qed "Bolzano_bisect_le"; - -Goal "a \\ b ==> \ -\ \\n. fst(Bolzano_bisect P a b n) \\ fst (Bolzano_bisect P a b (Suc n))"; -by (rtac allI 1); -by (induct_tac "n" 1); -by (auto_tac (claset(), - simpset() addsimps [Bolzano_bisect_le, Let_def, split_def])); -qed "Bolzano_bisect_fst_le_Suc"; - -Goal "a \\ b ==> \ -\ \\n. snd(Bolzano_bisect P a b (Suc n)) \\ snd (Bolzano_bisect P a b n)"; -by (rtac allI 1); -by (induct_tac "n" 1); -by (auto_tac (claset(), - simpset() addsimps [Bolzano_bisect_le, Let_def, split_def])); -qed "Bolzano_bisect_Suc_le_snd"; - -Goal "((x::real) = y / (2 * z)) = (2 * x = y/z)"; -by Auto_tac; -by (dres_inst_tac [("f","%u. (1/2)*u")] arg_cong 1); -by Auto_tac; -qed "eq_divide_2_times_iff"; - -Goal "a \\ b ==> \ -\ snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = \ -\ (b-a) / (2 ^ n)"; -by (induct_tac "n" 1); -by (auto_tac (claset(), - simpset() addsimps [eq_divide_2_times_iff, add_divide_distrib, - Let_def, split_def])); -by (auto_tac (claset(), - simpset() addsimps (add_ac@[Bolzano_bisect_le, real_diff_def]))); -qed "Bolzano_bisect_diff"; - -val Bolzano_nest_unique = - [Bolzano_bisect_fst_le_Suc, Bolzano_bisect_Suc_le_snd, Bolzano_bisect_le] - MRS lemma_nest_unique; - -(*P_prem is a looping simprule, so it works better if it isn't an assumption*) -val P_prem::notP_prem::rest = -Goal "[| !!a b c. [| P(a,b); P(b,c); a \\ b; b \\ c|] ==> P(a,c); \ -\ ~ P(a,b); a \\ b |] ==> \ -\ ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))"; -by (cut_facts_tac rest 1); -by (induct_tac "n" 1); -by (auto_tac (claset(), - simpset() delsimps [surjective_pairing RS sym] - addsimps [notP_prem, Let_def, split_def])); -by (swap_res_tac [P_prem] 1); -by (assume_tac 1); -by (auto_tac (claset(), simpset() addsimps [Bolzano_bisect_le])); -qed "not_P_Bolzano_bisect"; - -(*Now we re-package P_prem as a formula*) -Goal "[| \\a b c. P(a,b) & P(b,c) & a \\ b & b \\ c --> P(a,c); \ -\ ~ P(a,b); a \\ b |] ==> \ -\ \\n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))"; -by (blast_tac (claset() addSEs [not_P_Bolzano_bisect RSN (2,rev_notE)]) 1); -qed "not_P_Bolzano_bisect'"; - - -Goal "[| \\a b c. P(a,b) & P(b,c) & a \\ b & b \\ c --> P(a,c); \ -\ \\x. \\d::real. 0 < d & \ -\ (\\a b. a \\ x & x \\ b & (b - a) < d --> P(a,b)); \ -\ a \\ b |] \ -\ ==> P(a,b)"; -by (rtac (inst "P1" "P" Bolzano_nest_unique RS exE) 1); -by (REPEAT (assume_tac 1)); -by (rtac LIMSEQ_minus_cancel 1); -by (asm_simp_tac (simpset() addsimps [Bolzano_bisect_diff, - LIMSEQ_divide_realpow_zero]) 1); -by (rtac ccontr 1); -by (dtac not_P_Bolzano_bisect' 1); -by (REPEAT (assume_tac 1)); -by (rename_tac "l" 1); -by (dres_inst_tac [("x","l")] spec 1 THEN Clarify_tac 1); -by (rewtac LIMSEQ_def); -by (dres_inst_tac [("P", "%r. 0 ?Q r"), ("x","d/2")] spec 1); -by (dres_inst_tac [("P", "%r. 0 ?Q r"), ("x","d/2")] spec 1); -by (dtac real_less_half_sum 1); -by Safe_tac; -(*linear arithmetic bug if we just use Asm_simp_tac*) -by (ALLGOALS Asm_full_simp_tac); -by (dres_inst_tac [("x","fst(Bolzano_bisect P a b (no + noa))")] spec 1); -by (dres_inst_tac [("x","snd(Bolzano_bisect P a b (no + noa))")] spec 1); -by Safe_tac; -by (ALLGOALS Asm_simp_tac); -by (res_inst_tac [("y","abs(fst(Bolzano_bisect P a b(no + noa)) - l) + \ -\ abs(snd(Bolzano_bisect P a b(no + noa)) - l)")] - order_le_less_trans 1); -by (asm_simp_tac (simpset() addsimps [real_abs_def]) 1); -by (rtac (real_sum_of_halves RS subst) 1); -by (rtac add_strict_mono 1); -by (ALLGOALS - (asm_full_simp_tac (simpset() addsimps [symmetric real_diff_def]))); -qed "lemma_BOLZANO"; - - -Goal "((\\a b c. (a \\ b & b \\ c & P(a,b) & P(b,c)) --> P(a,c)) & \ -\ (\\x. \\d::real. 0 < d & \ -\ (\\a b. a \\ x & x \\ b & (b - a) < d --> P(a,b)))) \ -\ --> (\\a b. a \\ b --> P(a,b))"; -by (Clarify_tac 1); -by (blast_tac (claset() addIs [lemma_BOLZANO]) 1); -qed "lemma_BOLZANO2"; - - -(*----------------------------------------------------------------------------*) -(* Intermediate Value Theorem (prove contrapositive by bisection) *) -(*----------------------------------------------------------------------------*) - -Goal "[| f(a) \\ y & y \\ f(b); \ -\ a \\ b; \ -\ (\\x. a \\ x & x \\ b --> isCont f x) |] \ -\ ==> \\x. a \\ x & x \\ b & f(x) = y"; -by (rtac contrapos_pp 1); -by (assume_tac 1); -by (cut_inst_tac - [("P","%(u,v). a \\ u & u \\ v & v \\ b --> ~(f(u) \\ y & y \\ f(v))")] - lemma_BOLZANO2 1); -by Safe_tac; -by (ALLGOALS(Asm_full_simp_tac)); -by (asm_full_simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1); -by (rtac ccontr 1); -by (subgoal_tac "a \\ x & x \\ b" 1); -by (Asm_full_simp_tac 2); -by (dres_inst_tac [("P", "%d. 0 ?P d"),("x","1")] spec 2); -by (Step_tac 2); -by (Asm_full_simp_tac 2); -by (Asm_full_simp_tac 2); -by (REPEAT(blast_tac (claset() addIs [order_trans]) 2)); -by (REPEAT(dres_inst_tac [("x","x")] spec 1)); -by (Asm_full_simp_tac 1); -by (dres_inst_tac [("P", "%r. ?P r --> (\\s. 0 aa" 1); -by (ALLGOALS Asm_full_simp_tac); -by (dres_inst_tac [("z","x"),("w","aa")] real_le_anti_sym 1); -by (assume_tac 1 THEN Asm_full_simp_tac 1); -qed "IVT"; - - -Goal "[| f(b) \\ y & y \\ f(a); \ -\ a \\ b; \ -\ (\\x. a \\ x & x \\ b --> isCont f x) \ -\ |] ==> \\x. a \\ x & x \\ b & f(x) = y"; -by (subgoal_tac "- f a \\ -y & -y \\ - f b" 1); -by (thin_tac "f b \\ y & y \\ f a" 1); -by (dres_inst_tac [("f","%x. - f x")] IVT 1); -by (auto_tac (claset() addIs [isCont_minus],simpset())); -qed "IVT2"; - - -(*HOL style here: object-level formulations*) -Goal "(f(a) \\ y & y \\ f(b) & a \\ b & \ -\ (\\x. a \\ x & x \\ b --> isCont f x)) \ -\ --> (\\x. a \\ x & x \\ b & f(x) = y)"; -by (blast_tac (claset() addIs [IVT]) 1); -qed "IVT_objl"; - -Goal "(f(b) \\ y & y \\ f(a) & a \\ b & \ -\ (\\x. a \\ x & x \\ b --> isCont f x)) \ -\ --> (\\x. a \\ x & x \\ b & f(x) = y)"; -by (blast_tac (claset() addIs [IVT2]) 1); -qed "IVT2_objl"; - -(*---------------------------------------------------------------------------*) -(* By bisection, function continuous on closed interval is bounded above *) -(*---------------------------------------------------------------------------*) - -Goal "abs (real x) = real (x::nat)"; -by (auto_tac (claset() addIs [abs_eqI1], simpset())); -qed "abs_real_of_nat_cancel"; -Addsimps [abs_real_of_nat_cancel]; - -Goal "~ abs(x) + (1::real) < x"; -by (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1); -by (auto_tac (claset() addIs [abs_ge_self RS order_trans],simpset())); -qed "abs_add_one_not_less_self"; -Addsimps [abs_add_one_not_less_self]; - - -Goal "[| a \\ b; \\x. a \\ x & x \\ b --> isCont f x |]\ -\ ==> \\M. \\x. a \\ x & x \\ b --> f(x) \\ M"; -by (cut_inst_tac [("P","%(u,v). a \\ u & u \\ v & v \\ b --> \ -\ (\\M. \\x. u \\ x & x \\ v --> f x \\ M)")] - lemma_BOLZANO2 1); -by Safe_tac; -by (ALLGOALS Asm_full_simp_tac); -by (rename_tac "x xa ya M Ma" 1); -by (cut_inst_tac [("x","M"),("y","Ma")] linorder_linear 1); -by Safe_tac; -by (res_inst_tac [("x","Ma")] exI 1); -by (Clarify_tac 1); -by (cut_inst_tac [("x","xb"),("y","xa")] linorder_linear 1); -by (Force_tac 1); -by (res_inst_tac [("x","M")] exI 1); -by (Clarify_tac 1); -by (cut_inst_tac [("x","xb"),("y","xa")] linorder_linear 1); -by (Force_tac 1); -by (case_tac "a \\ x & x \\ b" 1); -by (res_inst_tac [("x","1")] exI 2); -by (Force_tac 2); -by (asm_full_simp_tac (simpset() addsimps [LIM_def,isCont_iff]) 1); -by (dres_inst_tac [("x","x")] spec 1 THEN Auto_tac); -by (thin_tac "\\M. \\x. a \\ x & x \\ b & ~ f x \\ M" 1); -by (dres_inst_tac [("x","1")] spec 1); -by Auto_tac; -by (res_inst_tac [("x","s")] exI 1 THEN Clarify_tac 1); -by (res_inst_tac [("x","abs(f x) + 1")] exI 1 THEN Clarify_tac 1); -by (dres_inst_tac [("x","xa - x")] spec 1); -by (auto_tac (claset(), simpset() addsimps [abs_ge_self])); -by (REPEAT (arith_tac 1)); -qed "isCont_bounded"; - -(*----------------------------------------------------------------------------*) -(* Refine the above to existence of least upper bound *) -(*----------------------------------------------------------------------------*) - -Goal "((\\x. x \\ S) & (\\y. isUb UNIV S (y::real))) --> \ -\ (\\t. isLub UNIV S t)"; -by (blast_tac (claset() addIs [reals_complete]) 1); -qed "lemma_reals_complete"; - -Goal "[| a \\ b; \\x. a \\ x & x \\ b --> isCont f x |] \ -\ ==> \\M. (\\x. a \\ x & x \\ b --> f(x) \\ M) & \ -\ (\\N. N < M --> (\\x. a \\ x & x \\ b & N < f(x)))"; -by (cut_inst_tac [("S","Collect (%y. \\x. a \\ x & x \\ b & y = f x)")] - lemma_reals_complete 1); -by Auto_tac; -by (dtac isCont_bounded 1 THEN assume_tac 1); -by (auto_tac (claset(),simpset() addsimps [isUb_def,leastP_def, - isLub_def,setge_def,setle_def])); -by (rtac exI 1 THEN Auto_tac); -by (REPEAT(dtac spec 1) THEN Auto_tac); -by (dres_inst_tac [("x","x")] spec 1); -by (auto_tac (claset() addSIs [(linorder_not_less RS iffD1)],simpset())); -qed "isCont_has_Ub"; - -(*----------------------------------------------------------------------------*) -(* Now show that it attains its upper bound *) -(*----------------------------------------------------------------------------*) - -Goal "[| a \\ b; \\x. a \\ x & x \\ b --> isCont f x |] \ -\ ==> \\M. (\\x. a \\ x & x \\ b --> f(x) \\ M) & \ -\ (\\x. a \\ x & x \\ b & f(x) = M)"; -by (ftac isCont_has_Ub 1 THEN assume_tac 1); -by (Clarify_tac 1); -by (res_inst_tac [("x","M")] exI 1); -by (Asm_full_simp_tac 1); -by (rtac ccontr 1); -by (subgoal_tac "\\x. a \\ x & x \\ b --> f x < M" 1 THEN Step_tac 1); -by (rtac ccontr 2 THEN dtac (linorder_not_less RS iffD1) 2); -by (dres_inst_tac [("z","M")] real_le_anti_sym 2); -by (REPEAT(Blast_tac 2)); -by (subgoal_tac "\\x. a \\ x & x \\ b --> isCont (%x. inverse(M - f x)) x" 1); -by Safe_tac; -by (EVERY[rtac isCont_inverse 2, rtac isCont_diff 2, rtac notI 4]); -by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [diff_eq_eq]))); -by (Blast_tac 2); -by (subgoal_tac - "\\k. \\x. a \\ x & x \\ b --> (%x. inverse(M - (f x))) x \\ k" 1); -by (rtac isCont_bounded 2); -by Safe_tac; -by (subgoal_tac "\\x. a \\ x & x \\ b --> 0 < inverse(M - f(x))" 1); -by (Asm_full_simp_tac 1); -by Safe_tac; -by (asm_full_simp_tac (simpset() addsimps [less_diff_eq]) 2); -by (subgoal_tac - "\\x. a \\ x & x \\ b --> (%x. inverse(M - (f x))) x < (k + 1)" 1); -by Safe_tac; -by (res_inst_tac [("y","k")] order_le_less_trans 2); -by (asm_full_simp_tac (simpset() addsimps [zero_less_one]) 3); -by (Asm_full_simp_tac 2); -by (subgoal_tac "\\x. a \\ x & x \\ b --> \ -\ inverse(k + 1) < inverse((%x. inverse(M - (f x))) x)" 1); -by Safe_tac; -by (rtac less_imp_inverse_less 2); -by (ALLGOALS Asm_full_simp_tac); -by (dres_inst_tac [("P", "%N. N ?Q N"), - ("x","M - inverse(k + 1)")] spec 1); -by (Step_tac 1 THEN dtac (linorder_not_less RS iffD1) 1); -by (dtac (le_diff_eq RS iffD1) 1); -by (REPEAT(dres_inst_tac [("x","a")] spec 1)); -by (Asm_full_simp_tac 1); -by (asm_full_simp_tac - (simpset() addsimps [inverse_eq_divide, pos_divide_le_eq]) 1); -by (cut_inst_tac [("a","k"),("b","M-f a")] zero_less_mult_iff 1); -by (Asm_full_simp_tac 1); -(*last one*) -by (REPEAT(dres_inst_tac [("x","x")] spec 1)); -by (Asm_full_simp_tac 1); -qed "isCont_eq_Ub"; - - -(*----------------------------------------------------------------------------*) -(* Same theorem for lower bound *) -(*----------------------------------------------------------------------------*) - -Goal "[| a \\ b; \\x. a \\ x & x \\ b --> isCont f x |] \ -\ ==> \\M. (\\x. a \\ x & x \\ b --> M \\ f(x)) & \ -\ (\\x. a \\ x & x \\ b & f(x) = M)"; -by (subgoal_tac "\\x. a \\ x & x \\ b --> isCont (%x. -(f x)) x" 1); -by (blast_tac (claset() addIs [isCont_minus]) 2); -by (dres_inst_tac [("f","(%x. -(f x))")] isCont_eq_Ub 1); -by Safe_tac; -by Auto_tac; -qed "isCont_eq_Lb"; - - -(* ------------------------------------------------------------------------- *) -(* Another version. *) -(* ------------------------------------------------------------------------- *) - -Goal "[|a \\ b; \\x. a \\ x & x \\ b --> isCont f x |] \ -\ ==> \\L M. (\\x. a \\ x & x \\ b --> L \\ f(x) & f(x) \\ M) & \ -\ (\\y. L \\ y & y \\ M --> (\\x. a \\ x & x \\ b & (f(x) = y)))"; -by (ftac isCont_eq_Lb 1); -by (ftac isCont_eq_Ub 2); -by (REPEAT(assume_tac 1)); -by Safe_tac; -by (res_inst_tac [("x","f x")] exI 1); -by (res_inst_tac [("x","f xa")] exI 1); -by (Asm_full_simp_tac 1); -by Safe_tac; -by (cut_inst_tac [("x","x"),("y","xa")] linorder_linear 1); -by Safe_tac; -by (cut_inst_tac [("f","f"),("a","x"),("b","xa"),("y","y")] IVT_objl 1); -by (cut_inst_tac [("f","f"),("a","xa"),("b","x"),("y","y")] IVT2_objl 2); -by Safe_tac; -by (res_inst_tac [("x","xb")] exI 2); -by (res_inst_tac [("x","xb")] exI 4); -by (ALLGOALS(Asm_full_simp_tac)); -qed "isCont_Lb_Ub"; - -(*----------------------------------------------------------------------------*) -(* If f'(x) > 0 then x is locally strictly increasing at the right *) -(*----------------------------------------------------------------------------*) - -Goalw [deriv_def,LIM_def] - "[| DERIV f x :> l; 0 < l |] \ -\ ==> \\d. 0 < d & (\\h. 0 < h & h < d --> f(x) < f(x + h))"; -by (dtac spec 1 THEN Auto_tac); -by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac); -by (subgoal_tac "0 < l*h" 1); -by (asm_full_simp_tac (simpset() addsimps [zero_less_mult_iff]) 2); -by (dres_inst_tac [("x","h")] spec 1); -by (asm_full_simp_tac - (simpset() addsimps [real_abs_def, inverse_eq_divide, - pos_le_divide_eq, pos_less_divide_eq] - addsplits [split_if_asm]) 1); -qed "DERIV_left_inc"; - -val prems = goalw (the_context()) [deriv_def,LIM_def] - "[| DERIV f x :> l; l < 0 |] ==> \ -\ \\d. 0 < d & (\\h. 0 < h & h < d --> f(x) < f(x - h))"; -by (cut_facts_tac prems 1); (*needed because arith removes the assumption l<0*) -by (dres_inst_tac [("x","-l")] spec 1 THEN Auto_tac); -by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac); -by (dres_inst_tac [("x","-h")] spec 1); -by (asm_full_simp_tac - (simpset() addsimps [real_abs_def, inverse_eq_divide, - pos_less_divide_eq, - symmetric real_diff_def] - addsplits [split_if_asm]) 1); -by (subgoal_tac "0 < (f (x - h) - f x)/h" 1); -by (asm_full_simp_tac (simpset() addsimps [pos_less_divide_eq]) 1); -by (cut_facts_tac prems 1); -by (arith_tac 1); -qed "DERIV_left_dec"; - -(*????previous proof, revealing arith problem: -by (dres_inst_tac [("x","-l")] spec 1 THEN Auto_tac); -by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac); -by (subgoal_tac "l*h < 0" 1); -by (asm_full_simp_tac (simpset() addsimps [mult_less_0_iff]) 2); -by (dres_inst_tac [("x","-h")] spec 1); -by (asm_full_simp_tac - (simpset() addsimps [real_abs_def, inverse_eq_divide, - pos_less_divide_eq, - symmetric real_diff_def] - addsplits [split_if_asm] - delsimprocs [fast_real_arith_simproc]) 1); -by (subgoal_tac "0 < (f (x - h) - f x)/h" 1); -by (arith_tac 2); -by (asm_full_simp_tac - (simpset() addsimps [pos_less_divide_eq]) 1); -qed "DERIV_left_dec"; -*) - - -Goal "[| DERIV f x :> l; \ -\ \\d. 0 < d & (\\y. abs(x - y) < d --> f(y) \\ f(x)) |] \ -\ ==> l = 0"; -by (res_inst_tac [("x","l"),("y","0")] linorder_cases 1); -by Safe_tac; -by (dtac DERIV_left_dec 1); -by (dtac DERIV_left_inc 3); -by Safe_tac; -by (dres_inst_tac [("d1.0","d"),("d2.0","da")] real_lbound_gt_zero 1); -by (dres_inst_tac [("d1.0","d"),("d2.0","da")] real_lbound_gt_zero 3); -by Safe_tac; -by (dres_inst_tac [("x","x - e")] spec 1); -by (dres_inst_tac [("x","x + e")] spec 2); -by (auto_tac (claset(), simpset() addsimps [real_abs_def])); -qed "DERIV_local_max"; - -(*----------------------------------------------------------------------------*) -(* Similar theorem for a local minimum *) -(*----------------------------------------------------------------------------*) - -Goal "[| DERIV f x :> l; \ -\ \\d::real. 0 < d & (\\y. abs(x - y) < d --> f(x) \\ f(y)) |] \ -\ ==> l = 0"; -by (dtac (DERIV_minus RS DERIV_local_max) 1); -by Auto_tac; -qed "DERIV_local_min"; - -(*----------------------------------------------------------------------------*) -(* In particular if a function is locally flat *) -(*----------------------------------------------------------------------------*) - -Goal "[| DERIV f x :> l; \ -\ \\d. 0 < d & (\\y. abs(x - y) < d --> f(x) = f(y)) |] \ -\ ==> l = 0"; -by (auto_tac (claset() addSDs [DERIV_local_max],simpset())); -qed "DERIV_local_const"; - -(*----------------------------------------------------------------------------*) -(* Lemma about introducing open ball in open interval *) -(*----------------------------------------------------------------------------*) - -Goal "[| a < x; x < b |] ==> \ -\ \\d::real. 0 < d & (\\y. abs(x - y) < d --> a < y & y < b)"; -by (simp_tac (simpset() addsimps [abs_interval_iff]) 1); -by (cut_inst_tac [("x","x - a"),("y","b - x")] linorder_linear 1); -by Safe_tac; -by (res_inst_tac [("x","x - a")] exI 1); -by (res_inst_tac [("x","b - x")] exI 2); -by Auto_tac; -by (auto_tac (claset(),simpset() addsimps [less_diff_eq])); -qed "lemma_interval_lt"; - -Goal "[| a < x; x < b |] ==> \ -\ \\d::real. 0 < d & (\\y. abs(x - y) < d --> a \\ y & y \\ b)"; -by (dtac lemma_interval_lt 1); -by Auto_tac; -by (auto_tac (claset() addSIs [exI] ,simpset())); -qed "lemma_interval"; - -(*----------------------------------------------------------------------- - Rolle's Theorem - If f is defined and continuous on the finite closed interval [a,b] - and differentiable a least on the open interval (a,b), and f(a) = f(b), - then x0 \\ (a,b) such that f'(x0) = 0 - ----------------------------------------------------------------------*) - -Goal "[| a < b; f(a) = f(b); \ -\ \\x. a \\ x & x \\ b --> isCont f x; \ -\ \\x. a < x & x < b --> f differentiable x \ -\ |] ==> \\z. a < z & z < b & DERIV f z :> 0"; -by (ftac (order_less_imp_le RS isCont_eq_Ub) 1); -by (EVERY1[assume_tac,Step_tac]); -by (ftac (order_less_imp_le RS isCont_eq_Lb) 1); -by (EVERY1[assume_tac,Step_tac]); -by (case_tac "a < x & x < b" 1 THEN etac conjE 1); -by (Asm_full_simp_tac 2); -by (forw_inst_tac [("a","a"),("x","x")] lemma_interval 1); -by (EVERY1[assume_tac,etac exE]); -by (res_inst_tac [("x","x")] exI 1 THEN Asm_full_simp_tac 1); -by (subgoal_tac "(\\l. DERIV f x :> l) & \ -\ (\\d. 0 < d & (\\y. abs(x - y) < d --> f(y) \\ f(x)))" 1); -by (Clarify_tac 1 THEN rtac conjI 2); -by (blast_tac (claset() addIs [differentiableD]) 2); -by (Blast_tac 2); -by (ftac DERIV_local_max 1); -by (EVERY1[Blast_tac,Blast_tac]); -by (case_tac "a < xa & xa < b" 1 THEN etac conjE 1); -by (Asm_full_simp_tac 2); -by (forw_inst_tac [("a","a"),("x","xa")] lemma_interval 1); -by (EVERY1[assume_tac,etac exE]); -by (res_inst_tac [("x","xa")] exI 1 THEN Asm_full_simp_tac 1); -by (subgoal_tac "(\\l. DERIV f xa :> l) & \ -\ (\\d. 0 < d & (\\y. abs(xa - y) < d --> f(xa) \\ f(y)))" 1); -by (Clarify_tac 1 THEN rtac conjI 2); -by (blast_tac (claset() addIs [differentiableD]) 2); -by (Blast_tac 2); -by (ftac DERIV_local_min 1); -by (EVERY1[Blast_tac,Blast_tac]); -by (subgoal_tac "\\x. a \\ x & x \\ b --> f(x) = f(b)" 1); -by (Clarify_tac 2); -by (rtac real_le_anti_sym 2); -by (subgoal_tac "f b = f x" 2); -by (Asm_full_simp_tac 2); -by (res_inst_tac [("x1","a"),("y1","x")] (order_le_imp_less_or_eq RS disjE) 2); -by (assume_tac 2); -by (dres_inst_tac [("z","x"),("w","b")] real_le_anti_sym 2); -by (subgoal_tac "f b = f xa" 5); -by (Asm_full_simp_tac 5); -by (res_inst_tac [("x1","a"),("y1","xa")] (order_le_imp_less_or_eq RS disjE) 5); -by (assume_tac 5); -by (dres_inst_tac [("z","xa"),("w","b")] real_le_anti_sym 5); -by (REPEAT(Asm_full_simp_tac 2)); -by (dtac real_dense 1 THEN etac exE 1); -by (res_inst_tac [("x","r")] exI 1 THEN Asm_simp_tac 1); -by (etac conjE 1); -by (forw_inst_tac [("a","a"),("x","r")] lemma_interval 1); -by (EVERY1[assume_tac, etac exE]); -by (subgoal_tac "(\\l. DERIV f r :> l) & \ -\ (\\d. 0 < d & (\\y. abs(r - y) < d --> f(r) = f(y)))" 1); -by (Clarify_tac 1 THEN rtac conjI 2); -by (blast_tac (claset() addIs [differentiableD]) 2); -by (EVERY1[ftac DERIV_local_const, Blast_tac, Blast_tac]); -by (res_inst_tac [("x","d")] exI 1); -by (EVERY1[rtac conjI, Blast_tac, rtac allI, rtac impI]); -by (res_inst_tac [("s","f b")] trans 1); -by (blast_tac (claset() addSDs [order_less_imp_le]) 1); -by (rtac sym 1 THEN Blast_tac 1); -qed "Rolle"; - -(*----------------------------------------------------------------------------*) -(* Mean value theorem *) -(*----------------------------------------------------------------------------*) - -Goal "f a - (f b - f a)/(b - a) * a = \ -\ f b - (f b - f a)/(b - a) * (b::real)"; -by (case_tac "a = b" 1); -by (Asm_full_simp_tac 1); -by (res_inst_tac [("c1","b - a")] (real_mult_left_cancel RS iffD1) 1); -by (arith_tac 1); -by (auto_tac (claset(), simpset() addsimps [right_diff_distrib])); -by (auto_tac (claset(), simpset() addsimps [left_diff_distrib])); -qed "lemma_MVT"; - -Goal "[| a < b; \ -\ \\x. a \\ x & x \\ b --> isCont f x; \ -\ \\x. a < x & x < b --> f differentiable x |] \ -\ ==> \\l z. a < z & z < b & DERIV f z :> l & \ -\ (f(b) - f(a) = (b - a) * l)"; -by (dres_inst_tac [("f","%x. f(x) - (((f(b) - f(a)) / (b - a)) * x)")] - Rolle 1); -by (rtac lemma_MVT 1); -by Safe_tac; -by (rtac isCont_diff 1 THEN Blast_tac 1); -by (rtac (isCont_const RS isCont_mult) 1); -by (rtac isCont_Id 1); -by (dres_inst_tac [("P", "%x. ?Pre x --> f differentiable x"), - ("x","x")] spec 1); -by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1); -by Safe_tac; -by (res_inst_tac [("x","xa - ((f(b) - f(a)) / (b - a))")] exI 1); -by (rtac DERIV_diff 1 THEN assume_tac 1); -(*derivative of a linear function is the constant...*) -by (subgoal_tac "(%x. (f b - f a) * x / (b - a)) = \ -\ op * ((f b - f a) / (b - a))" 1); -by (rtac ext 2 THEN Simp_tac 2); -by (Asm_full_simp_tac 1); -(*final case*) -by (res_inst_tac [("x","((f(b) - f(a)) / (b - a))")] exI 1); -by (res_inst_tac [("x","z")] exI 1); -by Safe_tac; -by (Asm_full_simp_tac 2); -by (subgoal_tac "DERIV (%x. ((f(b) - f(a)) / (b - a)) * x) z :> \ -\ ((f(b) - f(a)) / (b - a))" 1); -by (rtac DERIV_cmult_Id 2); -by (dtac DERIV_add 1 THEN assume_tac 1); -by (asm_full_simp_tac (simpset() addsimps [real_add_assoc, real_diff_def]) 1); -qed "MVT"; - -(*----------------------------------------------------------------------------*) -(* Theorem that function is constant if its derivative is 0 over an interval. *) -(*----------------------------------------------------------------------------*) - -Goal "[| a < b; \ -\ \\x. a \\ x & x \\ b --> isCont f x; \ -\ \\x. a < x & x < b --> DERIV f x :> 0 |] \ -\ ==> (f b = f a)"; -by (dtac MVT 1 THEN assume_tac 1); -by (blast_tac (claset() addIs [differentiableI]) 1); -by (auto_tac (claset() addSDs [DERIV_unique],simpset() - addsimps [diff_eq_eq])); -qed "DERIV_isconst_end"; - -Goal "[| a < b; \ -\ \\x. a \\ x & x \\ b --> isCont f x; \ -\ \\x. a < x & x < b --> DERIV f x :> 0 |] \ -\ ==> \\x. a \\ x & x \\ b --> f x = f a"; -by Safe_tac; -by (dres_inst_tac [("x","a")] order_le_imp_less_or_eq 1); -by Safe_tac; -by (dres_inst_tac [("b","x")] DERIV_isconst_end 1); -by Auto_tac; -qed "DERIV_isconst1"; - -Goal "[| a < b; \ -\ \\x. a \\ x & x \\ b --> isCont f x; \ -\ \\x. a < x & x < b --> DERIV f x :> 0; \ -\ a \\ x; x \\ b |] \ -\ ==> f x = f a"; -by (blast_tac (claset() addDs [DERIV_isconst1]) 1); -qed "DERIV_isconst2"; - -Goal "\\x. DERIV f x :> 0 ==> f(x) = f(y)"; -by (res_inst_tac [("x","x"),("y","y")] linorder_cases 1); -by (rtac sym 1); -by (auto_tac (claset() addIs [DERIV_isCont,DERIV_isconst_end],simpset())); -qed "DERIV_isconst_all"; - -Goal "[|a \\ b; \\x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b - a) * k"; -by (res_inst_tac [("x","a"),("y","b")] linorder_cases 1); -by Auto_tac; -by (ALLGOALS(dres_inst_tac [("f","f")] MVT)); -by (auto_tac (claset() addDs [DERIV_isCont,DERIV_unique],simpset() addsimps - [differentiable_def])); -by (auto_tac (claset() addDs [DERIV_unique], - simpset() addsimps [left_distrib, real_diff_def])); -qed "DERIV_const_ratio_const"; - -Goal "[|a \\ b; \\x. DERIV f x :> k |] ==> (f(b) - f(a))/(b - a) = k"; -by (res_inst_tac [("c1","b - a")] (real_mult_right_cancel RS iffD1) 1); -by (auto_tac (claset() addSDs [DERIV_const_ratio_const], - simpset() addsimps [real_mult_assoc])); -qed "DERIV_const_ratio_const2"; - -Goal "((a + b) /2 - a) = (b - a)/(2::real)"; -by Auto_tac; -qed "real_average_minus_first"; -Addsimps [real_average_minus_first]; - -Goal "((b + a)/2 - a) = (b - a)/(2::real)"; -by Auto_tac; -qed "real_average_minus_second"; -Addsimps [real_average_minus_second]; - - -(* Gallileo's "trick": average velocity = av. of end velocities *) -Goal "[|a \\ (b::real); \\x. DERIV v x :> k|] \ -\ ==> v((a + b)/2) = (v a + v b)/2"; -by (res_inst_tac [("x","a"),("y","b")] linorder_cases 1); -by Safe_tac; -by (ftac DERIV_const_ratio_const2 1 THEN assume_tac 1); -by (ftac DERIV_const_ratio_const2 2 THEN assume_tac 2); -by (dtac real_less_half_sum 1); -by (dtac real_gt_half_sum 2); -by (ftac (real_not_refl2 RS DERIV_const_ratio_const2) 1 THEN assume_tac 1); -by (dtac ((real_not_refl2 RS not_sym) RS DERIV_const_ratio_const2) 2 - THEN assume_tac 2); -by (ALLGOALS (dres_inst_tac [("f","%u. (b-a)*u")] arg_cong)); -by (auto_tac (claset(), simpset() addsimps [inverse_eq_divide])); -by (asm_full_simp_tac (simpset() addsimps [real_add_commute, eq_commute]) 1); -qed "DERIV_const_average"; - - -(* ------------------------------------------------------------------------ *) -(* Dull lemma that an continuous injection on an interval must have a strict*) -(* maximum at an end point, not in the middle. *) -(* ------------------------------------------------------------------------ *) - -Goal "[|0 < d; \\z. abs(z - x) \\ d --> g(f z) = z; \ -\ \\z. abs(z - x) \\ d --> isCont f z |] \ -\ ==> ~(\\z. abs(z - x) \\ d --> f(z) \\ f(x))"; -by (rtac notI 1); -by (rotate_tac 3 1); -by (forw_inst_tac [("x","x - d")] spec 1); -by (forw_inst_tac [("x","x + d")] spec 1); -by Safe_tac; -by (cut_inst_tac [("x","f(x - d)"),("y","f(x + d)")] - (ARITH_PROVE "x \\ y | y \\ (x::real)") 4); -by (etac disjE 4); -by (REPEAT(arith_tac 1)); -by (cut_inst_tac [("f","f"),("a","x - d"),("b","x"),("y","f(x + d)")] - IVT_objl 1); -by Safe_tac; -by (arith_tac 1); -by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1); -by (dres_inst_tac [("f","g")] arg_cong 1); -by (rotate_tac 2 1); -by (forw_inst_tac [("x","xa")] spec 1); -by (dres_inst_tac [("x","x + d")] spec 1); -by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1); -(* 2nd case: similar *) -by (cut_inst_tac [("f","f"),("a","x"),("b","x + d"),("y","f(x - d)")] - IVT2_objl 1); -by Safe_tac; -by (arith_tac 1); -by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1); -by (dres_inst_tac [("f","g")] arg_cong 1); -by (rotate_tac 2 1); -by (forw_inst_tac [("x","xa")] spec 1); -by (dres_inst_tac [("x","x - d")] spec 1); -by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1); -qed "lemma_isCont_inj"; - -(* ------------------------------------------------------------------------ *) -(* Similar version for lower bound *) -(* ------------------------------------------------------------------------ *) - -Goal "[|0 < d; \\z. abs(z - x) \\ d --> g(f z) = z; \ -\ \\z. abs(z - x) \\ d --> isCont f z |] \ -\ ==> ~(\\z. abs(z - x) \\ d --> f(x) \\ f(z))"; -by (auto_tac (claset() addSDs [(asm_full_simplify (simpset()) - (read_instantiate [("f","%x. - f x"),("g","%y. g(-y)"),("x","x"),("d","d")] - lemma_isCont_inj))],simpset() addsimps [isCont_minus])); -qed "lemma_isCont_inj2"; - -(* ------------------------------------------------------------------------ *) -(* Show there's an interval surrounding f(x) in f[[x - d, x + d]] *) -(* Also from John's theory *) -(* ------------------------------------------------------------------------ *) - -val lemma_le = ARITH_PROVE "0 \\ (d::real) ==> -d \\ d"; - -(* FIXME: awful proof - needs improvement *) -Goal "[| 0 < d; \\z. abs(z - x) \\ d --> g(f z) = z; \ -\ \\z. abs(z - x) \\ d --> isCont f z |] \ -\ ==> \\e. 0 < e & \ -\ (\\y. \ -\ abs(y - f(x)) \\ e --> \ -\ (\\z. abs(z - x) \\ d & (f z = y)))"; -by (ftac order_less_imp_le 1); -by (dtac (lemma_le RS (asm_full_simplify (simpset()) (read_instantiate - [("f","f"),("a","x - d"),("b","x + d")] isCont_Lb_Ub))) 1); -by Safe_tac; -by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1); -by (subgoal_tac "L \\ f x & f x \\ M" 1); -by (dres_inst_tac [("P", "%v. ?P v --> ?Q v & ?R v"), ("x","x")] spec 2); -by (Asm_full_simp_tac 2); -by (subgoal_tac "L < f x & f x < M" 1); -by Safe_tac; -by (dres_inst_tac [("x","L")] (ARITH_PROVE "x < y ==> 0 < y - (x::real)") 1); -by (dres_inst_tac [("x","f x")] (ARITH_PROVE "x < y ==> 0 < y - (x::real)") 1); -by (dres_inst_tac [("d1.0","f x - L"),("d2.0","M - f x")] - (real_lbound_gt_zero) 1); -by Safe_tac; -by (res_inst_tac [("x","e")] exI 1); -by Safe_tac; -by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1); -by (dres_inst_tac [("P","%v. ?PP v --> (\\xa. ?Q v xa)"),("x","y")] spec 1); -by (Step_tac 1 THEN REPEAT(arith_tac 1)); -by (res_inst_tac [("x","xa")] exI 1); -by (arith_tac 1); -by (ALLGOALS(etac (ARITH_PROVE "[|x \\ y; x \\ y |] ==> x < (y::real)"))); -by (ALLGOALS(rotate_tac 3)); -by (dtac lemma_isCont_inj2 1); -by (assume_tac 2); -by (dtac lemma_isCont_inj 3); -by (assume_tac 4); -by (TRYALL(assume_tac)); -by Safe_tac; -by (ALLGOALS(dres_inst_tac [("x","z")] spec)); -by (ALLGOALS(arith_tac)); -qed "isCont_inj_range"; - - -(* ------------------------------------------------------------------------ *) -(* Continuity of inverse function *) -(* ------------------------------------------------------------------------ *) - -Goal "[| 0 < d; \\z. abs(z - x) \\ d --> g(f(z)) = z; \ -\ \\z. abs(z - x) \\ d --> isCont f z |] \ -\ ==> isCont g (f x)"; -by (simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1); -by Safe_tac; -by (dres_inst_tac [("d1.0","r")] (real_lbound_gt_zero) 1); -by (assume_tac 1 THEN Step_tac 1); -by (subgoal_tac "\\z. abs(z - x) \\ e --> (g(f z) = z)" 1); -by (Force_tac 2); -by (subgoal_tac "\\z. abs(z - x) \\ e --> isCont f z" 1); -by (Force_tac 2); -by (dres_inst_tac [("d","e")] isCont_inj_range 1); -by (assume_tac 2 THEN assume_tac 1); -by Safe_tac; -by (res_inst_tac [("x","ea")] exI 1); -by Auto_tac; -by (rotate_tac 4 1); -by (dres_inst_tac [("x","f(x) + xa")] spec 1); -by Auto_tac; -by (dtac sym 1 THEN Auto_tac); -by (arith_tac 1); -qed "isCont_inverse_function"; - diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/Hyperreal/Lim.thy --- a/src/HOL/Hyperreal/Lim.thy Fri Mar 19 10:50:06 2004 +0100 +++ b/src/HOL/Hyperreal/Lim.thy Fri Mar 19 10:51:03 2004 +0100 @@ -1,66 +1,67 @@ (* Title : Lim.thy + ID : $Id$ Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge - Description : Theory of limits, continuity and - differentiation of real=>real functions + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 *) -Lim = SEQ + RealDef + +header{*Limits, Continuity and Differentiation*} -(*----------------------------------------------------------------------- - Limits, continuity and differentiation: standard and NS definitions - -----------------------------------------------------------------------*) +theory Lim = SEQ + RealDef: + +text{*Standard and Nonstandard Definitions*} constdefs - LIM :: [real=>real,real,real] => bool + LIM :: "[real=>real,real,real] => bool" ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) "f -- a --> L == - ALL r. 0 < r --> - (EX s. 0 < s & (ALL x. (x ~= a & (abs(x + -a) < s) - --> abs(f x + -L) < r)))" + \r. 0 < r --> + (\s. 0 < s & (\x. (x \ a & (\x + -a\ < s) + --> \f x + -L\ < r)))" - NSLIM :: [real=>real,real,real] => bool + NSLIM :: "[real=>real,real,real] => bool" ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60) - "f -- a --NS> L == (ALL x. (x ~= hypreal_of_real a & + "f -- a --NS> L == (\x. (x \ hypreal_of_real a & x @= hypreal_of_real a --> - ( *f* f) x @= hypreal_of_real L))" + ( *f* f) x @= hypreal_of_real L))" - isCont :: [real=>real,real] => bool - "isCont f a == (f -- a --> (f a))" + isCont :: "[real=>real,real] => bool" + "isCont f a == (f -- a --> (f a))" (* NS definition dispenses with limit notions *) - isNSCont :: [real=>real,real] => bool - "isNSCont f a == (ALL y. y @= hypreal_of_real a --> + isNSCont :: "[real=>real,real] => bool" + "isNSCont f a == (\y. y @= hypreal_of_real a --> ( *f* f) y @= hypreal_of_real (f a))" (* differentiation: D is derivative of function f at x *) - deriv:: [real=>real,real,real] => bool + deriv:: "[real=>real,real,real] => bool" ("(DERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60) - "DERIV f x :> D == ((%h. (f(x + h) + -f(x))/h) -- 0 --> D)" + "DERIV f x :> D == ((%h. (f(x + h) + -f x)/h) -- 0 --> D)" - nsderiv :: [real=>real,real,real] => bool + nsderiv :: "[real=>real,real,real] => bool" ("(NSDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60) - "NSDERIV f x :> D == (ALL h: Infinitesimal - {0}. - (( *f* f)(hypreal_of_real x + h) + + "NSDERIV f x :> D == (\h \ Infinitesimal - {0}. + (( *f* f)(hypreal_of_real x + h) + - hypreal_of_real (f x))/h @= hypreal_of_real D)" - differentiable :: [real=>real,real] => bool (infixl 60) - "f differentiable x == (EX D. DERIV f x :> D)" + differentiable :: "[real=>real,real] => bool" (infixl "differentiable" 60) + "f differentiable x == (\D. DERIV f x :> D)" - NSdifferentiable :: [real=>real,real] => bool (infixl 60) - "f NSdifferentiable x == (EX D. NSDERIV f x :> D)" + NSdifferentiable :: "[real=>real,real] => bool" + (infixl "NSdifferentiable" 60) + "f NSdifferentiable x == (\D. NSDERIV f x :> D)" - increment :: [real=>real,real,hypreal] => hypreal - "increment f x h == (@inc. f NSdifferentiable x & + increment :: "[real=>real,real,hypreal] => hypreal" + "increment f x h == (@inc. f NSdifferentiable x & inc = ( *f* f)(hypreal_of_real x + h) + -hypreal_of_real (f x))" - isUCont :: (real=>real) => bool - "isUCont f == (ALL r. 0 < r --> - (EX s. 0 < s & (ALL x y. abs(x + -y) < s - --> abs(f x + -f y) < r)))" + isUCont :: "(real=>real) => bool" + "isUCont f == (\r. 0 < r --> + (\s. 0 < s & (\x y. \x + -y\ < s + --> \f x + -f y\ < r)))" - isNSUCont :: (real=>real) => bool - "isNSUCont f == (ALL x y. x @= y --> ( *f* f) x @= ( *f* f) y)" + isNSUCont :: "(real=>real) => bool" + "isNSUCont f == (\x y. x @= y --> ( *f* f) x @= ( *f* f) y)" (*Used in the proof of the Bolzano theorem*) @@ -72,8 +73,2258 @@ "Bolzano_bisect P a b (Suc n) = (let (x,y) = Bolzano_bisect P a b n in if P(x, (x+y)/2) then ((x+y)/2, y) - else (x, (x+y)/2) )" - + else (x, (x+y)/2))" + + + +section{*Some Purely Standard Proofs*} + +lemma LIM_eq: + "f -- a --> L = + (\r. 0 (\s. 0 < s & (\x. x \ a & \x-a\ < s --> \f x - L\ < r)))" +by (simp add: LIM_def diff_minus) + +lemma LIM_D: + "[| f -- a --> L; 0 \s. 0 < s & (\x. x \ a & \x-a\ < s --> \f x - L\ < r)" +by (simp add: LIM_eq) + +lemma LIM_const: "(%x. k) -- x --> k" +by (simp add: LIM_def) +declare LIM_const [simp] + +lemma LIM_add: + assumes f: "f -- a --> L" and g: "g -- a --> M" + shows "(%x. f x + g(x)) -- a --> (L + M)" +proof (simp add: LIM_eq, clarify) + fix r :: real + assume r: "0x. x \ a & \x-a\ < fs --> \f x - L\ < r/2" + by blast + from LIM_D [OF g half_gt_zero [OF r]] + obtain gs + where gs: "0 < gs" + and gs_lt: "\x. x \ a & \x-a\ < gs --> \g x - M\ < r/2" + by blast + show "\s. 0 < s \ + (\x. x \ a \ \x-a\ < s \ \f x + g x - (L + M)\ < r)" + proof (intro exI conjI strip) + show "0 < min fs gs" by (simp add: fs gs) + fix x :: real + assume "x \ a \ \x-a\ < min fs gs" + with fs_lt gs_lt + have "\f x - L\ < r/2" and "\g x - M\ < r/2" by (auto simp add: fs_lt) + hence "\f x - L\ + \g x - M\ < r" by arith + thus "\f x + g x - (L + M)\ < r" + by (blast intro: abs_diff_triangle_ineq order_le_less_trans) + qed +qed + +lemma LIM_minus: "f -- a --> L ==> (%x. -f(x)) -- a --> -L" +apply (simp add: LIM_eq) +apply (subgoal_tac "\x. \- f x + L\ = \f x - L\") +apply (simp_all add: abs_if) +done + +lemma LIM_add_minus: + "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + -g(x)) -- x --> (l + -m)" +by (blast dest: LIM_add LIM_minus) + +lemma LIM_diff: + "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) - g(x)) -- x --> l-m" +by (simp add: diff_minus LIM_add_minus) + + +lemma LIM_const_not_eq: "k \ L ==> ~ ((%x. k) -- a --> L)" +proof (simp add: linorder_neq_iff LIM_eq, elim disjE) + assume k: "k < L" + show "\r. 0 < r \ + (\s. 0 < s \ (\x. (x < a \ a < x) \ \x-a\ < s) \ \ \k-L\ < r)" + proof (intro exI conjI strip) + show "0 < L-k" by (simp add: k) + fix s :: real + assume s: "0 a < s/2 + a" by arith + next + from s show "\s / 2 + a - a\ < s" by (simp add: abs_if) + next + from s show "~ \k-L\ < L-k" by (simp add: abs_if) } + qed +next + assume k: "L < k" + show "\r. 0 < r \ + (\s. 0 < s \ (\x. (x < a \ a < x) \ \x-a\ < s) \ \ \k-L\ < r)" + proof (intro exI conjI strip) + show "0 < k-L" by (simp add: k) + fix s :: real + assume s: "0 a < s/2 + a" by arith + next + from s show "\s / 2 + a - a\ < s" by (simp add: abs_if) + next + from s show "~ \k-L\ < k-L" by (simp add: abs_if) } + qed +qed + +lemma LIM_const_eq: "(%x. k) -- x --> L ==> k = L" +apply (rule ccontr) +apply (blast dest: LIM_const_not_eq) +done + +lemma LIM_unique: "[| f -- a --> L; f -- a --> M |] ==> L = M" +apply (drule LIM_diff, assumption) +apply (auto dest!: LIM_const_eq) +done + +lemma LIM_mult_zero: + assumes f: "f -- a --> 0" and g: "g -- a --> 0" + shows "(%x. f(x) * g(x)) -- a --> 0" +proof (simp add: LIM_eq, clarify) + fix r :: real + assume r: "0x. x \ a & \x-a\ < fs --> \f x\ < 1" + by auto + from LIM_D [OF g r] + obtain gs + where gs: "0 < gs" + and gs_lt: "\x. x \ a & \x-a\ < gs --> \g x\ < r" + by auto + show "\s. 0 < s \ (\x. x \ a \ \x-a\ < s \ \f x\ * \g x\ < r)" + proof (intro exI conjI strip) + show "0 < min fs gs" by (simp add: fs gs) + fix x :: real + assume "x \ a \ \x-a\ < min fs gs" + with fs_lt gs_lt + have "\f x\ < 1" and "\g x\ < r" by (auto simp add: fs_lt) + hence "\f x\ * \g x\ < 1*r" by (rule abs_mult_less) + thus "\f x\ * \g x\ < r" by simp + qed +qed + +lemma LIM_self: "(%x. x) -- a --> a" +by (auto simp add: LIM_def) + +text{*Limits are equal for functions equal except at limit point*} +lemma LIM_equal: + "[| \x. x \ a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)" +by (simp add: LIM_def) + +text{*Two uses in Hyperreal/Transcendental.ML*} +lemma LIM_trans: + "[| (%x. f(x) + -g(x)) -- a --> 0; g -- a --> l |] ==> f -- a --> l" +apply (drule LIM_add, assumption) +apply (auto simp add: add_assoc) +done + + +subsection{*Relationships Between Standard and Nonstandard Concepts*} + +text{*Standard and NS definitions of Limit*} (*NEEDS STRUCTURING*) +lemma LIM_NSLIM: + "f -- x --> L ==> f -- x --NS> L" +apply (simp add: LIM_def NSLIM_def approx_def) +apply (simp add: Infinitesimal_FreeUltrafilterNat_iff, safe) +apply (rule_tac z = xa in eq_Abs_hypreal) +apply (auto simp add: real_add_minus_iff starfun hypreal_minus hypreal_of_real_def hypreal_add) +apply (rule bexI, rule_tac [2] lemma_hyprel_refl, clarify) +apply (drule_tac x = u in spec, clarify) +apply (drule_tac x = s in spec, clarify) +apply (subgoal_tac "\n::nat. (xa n) \ x & abs ((xa n) + - x) < s --> abs (f (xa n) + - L) < u") +prefer 2 apply blast +apply (drule FreeUltrafilterNat_all, ultra) +done + +(*--------------------------------------------------------------------- + Limit: NS definition ==> standard definition + ---------------------------------------------------------------------*) + +lemma lemma_LIM: "\s. 0 < s --> (\xa. xa \ x & + \xa + - x\ < s & r \ \f xa + -L\) + ==> \n::nat. \xa. xa \ x & + \xa + -x\ < inverse(real(Suc n)) & r \ \f xa + -L\" +apply clarify +apply (cut_tac n1 = n in real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive], auto) +done + +lemma lemma_skolemize_LIM2: + "\s. 0 < s --> (\xa. xa \ x & + \xa + - x\ < s & r \ \f xa + -L\) + ==> \X. \n::nat. X n \ x & + \X n + -x\ < inverse(real(Suc n)) & r \ abs(f (X n) + -L)" +apply (drule lemma_LIM) +apply (drule choice, blast) +done + +lemma lemma_simp: "\n. X n \ x & + \X n + - x\ < inverse (real(Suc n)) & + r \ abs (f (X n) + - L) ==> + \n. \X n + - x\ < inverse (real(Suc n))" +by auto + + +(*------------------- + NSLIM => LIM + -------------------*) + +lemma NSLIM_LIM: "f -- x --NS> L ==> f -- x --> L" +apply (simp add: LIM_def NSLIM_def approx_def) +apply (simp add: Infinitesimal_FreeUltrafilterNat_iff, clarify) +apply (rule ccontr, simp) +apply (simp add: linorder_not_less) +apply (drule lemma_skolemize_LIM2, safe) +apply (drule_tac x = "Abs_hypreal (hyprel``{X}) " in spec) +apply (auto simp add: starfun hypreal_minus hypreal_of_real_def hypreal_add) +apply (drule lemma_simp [THEN real_seq_to_hypreal_Infinitesimal]) +apply (simp add: Infinitesimal_FreeUltrafilterNat_iff hypreal_of_real_def hypreal_minus hypreal_add, blast) +apply (drule spec, drule mp, assumption) +apply (drule FreeUltrafilterNat_all, ultra) +done + + +(**** Key result ****) +lemma LIM_NSLIM_iff: "(f -- x --> L) = (f -- x --NS> L)" +by (blast intro: LIM_NSLIM NSLIM_LIM) + +(*-------------------------------------------------------------------*) +(* Proving properties of limits using nonstandard definition and *) +(* hence, the properties hold for standard limits as well *) +(*-------------------------------------------------------------------*) +(*------------------------------------------------ + NSLIM_mult and hence (trivially) LIM_mult + ------------------------------------------------*) + +lemma NSLIM_mult: + "[| f -- x --NS> l; g -- x --NS> m |] + ==> (%x. f(x) * g(x)) -- x --NS> (l * m)" +apply (simp add: NSLIM_def) +apply (auto intro!: approx_mult_HFinite) +done + +lemma LIM_mult2: "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) * g(x)) -- x --> (l * m)" +by (simp add: LIM_NSLIM_iff NSLIM_mult) + +(*---------------------------------------------- + NSLIM_add and hence (trivially) LIM_add + Note the much shorter proof + ----------------------------------------------*) +lemma NSLIM_add: + "[| f -- x --NS> l; g -- x --NS> m |] + ==> (%x. f(x) + g(x)) -- x --NS> (l + m)" +apply (simp add: NSLIM_def) +apply (auto intro!: approx_add) +done + +lemma LIM_add2: "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + g(x)) -- x --> (l + m)" +by (simp add: LIM_NSLIM_iff NSLIM_add) + + +lemma NSLIM_const: "(%x. k) -- x --NS> k" +by (simp add: NSLIM_def) + +declare NSLIM_const [simp] + +lemma LIM_const2: "(%x. k) -- x --> k" +by (simp add: LIM_NSLIM_iff) + + +lemma NSLIM_minus: "f -- a --NS> L ==> (%x. -f(x)) -- a --NS> -L" +by (simp add: NSLIM_def) + +lemma LIM_minus2: "f -- a --> L ==> (%x. -f(x)) -- a --> -L" +by (simp add: LIM_NSLIM_iff NSLIM_minus) + + +lemma NSLIM_add_minus: "[| f -- x --NS> l; g -- x --NS> m |] ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)" +by (blast dest: NSLIM_add NSLIM_minus) + +lemma LIM_add_minus2: "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + -g(x)) -- x --> (l + -m)" +by (simp add: LIM_NSLIM_iff NSLIM_add_minus) + + +lemma NSLIM_inverse: + "[| f -- a --NS> L; L \ 0 |] + ==> (%x. inverse(f(x))) -- a --NS> (inverse L)" +apply (simp add: NSLIM_def, clarify) +apply (drule spec) +apply (auto simp add: hypreal_of_real_approx_inverse) +done + +lemma LIM_inverse: "[| f -- a --> L; L \ 0 |] ==> (%x. inverse(f(x))) -- a --> (inverse L)" +by (simp add: LIM_NSLIM_iff NSLIM_inverse) + + +lemma NSLIM_zero: + assumes f: "f -- a --NS> l" shows "(%x. f(x) + -l) -- a --NS> 0" +proof -; + have "(\x. f x + - l) -- a --NS> l + -l" + by (rule NSLIM_add_minus [OF f NSLIM_const]) + thus ?thesis by simp +qed + +lemma LIM_zero2: "f -- a --> l ==> (%x. f(x) + -l) -- a --> 0" +by (simp add: LIM_NSLIM_iff NSLIM_zero) + +lemma NSLIM_zero_cancel: "(%x. f(x) - l) -- x --NS> 0 ==> f -- x --NS> l" +apply (drule_tac g = "%x. l" and m = l in NSLIM_add) +apply (auto simp add: diff_minus add_assoc) +done + +lemma LIM_zero_cancel: "(%x. f(x) - l) -- x --> 0 ==> f -- x --> l" +apply (drule_tac g = "%x. l" and M = l in LIM_add) +apply (auto simp add: diff_minus add_assoc) +done + + + +lemma NSLIM_not_zero: "k \ 0 ==> ~ ((%x. k) -- x --NS> 0)" +apply (simp add: NSLIM_def) +apply (rule_tac x = "hypreal_of_real x + epsilon" in exI) +apply (auto intro: Infinitesimal_add_approx_self [THEN approx_sym] + simp add: hypreal_epsilon_not_zero) +done + +lemma NSLIM_const_not_eq: "k \ L ==> ~ ((%x. k) -- x --NS> L)" +apply (simp add: NSLIM_def) +apply (rule_tac x = "hypreal_of_real x + epsilon" in exI) +apply (auto intro: Infinitesimal_add_approx_self [THEN approx_sym] + simp add: hypreal_epsilon_not_zero) +done + +lemma NSLIM_const_eq: "(%x. k) -- x --NS> L ==> k = L" +apply (rule ccontr) +apply (blast dest: NSLIM_const_not_eq) +done + +(* can actually be proved more easily by unfolding def! *) +lemma NSLIM_unique: "[| f -- x --NS> L; f -- x --NS> M |] ==> L = M" +apply (drule NSLIM_minus) +apply (drule NSLIM_add, assumption) +apply (auto dest!: NSLIM_const_eq [symmetric]) +done + +lemma LIM_unique2: "[| f -- x --> L; f -- x --> M |] ==> L = M" +by (simp add: LIM_NSLIM_iff NSLIM_unique) + + +lemma NSLIM_mult_zero: "[| f -- x --NS> 0; g -- x --NS> 0 |] ==> (%x. f(x)*g(x)) -- x --NS> 0" +by (drule NSLIM_mult, auto) + +(* we can use the corresponding thm LIM_mult2 *) +(* for standard definition of limit *) + +lemma LIM_mult_zero2: "[| f -- x --> 0; g -- x --> 0 |] ==> (%x. f(x)*g(x)) -- x --> 0" +by (drule LIM_mult2, auto) + + +lemma NSLIM_self: "(%x. x) -- a --NS> a" +by (simp add: NSLIM_def) + + +(*----------------------------------------------------------------------------- + Derivatives and Continuity - NS and Standard properties + -----------------------------------------------------------------------------*) +text{*Continuity*} + +lemma isNSContD: "[| isNSCont f a; y \ hypreal_of_real a |] ==> ( *f* f) y \ hypreal_of_real (f a)" +by (simp add: isNSCont_def) + +lemma isNSCont_NSLIM: "isNSCont f a ==> f -- a --NS> (f a) " +by (simp add: isNSCont_def NSLIM_def) + +lemma NSLIM_isNSCont: "f -- a --NS> (f a) ==> isNSCont f a" +apply (simp add: isNSCont_def NSLIM_def, auto) +apply (rule_tac Q = "y = hypreal_of_real a" in excluded_middle [THEN disjE], auto) +done + +(*----------------------------------------------------- + NS continuity can be defined using NS Limit in + similar fashion to standard def of continuity + -----------------------------------------------------*) +lemma isNSCont_NSLIM_iff: "(isNSCont f a) = (f -- a --NS> (f a))" +by (blast intro: isNSCont_NSLIM NSLIM_isNSCont) + +(*---------------------------------------------- + Hence, NS continuity can be given + in terms of standard limit + ---------------------------------------------*) +lemma isNSCont_LIM_iff: "(isNSCont f a) = (f -- a --> (f a))" +by (simp add: LIM_NSLIM_iff isNSCont_NSLIM_iff) + +(*----------------------------------------------- + Moreover, it's trivial now that NS continuity + is equivalent to standard continuity + -----------------------------------------------*) +lemma isNSCont_isCont_iff: "(isNSCont f a) = (isCont f a)" +apply (simp add: isCont_def) +apply (rule isNSCont_LIM_iff) +done + +(*---------------------------------------- + Standard continuity ==> NS continuity + ----------------------------------------*) +lemma isCont_isNSCont: "isCont f a ==> isNSCont f a" +by (erule isNSCont_isCont_iff [THEN iffD2]) + +(*---------------------------------------- + NS continuity ==> Standard continuity + ----------------------------------------*) +lemma isNSCont_isCont: "isNSCont f a ==> isCont f a" +by (erule isNSCont_isCont_iff [THEN iffD1]) + +text{*Alternative definition of continuity*} +(* Prove equivalence between NS limits - *) +(* seems easier than using standard def *) +lemma NSLIM_h_iff: "(f -- a --NS> L) = ((%h. f(a + h)) -- 0 --NS> L)" +apply (simp add: NSLIM_def, auto) +apply (drule_tac x = "hypreal_of_real a + x" in spec) +apply (drule_tac [2] x = "-hypreal_of_real a + x" in spec, safe, simp) +apply (rule mem_infmal_iff [THEN iffD2, THEN Infinitesimal_add_approx_self [THEN approx_sym]]) +apply (rule_tac [4] approx_minus_iff2 [THEN iffD1]) + prefer 3 apply (simp add: add_commute) +apply (rule_tac [2] z = x in eq_Abs_hypreal) +apply (rule_tac [4] z = x in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_of_real_def hypreal_minus hypreal_add add_assoc approx_refl hypreal_zero_def) +done + +lemma NSLIM_isCont_iff: "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)" +by (rule NSLIM_h_iff) + +lemma LIM_isCont_iff: "(f -- a --> f a) = ((%h. f(a + h)) -- 0 --> f(a))" +by (simp add: LIM_NSLIM_iff NSLIM_isCont_iff) + +lemma isCont_iff: "(isCont f x) = ((%h. f(x + h)) -- 0 --> f(x))" +by (simp add: isCont_def LIM_isCont_iff) + +(*-------------------------------------------------------------------------- + Immediate application of nonstandard criterion for continuity can offer + very simple proofs of some standard property of continuous functions + --------------------------------------------------------------------------*) +text{*sum continuous*} +lemma isCont_add: "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) + g(x)) a" +by (auto intro: approx_add simp add: isNSCont_isCont_iff [symmetric] isNSCont_def) + +text{*mult continuous*} +lemma isCont_mult: "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) * g(x)) a" +by (auto intro!: starfun_mult_HFinite_approx + simp del: starfun_mult [symmetric] + simp add: isNSCont_isCont_iff [symmetric] isNSCont_def) + +(*------------------------------------------- + composition of continuous functions + Note very short straightforard proof! + ------------------------------------------*) +lemma isCont_o: "[| isCont f a; isCont g (f a) |] ==> isCont (g o f) a" +by (auto simp add: isNSCont_isCont_iff [symmetric] isNSCont_def starfun_o [symmetric]) + +lemma isCont_o2: "[| isCont f a; isCont g (f a) |] ==> isCont (%x. g (f x)) a" +by (auto dest: isCont_o simp add: o_def) + +lemma isNSCont_minus: "isNSCont f a ==> isNSCont (%x. - f x) a" +by (simp add: isNSCont_def) + +lemma isCont_minus: "isCont f a ==> isCont (%x. - f x) a" +by (auto simp add: isNSCont_isCont_iff [symmetric] isNSCont_minus) + +lemma isCont_inverse: + "[| isCont f x; f x \ 0 |] ==> isCont (%x. inverse (f x)) x" +apply (simp add: isCont_def) +apply (blast intro: LIM_inverse) +done + +lemma isNSCont_inverse: "[| isNSCont f x; f x \ 0 |] ==> isNSCont (%x. inverse (f x)) x" +by (auto intro: isCont_inverse simp add: isNSCont_isCont_iff) + +lemma isCont_diff: + "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) - g(x)) a" +apply (simp add: diff_minus) +apply (auto intro: isCont_add isCont_minus) +done + +lemma isCont_const: "isCont (%x. k) a" +by (simp add: isCont_def) +declare isCont_const [simp] + +lemma isNSCont_const: "isNSCont (%x. k) a" +by (simp add: isNSCont_def) +declare isNSCont_const [simp] + +lemma isNSCont_rabs: "isNSCont abs a" +apply (simp add: isNSCont_def) +apply (auto intro: approx_hrabs simp add: hypreal_of_real_hrabs [symmetric] starfun_rabs_hrabs) +done +declare isNSCont_rabs [simp] + +lemma isCont_rabs: "isCont abs a" +by (auto simp add: isNSCont_isCont_iff [symmetric]) +declare isCont_rabs [simp] + +(**************************************************************** +(%* Leave as commented until I add topology theory or remove? *%) +(%*------------------------------------------------------------ + Elementary topology proof for a characterisation of + continuity now: a function f is continuous if and only + if the inverse image, {x. f(x) \ A}, of any open set A + is always an open set + ------------------------------------------------------------*%) +Goal "[| isNSopen A; \x. isNSCont f x |] + ==> isNSopen {x. f x \ A}" +by (auto_tac (claset(),simpset() addsimps [isNSopen_iff1])); +by (dtac (mem_monad_approx RS approx_sym); +by (dres_inst_tac [("x","a")] spec 1); +by (dtac isNSContD 1 THEN assume_tac 1) +by (dtac bspec 1 THEN assume_tac 1) +by (dres_inst_tac [("x","( *f* f) x")] approx_mem_monad2 1); +by (blast_tac (claset() addIs [starfun_mem_starset]); +qed "isNSCont_isNSopen"; + +Goalw [isNSCont_def] + "\A. isNSopen A --> isNSopen {x. f x \ A} \ +\ ==> isNSCont f x"; +by (auto_tac (claset() addSIs [(mem_infmal_iff RS iffD1) RS + (approx_minus_iff RS iffD2)],simpset() addsimps + [Infinitesimal_def,SReal_iff])); +by (dres_inst_tac [("x","{z. abs(z + -f(x)) < ya}")] spec 1); +by (etac (isNSopen_open_interval RSN (2,impE)); +by (auto_tac (claset(),simpset() addsimps [isNSopen_def,isNSnbhd_def])); +by (dres_inst_tac [("x","x")] spec 1); +by (auto_tac (claset() addDs [approx_sym RS approx_mem_monad], + simpset() addsimps [hypreal_of_real_zero RS sym,STAR_starfun_rabs_add_minus])); +qed "isNSopen_isNSCont"; + +Goal "(\x. isNSCont f x) = \ +\ (\A. isNSopen A --> isNSopen {x. f(x) \ A})"; +by (blast_tac (claset() addIs [isNSCont_isNSopen, + isNSopen_isNSCont]); +qed "isNSCont_isNSopen_iff"; + +(%*------- Standard version of same theorem --------*%) +Goal "(\x. isCont f x) = \ +\ (\A. isopen A --> isopen {x. f(x) \ A})"; +by (auto_tac (claset() addSIs [isNSCont_isNSopen_iff], + simpset() addsimps [isNSopen_isopen_iff RS sym, + isNSCont_isCont_iff RS sym])); +qed "isCont_isopen_iff"; +*******************************************************************) + +text{*Uniform continuity*} +lemma isNSUContD: "[| isNSUCont f; x \ y|] ==> ( *f* f) x \ ( *f* f) y" +by (simp add: isNSUCont_def) + +lemma isUCont_isCont: "isUCont f ==> isCont f x" +by (simp add: isUCont_def isCont_def LIM_def, meson) + +lemma isUCont_isNSUCont: "isUCont f ==> isNSUCont f" +apply (simp add: isNSUCont_def isUCont_def approx_def) +apply (simp add: Infinitesimal_FreeUltrafilterNat_iff, safe) +apply (rule_tac z = x in eq_Abs_hypreal) +apply (rule_tac z = y in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_minus hypreal_add) +apply (rule bexI, rule_tac [2] lemma_hyprel_refl, safe) +apply (drule_tac x = u in spec, clarify) +apply (drule_tac x = s in spec, clarify) +apply (subgoal_tac "\n::nat. abs ((xa n) + - (xb n)) < s --> abs (f (xa n) + - f (xb n)) < u") +prefer 2 apply blast +apply (erule_tac V = "\x y. \x + - y\ < s --> \f x + - f y\ < u" in thin_rl) +apply (drule FreeUltrafilterNat_all, ultra) +done + +lemma lemma_LIMu: "\s. 0 < s --> (\z y. \z + - y\ < s & r \ \f z + -f y\) + ==> \n::nat. \z y. + \z + -y\ < inverse(real(Suc n)) & + r \ \f z + -f y\" +apply clarify +apply (cut_tac n1 = n in real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive], auto) +done + +lemma lemma_skolemize_LIM2u: "\s. 0 < s --> (\z y. \z + - y\ < s & r \ \f z + -f y\) + ==> \X Y. \n::nat. + abs(X n + -(Y n)) < inverse(real(Suc n)) & + r \ abs(f (X n) + -f (Y n))" +apply (drule lemma_LIMu) +apply (drule choice, safe) +apply (drule choice, blast) +done + +lemma lemma_simpu: "\n. \X n + -Y n\ < inverse (real(Suc n)) & + r \ abs (f (X n) + - f(Y n)) ==> + \n. \X n + - Y n\ < inverse (real(Suc n))" +apply auto +done + +lemma isNSUCont_isUCont: + "isNSUCont f ==> isUCont f" +apply (simp add: isNSUCont_def isUCont_def approx_def) +apply (simp add: Infinitesimal_FreeUltrafilterNat_iff, safe) +apply (rule ccontr, simp) +apply (simp add: linorder_not_less) +apply (drule lemma_skolemize_LIM2u, safe) +apply (drule_tac x = "Abs_hypreal (hyprel``{X}) " in spec) +apply (drule_tac x = "Abs_hypreal (hyprel``{Y}) " in spec) +apply (simp add: starfun hypreal_minus hypreal_add, auto) +apply (drule lemma_simpu [THEN real_seq_to_hypreal_Infinitesimal2]) +apply (simp add: Infinitesimal_FreeUltrafilterNat_iff hypreal_minus hypreal_add, blast) +apply (rotate_tac 2) +apply (drule_tac x = r in spec, clarify) +apply (drule FreeUltrafilterNat_all, ultra) +done + +text{*Derivatives*} +lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --> D)" +by (simp add: deriv_def) + +lemma DERIV_NS_iff: + "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --NS> D)" +by (simp add: deriv_def LIM_NSLIM_iff) + +lemma DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) + - f(x))/h) -- 0 --> D" +by (simp add: deriv_def) + +lemma NS_DERIV_D: "DERIV f x :> D ==> + (%h. (f(x + h) + - f(x))/h) -- 0 --NS> D" +by (simp add: deriv_def LIM_NSLIM_iff) + +subsubsection{*Uniqueness*} + +lemma DERIV_unique: + "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E" +apply (simp add: deriv_def) +apply (blast intro: LIM_unique) +done + +lemma NSDeriv_unique: + "[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E" +apply (simp add: nsderiv_def) +apply (cut_tac Infinitesimal_epsilon hypreal_epsilon_not_zero) +apply (auto dest!: bspec [where x=epsilon] + intro!: inj_hypreal_of_real [THEN injD] + dest: approx_trans3) +done + +subsubsection{*Differentiable*} + +lemma differentiableD: "f differentiable x ==> \D. DERIV f x :> D" +by (simp add: differentiable_def) + +lemma differentiableI: "DERIV f x :> D ==> f differentiable x" +by (force simp add: differentiable_def) + +lemma NSdifferentiableD: "f NSdifferentiable x ==> \D. NSDERIV f x :> D" +by (simp add: NSdifferentiable_def) + +lemma NSdifferentiableI: "NSDERIV f x :> D ==> f NSdifferentiable x" +by (force simp add: NSdifferentiable_def) + +subsubsection{*Alternative definition for differentiability*} + +lemma LIM_I: + "(!!r. 0 (\s. 0 < s & (\x. x \ a & \x-a\ < s --> \f x - L\ < r))) + ==> f -- a --> L" +by (simp add: LIM_eq) + +lemma DERIV_LIM_iff: + "((%h. (f(a + h) - f(a)) / h) -- 0 --> D) = + ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)" +proof (intro iffI LIM_I) + fix r::real + assume r: "0h. (f (a + h) - f a) / h) -- 0 --> D" + from LIM_D [OF this r] + obtain s + where s: "0 < s" + and s_lt: "\x. x \ 0 & \x\ < s --> \(f (a + x) - f a) / x - D\ < r" + by auto + show "\s. 0 < s \ + (\x. x \ a \ \x-a\ < s \ \(f x - f a) / (x-a) - D\ < r)" + proof (intro exI conjI strip) + show "0 < s" by (rule s) + next + fix x::real + assume "x \ a \ \x-a\ < s" + with s_lt [THEN spec [where x="x-a"]] + show "\(f x - f a) / (x-a) - D\ < r" by auto + qed +next + fix r::real + assume r: "0x. (f x - f a) / (x-a)) -- a --> D" + from LIM_D [OF this r] + obtain s + where s: "0 < s" + and s_lt: "\x. x \ a & \x-a\ < s --> \(f x - f a)/(x-a) - D\ < r" + by auto + show "\s. 0 < s \ + (\x. x \ 0 & \x - 0\ < s --> \(f (a + x) - f a) / x - D\ < r)" + proof (intro exI conjI strip) + show "0 < s" by (rule s) + next + fix x::real + assume "x \ 0 \ \x - 0\ < s" + with s_lt [THEN spec [where x="x+a"]] + show "\(f (a + x) - f a) / x - D\ < r" by (auto simp add: add_ac) + qed +qed + +lemma DERIV_iff2: "(DERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --> D)" +by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff) + + +subsection{*Equivalence of NS and standard definitions of differentiation*} + +text{*First NSDERIV in terms of NSLIM*} + +(*--- first equivalence ---*) +lemma NSDERIV_NSLIM_iff: + "(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --NS> D)" +apply (simp add: nsderiv_def NSLIM_def, auto) +apply (drule_tac x = xa in bspec) +apply (rule_tac [3] ccontr) +apply (drule_tac [3] x = h in spec) +apply (auto simp add: mem_infmal_iff starfun_lambda_cancel) +done + +(*--- second equivalence ---*) +lemma NSDERIV_NSLIM_iff2: + "(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --NS> D)" +by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff diff_minus [symmetric] + LIM_NSLIM_iff [symmetric]) + +(* while we're at it! *) +lemma NSDERIV_iff2: + "(NSDERIV f x :> D) = + (\w. + w \ hypreal_of_real x & w \ hypreal_of_real x --> + ( *f* (%z. (f z - f x) / (z-x))) w \ hypreal_of_real D)" +by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def) + +(*FIXME DELETE*) +lemma hypreal_not_eq_minus_iff: "(x \ a) = (x + -a \ (0::hypreal))" +by (auto dest: hypreal_eq_minus_iff [THEN iffD2]) + +lemma NSDERIVD5: + "(NSDERIV f x :> D) ==> + (\u. u \ hypreal_of_real x --> + ( *f* (%z. f z - f x)) u \ hypreal_of_real D * (u - hypreal_of_real x))" +apply (auto simp add: NSDERIV_iff2) +apply (case_tac "u = hypreal_of_real x", auto) +apply (drule_tac x = u in spec, auto) +apply (drule_tac c = "u - hypreal_of_real x" and b = "hypreal_of_real D" in approx_mult1) +apply (drule_tac [!] hypreal_not_eq_minus_iff [THEN iffD1]) +apply (subgoal_tac [2] "( *f* (%z. z-x)) u \ (0::hypreal) ") +apply (auto simp add: diff_minus + approx_minus_iff [THEN iffD1, THEN mem_infmal_iff [THEN iffD2]] + Infinitesimal_subset_HFinite [THEN subsetD]) +done + +lemma NSDERIVD4: + "(NSDERIV f x :> D) ==> + (\h \ Infinitesimal. + (( *f* f)(hypreal_of_real x + h) - + hypreal_of_real (f x))\ (hypreal_of_real D) * h)" +apply (auto simp add: nsderiv_def) +apply (case_tac "h = (0::hypreal) ") +apply (auto simp add: diff_minus) +apply (drule_tac x = h in bspec) +apply (drule_tac [2] c = h in approx_mult1) +apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] + simp add: diff_minus) +done + +lemma NSDERIVD3: + "(NSDERIV f x :> D) ==> + (\h \ Infinitesimal - {0}. + (( *f* f)(hypreal_of_real x + h) - + hypreal_of_real (f x))\ (hypreal_of_real D) * h)" +apply (auto simp add: nsderiv_def) +apply (rule ccontr, drule_tac x = h in bspec) +apply (drule_tac [2] c = h in approx_mult1) +apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] + simp add: mult_assoc diff_minus) +done + +text{*Now equivalence between NSDERIV and DERIV*} +lemma NSDERIV_DERIV_iff: "(NSDERIV f x :> D) = (DERIV f x :> D)" +by (simp add: deriv_def NSDERIV_NSLIM_iff LIM_NSLIM_iff) + +(*--------------------------------------------------- + Differentiability implies continuity + nice and simple "algebraic" proof + --------------------------------------------------*) +lemma NSDERIV_isNSCont: "NSDERIV f x :> D ==> isNSCont f x" +apply (auto simp add: nsderiv_def isNSCont_NSLIM_iff NSLIM_def) +apply (drule approx_minus_iff [THEN iffD1]) +apply (drule hypreal_not_eq_minus_iff [THEN iffD1]) +apply (drule_tac x = "-hypreal_of_real x + xa" in bspec) + prefer 2 apply (simp add: add_assoc [symmetric]) +apply (auto simp add: mem_infmal_iff [symmetric] hypreal_add_commute) +apply (drule_tac c = "xa + -hypreal_of_real x" in approx_mult1) +apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] + simp add: mult_assoc) +apply (drule_tac x3=D in + HFinite_hypreal_of_real [THEN [2] Infinitesimal_HFinite_mult, + THEN mem_infmal_iff [THEN iffD1]]) +apply (auto simp add: mult_commute + intro: approx_trans approx_minus_iff [THEN iffD2]) +done + +text{*Now Sandard proof*} +lemma DERIV_isCont: "DERIV f x :> D ==> isCont f x" +by (simp add: NSDERIV_DERIV_iff [symmetric] isNSCont_isCont_iff [symmetric] + NSDERIV_isNSCont) + + +(*---------------------------------------------------------------------------- + Differentiation rules for combinations of functions + follow from clear, straightforard, algebraic + manipulations + ----------------------------------------------------------------------------*) +text{*Constant function*} + +(* use simple constant nslimit theorem *) +lemma NSDERIV_const: "(NSDERIV (%x. k) x :> 0)" +by (simp add: NSDERIV_NSLIM_iff) +declare NSDERIV_const [simp] + +lemma DERIV_const: "(DERIV (%x. k) x :> 0)" +by (simp add: NSDERIV_DERIV_iff [symmetric]) +declare DERIV_const [simp] + +(*----------------------------------------------------- + Sum of functions- proved easily + ----------------------------------------------------*) + + +lemma NSDERIV_add: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] + ==> NSDERIV (%x. f x + g x) x :> Da + Db" +apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) +apply (auto simp add: add_divide_distrib dest!: spec) +apply (drule_tac b = "hypreal_of_real Da" and d = "hypreal_of_real Db" in approx_add) +apply (auto simp add: add_ac) +done + +(* Standard theorem *) +lemma DERIV_add: "[| DERIV f x :> Da; DERIV g x :> Db |] + ==> DERIV (%x. f x + g x) x :> Da + Db" +apply (simp add: NSDERIV_add NSDERIV_DERIV_iff [symmetric]) +done + +(*----------------------------------------------------- + Product of functions - Proof is trivial but tedious + and long due to rearrangement of terms + ----------------------------------------------------*) + +lemma lemma_nsderiv1: "((a::hypreal)*b) + -(c*d) = (b*(a + -c)) + (c*(b + -d))" +by (simp add: right_distrib) + +lemma lemma_nsderiv2: "[| (x + y) / z = hypreal_of_real D + yb; z \ 0; + z \ Infinitesimal; yb \ Infinitesimal |] + ==> x + y \ 0" +apply (frule_tac c1 = z in hypreal_mult_right_cancel [THEN iffD2], assumption) +apply (erule_tac V = " (x + y) / z = hypreal_of_real D + yb" in thin_rl) +apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add + simp add: hypreal_mult_assoc mem_infmal_iff [symmetric]) +apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) +done + + +lemma NSDERIV_mult: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] + ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" +apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) +apply (auto dest!: spec + simp add: starfun_lambda_cancel lemma_nsderiv1) +apply (simp (no_asm) add: add_divide_distrib) +apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+ +apply (auto simp del: times_divide_eq_right simp add: times_divide_eq_right [symmetric]) +apply (drule_tac D = Db in lemma_nsderiv2) +apply (drule_tac [4] + approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]]) +apply (auto intro!: approx_add_mono1 + simp add: left_distrib right_distrib mult_commute add_assoc) +apply (rule_tac b1 = "hypreal_of_real Db * hypreal_of_real (f x)" + in add_commute [THEN subst]) +apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym] + Infinitesimal_add Infinitesimal_mult + Infinitesimal_hypreal_of_real_mult + Infinitesimal_hypreal_of_real_mult2 + simp add: add_assoc [symmetric]) +done + +lemma DERIV_mult: + "[| DERIV f x :> Da; DERIV g x :> Db |] + ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" +by (simp add: NSDERIV_mult NSDERIV_DERIV_iff [symmetric]) + +text{*Multiplying by a constant*} +lemma NSDERIV_cmult: "NSDERIV f x :> D + ==> NSDERIV (%x. c * f x) x :> c*D" +apply (simp only: times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff + minus_mult_right right_distrib [symmetric]) +apply (erule NSLIM_const [THEN NSLIM_mult]) +done + +(* let's do the standard proof though theorem *) +(* LIM_mult2 follows from a NS proof *) + +lemma DERIV_cmult: + "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D" +apply (simp only: deriv_def times_divide_eq_right [symmetric] + NSDERIV_NSLIM_iff minus_mult_right right_distrib [symmetric]) +apply (erule LIM_const [THEN LIM_mult2]) +done + +text{*Negation of function*} +lemma NSDERIV_minus: "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D" +proof (simp add: NSDERIV_NSLIM_iff) + assume "(\h. (f (x + h) + - f x) / h) -- 0 --NS> D" + hence deriv: "(\h. - ((f(x+h) + - f x) / h)) -- 0 --NS> - D" + by (rule NSLIM_minus) + have "\h. - ((f (x + h) + - f x) / h) = (- f (x + h) + f x) / h" + by (simp add: minus_divide_left) + with deriv + show "(\h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp +qed + + +lemma DERIV_minus: "DERIV f x :> D ==> DERIV (%x. -(f x)) x :> -D" +by (simp add: NSDERIV_minus NSDERIV_DERIV_iff [symmetric]) + +text{*Subtraction*} +lemma NSDERIV_add_minus: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x + -g x) x :> Da + -Db" +by (blast dest: NSDERIV_add NSDERIV_minus) + +lemma DERIV_add_minus: "[| DERIV f x :> Da; DERIV g x :> Db |] ==> DERIV (%x. f x + -g x) x :> Da + -Db" +by (blast dest: DERIV_add DERIV_minus) + +lemma NSDERIV_diff: + "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] + ==> NSDERIV (%x. f x - g x) x :> Da-Db" +apply (simp add: diff_minus) +apply (blast intro: NSDERIV_add_minus) +done + +lemma DERIV_diff: + "[| DERIV f x :> Da; DERIV g x :> Db |] + ==> DERIV (%x. f x - g x) x :> Da-Db" +apply (simp add: diff_minus) +apply (blast intro: DERIV_add_minus) +done + +(*--------------------------------------------------------------- + (NS) Increment + ---------------------------------------------------------------*) +lemma incrementI: + "f NSdifferentiable x ==> + increment f x h = ( *f* f) (hypreal_of_real(x) + h) + + -hypreal_of_real (f x)" +by (simp add: increment_def) + +lemma incrementI2: "NSDERIV f x :> D ==> + increment f x h = ( *f* f) (hypreal_of_real(x) + h) + + -hypreal_of_real (f x)" +apply (erule NSdifferentiableI [THEN incrementI]) +done + +(* The Increment theorem -- Keisler p. 65 *) +lemma increment_thm: "[| NSDERIV f x :> D; h \ Infinitesimal; h \ 0 |] + ==> \e \ Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h" +apply (frule_tac h = h in incrementI2, simp add: nsderiv_def) +apply (drule bspec, auto) +apply (drule bex_Infinitesimal_iff2 [THEN iffD2], clarify) +apply (frule_tac b1 = "hypreal_of_real (D) + y" + in hypreal_mult_right_cancel [THEN iffD2]) +apply (erule_tac [2] V = "(( *f* f) (hypreal_of_real (x) + h) + - hypreal_of_real (f x)) / h = hypreal_of_real (D) + y" in thin_rl) +apply assumption +apply (simp add: times_divide_eq_right [symmetric] del: times_divide_eq_right) +apply (auto simp add: left_distrib) +done + +lemma increment_thm2: + "[| NSDERIV f x :> D; h \ 0; h \ 0 |] + ==> \e \ Infinitesimal. increment f x h = + hypreal_of_real(D)*h + e*h" +by (blast dest!: mem_infmal_iff [THEN iffD2] intro!: increment_thm) + + +lemma increment_approx_zero: "[| NSDERIV f x :> D; h \ 0; h \ 0 |] + ==> increment f x h \ 0" +apply (drule increment_thm2, + auto intro!: Infinitesimal_HFinite_mult2 HFinite_add simp add: left_distrib [symmetric] mem_infmal_iff [symmetric]) +apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) +done + +text{* Similarly to the above, the chain rule admits an entirely + straightforward derivation. Compare this with Harrison's + HOL proof of the chain rule, which proved to be trickier and + required an alternative characterisation of differentiability- + the so-called Carathedory derivative. Our main problem is + manipulation of terms.*} + + +(* lemmas *) +lemma NSDERIV_zero: + "[| NSDERIV g x :> D; + ( *f* g) (hypreal_of_real(x) + xa) = hypreal_of_real(g x); + xa \ Infinitesimal; + xa \ 0 + |] ==> D = 0" +apply (simp add: nsderiv_def) +apply (drule bspec, auto) +done + +(* can be proved differently using NSLIM_isCont_iff *) +lemma NSDERIV_approx: + "[| NSDERIV f x :> D; h \ Infinitesimal; h \ 0 |] + ==> ( *f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) \ 0" +apply (simp add: nsderiv_def) +apply (simp add: mem_infmal_iff [symmetric]) +apply (rule Infinitesimal_ratio) +apply (rule_tac [3] approx_hypreal_of_real_HFinite, auto) +done + +(*--------------------------------------------------------------- + from one version of differentiability + + f(x) - f(a) + --------------- \ Db + x - a + ---------------------------------------------------------------*) +lemma NSDERIVD1: "[| NSDERIV f (g x) :> Da; + ( *f* g) (hypreal_of_real(x) + xa) \ hypreal_of_real (g x); + ( *f* g) (hypreal_of_real(x) + xa) \ hypreal_of_real (g x) + |] ==> (( *f* f) (( *f* g) (hypreal_of_real(x) + xa)) + + - hypreal_of_real (f (g x))) + / (( *f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real (g x)) + \ hypreal_of_real(Da)" +by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def diff_minus [symmetric]) + +(*-------------------------------------------------------------- + from other version of differentiability + + f(x + h) - f(x) + ----------------- \ Db + h + --------------------------------------------------------------*) +lemma NSDERIVD2: "[| NSDERIV g x :> Db; xa \ Infinitesimal; xa \ 0 |] + ==> (( *f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real(g x)) / xa + \ hypreal_of_real(Db)" +by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff starfun_lambda_cancel) + +lemma lemma_chain: "(z::hypreal) \ 0 ==> x*y = (x*inverse(z))*(z*y)" +by auto + +(*------------------------------------------------------ + This proof uses both definitions of differentiability. + ------------------------------------------------------*) +lemma NSDERIV_chain: "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] + ==> NSDERIV (f o g) x :> Da * Db" +apply (simp (no_asm_simp) add: NSDERIV_NSLIM_iff NSLIM_def + mem_infmal_iff [symmetric]) +apply clarify +apply (frule_tac f = g in NSDERIV_approx) +apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric]) +apply (case_tac "( *f* g) (hypreal_of_real (x) + xa) = hypreal_of_real (g x) ") +apply (drule_tac g = g in NSDERIV_zero) +apply (auto simp add: divide_inverse) +apply (rule_tac z1 = "( *f* g) (hypreal_of_real (x) + xa) + -hypreal_of_real (g x) " and y1 = "inverse xa" in lemma_chain [THEN ssubst]) +apply (erule hypreal_not_eq_minus_iff [THEN iffD1]) +apply (rule approx_mult_hypreal_of_real) +apply (simp_all add: divide_inverse [symmetric]) +apply (blast intro: NSDERIVD1 approx_minus_iff [THEN iffD2]) +apply (blast intro: NSDERIVD2) +done + +(* standard version *) +lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db" +by (simp add: NSDERIV_DERIV_iff [symmetric] NSDERIV_chain) + +lemma DERIV_chain2: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db" +by (auto dest: DERIV_chain simp add: o_def) + +text{*Differentiation of natural number powers*} +lemma NSDERIV_Id: "NSDERIV (%x. x) x :> 1" +by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def starfun_Id) +declare NSDERIV_Id [simp] + +(*derivative of the identity function*) +lemma DERIV_Id: "DERIV (%x. x) x :> 1" +by (simp add: NSDERIV_DERIV_iff [symmetric]) +declare DERIV_Id [simp] + +lemmas isCont_Id = DERIV_Id [THEN DERIV_isCont, standard] + +(*derivative of linear multiplication*) +lemma DERIV_cmult_Id: "DERIV (op * c) x :> c" +by (cut_tac c = c and x = x in DERIV_Id [THEN DERIV_cmult], simp) +declare DERIV_cmult_Id [simp] + +lemma NSDERIV_cmult_Id: "NSDERIV (op * c) x :> c" +by (simp add: NSDERIV_DERIV_iff) +declare NSDERIV_cmult_Id [simp] + +lemma DERIV_pow: "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" +apply (induct_tac "n") +apply (drule_tac [2] DERIV_Id [THEN DERIV_mult]) +apply (auto simp add: real_of_nat_Suc left_distrib) +apply (case_tac "0 < n") +apply (drule_tac x = x in realpow_minus_mult) +apply (auto simp add: real_mult_assoc real_add_commute) +done + +(* NS version *) +lemma NSDERIV_pow: "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" +by (simp add: NSDERIV_DERIV_iff DERIV_pow) + +(*--------------------------------------------------------------- + Power of -1 + ---------------------------------------------------------------*) + +(*Can't get rid of x \ 0 because it isn't continuous at zero*) +lemma NSDERIV_inverse: + "x \ 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))" +apply (simp add: nsderiv_def) +apply (rule ballI, simp, clarify) +apply (frule Infinitesimal_add_not_zero) +prefer 2 apply (simp add: add_commute) +apply (auto simp add: starfun_inverse_inverse realpow_two + simp del: minus_mult_left [symmetric] minus_mult_right [symmetric]) +apply (simp add: inverse_add inverse_mult_distrib [symmetric] + inverse_minus_eq [symmetric] add_ac mult_ac + del: inverse_mult_distrib inverse_minus_eq + minus_mult_left [symmetric] minus_mult_right [symmetric]) +apply (simp (no_asm_simp) add: mult_assoc [symmetric] right_distrib + del: minus_mult_left [symmetric] minus_mult_right [symmetric]) +apply (rule_tac y = " inverse (- hypreal_of_real x * hypreal_of_real x) " in approx_trans) +apply (rule inverse_add_Infinitesimal_approx2) +apply (auto dest!: hypreal_of_real_HFinite_diff_Infinitesimal + simp add: inverse_minus_eq [symmetric] HFinite_minus_iff) +apply (rule Infinitesimal_HFinite_mult2, auto) +done + + + + +lemma DERIV_inverse: "x \ 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))" +by (simp add: NSDERIV_inverse NSDERIV_DERIV_iff [symmetric] del: realpow_Suc) + +text{*Derivative of inverse*} +lemma DERIV_inverse_fun: "[| DERIV f x :> d; f(x) \ 0 |] + ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" +apply (simp only: mult_commute [of d] minus_mult_left power_inverse) +apply (fold o_def) +apply (blast intro!: DERIV_chain DERIV_inverse) +done + +lemma NSDERIV_inverse_fun: "[| NSDERIV f x :> d; f(x) \ 0 |] + ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" +by (simp add: NSDERIV_DERIV_iff DERIV_inverse_fun del: realpow_Suc) + +text{*Derivative of quotient*} +lemma DERIV_quotient: "[| DERIV f x :> d; DERIV g x :> e; g(x) \ 0 |] + ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) + -(e*f(x))) / (g(x) ^ Suc (Suc 0))" +apply (drule_tac f = g in DERIV_inverse_fun) +apply (drule_tac [2] DERIV_mult) +apply (assumption+) +apply (simp add: divide_inverse right_distrib power_inverse minus_mult_left + mult_ac + del: realpow_Suc minus_mult_right [symmetric] minus_mult_left [symmetric]) +done + +lemma NSDERIV_quotient: "[| NSDERIV f x :> d; DERIV g x :> e; g(x) \ 0 |] + ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) + + -(e*f(x))) / (g(x) ^ Suc (Suc 0))" +by (simp add: NSDERIV_DERIV_iff DERIV_quotient del: realpow_Suc) + +(* ------------------------------------------------------------------------ *) +(* Caratheodory formulation of derivative at a point: standard proof *) +(* ------------------------------------------------------------------------ *) + +lemma CARAT_DERIV: + "(DERIV f x :> l) = + (\g. (\z. f z - f x = g z * (z-x)) & isCont g x & g x = l)" + (is "?lhs = ?rhs") +proof + assume der: "DERIV f x :> l" + show "\g. (\z. f z - f x = g z * (z-x)) \ isCont g x \ g x = l" + proof (intro exI conjI) + let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" + show "\z. f z - f x = ?g z * (z-x)" by simp + show "isCont ?g x" using der + by (simp add: isCont_iff DERIV_iff diff_minus + cong: LIM_equal [rule_format]) + show "?g x = l" by simp + qed +next + assume "?rhs" + then obtain g where + "(\z. f z - f x = g z * (z-x))" and "isCont g x" and "g x = l" by blast + thus "(DERIV f x :> l)" + by (auto simp add: isCont_iff DERIV_iff diff_minus + cong: LIM_equal [rule_format]) +qed + + +lemma CARAT_NSDERIV: "NSDERIV f x :> l ==> + \g. (\z. f z - f x = g z * (z-x)) & isNSCont g x & g x = l" +by (auto simp add: NSDERIV_DERIV_iff isNSCont_isCont_iff CARAT_DERIV) + +lemma hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))" +by auto + +lemma CARAT_DERIVD: + assumes all: "\z. f z - f x = g z * (z-x)" + and nsc: "isNSCont g x" + shows "NSDERIV f x :> g x" +proof - + from nsc + have "\w. w \ hypreal_of_real x \ w \ hypreal_of_real x \ + ( *f* g) w * (w - hypreal_of_real x) / (w - hypreal_of_real x) \ + hypreal_of_real (g x)" + by (simp add: diff_minus isNSCont_def) + thus ?thesis using all + by (simp add: NSDERIV_iff2 starfun_if_eq cong: if_cong) +qed + +(*--------------------------------------------------------------------------*) +(* Lemmas about nested intervals and proof by bisection (cf.Harrison) *) +(* All considerably tidied by lcp *) +(*--------------------------------------------------------------------------*) + +lemma lemma_f_mono_add [rule_format (no_asm)]: "(\n. (f::nat=>real) n \ f (Suc n)) --> f m \ f(m + no)" +apply (induct_tac "no") +apply (auto intro: order_trans) +done + +lemma f_inc_g_dec_Beq_f: "[| \n. f(n) \ f(Suc n); + \n. g(Suc n) \ g(n); + \n. f(n) \ g(n) |] + ==> Bseq f" +apply (rule_tac k = "f 0" and K = "g 0" in BseqI2, rule allI) +apply (induct_tac "n") +apply (auto intro: order_trans) +apply (rule_tac y = "g (Suc na) " in order_trans) +apply (induct_tac [2] "na") +apply (auto intro: order_trans) +done + +lemma f_inc_g_dec_Beq_g: "[| \n. f(n) \ f(Suc n); + \n. g(Suc n) \ g(n); + \n. f(n) \ g(n) |] + ==> Bseq g" +apply (subst Bseq_minus_iff [symmetric]) +apply (rule_tac g = "%x. - (f x) " in f_inc_g_dec_Beq_f) +apply auto +done + +lemma f_inc_imp_le_lim: "[| \n. f n \ f (Suc n); convergent f |] ==> f n \ lim f" +apply (rule linorder_not_less [THEN iffD1]) +apply (auto simp add: convergent_LIMSEQ_iff LIMSEQ_iff monoseq_Suc) +apply (drule real_less_sum_gt_zero) +apply (drule_tac x = "f n + - lim f" in spec, safe) +apply (drule_tac P = "%na. no\na --> ?Q na" and x = "no + n" in spec, auto) +apply (subgoal_tac "lim f \ f (no + n) ") +apply (induct_tac [2] "no") +apply (auto intro: order_trans simp add: diff_minus real_abs_def) +apply (drule_tac no=no and m=n in lemma_f_mono_add) +apply (auto simp add: add_commute) +done + +lemma lim_uminus: "convergent g ==> lim (%x. - g x) = - (lim g)" +apply (rule LIMSEQ_minus [THEN limI]) +apply (simp add: convergent_LIMSEQ_iff) +done + +lemma g_dec_imp_lim_le: "[| \n. g(Suc n) \ g(n); convergent g |] ==> lim g \ g n" +apply (subgoal_tac "- (g n) \ - (lim g) ") +apply (cut_tac [2] f = "%x. - (g x) " in f_inc_imp_le_lim) +apply (auto simp add: lim_uminus convergent_minus_iff [symmetric]) +done + +lemma lemma_nest: "[| \n. f(n) \ f(Suc n); + \n. g(Suc n) \ g(n); + \n. f(n) \ g(n) |] + ==> \l m. l \ m & ((\n. f(n) \ l) & f ----> l) & + ((\n. m \ g(n)) & g ----> m)" +apply (subgoal_tac "monoseq f & monoseq g") +prefer 2 apply (force simp add: LIMSEQ_iff monoseq_Suc) +apply (subgoal_tac "Bseq f & Bseq g") +prefer 2 apply (blast intro: f_inc_g_dec_Beq_f f_inc_g_dec_Beq_g) +apply (auto dest!: Bseq_monoseq_convergent simp add: convergent_LIMSEQ_iff) +apply (rule_tac x = "lim f" in exI) +apply (rule_tac x = "lim g" in exI) +apply (auto intro: LIMSEQ_le) +apply (auto simp add: f_inc_imp_le_lim g_dec_imp_lim_le convergent_LIMSEQ_iff) +done + +lemma lemma_nest_unique: "[| \n. f(n) \ f(Suc n); + \n. g(Suc n) \ g(n); + \n. f(n) \ g(n); + (%n. f(n) - g(n)) ----> 0 |] + ==> \l. ((\n. f(n) \ l) & f ----> l) & + ((\n. l \ g(n)) & g ----> l)" +apply (drule lemma_nest, auto) +apply (subgoal_tac "l = m") +apply (drule_tac [2] X = f in LIMSEQ_diff) +apply (auto intro: LIMSEQ_unique) +done + +text{*The universal quantifiers below are required for the declaration + of @{text Bolzano_nest_unique} below.*} + +lemma Bolzano_bisect_le: + "a \ b ==> \n. fst (Bolzano_bisect P a b n) \ snd (Bolzano_bisect P a b n)" +apply (rule allI) +apply (induct_tac "n") +apply (auto simp add: Let_def split_def) +done + +lemma Bolzano_bisect_fst_le_Suc: "a \ b ==> + \n. fst(Bolzano_bisect P a b n) \ fst (Bolzano_bisect P a b (Suc n))" +apply (rule allI) +apply (induct_tac "n") +apply (auto simp add: Bolzano_bisect_le Let_def split_def) +done + +lemma Bolzano_bisect_Suc_le_snd: "a \ b ==> + \n. snd(Bolzano_bisect P a b (Suc n)) \ snd (Bolzano_bisect P a b n)" +apply (rule allI) +apply (induct_tac "n") +apply (auto simp add: Bolzano_bisect_le Let_def split_def) +done + +lemma eq_divide_2_times_iff: "((x::real) = y / (2 * z)) = (2 * x = y/z)" +apply auto +apply (drule_tac f = "%u. (1/2) *u" in arg_cong) +apply auto +done + +lemma Bolzano_bisect_diff: + "a \ b ==> + snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = + (b-a) / (2 ^ n)" +apply (induct_tac "n") +apply (auto simp add: eq_divide_2_times_iff add_divide_distrib Let_def split_def) +apply (auto simp add: add_ac Bolzano_bisect_le diff_minus) +done + +lemmas Bolzano_nest_unique = + lemma_nest_unique + [OF Bolzano_bisect_fst_le_Suc Bolzano_bisect_Suc_le_snd Bolzano_bisect_le] + + +lemma not_P_Bolzano_bisect: + assumes P: "!!a b c. [| P(a,b); P(b,c); a \ b; b \ c|] ==> P(a,c)" + and notP: "~ P(a,b)" + and le: "a \ b" + shows "~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" +proof (induct n) + case 0 thus ?case by simp + next + case (Suc n) + thus ?case + by (auto simp del: surjective_pairing [symmetric] + simp add: Let_def split_def Bolzano_bisect_le [OF le] + P [of "fst (Bolzano_bisect P a b n)" _ "snd (Bolzano_bisect P a b n)"]) +qed + +(*Now we re-package P_prem as a formula*) +lemma not_P_Bolzano_bisect': + "[| \a b c. P(a,b) & P(b,c) & a \ b & b \ c --> P(a,c); + ~ P(a,b); a \ b |] ==> + \n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" +by (blast elim!: not_P_Bolzano_bisect [THEN [2] rev_notE]) + + + +lemma lemma_BOLZANO: + "[| \a b c. P(a,b) & P(b,c) & a \ b & b \ c --> P(a,c); + \x. \d::real. 0 < d & + (\a b. a \ x & x \ b & (b-a) < d --> P(a,b)); + a \ b |] + ==> P(a,b)" +apply (rule Bolzano_nest_unique [where P1=P, THEN exE], assumption+) +apply (rule LIMSEQ_minus_cancel) +apply (simp (no_asm_simp) add: Bolzano_bisect_diff LIMSEQ_divide_realpow_zero) +apply (rule ccontr) +apply (drule not_P_Bolzano_bisect', assumption+) +apply (rename_tac "l") +apply (drule_tac x = l in spec, clarify) +apply (simp add: LIMSEQ_def) +apply (drule_tac P = "%r. 0 ?Q r" and x = "d/2" in spec) +apply (drule_tac P = "%r. 0 ?Q r" and x = "d/2" in spec) +apply (drule real_less_half_sum, auto) +apply (drule_tac x = "fst (Bolzano_bisect P a b (no + noa))" in spec) +apply (drule_tac x = "snd (Bolzano_bisect P a b (no + noa))" in spec) +apply safe +apply (simp_all (no_asm_simp)) +apply (rule_tac y = "abs (fst (Bolzano_bisect P a b (no + noa)) - l) + abs (snd (Bolzano_bisect P a b (no + noa)) - l) " in order_le_less_trans) +apply (simp (no_asm_simp) add: abs_if) +apply (rule real_sum_of_halves [THEN subst]) +apply (rule add_strict_mono) +apply (simp_all add: diff_minus [symmetric]) +done + + +lemma lemma_BOLZANO2: "((\a b c. (a \ b & b \ c & P(a,b) & P(b,c)) --> P(a,c)) & + (\x. \d::real. 0 < d & + (\a b. a \ x & x \ b & (b-a) < d --> P(a,b)))) + --> (\a b. a \ b --> P(a,b))" +apply clarify +apply (blast intro: lemma_BOLZANO) +done + + +subsection{*Intermediate Value Theorem: Prove Contrapositive by Bisection*} + +lemma IVT: "[| f(a) \ y; y \ f(b); + a \ b; + (\x. a \ x & x \ b --> isCont f x) |] + ==> \x. a \ x & x \ b & f(x) = y" +apply (rule contrapos_pp, assumption) +apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> ~ (f (u) \ y & y \ f (v))" in lemma_BOLZANO2) +apply safe +apply simp_all +apply (simp add: isCont_iff LIM_def) +apply (rule ccontr) +apply (subgoal_tac "a \ x & x \ b") + prefer 2 + apply simp + apply (drule_tac P = "%d. 0 ?P d" and x = 1 in spec, arith) +apply (drule_tac x = x in spec)+ +apply simp +apply (drule_tac P = "%r. ?P r --> (\s. 0 aa", simp_all) +apply (drule_tac x = x and y = aa in order_antisym) +apply (assumption, simp) +done + +lemma IVT2: "[| f(b) \ y; y \ f(a); + a \ b; + (\x. a \ x & x \ b --> isCont f x) + |] ==> \x. a \ x & x \ b & f(x) = y" +apply (subgoal_tac "- f a \ -y & -y \ - f b", clarify) +apply (drule IVT [where f = "%x. - f x"], assumption) +apply (auto intro: isCont_minus) +done + +(*HOL style here: object-level formulations*) +lemma IVT_objl: "(f(a) \ y & y \ f(b) & a \ b & + (\x. a \ x & x \ b --> isCont f x)) + --> (\x. a \ x & x \ b & f(x) = y)" +apply (blast intro: IVT) +done + +lemma IVT2_objl: "(f(b) \ y & y \ f(a) & a \ b & + (\x. a \ x & x \ b --> isCont f x)) + --> (\x. a \ x & x \ b & f(x) = y)" +apply (blast intro: IVT2) +done + +(*---------------------------------------------------------------------------*) +(* By bisection, function continuous on closed interval is bounded above *) +(*---------------------------------------------------------------------------*) + + +lemma isCont_bounded: + "[| a \ b; \x. a \ x & x \ b --> isCont f x |] + ==> \M. \x. a \ x & x \ b --> f(x) \ M" +apply (cut_tac P = "% (u,v) . a \ u & u \ v & v \ b --> (\M. \x. u \ x & x \ v --> f x \ M) " in lemma_BOLZANO2) +apply safe +apply simp_all +apply (rename_tac x xa ya M Ma) +apply (cut_tac x = M and y = Ma in linorder_linear, safe) +apply (rule_tac x = Ma in exI, clarify) +apply (cut_tac x = xb and y = xa in linorder_linear, force) +apply (rule_tac x = M in exI, clarify) +apply (cut_tac x = xb and y = xa in linorder_linear, force) +apply (case_tac "a \ x & x \ b") +apply (rule_tac [2] x = 1 in exI) +prefer 2 apply force +apply (simp add: LIM_def isCont_iff) +apply (drule_tac x = x in spec, auto) +apply (erule_tac V = "\M. \x. a \ x & x \ b & ~ f x \ M" in thin_rl) +apply (drule_tac x = 1 in spec, auto) +apply (rule_tac x = s in exI, clarify) +apply (rule_tac x = "\f x\ + 1" in exI, clarify) +apply (drule_tac x = "xa-x" in spec) +apply (auto simp add: abs_ge_self, arith+) +done + +(*----------------------------------------------------------------------------*) +(* Refine the above to existence of least upper bound *) +(*----------------------------------------------------------------------------*) + +lemma lemma_reals_complete: "((\x. x \ S) & (\y. isUb UNIV S (y::real))) --> + (\t. isLub UNIV S t)" +apply (blast intro: reals_complete) +done + +lemma isCont_has_Ub: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] + ==> \M. (\x. a \ x & x \ b --> f(x) \ M) & + (\N. N < M --> (\x. a \ x & x \ b & N < f(x)))" +apply (cut_tac S = "Collect (%y. \x. a \ x & x \ b & y = f x) " in lemma_reals_complete) +apply auto +apply (drule isCont_bounded, assumption) +apply (auto simp add: isUb_def leastP_def isLub_def setge_def setle_def) +apply (rule exI, auto) +apply (auto dest!: spec simp add: linorder_not_less) +done + +(*----------------------------------------------------------------------------*) +(* Now show that it attains its upper bound *) +(*----------------------------------------------------------------------------*) + +lemma isCont_eq_Ub: + assumes le: "a \ b" + and con: "\x. a \ x & x \ b --> isCont f x" + shows "\M. (\x. a \ x & x \ b --> f(x) \ M) & + (\x. a \ x & x \ b & f(x) = M)" +proof - + from isCont_has_Ub [OF le con] + obtain M where M1: "\x. a \ x \ x \ b \ f x \ M" + and M2: "!!N. N \x. a \ x \ x \ b \ N < f x" by blast + show ?thesis + proof (intro exI, intro conjI) + show " \x. a \ x \ x \ b \ f x \ M" by (rule M1) + show "\x. a \ x \ x \ b \ f x = M" + proof (rule ccontr) + assume "\ (\x. a \ x \ x \ b \ f x = M)" + with M1 have M3: "\x. a \ x & x \ b --> f x < M" + by (auto simp add: linorder_not_le [symmetric] intro: order_antisym) + hence "\x. a \ x & x \ b --> isCont (%x. inverse (M - f x)) x" + by (auto simp add: isCont_inverse isCont_diff con) + from isCont_bounded [OF le this] + obtain k where k: "!!x. a \ x & x \ b --> inverse (M - f x) \ k" by auto + have Minv: "!!x. a \ x & x \ b --> 0 < inverse (M - f (x))" + by (simp add: M3) + have "!!x. a \ x & x \ b --> inverse (M - f x) < k+1" using k + by (auto intro: order_le_less_trans [of _ k]) + with Minv + have "!!x. a \ x & x \ b --> inverse(k+1) < inverse(inverse(M - f x))" + by (intro strip less_imp_inverse_less, simp_all) + hence invlt: "!!x. a \ x & x \ b --> inverse(k+1) < M - f x" + by simp + have "M - inverse (k+1) < M" using k [of a] Minv [of a] le + by (simp, arith) + from M2 [OF this] + obtain x where ax: "a \ x & x \ b & M - inverse(k+1) < f x" .. + thus False using invlt [of x] by force + qed + qed +qed + + + +(*----------------------------------------------------------------------------*) +(* Same theorem for lower bound *) +(*----------------------------------------------------------------------------*) + +lemma isCont_eq_Lb: "[| a \ b; \x. a \ x & x \ b --> isCont f x |] + ==> \M. (\x. a \ x & x \ b --> M \ f(x)) & + (\x. a \ x & x \ b & f(x) = M)" +apply (subgoal_tac "\x. a \ x & x \ b --> isCont (%x. - (f x)) x") +prefer 2 apply (blast intro: isCont_minus) +apply (drule_tac f = " (%x. - (f x))" in isCont_eq_Ub) +apply safe +apply auto +done + + +(* ------------------------------------------------------------------------- *) +(* Another version. *) +(* ------------------------------------------------------------------------- *) + +lemma isCont_Lb_Ub: "[|a \ b; \x. a \ x & x \ b --> isCont f x |] + ==> \L M. (\x. a \ x & x \ b --> L \ f(x) & f(x) \ M) & + (\y. L \ y & y \ M --> (\x. a \ x & x \ b & (f(x) = y)))" +apply (frule isCont_eq_Lb) +apply (frule_tac [2] isCont_eq_Ub) +apply (assumption+, safe) +apply (rule_tac x = "f x" in exI) +apply (rule_tac x = "f xa" in exI, simp, safe) +apply (cut_tac x = x and y = xa in linorder_linear, safe) +apply (cut_tac f = f and a = x and b = xa and y = y in IVT_objl) +apply (cut_tac [2] f = f and a = xa and b = x and y = y in IVT2_objl, safe) +apply (rule_tac [2] x = xb in exI) +apply (rule_tac [4] x = xb in exI, simp_all) +done + +(*----------------------------------------------------------------------------*) +(* If f'(x) > 0 then x is locally strictly increasing at the right *) +(*----------------------------------------------------------------------------*) + +lemma DERIV_left_inc: + "[| DERIV f x :> l; 0 < l |] + ==> \d. 0 < d & (\h. 0 < h & h < d --> f(x) < f(x + h))" +apply (simp add: deriv_def LIM_def) +apply (drule spec, auto) +apply (rule_tac x = s in exI, auto) +apply (subgoal_tac "0 < l*h") + prefer 2 apply (simp add: zero_less_mult_iff) +apply (drule_tac x = h in spec) +apply (simp add: real_abs_def pos_le_divide_eq pos_less_divide_eq + split add: split_if_asm) +done + +lemma DERIV_left_dec: + assumes der: "DERIV f x :> l" + and l: "l < 0" + shows "\d. 0 < d & (\h. 0 < h & h < d --> f(x) < f(x-h))" +proof - + from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]] + have "\s. 0 < s \ + (\z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < -l)" + by (simp add: diff_minus) + then obtain s + where s: "0 < s" + and all: "!!z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < -l" + by auto + thus ?thesis + proof (intro exI conjI strip) + show "0 h < s" + with all [of "-h"] show "f x < f (x-h)" + proof (simp add: real_abs_def pos_less_divide_eq diff_minus [symmetric] + split add: split_if_asm) + assume "~ l \ - ((f (x-h) - f x) / h)" and h: "0 < h" + with l + have "0 < (f (x-h) - f x) / h" by arith + thus "f x < f (x-h)" + by (simp add: pos_less_divide_eq h) + qed + qed +qed + +lemma DERIV_local_max: + assumes der: "DERIV f x :> l" + and d: "0 < d" + and le: "\y. \x-y\ < d --> f(y) \ f(x)" + shows "l = 0" +proof (cases rule: linorder_cases [of l 0]) + case equal show ?thesis . +next + case less + from DERIV_left_dec [OF der less] + obtain d' where d': "0 < d'" + and lt: "\h. 0 < h \ h < d' \ f x < f (x-h)" by blast + from real_lbound_gt_zero [OF d d'] + obtain e where "0 < e \ e < d \ e < d'" .. + with lt le [THEN spec [where x="x-e"]] + show ?thesis by (auto simp add: abs_if) +next + case greater + from DERIV_left_inc [OF der greater] + obtain d' where d': "0 < d'" + and lt: "\h. 0 < h \ h < d' \ f x < f (x + h)" by blast + from real_lbound_gt_zero [OF d d'] + obtain e where "0 < e \ e < d \ e < d'" .. + with lt le [THEN spec [where x="x+e"]] + show ?thesis by (auto simp add: abs_if) +qed + + +text{*Similar theorem for a local minimum*} +lemma DERIV_local_min: + "[| DERIV f x :> l; 0 < d; \y. \x-y\ < d --> f(x) \ f(y) |] ==> l = 0" +by (drule DERIV_minus [THEN DERIV_local_max], auto) + + +text{*In particular, if a function is locally flat*} +lemma DERIV_local_const: + "[| DERIV f x :> l; 0 < d; \y. \x-y\ < d --> f(x) = f(y) |] ==> l = 0" +by (auto dest!: DERIV_local_max) + +text{*Lemma about introducing open ball in open interval*} +lemma lemma_interval_lt: + "[| a < x; x < b |] + ==> \d::real. 0 < d & (\y. \x-y\ < d --> a < y & y < b)" +apply (simp add: abs_interval_iff) +apply (insert linorder_linear [of "x-a" "b-x"], safe) +apply (rule_tac x = "x-a" in exI) +apply (rule_tac [2] x = "b-x" in exI, auto) +done + +lemma lemma_interval: "[| a < x; x < b |] ==> + \d::real. 0 < d & (\y. \x-y\ < d --> a \ y & y \ b)" +apply (drule lemma_interval_lt, auto) +apply (auto intro!: exI) +done + +text{*Rolle's Theorem. + If @{term f} is defined and continuous on the closed interval + @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"}, + and @{term "f(a) = f(b)"}, + then there exists @{text "x0 \ (a,b)"} such that @{term "f'(x0) = 0"}*} +theorem Rolle: + assumes lt: "a < b" + and eq: "f(a) = f(b)" + and con: "\x. a \ x & x \ b --> isCont f x" + and dif [rule_format]: "\x. a < x & x < b --> f differentiable x" + shows "\z. a < z & z < b & DERIV f z :> 0" +proof - + have le: "a \ b" using lt by simp + from isCont_eq_Ub [OF le con] + obtain x where x_max: "\z. a \ z \ z \ b \ f z \ f x" + and alex: "a \ x" and xleb: "x \ b" + by blast + from isCont_eq_Lb [OF le con] + obtain x' where x'_min: "\z. a \ z \ z \ b \ f x' \ f z" + and alex': "a \ x'" and x'leb: "x' \ b" + by blast + show ?thesis + proof cases + assume axb: "a < x & x < b" + --{*@{term f} attains its maximum within the interval*} + hence ax: "ay. \x-y\ < d \ a \ y \ y \ b" + by blast + hence bound': "\y. \x-y\ < d \ f y \ f x" using x_max + by blast + from differentiableD [OF dif [OF axb]] + obtain l where der: "DERIV f x :> l" .. + have "l=0" by (rule DERIV_local_max [OF der d bound']) + --{*the derivative at a local maximum is zero*} + thus ?thesis using ax xb der by auto + next + assume notaxb: "~ (a < x & x < b)" + hence xeqab: "x=a | x=b" using alex xleb by arith + hence fb_eq_fx: "f b = f x" by (auto simp add: eq) + show ?thesis + proof cases + assume ax'b: "a < x' & x' < b" + --{*@{term f} attains its minimum within the interval*} + hence ax': "ay. \x'-y\ < d \ a \ y \ y \ b" + by blast + hence bound': "\y. \x'-y\ < d \ f x' \ f y" using x'_min + by blast + from differentiableD [OF dif [OF ax'b]] + obtain l where der: "DERIV f x' :> l" .. + have "l=0" by (rule DERIV_local_min [OF der d bound']) + --{*the derivative at a local minimum is zero*} + thus ?thesis using ax' x'b der by auto + next + assume notax'b: "~ (a < x' & x' < b)" + --{*@{term f} is constant througout the interval*} + hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith + hence fb_eq_fx': "f b = f x'" by (auto simp add: eq) + from dense [OF lt] + obtain r where ar: "a < r" and rb: "r < b" by blast + from lemma_interval [OF ar rb] + obtain d where d: "0y. \r-y\ < d \ a \ y \ y \ b" + by blast + have eq_fb: "\z. a \ z --> z \ b --> f z = f b" + proof (clarify) + fix z::real + assume az: "a \ z" and zb: "z \ b" + show "f z = f b" + proof (rule order_antisym) + show "f z \ f b" by (simp add: fb_eq_fx x_max az zb) + show "f b \ f z" by (simp add: fb_eq_fx' x'_min az zb) + qed + qed + have bound': "\y. \r-y\ < d \ f r = f y" + proof (intro strip) + fix y::real + assume lt: "\r-y\ < d" + hence "f y = f b" by (simp add: eq_fb bound) + thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le) + qed + from differentiableD [OF dif [OF conjI [OF ar rb]]] + obtain l where der: "DERIV f r :> l" .. + have "l=0" by (rule DERIV_local_const [OF der d bound']) + --{*the derivative of a constant function is zero*} + thus ?thesis using ar rb der by auto + qed + qed +qed + + +subsection{*Mean Value Theorem*} + +lemma lemma_MVT: + "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)" +proof cases + assume "a=b" thus ?thesis by simp +next + assume "a\b" + hence ba: "b-a \ 0" by arith + show ?thesis + by (rule real_mult_left_cancel [OF ba, THEN iffD1], + simp add: right_diff_distrib, simp add: left_diff_distrib) +qed + +theorem MVT: + assumes lt: "a < b" + and con: "\x. a \ x & x \ b --> isCont f x" + and dif [rule_format]: "\x. a < x & x < b --> f differentiable x" + shows "\l z. a < z & z < b & DERIV f z :> l & + (f(b) - f(a) = (b-a) * l)" +proof - + let ?F = "%x. f x - ((f b - f a) / (b-a)) * x" + have contF: "\x. a \ x \ x \ b \ isCont ?F x" using con + by (fast intro: isCont_diff isCont_const isCont_mult isCont_Id) + have difF: "\x. a < x \ x < b \ ?F differentiable x" + proof (clarify) + fix x::real + assume ax: "a < x" and xb: "x < b" + from differentiableD [OF dif [OF conjI [OF ax xb]]] + obtain l where der: "DERIV f x :> l" .. + show "?F differentiable x" + by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"], + blast intro: DERIV_diff DERIV_cmult_Id der) + qed + from Rolle [where f = ?F, OF lt lemma_MVT contF difF] + obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0" + by blast + have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)" + by (rule DERIV_cmult_Id) + hence derF: "DERIV (\x. ?F x + (f b - f a) / (b - a) * x) z + :> 0 + (f b - f a) / (b - a)" + by (rule DERIV_add [OF der]) + show ?thesis + proof (intro exI conjI) + show "a < z" . + show "z < b" . + show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by simp + show "DERIV f z :> ((f b - f a)/(b-a))" using derF by simp + qed +qed + + +text{*A function is constant if its derivative is 0 over an interval.*} + +lemma DERIV_isconst_end: "[| a < b; + \x. a \ x & x \ b --> isCont f x; + \x. a < x & x < b --> DERIV f x :> 0 |] + ==> (f b = f a)" +apply (drule MVT, assumption) +apply (blast intro: differentiableI) +apply (auto dest!: DERIV_unique simp add: diff_eq_eq) +done + +lemma DERIV_isconst1: "[| a < b; + \x. a \ x & x \ b --> isCont f x; + \x. a < x & x < b --> DERIV f x :> 0 |] + ==> \x. a \ x & x \ b --> f x = f a" +apply safe +apply (drule_tac x = a in order_le_imp_less_or_eq, safe) +apply (drule_tac b = x in DERIV_isconst_end, auto) +done + +lemma DERIV_isconst2: "[| a < b; + \x. a \ x & x \ b --> isCont f x; + \x. a < x & x < b --> DERIV f x :> 0; + a \ x; x \ b |] + ==> f x = f a" +apply (blast dest: DERIV_isconst1) +done + +lemma DERIV_isconst_all: "\x. DERIV f x :> 0 ==> f(x) = f(y)" +apply (rule linorder_cases [of x y]) +apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+ +done + +lemma DERIV_const_ratio_const: + "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k" +apply (rule linorder_cases [of a b], auto) +apply (drule_tac [!] f = f in MVT) +apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def) +apply (auto dest: DERIV_unique simp add: left_distrib diff_minus) +done + +lemma DERIV_const_ratio_const2: + "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k" +apply (rule_tac c1 = "b-a" in real_mult_right_cancel [THEN iffD1]) +apply (auto dest!: DERIV_const_ratio_const simp add: real_mult_assoc) +done + +lemma real_average_minus_first: "((a + b) /2 - a) = (b-a)/(2::real)" +by auto +declare real_average_minus_first [simp] + +lemma real_average_minus_second: "((b + a)/2 - a) = (b-a)/(2::real)" +by auto +declare real_average_minus_second [simp] + +text{*Gallileo's "trick": average velocity = av. of end velocities*} + +lemma DERIV_const_average: + assumes neq: "a \ (b::real)" + and der: "\x. DERIV v x :> k" + shows "v ((a + b)/2) = (v a + v b)/2" +proof (cases rule: linorder_cases [of a b]) + case equal with neq show ?thesis by simp +next + case less + have "(v b - v a) / (b - a) = k" + by (rule DERIV_const_ratio_const2 [OF neq der]) + hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp + moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" + by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) + ultimately show ?thesis using neq by force +next + case greater + have "(v b - v a) / (b - a) = k" + by (rule DERIV_const_ratio_const2 [OF neq der]) + hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp + moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" + by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) + ultimately show ?thesis using neq by (force simp add: add_commute) +qed + + +text{*Dull lemma: an continuous injection on an interval must have a +strict maximum at an end point, not in the middle.*} + +lemma lemma_isCont_inj: + assumes d: "0 < d" + and inj [rule_format]: "\z. \z-x\ \ d --> g(f z) = z" + and cont: "\z. \z-x\ \ d --> isCont f z" + shows "\z. \z-x\ \ d & f x < f z" +proof (rule ccontr) + assume "~ (\z. \z-x\ \ d & f x < f z)" + hence all [rule_format]: "\z. \z - x\ \ d --> f z \ f x" by auto + show False + proof (cases rule: linorder_le_cases [of "f(x-d)" "f(x+d)"]) + case le + from d cont all [of "x+d"] + have flef: "f(x+d) \ f x" + and xlex: "x - d \ x" + and cont': "\z. x - d \ z \ z \ x \ isCont f z" + by (auto simp add: abs_if) + from IVT [OF le flef xlex cont'] + obtain x' where "x-d \ x'" "x' \ x" "f x' = f(x+d)" by blast + moreover + hence "g(f x') = g (f(x+d))" by simp + ultimately show False using d inj [of x'] inj [of "x+d"] + by (simp add: abs_le_interval_iff) + next + case ge + from d cont all [of "x-d"] + have flef: "f(x-d) \ f x" + and xlex: "x \ x+d" + and cont': "\z. x \ z \ z \ x+d \ isCont f z" + by (auto simp add: abs_if) + from IVT2 [OF ge flef xlex cont'] + obtain x' where "x \ x'" "x' \ x+d" "f x' = f(x-d)" by blast + moreover + hence "g(f x') = g (f(x-d))" by simp + ultimately show False using d inj [of x'] inj [of "x-d"] + by (simp add: abs_le_interval_iff) + qed +qed + + +text{*Similar version for lower bound.*} + +lemma lemma_isCont_inj2: + "[|0 < d; \z. \z-x\ \ d --> g(f z) = z; + \z. \z-x\ \ d --> isCont f z |] + ==> \z. \z-x\ \ d & f z < f x" +apply (insert lemma_isCont_inj + [where f = "%x. - f x" and g = "%y. g(-y)" and x = x and d = d]) +apply (simp add: isCont_minus linorder_not_le) +done + +text{*Show there's an interval surrounding @{term "f(x)"} in +@{text "f[[x - d, x + d]]"} .*} + +lemma isCont_inj_range: + assumes d: "0 < d" + and inj: "\z. \z-x\ \ d --> g(f z) = z" + and cont: "\z. \z-x\ \ d --> isCont f z" + shows "\e. 0y. \y - f x\ \ e --> (\z. \z-x\ \ d & f z = y))" +proof - + have "x-d \ x+d" "\z. x-d \ z \ z \ x+d \ isCont f z" using cont d + by (auto simp add: abs_le_interval_iff) + from isCont_Lb_Ub [OF this] + obtain L M + where all1 [rule_format]: "\z. x-d \ z \ z \ x+d \ L \ f z \ f z \ M" + and all2 [rule_format]: + "\y. L \ y \ y \ M \ (\z. x-d \ z \ z \ x+d \ f z = y)" + by auto + with d have "L \ f x & f x \ M" by simp + moreover have "L \ f x" + proof - + from lemma_isCont_inj2 [OF d inj cont] + obtain u where "\u - x\ \ d" "f u < f x" by auto + thus ?thesis using all1 [of u] by arith + qed + moreover have "f x \ M" + proof - + from lemma_isCont_inj [OF d inj cont] + obtain u where "\u - x\ \ d" "f x < f u" by auto + thus ?thesis using all1 [of u] by arith + qed + ultimately have "L < f x & f x < M" by arith + hence "0 < f x - L" "0 < M - f x" by arith+ + from real_lbound_gt_zero [OF this] + obtain e where e: "0 < e" "e < f x - L" "e < M - f x" by auto + thus ?thesis + proof (intro exI conjI) + show "0y. \y - f x\ \ e \ (\z. \z - x\ \ d \ f z = y)" + proof (intro strip) + fix y::real + assume "\y - f x\ \ e" + with e have "L \ y \ y \ M" by arith + from all2 [OF this] + obtain z where "x - d \ z" "z \ x + d" "f z = y" by blast + thus "\z. \z - x\ \ d \ f z = y" + by (force simp add: abs_le_interval_iff) + qed + qed +qed + + +text{*Continuity of inverse function*} + +lemma isCont_inverse_function: + assumes d: "0 < d" + and inj: "\z. \z-x\ \ d --> g(f z) = z" + and cont: "\z. \z-x\ \ d --> isCont f z" + shows "isCont g (f x)" +proof (simp add: isCont_iff LIM_eq) + show "\r. 0 < r \ + (\s. 0 (\z. z\0 \ \z\ < s \ \g(f x + z) - g(f x)\ < r))" + proof (intro strip) + fix r::real + assume r: "0 e < d" by blast + with inj cont + have e_simps: "\z. \z-x\ \ e --> g (f z) = z" + "\z. \z-x\ \ e --> isCont f z" by auto + from isCont_inj_range [OF e this] + obtain e' where e': "0 < e'" + and all: "\y. \y - f x\ \ e' \ (\z. \z - x\ \ e \ f z = y)" + by blast + show "\s. 0 (\z. z\0 \ \z\ < s \ \g(f x + z) - g(f x)\ < r)" + proof (intro exI conjI) + show "0z. z \ 0 \ \z\ < e' \ \g (f x + z) - g (f x)\ < r" + proof (intro strip) + fix z::real + assume z: "z \ 0 \ \z\ < e'" + with e e_lt e_simps all [rule_format, of "f x + z"] + show "\g (f x + z) - g (f x)\ < r" by force + qed + qed + qed +qed + +ML +{* +val LIM_def = thm"LIM_def"; +val NSLIM_def = thm"NSLIM_def"; +val isCont_def = thm"isCont_def"; +val isNSCont_def = thm"isNSCont_def"; +val deriv_def = thm"deriv_def"; +val nsderiv_def = thm"nsderiv_def"; +val differentiable_def = thm"differentiable_def"; +val NSdifferentiable_def = thm"NSdifferentiable_def"; +val increment_def = thm"increment_def"; +val isUCont_def = thm"isUCont_def"; +val isNSUCont_def = thm"isNSUCont_def"; + +val half_gt_zero_iff = thm "half_gt_zero_iff"; +val half_gt_zero = thms "half_gt_zero"; +val abs_diff_triangle_ineq = thm "abs_diff_triangle_ineq"; +val LIM_eq = thm "LIM_eq"; +val LIM_D = thm "LIM_D"; +val LIM_const = thm "LIM_const"; +val LIM_add = thm "LIM_add"; +val LIM_minus = thm "LIM_minus"; +val LIM_add_minus = thm "LIM_add_minus"; +val LIM_diff = thm "LIM_diff"; +val LIM_const_not_eq = thm "LIM_const_not_eq"; +val LIM_const_eq = thm "LIM_const_eq"; +val LIM_unique = thm "LIM_unique"; +val LIM_mult_zero = thm "LIM_mult_zero"; +val LIM_self = thm "LIM_self"; +val LIM_equal = thm "LIM_equal"; +val LIM_trans = thm "LIM_trans"; +val LIM_NSLIM = thm "LIM_NSLIM"; +val NSLIM_LIM = thm "NSLIM_LIM"; +val LIM_NSLIM_iff = thm "LIM_NSLIM_iff"; +val NSLIM_mult = thm "NSLIM_mult"; +val LIM_mult2 = thm "LIM_mult2"; +val NSLIM_add = thm "NSLIM_add"; +val LIM_add2 = thm "LIM_add2"; +val NSLIM_const = thm "NSLIM_const"; +val LIM_const2 = thm "LIM_const2"; +val NSLIM_minus = thm "NSLIM_minus"; +val LIM_minus2 = thm "LIM_minus2"; +val NSLIM_add_minus = thm "NSLIM_add_minus"; +val LIM_add_minus2 = thm "LIM_add_minus2"; +val NSLIM_inverse = thm "NSLIM_inverse"; +val LIM_inverse = thm "LIM_inverse"; +val NSLIM_zero = thm "NSLIM_zero"; +val LIM_zero2 = thm "LIM_zero2"; +val NSLIM_zero_cancel = thm "NSLIM_zero_cancel"; +val LIM_zero_cancel = thm "LIM_zero_cancel"; +val NSLIM_not_zero = thm "NSLIM_not_zero"; +val NSLIM_const_not_eq = thm "NSLIM_const_not_eq"; +val NSLIM_const_eq = thm "NSLIM_const_eq"; +val NSLIM_unique = thm "NSLIM_unique"; +val LIM_unique2 = thm "LIM_unique2"; +val NSLIM_mult_zero = thm "NSLIM_mult_zero"; +val LIM_mult_zero2 = thm "LIM_mult_zero2"; +val NSLIM_self = thm "NSLIM_self"; +val isNSContD = thm "isNSContD"; +val isNSCont_NSLIM = thm "isNSCont_NSLIM"; +val NSLIM_isNSCont = thm "NSLIM_isNSCont"; +val isNSCont_NSLIM_iff = thm "isNSCont_NSLIM_iff"; +val isNSCont_LIM_iff = thm "isNSCont_LIM_iff"; +val isNSCont_isCont_iff = thm "isNSCont_isCont_iff"; +val isCont_isNSCont = thm "isCont_isNSCont"; +val isNSCont_isCont = thm "isNSCont_isCont"; +val NSLIM_h_iff = thm "NSLIM_h_iff"; +val NSLIM_isCont_iff = thm "NSLIM_isCont_iff"; +val LIM_isCont_iff = thm "LIM_isCont_iff"; +val isCont_iff = thm "isCont_iff"; +val isCont_add = thm "isCont_add"; +val isCont_mult = thm "isCont_mult"; +val isCont_o = thm "isCont_o"; +val isCont_o2 = thm "isCont_o2"; +val isNSCont_minus = thm "isNSCont_minus"; +val isCont_minus = thm "isCont_minus"; +val isCont_inverse = thm "isCont_inverse"; +val isNSCont_inverse = thm "isNSCont_inverse"; +val isCont_diff = thm "isCont_diff"; +val isCont_const = thm "isCont_const"; +val isNSCont_const = thm "isNSCont_const"; +val isNSCont_rabs = thm "isNSCont_rabs"; +val isCont_rabs = thm "isCont_rabs"; +val isNSUContD = thm "isNSUContD"; +val isUCont_isCont = thm "isUCont_isCont"; +val isUCont_isNSUCont = thm "isUCont_isNSUCont"; +val isNSUCont_isUCont = thm "isNSUCont_isUCont"; +val DERIV_iff = thm "DERIV_iff"; +val DERIV_NS_iff = thm "DERIV_NS_iff"; +val DERIV_D = thm "DERIV_D"; +val NS_DERIV_D = thm "NS_DERIV_D"; +val DERIV_unique = thm "DERIV_unique"; +val NSDeriv_unique = thm "NSDeriv_unique"; +val differentiableD = thm "differentiableD"; +val differentiableI = thm "differentiableI"; +val NSdifferentiableD = thm "NSdifferentiableD"; +val NSdifferentiableI = thm "NSdifferentiableI"; +val LIM_I = thm "LIM_I"; +val DERIV_LIM_iff = thm "DERIV_LIM_iff"; +val DERIV_iff2 = thm "DERIV_iff2"; +val NSDERIV_NSLIM_iff = thm "NSDERIV_NSLIM_iff"; +val NSDERIV_NSLIM_iff2 = thm "NSDERIV_NSLIM_iff2"; +val NSDERIV_iff2 = thm "NSDERIV_iff2"; +val hypreal_not_eq_minus_iff = thm "hypreal_not_eq_minus_iff"; +val NSDERIVD5 = thm "NSDERIVD5"; +val NSDERIVD4 = thm "NSDERIVD4"; +val NSDERIVD3 = thm "NSDERIVD3"; +val NSDERIV_DERIV_iff = thm "NSDERIV_DERIV_iff"; +val NSDERIV_isNSCont = thm "NSDERIV_isNSCont"; +val DERIV_isCont = thm "DERIV_isCont"; +val NSDERIV_const = thm "NSDERIV_const"; +val DERIV_const = thm "DERIV_const"; +val NSDERIV_add = thm "NSDERIV_add"; +val DERIV_add = thm "DERIV_add"; +val NSDERIV_mult = thm "NSDERIV_mult"; +val DERIV_mult = thm "DERIV_mult"; +val NSDERIV_cmult = thm "NSDERIV_cmult"; +val DERIV_cmult = thm "DERIV_cmult"; +val NSDERIV_minus = thm "NSDERIV_minus"; +val DERIV_minus = thm "DERIV_minus"; +val NSDERIV_add_minus = thm "NSDERIV_add_minus"; +val DERIV_add_minus = thm "DERIV_add_minus"; +val NSDERIV_diff = thm "NSDERIV_diff"; +val DERIV_diff = thm "DERIV_diff"; +val incrementI = thm "incrementI"; +val incrementI2 = thm "incrementI2"; +val increment_thm = thm "increment_thm"; +val increment_thm2 = thm "increment_thm2"; +val increment_approx_zero = thm "increment_approx_zero"; +val NSDERIV_zero = thm "NSDERIV_zero"; +val NSDERIV_approx = thm "NSDERIV_approx"; +val NSDERIVD1 = thm "NSDERIVD1"; +val NSDERIVD2 = thm "NSDERIVD2"; +val NSDERIV_chain = thm "NSDERIV_chain"; +val DERIV_chain = thm "DERIV_chain"; +val DERIV_chain2 = thm "DERIV_chain2"; +val NSDERIV_Id = thm "NSDERIV_Id"; +val DERIV_Id = thm "DERIV_Id"; +val isCont_Id = thms "isCont_Id"; +val DERIV_cmult_Id = thm "DERIV_cmult_Id"; +val NSDERIV_cmult_Id = thm "NSDERIV_cmult_Id"; +val DERIV_pow = thm "DERIV_pow"; +val NSDERIV_pow = thm "NSDERIV_pow"; +val NSDERIV_inverse = thm "NSDERIV_inverse"; +val DERIV_inverse = thm "DERIV_inverse"; +val DERIV_inverse_fun = thm "DERIV_inverse_fun"; +val NSDERIV_inverse_fun = thm "NSDERIV_inverse_fun"; +val DERIV_quotient = thm "DERIV_quotient"; +val NSDERIV_quotient = thm "NSDERIV_quotient"; +val CARAT_DERIV = thm "CARAT_DERIV"; +val CARAT_NSDERIV = thm "CARAT_NSDERIV"; +val hypreal_eq_minus_iff3 = thm "hypreal_eq_minus_iff3"; +val starfun_if_eq = thm "starfun_if_eq"; +val CARAT_DERIVD = thm "CARAT_DERIVD"; +val f_inc_g_dec_Beq_f = thm "f_inc_g_dec_Beq_f"; +val f_inc_g_dec_Beq_g = thm "f_inc_g_dec_Beq_g"; +val f_inc_imp_le_lim = thm "f_inc_imp_le_lim"; +val lim_uminus = thm "lim_uminus"; +val g_dec_imp_lim_le = thm "g_dec_imp_lim_le"; +val Bolzano_bisect_le = thm "Bolzano_bisect_le"; +val Bolzano_bisect_fst_le_Suc = thm "Bolzano_bisect_fst_le_Suc"; +val Bolzano_bisect_Suc_le_snd = thm "Bolzano_bisect_Suc_le_snd"; +val eq_divide_2_times_iff = thm "eq_divide_2_times_iff"; +val Bolzano_bisect_diff = thm "Bolzano_bisect_diff"; +val Bolzano_nest_unique = thms "Bolzano_nest_unique"; +val not_P_Bolzano_bisect = thm "not_P_Bolzano_bisect"; +val not_P_Bolzano_bisect = thm "not_P_Bolzano_bisect"; +val lemma_BOLZANO2 = thm "lemma_BOLZANO2"; +val IVT = thm "IVT"; +val IVT2 = thm "IVT2"; +val IVT_objl = thm "IVT_objl"; +val IVT2_objl = thm "IVT2_objl"; +val isCont_bounded = thm "isCont_bounded"; +val isCont_has_Ub = thm "isCont_has_Ub"; +val isCont_eq_Ub = thm "isCont_eq_Ub"; +val isCont_eq_Lb = thm "isCont_eq_Lb"; +val isCont_Lb_Ub = thm "isCont_Lb_Ub"; +val DERIV_left_inc = thm "DERIV_left_inc"; +val DERIV_left_dec = thm "DERIV_left_dec"; +val DERIV_local_max = thm "DERIV_local_max"; +val DERIV_local_min = thm "DERIV_local_min"; +val DERIV_local_const = thm "DERIV_local_const"; +val Rolle = thm "Rolle"; +val MVT = thm "MVT"; +val DERIV_isconst_end = thm "DERIV_isconst_end"; +val DERIV_isconst1 = thm "DERIV_isconst1"; +val DERIV_isconst2 = thm "DERIV_isconst2"; +val DERIV_isconst_all = thm "DERIV_isconst_all"; +val DERIV_const_ratio_const = thm "DERIV_const_ratio_const"; +val DERIV_const_ratio_const2 = thm "DERIV_const_ratio_const2"; +val real_average_minus_first = thm "real_average_minus_first"; +val real_average_minus_second = thm "real_average_minus_second"; +val DERIV_const_average = thm "DERIV_const_average"; +val isCont_inj_range = thm "isCont_inj_range"; +val isCont_inverse_function = thm "isCont_inverse_function"; +*} + end diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/Hyperreal/NSA.thy --- a/src/HOL/Hyperreal/NSA.thy Fri Mar 19 10:50:06 2004 +0100 +++ b/src/HOL/Hyperreal/NSA.thy Fri Mar 19 10:51:03 2004 +0100 @@ -154,7 +154,7 @@ lemma SReal_dense: "[| (x::hypreal) \ Reals; y \ Reals; x \r \ Reals. x 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 (drule half_gt_zero) apply (blast intro: hrabs_add_less hrabs_add_less SReal_divide_number_of) done @@ -488,6 +485,9 @@ by (simp add: approx_def Infinitesimal_def) declare approx_refl [iff] +lemma hypreal_minus_distrib1: "-(y + -(x::hypreal)) = x + -y" +by (simp add: hypreal_add_commute) + lemma approx_sym: "x @= y ==> y @= x" apply (simp add: approx_def) apply (rule hypreal_minus_distrib1 [THEN subst]) @@ -562,7 +562,6 @@ 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"; diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/Hyperreal/NthRoot.thy --- a/src/HOL/Hyperreal/NthRoot.thy Fri Mar 19 10:50:06 2004 +0100 +++ b/src/HOL/Hyperreal/NthRoot.thy Fri Mar 19 10:51:03 2004 +0100 @@ -1,8 +1,7 @@ (* Title : NthRoot.thy Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge - Description : Existence of nth root. Adapted from - http://www.math.unl.edu/~webnotes + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 *) header{*Existence of Nth Root*} @@ -17,10 +16,9 @@ lemma lemma_nth_realpow_non_empty: "[| (0::real) < a; 0 < n |] ==> \s. s : {x. x ^ n <= a & 0 < x}" apply (case_tac "1 <= a") -apply (rule_tac x = "1" in exI) +apply (rule_tac x = 1 in exI) apply (drule_tac [2] linorder_not_le [THEN iffD1]) -apply (drule_tac [2] less_not_refl2 [THEN not0_implies_Suc]) -apply (simp add: ); +apply (drule_tac [2] less_not_refl2 [THEN not0_implies_Suc], simp) apply (force intro!: realpow_Suc_le_self simp del: realpow_Suc) done @@ -32,23 +30,19 @@ "[| (0::real) < a; 0 < n |] ==> \u. isUb (UNIV::real set) {x. x ^ n <= a & 0 < x} u" apply (case_tac "1 <= a") -apply (rule_tac x = "a" in exI) +apply (rule_tac x = a in exI) apply (drule_tac [2] linorder_not_le [THEN iffD1]) -apply (rule_tac [2] x = "1" in exI) -apply (rule_tac [!] setleI [THEN isUbI]) -apply safe +apply (rule_tac [2] x = 1 in exI) +apply (rule_tac [!] setleI [THEN isUbI], safe) apply (simp_all (no_asm)) apply (rule_tac [!] ccontr) apply (drule_tac [!] linorder_not_le [THEN iffD1]) -apply (drule realpow_ge_self2 , assumption) -apply (drule_tac n = "n" in realpow_less) +apply (drule realpow_ge_self2, assumption) +apply (drule_tac n = n in realpow_less) apply (assumption+) -apply (drule real_le_trans , assumption) -apply (drule_tac y = "y ^ n" in order_less_le_trans) -apply (assumption) -apply (simp); -apply (drule_tac n = "n" in zero_less_one [THEN realpow_less]) -apply auto +apply (drule real_le_trans, assumption) +apply (drule_tac y = "y ^ n" in order_less_le_trans, assumption, simp) +apply (drule_tac n = n in zero_less_one [THEN realpow_less], auto) done lemma nth_realpow_isLub_ex: @@ -62,23 +56,23 @@ lemma lemma_nth_realpow_seq: "isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u ==> u + inverse(real (Suc k)) ~: {x. x ^ n <= a & 0 < x}" -apply (safe , drule isLubD2 , blast) +apply (safe, drule isLubD2, blast) apply (simp add: linorder_not_less [symmetric]) done lemma lemma_nth_realpow_isLub_gt_zero: "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u; 0 < a; 0 < n |] ==> 0 < u" -apply (drule lemma_nth_realpow_non_empty , auto) -apply (drule_tac y = "s" in isLub_isUb [THEN isUbD]) +apply (drule lemma_nth_realpow_non_empty, auto) +apply (drule_tac y = s in isLub_isUb [THEN isUbD]) apply (auto intro: order_less_le_trans) done lemma lemma_nth_realpow_isLub_ge: "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u; 0 < a; 0 < n |] ==> ALL k. a <= (u + inverse(real (Suc k))) ^ n" -apply (safe) -apply (frule lemma_nth_realpow_seq , safe) +apply safe +apply (frule lemma_nth_realpow_seq, safe) apply (auto elim: order_less_asym simp add: linorder_not_less [symmetric]) apply (simp add: linorder_not_less) apply (rule order_less_trans [of _ 0]) @@ -90,7 +84,7 @@ "[| (0::real) < a; 0 < n; isLub (UNIV::real set) {x. x ^ n <= a & 0 < x} u |] ==> a <= u ^ n" -apply (frule lemma_nth_realpow_isLub_ge , safe) +apply (frule lemma_nth_realpow_isLub_ge, safe) apply (rule LIMSEQ_inverse_real_of_nat_add [THEN LIMSEQ_pow, THEN LIMSEQ_le_const]) apply (auto simp add: real_of_nat_def) done @@ -100,16 +94,14 @@ lemma less_isLub_not_isUb: "[| isLub (UNIV::real set) S u; x < u |] ==> ~ isUb (UNIV::real set) S x" -apply (safe) -apply (drule isLub_le_isUb) -apply assumption -apply (drule order_less_le_trans) -apply (auto) +apply safe +apply (drule isLub_le_isUb, assumption) +apply (drule order_less_le_trans, auto) done lemma not_isUb_less_ex: "~ isUb (UNIV::real set) S u ==> \x \ S. u < x" -apply (rule ccontr , erule swap) +apply (rule ccontr, erule swap) apply (rule setleI [THEN isUbI]) apply (auto simp add: linorder_not_less [symmetric]) done @@ -129,14 +121,12 @@ lemma lemma_nth_realpow_isLub_le: "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u; 0 < a; 0 < n |] ==> ALL k. (u*(1 + -inverse(real (Suc k)))) ^ n <= a" -apply (safe) +apply safe apply (frule less_isLub_not_isUb [THEN not_isUb_less_ex]) -apply (rule_tac n = "k" in real_mult_less_self) -apply (blast intro: lemma_nth_realpow_isLub_gt_zero) -apply (safe) -apply (drule_tac n = "k" in - lemma_nth_realpow_isLub_gt_zero [THEN real_mult_add_one_minus_ge_zero]) -apply assumption+ +apply (rule_tac n = k in real_mult_less_self) +apply (blast intro: lemma_nth_realpow_isLub_gt_zero, safe) +apply (drule_tac n = k in + lemma_nth_realpow_isLub_gt_zero [THEN real_mult_add_one_minus_ge_zero], assumption+) apply (blast intro: order_trans order_less_imp_le power_mono) done @@ -145,7 +135,7 @@ "[| (0::real) < a; 0 < n; isLub (UNIV::real set) {x. x ^ n <= a & 0 < x} u |] ==> u ^ n <= a" -apply (frule lemma_nth_realpow_isLub_le , safe) +apply (frule lemma_nth_realpow_isLub_le, safe) apply (rule LIMSEQ_inverse_real_of_nat_add_minus_mult [THEN LIMSEQ_pow, THEN LIMSEQ_le_const2]) apply (auto simp add: real_of_nat_def) @@ -153,29 +143,26 @@ text{*The theorem at last!*} lemma realpow_nth: "[| (0::real) < a; 0 < n |] ==> \r. r ^ n = a" -apply (frule nth_realpow_isLub_ex , auto) -apply (auto intro: realpow_nth_le realpow_nth_ge real_le_anti_sym) +apply (frule nth_realpow_isLub_ex, auto) +apply (auto intro: realpow_nth_le realpow_nth_ge order_antisym) done (* positive only *) lemma realpow_pos_nth: "[| (0::real) < a; 0 < n |] ==> \r. 0 < r & r ^ n = a" -apply (frule nth_realpow_isLub_ex , auto) -apply (auto intro: realpow_nth_le realpow_nth_ge real_le_anti_sym lemma_nth_realpow_isLub_gt_zero) +apply (frule nth_realpow_isLub_ex, auto) +apply (auto intro: realpow_nth_le realpow_nth_ge order_antisym lemma_nth_realpow_isLub_gt_zero) done lemma realpow_pos_nth2: "(0::real) < a ==> \r. 0 < r & r ^ Suc n = a" -apply (blast intro: realpow_pos_nth) -done +by (blast intro: realpow_pos_nth) (* uniqueness of nth positive root *) lemma realpow_pos_nth_unique: "[| (0::real) < a; 0 < n |] ==> EX! r. 0 < r & r ^ n = a" apply (auto intro!: realpow_pos_nth) -apply (cut_tac x = "r" and y = "y" in linorder_less_linear) -apply auto -apply (drule_tac x = "r" in realpow_less) -apply (drule_tac [4] x = "y" in realpow_less) -apply (auto) +apply (cut_tac x = r and y = y in linorder_less_linear, auto) +apply (drule_tac x = r in realpow_less) +apply (drule_tac [4] x = y in realpow_less, auto) done ML diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/Hyperreal/SEQ.ML --- a/src/HOL/Hyperreal/SEQ.ML Fri Mar 19 10:50:06 2004 +0100 +++ b/src/HOL/Hyperreal/SEQ.ML Fri Mar 19 10:51:03 2004 +0100 @@ -618,7 +618,7 @@ \ |] ==> ALL n. ma <= n --> X n = X ma"; by (Step_tac 1); by (dres_inst_tac [("y","X n")] isLubD2 1); -by (ALLGOALS(blast_tac (claset() addDs [real_le_anti_sym]))); +by (ALLGOALS(blast_tac (claset() addDs [order_antisym]))); qed "lemma_converg1"; (*------------------------------------------------------------------- diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/Hyperreal/Star.thy --- a/src/HOL/Hyperreal/Star.thy Fri Mar 19 10:50:06 2004 +0100 +++ b/src/HOL/Hyperreal/Star.thy Fri Mar 19 10:51:03 2004 +0100 @@ -180,7 +180,6 @@ *) lemma hrabs_is_starext_rabs: "is_starext abs abs" - apply (simp add: is_starext_def, safe) apply (rule_tac z = x in eq_Abs_hypreal) apply (rule_tac z = y in eq_Abs_hypreal, auto) @@ -190,9 +189,10 @@ apply (arith | ultra)+ done -lemma Rep_hypreal_FreeUltrafilterNat: "[| X \ Rep_hypreal z; Y \ Rep_hypreal z |] +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 (cases z) apply (auto, ultra) done @@ -212,22 +212,25 @@ UN_equiv_class [OF equiv_hyprel starfun_congruent]) done +lemma starfun_if_eq: + "w \ hypreal_of_real x + ==> ( *f* (\z. if z = x then a else g z)) w = ( *f* g) w" +apply (cases w) +apply (simp add: hypreal_of_real_def starfun, ultra) +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 +by (cases xa, simp add: starfun hypreal_mult) 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 +by (cases xa, simp add: starfun hypreal_add) declare starfun_add [symmetric, simp] (*-------------------------------------------- @@ -235,7 +238,7 @@ -------------------------------------------*) lemma starfun_minus: "- ( *f* f) x = ( *f* (%x. - f x)) x" -apply (rule_tac z = x in eq_Abs_hypreal) +apply (cases x) apply (auto simp add: starfun hypreal_minus) done declare starfun_minus [symmetric, simp] @@ -271,7 +274,7 @@ 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 (cases xa) apply (auto simp add: starfun hypreal_of_real_def) done @@ -282,12 +285,12 @@ ----------------------------------------------------*) 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 (cases x) apply (auto simp add: starfun) done lemma starfun_Id: "( *f* (%x. x)) x = x" -apply (rule_tac z = x in eq_Abs_hypreal) +apply (cases x) apply (auto simp add: starfun) done declare starfun_Id [simp] @@ -297,7 +300,6 @@ ----------------------------------------------------------------------*) lemma is_starext_starfun: "is_starext ( *f* f) f" - apply (simp add: is_starext_def, auto) apply (rule_tac z = x in eq_Abs_hypreal) apply (rule_tac z = y in eq_Abs_hypreal) @@ -309,7 +311,6 @@ ----------------------------------------------------------------------*) lemma is_starfun_starext: "is_starext F f ==> F = *f* f" - apply (simp add: is_starext_def) apply (rule ext) apply (rule_tac z = x in eq_Abs_hypreal) @@ -336,12 +337,12 @@ (* 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 (cases xa) 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 (cases xa) apply (auto simp add: starfun hypreal_of_real_def hypreal_add) done @@ -370,13 +371,13 @@ by (rule hrabs_is_starext_rabs [THEN is_starext_starfun_iff [THEN iffD1], symmetric]) lemma starfun_inverse_inverse: "( *f* inverse) x = inverse(x)" -apply (rule_tac z = x in eq_Abs_hypreal) +apply (cases x) 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 (cases x) apply (auto simp add: starfun hypreal_inverse) done declare starfun_inverse [symmetric, simp] @@ -386,7 +387,7 @@ 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 (cases x) apply (auto intro: FreeUltrafilterNat_subset dest!: FreeUltrafilterNat_Compl_mem simp add: starfun hypreal_inverse hypreal_zero_def) done @@ -397,7 +398,7 @@ lemma starfun_mem_starset: "( *f* f) x : *s* A ==> x : *s* {x. f x \ A}" apply (simp add: starset_def) -apply (rule_tac z = x in eq_Abs_hypreal) +apply (cases x) apply (auto simp add: starfun) apply (rename_tac "X") apply (drule_tac x = "%n. f (X n) " in bspec) @@ -446,7 +447,7 @@ (\X \ Rep_hypreal(x). \m. {n. abs(X n) < inverse(real(Suc m))} \ FreeUltrafilterNat)" -apply (rule eq_Abs_hypreal [of x]) +apply (cases x) 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_eq) diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/Hyperreal/Transcendental.ML --- a/src/HOL/Hyperreal/Transcendental.ML Fri Mar 19 10:50:06 2004 +0100 +++ b/src/HOL/Hyperreal/Transcendental.ML Fri Mar 19 10:51:03 2004 +0100 @@ -5,6 +5,9 @@ Description : Power Series *) +fun ARITH_PROVE str = prove_goal thy str + (fn prems => [cut_facts_tac prems 1,arith_tac 1]); + fun multr_by_tac x i = let val cancel_thm = CLAIM "[| (0::real) x (root(Suc n) x = root(Suc n) y) = (x = y)"; -by (auto_tac (claset() addSIs [real_le_anti_sym],simpset())); +by (auto_tac (claset() addSIs [order_antisym],simpset())); by (res_inst_tac [("n1","n")] (real_root_le_iff RS iffD1) 1); by (res_inst_tac [("n1","n")] (real_root_le_iff RS iffD1) 4); by Auto_tac; diff -r 758e7acdea2f -r cc61fd03e589 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Mar 19 10:50:06 2004 +0100 +++ b/src/HOL/IsaMakefile Fri Mar 19 10:51:03 2004 +0100 @@ -148,7 +148,7 @@ Hyperreal/HTranscendental.thy Hyperreal/HyperArith.thy\ Hyperreal/HyperDef.thy Hyperreal/HyperNat.thy\ Hyperreal/HyperPow.thy Hyperreal/Hyperreal.thy Hyperreal/IntFloor.thy\ - Hyperreal/Lim.ML Hyperreal/Lim.thy Hyperreal/Log.thy\ + Hyperreal/Lim.thy Hyperreal/Log.thy\ Hyperreal/MacLaurin.ML Hyperreal/MacLaurin.thy Hyperreal/NatStar.thy\ Hyperreal/NSA.thy Hyperreal/NthRoot.thy Hyperreal/Poly.thy\ Hyperreal/SEQ.ML Hyperreal/SEQ.thy Hyperreal/Series.thy\