# HG changeset patch # User paulson # Date 1090856092 -7200 # Node ID 89840837108e20b5552986677d1d29dd8a5caa75 # Parent 4b3d280ef06a85c1ff8881cd11fa3146363d62d8 converting Hyperreal/Transcendental to Isar script diff -r 4b3d280ef06a -r 89840837108e src/HOL/Hyperreal/HTranscendental.thy --- a/src/HOL/Hyperreal/HTranscendental.thy Mon Jul 26 15:48:50 2004 +0200 +++ b/src/HOL/Hyperreal/HTranscendental.thy Mon Jul 26 17:34:52 2004 +0200 @@ -60,7 +60,7 @@ lemma hypreal_sqrt_gt_zero_pow2: "0 < x ==> ( *f* sqrt) (x) ^ 2 = x" apply (cases x) -apply (auto intro: FreeUltrafilterNat_subset real_sqrt_gt_zero_pow2 +apply (auto intro: FreeUltrafilterNat_subset simp add: hypreal_less starfun hrealpow hypreal_zero_num simp del: hpowr_Suc realpow_Suc) done diff -r 4b3d280ef06a -r 89840837108e src/HOL/Hyperreal/Lim.thy --- a/src/HOL/Hyperreal/Lim.thy Mon Jul 26 15:48:50 2004 +0200 +++ b/src/HOL/Hyperreal/Lim.thy Mon Jul 26 17:34:52 2004 +0200 @@ -35,11 +35,11 @@ (* differentiation: D is derivative of function f at x *) deriv:: "[real=>real,real,real] => bool" - ("(DERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60) + ("(DERIV (_)/ (_)/ :> (_))" [60, 60, 60] 60) "DERIV f x :> D == ((%h. (f(x + h) + -f x)/h) -- 0 --> D)" nsderiv :: "[real=>real,real,real] => bool" - ("(NSDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60) + ("(NSDERIV (_)/ (_)/ :> (_))" [60, 60, 60] 60) "NSDERIV f x :> D == (\h \ Infinitesimal - {0}. (( *f* f)(hypreal_of_real x + h) + - hypreal_of_real (f x))/h @= hypreal_of_real D)" diff -r 4b3d280ef06a -r 89840837108e src/HOL/Hyperreal/MacLaurin_lemmas.ML --- a/src/HOL/Hyperreal/MacLaurin_lemmas.ML Mon Jul 26 15:48:50 2004 +0200 +++ b/src/HOL/Hyperreal/MacLaurin_lemmas.ML Mon Jul 26 17:34:52 2004 +0200 @@ -4,6 +4,35 @@ Description : MacLaurin series *) +val DERIV_intros = thms"DERIV_intros"; + +val lemma_DERIV_subst = thm"lemma_DERIV_subst"; + +fun ARITH_PROVE str = prove_goal thy str + (fn prems => [cut_facts_tac prems 1,arith_tac 1]); + + +(* FIXME: remove this quick, crude tactic *) +exception DERIV_name; +fun get_fun_name (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _)) = f +| get_fun_name (_ $ (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _))) = f +| get_fun_name _ = raise DERIV_name; + +val deriv_rulesI = [DERIV_Id,DERIV_const,DERIV_cos,DERIV_cmult, + DERIV_sin, DERIV_exp, DERIV_inverse,DERIV_pow, + DERIV_add, DERIV_diff, DERIV_mult, DERIV_minus, + DERIV_inverse_fun,DERIV_quotient,DERIV_fun_pow, + DERIV_fun_exp,DERIV_fun_sin,DERIV_fun_cos, + DERIV_Id,DERIV_const,DERIV_cos]; + + +fun deriv_tac i = (resolve_tac deriv_rulesI i) ORELSE + ((rtac (read_instantiate [("f",get_fun_name (getgoal i))] + DERIV_chain2) i) handle DERIV_name => no_tac); + +val DERIV_tac = ALLGOALS(fn i => REPEAT(deriv_tac i)); + + Goal "sumr 0 n (%m. f (m + k)) = sumr 0 (n + k) f - sumr 0 k f"; by (induct_tac "n" 1); by Auto_tac; @@ -35,6 +64,29 @@ (* Maclaurin's theorem with Lagrange form of remainder *) (*---------------------------------------------------------------------------*) + + +(* FIXME: remove this quick, crude tactic *) +exception DERIV_name; +fun get_fun_name (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _)) = f +| get_fun_name (_ $ (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _))) = f +| get_fun_name _ = raise DERIV_name; + +val deriv_rulesI = [DERIV_Id,DERIV_const,DERIV_cos,DERIV_cmult, + DERIV_sin, DERIV_exp, DERIV_inverse,DERIV_pow, + DERIV_add, DERIV_diff, DERIV_mult, DERIV_minus, + DERIV_inverse_fun,DERIV_quotient,DERIV_fun_pow, + DERIV_fun_exp,DERIV_fun_sin,DERIV_fun_cos, + DERIV_Id,DERIV_const,DERIV_cos]; + + +fun deriv_tac i = (resolve_tac deriv_rulesI i) ORELSE + ((rtac (read_instantiate [("f",get_fun_name (getgoal i))] + DERIV_chain2) i) handle DERIV_name => no_tac); + +val DERIV_tac = ALLGOALS(fn i => REPEAT(deriv_tac i)); + + (* Annoying: Proof is now even longer due mostly to change in behaviour of simplifier since Isabelle99 *) Goal " [| 0 < h; 0 < n; diff 0 = f; \ @@ -123,7 +175,7 @@ delsimps [fact_Suc,realpow_Suc]) 2); by (rtac DERIV_cmult 2); by (rtac lemma_DERIV_subst 2); -by DERIV_tac; +by (best_tac (claset() addIs [DERIV_chain2] addSIs DERIV_intros) 2); by (stac fact_Suc 2); by (stac real_of_nat_mult 2); by (simp_tac (simpset() addsimps [inverse_mult_distrib] @ diff -r 4b3d280ef06a -r 89840837108e src/HOL/Hyperreal/Transcendental.ML --- a/src/HOL/Hyperreal/Transcendental.ML Mon Jul 26 15:48:50 2004 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,3166 +0,0 @@ -(* Title : Transcendental.ML - Author : Jacques D. Fleuriot - Copyright : 1998,1999 University of Cambridge - 1999 University of Edinburgh - 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) ^ (Suc n) = x"; -by (dres_inst_tac [("n","n")] realpow_pos_nth2 1); -by (auto_tac (claset() addIs [someI2],simpset())); -qed "real_root_pow_pos"; - -Goal "0 <= x ==> (root(Suc n) x) ^ (Suc n) = x"; -by (auto_tac (claset() addSDs [real_le_imp_less_or_eq] - addDs [real_root_pow_pos],simpset())); -qed "real_root_pow_pos2"; - -Goalw [root_def] - "0 < x ==> root(Suc n) (x ^ (Suc n)) = x"; -by (rtac some_equality 1); -by (forw_inst_tac [("n","n")] zero_less_power 2); -by (auto_tac (claset(),simpset() addsimps [zero_less_mult_iff])); -by (res_inst_tac [("x","u"),("y","x")] linorder_cases 1); -by (dres_inst_tac [("n1","n"),("x","u")] (zero_less_Suc RSN (3, realpow_less)) 1); -by (dres_inst_tac [("n1","n"),("x","x")] (zero_less_Suc RSN (3, realpow_less)) 4); -by (auto_tac (claset(),simpset() addsimps [order_less_irrefl])); -qed "real_root_pos"; - -Goal "0 <= x ==> root(Suc n) (x ^ (Suc n)) = x"; -by (auto_tac (claset() addSDs [real_le_imp_less_or_eq, - real_root_pos],simpset())); -qed "real_root_pos2"; - -Goalw [root_def] - "0 < x ==> 0 <= root(Suc n) x"; -by (dres_inst_tac [("n","n")] realpow_pos_nth2 1); -by (Safe_tac THEN rtac someI2 1); -by (auto_tac (claset() addSIs [order_less_imp_le] - addDs [zero_less_power],simpset() addsimps [zero_less_mult_iff])); -qed "real_root_pos_pos"; - -Goal "0 <= x ==> 0 <= root(Suc n) x"; -by (auto_tac (claset() addSDs [real_le_imp_less_or_eq] - addDs [real_root_pos_pos],simpset())); -qed "real_root_pos_pos_le"; - -Goalw [root_def] "root (Suc n) 1 = 1"; -by (rtac some_equality 1); -by Auto_tac; -by (rtac ccontr 1); -by (res_inst_tac [("x","u"),("y","1")] linorder_cases 1); -by (dres_inst_tac [("n","n")] realpow_Suc_less_one 1); -by (dres_inst_tac [("n","n")] power_gt1_lemma 4); -by (auto_tac (claset(),simpset() addsimps [order_less_irrefl])); -qed "real_root_one"; -Addsimps [real_root_one]; - -(*----------------------------------------------------------------------*) -(* Square root *) -(*----------------------------------------------------------------------*) - -(*lcp: needed now because 2 is a binary numeral!*) -Goal "root 2 = root (Suc (Suc 0))"; -by (simp_tac (simpset() delsimps [nat_numeral_0_eq_0, nat_numeral_1_eq_1] - addsimps [nat_numeral_0_eq_0 RS sym]) 1); -qed "root_2_eq"; -Addsimps [root_2_eq]; - -Goalw [sqrt_def] "sqrt 0 = 0"; -by (Auto_tac); -qed "real_sqrt_zero"; -Addsimps [real_sqrt_zero]; - -Goalw [sqrt_def] "sqrt 1 = 1"; -by (Auto_tac); -qed "real_sqrt_one"; -Addsimps [real_sqrt_one]; - -Goalw [sqrt_def] "(sqrt(x) ^ 2 = x) = (0 <= x)"; -by (Step_tac 1); -by (cut_inst_tac [("r","root 2 x")] realpow_two_le 1); -by (stac numeral_2_eq_2 2); -by (rtac real_root_pow_pos2 2); -by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); -qed "real_sqrt_pow2_iff"; -Addsimps [real_sqrt_pow2_iff]; - - -Addsimps [realpow_two_le_add_order RS (real_sqrt_pow2_iff RS iffD2)]; -Addsimps [simplify (simpset()) (realpow_two_le_add_order RS - (real_sqrt_pow2_iff RS iffD2))]; - -Goalw [sqrt_def] "0 < x ==> sqrt(x) ^ 2 = x"; -by (stac numeral_2_eq_2 1); -by (etac real_root_pow_pos 1); -qed "real_sqrt_gt_zero_pow2"; - -Goal "(sqrt(abs(x)) ^ 2 = abs x)"; -by (rtac (real_sqrt_pow2_iff RS iffD2) 1); -by (arith_tac 1); -qed "real_sqrt_abs_abs"; -Addsimps [real_sqrt_abs_abs]; - -Goalw [sqrt_def] - "0 <= x ==> sqrt(x) ^ 2 = sqrt(x ^ 2)"; -by (stac numeral_2_eq_2 1); -by (auto_tac (claset() addIs [real_root_pow_pos2 - RS ssubst, real_root_pos2 RS ssubst], - simpset() delsimps [realpow_Suc])); -qed "real_pow_sqrt_eq_sqrt_pow"; - -Goal "0 <= x ==> sqrt(x) ^ 2 = sqrt(abs(x) ^ 2)"; -by (asm_full_simp_tac (simpset() addsimps [real_pow_sqrt_eq_sqrt_pow]) 1); -qed "real_pow_sqrt_eq_sqrt_abs_pow2"; - -Goal "0 <= x ==> sqrt(x) ^ 2 = abs(x)"; -by (rtac (real_sqrt_abs_abs RS subst) 1); -by (res_inst_tac [("x1","x")] - (real_pow_sqrt_eq_sqrt_abs_pow2 RS ssubst) 1); -by (rtac (real_pow_sqrt_eq_sqrt_pow RS sym) 2); -by (assume_tac 1 THEN arith_tac 1); -qed "real_sqrt_pow_abs"; - -Goal "(~ (0::real) < x*x) = (x = 0)"; -by Auto_tac; -by (rtac ccontr 1); -by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1); -by Auto_tac; -by (ftac (real_mult_order) 2); -by (asm_full_simp_tac (simpset() addsimps [zero_less_mult_iff]) 1); -by Auto_tac; -qed "not_real_square_gt_zero"; -Addsimps [not_real_square_gt_zero]; - - -(* proof used to be simpler *) -Goalw [sqrt_def,root_def] - "[| 0 < x; 0 < y |] ==>sqrt(x*y) = sqrt(x) * sqrt(y)"; -by (dres_inst_tac [("n","1")] realpow_pos_nth2 1); -by (dres_inst_tac [("n","1")] realpow_pos_nth2 1); -by (asm_full_simp_tac (simpset() delsimps [realpow_Suc] - addsimps [numeral_2_eq_2]) 1); -by (Step_tac 1); -by (rtac someI2 1 THEN Step_tac 1 THEN Blast_tac 2); -by (Asm_full_simp_tac 1 THEN Asm_full_simp_tac 1); -by (rtac someI2 1 THEN Step_tac 1 THEN Blast_tac 2); -by (Asm_full_simp_tac 1 THEN Asm_full_simp_tac 1); -by (res_inst_tac [("a","xa * x")] someI2 1); -by (auto_tac (claset() addEs [order_less_asym], - simpset() addsimps mult_ac@[power_mult_distrib RS sym,realpow_two_disj, - zero_less_power, real_mult_order] delsimps [realpow_Suc])); -qed "real_sqrt_mult_distrib"; - -Goal "[|0<=x; 0<=y |] ==> sqrt(x*y) = sqrt(x) * sqrt(y)"; -by (auto_tac (claset() addIs [ real_sqrt_mult_distrib], - simpset() addsimps [order_le_less])); -qed "real_sqrt_mult_distrib2"; - -Goal "(r * r = 0) = (r = (0::real))"; -by Auto_tac; -qed "real_mult_self_eq_zero_iff"; -Addsimps [real_mult_self_eq_zero_iff]; - -Goalw [sqrt_def,root_def] "0 < x ==> 0 < sqrt(x)"; -by (stac numeral_2_eq_2 1); -by (dtac realpow_pos_nth2 1 THEN Step_tac 1); -by (rtac someI2 1 THEN Step_tac 1 THEN Blast_tac 2); -by Auto_tac; -qed "real_sqrt_gt_zero"; - -Goal "0 <= x ==> 0 <= sqrt(x)"; -by (auto_tac (claset() addIs [real_sqrt_gt_zero], - simpset() addsimps [order_le_less])); -qed "real_sqrt_ge_zero"; - -Goal "0 <= sqrt (x ^ 2 + y ^ 2)"; -by (auto_tac (claset() addSIs [real_sqrt_ge_zero],simpset())); -qed "real_sqrt_sum_squares_ge_zero"; -Addsimps [real_sqrt_sum_squares_ge_zero]; - -Goal "0 <= sqrt ((x ^ 2 + y ^ 2)*(xa ^ 2 + ya ^ 2))"; -by (auto_tac (claset() addSIs [real_sqrt_ge_zero],simpset() - addsimps [zero_le_mult_iff])); -qed "real_sqrt_sum_squares_mult_ge_zero"; -Addsimps [real_sqrt_sum_squares_mult_ge_zero]; - -Goal "sqrt ((x ^ 2 + y ^ 2) * (xa ^ 2 + ya ^ 2)) ^ 2 = \ -\ (x ^ 2 + y ^ 2) * (xa ^ 2 + ya ^ 2)"; -by (auto_tac (claset(),simpset() addsimps [real_sqrt_pow2_iff, - zero_le_mult_iff] delsimps [realpow_Suc])); -qed "real_sqrt_sum_squares_mult_squared_eq"; -Addsimps [real_sqrt_sum_squares_mult_squared_eq]; - -Goal "sqrt(x ^ 2) = abs(x)"; -by (rtac (abs_realpow_two RS subst) 1); -by (rtac (real_sqrt_abs_abs RS subst) 1); -by (stac real_pow_sqrt_eq_sqrt_pow 1); -by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2, abs_mult])); -qed "real_sqrt_abs"; -Addsimps [real_sqrt_abs]; - -Goal "sqrt(x*x) = abs(x)"; -by (rtac (realpow_two RS subst) 1); -by (stac (numeral_2_eq_2 RS sym) 1); -by (rtac real_sqrt_abs 1); -qed "real_sqrt_abs2"; -Addsimps [real_sqrt_abs2]; - -Goal "0 < x ==> 0 < sqrt(x) ^ 2"; -by (asm_full_simp_tac (simpset() addsimps [real_sqrt_gt_zero_pow2]) 1); -qed "real_sqrt_pow2_gt_zero"; - -Goal "0 < x ==> sqrt x ~= 0"; -by (ftac real_sqrt_pow2_gt_zero 1); -by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2, order_less_irrefl])); -qed "real_sqrt_not_eq_zero"; - -Goal "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x"; -by (cut_inst_tac [("n1","2"),("a1","sqrt x")] (power_inverse RS sym) 1); -by (auto_tac (claset() addDs [real_sqrt_gt_zero_pow2],simpset())); -qed "real_inv_sqrt_pow2"; - -Goal "[| 0 <= x; sqrt(x) = 0|] ==> x = 0"; -by (dtac real_le_imp_less_or_eq 1); -by (auto_tac (claset() addDs [real_sqrt_not_eq_zero],simpset())); -qed "real_sqrt_eq_zero_cancel"; - -Goal "0 <= x ==> ((sqrt x = 0) = (x = 0))"; -by (auto_tac (claset(),simpset() addsimps [real_sqrt_eq_zero_cancel])); -qed "real_sqrt_eq_zero_cancel_iff"; -Addsimps [real_sqrt_eq_zero_cancel_iff]; - -Goal "x <= sqrt(x ^ 2 + y ^ 2)"; -by (subgoal_tac "x <= 0 | 0 <= x" 1); -by (Step_tac 1); -by (rtac real_le_trans 1); -by (auto_tac (claset(),simpset() delsimps [realpow_Suc])); -by (res_inst_tac [("n","1")] realpow_increasing 1); -by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2 RS sym] - delsimps [realpow_Suc])); -qed "real_sqrt_sum_squares_ge1"; -Addsimps [real_sqrt_sum_squares_ge1]; - -Goal "y <= sqrt(z ^ 2 + y ^ 2)"; -by (simp_tac (simpset() addsimps [real_add_commute] - delsimps [realpow_Suc]) 1); -qed "real_sqrt_sum_squares_ge2"; -Addsimps [real_sqrt_sum_squares_ge2]; - -Goal "1 <= x ==> 1 <= sqrt x"; -by (res_inst_tac [("n","1")] realpow_increasing 1); -by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2 RS sym, real_sqrt_gt_zero_pow2, - real_sqrt_ge_zero] delsimps [realpow_Suc])); -qed "real_sqrt_ge_one"; - -(*-------------------------------------------------------------------------*) -(* Exponential function *) -(*-------------------------------------------------------------------------*) - -Goal "summable (%n. inverse (real (fact n)) * x ^ n)"; -by (cut_inst_tac [("'a","real")] (zero_less_one RS dense) 1); -by (Step_tac 1); -by (cut_inst_tac [("x","r")] reals_Archimedean3 1); -by Auto_tac; -by (dres_inst_tac [("x","abs x")] spec 1 THEN Safe_tac); -by (res_inst_tac [("N","n"),("c","r")] ratio_test 1); -by (auto_tac (claset(), - simpset() addsimps [abs_mult,mult_assoc RS sym] delsimps [fact_Suc])); -by (rtac mult_right_mono 1); -by (res_inst_tac [("b1","abs x")] (mult_commute RS ssubst) 1); -by (stac fact_Suc 1); -by (stac real_of_nat_mult 1); -by (auto_tac (claset(),simpset() addsimps [abs_mult,inverse_mult_distrib])); -by (auto_tac (claset(), simpset() addsimps - [mult_assoc RS sym, abs_eqI2, positive_imp_inverse_positive])); -by (rtac order_less_imp_le 1); -by (res_inst_tac [("z1","real (Suc na)")] (real_mult_less_iff1 - RS iffD1) 1); -by (auto_tac (claset(),simpset() addsimps [real_not_refl2 RS not_sym, - mult_assoc,abs_inverse])); -by (etac order_less_trans 1); -by (auto_tac (claset(),simpset() addsimps [mult_less_cancel_left]@mult_ac)); -qed "summable_exp"; - -Addsimps [real_of_nat_fact_gt_zero, - real_of_nat_fact_ge_zero,inv_real_of_nat_fact_gt_zero, - inv_real_of_nat_fact_ge_zero]; - -Goalw [real_divide_def] - "summable (%n. \ -\ (if even n then 0 \ -\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * \ -\ x ^ n)"; -by (res_inst_tac [("g","(%n. inverse (real (fact n)) * abs(x) ^ n)")] - summable_comparison_test 1); -by (rtac summable_exp 2); -by (res_inst_tac [("x","0")] exI 1); -by (auto_tac (claset(), simpset() addsimps [power_abs RS sym, - abs_mult,zero_le_mult_iff])); -by (auto_tac (claset() addIs [mult_right_mono], - simpset() addsimps [positive_imp_inverse_positive,abs_eqI2])); -qed "summable_sin"; - -Goalw [real_divide_def] - "summable (%n. \ -\ (if even n then \ -\ (- 1) ^ (n div 2)/(real (fact n)) else 0) * x ^ n)"; -by (res_inst_tac [("g","(%n. inverse (real (fact n)) * abs(x) ^ n)")] - summable_comparison_test 1); -by (rtac summable_exp 2); -by (res_inst_tac [("x","0")] exI 1); -by (auto_tac (claset(), simpset() addsimps [power_abs RS sym,abs_mult, - zero_le_mult_iff])); -by (auto_tac (claset() addSIs [mult_right_mono], - simpset() addsimps [positive_imp_inverse_positive,abs_eqI2])); -qed "summable_cos"; - -Goal "(if even n then 0 \ -\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0"; -by (induct_tac "n" 1); -by (Auto_tac); -qed "lemma_STAR_sin"; -Addsimps [lemma_STAR_sin]; - -Goal "0 < n --> \ -\ (- 1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0"; -by (induct_tac "n" 1); -by (Auto_tac); -qed "lemma_STAR_cos"; -Addsimps [lemma_STAR_cos]; - -Goal "0 < n --> \ -\ (-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0"; -by (induct_tac "n" 1); -by (Auto_tac); -qed "lemma_STAR_cos1"; -Addsimps [lemma_STAR_cos1]; - -Goal "sumr 1 n (%n. if even n \ -\ then (- 1) ^ (n div 2)/(real (fact n)) * \ -\ 0 ^ n \ -\ else 0) = 0"; -by (induct_tac "n" 1); -by (case_tac "n" 2); -by (Auto_tac); -qed "lemma_STAR_cos2"; -Addsimps [lemma_STAR_cos2]; - -Goalw [exp_def] "(%n. inverse (real (fact n)) * x ^ n) sums exp(x)"; -by (rtac (summable_exp RS summable_sums) 1); -qed "exp_converges"; - -Goalw [sin_def] - "(%n. (if even n then 0 \ -\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * \ -\ x ^ n) sums sin(x)"; -by (rtac (summable_sin RS summable_sums) 1); -qed "sin_converges"; - -Goalw [cos_def] - "(%n. (if even n then \ -\ (- 1) ^ (n div 2)/(real (fact n)) \ -\ else 0) * x ^ n) sums cos(x)"; -by (rtac (summable_cos RS summable_sums) 1); -qed "cos_converges"; - -Goal "p <= n --> y ^ (Suc n - p) = ((y::real) ^ (n - p)) * y"; -by (induct_tac "n" 1 THEN Auto_tac); -by (subgoal_tac "p = Suc n" 1); -by (Asm_simp_tac 1 THEN Auto_tac); -by (dtac sym 1 THEN asm_full_simp_tac (simpset() addsimps - [Suc_diff_le,real_mult_commute,realpow_Suc RS sym] - delsimps [realpow_Suc]) 1); -qed_spec_mp "lemma_realpow_diff"; - -(*--------------------------------------------------------------------------*) -(* Properties of power series *) -(*--------------------------------------------------------------------------*) - -Goal "sumr 0 (Suc n) (%p. (x ^ p) * y ^ ((Suc n) - p)) = \ -\ y * sumr 0 (Suc n) (%p. (x ^ p) * (y ^ (n - p)))"; -by (auto_tac (claset(),simpset() addsimps [sumr_mult] delsimps [sumr_Suc])); -by (rtac sumr_subst 1); -by (strip_tac 1); -by (stac lemma_realpow_diff 1); -by (auto_tac (claset(),simpset() addsimps mult_ac)); -qed "lemma_realpow_diff_sumr"; - -Goal "x ^ (Suc n) - y ^ (Suc n) = \ -\ (x - y) * sumr 0 (Suc n) (%p. (x ^ p) * (y ^(n - p)))"; -by (induct_tac "n" 1); -by (Asm_full_simp_tac 1); -by (auto_tac (claset(),simpset() delsimps [sumr_Suc])); -by (stac sumr_Suc 1); -by (dtac sym 1); -by (auto_tac (claset(),simpset() addsimps [lemma_realpow_diff_sumr, - right_distrib,real_diff_def] @ - mult_ac delsimps [sumr_Suc])); -qed "lemma_realpow_diff_sumr2"; - -Goal "sumr 0 (Suc n) (%p. (x ^ p) * (y ^ (n - p))) = \ -\ sumr 0 (Suc n) (%p. (x ^ (n - p)) * (y ^ p))"; -by (case_tac "x = y" 1); -by (auto_tac (claset(),simpset() addsimps [real_mult_commute, - power_add RS sym] delsimps [sumr_Suc])); -by (res_inst_tac [("c1","x - y")] (real_mult_left_cancel RS iffD1) 1); -by (rtac (minus_minus RS subst) 2); -by (stac minus_mult_left 2); -by (auto_tac (claset(),simpset() addsimps [lemma_realpow_diff_sumr2 - RS sym] delsimps [sumr_Suc])); -qed "lemma_realpow_rev_sumr"; - -(* ------------------------------------------------------------------------ *) -(* Power series has a `circle` of convergence, *) -(* i.e. if it sums for x, then it sums absolutely for z with |z| < |x|. *) -(* ------------------------------------------------------------------------ *) - -Goal "[| summable (%n. f(n) * (x ^ n)); abs(z) < abs(x) |] \ -\ ==> summable (%n. abs(f(n)) * (z ^ n))"; -by (dtac summable_LIMSEQ_zero 1); -by (dtac convergentI 1); -by (asm_full_simp_tac (simpset() addsimps [Cauchy_convergent_iff RS sym]) 1); -by (dtac Cauchy_Bseq 1); -by (asm_full_simp_tac (simpset() addsimps [Bseq_def]) 1); -by (Step_tac 1); -by (res_inst_tac [("g","%n. K * abs(z ^ n) * inverse (abs(x ^ n))")] - summable_comparison_test 1); -by (res_inst_tac [("x","0")] exI 1 THEN Step_tac 1); -by (subgoal_tac "0 < abs (x ^ n)" 1); -by (res_inst_tac [("z","abs (x ^ n)")] (CLAIM_SIMP - "[| (0::real) x<=y" [mult_le_cancel_left]) 1); -by (auto_tac (claset(), - simpset() addsimps [mult_assoc,power_abs])); -by (dres_inst_tac [("x","0")] spec 2 THEN Force_tac 2); -by (auto_tac (claset(),simpset() addsimps [abs_mult,power_abs] @ mult_ac)); -by (res_inst_tac [("a2","z ^ n")] (abs_ge_zero RS real_le_imp_less_or_eq - RS disjE) 1 THEN dtac sym 2); -by (auto_tac (claset() addSIs [mult_right_mono], - simpset() addsimps [mult_assoc RS sym, power_abs,summable_def, power_0_left])); -by (res_inst_tac [("x","K * inverse(1 - (abs(z) * inverse(abs x)))")] exI 1); -by (auto_tac (claset() addSIs [sums_mult],simpset() addsimps [mult_assoc])); -by (subgoal_tac - "abs(z ^ n) * inverse(abs x ^ n) = (abs(z) * inverse(abs x)) ^ n" 1); -by (auto_tac (claset(),simpset() addsimps [power_abs RS sym])); -by (subgoal_tac "x ~= 0" 1); -by (subgoal_tac "x ~= 0" 3); -by (auto_tac (claset(), - simpset() delsimps [abs_inverse, abs_mult] - addsimps [abs_inverse RS sym, realpow_not_zero, abs_mult RS sym, - power_inverse, power_mult_distrib RS sym])); -by (auto_tac (claset() addSIs [geometric_sums], - simpset() addsimps [power_abs, inverse_eq_divide])); -by (res_inst_tac [("z","abs(x)")] (CLAIM_SIMP - "[|(0::real) x summable (%n. f(n) * (z ^ n))"; -by (dres_inst_tac [("z","abs z")] powser_insidea 1); -by (auto_tac (claset() addIs [summable_rabs_cancel], - simpset() addsimps [power_abs RS sym])); -qed "powser_inside"; - -(* ------------------------------------------------------------------------ *) -(* Differentiation of power series *) -(* ------------------------------------------------------------------------ *) - -(* Lemma about distributing negation over it *) -Goalw [diffs_def] "diffs (%n. - c n) = (%n. - diffs c n)"; -by Auto_tac; -qed "diffs_minus"; - -(* ------------------------------------------------------------------------ *) -(* Show that we can shift the terms down one *) -(* ------------------------------------------------------------------------ *) - -Goal "sumr 0 n (%n. (diffs c)(n) * (x ^ n)) = \ -\ sumr 0 n (%n. real n * c(n) * (x ^ (n - Suc 0))) + \ -\ (real n * c(n) * x ^ (n - Suc 0))"; -by (induct_tac "n" 1); -by (auto_tac (claset(),simpset() addsimps [real_mult_assoc, - real_add_assoc RS sym,diffs_def])); -qed "lemma_diffs"; - -Goal "sumr 0 n (%n. real n * c(n) * (x ^ (n - Suc 0))) = \ -\ sumr 0 n (%n. (diffs c)(n) * (x ^ n)) - \ -\ (real n * c(n) * x ^ (n - Suc 0))"; -by (auto_tac (claset(),simpset() addsimps [lemma_diffs])); -qed "lemma_diffs2"; - -Goal "summable (%n. (diffs c)(n) * (x ^ n)) ==> \ -\ (%n. real n * c(n) * (x ^ (n - Suc 0))) sums \ -\ (suminf(%n. (diffs c)(n) * (x ^ n)))"; -by (ftac summable_LIMSEQ_zero 1); -by (subgoal_tac "(%n. real n * c(n) * (x ^ (n - Suc 0))) ----> 0" 1); -by (rtac LIMSEQ_imp_Suc 2); -by (dtac summable_sums 1); -by (auto_tac (claset(),simpset() addsimps [sums_def])); -by (thin_tac "(%n. diffs c n * x ^ n) ----> 0" 1); -by (rotate_tac 1 1); -by (dtac LIMSEQ_diff 1); -by (auto_tac (claset(),simpset() addsimps [lemma_diffs2 RS sym, - symmetric diffs_def])); -by (asm_full_simp_tac (simpset() addsimps [diffs_def]) 1); -qed "diffs_equiv"; - -(* -------------------------------------------------------------------------*) -(* Term-by-term differentiability of power series *) -(* -------------------------------------------------------------------------*) - -Goal "sumr 0 m (%p. (((z + h) ^ (m - p)) * (z ^ p)) - (z ^ m)) = \ -\ sumr 0 m (%p. (z ^ p) * \ -\ (((z + h) ^ (m - p)) - (z ^ (m - p))))"; -by (rtac sumr_subst 1); -by (auto_tac (claset(),simpset() addsimps [right_distrib, - real_diff_def,power_add RS sym] - @ mult_ac)); -qed "lemma_termdiff1"; - -(* proved elsewhere? *) -Goal "m < n --> (EX d. n = m + d + Suc 0)"; -by (induct_tac "m" 1 THEN Auto_tac); -by (case_tac "n" 1); -by (case_tac "d" 3); -by (Auto_tac); -qed_spec_mp "less_add_one"; - -Goal " h ~= 0 ==> \ -\ (((z + h) ^ n) - (z ^ n)) * inverse h - \ -\ real n * (z ^ (n - Suc 0)) = \ -\ h * sumr 0 (n - Suc 0) (%p. (z ^ p) * \ -\ sumr 0 ((n - Suc 0) - p) \ -\ (%q. ((z + h) ^ q) * (z ^ (((n - 2) - p) - q))))"; -by (rtac (real_mult_left_cancel RS iffD1) 1 THEN Asm_simp_tac 1); -by (asm_full_simp_tac (simpset() addsimps [right_diff_distrib] - @ mult_ac) 1); -by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1); -by (case_tac "n" 1 THEN auto_tac (claset(),simpset() - addsimps [lemma_realpow_diff_sumr2, - right_diff_distrib RS sym,real_mult_assoc] - delsimps [realpow_Suc,sumr_Suc])); -by (auto_tac (claset(),simpset() addsimps [lemma_realpow_rev_sumr] - delsimps [sumr_Suc])); -by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc,sumr_diff_mult_const, - left_distrib,CLAIM "(a + b) - (c + d) = a - c + b - (d::real)", - lemma_termdiff1,sumr_mult])); -by (auto_tac (claset() addSIs [sumr_subst],simpset() addsimps - [real_diff_def,real_add_assoc])); -by (fold_tac [real_diff_def] THEN dtac less_add_one 1); -by (auto_tac (claset(),simpset() addsimps [sumr_mult,lemma_realpow_diff_sumr2] - @ mult_ac delsimps [sumr_Suc,realpow_Suc])); -qed "lemma_termdiff2"; - -Goal "[| h ~= 0; abs z <= K; abs (z + h) <= K |] \ -\ ==> abs (((z + h) ^ n - z ^ n) * inverse h - real n * z ^ (n - Suc 0)) \ -\ <= real n * real (n - Suc 0) * K ^ (n - 2) * abs h"; -by (stac lemma_termdiff2 1); -by (asm_full_simp_tac (simpset() addsimps [abs_mult,real_mult_commute]) 2); -by (stac real_mult_commute 2); -by (rtac (sumr_rabs RS real_le_trans) 2); -by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 2); -by (rtac (real_mult_commute RS subst) 2); -by (auto_tac (claset() addSIs [sumr_bound2],simpset() addsimps [abs_mult])); -by (case_tac "n" 1 THEN Auto_tac); -by (dtac less_add_one 1); -by (auto_tac (claset(),simpset() addsimps [power_add,real_add_assoc RS sym, - CLAIM_SIMP "(a * b) * c = a * (c * (b::real))" mult_ac] - delsimps [sumr_Suc])); -by (auto_tac (claset() addSIs [mult_mono],simpset()delsimps [sumr_Suc])); -by (auto_tac (claset() addSIs [power_mono], - simpset() addsimps [power_abs] delsimps [sumr_Suc] )); -by (res_inst_tac [("j","real (Suc d) * (K ^ d)")] real_le_trans 1); -by (subgoal_tac "0 <= K" 2); -by (arith_tac 3); -by (dres_inst_tac [("n","d")] zero_le_power 2); -by (auto_tac (claset(),simpset() delsimps [sumr_Suc] )); -by (rtac (sumr_rabs RS real_le_trans) 1); -by (rtac sumr_bound2 1 THEN - auto_tac (claset() addSDs [less_add_one] - addSIs [mult_mono], simpset() addsimps [abs_mult, power_add])); -by (auto_tac (claset() addSIs [power_mono,zero_le_power], - simpset() addsimps [power_abs])); -by (ALLGOALS(arith_tac)); -qed "lemma_termdiff3"; - -Goalw [LIM_def] - "[| 0 < k; \ -\ (ALL h. 0 < abs(h) & abs(h) < k --> abs(f h) <= K * abs(h)) |] \ -\ ==> f -- 0 --> 0"; -by (Auto_tac); -by (subgoal_tac "0 <= K" 1); -by (dres_inst_tac [("x","k/2")] spec 2); -by (ftac real_less_half_sum 2); -by (dtac real_gt_half_sum 2); -by (auto_tac (claset(),simpset() addsimps [abs_eqI2])); -by (res_inst_tac [("z","k/2")] (CLAIM_SIMP - "[| (0::real) x<=y" [mult_le_cancel_left]) 2); -by (auto_tac (claset() addIs [abs_ge_zero RS real_le_trans],simpset())); -by (dtac real_le_imp_less_or_eq 1); -by Auto_tac; -by (subgoal_tac "0 < (r * inverse K) * inverse 2" 1); -by (REPEAT(rtac (real_mult_order) 2)); -by (dres_inst_tac [("d1.0","r * inverse K * inverse 2"),("d2.0","k")] - real_lbound_gt_zero 1); -by (auto_tac (claset(),simpset() addsimps [positive_imp_inverse_positive, - zero_less_mult_iff])); -by (rtac real_le_trans 2 THEN assume_tac 3 THEN Auto_tac); -by (res_inst_tac [("x","e")] exI 1 THEN Auto_tac); -by (res_inst_tac [("y","K * abs x")] order_le_less_trans 1); -by (res_inst_tac [("y","K * e")] order_less_trans 2); -by (res_inst_tac [("z","inverse K")] (CLAIM_SIMP - "[|(0::real) x \ -\ (ALL n. abs(g(h) (n::nat)) <= (f(n) * abs(h))) |] \ -\ ==> (%h. suminf(g h)) -- 0 --> 0"; -by (dtac summable_sums 1); -by (subgoal_tac "ALL h. 0 < abs h & abs h < k --> \ -\ abs(suminf (g h)) <= suminf f * abs h" 1); -by (Auto_tac); -by (subgoal_tac "summable (%n. f n * abs h)" 2); -by (simp_tac (simpset() addsimps [summable_def]) 3); -by (res_inst_tac [("x","suminf f * abs h")] exI 3); -by (dres_inst_tac [("c","abs h")] sums_mult 3); -by (asm_full_simp_tac (simpset() addsimps mult_ac) 3); -by (subgoal_tac "summable (%n. abs(g(h::real)(n::nat)))" 2); -by (res_inst_tac [("g","%n. f(n::nat) * abs(h)")] summable_comparison_test 3); -by (res_inst_tac [("x","0")] exI 3); -by Auto_tac; -by (res_inst_tac [("j","suminf(%n. abs(g h n))")] real_le_trans 2); -by (auto_tac (claset() addIs [summable_rabs,summable_le],simpset() addsimps - [sums_summable RS suminf_mult])); -by (auto_tac (claset() addSIs [lemma_termdiff4],simpset() addsimps - [(sums_summable RS suminf_mult) RS sym])); -qed "lemma_termdiff5"; - -(* FIXME: Long proof *) -Goalw [deriv_def] - "[| summable(%n. c(n) * (K ^ n)); \ -\ summable(%n. (diffs c)(n) * (K ^ n)); \ -\ summable(%n. (diffs(diffs c))(n) * (K ^ n)); \ -\ abs(x) < abs(K) |] \ -\ ==> DERIV (%x. suminf (%n. c(n) * (x ^ n))) x :> \ -\ suminf (%n. (diffs c)(n) * (x ^ n))"; - -by (res_inst_tac [("g","%h. suminf(%n. ((c(n) * ((x + h) ^ n)) - \ -\ (c(n) * (x ^ n))) * inverse h)")] LIM_trans 1); -by (asm_full_simp_tac (simpset() addsimps [LIM_def]) 1); -by (Step_tac 1); -by (res_inst_tac [("x","abs K - abs x")] exI 1); -by (auto_tac (claset(),simpset() addsimps [less_diff_eq])); -by (dtac (abs_triangle_ineq RS order_le_less_trans) 1); -by (res_inst_tac [("y","0")] order_le_less_trans 1); -by Auto_tac; -by (subgoal_tac "(%n. (c n) * (x ^ n)) sums \ -\ (suminf(%n. (c n) * (x ^ n))) & \ -\ (%n. (c n) * ((x + xa) ^ n)) sums \ -\ (suminf(%n. (c n) * ((x + xa) ^ n)))" 1); -by (auto_tac (claset() addSIs [summable_sums],simpset())); -by (rtac powser_inside 2 THEN rtac powser_inside 4); -by (auto_tac (claset(),simpset() addsimps [real_add_commute])); -by (EVERY1[rotate_tac 8, dtac sums_diff, assume_tac]); -by (dres_inst_tac [("x","(%n. c n * (xa + x) ^ n - c n * x ^ n)"), - ("c","inverse xa")] sums_mult 1); -by (rtac (sums_unique RS sym) 1); -by (asm_full_simp_tac (simpset() addsimps [real_diff_def, - real_divide_def] @ add_ac @ mult_ac) 1); -by (rtac LIM_zero_cancel 1); -by (res_inst_tac [("g","%h. suminf (%n. c(n) * (((((x + h) ^ n) - \ -\ (x ^ n)) * inverse h) - (real n * (x ^ (n - Suc 0)))))")] LIM_trans 1); -by (asm_full_simp_tac (simpset() addsimps [LIM_def]) 1); -by (Step_tac 1); -by (res_inst_tac [("x","abs K - abs x")] exI 1); -by (auto_tac (claset(),simpset() addsimps [less_diff_eq])); -by (dtac (abs_triangle_ineq RS order_le_less_trans) 1); -by (res_inst_tac [("y","0")] order_le_less_trans 1); -by Auto_tac; -by (subgoal_tac "summable(%n. (diffs c)(n) * (x ^ n))" 1); -by (rtac powser_inside 2); -by (Auto_tac); -by (dres_inst_tac [("c","c"),("x","x")] diffs_equiv 1); -by (ftac sums_unique 1 THEN Auto_tac); -by (subgoal_tac "(%n. (c n) * (x ^ n)) sums \ -\ (suminf(%n. (c n) * (x ^ n))) & \ -\ (%n. (c n) * ((x + xa) ^ n)) sums \ -\ (suminf(%n. (c n) * ((x + xa) ^ n)))" 1); -by (Step_tac 1); -by (auto_tac (claset() addSIs [summable_sums],simpset())); -by (rtac powser_inside 2 THEN rtac powser_inside 4); -by (auto_tac (claset(),simpset() addsimps [real_add_commute])); -by (forw_inst_tac [("x","(%n. c n * (xa + x) ^ n)"), - ("y","(%n. c n * x ^ n)")] sums_diff 1 THEN assume_tac 1); -by (asm_full_simp_tac (simpset() addsimps [[sums_summable,sums_summable] - MRS suminf_diff,right_diff_distrib RS sym]) 1); -by (forw_inst_tac [("x","(%n. c n * ((xa + x) ^ n - x ^ n))"), - ("c","inverse xa")] sums_mult 1); -by (asm_full_simp_tac (simpset() addsimps [sums_summable RS suminf_mult2]) 1); -by (forw_inst_tac [("x","(%n. inverse xa * (c n * ((xa + x) ^ n - x ^ n)))"), - ("y","(%n. real n * c n * x ^ (n - Suc 0))")] sums_diff 1); -by (assume_tac 1); -by (rtac (ARITH_PROVE "z - y = x ==> - x = (y::real) - z") 1); -by (asm_full_simp_tac (simpset() addsimps [[sums_summable,sums_summable] - MRS suminf_diff] @ add_ac @ mult_ac ) 1); -by (res_inst_tac [("f","suminf")] arg_cong 1); -by (rtac ext 1); -by (asm_full_simp_tac (simpset() addsimps [real_diff_def, - right_distrib] @ add_ac @ mult_ac) 1); -(* 46 *) -by (dtac dense 1 THEN Step_tac 1); -by (ftac (real_less_sum_gt_zero) 1); -by (dres_inst_tac [("f","%n. abs(c n) * real n * \ -\ real (n - Suc 0) * (r ^ (n - 2))"), - ("g","%h n. c(n) * (((((x + h) ^ n) - (x ^ n)) * inverse h) - \ -\ (real n * (x ^ (n - Suc 0))))")] lemma_termdiff5 1); -by (auto_tac (claset(),simpset() addsimps [real_add_commute])); -by (subgoal_tac "summable(%n. abs(diffs(diffs c) n) * (r ^ n))" 1); -by (res_inst_tac [("x","K")] powser_insidea 2 THEN Auto_tac); -by (subgoal_tac "abs r = r" 2 THEN Auto_tac); -by (res_inst_tac [("j1","abs x")] (real_le_trans RS abs_eqI1) 2); -by Auto_tac; -by (asm_full_simp_tac (simpset() addsimps [diffs_def,abs_mult, - real_mult_assoc RS sym]) 1); -by (subgoal_tac "ALL n. real (Suc n) * real (Suc(Suc n)) * \ -\ abs(c(Suc(Suc n))) * (r ^ n) = diffs(diffs (%n. abs(c n))) n * (r ^ n)" 1); -by (dres_inst_tac [("P","summable")] - (CLAIM "[|ALL n. f(n) = g(n); P(%n. f n)|] ==> P(%n. g(n))") 1); -by (Auto_tac); -by (asm_full_simp_tac (simpset() addsimps [diffs_def]) 2 - THEN asm_full_simp_tac (simpset() addsimps [diffs_def]) 2); -by (dtac diffs_equiv 1); -by (dtac sums_summable 1); -by (asm_full_simp_tac (simpset() addsimps [diffs_def] @ mult_ac) 1); -by (subgoal_tac "(%n. real n * (real (Suc n) * (abs(c(Suc n)) * \ -\ (r ^ (n - Suc 0))))) = (%n. diffs(%m. real (m - Suc 0) * \ -\ abs(c m) * inverse r) n * (r ^ n))" 1); -by (Auto_tac); -by (rtac ext 2); -by (asm_full_simp_tac (simpset() addsimps [diffs_def]) 2); -by (case_tac "n" 2); -by Auto_tac; -(* 69 *) -by (dtac (abs_ge_zero RS order_le_less_trans) 2); -by (asm_full_simp_tac (simpset() addsimps mult_ac) 2); -by (dtac diffs_equiv 1); -by (dtac sums_summable 1); -by (res_inst_tac [("a","summable (%n. real n * \ -\ (real (n - Suc 0) * abs (c n) * inverse r) * r ^ (n - Suc 0))")] - (CLAIM "(a = b) ==> a ==> b") 1 THEN assume_tac 2); -by (res_inst_tac [("f","summable")] arg_cong 1 THEN rtac ext 1); -by (dtac (abs_ge_zero RS order_le_less_trans) 2); -by (asm_full_simp_tac (simpset() addsimps mult_ac) 2); -(* 77 *) -by (case_tac "n" 1); -by Auto_tac; -by (case_tac "nat" 1); -by Auto_tac; -by (dtac (abs_ge_zero RS order_le_less_trans) 1); -by (auto_tac (claset(),simpset() addsimps [CLAIM_SIMP - "(a::real) * (b * (c * d)) = a * (b * c) * d" - mult_ac])); -by (dtac (abs_ge_zero RS order_le_less_trans) 1); -by (asm_full_simp_tac (simpset() addsimps [abs_mult,real_mult_assoc]) 1); -by (rtac mult_left_mono 1); -by (rtac (add_commute RS subst) 1); -by (simp_tac (simpset() addsimps [mult_assoc RS sym]) 1); -by (rtac lemma_termdiff3 1); -by (auto_tac (claset() addIs [(abs_triangle_ineq RS real_le_trans)], - simpset())); -by (arith_tac 1); -qed "termdiffs"; - -(* ------------------------------------------------------------------------ *) -(* Formal derivatives of exp, sin, and cos series *) -(* ------------------------------------------------------------------------ *) - -Goalw [diffs_def] - "diffs (%n. inverse(real (fact n))) = (%n. inverse(real (fact n)))"; -by (rtac ext 1); -by (stac fact_Suc 1); -by (stac real_of_nat_mult 1); -by (stac inverse_mult_distrib 1); -by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym])); -qed "exp_fdiffs"; - -Goalw [diffs_def,real_divide_def] - "diffs(%n. if even n then 0 \ -\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) \ -\ = (%n. if even n then \ -\ (- 1) ^ (n div 2)/(real (fact n)) \ -\ else 0)"; -by (rtac ext 1); -by (stac fact_Suc 1); -by (stac real_of_nat_mult 1); -by (stac even_nat_Suc 1); -by (stac inverse_mult_distrib 1); -by Auto_tac; -qed "sin_fdiffs"; - -Goalw [diffs_def,real_divide_def] - "diffs(%n. if even n then 0 \ -\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) n \ -\ = (if even n then \ -\ (- 1) ^ (n div 2)/(real (fact n)) \ -\ else 0)"; -by (stac fact_Suc 1); -by (stac real_of_nat_mult 1); -by (stac even_nat_Suc 1); -by (stac inverse_mult_distrib 1); -by Auto_tac; -qed "sin_fdiffs2"; - -Goalw [diffs_def,real_divide_def] - "diffs(%n. if even n then \ -\ (- 1) ^ (n div 2)/(real (fact n)) else 0) \ -\ = (%n. - (if even n then 0 \ -\ else (- 1) ^ ((n - Suc 0)div 2)/(real (fact n))))"; -by (rtac ext 1); -by (stac fact_Suc 1); -by (stac real_of_nat_mult 1); -by (stac even_nat_Suc 1); -by (stac inverse_mult_distrib 1); -by (res_inst_tac [("a1","real (Suc n)")] (mult_commute RS ssubst) 1); -by (res_inst_tac [("a1","inverse(real (Suc n))")] - (mult_commute RS ssubst) 1); -by (auto_tac (claset(),simpset() addsimps [real_mult_assoc, - odd_Suc_mult_two_ex])); -qed "cos_fdiffs"; - - -Goalw [diffs_def,real_divide_def] - "diffs(%n. if even n then \ -\ (- 1) ^ (n div 2)/(real (fact n)) else 0) n\ -\ = - (if even n then 0 \ -\ else (- 1) ^ ((n - Suc 0)div 2)/(real (fact n)))"; -by (stac fact_Suc 1); -by (stac real_of_nat_mult 1); -by (stac even_nat_Suc 1); -by (stac inverse_mult_distrib 1); -by (res_inst_tac [("z1","real (Suc n)")] (real_mult_commute RS ssubst) 1); -by (res_inst_tac [("z1","inverse (real (Suc n))")] - (real_mult_commute RS ssubst) 1); -by (auto_tac (claset(),simpset() addsimps [real_mult_assoc, - odd_Suc_mult_two_ex])); -qed "cos_fdiffs2"; - -(* ------------------------------------------------------------------------ *) -(* Now at last we can get the derivatives of exp, sin and cos *) -(* ------------------------------------------------------------------------ *) - -Goal "- sin x = suminf(%n. - ((if even n then 0 \ -\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * x ^ n))"; -by (auto_tac (claset() addSIs [sums_unique,sums_minus,sin_converges], - simpset())); -qed "lemma_sin_minus"; - -Goal "exp = (%x. suminf (%n. inverse (real (fact n)) * x ^ n))"; -by (auto_tac (claset() addSIs [ext],simpset() addsimps [exp_def])); -qed "lemma_exp_ext"; - -Goalw [exp_def] "DERIV exp x :> exp(x)"; -by (stac lemma_exp_ext 1); -by (subgoal_tac "DERIV (%u. suminf (%n. inverse (real (fact n)) * u ^ n)) x \ -\ :> suminf (%n. diffs (%n. inverse (real (fact n))) n * x ^ n)" 1); -by (res_inst_tac [("K","1 + abs(x)")] termdiffs 2); -by (auto_tac (claset() addIs [exp_converges RS sums_summable], - simpset() addsimps [exp_fdiffs])); -by (arith_tac 1); -qed "DERIV_exp"; -Addsimps [DERIV_exp]; - -Goal "sin = (%x. suminf(%n. (if even n then 0 \ -\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * \ -\ x ^ n))"; -by (auto_tac (claset() addSIs [ext],simpset() addsimps [sin_def])); -qed "lemma_sin_ext"; - -Goal "cos = (%x. suminf(%n. (if even n then \ -\ (- 1) ^ (n div 2)/(real (fact n)) \ -\ else 0) * x ^ n))"; -by (auto_tac (claset() addSIs [ext],simpset() addsimps [cos_def])); -qed "lemma_cos_ext"; - -Goalw [cos_def] "DERIV sin x :> cos(x)"; -by (stac lemma_sin_ext 1); -by (auto_tac (claset(),simpset() addsimps [sin_fdiffs2 RS sym])); -by (res_inst_tac [("K","1 + abs(x)")] termdiffs 1); -by (auto_tac (claset() addIs [sin_converges, cos_converges, sums_summable] - addSIs [sums_minus RS sums_summable], - simpset() addsimps [cos_fdiffs,sin_fdiffs])); -by (arith_tac 1); -qed "DERIV_sin"; -Addsimps [DERIV_sin]; - -Goal "DERIV cos x :> -sin(x)"; -by (stac lemma_cos_ext 1); -by (auto_tac (claset(),simpset() addsimps [lemma_sin_minus, - cos_fdiffs2 RS sym,minus_mult_left])); -by (res_inst_tac [("K","1 + abs(x)")] termdiffs 1); -by (auto_tac (claset() addIs [sin_converges,cos_converges, sums_summable] - addSIs [sums_minus RS sums_summable], - simpset() addsimps [cos_fdiffs,sin_fdiffs,diffs_minus])); -by (arith_tac 1); -qed "DERIV_cos"; -Addsimps [DERIV_cos]; - -(* ------------------------------------------------------------------------ *) -(* Properties of the exponential function *) -(* ------------------------------------------------------------------------ *) - -Goalw [exp_def] "exp 0 = 1"; -by (rtac (CLAIM_SIMP "sumr 0 1 (%n. inverse (real (fact n)) * 0 ^ n) = 1" - [real_of_nat_one] RS subst) 1); -by (rtac ((series_zero RS sums_unique) RS sym) 1); -by (Step_tac 1); -by (case_tac "m" 1); -by (Auto_tac); -qed "exp_zero"; -Addsimps [exp_zero]; - -Goal "0 <= x ==> (1 + x) <= exp(x)"; -by (dtac real_le_imp_less_or_eq 1 THEN Auto_tac); -by (rewtac exp_def); -by (rtac real_le_trans 1); -by (res_inst_tac [("n","2"),("f","(%n. inverse (real (fact n)) * x ^ n)")] - series_pos_le 2); -by (auto_tac (claset() addIs [summable_exp],simpset() - addsimps [numeral_2_eq_2,zero_le_power,zero_le_mult_iff])); -qed "exp_ge_add_one_self"; -Addsimps [exp_ge_add_one_self]; - -Goal "0 < x ==> 1 < exp x"; -by (rtac order_less_le_trans 1); -by (rtac exp_ge_add_one_self 2); -by (Auto_tac); -qed "exp_gt_one"; -Addsimps [exp_gt_one]; - -Goal "DERIV (%x. exp (x + y)) x :> exp(x + y)"; -by (auto_tac (claset(),simpset() addsimps - [CLAIM_SIMP "(%x. exp (x + y)) = exp o (%x. x + y)" [ext]])); -by (rtac (real_mult_1_right RS subst) 1); -by (rtac DERIV_chain 1); -by (rtac (add_zero_right RS subst) 2); -by (rtac DERIV_add 2); -by Auto_tac; -qed "DERIV_exp_add_const"; -Addsimps [DERIV_exp_add_const]; - -Goal "DERIV (%x. exp (-x)) x :> - exp(-x)"; -by (auto_tac (claset(),simpset() addsimps - [CLAIM_SIMP "(%x. exp(-x)) = exp o (%x. - x)" [ext]])); -by (rtac (real_mult_1_right RS subst) 1); -by (rtac (minus_mult_left RS subst) 1); -by (stac minus_mult_right 1); -by (rtac DERIV_chain 1); -by (rtac DERIV_minus 2); -by Auto_tac; -qed "DERIV_exp_minus"; -Addsimps [DERIV_exp_minus]; - -Goal "DERIV (%x. exp (x + y) * exp (- x)) x :> 0"; -by (cut_inst_tac [("x","x"),("y2","y")] ([DERIV_exp_add_const, - DERIV_exp_minus] MRS DERIV_mult) 1); -by (auto_tac (claset(),simpset() addsimps mult_ac)); -qed "DERIV_exp_exp_zero"; -Addsimps [DERIV_exp_exp_zero]; - -Goal "exp(x + y)*exp(-x) = exp(y)"; -by (cut_inst_tac [("x","x"),("y2","y"),("y","0")] - ((CLAIM "ALL x. DERIV (%x. exp (x + y) * exp (- x)) x :> 0") RS - DERIV_isconst_all) 1); -by (Auto_tac); -qed "exp_add_mult_minus"; -Addsimps [exp_add_mult_minus]; - -Goal "exp(x)*exp(-x) = 1"; -by (cut_inst_tac [("x","x"),("y","0")] exp_add_mult_minus 1); -by (Asm_full_simp_tac 1); -qed "exp_mult_minus"; -Addsimps [exp_mult_minus]; - -Goal "exp(-x)*exp(x) = 1"; -by (simp_tac (simpset() addsimps [real_mult_commute]) 1); -qed "exp_mult_minus2"; -Addsimps [exp_mult_minus2]; - -Goal "exp(-x) = inverse(exp(x))"; -by (auto_tac (claset() addIs [real_inverse_unique],simpset())); -qed "exp_minus"; - -Goal "exp(x + y) = exp(x) * exp(y)"; -by (cut_inst_tac [("x1","x"),("y1","y"),("z","exp x")] - (exp_add_mult_minus RS (CLAIM "x = y ==> z * y = z * (x::real)")) 1); -by (asm_full_simp_tac HOL_ss 1); -by (asm_full_simp_tac (simpset() delsimps [exp_add_mult_minus] - addsimps mult_ac) 1); -qed "exp_add"; - -Goal "0 <= exp x"; -by (res_inst_tac [("t","x")] (real_sum_of_halves RS subst) 1); -by (stac exp_add 1 THEN Auto_tac); -qed "exp_ge_zero"; -Addsimps [exp_ge_zero]; - -Goal "exp x ~= 0"; -by (cut_inst_tac [("x","x")] exp_mult_minus2 1); -by (auto_tac (claset(),simpset() delsimps [exp_mult_minus2])); -qed "exp_not_eq_zero"; -Addsimps [exp_not_eq_zero]; - -Goal "0 < exp x"; -by (simp_tac (simpset() addsimps - [CLAIM_SIMP "(x < y) = (x <= y & y ~= (x::real))" [order_le_less]]) 1); -qed "exp_gt_zero"; -Addsimps [exp_gt_zero]; - -Goal "0 < inverse(exp x)"; -by (auto_tac (claset() addIs [positive_imp_inverse_positive],simpset())); -qed "inv_exp_gt_zero"; -Addsimps [inv_exp_gt_zero]; - -Goal "abs(exp x) = exp x"; -by (auto_tac (claset(),simpset() addsimps [abs_eqI2])); -qed "abs_exp_cancel"; -Addsimps [abs_exp_cancel]; - -Goal "exp(real n * x) = exp(x) ^ n"; -by (induct_tac "n" 1); -by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc, - right_distrib,exp_add,real_mult_commute])); -qed "exp_real_of_nat_mult"; - -Goalw [real_diff_def,real_divide_def] - "exp(x - y) = exp(x)/(exp y)"; -by (simp_tac (simpset() addsimps [exp_add,exp_minus]) 1); -qed "exp_diff"; - -Goal "x < y ==> exp x < exp y"; -by (dtac ((real_less_sum_gt_zero) RS exp_gt_one) 1); -by (multr_by_tac "inverse(exp x)" 1); -by (auto_tac (claset(),simpset() addsimps [exp_add,exp_minus])); -qed "exp_less_mono"; - -Goal "exp x < exp y ==> x < y"; -by (EVERY1[rtac ccontr, dtac (linorder_not_less RS iffD1), dtac real_le_imp_less_or_eq]); -by (auto_tac (claset() addDs [exp_less_mono],simpset())); -qed "exp_less_cancel"; - -Goal "(exp(x) < exp(y)) = (x < y)"; -by (auto_tac (claset() addIs [exp_less_mono,exp_less_cancel],simpset())); -qed "exp_less_cancel_iff"; -AddIffs [exp_less_cancel_iff]; - -Goal "(exp(x) <= exp(y)) = (x <= y)"; -by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym])); -qed "exp_le_cancel_iff"; -AddIffs [exp_le_cancel_iff]; - -Goal "(exp x = exp y) = (x = y)"; -by (auto_tac (claset(),simpset() addsimps - [CLAIM "(x = (y::real)) = (x <= y & y <= x)"])); -qed "exp_inj_iff"; -AddIffs [exp_inj_iff]; - -Goal "1 <= y ==> EX x. 0 <= x & x <= y - 1 & exp(x) = y"; -by (rtac IVT 1); -by (auto_tac (claset() addIs [DERIV_exp RS DERIV_isCont], - simpset() addsimps [le_diff_eq])); -by (dtac (CLAIM_SIMP "x <= y ==> (0::real) <= y - x" [le_diff_eq]) 1); -by (dtac exp_ge_add_one_self 1); -by (Asm_full_simp_tac 1); -qed "lemma_exp_total"; - -Goal "0 < y ==> EX x. exp x = y"; -by (res_inst_tac [("x","1"),("y","y")] linorder_cases 1); -by (dtac (order_less_imp_le RS lemma_exp_total) 1); -by (res_inst_tac [("x","0")] exI 2); -by (ftac real_inverse_gt_one 3); -by (dtac (order_less_imp_le RS lemma_exp_total) 4); -by (Step_tac 3); -by (res_inst_tac [("x","-x")] exI 3); -by (auto_tac (claset(),simpset() addsimps [exp_minus])); -qed "exp_total"; - -(* ------------------------------------------------------------------------ *) -(* Properties of the logarithmic function *) -(* ------------------------------------------------------------------------ *) - -Goalw [ln_def] "ln(exp x) = x"; -by (Simp_tac 1); -qed "ln_exp"; -Addsimps [ln_exp]; - -Goal "(exp(ln x) = x) = (0 < x)"; -by (auto_tac (claset() addDs [exp_total],simpset())); -by (dtac subst 1); -by (Auto_tac); -qed "exp_ln_iff"; -Addsimps [exp_ln_iff]; - -Goal "[| 0 < x; 0 < y |] ==> ln(x * y) = ln(x) + ln(y)"; -by (rtac (exp_inj_iff RS iffD1) 1); -by (ftac (real_mult_order) 1); -by (auto_tac (claset(),simpset() addsimps [exp_add,exp_ln_iff RS sym] - delsimps [exp_inj_iff,exp_ln_iff])); -qed "ln_mult"; - -Goal "[| 0 < x; 0 < y |] ==> (ln x = ln y) = (x = y)"; -by (auto_tac (claset() addSDs [(exp_ln_iff RS iffD2 RS sym)],simpset())); -qed "ln_inj_iff"; -Addsimps [ln_inj_iff]; - -Goal "ln 1 = 0"; -by (rtac (exp_inj_iff RS iffD1) 1); -by Auto_tac; -qed "ln_one"; -Addsimps [ln_one]; - -Goal "0 < x ==> ln(inverse x) = - ln x"; -by (res_inst_tac [("a1","ln x")] (add_left_cancel RS iffD1) 1); -by (auto_tac (claset(),simpset() addsimps [positive_imp_inverse_positive,ln_mult RS sym])); -qed "ln_inverse"; - -Goalw [real_divide_def] - "[|0 < x; 0 < y|] ==> ln(x/y) = ln x - ln y"; -by (auto_tac (claset(),simpset() addsimps [positive_imp_inverse_positive, - ln_mult,ln_inverse])); -qed "ln_div"; - -Goal "[| 0 < x; 0 < y|] ==> (ln x < ln y) = (x < y)"; -by (REPEAT(dtac (exp_ln_iff RS iffD2) 1)); -by (REPEAT(dtac subst 1 THEN assume_tac 2)); -by (Simp_tac 1); -qed "ln_less_cancel_iff"; -Addsimps [ln_less_cancel_iff]; - -Goal "[| 0 < x; 0 < y|] ==> (ln x <= ln y) = (x <= y)"; -by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym])); -by (Auto_tac); -qed "ln_le_cancel_iff"; -Addsimps [ln_le_cancel_iff]; - -Goal "0 < x ==> ln(x ^ n) = real n * ln(x)"; -by (auto_tac (claset() addSDs [exp_total],simpset() - addsimps [exp_real_of_nat_mult RS sym])); -qed "ln_realpow"; - -Goal "0 <= x ==> ln(1 + x) <= x"; -by (rtac (ln_exp RS subst) 1); -by (rtac (ln_le_cancel_iff RS iffD2) 1); -by Auto_tac; -qed "ln_add_one_self_le_self"; -Addsimps [ln_add_one_self_le_self]; - -Goal "0 < x ==> ln x < x"; -by (rtac order_less_le_trans 1); -by (rtac ln_add_one_self_le_self 2); -by (rtac (ln_less_cancel_iff RS iffD2) 1); -by Auto_tac; -qed "ln_less_self"; -Addsimps [ln_less_self]; - -Goal "1 <= x ==> 0 <= ln x"; -by (subgoal_tac "0 < x" 1); -by (rtac order_less_le_trans 2 THEN assume_tac 3); -by (rtac (exp_le_cancel_iff RS iffD1) 1); -by (auto_tac (claset(),simpset() addsimps - [exp_ln_iff RS sym] delsimps [exp_ln_iff])); -qed "ln_ge_zero"; -Addsimps [ln_ge_zero]; - -Goal "1 < x ==> 0 < ln x"; -by (rtac (exp_less_cancel_iff RS iffD1) 1); -by (rtac (exp_ln_iff RS iffD2 RS ssubst) 1); -by Auto_tac; -qed "ln_gt_zero"; -Addsimps [ln_gt_zero]; - -Goal "[| 0 < x; x ~= 1 |] ==> ln x ~= 0"; -by (Step_tac 1); -by (dtac (exp_inj_iff RS iffD2) 1); -by (dtac (exp_ln_iff RS iffD2) 1); -by Auto_tac; -qed "ln_not_eq_zero"; -Addsimps [ln_not_eq_zero]; - -Goal "[| 0 < x; x < 1 |] ==> ln x < 0"; -by (rtac (exp_less_cancel_iff RS iffD1) 1); -by (auto_tac (claset(),simpset() addsimps [exp_ln_iff RS sym] - delsimps [exp_ln_iff])); -qed "ln_less_zero"; - -Goal "exp u = x ==> ln x = u"; -by Auto_tac; -qed "exp_ln_eq"; - - -(* ------------------------------------------------------------------------ *) -(* Basic properties of the trig functions *) -(* ------------------------------------------------------------------------ *) - -Goalw [sin_def] "sin 0 = 0"; -by (auto_tac (claset() addSIs [sums_unique RS sym, LIMSEQ_const], - simpset() addsimps [sums_def] delsimps [power_0_left])); -qed "sin_zero"; -Addsimps [sin_zero]; - -Goal "(ALL m. n <= m --> f m = 0) --> f sums sumr 0 n f"; -by (auto_tac (claset() addIs [series_zero],simpset())); -qed "lemma_series_zero2"; - -Goalw [cos_def] "cos 0 = 1"; -by (rtac (sums_unique RS sym) 1); -by (cut_inst_tac [("n","1"),("f","(%n. (if even n then (- 1) ^ (n div 2)/ \ -\ (real (fact n)) else 0) * 0 ^ n)")] lemma_series_zero2 1); -by Auto_tac; -qed "cos_zero"; -Addsimps [cos_zero]; - -Goal "DERIV (%x. sin(x)*sin(x)) x :> cos(x) * sin(x) + cos(x) * sin(x)"; -by (rtac DERIV_mult 1 THEN Auto_tac); -qed "DERIV_sin_sin_mult"; -Addsimps [DERIV_sin_sin_mult]; - -Goal "DERIV (%x. sin(x)*sin(x)) x :> 2 * cos(x) * sin(x)"; -by (cut_inst_tac [("x","x")] DERIV_sin_sin_mult 1); -by (auto_tac (claset(),simpset() addsimps [real_mult_assoc])); -qed "DERIV_sin_sin_mult2"; -Addsimps [DERIV_sin_sin_mult2]; - -Goal "DERIV (%x. sin(x) ^ 2) x :> cos(x) * sin(x) + cos(x) * sin(x)"; -by (auto_tac (claset(), - simpset() addsimps [numeral_2_eq_2, real_mult_assoc RS sym])); -qed "DERIV_sin_realpow2"; -Addsimps [DERIV_sin_realpow2]; - -Goal "DERIV (%x. sin(x) ^ 2) x :> 2 * cos(x) * sin(x)"; -by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); -qed "DERIV_sin_realpow2a"; -Addsimps [ DERIV_sin_realpow2a]; - -Goal "DERIV (%x. cos(x)*cos(x)) x :> -sin(x) * cos(x) + -sin(x) * cos(x)"; -by (rtac DERIV_mult 1 THEN Auto_tac); -qed "DERIV_cos_cos_mult"; -Addsimps [DERIV_cos_cos_mult]; - -Goal "DERIV (%x. cos(x)*cos(x)) x :> -2 * cos(x) * sin(x)"; -by (cut_inst_tac [("x","x")] DERIV_cos_cos_mult 1); -by (auto_tac (claset(),simpset() addsimps mult_ac)); -qed "DERIV_cos_cos_mult2"; -Addsimps [DERIV_cos_cos_mult2]; - -Goal "DERIV (%x. cos(x) ^ 2) x :> -sin(x) * cos(x) + -sin(x) * cos(x)"; -by (auto_tac (claset(), - simpset() addsimps [numeral_2_eq_2, real_mult_assoc RS sym])); -qed "DERIV_cos_realpow2"; -Addsimps [DERIV_cos_realpow2]; - -Goal "DERIV (%x. cos(x) ^ 2) x :> -2 * cos(x) * sin(x)"; -by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); -qed "DERIV_cos_realpow2a"; -Addsimps [DERIV_cos_realpow2a]; - -Goal "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E"; -by (Auto_tac); -qed "lemma_DERIV_subst"; - -Goal "DERIV (%x. cos(x) ^ 2) x :> -(2 * cos(x) * sin(x))"; -by (rtac lemma_DERIV_subst 1); -by (rtac DERIV_cos_realpow2a 1); -by Auto_tac; -qed "DERIV_cos_realpow2b"; - -(* most useful *) -Goal "DERIV (%x. cos(x)*cos(x)) x :> -(2 * cos(x) * sin(x))"; -by (rtac lemma_DERIV_subst 1); -by (rtac DERIV_cos_cos_mult2 1); -by Auto_tac; -qed "DERIV_cos_cos_mult3"; -Addsimps [DERIV_cos_cos_mult3]; - -Goalw [real_diff_def] - "ALL x. DERIV (%x. sin(x) ^ 2 + cos(x) ^ 2) x :> \ -\ (2*cos(x)*sin(x) - 2*cos(x)*sin(x))"; -by (Step_tac 1); -by (rtac DERIV_add 1); -by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); -qed "DERIV_sin_circle_all"; - -Goal "ALL x. DERIV (%x. sin(x) ^ 2 + cos(x) ^ 2) x :> 0"; -by (cut_facts_tac [DERIV_sin_circle_all] 1); -by Auto_tac; -qed "DERIV_sin_circle_all_zero"; -Addsimps [DERIV_sin_circle_all_zero]; - -Goal "(sin(x) ^ 2) + (cos(x) ^ 2) = 1"; -by (cut_inst_tac [("x","x"),("y","0")] - (DERIV_sin_circle_all_zero RS DERIV_isconst_all) 1); -by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); -qed "sin_cos_squared_add"; -Addsimps [sin_cos_squared_add]; - -Goal "(cos(x) ^ 2) + (sin(x) ^ 2) = 1"; -by (stac real_add_commute 1); -by (simp_tac (simpset() delsimps [realpow_Suc]) 1); -qed "sin_cos_squared_add2"; -Addsimps [sin_cos_squared_add2]; - -Goal "cos x * cos x + sin x * sin x = 1"; -by (cut_inst_tac [("x","x")] sin_cos_squared_add2 1); -by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); -qed "sin_cos_squared_add3"; -Addsimps [sin_cos_squared_add3]; - -Goal "(sin(x) ^ 2) = 1 - (cos(x) ^ 2)"; -by (res_inst_tac [("a1","(cos(x) ^ 2)")] (add_right_cancel RS iffD1) 1); -by (simp_tac (simpset() delsimps [realpow_Suc]) 1); -qed "sin_squared_eq"; - -Goal "(cos(x) ^ 2) = 1 - (sin(x) ^ 2)"; -by (res_inst_tac [("a1","(sin(x) ^ 2)")] (add_right_cancel RS iffD1) 1); -by (simp_tac (simpset() delsimps [realpow_Suc]) 1); -qed "cos_squared_eq"; - -Goal "[| 1 < x; 0 <= y |] ==> 1 < x + (y::real)"; -by (arith_tac 1); -qed "real_gt_one_ge_zero_add_less"; - -Goal "abs(sin x) <= 1"; -by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym])); -by (dres_inst_tac [("n","Suc 0")] power_gt1 1); -by (auto_tac (claset(),simpset() delsimps [realpow_Suc])); -by (dres_inst_tac [("r1","cos x")] (realpow_two_le RSN - (2, real_gt_one_ge_zero_add_less)) 1); -by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2 RS sym] - delsimps [realpow_Suc]) 1); -qed "abs_sin_le_one"; -Addsimps [abs_sin_le_one]; - -Goal "- 1 <= sin x"; -by (full_simp_tac (simpset() addsimps [simplify (simpset()) (abs_sin_le_one RS - (abs_le_interval_iff RS iffD1))]) 1); -qed "sin_ge_minus_one"; -Addsimps [sin_ge_minus_one]; - -Goal "-1 <= sin x"; -by (rtac (simplify (simpset()) sin_ge_minus_one) 1); -qed "sin_ge_minus_one2"; -Addsimps [sin_ge_minus_one2]; - -Goal "sin x <= 1"; -by (full_simp_tac (simpset() addsimps [abs_sin_le_one RS - (abs_le_interval_iff RS iffD1)]) 1); -qed "sin_le_one"; -Addsimps [sin_le_one]; - -Goal "abs(cos x) <= 1"; -by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym])); -by (dres_inst_tac [("n","Suc 0")] power_gt1 1); -by (auto_tac (claset(),simpset() delsimps [realpow_Suc])); -by (dres_inst_tac [("r1","sin x")] (realpow_two_le RSN - (2, real_gt_one_ge_zero_add_less)) 1); -by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2 RS sym] - delsimps [realpow_Suc]) 1); -qed "abs_cos_le_one"; -Addsimps [abs_cos_le_one]; - -Goal "- 1 <= cos x"; -by (full_simp_tac (simpset() addsimps [simplify (simpset())(abs_cos_le_one RS - (abs_le_interval_iff RS iffD1))]) 1); -qed "cos_ge_minus_one"; -Addsimps [cos_ge_minus_one]; - -Goal "-1 <= cos x"; -by (rtac (simplify (simpset()) cos_ge_minus_one) 1); -qed "cos_ge_minus_one2"; -Addsimps [cos_ge_minus_one2]; - -Goal "cos x <= 1"; -by (full_simp_tac (simpset() addsimps [abs_cos_le_one RS - (abs_le_interval_iff RS iffD1)]) 1); -qed "cos_le_one"; -Addsimps [cos_le_one]; - -Goal "DERIV g x :> m ==> \ -\ DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m"; -by (rtac lemma_DERIV_subst 1); -by (res_inst_tac [("f","(%x. x ^ n)")] DERIV_chain2 1); -by (rtac DERIV_pow 1 THEN Auto_tac); -qed "DERIV_fun_pow"; - -Goal "DERIV g x :> m ==> DERIV (%x. exp(g x)) x :> exp(g x) * m"; -by (rtac lemma_DERIV_subst 1); -by (res_inst_tac [("f","exp")] DERIV_chain2 1); -by (rtac DERIV_exp 1 THEN Auto_tac); -qed "DERIV_fun_exp"; - -Goal "DERIV g x :> m ==> DERIV (%x. sin(g x)) x :> cos(g x) * m"; -by (rtac lemma_DERIV_subst 1); -by (res_inst_tac [("f","sin")] DERIV_chain2 1); -by (rtac DERIV_sin 1 THEN Auto_tac); -qed "DERIV_fun_sin"; - -Goal "DERIV g x :> m ==> DERIV (%x. cos(g x)) x :> -sin(g x) * m"; -by (rtac lemma_DERIV_subst 1); -by (res_inst_tac [("f","cos")] DERIV_chain2 1); -by (rtac DERIV_cos 1 THEN Auto_tac); -qed "DERIV_fun_cos"; - -(* FIXME: remove this quick, crude tactic *) -exception DERIV_name; -fun get_fun_name (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _)) = f -| get_fun_name (_ $ (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _))) = f -| get_fun_name _ = raise DERIV_name; - -val deriv_rulesI = [DERIV_Id,DERIV_const,DERIV_cos,DERIV_cmult, - DERIV_sin, DERIV_exp, DERIV_inverse,DERIV_pow, - DERIV_add, DERIV_diff, DERIV_mult, DERIV_minus, - DERIV_inverse_fun,DERIV_quotient,DERIV_fun_pow, - DERIV_fun_exp,DERIV_fun_sin,DERIV_fun_cos, - DERIV_Id,DERIV_const,DERIV_cos]; - - -fun deriv_tac i = (resolve_tac deriv_rulesI i) ORELSE - ((rtac (read_instantiate [("f",get_fun_name (getgoal i))] - DERIV_chain2) i) handle DERIV_name => no_tac); - -val DERIV_tac = ALLGOALS(fn i => REPEAT(deriv_tac i)); - -(* lemma *) -Goal "ALL x. \ -\ DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + \ -\ (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2) x :> 0"; -by (Step_tac 1 THEN rtac lemma_DERIV_subst 1); -by DERIV_tac; -by (auto_tac (claset(),simpset() addsimps [real_diff_def, - left_distrib,right_distrib] @ - mult_ac @ add_ac)); -qed "lemma_DERIV_sin_cos_add"; - -Goal "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + \ -\ (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = 0"; -by (cut_inst_tac [("y","0"),("x","x"),("y7","y")] - (lemma_DERIV_sin_cos_add RS DERIV_isconst_all) 1); -by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); -qed "sin_cos_add"; -Addsimps [sin_cos_add]; - -Goal "sin (x + y) = sin x * cos y + cos x * sin y"; -by (cut_inst_tac [("x","x"),("y","y")] sin_cos_add 1); -by (auto_tac (claset() addSDs [real_sum_squares_cancel_a], - simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_add])); -qed "sin_add"; - -Goal "cos (x + y) = cos x * cos y - sin x * sin y"; -by (cut_inst_tac [("x","x"),("y","y")] sin_cos_add 1); -by (auto_tac (claset() addSDs [real_sum_squares_cancel_a],simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_add])); -qed "cos_add"; - -Goal "ALL x. \ -\ DERIV (%x. (sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2) x :> 0"; -by (Step_tac 1 THEN rtac lemma_DERIV_subst 1); -by DERIV_tac; -by (auto_tac (claset(),simpset() addsimps [real_diff_def, - left_distrib,right_distrib] - @ mult_ac @ add_ac)); -qed "lemma_DERIV_sin_cos_minus"; - -Goal "(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0"; -by (cut_inst_tac [("y","0"),("x","x")] - (lemma_DERIV_sin_cos_minus RS DERIV_isconst_all) 1); -by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); -qed "sin_cos_minus"; -Addsimps [sin_cos_minus]; - -Goal "sin (-x) = -sin(x)"; -by (cut_inst_tac [("x","x")] sin_cos_minus 1); -by (auto_tac (claset() addSDs [real_sum_squares_cancel_a], - simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_minus])); -qed "sin_minus"; -Addsimps [sin_minus]; - -Goal "cos (-x) = cos(x)"; -by (cut_inst_tac [("x","x")] sin_cos_minus 1); -by (auto_tac (claset() addSDs [real_sum_squares_cancel_a], - simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_minus])); -qed "cos_minus"; -Addsimps [cos_minus]; - -Goalw [real_diff_def] "sin (x - y) = sin x * cos y - cos x * sin y"; -by (simp_tac (simpset() addsimps [sin_add]) 1); -qed "sin_diff"; - -Goal "sin (x - y) = cos y * sin x - sin y * cos x"; -by (simp_tac (simpset() addsimps [sin_diff,real_mult_commute]) 1); -qed "sin_diff2"; - -Goalw [real_diff_def] "cos (x - y) = cos x * cos y + sin x * sin y"; -by (simp_tac (simpset() addsimps [cos_add]) 1); -qed "cos_diff"; - -Goal "cos (x - y) = cos y * cos x + sin y * sin x"; -by (simp_tac (simpset() addsimps [cos_diff,real_mult_commute]) 1); -qed "cos_diff2"; - -Goal "sin(2 * x) = 2* sin x * cos x"; -by (cut_inst_tac [("x","x"),("y","x")] sin_add 1); -by Auto_tac; -qed "sin_double"; - -Addsimps [sin_double]; - -Goal "cos(2* x) = (cos(x) ^ 2) - (sin(x) ^ 2)"; -by (cut_inst_tac [("x","x"),("y","x")] cos_add 1); -by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); -qed "cos_double"; - -(* ------------------------------------------------------------------------ *) -(* Show that there's a least positive x with cos(x) = 0; hence define pi *) -(* ------------------------------------------------------------------------ *) - -Goal "(%n. (- 1) ^ n /(real (fact (2 * n + 1))) * \ -\ x ^ (2 * n + 1)) sums sin x"; -by (cut_inst_tac [("x2","x")] (CLAIM "0 < (2::nat)" RS ((sin_converges - RS sums_summable) RS sums_group)) 1); -by (auto_tac (claset(),simpset() addsimps mult_ac@[sin_def])); -qed "sin_paired"; - -Goal "real (Suc (Suc (Suc (Suc 2)))) = \ -\ real (2::nat) * real (Suc 2)"; -by (simp_tac (simpset() addsimps [numeral_2_eq_2, real_of_nat_Suc]) 1); -qed "lemma_real_of_nat_six_mult"; - -Goal "[|0 < x; x < 2 |] ==> 0 < sin x"; -by (cut_inst_tac [("x2","x")] (CLAIM "0 < (2::nat)" RS ((sin_paired - RS sums_summable) RS sums_group)) 1); -by (rotate_tac 2 1); -by (dtac ((sin_paired RS sums_unique) RS ssubst) 1); -by (auto_tac (claset(),simpset() delsimps [fact_Suc,realpow_Suc])); -by (ftac sums_unique 1); -by (auto_tac (claset(),simpset() delsimps [fact_Suc,realpow_Suc])); -by (res_inst_tac [("n1","0")] (series_pos_less RSN (2,order_le_less_trans)) 1); -by (auto_tac (claset(),simpset() delsimps [fact_Suc,realpow_Suc])); -by (etac sums_summable 1); -by (case_tac "m=0" 1); -by (Asm_simp_tac 1); -by (subgoal_tac "6 * (x * (x * x) / real (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) < 6 * x" 1); -by (asm_full_simp_tac (HOL_ss addsimps [mult_less_cancel_left]) 1); -by (asm_full_simp_tac (simpset() addsimps []) 1); -by (asm_simp_tac (simpset() addsimps [numeral_2_eq_2 RS sym, real_mult_assoc RS sym]) 1); -by (stac (CLAIM "6 = 2 * (3::real)") 1); -by (rtac real_mult_less_mono 1); (*mult_strict_mono would be stronger*) -by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc] delsimps [fact_Suc])); -by (stac fact_Suc 1); -by (stac fact_Suc 1); -by (stac fact_Suc 1); -by (stac fact_Suc 1); -by (stac real_of_nat_mult 1); -by (stac real_of_nat_mult 1); -by (stac real_of_nat_mult 1); -by (stac real_of_nat_mult 1); -by (simp_tac (simpset() addsimps [real_divide_def, - inverse_mult_distrib] delsimps [fact_Suc]) 1); -by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym] - delsimps [fact_Suc])); -by (multr_by_tac "real (Suc (Suc (4*m)))" 1); -by (auto_tac (claset(),simpset() addsimps [real_mult_assoc] - delsimps [fact_Suc])); -by (multr_by_tac "real (Suc (Suc (Suc (4*m))))" 1); -by (auto_tac (claset(),simpset() addsimps [mult_assoc,mult_less_cancel_left] - delsimps [fact_Suc])); -by (auto_tac (claset(),simpset() addsimps [ - CLAIM "x * (x * x ^ (4*m)) = (x ^ (4*m)) * (x * (x::real))"] - delsimps [fact_Suc])); -by (subgoal_tac "0 < x ^ (4 * m)" 1); -by (asm_simp_tac (simpset() addsimps [zero_less_power]) 2); -by (asm_simp_tac (simpset() addsimps [mult_less_cancel_left]) 1); -by (rtac real_mult_less_mono 1); (*mult_strict_mono would be stronger*) -by (ALLGOALS(Asm_simp_tac)); -by (TRYALL(rtac order_less_trans)); -by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc] delsimps [fact_Suc])); -qed "sin_gt_zero"; - -Goal "[|0 < x; x < 2 |] ==> 0 < sin x"; -by (auto_tac (claset() addIs [sin_gt_zero],simpset())); -qed "sin_gt_zero1"; - -Goal "[| 0 < x; x < 2 |] ==> cos (2 * x) < 1"; -by (cut_inst_tac [("x","x")] sin_gt_zero1 1); -by (auto_tac (claset(), simpset() addsimps [cos_squared_eq, cos_double])); -qed "cos_double_less_one"; - -Goal "(%n. (- 1) ^ n /(real (fact (2 * n))) * x ^ (2 * n)) \ -\ sums cos x"; -by (cut_inst_tac [("x2","x")] (CLAIM "0 < (2::nat)" RS ((cos_converges - RS sums_summable) RS sums_group)) 1); -by (auto_tac (claset(),simpset() addsimps mult_ac@[cos_def])); -qed "cos_paired"; - -Addsimps [zero_less_power]; - -Goal "cos (2) < 0"; -by (cut_inst_tac [("x","2")] cos_paired 1); -by (dtac sums_minus 1); -by (rtac (CLAIM "- x < -y ==> (y::real) < x") 1); -by (ftac sums_unique 1 THEN Auto_tac); -by (res_inst_tac [("y", - "sumr 0 (Suc (Suc (Suc 0))) (%n. -((- 1) ^ n /(real (fact(2 * n))) \ -\ * 2 ^ (2 * n)))")] order_less_trans 1); -by (simp_tac (simpset() addsimps [fact_num_eq_if,realpow_num_eq_if] - delsimps [fact_Suc,realpow_Suc]) 1); -by (simp_tac (simpset() addsimps [real_mult_assoc] - delsimps [sumr_Suc]) 1); -by (rtac sumr_pos_lt_pair 1); -by (etac sums_summable 1); -by (Step_tac 1); -by (simp_tac (simpset() addsimps [real_divide_def,real_mult_assoc RS sym] - delsimps [fact_Suc]) 1); -by (rtac real_mult_inverse_cancel2 1); -by (TRYALL(rtac (real_of_nat_fact_gt_zero))); -by (simp_tac (simpset() addsimps [real_mult_assoc RS sym] - delsimps [fact_Suc]) 1); -by (rtac ((CLAIM "real(n::nat) * 4 = real(4 * n)") RS ssubst) 1); -by (stac fact_Suc 1); -by (stac real_of_nat_mult 1); -by (stac real_of_nat_mult 1); -by (rtac real_mult_less_mono 1); (*mult_strict_mono would be stronger*) -by (Force_tac 1); -by (Force_tac 2); -by (rtac real_of_nat_fact_gt_zero 2); -by (rtac (real_of_nat_less_iff RS iffD2) 1); -by (rtac fact_less_mono 1); -by Auto_tac; -qed "cos_two_less_zero"; -Addsimps [cos_two_less_zero]; -Addsimps [cos_two_less_zero RS real_not_refl2]; -Addsimps [cos_two_less_zero RS order_less_imp_le]; - -Goal "EX! x. 0 <= x & x <= 2 & cos x = 0"; -by (subgoal_tac "EX x. 0 <= x & x <= 2 & cos x = 0" 1); -by (rtac IVT2 2); -by (auto_tac (claset() addIs [DERIV_isCont,DERIV_cos],simpset ())); -by (cut_inst_tac [("x","xa"),("y","y")] linorder_less_linear 1); -by (rtac ccontr 1); -by (subgoal_tac "(ALL x. cos differentiable x) & \ -\ (ALL x. isCont cos x)" 1); -by (auto_tac (claset() addIs [DERIV_cos,DERIV_isCont],simpset() - addsimps [differentiable_def])); -by (dres_inst_tac [("f","cos")] Rolle 1); -by (dres_inst_tac [("f","cos")] Rolle 5); -by (auto_tac (claset() addSDs [DERIV_cos RS DERIV_unique], - simpset() addsimps [differentiable_def])); -by (dres_inst_tac [("y1","xa")] (order_le_less_trans RS sin_gt_zero) 1); -by (assume_tac 1 THEN rtac order_less_le_trans 1); -by (dres_inst_tac [("y1","y")] (order_le_less_trans RS sin_gt_zero) 4); -by (assume_tac 4 THEN rtac order_less_le_trans 4); -by (assume_tac 1 THEN assume_tac 3); -by (ALLGOALS (Asm_full_simp_tac)); -qed "cos_is_zero"; - -Goalw [pi_def] "pi/2 = (@x. 0 <= x & x <= 2 & cos x = 0)"; -by Auto_tac; -qed "pi_half"; - -Goal "cos (pi / 2) = 0"; -by (rtac (cos_is_zero RS ex1E) 1); -by (auto_tac (claset() addSIs [someI2], - simpset() addsimps [pi_half])); -qed "cos_pi_half"; -Addsimps [cos_pi_half]; - -Goal "0 < pi / 2"; -by (rtac (cos_is_zero RS ex1E) 1); -by (auto_tac (claset(),simpset() addsimps [pi_half])); -by (rtac someI2 1); -by (Blast_tac 1); -by (Step_tac 1); -by (dres_inst_tac [("y","xa")] real_le_imp_less_or_eq 1); -by (Step_tac 1 THEN Asm_full_simp_tac 1); -qed "pi_half_gt_zero"; -Addsimps [pi_half_gt_zero]; -Addsimps [(pi_half_gt_zero RS real_not_refl2) RS not_sym]; -Addsimps [pi_half_gt_zero RS order_less_imp_le]; - -Goal "pi / 2 < 2"; -by (rtac (cos_is_zero RS ex1E) 1); -by (auto_tac (claset(),simpset() addsimps [pi_half])); -by (rtac someI2 1); -by (Blast_tac 1); -by (Step_tac 1); -by (dres_inst_tac [("x","xa")] order_le_imp_less_or_eq 1); -by (Step_tac 1 THEN Asm_full_simp_tac 1); -qed "pi_half_less_two"; -Addsimps [pi_half_less_two]; -Addsimps [pi_half_less_two RS real_not_refl2]; -Addsimps [pi_half_less_two RS order_less_imp_le]; - -Goal "0 < pi"; -by (multr_by_tac "inverse 2" 1); -by (Simp_tac 1); -by (cut_facts_tac [pi_half_gt_zero] 1); -by (full_simp_tac (HOL_ss addsimps [mult_zero_left, real_divide_def]) 1); -qed "pi_gt_zero"; -Addsimps [pi_gt_zero]; -Addsimps [(pi_gt_zero RS real_not_refl2) RS not_sym]; -Addsimps [pi_gt_zero RS CLAIM "(x::real) < y ==> ~ y < x"]; - -Goal "0 <= pi"; -by (auto_tac (claset() addIs [order_less_imp_le],simpset())); -qed "pi_ge_zero"; -Addsimps [pi_ge_zero]; - -Goal "-(pi/2) < 0"; -by Auto_tac; -qed "minus_pi_half_less_zero"; -Addsimps [minus_pi_half_less_zero]; - -Goal "sin(pi/2) = 1"; -by (cut_inst_tac [("x","pi/2")] sin_cos_squared_add2 1); -by (cut_facts_tac [[pi_half_gt_zero,pi_half_less_two] MRS sin_gt_zero] 1); -by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); -qed "sin_pi_half"; -Addsimps [sin_pi_half]; - -Goal "cos pi = - 1"; -by (cut_inst_tac [("x","pi/2"),("y","pi/2")] cos_add 1); -by (Asm_full_simp_tac 1); -qed "cos_pi"; -Addsimps [cos_pi]; - -Goal "sin pi = 0"; -by (cut_inst_tac [("x","pi/2"),("y","pi/2")] sin_add 1); -by (Asm_full_simp_tac 1); -qed "sin_pi"; -Addsimps [sin_pi]; - -Goalw [real_diff_def] "sin x = cos (pi/2 - x)"; -by (simp_tac (simpset() addsimps [cos_add]) 1); -qed "sin_cos_eq"; - -Goal "-sin x = cos (x + pi/2)"; -by (simp_tac (simpset() addsimps [cos_add]) 1); -qed "minus_sin_cos_eq"; -Addsimps [minus_sin_cos_eq RS sym]; - -Goalw [real_diff_def] "cos x = sin (pi/2 - x)"; -by (simp_tac (simpset() addsimps [sin_add]) 1); -qed "cos_sin_eq"; -Addsimps [sin_cos_eq RS sym, cos_sin_eq RS sym]; - -Goal "sin (x + pi) = - sin x"; -by (simp_tac (simpset() addsimps [sin_add]) 1); -qed "sin_periodic_pi"; -Addsimps [sin_periodic_pi]; - -Goal "sin (pi + x) = - sin x"; -by (simp_tac (simpset() addsimps [sin_add]) 1); -qed "sin_periodic_pi2"; -Addsimps [sin_periodic_pi2]; - -Goal "cos (x + pi) = - cos x"; -by (simp_tac (simpset() addsimps [cos_add]) 1); -qed "cos_periodic_pi"; -Addsimps [cos_periodic_pi]; - -Goal "sin (x + 2*pi) = sin x"; -by (simp_tac (simpset() addsimps [sin_add,cos_double,numeral_2_eq_2]) 1); - (*FIXME: just needs x^n for literals!*) -qed "sin_periodic"; -Addsimps [sin_periodic]; - -Goal "cos (x + 2*pi) = cos x"; -by (simp_tac (simpset() addsimps [cos_add,cos_double,numeral_2_eq_2]) 1); - (*FIXME: just needs x^n for literals!*) -qed "cos_periodic"; -Addsimps [cos_periodic]; - -Goal "cos (real n * pi) = (-(1::real)) ^ n"; -by (induct_tac "n" 1); -by (auto_tac (claset(),simpset() addsimps - [real_of_nat_Suc,left_distrib])); -qed "cos_npi"; -Addsimps [cos_npi]; - -Goal "sin (real (n::nat) * pi) = 0"; -by (induct_tac "n" 1); -by (auto_tac (claset(),simpset() addsimps - [real_of_nat_Suc,left_distrib])); -qed "sin_npi"; -Addsimps [sin_npi]; - -Goal "sin (pi * real (n::nat)) = 0"; -by (cut_inst_tac [("n","n")] sin_npi 1); -by (auto_tac (claset(),simpset() addsimps [real_mult_commute] - delsimps [sin_npi])); -qed "sin_npi2"; -Addsimps [sin_npi2]; - -Goal "cos (2 * pi) = 1"; -by (simp_tac (simpset() addsimps [cos_double,numeral_2_eq_2]) 1); - (*FIXME: just needs x^n for literals!*) -qed "cos_two_pi"; -Addsimps [cos_two_pi]; - -Goal "sin (2 * pi) = 0"; -by (Simp_tac 1); -qed "sin_two_pi"; -Addsimps [sin_two_pi]; - -Goal "[| 0 < x; x < pi/2 |] ==> 0 < sin x"; -by (rtac sin_gt_zero 1); -by (assume_tac 1); -by (rtac order_less_trans 1 THEN assume_tac 1); -by (rtac pi_half_less_two 1); -qed "sin_gt_zero2"; - -Goal "[| - pi/2 < x; x < 0 |] ==> sin x < 0"; -by (rtac (CLAIM "(0::real) < - x ==> x < 0") 1); -by (rtac (sin_minus RS subst) 1); -by (rtac sin_gt_zero2 1); -by (rtac (CLAIM "-y < x ==> -x < (y::real)") 2); -by Auto_tac; -qed "sin_less_zero"; - -Goal "pi < 4"; -by (cut_facts_tac [pi_half_less_two] 1); -by Auto_tac; -qed "pi_less_4"; - -Goal "[| 0 < x; x < pi/2 |] ==> 0 < cos x"; -by (cut_facts_tac [pi_less_4] 1); -by (cut_inst_tac [("f","cos"),("a","0"),("b","x"),("y","0")] IVT2_objl 1); -by (Step_tac 1); -by (cut_facts_tac [cos_is_zero] 5); -by (Step_tac 5); -by (dres_inst_tac [("x","xa")] spec 5); -by (dres_inst_tac [("x","pi/2")] spec 5); -by (force_tac (claset(), simpset() addsimps []) 1); -by (force_tac (claset(), simpset() addsimps []) 1); -by (force_tac (claset(), simpset() addsimps []) 1); -by (auto_tac (claset() addSDs [ pi_half_less_two RS order_less_trans, - CLAIM "~ m <= n ==> n < (m::real)"] - addIs [DERIV_isCont,DERIV_cos],simpset())); -qed "cos_gt_zero"; - -Goal "[| -(pi/2) < x; x < pi/2 |] ==> 0 < cos x"; -by (res_inst_tac [("x","x"),("y","0")] linorder_cases 1); -by (rtac (cos_minus RS subst) 1); -by (rtac cos_gt_zero 1); -by (rtac (CLAIM "-y < x ==> -x < (y::real)") 2); -by (auto_tac (claset() addIs [cos_gt_zero],simpset())); -qed "cos_gt_zero_pi"; - -Goal "[| -(pi/2) <= x; x <= pi/2 |] ==> 0 <= cos x"; -by (auto_tac (claset(),HOL_ss addsimps [order_le_less, - cos_gt_zero_pi])); -by Auto_tac; -qed "cos_ge_zero"; - -Goal "[| 0 < x; x < pi |] ==> 0 < sin x"; -by (stac sin_cos_eq 1); -by (rotate_tac 1 1); -by (dtac (real_sum_of_halves RS ssubst) 1); -by (auto_tac (claset() addSIs [cos_gt_zero_pi], - simpset() delsimps [sin_cos_eq RS sym])); -qed "sin_gt_zero_pi"; - -Goal "[| 0 <= x; x <= pi |] ==> 0 <= sin x"; -by (auto_tac (claset(),simpset() addsimps [order_le_less, - sin_gt_zero_pi])); -qed "sin_ge_zero"; - -Goal "[| - 1 <= y; y <= 1 |] ==> EX! x. 0 <= x & x <= pi & (cos x = y)"; -by (subgoal_tac "EX x. 0 <= x & x <= pi & cos x = y" 1); -by (rtac IVT2 2); -by (auto_tac (claset() addIs [order_less_imp_le,DERIV_isCont,DERIV_cos], - simpset ())); -by (cut_inst_tac [("x","xa"),("y","y")] linorder_less_linear 1); -by (rtac ccontr 1 THEN Auto_tac); -by (dres_inst_tac [("f","cos")] Rolle 1); -by (dres_inst_tac [("f","cos")] Rolle 5); -by (auto_tac (claset() addIs [order_less_imp_le,DERIV_isCont,DERIV_cos] - addSDs [DERIV_cos RS DERIV_unique],simpset() addsimps [differentiable_def])); -by (auto_tac (claset() addDs [[order_le_less_trans,order_less_le_trans] MRS - sin_gt_zero_pi],simpset())); -qed "cos_total"; - -Goal "[| - 1 <= y; y <= 1 |] ==> \ -\ EX! x. -(pi/2) <= x & x <= pi/2 & (sin x = y)"; -by (rtac ccontr 1); -by (subgoal_tac "ALL x. (-(pi/2) <= x & x <= pi/2 & (sin x = y)) \ -\ = (0 <= (x + pi/2) & (x + pi/2) <= pi & \ -\ (cos(x + pi/2) = -y))" 1); -by (etac swap 1); -by (asm_full_simp_tac (simpset() delsimps [minus_sin_cos_eq RS sym]) 1); -by (dtac (CLAIM "(x::real) <= y ==> -y <= -x") 1); -by (dtac (CLAIM "(x::real) <= y ==> -y <= -x") 1); -by (dtac cos_total 1); -by (Asm_full_simp_tac 1); -by (etac ex1E 1); -by (res_inst_tac [("a","x - (pi/2)")] ex1I 1); -by (simp_tac (simpset() addsimps [real_add_assoc]) 1); -by (rotate_tac 3 1); -by (dres_inst_tac [("x","xa + pi/2")] spec 1); -by (Step_tac 1); -by (TRYALL(Asm_full_simp_tac)); -by (auto_tac (claset(),simpset() addsimps [CLAIM "(-x <= y) = (-y <= (x::real))"])); -qed "sin_total"; - -Goal "(EX n. P (n::nat)) = (EX n. P n & (ALL m. m < n --> ~ P m))"; -by (rtac iffI 1); -by (rtac contrapos_pp 1 THEN assume_tac 1); -by (EVERY1[Simp_tac, rtac allI, rtac nat_less_induct]); -by (Auto_tac); -qed "less_induct_ex_iff"; - -Goal "[| 0 < y; 0 <= x |] ==> \ -\ EX n. real n * y <= x & x < real (Suc n) * y"; -by (auto_tac (claset() addSDs [reals_Archimedean3],simpset())); -by (dres_inst_tac [("x","x")] spec 1); -by (dtac (less_induct_ex_iff RS iffD1) 1 THEN Step_tac 1); -by (case_tac "n" 1); -by (res_inst_tac [("x","nat")] exI 2); -by Auto_tac; -qed "reals_Archimedean4"; - -(* Pre Isabelle99-2 proof was simpler- numerals arithmetic - now causes some unwanted re-arrangements of literals! *) -Goal "[| 0 <= x; cos x = 0 |] ==> \ -\ EX n::nat. ~even n & x = real n * (pi/2)"; -by (dtac (pi_gt_zero RS reals_Archimedean4) 1); -by (Step_tac 1); -by (subgoal_tac - "0 <= x - real n * pi & (x - real n * pi) <= pi & \ -\ (cos(x - real n * pi) = 0)" 1); -by (Step_tac 1); -by (asm_full_simp_tac (simpset() addsimps [real_of_nat_Suc, - left_distrib]) 2); -by (asm_full_simp_tac (simpset() addsimps [cos_diff]) 1); -by (asm_full_simp_tac (simpset() addsimps [cos_diff]) 2); -by (subgoal_tac "EX! x. 0 <= x & x <= pi & cos x = 0" 1); -by (rtac cos_total 2); -by (Step_tac 1); -by (dres_inst_tac [("x","x - real n * pi")] spec 1); -by (dres_inst_tac [("x","pi/2")] spec 1); -by (asm_full_simp_tac (simpset() addsimps [cos_diff]) 1); -by (res_inst_tac [("x","Suc (2 * n)")] exI 1); -by (asm_full_simp_tac (simpset() addsimps [real_of_nat_Suc, - left_distrib]) 1); -by Auto_tac; -qed "cos_zero_lemma"; - -Goal "[| 0 <= x; sin x = 0 |] ==> \ -\ EX n::nat. even n & x = real n * (pi/2)"; -by (subgoal_tac - "EX n. ~ even n & x + pi/2 = real n * (pi/2)" 1); -by (rtac cos_zero_lemma 2); -by (Step_tac 1); -by (res_inst_tac [("x","n - 1")] exI 1); -by (rtac (CLAIM "-y <= x ==> -x <= (y::real)") 2); -by (rtac real_le_trans 2 THEN assume_tac 3); -by (auto_tac (claset(),simpset() addsimps [ - odd_Suc_mult_two_ex,real_of_nat_Suc, - left_distrib,real_mult_assoc RS sym])); -qed "sin_zero_lemma"; - -(* also spoilt by numeral arithmetic *) -Goal "(cos x = 0) = \ -\ ((EX n::nat. ~even n & (x = real n * (pi/2))) | \ -\ (EX n::nat. ~even n & (x = -(real n * (pi/2)))))"; -by (rtac iffI 1); -by (cut_inst_tac [("x","x")] (CLAIM "0 <= (x::real) | x <= 0") 1); -by (Step_tac 1); -by (dtac cos_zero_lemma 1); -by (dtac (CLAIM "(x::real) <= 0 ==> 0 <= -x") 3); -by (dtac cos_zero_lemma 3); -by (Step_tac 1); -by (dtac (CLAIM "-x = y ==> x = -(y::real)") 2); -by (auto_tac (claset(),HOL_ss addsimps [ - odd_Suc_mult_two_ex,real_of_nat_Suc,left_distrib])); -by (auto_tac (claset(),simpset() addsimps [cos_add])); -qed "cos_zero_iff"; - -(* ditto: but to a lesser extent *) -Goal "(sin x = 0) = \ -\ ((EX n::nat. even n & (x = real n * (pi/2))) | \ -\ (EX n::nat. even n & (x = -(real n * (pi/2)))))"; -by (rtac iffI 1); -by (cut_inst_tac [("x","x")] (CLAIM "0 <= (x::real) | x <= 0") 1); -by (Step_tac 1); -by (dtac sin_zero_lemma 1); -by (dtac (CLAIM "(x::real) <= 0 ==> 0 <= -x") 3); -by (dtac sin_zero_lemma 3); -by (Step_tac 1); -by (dtac (CLAIM "-x = y ==> x = -(y::real)") 2); -by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex])); -qed "sin_zero_iff"; - -(* ------------------------------------------------------------------------ *) -(* Tangent *) -(* ------------------------------------------------------------------------ *) - -Goalw [tan_def] "tan 0 = 0"; -by (Simp_tac 1); -qed "tan_zero"; -Addsimps [tan_zero]; - -Goalw [tan_def] "tan pi = 0"; -by (Simp_tac 1); -qed "tan_pi"; -Addsimps [tan_pi]; - -Goalw [tan_def] "tan (real (n::nat) * pi) = 0"; -by (Simp_tac 1); -qed "tan_npi"; -Addsimps [tan_npi]; - -Goalw [tan_def] "tan (-x) = - tan x"; -by (simp_tac (simpset() addsimps [minus_mult_left]) 1); -qed "tan_minus"; -Addsimps [tan_minus]; - -Goalw [tan_def] "tan (x + 2*pi) = tan x"; -by (Simp_tac 1); -qed "tan_periodic"; -Addsimps [tan_periodic]; - -Goalw [tan_def,real_divide_def] - "[| cos x ~= 0; cos y ~= 0 |] \ -\ ==> 1 - tan(x)*tan(y) = cos (x + y)/(cos x * cos y)"; -by (auto_tac (claset(), - simpset() delsimps [inverse_mult_distrib] - addsimps [inverse_mult_distrib RS sym] @ mult_ac)); -by (res_inst_tac [("c1","cos x * cos y")] (real_mult_right_cancel RS subst) 1); -by (auto_tac (claset(), - simpset() delsimps [inverse_mult_distrib] - addsimps [mult_assoc, left_diff_distrib,cos_add])); -qed "lemma_tan_add1"; -Addsimps [lemma_tan_add1]; - -Goalw [tan_def] - "[| cos x ~= 0; cos y ~= 0 |] \ -\ ==> tan x + tan y = sin(x + y)/(cos x * cos y)"; -by (res_inst_tac [("c1","cos x * cos y")] (real_mult_right_cancel RS subst) 1); -by (auto_tac (claset(), simpset() addsimps [mult_assoc, left_distrib])); -by (simp_tac (simpset() addsimps [sin_add]) 1); -qed "add_tan_eq"; - -Goal "[| cos x ~= 0; cos y ~= 0; cos (x + y) ~= 0 |] \ -\ ==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))"; -by (asm_simp_tac (simpset() addsimps [add_tan_eq]) 1); -by (simp_tac (simpset() addsimps [tan_def]) 1); -qed "tan_add"; - -Goal "[| cos x ~= 0; cos (2 * x) ~= 0 |] \ -\ ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))"; -by (auto_tac (claset(),simpset() addsimps [asm_full_simplify - (simpset() addsimps [thm"mult_2" RS sym] delsimps [lemma_tan_add1]) - (read_instantiate [("x","x"),("y","x")] tan_add),numeral_2_eq_2] - delsimps [lemma_tan_add1])); -qed "tan_double"; - -Goalw [tan_def,real_divide_def] "[| 0 < x; x < pi/2 |] ==> 0 < tan x"; -by (auto_tac (claset() addSIs [sin_gt_zero2,cos_gt_zero_pi] - addSIs [real_mult_order, positive_imp_inverse_positive],simpset())); -qed "tan_gt_zero"; - -Goal "[| - pi/2 < x; x < 0 |] ==> tan x < 0"; -by (rtac (CLAIM "(0::real) < - x ==> x < 0") 1); -by (rtac (tan_minus RS subst) 1); -by (rtac tan_gt_zero 1); -by (rtac (CLAIM "-x < y ==> -y < (x::real)") 2 THEN Auto_tac); -qed "tan_less_zero"; - -Goal "cos x ~= 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse(cos x ^ 2)"; -by (rtac lemma_DERIV_subst 1); -by DERIV_tac; -by (auto_tac (claset(),simpset() addsimps [real_divide_def,numeral_2_eq_2])); -qed "lemma_DERIV_tan"; - -Goal "cos x ~= 0 ==> DERIV tan x :> inverse(cos(x) ^ 2)"; -by (auto_tac (claset() addDs [lemma_DERIV_tan],simpset() - addsimps [(tan_def RS meta_eq_to_obj_eq) RS sym])); -qed "DERIV_tan"; -Addsimps [DERIV_tan]; - -Goalw [real_divide_def] - "(%x. cos(x)/sin(x)) -- pi/2 --> 0"; -by (res_inst_tac [("a1","1")] ((mult_zero_left) RS subst) 1); -by (rtac LIM_mult2 1); -by (rtac (inverse_1 RS subst) 2); -by (rtac LIM_inverse 2); -by (fold_tac [real_divide_def]); -by (auto_tac (claset() addSIs [DERIV_isCont],simpset() - addsimps [(isCont_def RS meta_eq_to_obj_eq) - RS sym, cos_pi_half RS sym, sin_pi_half RS sym] - delsimps [cos_pi_half,sin_pi_half])); -by (DERIV_tac THEN Auto_tac); -qed "LIM_cos_div_sin"; -Addsimps [LIM_cos_div_sin]; - -Goal "0 < y ==> EX x. 0 < x & x < pi/2 & y < tan x"; -by (cut_facts_tac [LIM_cos_div_sin] 1); -by (asm_full_simp_tac (HOL_ss addsimps [LIM_def]) 1); -by (dres_inst_tac [("x","inverse y")] spec 1); -by (Step_tac 1); -by (Force_tac 1); -by (dres_inst_tac [("d1.0","s")] - (pi_half_gt_zero RSN (2,real_lbound_gt_zero)) 1); -by (Step_tac 1); -by (res_inst_tac [("x","(pi/2) - e")] exI 1); -by (Asm_simp_tac 1); -by (dres_inst_tac [("x","(pi/2) - e")] spec 1); -by (auto_tac (claset(),simpset() addsimps [abs_eqI2,tan_def])); -by (rtac (inverse_less_iff_less RS iffD1) 1); -by (auto_tac (claset(),simpset() addsimps [real_divide_def])); -by (rtac (real_mult_order) 1); -by (subgoal_tac "0 < sin e" 3); -by (subgoal_tac "0 < cos e" 3); -by (auto_tac (claset() addIs [cos_gt_zero,sin_gt_zero2],simpset() - addsimps [inverse_mult_distrib,abs_mult])); -by (dres_inst_tac [("a","cos e")] (positive_imp_inverse_positive) 1); -by (dres_inst_tac [("x","inverse (cos e)")] abs_eqI2 1); -by (auto_tac (claset() addSDs [abs_eqI2],simpset() addsimps mult_ac)); -qed "lemma_tan_total"; - - -Goal "0 <= y ==> EX x. 0 <= x & x < pi/2 & tan x = y"; -by (ftac real_le_imp_less_or_eq 1); -by (Step_tac 1 THEN Force_tac 2); -by (dtac lemma_tan_total 1 THEN Step_tac 1); -by (cut_inst_tac [("f","tan"),("a","0"),("b","x"),("y","y")] IVT_objl 1); -by (auto_tac (claset() addSIs [DERIV_tan RS DERIV_isCont],simpset())); -by (dres_inst_tac [("y","xa")] order_le_imp_less_or_eq 1); -by (auto_tac (claset() addDs [cos_gt_zero],simpset())); -qed "tan_total_pos"; - -Goal "EX x. -(pi/2) < x & x < (pi/2) & tan x = y"; -by (cut_inst_tac [("y","y")] (CLAIM "0 <= (y::real) | 0 <= -y") 1); -by (Step_tac 1); -by (dtac tan_total_pos 1); -by (dtac tan_total_pos 2); -by (Step_tac 1); -by (res_inst_tac [("x","-x")] exI 2); -by (auto_tac (claset() addSIs [exI],simpset())); -qed "lemma_tan_total1"; - -Goal "EX! x. -(pi/2) < x & x < (pi/2) & tan x = y"; -by (cut_inst_tac [("y","y")] lemma_tan_total1 1); -by (Auto_tac); -by (cut_inst_tac [("x","xa"),("y","y")] linorder_less_linear 1); -by (Auto_tac); -by (subgoal_tac "EX z. xa < z & z < y & DERIV tan z :> 0" 1); -by (subgoal_tac "EX z. y < z & z < xa & DERIV tan z :> 0" 3); -by (rtac Rolle 2); -by (rtac Rolle 7); -by (auto_tac (claset() addSIs [DERIV_tan,DERIV_isCont,exI],simpset() - addsimps [differentiable_def])); -by (TRYALL(rtac DERIV_tan)); -by (TRYALL(dtac (DERIV_tan RSN (2,DERIV_unique)))); -by (TRYALL(rtac (real_not_refl2 RS not_sym))); -by (auto_tac (claset() addSIs [cos_gt_zero_pi],simpset())); -by (ALLGOALS(subgoal_tac "0 < cos z")); -by (Force_tac 1 THEN Force_tac 2); -by (ALLGOALS(thin_tac "cos z = 0")); -by (auto_tac (claset() addSIs [cos_gt_zero_pi],simpset())); -qed "tan_total"; - -Goal "[| - 1 <= y; y <= 1 |] \ -\ ==> -(pi/2) <= arcsin y & arcsin y <= pi & sin(arcsin y) = y"; -by (dtac sin_total 1); -by (etac ex1E 2); -by (rewtac arcsin_def); -by (rtac someI2 2); -by (EVERY1[assume_tac, Blast_tac, Step_tac]); -by (rtac real_le_trans 1 THEN assume_tac 1); -by (Force_tac 1); -qed "arcsin_pi"; - -Goal "[| - 1 <= y; y <= 1 |] \ -\ ==> -(pi/2) <= arcsin y & \ -\ arcsin y <= pi/2 & sin(arcsin y) = y"; -by (dtac sin_total 1 THEN assume_tac 1); -by (etac ex1E 1); -by (rewtac arcsin_def); -by (rtac someI2 1); -by (ALLGOALS(Blast_tac)); -qed "arcsin"; - -Goal "[| - 1 <= y; y <= 1 |] ==> sin(arcsin y) = y"; -by (blast_tac (claset() addDs [arcsin]) 1); -qed "sin_arcsin"; -Addsimps [sin_arcsin]; - -Goal "[| -1 <= y; y <= 1 |] ==> sin(arcsin y) = y"; -by (auto_tac (claset() addIs [sin_arcsin],simpset())); -qed "sin_arcsin2"; -Addsimps [sin_arcsin2]; - -Goal "[| - 1 <= y; y <= 1 |] \ -\ ==> -(pi/2) <= arcsin y & arcsin y <= pi/2"; -by (blast_tac (claset() addDs [arcsin]) 1); -qed "arcsin_bounded"; - -Goal "[| - 1 <= y; y <= 1 |] ==> -(pi/2) <= arcsin y"; -by (blast_tac (claset() addDs [arcsin]) 1); -qed "arcsin_lbound"; - -Goal "[| - 1 <= y; y <= 1 |] ==> arcsin y <= pi/2"; -by (blast_tac (claset() addDs [arcsin]) 1); -qed "arcsin_ubound"; - -Goal "[| - 1 < y; y < 1 |] \ -\ ==> -(pi/2) < arcsin y & arcsin y < pi/2"; -by (ftac order_less_imp_le 1); -by (forw_inst_tac [("y","y")] order_less_imp_le 1); -by (ftac arcsin_bounded 1); -by (Step_tac 1 THEN Asm_full_simp_tac 1); -by (dres_inst_tac [("y","arcsin y")] order_le_imp_less_or_eq 1); -by (dres_inst_tac [("y","pi/2")] order_le_imp_less_or_eq 2); -by (Step_tac 1); -by (ALLGOALS(dres_inst_tac [("f","sin")] arg_cong)); -by (Auto_tac); -qed "arcsin_lt_bounded"; - -Goalw [arcsin_def] - "[|-(pi/2) <= x; x <= pi/2 |] ==> arcsin(sin x) = x"; -by (rtac some1_equality 1); -by (rtac sin_total 1); -by Auto_tac; -qed "arcsin_sin"; - -Goal "[| - 1 <= y; y <= 1 |] \ -\ ==> 0 <= arcos y & arcos y <= pi & cos(arcos y) = y"; -by (dtac cos_total 1 THEN assume_tac 1); -by (etac ex1E 1); -by (rewtac arcos_def); -by (rtac someI2 1); -by (ALLGOALS(Blast_tac)); -qed "arcos"; - -Goal "[| - 1 <= y; y <= 1 |] ==> cos(arcos y) = y"; -by (blast_tac (claset() addDs [arcos]) 1); -qed "cos_arcos"; -Addsimps [cos_arcos]; - -Goal "[| -1 <= y; y <= 1 |] ==> cos(arcos y) = y"; -by (auto_tac (claset() addIs [cos_arcos],simpset())); -qed "cos_arcos2"; -Addsimps [cos_arcos2]; - -Goal "[| - 1 <= y; y <= 1 |] ==> 0 <= arcos y & arcos y <= pi"; -by (blast_tac (claset() addDs [arcos]) 1); -qed "arcos_bounded"; - -Goal "[| - 1 <= y; y <= 1 |] ==> 0 <= arcos y"; -by (blast_tac (claset() addDs [arcos]) 1); -qed "arcos_lbound"; - -Goal "[| - 1 <= y; y <= 1 |] ==> arcos y <= pi"; -by (blast_tac (claset() addDs [arcos]) 1); -qed "arcos_ubound"; - -Goal "[| - 1 < y; y < 1 |] \ -\ ==> 0 < arcos y & arcos y < pi"; -by (ftac order_less_imp_le 1); -by (forw_inst_tac [("y","y")] order_less_imp_le 1); -by (ftac arcos_bounded 1); -by (Auto_tac); -by (dres_inst_tac [("y","arcos y")] order_le_imp_less_or_eq 1); -by (dres_inst_tac [("y","pi")] order_le_imp_less_or_eq 2); -by (Auto_tac); -by (ALLGOALS(dres_inst_tac [("f","cos")] arg_cong)); -by (Auto_tac); -qed "arcos_lt_bounded"; - -Goalw [arcos_def] "[|0 <= x; x <= pi |] ==> arcos(cos x) = x"; -by (auto_tac (claset() addSIs [some1_equality,cos_total],simpset())); -qed "arcos_cos"; - -Goalw [arcos_def] "[|x <= 0; -pi <= x |] ==> arcos(cos x) = -x"; -by (auto_tac (claset() addSIs [some1_equality,cos_total],simpset())); -qed "arcos_cos2"; - -Goal "- (pi/2) < arctan y & \ -\ arctan y < pi/2 & tan (arctan y) = y"; -by (cut_inst_tac [("y","y")] tan_total 1); -by (etac ex1E 1); -by (rewtac arctan_def); -by (rtac someI2 1); -by (ALLGOALS(Blast_tac)); -qed "arctan"; -Addsimps [arctan]; - -Goal "tan(arctan y) = y"; -by (Auto_tac); -qed "tan_arctan"; - -Goal "- (pi/2) < arctan y & arctan y < pi/2"; -by (asm_full_simp_tac (HOL_ss addsimps [arctan]) 1); -qed "arctan_bounded"; - -Goal "- (pi/2) < arctan y"; -by (Auto_tac); -qed "arctan_lbound"; - -Goal "arctan y < pi/2"; -by (asm_full_simp_tac (HOL_ss addsimps [arctan]) 1); -qed "arctan_ubound"; - -Goalw [arctan_def] - "[|-(pi/2) < x; x < pi/2 |] ==> arctan(tan x) = x"; -by (rtac some1_equality 1); -by (rtac tan_total 1); -by Auto_tac; -qed "arctan_tan"; - -Goal "arctan 0 = 0"; -by (rtac (asm_full_simplify (simpset()) - (read_instantiate [("x","0")] arctan_tan)) 1); -qed "arctan_zero_zero"; -Addsimps [arctan_zero_zero]; - -(* ------------------------------------------------------------------------- *) -(* Differentiation of arctan. *) -(* ------------------------------------------------------------------------- *) - -Goal "cos(arctan x) ~= 0"; -by (auto_tac (claset(),simpset() addsimps [cos_zero_iff])); -by (case_tac "n" 1); -by (case_tac "n" 3); -by (cut_inst_tac [("y","x")] arctan_ubound 2); -by (cut_inst_tac [("y","x")] arctan_lbound 4); -by (auto_tac (claset(), - simpset() addsimps [real_of_nat_Suc, left_distrib,linorder_not_less RS sym, mult_less_0_iff] - delsimps [arctan])); -qed "cos_arctan_not_zero"; -Addsimps [cos_arctan_not_zero]; - -Goal "cos x ~= 0 ==> 1 + tan(x) ^ 2 = inverse(cos x) ^ 2"; -by (rtac (power_inverse RS subst) 1); -by (res_inst_tac [("c1","cos(x) ^ 2")] (real_mult_right_cancel RS iffD1) 1); -by (auto_tac (claset() addDs [realpow_not_zero], simpset() addsimps - [power_mult_distrib,left_distrib,realpow_divide, - tan_def,real_mult_assoc,power_inverse RS sym] - delsimps [realpow_Suc])); -qed "tan_sec"; - - -(*--------------------------------------------------------------------------*) -(* Some more theorems- developed while at ICASE (07/2001) *) -(*--------------------------------------------------------------------------*) - -Goal "sin (xa + 1 / 2 * real (Suc m) * pi) = \ -\ cos (xa + 1 / 2 * real (m) * pi)"; -by (simp_tac (HOL_ss addsimps [cos_add,sin_add, - real_of_nat_Suc,left_distrib,right_distrib]) 1); -by Auto_tac; -qed "lemma_sin_cos_eq"; -Addsimps [lemma_sin_cos_eq]; - -Goal "sin (xa + real (Suc m) * pi / 2) = \ -\ cos (xa + real (m) * pi / 2)"; -by (simp_tac (HOL_ss addsimps [cos_add,sin_add,real_divide_def, - real_of_nat_Suc,left_distrib,right_distrib]) 1); -by Auto_tac; -qed "lemma_sin_cos_eq2"; -Addsimps [lemma_sin_cos_eq2]; - -Goal "DERIV (%x. sin (x + k)) xa :> cos (xa + k)"; -by (rtac lemma_DERIV_subst 1); -by (res_inst_tac [("f","sin"),("g","%x. x + k")] DERIV_chain2 1); -by DERIV_tac; -by (Simp_tac 1); -qed "DERIV_sin_add"; -Addsimps [DERIV_sin_add]; - -(* which further simplifies to (- 1 ^ m) !! *) -Goal "sin ((real m + 1/2) * pi) = cos (real m * pi)"; -by (auto_tac (claset(),simpset() addsimps [right_distrib, - sin_add,left_distrib] @ mult_ac)); -qed "sin_cos_npi"; -Addsimps [sin_cos_npi]; - -Goal "sin (real (Suc (2 * n)) * pi / 2) = (- 1) ^ n"; -by (cut_inst_tac [("m","n")] sin_cos_npi 1); -by (auto_tac (claset(),HOL_ss addsimps [real_of_nat_Suc, - left_distrib,real_divide_def])); -by Auto_tac; -qed "sin_cos_npi2"; -Addsimps [ sin_cos_npi2]; - -Goal "cos (2 * real (n::nat) * pi) = 1"; -by (auto_tac (claset(),simpset() addsimps [cos_double, - real_mult_assoc,power_add RS sym,numeral_2_eq_2])); - (*FIXME: just needs x^n for literals!*) -qed "cos_2npi"; -Addsimps [cos_2npi]; - -Goal "cos (3 / 2 * pi) = 0"; -by (rtac (CLAIM "(1::real) + 1/2 = 3/2" RS subst) 1); -by (stac left_distrib 1); -by (auto_tac (claset(),simpset() addsimps [cos_add] @ mult_ac)); -qed "cos_3over2_pi"; -Addsimps [cos_3over2_pi]; - -Goal "sin (2 * real (n::nat) * pi) = 0"; -by (auto_tac (claset(),simpset() addsimps [real_mult_assoc])); -qed "sin_2npi"; -Addsimps [sin_2npi]; - -Goal "sin (3 / 2 * pi) = - 1"; -by (rtac (CLAIM "(1::real) + 1/2 = 3/2" RS subst) 1); -by (stac left_distrib 1); -by (auto_tac (claset(),simpset() addsimps [sin_add] @mult_ac)); -qed "sin_3over2_pi"; -Addsimps [sin_3over2_pi]; - -Goal "cos(xa + 1 / 2 * real (Suc m) * pi) = \ -\ -sin (xa + 1 / 2 * real (m) * pi)"; -by (simp_tac (HOL_ss addsimps [cos_add,sin_add, - real_of_nat_Suc,right_distrib,left_distrib, - minus_mult_right]) 1); -by Auto_tac; -qed "lemma_cos_sin_eq"; -Addsimps [lemma_cos_sin_eq]; - -Goal "cos (xa + real (Suc m) * pi / 2) = \ -\ -sin (xa + real (m) * pi / 2)"; -by (simp_tac (HOL_ss addsimps [cos_add,sin_add,real_divide_def, - real_of_nat_Suc,left_distrib,right_distrib]) 1); -by Auto_tac; -qed "lemma_cos_sin_eq2"; -Addsimps [lemma_cos_sin_eq2]; - -Goal "cos (pi * real (Suc (2 * m)) / 2) = 0"; -by (simp_tac (HOL_ss addsimps [cos_add,sin_add,real_divide_def, - real_of_nat_Suc,left_distrib,right_distrib]) 1); -by Auto_tac; -qed "cos_pi_eq_zero"; -Addsimps [cos_pi_eq_zero]; - -Goal "DERIV (%x. cos (x + k)) xa :> - sin (xa + k)"; -by (rtac lemma_DERIV_subst 1); -by (res_inst_tac [("f","cos"),("g","%x. x + k")] DERIV_chain2 1); -by DERIV_tac; -by (Simp_tac 1); -qed "DERIV_cos_add"; -Addsimps [DERIV_cos_add]; - -Goal "isCont cos x"; -by (rtac (DERIV_cos RS DERIV_isCont) 1); -qed "isCont_cos"; - -Goal "isCont sin x"; -by (rtac (DERIV_sin RS DERIV_isCont) 1); -qed "isCont_sin"; - -Goal "isCont exp x"; -by (rtac (DERIV_exp RS DERIV_isCont) 1); -qed "isCont_exp"; - -val isCont_simp = [isCont_exp,isCont_sin,isCont_cos]; -Addsimps isCont_simp; - -(** more theorems: e.g. used in complex geometry **) - -Goal "sin x = 0 ==> abs(cos x) = 1"; -by (auto_tac (claset(),simpset() addsimps [sin_zero_iff,even_mult_two_ex])); -qed "sin_zero_abs_cos_one"; - -Goal "(exp x = 1) = (x = 0)"; -by Auto_tac; -by (dres_inst_tac [("f","ln")] arg_cong 1); -by Auto_tac; -qed "exp_eq_one_iff"; -Addsimps [exp_eq_one_iff]; - -Goal "cos x = 1 ==> sin x = 0"; -by (cut_inst_tac [("x","x")] sin_cos_squared_add3 1); -by Auto_tac; -qed "cos_one_sin_zero"; - -(*-------------------------------------------------------------------------------*) -(* A few extra theorems *) -(*-------------------------------------------------------------------------------*) - -Goal "[| 0 <= x; x < y |] ==> root(Suc n) x < root(Suc n) y"; -by (ftac order_le_less_trans 1); -by (assume_tac 1); -by (forw_inst_tac [("n1","n")] (real_root_pow_pos2 RS ssubst) 1); -by (rotate_tac 1 1); -by (assume_tac 1); -by (forw_inst_tac [("n1","n")] (real_root_pow_pos RS ssubst) 1); -by (rotate_tac 3 1 THEN assume_tac 1); -by (dres_inst_tac [("y","root (Suc n) y ^ Suc n")] order_less_imp_le 1 ); -by (forw_inst_tac [("n","n")] real_root_pos_pos_le 1); -by (forw_inst_tac [("n","n")] real_root_pos_pos 1); -by (dres_inst_tac [("x","root (Suc n) x"), - ("y","root (Suc n) y")] realpow_increasing 1); -by (assume_tac 1 THEN assume_tac 1); -by (dres_inst_tac [("x","root (Suc n) x")] order_le_imp_less_or_eq 1); -by Auto_tac; -by (dres_inst_tac [("f","%x. x ^ (Suc n)")] arg_cong 1); -by (auto_tac (claset(),simpset() addsimps [real_root_pow_pos2] - delsimps [realpow_Suc])); -qed "real_root_less_mono"; - -Goal "[| 0 <= x; x <= y |] ==> root(Suc n) x <= root(Suc n) y"; -by (dres_inst_tac [("y","y")] order_le_imp_less_or_eq 1 ); -by (auto_tac (claset() addDs [real_root_less_mono] - addIs [order_less_imp_le],simpset())); -qed "real_root_le_mono"; - -Goal "[| 0 <= x; 0 <= y |] ==> (root(Suc n) x < root(Suc n) y) = (x < y)"; -by (auto_tac (claset() addIs [real_root_less_mono],simpset())); -by (rtac ccontr 1 THEN dtac (linorder_not_less RS iffD1) 1); -by (dres_inst_tac [("x","y"),("n","n")] real_root_le_mono 1); -by Auto_tac; -qed "real_root_less_iff"; -Addsimps [real_root_less_iff]; - -Goal "[| 0 <= x; 0 <= y |] ==> (root(Suc n) x <= root(Suc n) y) = (x <= y)"; -by (auto_tac (claset() addIs [real_root_le_mono],simpset())); -by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); -by Auto_tac; -by (dres_inst_tac [("x","y"),("n","n")] real_root_less_mono 1); -by Auto_tac; -qed "real_root_le_iff"; -Addsimps [real_root_le_iff]; - -Goal "[| 0 <= x; 0 <= y |] ==> (root(Suc n) x = root(Suc n) y) = (x = y)"; -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; -qed "real_root_eq_iff"; -Addsimps [real_root_eq_iff]; - -Goal "[| 0 <= x; 0 <= y; y ^ (Suc n) = x |] ==> root (Suc n) x = y"; -by (auto_tac (claset() addDs [real_root_pos2], - simpset() delsimps [realpow_Suc])); -qed "real_root_pos_unique"; - -Goal "[| 0 <= x; 0 <= y |]\ -\ ==> root(Suc n) (x * y) = root(Suc n) x * root(Suc n) y"; -by (rtac real_root_pos_unique 1); -by (auto_tac (claset() addSIs [real_root_pos_pos_le],simpset() - addsimps [power_mult_distrib,zero_le_mult_iff, - real_root_pow_pos2] delsimps [realpow_Suc])); -qed "real_root_mult"; - -Goal "0 <= x ==> (root(Suc n) (inverse x) = inverse(root(Suc n) x))"; -by (rtac real_root_pos_unique 1); -by (auto_tac (claset() addIs [real_root_pos_pos_le],simpset() - addsimps [power_inverse RS sym,real_root_pow_pos2] - delsimps [realpow_Suc])); -qed "real_root_inverse"; - -Goalw [real_divide_def] - "[| 0 <= x; 0 <= y |] \ -\ ==> (root(Suc n) (x / y) = root(Suc n) x / root(Suc n) y)"; -by (auto_tac (claset(),simpset() addsimps [real_root_mult, - real_root_inverse])); -qed "real_root_divide"; - -Goalw [sqrt_def] "[| 0 <= x; x < y |] ==> sqrt(x) < sqrt(y)"; -by (auto_tac (claset() addIs [real_root_less_mono],simpset())); -qed "real_sqrt_less_mono"; - -Goalw [sqrt_def] "[| 0 <= x; x <= y |] ==> sqrt(x) <= sqrt(y)"; -by (auto_tac (claset() addIs [real_root_le_mono],simpset())); -qed "real_sqrt_le_mono"; - -Goalw [sqrt_def] "[| 0 <= x; 0 <= y |] ==> (sqrt(x) < sqrt(y)) = (x < y)"; -by Auto_tac; -qed "real_sqrt_less_iff"; -Addsimps [real_sqrt_less_iff]; - -Goalw [sqrt_def] "[| 0 <= x; 0 <= y |] ==> (sqrt(x) <= sqrt(y)) = (x <= y)"; -by Auto_tac; -qed "real_sqrt_le_iff"; -Addsimps [real_sqrt_le_iff]; - -Goalw [sqrt_def] "[| 0 <= x; 0 <= y |] ==> (sqrt(x) = sqrt(y)) = (x = y)"; -by Auto_tac; -qed "real_sqrt_eq_iff"; -Addsimps [real_sqrt_eq_iff]; - -Goal "(sqrt(x ^ 2 + y ^ 2) < 1) = (x ^ 2 + y ^ 2 < 1)"; -by (rtac (real_sqrt_one RS subst) 1); -by (Step_tac 1); -by (rtac real_sqrt_less_mono 2); -by (dtac (rotate_prems 2 (real_sqrt_less_iff RS iffD1)) 1); -by Auto_tac; -qed "real_sqrt_sos_less_one_iff"; -Addsimps [real_sqrt_sos_less_one_iff]; - -Goal "(sqrt(x ^ 2 + y ^ 2) = 1) = (x ^ 2 + y ^ 2 = 1)"; -by (rtac (real_sqrt_one RS subst) 1); -by (Step_tac 1); -by (dtac (rotate_prems 2 (real_sqrt_eq_iff RS iffD1)) 1); -by Auto_tac; -qed "real_sqrt_sos_eq_one_iff"; -Addsimps [real_sqrt_sos_eq_one_iff]; - -Goalw [real_divide_def] "(((r::real) * a) / (r * r)) = a / r"; -by (case_tac "r=0" 1); -by (auto_tac (claset(),simpset() addsimps [inverse_mult_distrib] @ mult_ac)); -qed "real_divide_square_eq"; -Addsimps [real_divide_square_eq]; - -(*-------------------------------------------------------------------------------*) -(* More theorems about sqrt, transcendental functions etc. needed in Complex.ML *) -(*-------------------------------------------------------------------------------*) - - -Goalw [real_divide_def] - "0 < x ==> 0 <= x/(sqrt (x * x + y * y))"; -by (ftac ((real_sqrt_sum_squares_ge1 RSN (2,order_less_le_trans)) - RS (CLAIM "0 < x ==> 0 < inverse (x::real)")) 1); -by (rtac (real_mult_order RS order_less_imp_le) 1); -by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2])); -qed "lemma_real_divide_sqrt"; - -Goal "0 < x ==> -(1::real) <= x/(sqrt (x * x + y * y))"; -by (rtac real_le_trans 1); -by (rtac lemma_real_divide_sqrt 2); -by Auto_tac; -qed "lemma_real_divide_sqrt_ge_minus_one"; - -Goal "x < 0 ==> 0 < sqrt (x * x + y * y)"; -by (rtac real_sqrt_gt_zero 1); -by (rtac (ARITH_PROVE "[| 0 < x; 0 <= y |] ==> (0::real) < x + y") 1); -by (auto_tac (claset(),simpset() addsimps [zero_less_mult_iff])); -qed "real_sqrt_sum_squares_gt_zero1"; - -Goal "0 < x ==> 0 < sqrt (x * x + y * y)"; -by (rtac real_sqrt_gt_zero 1); -by (rtac (ARITH_PROVE "[| 0 < x; 0 <= y |] ==> (0::real) < x + y") 1); -by (auto_tac (claset(),simpset() addsimps [zero_less_mult_iff])); -qed "real_sqrt_sum_squares_gt_zero2"; - -Goal "x ~= 0 ==> 0 < sqrt(x ^ 2 + y ^ 2)"; -by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1); -by (auto_tac (claset() addIs [real_sqrt_sum_squares_gt_zero2, - real_sqrt_sum_squares_gt_zero1],simpset() addsimps [numeral_2_eq_2])); -qed "real_sqrt_sum_squares_gt_zero3"; - -Goal "y ~= 0 ==> 0 < sqrt(x ^ 2 + y ^ 2)"; -by (dres_inst_tac [("y","x")] real_sqrt_sum_squares_gt_zero3 1); -by (auto_tac (claset(),simpset() addsimps [real_add_commute])); -qed "real_sqrt_sum_squares_gt_zero3a"; - -Goal "sqrt(x ^ 2 + y ^ 2) = x ==> y = 0"; -by (rtac ccontr 1); -by (forw_inst_tac [("x","x")] real_sum_squares_not_zero2 1); -by (dres_inst_tac [("f","%x. x ^ 2")] arg_cong 1); -by (forw_inst_tac [("x","x"),("y","y")] real_sum_square_gt_zero2 1); -by (dtac real_sqrt_gt_zero_pow2 1); -by Auto_tac; -qed "real_sqrt_sum_squares_eq_cancel"; -Addsimps [real_sqrt_sum_squares_eq_cancel]; - -Goal "sqrt(x ^ 2 + y ^ 2) = y ==> x = 0"; -by (res_inst_tac [("x","y")] real_sqrt_sum_squares_eq_cancel 1); -by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1); -qed "real_sqrt_sum_squares_eq_cancel2"; -Addsimps [real_sqrt_sum_squares_eq_cancel2]; - -Goal "x < 0 ==> x/(sqrt (x * x + y * y)) <= 1"; -by (dtac (ARITH_PROVE "x < 0 ==> (0::real) < -x") 1); -by (dres_inst_tac [("y","y")] - lemma_real_divide_sqrt_ge_minus_one 1); -by (dtac (ARITH_PROVE "x <= y ==> -y <= -(x::real)") 1); -by Auto_tac; -qed "lemma_real_divide_sqrt_le_one"; - -Goal "x < 0 ==> -(1::real) <= x/(sqrt (x * x + y * y))"; -by (case_tac "y = 0" 1); -by Auto_tac; -by (ftac abs_minus_eqI2 1); -by (auto_tac (claset(),simpset() addsimps [inverse_minus_eq])); -by (rtac order_less_imp_le 1); -by (res_inst_tac [("z1","sqrt(x * x + y * y)")] - (real_mult_less_iff1 RS iffD1) 1); -by (forw_inst_tac [("y2","y")] - (real_sqrt_sum_squares_gt_zero1 RS real_not_refl2 - RS not_sym) 2); -by (auto_tac (claset() addIs [real_sqrt_sum_squares_gt_zero1], - simpset() addsimps mult_ac)); -by (rtac (ARITH_PROVE "-x < y ==> -y < (x::real)") 1); -by (cut_inst_tac [("x","-x"),("y","y")] real_sqrt_sum_squares_ge1 1); -by (dtac real_le_imp_less_or_eq 1); -by (Step_tac 1); -by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2]) 1); -by (dtac (sym RS real_sqrt_sum_squares_eq_cancel) 1); -by Auto_tac; -qed "lemma_real_divide_sqrt_ge_minus_one2"; - -Goal "0 < x ==> x/(sqrt (x * x + y * y)) <= 1"; -by (dtac (ARITH_PROVE "0 < x ==> -x < (0::real)") 1); -by (dres_inst_tac [("y","y")] - lemma_real_divide_sqrt_ge_minus_one2 1); -by (dtac (ARITH_PROVE "x <= y ==> -y <= -(x::real)") 1); -by Auto_tac; -qed "lemma_real_divide_sqrt_le_one2"; -(* was qed "lemma_real_mult_self_rinv_sqrt_squared5" *) - -Goal "-(1::real)<= x / sqrt (x * x + y * y)"; -by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1); -by (Step_tac 1); -by (rtac lemma_real_divide_sqrt_ge_minus_one2 1); -by (rtac lemma_real_divide_sqrt_ge_minus_one 3); -by Auto_tac; -qed "cos_x_y_ge_minus_one"; -Addsimps [cos_x_y_ge_minus_one]; - -Goal "-(1::real)<= y / sqrt (x * x + y * y)"; -by (cut_inst_tac [("x","y"),("y","x")] cos_x_y_ge_minus_one 1); -by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1); -qed "cos_x_y_ge_minus_one1a"; -Addsimps [cos_x_y_ge_minus_one1a, - simplify (simpset()) cos_x_y_ge_minus_one1a]; - -Goal "x / sqrt (x * x + y * y) <= 1"; -by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1); -by (Step_tac 1); -by (rtac lemma_real_divide_sqrt_le_one 1); -by (rtac lemma_real_divide_sqrt_le_one2 3); -by Auto_tac; -qed "cos_x_y_le_one"; -Addsimps [cos_x_y_le_one]; - -Goal "y / sqrt (x * x + y * y) <= 1"; -by (cut_inst_tac [("x","y"),("y","x")] cos_x_y_le_one 1); -by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1); -qed "cos_x_y_le_one2"; -Addsimps [cos_x_y_le_one2]; - -Addsimps [[cos_x_y_ge_minus_one,cos_x_y_le_one] MRS cos_arcos]; -Addsimps [[cos_x_y_ge_minus_one,cos_x_y_le_one] MRS arcos_bounded]; - -Addsimps [[cos_x_y_ge_minus_one1a,cos_x_y_le_one2] MRS cos_arcos]; -Addsimps [[cos_x_y_ge_minus_one1a,cos_x_y_le_one2] MRS arcos_bounded]; - -Goal "-(1::real) <= abs(x) / sqrt (x * x + y * y)"; -by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1); -by (auto_tac (claset(),simpset() addsimps [abs_minus_eqI2,abs_eqI2])); -by (dtac lemma_real_divide_sqrt_ge_minus_one 1 THEN Force_tac 1); -qed "cos_rabs_x_y_ge_minus_one"; - -Addsimps [cos_rabs_x_y_ge_minus_one, - simplify (simpset()) cos_rabs_x_y_ge_minus_one]; - -Goal "abs(x) / sqrt (x * x + y * y) <= 1"; -by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1); -by (auto_tac (claset(),simpset() addsimps [abs_minus_eqI2,abs_eqI2])); -by (dtac lemma_real_divide_sqrt_ge_minus_one2 1 THEN Force_tac 1); -qed "cos_rabs_x_y_le_one"; -Addsimps [cos_rabs_x_y_le_one]; - -Addsimps [[cos_rabs_x_y_ge_minus_one,cos_rabs_x_y_le_one] MRS cos_arcos]; -Addsimps [[cos_rabs_x_y_ge_minus_one,cos_rabs_x_y_le_one] MRS arcos_bounded]; - -Goal "-pi < 0"; -by (Simp_tac 1); -qed "minus_pi_less_zero"; -Addsimps [minus_pi_less_zero]; -Addsimps [minus_pi_less_zero RS order_less_imp_le]; - -Goal "[| -(1::real) <= y; y <= 1 |] ==> -pi <= arcos y"; -by (rtac real_le_trans 1); -by (rtac arcos_lbound 2); -by Auto_tac; -qed "arcos_ge_minus_pi"; - -Addsimps [[cos_x_y_ge_minus_one,cos_x_y_le_one] MRS arcos_ge_minus_pi]; - -(* How tedious! *) -Goal "[| x + (y::real) ~= 0; 1 - z = x/(x + y) \ -\ |] ==> z = y/(x + y)"; -by (res_inst_tac [("c1","x + y")] (real_mult_right_cancel RS iffD1) 1); -by (forw_inst_tac [("c1","x + y")] (real_mult_right_cancel RS iffD2) 2); -by (assume_tac 2); -by (rotate_tac 2 2); -by (dtac (real_mult_assoc RS subst) 2); -by (rotate_tac 2 2); -by (ftac (left_inverse RS subst) 2); -by (assume_tac 2); -by (thin_tac "(1 - z) * (x + y) = x /(x + y) * (x + y)" 2); -by (thin_tac "1 - z = x /(x + y)" 2); -by (auto_tac (claset(),simpset() addsimps [mult_assoc])); -by (auto_tac (claset(),simpset() addsimps [right_distrib, - left_diff_distrib])); -qed "lemma_divide_rearrange"; - -Goal "[| 0 < x * x + y * y; \ -\ 1 - sin xa ^ 2 = (x / sqrt (x * x + y * y)) ^ 2 \ -\ |] ==> sin xa ^ 2 = (y / sqrt (x * x + y * y)) ^ 2"; -by (auto_tac (claset() addIs [lemma_divide_rearrange],simpset() - addsimps [realpow_divide,real_sqrt_gt_zero_pow2, - power2_eq_square RS sym])); -qed "lemma_cos_sin_eq"; - -Goal "[| 0 < x * x + y * y; \ -\ 1 - cos xa ^ 2 = (y / sqrt (x * x + y * y)) ^ 2 \ -\ |] ==> cos xa ^ 2 = (x / sqrt (x * x + y * y)) ^ 2"; -by (auto_tac (claset(), - simpset() addsimps [realpow_divide, - real_sqrt_gt_zero_pow2,power2_eq_square RS sym])); -by (rtac (real_add_commute RS subst) 1); -by (rtac lemma_divide_rearrange 1); -by (asm_full_simp_tac (simpset() addsimps []) 1); -by (asm_full_simp_tac (simpset() addsimps [add_commute]) 1); -qed "lemma_sin_cos_eq"; - -Goal "[| x ~= 0; \ -\ cos xa = x / sqrt (x * x + y * y) \ -\ |] ==> sin xa = y / sqrt (x * x + y * y) | \ -\ sin xa = - y / sqrt (x * x + y * y)"; -by (dres_inst_tac [("f","%x. x ^ 2")] arg_cong 1); -by (forw_inst_tac [("y","y")] real_sum_square_gt_zero 1); -by (asm_full_simp_tac (simpset() addsimps [cos_squared_eq]) 1); -by (subgoal_tac "sin xa ^ 2 = (y / sqrt (x * x + y * y)) ^ 2" 1); -by (rtac lemma_cos_sin_eq 2); -by (Force_tac 2); -by (Asm_full_simp_tac 2); -by (auto_tac (claset(),simpset() addsimps [realpow_two_disj,numeral_2_eq_2] - delsimps [realpow_Suc])); -qed "sin_x_y_disj"; - -(*FIXME: remove real_sqrt_gt_zero_pow2*) -Goal "0 <= x ==> sqrt(x) ^ 2 = x"; -by (asm_full_simp_tac (simpset() addsimps [real_sqrt_pow_abs,abs_if]) 1); -qed "real_sqrt_ge_zero_pow2"; - -Goal "y ~= 0 ==> x / sqrt (x * x + y * y) ~= -(1::real)"; -by Auto_tac; -by (dres_inst_tac [("f","%x. x ^ 2")] arg_cong 1); -by (asm_full_simp_tac (simpset() addsimps [power_divide,thm"real_mult_self_sum_ge_zero",real_sqrt_ge_zero_pow2]) 1); -by (asm_full_simp_tac (simpset() addsimps [inst "a" "1" divide_eq_eq, power2_eq_square] addsplits [split_if_asm]) 1); -qed "cos_not_eq_minus_one"; - -Goalw [arcos_def] - "arcos (x / sqrt (x * x + y * y)) = pi ==> y = 0"; -by (rtac ccontr 1); -by (rtac swap 1 THEN assume_tac 2); -by (rtac (([cos_x_y_ge_minus_one,cos_x_y_le_one] MRS cos_total) RS - ((CLAIM "EX! x. P x ==> EX x. P x") RS someI2_ex)) 1); -by (auto_tac (claset() addDs [cos_not_eq_minus_one],simpset())); -qed "arcos_eq_pi_cancel"; - -Goalw [real_divide_def] "x ~= 0 ==> x / sqrt (x * x + y * y) ~= 0"; -by (forw_inst_tac [("y3","y")] (real_sqrt_sum_squares_gt_zero3 - RS real_not_refl2 RS not_sym RS nonzero_imp_inverse_nonzero) 1); -by (auto_tac (claset(),simpset() addsimps [power2_eq_square])); -qed "lemma_cos_not_eq_zero"; - -Goal "[| x ~= 0; \ -\ sin xa = y / sqrt (x * x + y * y) \ -\ |] ==> cos xa = x / sqrt (x * x + y * y) | \ -\ cos xa = - x / sqrt (x * x + y * y)"; -by (dres_inst_tac [("f","%x. x ^ 2")] arg_cong 1); -by (forw_inst_tac [("y","y")] real_sum_square_gt_zero 1); -by (asm_full_simp_tac (simpset() addsimps [sin_squared_eq] - delsimps [realpow_Suc]) 1); -by (subgoal_tac "cos xa ^ 2 = (x / sqrt (x * x + y * y)) ^ 2" 1); -by (rtac lemma_sin_cos_eq 2); -by (Force_tac 2); -by (Asm_full_simp_tac 2); -by (auto_tac (claset(),simpset() addsimps [realpow_two_disj, - numeral_2_eq_2] delsimps [realpow_Suc])); -qed "cos_x_y_disj"; - -Goal "0 < y ==> - y / sqrt (x * x + y * y) < 0"; -by (case_tac "x = 0" 1); -by (auto_tac (claset(),simpset() addsimps [abs_eqI2])); -by (dres_inst_tac [("y","y")] real_sqrt_sum_squares_gt_zero3 1); -by (auto_tac (claset(),simpset() addsimps [zero_less_mult_iff, - real_divide_def,power2_eq_square])); -qed "real_sqrt_divide_less_zero"; - -Goal "[| x ~= 0; 0 < y |] ==> EX r a. x = r * cos a & y = r * sin a"; -by (res_inst_tac [("x","sqrt(x ^ 2 + y ^ 2)")] exI 1); -by (res_inst_tac [("x","arcos(x / sqrt (x * x + y * y))")] exI 1); -by Auto_tac; -by (dres_inst_tac [("y2","y")] (real_sqrt_sum_squares_gt_zero3 - RS real_not_refl2 RS not_sym) 1); -by (auto_tac (claset(),simpset() addsimps [power2_eq_square])); -by (rewtac arcos_def); -by (cut_inst_tac [("x1","x"),("y1","y")] ([cos_x_y_ge_minus_one,cos_x_y_le_one] - MRS cos_total) 1); -by (rtac someI2_ex 1 THEN Blast_tac 1); -by (thin_tac - "EX! xa. 0 <= xa & xa <= pi & cos xa = x / sqrt (x * x + y * y)" 1); -by (ftac sin_x_y_disj 1 THEN Blast_tac 1); -by (dres_inst_tac [("y2","y")] (real_sqrt_sum_squares_gt_zero3 - RS real_not_refl2 RS not_sym) 1); -by (auto_tac (claset(),simpset() addsimps [power2_eq_square])); -by (dtac sin_ge_zero 1 THEN assume_tac 1); -by (dres_inst_tac [("x","x")] real_sqrt_divide_less_zero 1 THEN Auto_tac); -qed "polar_ex1"; - -Goal "x * x = -(y * y) ==> y = (0::real)"; -by (auto_tac (claset() addIs [real_sum_squares_cancel],simpset())); -qed "real_sum_squares_cancel2a"; - -Goal "[| x ~= 0; y < 0 |] ==> EX r a. x = r * cos a & y = r * sin a"; -by (cut_inst_tac [("x","0"),("y","x")] linorder_less_linear 1); -by Auto_tac; -by (res_inst_tac [("x","sqrt(x ^ 2 + y ^ 2)")] exI 1); -by (res_inst_tac [("x","arcsin(y / sqrt (x * x + y * y))")] exI 1); -by (auto_tac (claset() addDs [real_sum_squares_cancel2a], - simpset() addsimps [power2_eq_square])); -by (rewtac arcsin_def); -by (cut_inst_tac [("x1","x"),("y1","y")] ([cos_x_y_ge_minus_one1a, - cos_x_y_le_one2] MRS sin_total) 1); -by (rtac someI2_ex 1 THEN Blast_tac 1); -by (thin_tac "EX! xa. - (pi/2) <= xa & \ -\ xa <= pi/2 & sin xa = y / sqrt (x * x + y * y)" 1); -by (ftac ((CLAIM "0 < x ==> (x::real) ~= 0") RS cos_x_y_disj) 1 THEN Blast_tac 1); -by Auto_tac; -by (dtac cos_ge_zero 1 THEN Force_tac 1); -by (dres_inst_tac [("x","y")] real_sqrt_divide_less_zero 1); -by (auto_tac (claset(),simpset() addsimps [real_add_commute])); -by (dtac (ARITH_PROVE "(y::real) < 0 ==> 0 < - y") 1); -by (dtac (CLAIM "x < (0::real) ==> x ~= 0" RS polar_ex1) 1 THEN assume_tac 1); -by (REPEAT(etac exE 1)); -by (res_inst_tac [("x","r")] exI 1); -by (res_inst_tac [("x","-a")] exI 1); -by Auto_tac; -qed "polar_ex2"; - -Goal "EX r a. x = r * cos a & y = r * sin a"; -by (case_tac "x = 0" 1); -by Auto_tac; -by (res_inst_tac [("x","y")] exI 1); -by (res_inst_tac [("x","pi/2")] exI 1 THEN Auto_tac); -by (cut_inst_tac [("x","0"),("y","y")] linorder_less_linear 1); -by Auto_tac; -by (res_inst_tac [("x","x")] exI 2); -by (res_inst_tac [("x","0")] exI 2 THEN Auto_tac); -by (ALLGOALS(blast_tac (claset() addIs [polar_ex1,polar_ex2]))); -qed "polar_Ex"; - -Goal "abs x <= sqrt (x ^ 2 + y ^ 2)"; -by (res_inst_tac [("n","1")] realpow_increasing 1); -by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2 RS sym])); -qed "real_sqrt_ge_abs1"; - -Goal "abs y <= sqrt (x ^ 2 + y ^ 2)"; -by (rtac (real_add_commute RS subst) 1); -by (rtac real_sqrt_ge_abs1 1); -qed "real_sqrt_ge_abs2"; -Addsimps [real_sqrt_ge_abs1,real_sqrt_ge_abs2]; - -Goal "0 < sqrt 2"; -by (auto_tac (claset() addIs [real_sqrt_gt_zero],simpset())); -qed "real_sqrt_two_gt_zero"; -Addsimps [real_sqrt_two_gt_zero]; - -Goal "0 <= sqrt 2"; -by (auto_tac (claset() addIs [real_sqrt_ge_zero],simpset())); -qed "real_sqrt_two_ge_zero"; -Addsimps [real_sqrt_two_ge_zero]; - -Goal "1 < sqrt 2"; -by (res_inst_tac [("y","7/5")] order_less_le_trans 1); -by (res_inst_tac [("n","1")] realpow_increasing 2); -by (auto_tac (claset(),simpset() addsimps [real_sqrt_gt_zero_pow2,numeral_2_eq_2 RS sym] - delsimps [realpow_Suc])); -by (simp_tac (simpset() addsimps [numeral_2_eq_2]) 1); -qed "real_sqrt_two_gt_one"; -Addsimps [real_sqrt_two_gt_one]; - -Goal "0 < u ==> u / sqrt 2 < u"; -by (res_inst_tac [("z1","inverse u")] (real_mult_less_iff1 RS iffD1) 1); -by Auto_tac; -by (res_inst_tac [("z1","sqrt 2")] (real_mult_less_iff1 RS iffD1) 1); -by Auto_tac; -qed "lemma_real_divide_sqrt_less"; - -(* needed for infinitely close relation over the nonstandard complex numbers *) -Goal "[| 0 < u; x < u/2; y < u/2; 0 <= x; 0 <= y |] ==> sqrt (x ^ 2 + y ^ 2) < u"; -by (res_inst_tac [("y","u/sqrt 2")] order_le_less_trans 1); -by (etac lemma_real_divide_sqrt_less 2); -by (res_inst_tac [("n","1")] realpow_increasing 1); -by (auto_tac (claset(), - simpset() addsimps [real_0_le_divide_iff,realpow_divide, - real_sqrt_gt_zero_pow2,numeral_2_eq_2 RS sym] delsimps [realpow_Suc])); -by (res_inst_tac [("t","u ^ 2")] (real_sum_of_halves RS subst) 1); -by (rtac add_mono 1); -by (auto_tac (claset(),simpset() delsimps [realpow_Suc])); -by (ALLGOALS(rtac ((CLAIM "(2::real) ^ 2 = 4") RS subst))); -by (ALLGOALS(rtac (power_mult_distrib RS subst))); -by (ALLGOALS(rtac power_mono)); -by Auto_tac; -qed "lemma_sqrt_hcomplex_capprox"; - -Addsimps [real_sqrt_sum_squares_ge_zero RS abs_eqI1]; - -(* A few theorems involving ln and derivatives, etc *) - -Goal "DERIV ln z :> l ==> DERIV (%x. exp (ln x)) z :> exp (ln z) * l"; -by (etac DERIV_fun_exp 1); -qed "lemma_DERIV_ln"; - -Goal "0 < z ==> ( *f* (%x. exp (ln x))) z = z"; -by (res_inst_tac [("z","z")] eq_Abs_hypreal 1); -by (auto_tac (claset(),simpset() addsimps [starfun, - hypreal_zero_def,hypreal_less])); -qed "STAR_exp_ln"; - -Goal "[|e : Infinitesimal; 0 < x |] ==> 0 < hypreal_of_real x + e"; -by (res_inst_tac [("c1","-e")] (add_less_cancel_right RS iffD1) 1); -by (auto_tac (claset() addIs [Infinitesimal_less_SReal],simpset())); -qed "hypreal_add_Infinitesimal_gt_zero"; - -Goalw [nsderiv_def,NSLIM_def] "0 < z ==> NSDERIV (%x. exp (ln x)) z :> 1"; -by Auto_tac; -by (rtac ccontr 1); -by (subgoal_tac "0 < hypreal_of_real z + h" 1); -by (dtac STAR_exp_ln 1); -by (rtac hypreal_add_Infinitesimal_gt_zero 2); -by (dtac (CLAIM "h ~= 0 ==> h/h = (1::hypreal)") 1); -by (auto_tac (claset(),simpset() addsimps [exp_ln_iff RS sym] - delsimps [exp_ln_iff])); -qed "NSDERIV_exp_ln_one"; - -Goal "0 < z ==> DERIV (%x. exp (ln x)) z :> 1"; -by (auto_tac (claset() addIs [NSDERIV_exp_ln_one], - simpset() addsimps [NSDERIV_DERIV_iff RS sym])); -qed "DERIV_exp_ln_one"; - -Goal "[| 0 < z; DERIV ln z :> l |] ==> exp (ln z) * l = 1"; -by (rtac DERIV_unique 1); -by (rtac lemma_DERIV_ln 1); -by (rtac DERIV_exp_ln_one 2); -by Auto_tac; -qed "lemma_DERIV_ln2"; - -Goal "[| 0 < z; DERIV ln z :> l |] ==> l = 1/(exp (ln z))"; -by (res_inst_tac [("c1","exp(ln z)")] (real_mult_left_cancel RS iffD1) 1); -by (auto_tac (claset() addIs [lemma_DERIV_ln2],simpset())); -qed "lemma_DERIV_ln3"; - -Goal "[| 0 < z; DERIV ln z :> l |] ==> l = 1/z"; -by (res_inst_tac [("t","z")] (exp_ln_iff RS iffD2 RS subst) 1); -by (auto_tac (claset() addIs [lemma_DERIV_ln3],simpset())); -qed "lemma_DERIV_ln4"; - -(* need to rename second isCont_inverse *) - -Goal "[| 0 < d; ALL z. abs(z - x) <= d --> g(f(z)) = z; \ -\ ALL z. abs(z - x) <= d --> isCont f z |] \ -\ ==> isCont g (f x)"; -by (simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1); -by (Step_tac 1); -by (dres_inst_tac [("d1.0","r")] real_lbound_gt_zero 1); -by (assume_tac 1 THEN Step_tac 1); -by (subgoal_tac "ALL z. abs(z - x) <= e --> (g(f z) = z)" 1); -by (Force_tac 2); -by (subgoal_tac "ALL 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 (Step_tac 1); -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_inv_fun"; - -(* -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 "[| 0 < d; \ -\ ALL z. abs(z - x) <= d --> g(f(z)) = z; \ -\ ALL z. abs(z - x) <= d --> isCont f z |] \ -\ ==> EX e. 0 < (e::real) & \ -\ (ALL y. 0 < abs(y - f(x::real)) & abs(y - f(x)) < e --> f(g(y)) = y)"; -by (dtac isCont_inj_range 1); -by (assume_tac 2 THEN assume_tac 1); -by Auto_tac; -by (res_inst_tac [("x","e")] exI 1 THEN Auto_tac); -by (rotate_tac 2 1); -by (dres_inst_tac [("x","y")] spec 1 THEN Auto_tac); -qed "isCont_inv_fun_inv"; - - -(* Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*) -Goal "[| f -- c --> l; 0 < l |] \ -\ ==> EX r. 0 < r & (ALL x. x ~= c & abs (c - x) < r --> 0 < f x)"; -by (auto_tac (claset(),simpset() addsimps [LIM_def])); -by (dres_inst_tac [("x","l/2")] spec 1); -by (Step_tac 1); -by (Force_tac 1); -by (res_inst_tac [("x","s")] exI 1); -by (Step_tac 1); -by (rotate_tac 2 1); -by (dres_inst_tac [("x","x")] spec 1); -by (auto_tac (claset(),HOL_ss addsimps [abs_interval_iff])); -by (auto_tac (claset(),simpset() addsimps [CLAIM "(l::real) + -(l/2) = l/2", - CLAIM "(a < f + - l) = (l + a < (f::real))"])); -qed "LIM_fun_gt_zero"; - -Goal "[| f -- c --> l; l < 0 |] \ -\ ==> EX r. 0 < r & (ALL x. x ~= c & abs (c - x) < r --> f x < 0)"; -by (auto_tac (claset(),simpset() addsimps [LIM_def])); -by (dres_inst_tac [("x","-l/2")] spec 1); -by (Step_tac 1); -by (Force_tac 1); -by (res_inst_tac [("x","s")] exI 1); -by (Step_tac 1); -by (rotate_tac 2 1); -by (dres_inst_tac [("x","x")] spec 1); -by (auto_tac (claset(),HOL_ss addsimps [abs_interval_iff])); -by (auto_tac (claset(),simpset() addsimps [CLAIM "(l::real) + -(l/2) = l/2", - CLAIM "(f + - l < a) = ((f::real) < l + a)"])); -qed "LIM_fun_less_zero"; - - -Goal "[| f -- c --> l; l ~= 0 |] \ -\ ==> EX r. 0 < r & (ALL x. x ~= c & abs (c - x) < r --> f x ~= 0)"; -by (cut_inst_tac [("x","l"),("y","0")] linorder_less_linear 1); -by Auto_tac; -by (dtac LIM_fun_less_zero 1); -by (dtac LIM_fun_gt_zero 3); -by Auto_tac; -by (ALLGOALS(res_inst_tac [("x","r")] exI)); -by Auto_tac; -qed "LIM_fun_not_zero"; diff -r 4b3d280ef06a -r 89840837108e src/HOL/Hyperreal/Transcendental.thy --- a/src/HOL/Hyperreal/Transcendental.thy Mon Jul 26 15:48:50 2004 +0200 +++ b/src/HOL/Hyperreal/Transcendental.thy Mon Jul 26 17:34:52 2004 +0200 @@ -2,10 +2,14 @@ Author : Jacques D. Fleuriot Copyright : 1998,1999 University of Cambridge 1999,2001 University of Edinburgh - Description : Power Series, transcendental functions etc. + Conversion to Isar and new proofs by Lawrence C Paulson, 2004 *) -Transcendental = NthRoot + Fact + HSeries + EvenOdd + Lim + +header{*Power Series, Transcendental Functions etc.*} + +theory Transcendental = NthRoot + Fact + HSeries + EvenOdd + Lim: + +(*????FOR RING_AND_FIELD*) constdefs root :: "[nat,real] => real" @@ -32,18 +36,2854 @@ "ln x == (@u. exp u = x)" pi :: "real" - "pi == 2 * (@x. 0 <= (x::real) & x <= 2 & cos x = 0)" + "pi == 2 * (@x. 0 \ (x::real) & x \ 2 & cos x = 0)" tan :: "real => real" "tan x == (sin x)/(cos x)" arcsin :: "real => real" - "arcsin y == (@x. -(pi/2) <= x & x <= pi/2 & sin x = y)" + "arcsin y == (@x. -(pi/2) \ x & x \ pi/2 & sin x = y)" arcos :: "real => real" - "arcos y == (@x. 0 <= x & x <= pi & cos x = y)" + "arcos y == (@x. 0 \ x & x \ pi & cos x = y)" arctan :: "real => real" "arctan y == (@x. -(pi/2) < x & x < pi/2 & tan x = y)" + + +lemma real_root_zero [simp]: "root (Suc n) 0 = 0" +apply (unfold root_def) +apply (safe intro!: some_equality power_0_Suc elim!: realpow_zero_zero) +done + +lemma real_root_pow_pos: + "0 < x ==> (root(Suc n) x) ^ (Suc n) = x" +apply (unfold root_def) +apply (drule_tac n = n in realpow_pos_nth2) +apply (auto intro: someI2) +done + +lemma real_root_pow_pos2: "0 \ x ==> (root(Suc n) x) ^ (Suc n) = x" +by (auto dest!: real_le_imp_less_or_eq dest: real_root_pow_pos) + +lemma real_root_pos: + "0 < x ==> root(Suc n) (x ^ (Suc n)) = x" +apply (unfold root_def) +apply (rule some_equality) +apply (frule_tac [2] n = n in zero_less_power) +apply (auto simp add: zero_less_mult_iff) +apply (rule_tac x = u and y = x in linorder_cases) +apply (drule_tac n1 = n and x = u in zero_less_Suc [THEN [3] realpow_less]) +apply (drule_tac [4] n1 = n and x = x in zero_less_Suc [THEN [3] realpow_less]) +apply (auto simp add: order_less_irrefl) +done + +lemma real_root_pos2: "0 \ x ==> root(Suc n) (x ^ (Suc n)) = x" +by (auto dest!: real_le_imp_less_or_eq real_root_pos) + +lemma real_root_pos_pos: + "0 < x ==> 0 \ root(Suc n) x" +apply (unfold root_def) +apply (drule_tac n = n in realpow_pos_nth2) +apply (safe, rule someI2) +apply (auto intro!: order_less_imp_le dest: zero_less_power simp add: zero_less_mult_iff) +done + +lemma real_root_pos_pos_le: "0 \ x ==> 0 \ root(Suc n) x" +by (auto dest!: real_le_imp_less_or_eq dest: real_root_pos_pos) + +lemma real_root_one [simp]: "root (Suc n) 1 = 1" +apply (unfold root_def) +apply (rule some_equality, auto) +apply (rule ccontr) +apply (rule_tac x = u and y = 1 in linorder_cases) +apply (drule_tac n = n in realpow_Suc_less_one) +apply (drule_tac [4] n = n in power_gt1_lemma) +apply (auto simp add: order_less_irrefl) +done + + +subsection{*Square Root*} + +(*lcp: needed now because 2 is a binary numeral!*) +lemma root_2_eq [simp]: "root 2 = root (Suc (Suc 0))" +apply (simp (no_asm) del: nat_numeral_0_eq_0 nat_numeral_1_eq_1 add: nat_numeral_0_eq_0 [symmetric]) +done + +lemma real_sqrt_zero [simp]: "sqrt 0 = 0" +by (unfold sqrt_def, auto) + +lemma real_sqrt_one [simp]: "sqrt 1 = 1" +by (unfold sqrt_def, auto) + +lemma real_sqrt_pow2_iff [simp]: "((sqrt x)\ = x) = (0 \ x)" +apply (unfold sqrt_def) +apply (rule iffI) + apply (cut_tac r = "root 2 x" in realpow_two_le) + apply (simp add: numeral_2_eq_2) +apply (subst numeral_2_eq_2) +apply (erule real_root_pow_pos2) +done + +lemma [simp]: "(sqrt(u2\ + v2\))\ = u2\ + v2\" +by (rule realpow_two_le_add_order [THEN real_sqrt_pow2_iff [THEN iffD2]]) + +lemma real_sqrt_pow2 [simp]: "0 \ x ==> (sqrt x)\ = x" +by (simp add: real_sqrt_pow2_iff) + +lemma real_sqrt_abs_abs [simp]: "sqrt\x\ ^ 2 = \x\" +by (rule real_sqrt_pow2_iff [THEN iffD2], arith) + +lemma real_pow_sqrt_eq_sqrt_pow: + "0 \ x ==> (sqrt x)\ = sqrt(x\)" +apply (unfold sqrt_def) +apply (subst numeral_2_eq_2) +apply (auto intro: real_root_pow_pos2 [THEN ssubst] real_root_pos2 [THEN ssubst] simp del: realpow_Suc) +done + +lemma real_pow_sqrt_eq_sqrt_abs_pow2: + "0 \ x ==> (sqrt x)\ = sqrt(\x\ ^ 2)" +by (simp add: real_pow_sqrt_eq_sqrt_pow [symmetric]) + +lemma real_sqrt_pow_abs: "0 \ x ==> (sqrt x)\ = \x\" +apply (rule real_sqrt_abs_abs [THEN subst]) +apply (rule_tac x1 = x in real_pow_sqrt_eq_sqrt_abs_pow2 [THEN ssubst]) +apply (rule_tac [2] real_pow_sqrt_eq_sqrt_pow [symmetric]) +apply (assumption, arith) +done + +lemma not_real_square_gt_zero [simp]: "(~ (0::real) < x*x) = (x = 0)" +apply auto +apply (cut_tac x = x and y = 0 in linorder_less_linear) +apply (simp add: zero_less_mult_iff) +done + +lemma real_mult_self_eq_zero_iff [simp]: "(r * r = 0) = (r = (0::real))" +by auto + +lemma real_sqrt_gt_zero: "0 < x ==> 0 < sqrt(x)" +apply (unfold sqrt_def root_def) +apply (subst numeral_2_eq_2) +apply (drule realpow_pos_nth2 [where n=1], safe) +apply (rule someI2, auto) +done + +lemma real_sqrt_ge_zero: "0 \ x ==> 0 \ sqrt(x)" +by (auto intro: real_sqrt_gt_zero simp add: order_le_less) + + +(*we need to prove something like this: +lemma "[|r ^ n = a; 0 0 < r|] ==> root n a = r" +apply (case_tac n, simp) +apply (unfold root_def) +apply (rule someI2 [of _ r], safe) +apply (auto simp del: realpow_Suc dest: power_inject_base) +*) + +lemma sqrt_eqI: "[|r\ = a; 0 \ r|] ==> sqrt a = r" +apply (unfold sqrt_def root_def) +apply (rule someI2 [of _ r], auto) +apply (auto simp add: numeral_2_eq_2 simp del: realpow_Suc + dest: power_inject_base) +done + +lemma real_sqrt_mult_distrib: + "[| 0 \ x; 0 \ y |] ==> sqrt(x*y) = sqrt(x) * sqrt(y)" +apply (rule sqrt_eqI) +apply (simp add: power_mult_distrib) +apply (simp add: zero_le_mult_iff real_sqrt_ge_zero) +done + +lemma real_sqrt_mult_distrib2: "[|0\x; 0\y |] ==> sqrt(x*y) = sqrt(x) * sqrt(y)" +by (auto intro: real_sqrt_mult_distrib simp add: order_le_less) + +lemma real_sqrt_sum_squares_ge_zero [simp]: "0 \ sqrt (x\ + y\)" +by (auto intro!: real_sqrt_ge_zero) + +lemma real_sqrt_sum_squares_mult_ge_zero [simp]: "0 \ sqrt ((x\ + y\)*(xa\ + ya\))" +by (auto intro!: real_sqrt_ge_zero simp add: zero_le_mult_iff) + +lemma real_sqrt_sum_squares_mult_squared_eq [simp]: + "sqrt ((x\ + y\) * (xa\ + ya\)) ^ 2 = (x\ + y\) * (xa\ + ya\)" +by (auto simp add: real_sqrt_pow2_iff zero_le_mult_iff simp del: realpow_Suc) + +lemma real_sqrt_abs [simp]: "sqrt(x\) = \x\" +apply (rule abs_realpow_two [THEN subst]) +apply (rule real_sqrt_abs_abs [THEN subst]) +apply (subst real_pow_sqrt_eq_sqrt_pow) +apply (auto simp add: numeral_2_eq_2 abs_mult) +done + +lemma real_sqrt_abs2 [simp]: "sqrt(x*x) = \x\" +apply (rule realpow_two [THEN subst]) +apply (subst numeral_2_eq_2 [symmetric]) +apply (rule real_sqrt_abs) +done + +lemma real_sqrt_pow2_gt_zero: "0 < x ==> 0 < (sqrt x)\" +by simp + +lemma real_sqrt_not_eq_zero: "0 < x ==> sqrt x \ 0" +apply (frule real_sqrt_pow2_gt_zero) +apply (auto simp add: numeral_2_eq_2 order_less_irrefl) +done + +lemma real_inv_sqrt_pow2: "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x" +by (cut_tac n1 = 2 and a1 = "sqrt x" in power_inverse [symmetric], auto) + +lemma real_sqrt_eq_zero_cancel: "[| 0 \ x; sqrt(x) = 0|] ==> x = 0" +apply (drule real_le_imp_less_or_eq) +apply (auto dest: real_sqrt_not_eq_zero) +done + +lemma real_sqrt_eq_zero_cancel_iff [simp]: "0 \ x ==> ((sqrt x = 0) = (x = 0))" +by (auto simp add: real_sqrt_eq_zero_cancel) + +lemma real_sqrt_sum_squares_ge1 [simp]: "x \ sqrt(x\ + y\)" +apply (subgoal_tac "x \ 0 | 0 \ x", safe) +apply (rule real_le_trans) +apply (auto simp del: realpow_Suc) +apply (rule_tac n = 1 in realpow_increasing) +apply (auto simp add: numeral_2_eq_2 [symmetric] simp del: realpow_Suc) +done + +lemma real_sqrt_sum_squares_ge2 [simp]: "y \ sqrt(z\ + y\)" +apply (simp (no_asm) add: real_add_commute del: realpow_Suc) +done + +lemma real_sqrt_ge_one: "1 \ x ==> 1 \ sqrt x" +apply (rule_tac n = 1 in realpow_increasing) +apply (auto simp add: numeral_2_eq_2 [symmetric] real_sqrt_ge_zero simp + del: realpow_Suc) +done + + +subsection{*Exponential Function*} + +lemma summable_exp: "summable (%n. inverse (real (fact n)) * x ^ n)" +apply (cut_tac 'a = real in zero_less_one [THEN dense], safe) +apply (cut_tac x = r in reals_Archimedean3, auto) +apply (drule_tac x = "\x\" in spec, safe) +apply (rule_tac N = n and c = r in ratio_test) +apply (auto simp add: abs_mult mult_assoc [symmetric] simp del: fact_Suc) +apply (rule mult_right_mono) +apply (rule_tac b1 = "\x\" in mult_commute [THEN ssubst]) +apply (subst fact_Suc) +apply (subst real_of_nat_mult) +apply (auto simp add: abs_mult inverse_mult_distrib) +apply (auto simp add: mult_assoc [symmetric] abs_eqI2 positive_imp_inverse_positive) +apply (rule order_less_imp_le) +apply (rule_tac z1 = "real (Suc na) " in real_mult_less_iff1 [THEN iffD1]) +apply (auto simp add: real_not_refl2 [THEN not_sym] mult_assoc abs_inverse) +apply (erule order_less_trans) +apply (auto simp add: mult_less_cancel_left mult_ac) +done + + +lemma summable_sin: + "summable (%n. + (if even n then 0 + else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * + x ^ n)" +apply (unfold real_divide_def) +apply (rule_tac g = " (%n. inverse (real (fact n)) * \x\ ^ n) " in summable_comparison_test) +apply (rule_tac [2] summable_exp) +apply (rule_tac x = 0 in exI) +apply (auto simp add: power_abs [symmetric] abs_mult zero_le_mult_iff) +done + +lemma summable_cos: + "summable (%n. + (if even n then + (- 1) ^ (n div 2)/(real (fact n)) else 0) * x ^ n)" +apply (unfold real_divide_def) +apply (rule_tac g = " (%n. inverse (real (fact n)) * \x\ ^ n) " in summable_comparison_test) +apply (rule_tac [2] summable_exp) +apply (rule_tac x = 0 in exI) +apply (auto simp add: power_abs [symmetric] abs_mult zero_le_mult_iff) +done + +lemma lemma_STAR_sin [simp]: "(if even n then 0 + else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0" +apply (induct_tac "n", auto) +done + +lemma lemma_STAR_cos [simp]: "0 < n --> + (- 1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" +apply (induct_tac "n", auto) +done + +lemma lemma_STAR_cos1 [simp]: "0 < n --> + (-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" +apply (induct_tac "n", auto) +done + +lemma lemma_STAR_cos2 [simp]: "sumr 1 n (%n. if even n + then (- 1) ^ (n div 2)/(real (fact n)) * + 0 ^ n + else 0) = 0" +apply (induct_tac "n") +apply (case_tac [2] "n", auto) +done + +lemma exp_converges: "(%n. inverse (real (fact n)) * x ^ n) sums exp(x)" +apply (unfold exp_def) +apply (rule summable_exp [THEN summable_sums]) +done + +lemma sin_converges: + "(%n. (if even n then 0 + else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * + x ^ n) sums sin(x)" +apply (unfold sin_def) +apply (rule summable_sin [THEN summable_sums]) +done + +lemma cos_converges: + "(%n. (if even n then + (- 1) ^ (n div 2)/(real (fact n)) + else 0) * x ^ n) sums cos(x)" +apply (unfold cos_def) +apply (rule summable_cos [THEN summable_sums]) +done + +lemma lemma_realpow_diff [rule_format (no_asm)]: "p \ n --> y ^ (Suc n - p) = ((y::real) ^ (n - p)) * y" +apply (induct_tac "n", auto) +apply (subgoal_tac "p = Suc n") +apply (simp (no_asm_simp), auto) +apply (drule sym) +apply (simp add: Suc_diff_le mult_commute realpow_Suc [symmetric] + del: realpow_Suc) +done + + +subsection{*Properties of Power Series*} + +lemma lemma_realpow_diff_sumr: + "sumr 0 (Suc n) (%p. (x ^ p) * y ^ ((Suc n) - p)) = + y * sumr 0 (Suc n) (%p. (x ^ p) * (y ^ (n - p)))" +apply (auto simp add: sumr_mult simp del: sumr_Suc) +apply (rule sumr_subst) +apply (intro strip) +apply (subst lemma_realpow_diff) +apply (auto simp add: mult_ac) +done + +lemma lemma_realpow_diff_sumr2: "x ^ (Suc n) - y ^ (Suc n) = + (x - y) * sumr 0 (Suc n) (%p. (x ^ p) * (y ^(n - p)))" +apply (induct_tac "n", simp) +apply (auto simp del: sumr_Suc) +apply (subst sumr_Suc) +apply (drule sym) +apply (auto simp add: lemma_realpow_diff_sumr right_distrib real_diff_def mult_ac simp del: sumr_Suc) +done + +lemma lemma_realpow_rev_sumr: "sumr 0 (Suc n) (%p. (x ^ p) * (y ^ (n - p))) = + sumr 0 (Suc n) (%p. (x ^ (n - p)) * (y ^ p))" +apply (case_tac "x = y") +apply (auto simp add: mult_commute power_add [symmetric] simp del: sumr_Suc) +apply (rule_tac c1 = "x - y" in real_mult_left_cancel [THEN iffD1]) +apply (rule_tac [2] minus_minus [THEN subst], simp) +apply (subst minus_mult_left) +apply (simp add: lemma_realpow_diff_sumr2 [symmetric] del: sumr_Suc) +done + +text{*Power series has a `circle` of convergence, i.e. if it sums for @{term +x}, then it sums absolutely for @{term z} with @{term "\z\ < \x\"}.*} + +lemma powser_insidea: + "[| summable (%n. f(n) * (x ^ n)); \z\ < \x\ |] + ==> summable (%n. abs(f(n)) * (z ^ n))" +apply (drule summable_LIMSEQ_zero) +apply (drule convergentI) +apply (simp add: Cauchy_convergent_iff [symmetric]) +apply (drule Cauchy_Bseq) +apply (simp add: Bseq_def, safe) +apply (rule_tac g = "%n. K * abs (z ^ n) * inverse (abs (x ^ n))" in summable_comparison_test) +apply (rule_tac x = 0 in exI, safe) +apply (subgoal_tac "0 < abs (x ^ n) ") +apply (rule_tac c="abs (x ^ n)" in mult_right_le_imp_le) +apply (auto simp add: mult_assoc power_abs) + prefer 2 + apply (drule_tac x = 0 in spec, force) +apply (auto simp add: abs_mult power_abs mult_ac) +apply (rule_tac a2 = "z ^ n" + in abs_ge_zero [THEN real_le_imp_less_or_eq, THEN disjE]) +apply (auto intro!: mult_right_mono simp add: mult_assoc [symmetric] power_abs summable_def power_0_left) +apply (rule_tac x = "K * inverse (1 - (\z\ * inverse (\x\))) " in exI) +apply (auto intro!: sums_mult simp add: mult_assoc) +apply (subgoal_tac "abs (z ^ n) * inverse (\x\ ^ n) = (\z\ * inverse (\x\)) ^ n") +apply (auto simp add: power_abs [symmetric]) +apply (subgoal_tac "x \ 0") +apply (subgoal_tac [3] "x \ 0") +apply (auto simp del: abs_inverse abs_mult simp add: abs_inverse [symmetric] realpow_not_zero abs_mult [symmetric] power_inverse power_mult_distrib [symmetric]) +apply (auto intro!: geometric_sums simp add: power_abs inverse_eq_divide) +apply (rule_tac c="\x\" in mult_right_less_imp_less) +apply (auto simp add: abs_mult [symmetric] mult_assoc) +done + +lemma powser_inside: "[| summable (%n. f(n) * (x ^ n)); \z\ < \x\ |] + ==> summable (%n. f(n) * (z ^ n))" +apply (drule_tac z = "\z\" in powser_insidea) +apply (auto intro: summable_rabs_cancel simp add: power_abs [symmetric]) +done + + +subsection{*Differentiation of Power Series*} + +text{*Lemma about distributing negation over it*} +lemma diffs_minus: "diffs (%n. - c n) = (%n. - diffs c n)" +by (simp add: diffs_def) + +text{*Show that we can shift the terms down one*} +lemma lemma_diffs: + "sumr 0 n (%n. (diffs c)(n) * (x ^ n)) = + sumr 0 n (%n. real n * c(n) * (x ^ (n - Suc 0))) + + (real n * c(n) * x ^ (n - Suc 0))" +apply (induct_tac "n") +apply (auto simp add: mult_assoc add_assoc [symmetric] diffs_def) +done + +lemma lemma_diffs2: "sumr 0 n (%n. real n * c(n) * (x ^ (n - Suc 0))) = + sumr 0 n (%n. (diffs c)(n) * (x ^ n)) - + (real n * c(n) * x ^ (n - Suc 0))" +by (auto simp add: lemma_diffs) + + +lemma diffs_equiv: "summable (%n. (diffs c)(n) * (x ^ n)) ==> + (%n. real n * c(n) * (x ^ (n - Suc 0))) sums + (suminf(%n. (diffs c)(n) * (x ^ n)))" +apply (subgoal_tac " (%n. real n * c (n) * (x ^ (n - Suc 0))) ----> 0") +apply (rule_tac [2] LIMSEQ_imp_Suc) +apply (drule summable_sums) +apply (auto simp add: sums_def) +apply (drule_tac X="(\n. \n = 0.. (\d. n = m + d + Suc 0)" +by (simp add: less_iff_Suc_add) + +lemma sumdiff: "a + b - (c + d) = a - c + b - (d::real)" +by arith + + +lemma lemma_termdiff2: " h \ 0 ==> + (((z + h) ^ n) - (z ^ n)) * inverse h - + real n * (z ^ (n - Suc 0)) = + h * sumr 0 (n - Suc 0) (%p. (z ^ p) * + sumr 0 ((n - Suc 0) - p) + (%q. ((z + h) ^ q) * (z ^ (((n - 2) - p) - q))))" +apply (rule real_mult_left_cancel [THEN iffD1], simp (no_asm_simp)) +apply (simp add: right_diff_distrib mult_ac) +apply (simp add: mult_assoc [symmetric]) +apply (case_tac "n") +apply (auto simp add: lemma_realpow_diff_sumr2 + right_diff_distrib [symmetric] mult_assoc + simp del: realpow_Suc sumr_Suc) +apply (auto simp add: lemma_realpow_rev_sumr simp del: sumr_Suc) +apply (auto simp add: real_of_nat_Suc sumr_diff_mult_const left_distrib + sumdiff lemma_termdiff1 sumr_mult) +apply (auto intro!: sumr_subst simp add: real_diff_def real_add_assoc) +apply (simp add: diff_minus [symmetric] less_iff_Suc_add) +apply (auto simp add: sumr_mult lemma_realpow_diff_sumr2 mult_ac simp + del: sumr_Suc realpow_Suc) +done + +lemma lemma_termdiff3: "[| h \ 0; \z\ \ K; abs (z + h) \ K |] + ==> abs (((z + h) ^ n - z ^ n) * inverse h - real n * z ^ (n - Suc 0)) + \ real n * real (n - Suc 0) * K ^ (n - 2) * \h\" +apply (subst lemma_termdiff2, assumption) +apply (simp add: abs_mult mult_commute) +apply (simp add: mult_commute [of _ "K ^ (n - 2)"]) +apply (rule sumr_rabs [THEN real_le_trans]) +apply (simp add: mult_assoc [symmetric]) +apply (simp add: mult_commute [of _ "real (n - Suc 0)"]) +apply (auto intro!: sumr_bound2 simp add: abs_mult) +apply (case_tac "n", auto) +apply (drule less_add_one) +(*CLAIM_SIMP " (a * b * c = a * (c * (b::real))" mult_ac]*) +apply clarify +apply (subgoal_tac "K ^ p * K ^ d * real (Suc (Suc (p + d))) = + K ^ p * (real (Suc (Suc (p + d))) * K ^ d)") +apply (simp (no_asm_simp) add: power_add del: sumr_Suc) +apply (auto intro!: mult_mono simp del: sumr_Suc) +apply (auto intro!: power_mono simp add: power_abs simp del: sumr_Suc) +apply (rule_tac j = "real (Suc d) * (K ^ d) " in real_le_trans) +apply (subgoal_tac [2] "0 \ K") +apply (drule_tac [2] n = d in zero_le_power) +apply (auto simp del: sumr_Suc) +apply (rule sumr_rabs [THEN real_le_trans]) +apply (rule sumr_bound2, auto dest!: less_add_one intro!: mult_mono simp add: abs_mult power_add) +apply (auto intro!: power_mono zero_le_power simp add: power_abs, arith+) +done + +lemma lemma_termdiff4: + "[| 0 < k; + (\h. 0 < \h\ & \h\ < k --> abs(f h) \ K * \h\) |] + ==> f -- 0 --> 0" +apply (unfold LIM_def, auto) +apply (subgoal_tac "0 \ K") +apply (drule_tac [2] x = "k/2" in spec) +apply (frule_tac [2] real_less_half_sum) +apply (drule_tac [2] real_gt_half_sum) +apply (auto simp add: abs_eqI2) +apply (rule_tac [2] c = "k/2" in mult_right_le_imp_le) +apply (auto intro: abs_ge_zero [THEN real_le_trans]) +apply (drule real_le_imp_less_or_eq, auto) +apply (subgoal_tac "0 < (r * inverse K) * inverse 2") +apply (rule_tac [2] real_mult_order)+ +apply (drule_tac ?d1.0 = "r * inverse K * inverse 2" and ?d2.0 = k in real_lbound_gt_zero) +apply (auto simp add: positive_imp_inverse_positive zero_less_mult_iff) +apply (rule_tac [2] y="\f (k / 2)\ * 2" in order_trans, auto) +apply (rule_tac x = e in exI, auto) +apply (rule_tac y = "K * \x\" in order_le_less_trans) +apply (rule_tac [2] y = "K * e" in order_less_trans) +apply (rule_tac [3] c = "inverse K" in mult_right_less_imp_less, force) +apply (simp add: mult_less_cancel_left) +apply (auto simp add: mult_ac) +done + +lemma lemma_termdiff5: "[| 0 < k; + summable f; + \h. 0 < \h\ & \h\ < k --> + (\n. abs(g(h) (n::nat)) \ (f(n) * \h\)) |] + ==> (%h. suminf(g h)) -- 0 --> 0" +apply (drule summable_sums) +apply (subgoal_tac "\h. 0 < \h\ & \h\ < k --> abs (suminf (g h)) \ suminf f * \h\") +apply (auto intro!: lemma_termdiff4 simp add: sums_summable [THEN suminf_mult, symmetric]) +apply (subgoal_tac "summable (%n. f n * \h\) ") + prefer 2 + apply (simp add: summable_def) + apply (rule_tac x = "suminf f * \h\" in exI) + apply (drule_tac c = "\h\" in sums_mult) + apply (simp add: mult_ac) +apply (subgoal_tac "summable (%n. abs (g (h::real) (n::nat))) ") + apply (rule_tac [2] g = "%n. f n * \h\" in summable_comparison_test) + apply (rule_tac [2] x = 0 in exI, auto) +apply (rule_tac j = "suminf (%n. abs (g h n))" in real_le_trans) +apply (auto intro: summable_rabs summable_le simp add: sums_summable [THEN suminf_mult]) +done + + + +text{* FIXME: Long proofs*} + +lemma termdiffs_aux: + "[|summable (\n. diffs (diffs c) n * K ^ n); \x\ < \K\ |] + ==> (\h. suminf + (\n. c n * + (((x + h) ^ n - x ^ n) * inverse h - + real n * x ^ (n - Suc 0)))) + -- 0 --> 0" +apply (drule dense, safe) +apply (frule real_less_sum_gt_zero) +apply (drule_tac + f = "%n. abs (c n) * real n * real (n - Suc 0) * (r ^ (n - 2))" + and g = "%h n. c (n) * ((( ((x + h) ^ n) - (x ^ n)) * inverse h) + - (real n * (x ^ (n - Suc 0))))" + in lemma_termdiff5) +apply (auto simp add: add_commute) +apply (subgoal_tac "summable (%n. \diffs (diffs c) n\ * (r ^ n))") +apply (rule_tac [2] x = K in powser_insidea, auto) +apply (subgoal_tac [2] "\r\ = r", auto) +apply (rule_tac [2] y1 = "\x\" in order_trans [THEN abs_eqI1], auto) +apply (simp add: diffs_def mult_assoc [symmetric]) +apply (subgoal_tac + "\n. real (Suc n) * real (Suc (Suc n)) * \c (Suc (Suc n))\ * (r ^ n) + = diffs (diffs (%n. \c n\)) n * (r ^ n) ") +apply auto +apply (drule diffs_equiv) +apply (drule sums_summable) +apply (simp_all add: diffs_def) +apply (simp add: diffs_def mult_ac) +apply (subgoal_tac " (%n. real n * (real (Suc n) * (abs (c (Suc n)) * (r ^ (n - Suc 0))))) = (%n. diffs (%m. real (m - Suc 0) * abs (c m) * inverse r) n * (r ^ n))") +apply auto + prefer 2 + apply (rule ext) + apply (simp add: diffs_def) + apply (case_tac "n", auto) +txt{*23*} + apply (drule abs_ge_zero [THEN order_le_less_trans]) + apply (simp add: mult_ac) + apply (drule abs_ge_zero [THEN order_le_less_trans]) + apply (simp add: mult_ac) + apply (drule diffs_equiv) + apply (drule sums_summable) +apply (subgoal_tac + "summable + (\n. real n * (real (n - Suc 0) * \c n\ * inverse r) * + r ^ (n - Suc 0)) = + summable + (\n. real n * (\c n\ * (real (n - Suc 0) * r ^ (n - 2))))") +apply simp +apply (rule_tac f = summable in arg_cong, rule ext) +txt{*33*} +apply (case_tac "n", auto) +apply (case_tac "nat", auto) +apply (drule abs_ge_zero [THEN order_le_less_trans], auto) +apply (drule abs_ge_zero [THEN order_le_less_trans]) +apply (simp add: mult_assoc) +apply (rule mult_left_mono) +apply (rule add_commute [THEN subst]) +apply (simp (no_asm) add: mult_assoc [symmetric]) +apply (rule lemma_termdiff3) +apply (auto intro: abs_triangle_ineq [THEN order_trans], arith) +done + + +lemma termdiffs: + "[| summable(%n. c(n) * (K ^ n)); + summable(%n. (diffs c)(n) * (K ^ n)); + summable(%n. (diffs(diffs c))(n) * (K ^ n)); + \x\ < \K\ |] + ==> DERIV (%x. suminf (%n. c(n) * (x ^ n))) x :> + suminf (%n. (diffs c)(n) * (x ^ n))" +apply (unfold deriv_def) +apply (rule_tac g = "%h. suminf (%n. ((c (n) * ( (x + h) ^ n)) - (c (n) * (x ^ n))) * inverse h) " in LIM_trans) +apply (simp add: LIM_def, safe) +apply (rule_tac x = "\K\ - \x\" in exI) +apply (auto simp add: less_diff_eq) +apply (drule abs_triangle_ineq [THEN order_le_less_trans]) +apply (rule_tac y = 0 in order_le_less_trans, auto) +apply (subgoal_tac " (%n. (c n) * (x ^ n)) sums (suminf (%n. (c n) * (x ^ n))) & (%n. (c n) * ((x + xa) ^ n)) sums (suminf (%n. (c n) * ( (x + xa) ^ n))) ") +apply (auto intro!: summable_sums) +apply (rule_tac [2] powser_inside, rule_tac [4] powser_inside) +apply (auto simp add: add_commute) +apply (drule_tac x="(\n. c n * (xa + x) ^ n)" in sums_diff, assumption) +apply (drule_tac x = " (%n. c n * (xa + x) ^ n - c n * x ^ n) " and c = "inverse xa" in sums_mult) +apply (rule sums_unique [symmetric]) +apply (simp add: diff_def real_divide_def add_ac mult_ac) +apply (rule LIM_zero_cancel) +apply (rule_tac g = "%h. suminf (%n. c (n) * ((( ((x + h) ^ n) - (x ^ n)) * inverse h) - (real n * (x ^ (n - Suc 0))))) " in LIM_trans) + prefer 2 apply (blast intro: termdiffs_aux) +apply (simp add: LIM_def, safe) +apply (rule_tac x = "\K\ - \x\" in exI) +apply (auto simp add: less_diff_eq) +apply (drule abs_triangle_ineq [THEN order_le_less_trans]) +apply (rule_tac y = 0 in order_le_less_trans, auto) +apply (subgoal_tac "summable (%n. (diffs c) (n) * (x ^ n))") +apply (rule_tac [2] powser_inside, auto) +apply (drule_tac c = c and x = x in diffs_equiv) +apply (frule sums_unique, auto) +apply (subgoal_tac " (%n. (c n) * (x ^ n)) sums (suminf (%n. (c n) * (x ^ n))) & (%n. (c n) * ((x + xa) ^ n)) sums (suminf (%n. (c n) * ( (x + xa) ^ n))) ") +apply safe +apply (auto intro!: summable_sums) +apply (rule_tac [2] powser_inside, rule_tac [4] powser_inside) +apply (auto simp add: add_commute) +apply (frule_tac x = " (%n. c n * (xa + x) ^ n) " and y = " (%n. c n * x ^ n) " in sums_diff, assumption) +apply (simp add: suminf_diff [OF sums_summable sums_summable] + right_diff_distrib [symmetric]) +apply (frule_tac x = " (%n. c n * ((xa + x) ^ n - x ^ n))" and c = "inverse xa" in sums_mult) +apply (simp add: sums_summable [THEN suminf_mult2]) +apply (frule_tac x = " (%n. inverse xa * (c n * ((xa + x) ^ n - x ^ n))) " and y = " (%n. real n * c n * x ^ (n - Suc 0))" in sums_diff) +apply assumption +apply (subst minus_equation_iff, simp (no_asm)) +apply (simp add: suminf_diff [OF sums_summable sums_summable] add_ac mult_ac) +apply (rule_tac f = suminf in arg_cong) +apply (rule ext) +apply (simp add: diff_def right_distrib add_ac mult_ac) +done + + +subsection{*Formal Derivatives of Exp, Sin, and Cos Series*} + +lemma exp_fdiffs: + "diffs (%n. inverse(real (fact n))) = (%n. inverse(real (fact n)))" +apply (unfold diffs_def) +apply (rule ext) +apply (subst fact_Suc) +apply (subst real_of_nat_mult) +apply (subst inverse_mult_distrib) +apply (auto simp add: mult_assoc [symmetric]) +done + +lemma sin_fdiffs: + "diffs(%n. if even n then 0 + else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) + = (%n. if even n then + (- 1) ^ (n div 2)/(real (fact n)) + else 0)" +apply (unfold diffs_def real_divide_def) +apply (rule ext) +apply (subst fact_Suc) +apply (subst real_of_nat_mult) +apply (subst even_nat_Suc) +apply (subst inverse_mult_distrib, auto) +done + +lemma sin_fdiffs2: + "diffs(%n. if even n then 0 + else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) n + = (if even n then + (- 1) ^ (n div 2)/(real (fact n)) + else 0)" +apply (unfold diffs_def real_divide_def) +apply (subst fact_Suc) +apply (subst real_of_nat_mult) +apply (subst even_nat_Suc) +apply (subst inverse_mult_distrib, auto) +done + +lemma cos_fdiffs: + "diffs(%n. if even n then + (- 1) ^ (n div 2)/(real (fact n)) else 0) + = (%n. - (if even n then 0 + else (- 1) ^ ((n - Suc 0)div 2)/(real (fact n))))" +apply (unfold diffs_def real_divide_def) +apply (rule ext) +apply (subst fact_Suc) +apply (subst real_of_nat_mult) +apply (auto simp add: mult_assoc odd_Suc_mult_two_ex) +done + + +lemma cos_fdiffs2: + "diffs(%n. if even n then + (- 1) ^ (n div 2)/(real (fact n)) else 0) n + = - (if even n then 0 + else (- 1) ^ ((n - Suc 0)div 2)/(real (fact n)))" +apply (unfold diffs_def real_divide_def) +apply (subst fact_Suc) +apply (subst real_of_nat_mult) +apply (auto simp add: mult_assoc odd_Suc_mult_two_ex) +done + +text{*Now at last we can get the derivatives of exp, sin and cos*} + +lemma lemma_sin_minus: + "- sin x = + suminf(%n. - ((if even n then 0 + else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * x ^ n))" +by (auto intro!: sums_unique sums_minus sin_converges) + +lemma lemma_exp_ext: "exp = (%x. suminf (%n. inverse (real (fact n)) * x ^ n))" +by (auto intro!: ext simp add: exp_def) + +lemma DERIV_exp [simp]: "DERIV exp x :> exp(x)" +apply (unfold exp_def) +apply (subst lemma_exp_ext) +apply (subgoal_tac "DERIV (%u. suminf (%n. inverse (real (fact n)) * u ^ n)) x :> suminf (%n. diffs (%n. inverse (real (fact n))) n * x ^ n) ") +apply (rule_tac [2] K = "1 + \x\ " in termdiffs) +apply (auto intro: exp_converges [THEN sums_summable] simp add: exp_fdiffs, arith) +done + +lemma lemma_sin_ext: + "sin = (%x. suminf(%n. + (if even n then 0 + else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * + x ^ n))" +by (auto intro!: ext simp add: sin_def) + +lemma lemma_cos_ext: + "cos = (%x. suminf(%n. + (if even n then (- 1) ^ (n div 2)/(real (fact n)) else 0) * + x ^ n))" +by (auto intro!: ext simp add: cos_def) + +lemma DERIV_sin [simp]: "DERIV sin x :> cos(x)" +apply (unfold cos_def) +apply (subst lemma_sin_ext) +apply (auto simp add: sin_fdiffs2 [symmetric]) +apply (rule_tac K = "1 + \x\ " in termdiffs) +apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs, arith) +done + +lemma DERIV_cos [simp]: "DERIV cos x :> -sin(x)" +apply (subst lemma_cos_ext) +apply (auto simp add: lemma_sin_minus cos_fdiffs2 [symmetric] minus_mult_left) +apply (rule_tac K = "1 + \x\ " in termdiffs) +apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs diffs_minus, arith) +done + + +subsection{*Properties of the Exponential Function*} + +lemma exp_zero [simp]: "exp 0 = 1" +proof - + have "(\n = 0..<1. inverse (real (fact n)) * 0 ^ n) = + suminf (\n. inverse (real (fact n)) * 0 ^ n)" + by (rule series_zero [rule_format, THEN sums_unique], + case_tac "m", auto) + thus ?thesis by (simp add: exp_def) +qed + +lemma exp_ge_add_one_self [simp]: "0 \ x ==> (1 + x) \ exp(x)" +apply (drule real_le_imp_less_or_eq, auto) +apply (unfold exp_def) +apply (rule real_le_trans) +apply (rule_tac [2] n = 2 and f = " (%n. inverse (real (fact n)) * x ^ n) " in series_pos_le) +apply (auto intro: summable_exp simp add: numeral_2_eq_2 zero_le_power zero_le_mult_iff) +done + +lemma exp_gt_one [simp]: "0 < x ==> 1 < exp x" +apply (rule order_less_le_trans) +apply (rule_tac [2] exp_ge_add_one_self, auto) +done + +lemma DERIV_exp_add_const: "DERIV (%x. exp (x + y)) x :> exp(x + y)" +proof - + have "DERIV (exp \ (\x. x + y)) x :> exp (x + y) * (1+0)" + by (fast intro: DERIV_chain DERIV_add DERIV_exp DERIV_Id DERIV_const) + thus ?thesis by (simp add: o_def) +qed + +lemma DERIV_exp_minus [simp]: "DERIV (%x. exp (-x)) x :> - exp(-x)" +proof - + have "DERIV (exp \ uminus) x :> exp (- x) * - 1" + by (fast intro: DERIV_chain DERIV_minus DERIV_exp DERIV_Id) + thus ?thesis by (simp add: o_def) +qed + +lemma DERIV_exp_exp_zero [simp]: "DERIV (%x. exp (x + y) * exp (- x)) x :> 0" +proof - + have "DERIV (\x. exp (x + y) * exp (- x)) x + :> exp (x + y) * exp (- x) + - exp (- x) * exp (x + y)" + by (fast intro: DERIV_exp_add_const DERIV_exp_minus DERIV_mult) + thus ?thesis by simp +qed + +lemma exp_add_mult_minus [simp]: "exp(x + y)*exp(-x) = exp(y)" +proof - + have "\x. DERIV (%x. exp (x + y) * exp (- x)) x :> 0" by simp + hence "exp (x + y) * exp (- x) = exp (0 + y) * exp (- 0)" + by (rule DERIV_isconst_all) + thus ?thesis by simp +qed + +lemma exp_mult_minus [simp]: "exp x * exp(-x) = 1" +proof - + have "exp (x + 0) * exp (- x) = exp 0" by (rule exp_add_mult_minus) + thus ?thesis by simp +qed + +lemma exp_mult_minus2 [simp]: "exp(-x)*exp(x) = 1" +by (simp add: mult_commute) + + +lemma exp_minus: "exp(-x) = inverse(exp(x))" +by (auto intro: inverse_unique [symmetric]) + +lemma exp_add: "exp(x + y) = exp(x) * exp(y)" +proof - + have "exp x * exp y = exp x * (exp (x + y) * exp (- x))" by simp + thus ?thesis by (simp (no_asm_simp) add: mult_ac) +qed + +text{*Proof: because every exponential can be seen as a square.*} +lemma exp_ge_zero [simp]: "0 \ exp x" +apply (rule_tac t = x in real_sum_of_halves [THEN subst]) +apply (subst exp_add, auto) +done + +lemma exp_not_eq_zero [simp]: "exp x \ 0" +apply (cut_tac x = x in exp_mult_minus2) +apply (auto simp del: exp_mult_minus2) +done + +lemma exp_gt_zero [simp]: "0 < exp x" +by (simp add: order_less_le) + +lemma inv_exp_gt_zero [simp]: "0 < inverse(exp x)" +by (auto intro: positive_imp_inverse_positive) + +lemma abs_exp_cancel [simp]: "abs(exp x) = exp x" +by (auto simp add: abs_eqI2) + +lemma exp_real_of_nat_mult: "exp(real n * x) = exp(x) ^ n" +apply (induct_tac "n") +apply (auto simp add: real_of_nat_Suc right_distrib exp_add mult_commute) +done + +lemma exp_diff: "exp(x - y) = exp(x)/(exp y)" +apply (unfold real_diff_def real_divide_def) +apply (simp (no_asm) add: exp_add exp_minus) +done + + +lemma exp_less_mono: + assumes xy: "x < y" shows "exp x < exp y" +proof - + have "1 < exp (y + - x)" + by (rule real_less_sum_gt_zero [THEN exp_gt_one]) + hence "exp x * inverse (exp x) < exp y * inverse (exp x)" + by (auto simp add: exp_add exp_minus) + thus ?thesis + by (simp add: divide_inverse [symmetric] pos_less_divide_eq) +qed + +lemma exp_less_cancel: "exp x < exp y ==> x < y" +apply (rule ccontr) +apply (simp add: linorder_not_less order_le_less) +apply (auto dest: exp_less_mono) +done + +lemma exp_less_cancel_iff [iff]: "(exp(x) < exp(y)) = (x < y)" +by (auto intro: exp_less_mono exp_less_cancel) + +lemma exp_le_cancel_iff [iff]: "(exp(x) \ exp(y)) = (x \ y)" +by (auto simp add: linorder_not_less [symmetric]) + +lemma exp_inj_iff [iff]: "(exp x = exp y) = (x = y)" +by (simp add: order_eq_iff) + +lemma lemma_exp_total: "1 \ y ==> \x. 0 \ x & x \ y - 1 & exp(x) = y" +apply (rule IVT) +apply (auto intro: DERIV_exp [THEN DERIV_isCont] simp add: le_diff_eq) +apply (subgoal_tac "1 + (y - 1) \ exp (y - 1)") +apply simp +apply (rule exp_ge_add_one_self, simp) +done + +lemma exp_total: "0 < y ==> \x. exp x = y" +apply (rule_tac x = 1 and y = y in linorder_cases) +apply (drule order_less_imp_le [THEN lemma_exp_total]) +apply (rule_tac [2] x = 0 in exI) +apply (frule_tac [3] real_inverse_gt_one) +apply (drule_tac [4] order_less_imp_le [THEN lemma_exp_total], auto) +apply (rule_tac x = "-x" in exI) +apply (simp add: exp_minus) +done + + +subsection{*Properties of the Logarithmic Function*} + +lemma ln_exp[simp]: "ln(exp x) = x" +by (simp add: ln_def) + +lemma exp_ln_iff[simp]: "(exp(ln x) = x) = (0 < x)" +apply (auto dest: exp_total) +apply (erule subst, simp) +done + +lemma ln_mult: "[| 0 < x; 0 < y |] ==> ln(x * y) = ln(x) + ln(y)" +apply (rule exp_inj_iff [THEN iffD1]) +apply (frule real_mult_order) +apply (auto simp add: exp_add exp_ln_iff [symmetric] simp del: exp_inj_iff exp_ln_iff) +done + +lemma ln_inj_iff[simp]: "[| 0 < x; 0 < y |] ==> (ln x = ln y) = (x = y)" +apply (simp only: exp_ln_iff [symmetric]) +apply (erule subst)+ +apply simp +done + +lemma ln_one[simp]: "ln 1 = 0" +by (rule exp_inj_iff [THEN iffD1], auto) + +lemma ln_inverse: "0 < x ==> ln(inverse x) = - ln x" +apply (rule_tac a1 = "ln x" in add_left_cancel [THEN iffD1]) +apply (auto simp add: positive_imp_inverse_positive ln_mult [symmetric]) +done + +lemma ln_div: + "[|0 < x; 0 < y|] ==> ln(x/y) = ln x - ln y" +apply (unfold real_divide_def) +apply (auto simp add: positive_imp_inverse_positive ln_mult ln_inverse) +done + +lemma ln_less_cancel_iff[simp]: "[| 0 < x; 0 < y|] ==> (ln x < ln y) = (x < y)" +apply (simp only: exp_ln_iff [symmetric]) +apply (erule subst)+ +apply simp +done + +lemma ln_le_cancel_iff[simp]: "[| 0 < x; 0 < y|] ==> (ln x \ ln y) = (x \ y)" +by (auto simp add: linorder_not_less [symmetric]) + +lemma ln_realpow: "0 < x ==> ln(x ^ n) = real n * ln(x)" +by (auto dest!: exp_total simp add: exp_real_of_nat_mult [symmetric]) + +lemma ln_add_one_self_le_self [simp]: "0 \ x ==> ln(1 + x) \ x" +apply (rule ln_exp [THEN subst]) +apply (rule ln_le_cancel_iff [THEN iffD2], auto) +done + +lemma ln_less_self [simp]: "0 < x ==> ln x < x" +apply (rule order_less_le_trans) +apply (rule_tac [2] ln_add_one_self_le_self) +apply (rule ln_less_cancel_iff [THEN iffD2], auto) +done + +lemma ln_ge_zero: + assumes x: "1 \ x" shows "0 \ ln x" +proof - + have "0 < x" using x by arith + hence "exp 0 \ exp (ln x)" + by (simp add: x exp_ln_iff [symmetric] del: exp_ln_iff) + thus ?thesis by (simp only: exp_le_cancel_iff) +qed + +lemma ln_ge_zero_imp_ge_one: + assumes ln: "0 \ ln x" + and x: "0 < x" + shows "1 \ x" +proof - + from ln have "ln 1 \ ln x" by simp + thus ?thesis by (simp add: x del: ln_one) +qed + +lemma ln_ge_zero_iff [simp]: "0 < x ==> (0 \ ln x) = (1 \ x)" +by (blast intro: ln_ge_zero ln_ge_zero_imp_ge_one) + +lemma ln_gt_zero: + assumes x: "1 < x" shows "0 < ln x" +proof - + have "0 < x" using x by arith + hence "exp 0 < exp (ln x)" + by (simp add: x exp_ln_iff [symmetric] del: exp_ln_iff) + thus ?thesis by (simp only: exp_less_cancel_iff) +qed + +lemma ln_gt_zero_imp_gt_one: + assumes ln: "0 < ln x" + and x: "0 < x" + shows "1 < x" +proof - + from ln have "ln 1 < ln x" by simp + thus ?thesis by (simp add: x del: ln_one) +qed + +lemma ln_gt_zero_iff [simp]: "0 < x ==> (0 < ln x) = (1 < x)" +by (blast intro: ln_gt_zero ln_gt_zero_imp_gt_one) + +lemma ln_not_eq_zero [simp]: "[| 0 < x; x \ 1 |] ==> ln x \ 0" +apply safe +apply (drule exp_inj_iff [THEN iffD2]) +apply (drule exp_ln_iff [THEN iffD2], auto) +done + +lemma ln_less_zero: "[| 0 < x; x < 1 |] ==> ln x < 0" +apply (rule exp_less_cancel_iff [THEN iffD1]) +apply (auto simp add: exp_ln_iff [symmetric] simp del: exp_ln_iff) +done + +lemma exp_ln_eq: "exp u = x ==> ln x = u" +by auto + + +subsection{*Basic Properties of the Trigonometric Functions*} + +lemma sin_zero [simp]: "sin 0 = 0" +by (auto intro!: sums_unique [symmetric] LIMSEQ_const + simp add: sin_def sums_def simp del: power_0_left) + +lemma lemma_series_zero2: "(\m. n \ m --> f m = 0) --> f sums sumr 0 n f" +by (auto intro: series_zero) + +lemma cos_zero [simp]: "cos 0 = 1" +apply (unfold cos_def) +apply (rule sums_unique [symmetric]) +apply (cut_tac n = 1 and f = " (%n. (if even n then (- 1) ^ (n div 2) / (real (fact n)) else 0) * 0 ^ n) " in lemma_series_zero2) +apply auto +done + +lemma DERIV_sin_sin_mult [simp]: + "DERIV (%x. sin(x)*sin(x)) x :> cos(x) * sin(x) + cos(x) * sin(x)" +by (rule DERIV_mult, auto) + +lemma DERIV_sin_sin_mult2 [simp]: + "DERIV (%x. sin(x)*sin(x)) x :> 2 * cos(x) * sin(x)" +apply (cut_tac x = x in DERIV_sin_sin_mult) +apply (auto simp add: mult_assoc) +done + +lemma DERIV_sin_realpow2 [simp]: + "DERIV (%x. (sin x)\) x :> cos(x) * sin(x) + cos(x) * sin(x)" +by (auto simp add: numeral_2_eq_2 real_mult_assoc [symmetric]) + +lemma DERIV_sin_realpow2a [simp]: + "DERIV (%x. (sin x)\) x :> 2 * cos(x) * sin(x)" +by (auto simp add: numeral_2_eq_2) + +lemma DERIV_cos_cos_mult [simp]: + "DERIV (%x. cos(x)*cos(x)) x :> -sin(x) * cos(x) + -sin(x) * cos(x)" +by (rule DERIV_mult, auto) + +lemma DERIV_cos_cos_mult2 [simp]: + "DERIV (%x. cos(x)*cos(x)) x :> -2 * cos(x) * sin(x)" +apply (cut_tac x = x in DERIV_cos_cos_mult) +apply (auto simp add: mult_ac) +done + +lemma DERIV_cos_realpow2 [simp]: + "DERIV (%x. (cos x)\) x :> -sin(x) * cos(x) + -sin(x) * cos(x)" +by (auto simp add: numeral_2_eq_2 real_mult_assoc [symmetric]) + +lemma DERIV_cos_realpow2a [simp]: + "DERIV (%x. (cos x)\) x :> -2 * cos(x) * sin(x)" +by (auto simp add: numeral_2_eq_2) + +lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" +by auto + +lemma DERIV_cos_realpow2b: "DERIV (%x. (cos x)\) x :> -(2 * cos(x) * sin(x))" +apply (rule lemma_DERIV_subst) +apply (rule DERIV_cos_realpow2a, auto) +done + +(* most useful *) +lemma DERIV_cos_cos_mult3 [simp]: "DERIV (%x. cos(x)*cos(x)) x :> -(2 * cos(x) * sin(x))" +apply (rule lemma_DERIV_subst) +apply (rule DERIV_cos_cos_mult2, auto) +done + +lemma DERIV_sin_circle_all: + "\x. DERIV (%x. (sin x)\ + (cos x)\) x :> + (2*cos(x)*sin(x) - 2*cos(x)*sin(x))" +apply (unfold real_diff_def, safe) +apply (rule DERIV_add) +apply (auto simp add: numeral_2_eq_2) +done + +lemma DERIV_sin_circle_all_zero [simp]: "\x. DERIV (%x. (sin x)\ + (cos x)\) x :> 0" +by (cut_tac DERIV_sin_circle_all, auto) + +lemma sin_cos_squared_add [simp]: "((sin x)\) + ((cos x)\) = 1" +apply (cut_tac x = x and y = 0 in DERIV_sin_circle_all_zero [THEN DERIV_isconst_all]) +apply (auto simp add: numeral_2_eq_2) +done + +lemma sin_cos_squared_add2 [simp]: "((cos x)\) + ((sin x)\) = 1" +apply (subst real_add_commute) +apply (simp (no_asm) del: realpow_Suc) +done + +lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1" +apply (cut_tac x = x in sin_cos_squared_add2) +apply (auto simp add: numeral_2_eq_2) +done + +lemma sin_squared_eq: "(sin x)\ = 1 - (cos x)\" +apply (rule_tac a1 = "(cos x)\ " in add_right_cancel [THEN iffD1]) +apply (simp del: realpow_Suc) +done + +lemma cos_squared_eq: "(cos x)\ = 1 - (sin x)\" +apply (rule_tac a1 = "(sin x)\" in add_right_cancel [THEN iffD1]) +apply (simp del: realpow_Suc) +done + +lemma real_gt_one_ge_zero_add_less: "[| 1 < x; 0 \ y |] ==> 1 < x + (y::real)" +by arith + +lemma abs_sin_le_one [simp]: "abs(sin x) \ 1" +apply (auto simp add: linorder_not_less [symmetric]) +apply (drule_tac n = "Suc 0" in power_gt1) +apply (auto simp del: realpow_Suc) +apply (drule_tac r1 = "cos x" in realpow_two_le [THEN [2] real_gt_one_ge_zero_add_less]) +apply (simp add: numeral_2_eq_2 [symmetric] del: realpow_Suc) +done + +lemma sin_ge_minus_one [simp]: "-1 \ sin x" +apply (insert abs_sin_le_one [of x]) +apply (simp add: abs_le_interval_iff del: abs_sin_le_one) +done + +lemma sin_le_one [simp]: "sin x \ 1" +apply (insert abs_sin_le_one [of x]) +apply (simp add: abs_le_interval_iff del: abs_sin_le_one) +done + +lemma abs_cos_le_one [simp]: "abs(cos x) \ 1" +apply (auto simp add: linorder_not_less [symmetric]) +apply (drule_tac n = "Suc 0" in power_gt1) +apply (auto simp del: realpow_Suc) +apply (drule_tac r1 = "sin x" in realpow_two_le [THEN [2] real_gt_one_ge_zero_add_less]) +apply (simp add: numeral_2_eq_2 [symmetric] del: realpow_Suc) +done + +lemma cos_ge_minus_one [simp]: "-1 \ cos x" +apply (insert abs_cos_le_one [of x]) +apply (simp add: abs_le_interval_iff del: abs_cos_le_one) +done + +lemma cos_le_one [simp]: "cos x \ 1" +apply (insert abs_cos_le_one [of x]) +apply (simp add: abs_le_interval_iff del: abs_cos_le_one) +done + +lemma DERIV_fun_pow: "DERIV g x :> m ==> + DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = " (%x. x ^ n) " in DERIV_chain2) +apply (rule DERIV_pow, auto) +done + +lemma DERIV_fun_exp: "DERIV g x :> m ==> DERIV (%x. exp(g x)) x :> exp(g x) * m" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = exp in DERIV_chain2) +apply (rule DERIV_exp, auto) +done + +lemma DERIV_fun_sin: "DERIV g x :> m ==> DERIV (%x. sin(g x)) x :> cos(g x) * m" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = sin in DERIV_chain2) +apply (rule DERIV_sin, auto) +done + +lemma DERIV_fun_cos: "DERIV g x :> m ==> DERIV (%x. cos(g x)) x :> -sin(g x) * m" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = cos in DERIV_chain2) +apply (rule DERIV_cos, auto) +done + +lemmas DERIV_intros = DERIV_Id DERIV_const DERIV_cos DERIV_cmult + DERIV_sin DERIV_exp DERIV_inverse DERIV_pow + DERIV_add DERIV_diff DERIV_mult DERIV_minus + DERIV_inverse_fun DERIV_quotient DERIV_fun_pow + DERIV_fun_exp DERIV_fun_sin DERIV_fun_cos + DERIV_Id DERIV_const DERIV_cos + +(* lemma *) +lemma lemma_DERIV_sin_cos_add: "\x. + DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + + (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2) x :> 0" +apply (safe, rule lemma_DERIV_subst) +apply (best intro!: DERIV_intros intro: DERIV_chain2) + --{*replaces the old @{text DERIV_tac}*} +apply (auto simp add: real_diff_def left_distrib right_distrib mult_ac add_ac) +done + +lemma sin_cos_add [simp]: + "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + + (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = 0" +apply (cut_tac y = 0 and x = x and y7 = y + in lemma_DERIV_sin_cos_add [THEN DERIV_isconst_all]) +apply (auto simp add: numeral_2_eq_2) +done + +lemma sin_add: "sin (x + y) = sin x * cos y + cos x * sin y" +apply (cut_tac x = x and y = y in sin_cos_add) +apply (auto dest!: real_sum_squares_cancel_a + simp add: numeral_2_eq_2 simp del: sin_cos_add) +done + +lemma cos_add: "cos (x + y) = cos x * cos y - sin x * sin y" +apply (cut_tac x = x and y = y in sin_cos_add) +apply (auto dest!: real_sum_squares_cancel_a + simp add: numeral_2_eq_2 simp del: sin_cos_add) +done + +lemma lemma_DERIV_sin_cos_minus: "\x. + DERIV (%x. (sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2) x :> 0" +apply (safe, rule lemma_DERIV_subst) +apply (best intro!: DERIV_intros intro: DERIV_chain2) +apply (auto simp add: real_diff_def left_distrib right_distrib mult_ac add_ac) +done + +lemma sin_cos_minus [simp]: "(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0" +apply (cut_tac y = 0 and x = x in lemma_DERIV_sin_cos_minus [THEN DERIV_isconst_all]) +apply (auto simp add: numeral_2_eq_2) +done + +lemma sin_minus [simp]: "sin (-x) = -sin(x)" +apply (cut_tac x = x in sin_cos_minus) +apply (auto dest!: real_sum_squares_cancel_a simp add: numeral_2_eq_2 simp del: sin_cos_minus) +done + +lemma cos_minus [simp]: "cos (-x) = cos(x)" +apply (cut_tac x = x in sin_cos_minus) +apply (auto dest!: real_sum_squares_cancel_a simp add: numeral_2_eq_2 simp del: sin_cos_minus) +done + +lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y" +apply (unfold real_diff_def) +apply (simp (no_asm) add: sin_add) +done + +lemma sin_diff2: "sin (x - y) = cos y * sin x - sin y * cos x" +by (simp add: sin_diff mult_commute) + +lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y" +apply (unfold real_diff_def) +apply (simp (no_asm) add: cos_add) +done + +lemma cos_diff2: "cos (x - y) = cos y * cos x + sin y * sin x" +by (simp add: cos_diff mult_commute) + +lemma sin_double [simp]: "sin(2 * x) = 2* sin x * cos x" +by (cut_tac x = x and y = x in sin_add, auto) + + +lemma cos_double: "cos(2* x) = ((cos x)\) - ((sin x)\)" +apply (cut_tac x = x and y = x in cos_add) +apply (auto simp add: numeral_2_eq_2) +done + + +subsection{*The Constant Pi*} + +text{*Show that there's a least positive @{term x} with @{term "cos(x) = 0"}; + hence define pi.*} + +lemma sin_paired: + "(%n. (- 1) ^ n /(real (fact (2 * n + 1))) * x ^ (2 * n + 1)) + sums sin x" +proof - + have "(\n. \k = n * 2..n. (if even n then 0 + else (- 1) ^ ((n - Suc 0) div 2) / real (fact n)) * + x ^ n)" + by (rule sin_converges [THEN sums_summable, THEN sums_group], simp) + thus ?thesis by (simp add: mult_ac sin_def) +qed + +lemma sin_gt_zero: "[|0 < x; x < 2 |] ==> 0 < sin x" +apply (subgoal_tac + "(\n. \k = n * 2..n. (- 1) ^ n / real (fact (2 * n + 1)) * x ^ (2 * n + 1))") + prefer 2 + apply (rule sin_paired [THEN sums_summable, THEN sums_group], simp) +apply (rotate_tac 2) +apply (drule sin_paired [THEN sums_unique, THEN ssubst]) +apply (auto simp del: fact_Suc realpow_Suc) +apply (frule sums_unique) +apply (auto simp del: fact_Suc realpow_Suc) +apply (rule_tac n1 = 0 in series_pos_less [THEN [2] order_le_less_trans]) +apply (auto simp del: fact_Suc realpow_Suc) +apply (erule sums_summable) +apply (case_tac "m=0") +apply (simp (no_asm_simp)) +apply (subgoal_tac "6 * (x * (x * x) / real (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) < 6 * x") +apply (simp only: mult_less_cancel_left, simp) +apply (simp (no_asm_simp) add: numeral_2_eq_2 [symmetric] mult_assoc [symmetric]) +apply (subgoal_tac "x*x < 2*3", simp) +apply (rule mult_strict_mono) +apply (auto simp add: real_of_nat_Suc simp del: fact_Suc) +apply (subst fact_Suc) +apply (subst fact_Suc) +apply (subst fact_Suc) +apply (subst fact_Suc) +apply (subst real_of_nat_mult) +apply (subst real_of_nat_mult) +apply (subst real_of_nat_mult) +apply (subst real_of_nat_mult) +apply (simp (no_asm) add: real_divide_def inverse_mult_distrib del: fact_Suc) +apply (auto simp add: mult_assoc [symmetric] simp del: fact_Suc) +apply (rule_tac c="real (Suc (Suc (4*m)))" in mult_less_imp_less_right) +apply (auto simp add: mult_assoc simp del: fact_Suc) +apply (rule_tac c="real (Suc (Suc (Suc (4*m))))" in mult_less_imp_less_right) +apply (auto simp add: mult_assoc mult_less_cancel_left simp del: fact_Suc) +apply (subgoal_tac "x * (x * x ^ (4*m)) = (x ^ (4*m)) * (x * x)") +apply (erule ssubst)+ +apply (auto simp del: fact_Suc) +apply (subgoal_tac "0 < x ^ (4 * m) ") + prefer 2 apply (simp only: zero_less_power) +apply (simp (no_asm_simp) add: mult_less_cancel_left) +apply (rule mult_strict_mono) +apply (simp_all (no_asm_simp)) +done + +lemma sin_gt_zero1: "[|0 < x; x < 2 |] ==> 0 < sin x" +by (auto intro: sin_gt_zero) + +lemma cos_double_less_one: "[| 0 < x; x < 2 |] ==> cos (2 * x) < 1" +apply (cut_tac x = x in sin_gt_zero1) +apply (auto simp add: cos_squared_eq cos_double) +done + +lemma cos_paired: + "(%n. (- 1) ^ n /(real (fact (2 * n))) * x ^ (2 * n)) sums cos x" +proof - + have "(\n. \k = n * 2..n. (if even n then (- 1) ^ (n div 2) / real (fact n) else 0) * + x ^ n)" + by (rule cos_converges [THEN sums_summable, THEN sums_group], simp) + thus ?thesis by (simp add: mult_ac cos_def) +qed + +declare zero_less_power [simp] + +lemma fact_lemma: "real (n::nat) * 4 = real (4 * n)" +by simp + +lemma cos_two_less_zero: "cos (2) < 0" +apply (cut_tac x = 2 in cos_paired) +apply (drule sums_minus) +apply (rule neg_less_iff_less [THEN iffD1]) +apply (frule sums_unique, auto) +apply (rule_tac y = "sumr 0 (Suc (Suc (Suc 0))) (%n. - ((- 1) ^ n / (real (fact (2 * n))) * 2 ^ (2 * n))) " in order_less_trans) +apply (simp (no_asm) add: fact_num_eq_if realpow_num_eq_if del: fact_Suc realpow_Suc) +apply (simp (no_asm) add: mult_assoc del: sumr_Suc) +apply (rule sumr_pos_lt_pair) +apply (erule sums_summable, safe) +apply (simp (no_asm) add: real_divide_def mult_assoc [symmetric] del: fact_Suc) +apply (rule real_mult_inverse_cancel2) +apply (rule real_of_nat_fact_gt_zero)+ +apply (simp (no_asm) add: mult_assoc [symmetric] del: fact_Suc) +apply (subst fact_lemma) +apply (subst fact_Suc) +apply (subst real_of_nat_mult) +apply (erule ssubst, subst real_of_nat_mult) +apply (rule real_mult_less_mono, force) +prefer 2 apply force +apply (rule_tac [2] real_of_nat_fact_gt_zero) +apply (rule real_of_nat_less_iff [THEN iffD2]) +apply (rule fact_less_mono, auto) +done +declare cos_two_less_zero [simp] +declare cos_two_less_zero [THEN real_not_refl2, simp] +declare cos_two_less_zero [THEN order_less_imp_le, simp] + +lemma cos_is_zero: "EX! x. 0 \ x & x \ 2 & cos x = 0" +apply (subgoal_tac "\x. 0 \ x & x \ 2 & cos x = 0") +apply (rule_tac [2] IVT2) +apply (auto intro: DERIV_isCont DERIV_cos) +apply (cut_tac x = xa and y = y in linorder_less_linear) +apply (rule ccontr) +apply (subgoal_tac " (\x. cos differentiable x) & (\x. isCont cos x) ") +apply (auto intro: DERIV_cos DERIV_isCont simp add: differentiable_def) +apply (drule_tac f = cos in Rolle) +apply (drule_tac [5] f = cos in Rolle) +apply (auto dest!: DERIV_cos [THEN DERIV_unique] simp add: differentiable_def) +apply (drule_tac y1 = xa in order_le_less_trans [THEN sin_gt_zero]) +apply (assumption, rule_tac y=y in order_less_le_trans, simp_all) +apply (drule_tac y1 = y in order_le_less_trans [THEN sin_gt_zero], assumption, simp_all) +done + +lemma pi_half: "pi/2 = (@x. 0 \ x & x \ 2 & cos x = 0)" +by (simp add: pi_def) + +lemma cos_pi_half [simp]: "cos (pi / 2) = 0" +apply (rule cos_is_zero [THEN ex1E]) +apply (auto intro!: someI2 simp add: pi_half) +done + +lemma pi_half_gt_zero: "0 < pi / 2" +apply (rule cos_is_zero [THEN ex1E]) +apply (auto simp add: pi_half) +apply (rule someI2, blast, safe) +apply (drule_tac y = xa in real_le_imp_less_or_eq) +apply (safe, simp) +done +declare pi_half_gt_zero [simp] +declare pi_half_gt_zero [THEN real_not_refl2, THEN not_sym, simp] +declare pi_half_gt_zero [THEN order_less_imp_le, simp] + +lemma pi_half_less_two: "pi / 2 < 2" +apply (rule cos_is_zero [THEN ex1E]) +apply (auto simp add: pi_half) +apply (rule someI2, blast, safe) +apply (drule_tac x = xa in order_le_imp_less_or_eq) +apply (safe, simp) +done +declare pi_half_less_two [simp] +declare pi_half_less_two [THEN real_not_refl2, simp] +declare pi_half_less_two [THEN order_less_imp_le, simp] + +lemma pi_gt_zero [simp]: "0 < pi" +apply (rule_tac c="inverse 2" in mult_less_imp_less_right) +apply (cut_tac pi_half_gt_zero, simp_all) +done + +lemma pi_neq_zero [simp]: "pi \ 0" +by (rule pi_gt_zero [THEN real_not_refl2, THEN not_sym]) + +lemma pi_not_less_zero [simp]: "~ (pi < 0)" +apply (insert pi_gt_zero) +apply (blast elim: order_less_asym) +done + +lemma pi_ge_zero [simp]: "0 \ pi" +by (auto intro: order_less_imp_le) + +lemma minus_pi_half_less_zero [simp]: "-(pi/2) < 0" +by auto + +lemma sin_pi_half [simp]: "sin(pi/2) = 1" +apply (cut_tac x = "pi/2" in sin_cos_squared_add2) +apply (cut_tac sin_gt_zero [OF pi_half_gt_zero pi_half_less_two]) +apply (auto simp add: numeral_2_eq_2) +done + +lemma cos_pi [simp]: "cos pi = -1" +by (cut_tac x = "pi/2" and y = "pi/2" in cos_add, simp) + +lemma sin_pi [simp]: "sin pi = 0" +by (cut_tac x = "pi/2" and y = "pi/2" in sin_add, simp) + +lemma sin_cos_eq: "sin x = cos (pi/2 - x)" +apply (unfold real_diff_def) +apply (simp (no_asm) add: cos_add) +done + +lemma minus_sin_cos_eq: "-sin x = cos (x + pi/2)" +apply (simp (no_asm) add: cos_add) +done +declare minus_sin_cos_eq [symmetric, simp] + +lemma cos_sin_eq: "cos x = sin (pi/2 - x)" +apply (unfold real_diff_def) +apply (simp (no_asm) add: sin_add) +done +declare sin_cos_eq [symmetric, simp] cos_sin_eq [symmetric, simp] + +lemma sin_periodic_pi [simp]: "sin (x + pi) = - sin x" +apply (simp (no_asm) add: sin_add) +done + +lemma sin_periodic_pi2 [simp]: "sin (pi + x) = - sin x" +apply (simp (no_asm) add: sin_add) +done + +lemma cos_periodic_pi [simp]: "cos (x + pi) = - cos x" +apply (simp (no_asm) add: cos_add) +done + +lemma sin_periodic [simp]: "sin (x + 2*pi) = sin x" +by (simp add: sin_add cos_double) + +lemma cos_periodic [simp]: "cos (x + 2*pi) = cos x" +by (simp add: cos_add cos_double) + +lemma cos_npi [simp]: "cos (real n * pi) = -1 ^ n" +apply (induct_tac "n") +apply (auto simp add: real_of_nat_Suc left_distrib) +done + +lemma sin_npi [simp]: "sin (real (n::nat) * pi) = 0" +apply (induct_tac "n") +apply (auto simp add: real_of_nat_Suc left_distrib) +done + +lemma sin_npi2 [simp]: "sin (pi * real (n::nat)) = 0" +apply (cut_tac n = n in sin_npi) +apply (auto simp add: mult_commute simp del: sin_npi) +done + +lemma cos_two_pi [simp]: "cos (2 * pi) = 1" +by (simp add: cos_double) + +lemma sin_two_pi [simp]: "sin (2 * pi) = 0" +apply (simp (no_asm)) +done + +lemma sin_gt_zero2: "[| 0 < x; x < pi/2 |] ==> 0 < sin x" +apply (rule sin_gt_zero, assumption) +apply (rule order_less_trans, assumption) +apply (rule pi_half_less_two) +done + +lemma sin_less_zero: + assumes lb: "- pi/2 < x" and "x < 0" shows "sin x < 0" +proof - + have "0 < sin (- x)" using prems by (simp only: sin_gt_zero2) + thus ?thesis by simp +qed + +lemma pi_less_4: "pi < 4" +by (cut_tac pi_half_less_two, auto) + +lemma cos_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < cos x" +apply (cut_tac pi_less_4) +apply (cut_tac f = cos and a = 0 and b = x and y = 0 in IVT2_objl, safe, simp_all) +apply (force intro: DERIV_isCont DERIV_cos) +apply (cut_tac cos_is_zero, safe) +apply (rename_tac y z) +apply (drule_tac x = y in spec) +apply (drule_tac x = "pi/2" in spec, simp) +done + +lemma cos_gt_zero_pi: "[| -(pi/2) < x; x < pi/2 |] ==> 0 < cos x" +apply (rule_tac x = x and y = 0 in linorder_cases) +apply (rule cos_minus [THEN subst]) +apply (rule cos_gt_zero) +apply (auto intro: cos_gt_zero) +done + +lemma cos_ge_zero: "[| -(pi/2) \ x; x \ pi/2 |] ==> 0 \ cos x" +apply (auto simp add: order_le_less cos_gt_zero_pi) +apply (subgoal_tac "x = pi/2", auto) +done + +lemma sin_gt_zero_pi: "[| 0 < x; x < pi |] ==> 0 < sin x" +apply (subst sin_cos_eq) +apply (rotate_tac 1) +apply (drule real_sum_of_halves [THEN ssubst]) +apply (auto intro!: cos_gt_zero_pi simp del: sin_cos_eq [symmetric]) +done + +lemma sin_ge_zero: "[| 0 \ x; x \ pi |] ==> 0 \ sin x" +by (auto simp add: order_le_less sin_gt_zero_pi) + +lemma cos_total: "[| -1 \ y; y \ 1 |] ==> EX! x. 0 \ x & x \ pi & (cos x = y)" +apply (subgoal_tac "\x. 0 \ x & x \ pi & cos x = y") +apply (rule_tac [2] IVT2) +apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos) +apply (cut_tac x = xa and y = y in linorder_less_linear) +apply (rule ccontr, auto) +apply (drule_tac f = cos in Rolle) +apply (drule_tac [5] f = cos in Rolle) +apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos + dest!: DERIV_cos [THEN DERIV_unique] + simp add: differentiable_def) +apply (auto dest: sin_gt_zero_pi [OF order_le_less_trans order_less_le_trans]) +done + +lemma sin_total: + "[| -1 \ y; y \ 1 |] ==> EX! x. -(pi/2) \ x & x \ pi/2 & (sin x = y)" +apply (rule ccontr) +apply (subgoal_tac "\x. (- (pi/2) \ x & x \ pi/2 & (sin x = y)) = (0 \ (x + pi/2) & (x + pi/2) \ pi & (cos (x + pi/2) = -y))") +apply (erule swap) +apply (simp del: minus_sin_cos_eq [symmetric]) +apply (cut_tac y="-y" in cos_total, simp) apply simp +apply (erule ex1E) +apply (rule_tac a = "x - (pi/2) " in ex1I) +apply (simp (no_asm) add: real_add_assoc) +apply (rotate_tac 3) +apply (drule_tac x = "xa + pi/2" in spec, safe, simp_all) +done + +lemma reals_Archimedean4: + "[| 0 < y; 0 \ x |] ==> \n. real n * y \ x & x < real (Suc n) * y" +apply (auto dest!: reals_Archimedean3) +apply (drule_tac x = x in spec, clarify) +apply (subgoal_tac "x < real(LEAST m::nat. x < real m * y) * y") + prefer 2 apply (erule LeastI) +apply (case_tac "LEAST m::nat. x < real m * y", simp) +apply (subgoal_tac "~ x < real nat * y") + prefer 2 apply (rule not_less_Least, simp, force) +done + +(* Pre Isabelle99-2 proof was simpler- numerals arithmetic + now causes some unwanted re-arrangements of literals! *) +lemma cos_zero_lemma: "[| 0 \ x; cos x = 0 |] ==> + \n::nat. ~even n & x = real n * (pi/2)" +apply (drule pi_gt_zero [THEN reals_Archimedean4], safe) +apply (subgoal_tac "0 \ x - real n * pi & (x - real n * pi) \ pi & (cos (x - real n * pi) = 0) ") +apply safe + prefer 3 apply (simp add: cos_diff) + prefer 2 apply (simp add: real_of_nat_Suc left_distrib) +apply (simp add: cos_diff) +apply (subgoal_tac "EX! x. 0 \ x & x \ pi & cos x = 0") +apply (rule_tac [2] cos_total, safe) +apply (drule_tac x = "x - real n * pi" in spec) +apply (drule_tac x = "pi/2" in spec) +apply (simp add: cos_diff) +apply (rule_tac x = "Suc (2 * n) " in exI) +apply (simp add: real_of_nat_Suc left_distrib, auto) +done + +lemma sin_zero_lemma: "[| 0 \ x; sin x = 0 |] ==> + \n::nat. even n & x = real n * (pi/2)" +apply (subgoal_tac "\n::nat. ~ even n & x + pi/2 = real n * (pi/2) ") + apply (clarify, rule_tac x = "n - 1" in exI) + apply (force simp add: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib) +apply (rule cos_zero_lemma, clarify) +apply (rule minus_le_iff [THEN iffD1]) +apply (rule_tac y=0 in order_trans, auto) +done + + +(* also spoilt by numeral arithmetic *) +lemma cos_zero_iff: "(cos x = 0) = + ((\n::nat. ~even n & (x = real n * (pi/2))) | + (\n::nat. ~even n & (x = -(real n * (pi/2)))))" +apply (rule iffI) +apply (cut_tac linorder_linear [of 0 x], safe) +apply (drule cos_zero_lemma, assumption+) +apply (cut_tac x="-x" in cos_zero_lemma, simp, simp) +apply (force simp add: minus_equation_iff [of x]) +apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib) +apply (auto simp add: cos_add) +done + +(* ditto: but to a lesser extent *) +lemma sin_zero_iff: "(sin x = 0) = + ((\n::nat. even n & (x = real n * (pi/2))) | + (\n::nat. even n & (x = -(real n * (pi/2)))))" +apply (rule iffI) +apply (cut_tac linorder_linear [of 0 x], safe) +apply (drule sin_zero_lemma, assumption+) +apply (cut_tac x="-x" in sin_zero_lemma, simp, simp, safe) +apply (force simp add: minus_equation_iff [of x]) +apply (auto simp add: even_mult_two_ex) +done + + +subsection{*Tangent*} + +lemma tan_zero [simp]: "tan 0 = 0" +by (simp add: tan_def) + +lemma tan_pi [simp]: "tan pi = 0" +by (simp add: tan_def) + +lemma tan_npi [simp]: "tan (real (n::nat) * pi) = 0" +by (simp add: tan_def) + +lemma tan_minus [simp]: "tan (-x) = - tan x" +by (simp add: tan_def minus_mult_left) + +lemma tan_periodic [simp]: "tan (x + 2*pi) = tan x" +by (simp add: tan_def) + +lemma lemma_tan_add1: + "[| cos x \ 0; cos y \ 0 |] + ==> 1 - tan(x)*tan(y) = cos (x + y)/(cos x * cos y)" +apply (unfold tan_def real_divide_def) +apply (auto simp del: inverse_mult_distrib simp add: inverse_mult_distrib [symmetric] mult_ac) +apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst]) +apply (auto simp del: inverse_mult_distrib simp add: mult_assoc left_diff_distrib cos_add) +done + +lemma add_tan_eq: + "[| cos x \ 0; cos y \ 0 |] + ==> tan x + tan y = sin(x + y)/(cos x * cos y)" +apply (unfold tan_def) +apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst]) +apply (auto simp add: mult_assoc left_distrib) +apply (simp (no_asm) add: sin_add) +done + +lemma tan_add: "[| cos x \ 0; cos y \ 0; cos (x + y) \ 0 |] + ==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))" +apply (simp (no_asm_simp) add: add_tan_eq lemma_tan_add1) +apply (simp add: tan_def) +done + +lemma tan_double: "[| cos x \ 0; cos (2 * x) \ 0 |] + ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))" +apply (insert tan_add [of x x]) +apply (simp add: mult_2 [symmetric]) +apply (auto simp add: numeral_2_eq_2) +done + +lemma tan_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < tan x" +apply (unfold tan_def real_divide_def) +apply (auto intro!: sin_gt_zero2 cos_gt_zero_pi intro!: real_mult_order positive_imp_inverse_positive) +done + +lemma tan_less_zero: + assumes lb: "- pi/2 < x" and "x < 0" shows "tan x < 0" +proof - + have "0 < tan (- x)" using prems by (simp only: tan_gt_zero) + thus ?thesis by simp +qed + +lemma lemma_DERIV_tan: + "cos x \ 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse((cos x)\)" +apply (rule lemma_DERIV_subst) +apply (best intro!: DERIV_intros intro: DERIV_chain2) +apply (auto simp add: real_divide_def numeral_2_eq_2) +done + +lemma DERIV_tan [simp]: "cos x \ 0 ==> DERIV tan x :> inverse((cos x)\)" +by (auto dest: lemma_DERIV_tan simp add: tan_def [symmetric]) + +lemma LIM_cos_div_sin [simp]: "(%x. cos(x)/sin(x)) -- pi/2 --> 0" +apply (unfold real_divide_def) +apply (subgoal_tac "(\x. cos x * inverse (sin x)) -- pi * inverse 2 --> 0*1") +apply simp +apply (rule LIM_mult2) +apply (rule_tac [2] inverse_1 [THEN subst]) +apply (rule_tac [2] LIM_inverse) +apply (simp_all add: divide_inverse [symmetric]) +apply (simp_all only: isCont_def [symmetric] cos_pi_half [symmetric] sin_pi_half [symmetric]) +apply (blast intro!: DERIV_isCont DERIV_sin DERIV_cos)+ +done + +lemma lemma_tan_total: "0 < y ==> \x. 0 < x & x < pi/2 & y < tan x" +apply (cut_tac LIM_cos_div_sin) +apply (simp only: LIM_def) +apply (drule_tac x = "inverse y" in spec, safe, force) +apply (drule_tac ?d1.0 = s in pi_half_gt_zero [THEN [2] real_lbound_gt_zero], safe) +apply (rule_tac x = " (pi/2) - e" in exI) +apply (simp (no_asm_simp)) +apply (drule_tac x = " (pi/2) - e" in spec) +apply (auto simp add: abs_eqI2 tan_def) +apply (rule inverse_less_iff_less [THEN iffD1]) +apply (auto simp add: real_divide_def) +apply (rule real_mult_order) +apply (subgoal_tac [3] "0 < sin e") +apply (subgoal_tac [3] "0 < cos e") +apply (auto intro: cos_gt_zero sin_gt_zero2 simp add: inverse_mult_distrib abs_mult) +apply (drule_tac a = "cos e" in positive_imp_inverse_positive) +apply (drule_tac x = "inverse (cos e) " in abs_eqI2) +apply (auto dest!: abs_eqI2 simp add: mult_ac) +done + +lemma tan_total_pos: "0 \ y ==> \x. 0 \ x & x < pi/2 & tan x = y" +apply (frule real_le_imp_less_or_eq, safe) + prefer 2 apply force +apply (drule lemma_tan_total, safe) +apply (cut_tac f = tan and a = 0 and b = x and y = y in IVT_objl) +apply (auto intro!: DERIV_tan [THEN DERIV_isCont]) +apply (drule_tac y = xa in order_le_imp_less_or_eq) +apply (auto dest: cos_gt_zero) +done + +lemma lemma_tan_total1: "\x. -(pi/2) < x & x < (pi/2) & tan x = y" +apply (cut_tac linorder_linear [of 0 y], safe) +apply (drule tan_total_pos) +apply (cut_tac [2] y="-y" in tan_total_pos, safe) +apply (rule_tac [3] x = "-x" in exI) +apply (auto intro!: exI) +done + +lemma tan_total: "EX! x. -(pi/2) < x & x < (pi/2) & tan x = y" +apply (cut_tac y = y in lemma_tan_total1, auto) +apply (cut_tac x = xa and y = y in linorder_less_linear, auto) +apply (subgoal_tac [2] "\z. y < z & z < xa & DERIV tan z :> 0") +apply (subgoal_tac "\z. xa < z & z < y & DERIV tan z :> 0") +apply (rule_tac [4] Rolle) +apply (rule_tac [2] Rolle) +apply (auto intro!: DERIV_tan DERIV_isCont exI + simp add: differentiable_def) +txt{*Now, simulate TRYALL*} +apply (rule_tac [!] DERIV_tan asm_rl) +apply (auto dest!: DERIV_unique [OF _ DERIV_tan] + simp add: cos_gt_zero_pi [THEN real_not_refl2, THEN not_sym]) +done + +lemma arcsin_pi: "[| -1 \ y; y \ 1 |] + ==> -(pi/2) \ arcsin y & arcsin y \ pi & sin(arcsin y) = y" +apply (drule sin_total, assumption) +apply (erule ex1E) +apply (unfold arcsin_def) +apply (rule someI2, blast) +apply (force intro: order_trans) +done + +lemma arcsin: "[| -1 \ y; y \ 1 |] + ==> -(pi/2) \ arcsin y & + arcsin y \ pi/2 & sin(arcsin y) = y" +apply (unfold arcsin_def) +apply (drule sin_total, assumption) +apply (fast intro: someI2) +done + +lemma sin_arcsin [simp]: "[| -1 \ y; y \ 1 |] ==> sin(arcsin y) = y" +by (blast dest: arcsin) + +lemma arcsin_bounded: + "[| -1 \ y; y \ 1 |] ==> -(pi/2) \ arcsin y & arcsin y \ pi/2" +by (blast dest: arcsin) + +lemma arcsin_lbound: "[| -1 \ y; y \ 1 |] ==> -(pi/2) \ arcsin y" +by (blast dest: arcsin) + +lemma arcsin_ubound: "[| -1 \ y; y \ 1 |] ==> arcsin y \ pi/2" +by (blast dest: arcsin) + +lemma arcsin_lt_bounded: + "[| -1 < y; y < 1 |] ==> -(pi/2) < arcsin y & arcsin y < pi/2" +apply (frule order_less_imp_le) +apply (frule_tac y = y in order_less_imp_le) +apply (frule arcsin_bounded) +apply (safe, simp) +apply (drule_tac y = "arcsin y" in order_le_imp_less_or_eq) +apply (drule_tac [2] y = "pi/2" in order_le_imp_less_or_eq, safe) +apply (drule_tac [!] f = sin in arg_cong, auto) +done + +lemma arcsin_sin: "[|-(pi/2) \ x; x \ pi/2 |] ==> arcsin(sin x) = x" +apply (unfold arcsin_def) +apply (rule some1_equality) +apply (rule sin_total, auto) +done + +lemma arcos: "[| -1 \ y; y \ 1 |] + ==> 0 \ arcos y & arcos y \ pi & cos(arcos y) = y" +apply (unfold arcos_def) +apply (drule cos_total, assumption) +apply (fast intro: someI2) +done + +lemma cos_arcos [simp]: "[| -1 \ y; y \ 1 |] ==> cos(arcos y) = y" +by (blast dest: arcos) + +lemma arcos_bounded: "[| -1 \ y; y \ 1 |] ==> 0 \ arcos y & arcos y \ pi" +by (blast dest: arcos) + +lemma arcos_lbound: "[| -1 \ y; y \ 1 |] ==> 0 \ arcos y" +by (blast dest: arcos) + +lemma arcos_ubound: "[| -1 \ y; y \ 1 |] ==> arcos y \ pi" +by (blast dest: arcos) + +lemma arcos_lt_bounded: "[| -1 < y; y < 1 |] + ==> 0 < arcos y & arcos y < pi" +apply (frule order_less_imp_le) +apply (frule_tac y = y in order_less_imp_le) +apply (frule arcos_bounded, auto) +apply (drule_tac y = "arcos y" in order_le_imp_less_or_eq) +apply (drule_tac [2] y = pi in order_le_imp_less_or_eq, auto) +apply (drule_tac [!] f = cos in arg_cong, auto) +done + +lemma arcos_cos: "[|0 \ x; x \ pi |] ==> arcos(cos x) = x" +apply (unfold arcos_def) +apply (auto intro!: some1_equality cos_total) +done + +lemma arcos_cos2: "[|x \ 0; -pi \ x |] ==> arcos(cos x) = -x" +apply (unfold arcos_def) +apply (auto intro!: some1_equality cos_total) +done + +lemma arctan [simp]: + "- (pi/2) < arctan y & arctan y < pi/2 & tan (arctan y) = y" +apply (cut_tac y = y in tan_total) +apply (unfold arctan_def) +apply (fast intro: someI2) +done + +lemma tan_arctan: "tan(arctan y) = y" +by auto + +lemma arctan_bounded: "- (pi/2) < arctan y & arctan y < pi/2" +by (auto simp only: arctan) + +lemma arctan_lbound: "- (pi/2) < arctan y" +by auto + +lemma arctan_ubound: "arctan y < pi/2" +by (auto simp only: arctan) + +lemma arctan_tan: + "[|-(pi/2) < x; x < pi/2 |] ==> arctan(tan x) = x" +apply (unfold arctan_def) +apply (rule some1_equality) +apply (rule tan_total, auto) +done + +lemma arctan_zero_zero [simp]: "arctan 0 = 0" +by (insert arctan_tan [of 0], simp) + +lemma cos_arctan_not_zero [simp]: "cos(arctan x) \ 0" +apply (auto simp add: cos_zero_iff) +apply (case_tac "n") +apply (case_tac [3] "n") +apply (cut_tac [2] y = x in arctan_ubound) +apply (cut_tac [4] y = x in arctan_lbound) +apply (auto simp add: real_of_nat_Suc left_distrib mult_less_0_iff) +done + +lemma tan_sec: "cos x \ 0 ==> 1 + tan(x) ^ 2 = inverse(cos x) ^ 2" +apply (rule power_inverse [THEN subst]) +apply (rule_tac c1 = "(cos x)\" in real_mult_right_cancel [THEN iffD1]) +apply (auto dest: realpow_not_zero + simp add: power_mult_distrib left_distrib realpow_divide tan_def + mult_assoc power_inverse [symmetric] + simp del: realpow_Suc) +done + +lemma lemma_sin_cos_eq [simp]: "sin (xa + 1 / 2 * real (Suc m) * pi) = + cos (xa + 1 / 2 * real (m) * pi)" +apply (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib, auto) +done + +lemma lemma_sin_cos_eq2 [simp]: "sin (xa + real (Suc m) * pi / 2) = + cos (xa + real (m) * pi / 2)" +apply (simp only: cos_add sin_add real_divide_def real_of_nat_Suc left_distrib right_distrib, auto) +done + +lemma DERIV_sin_add [simp]: "DERIV (%x. sin (x + k)) xa :> cos (xa + k)" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = sin and g = "%x. x + k" in DERIV_chain2) +apply (best intro!: DERIV_intros intro: DERIV_chain2)+ +apply (simp (no_asm)) +done + +(* which further simplifies to (- 1 ^ m) !! *) +lemma sin_cos_npi [simp]: "sin ((real m + 1/2) * pi) = cos (real m * pi)" +by (auto simp add: right_distrib sin_add left_distrib mult_ac) + +lemma sin_cos_npi2 [simp]: "sin (real (Suc (2 * n)) * pi / 2) = (-1) ^ n" +apply (cut_tac m = n in sin_cos_npi) +apply (simp only: real_of_nat_Suc left_distrib real_divide_def, auto) +done + +lemma cos_2npi [simp]: "cos (2 * real (n::nat) * pi) = 1" +by (simp add: cos_double mult_assoc power_add [symmetric] numeral_2_eq_2) + +lemma cos_3over2_pi [simp]: "cos (3 / 2 * pi) = 0" +apply (subgoal_tac "3/2 = (1+1 / 2::real)") +apply (simp only: left_distrib) +apply (auto simp add: cos_add mult_ac) +done + +lemma sin_2npi [simp]: "sin (2 * real (n::nat) * pi) = 0" +by (auto simp add: mult_assoc) + +lemma sin_3over2_pi [simp]: "sin (3 / 2 * pi) = - 1" +apply (subgoal_tac "3/2 = (1+1 / 2::real)") +apply (simp only: left_distrib) +apply (auto simp add: sin_add mult_ac) +done + +(*NEEDED??*) +lemma [simp]: "cos(x + 1 / 2 * real(Suc m) * pi) = -sin (x + 1 / 2 * real m * pi)" +apply (simp only: cos_add sin_add real_of_nat_Suc right_distrib left_distrib minus_mult_right, auto) +done + +(*NEEDED??*) +lemma [simp]: "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)" +apply (simp only: cos_add sin_add real_divide_def real_of_nat_Suc left_distrib right_distrib, auto) +done + +lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0" +by (simp only: cos_add sin_add real_divide_def real_of_nat_Suc left_distrib right_distrib, auto) + +lemma DERIV_cos_add [simp]: "DERIV (%x. cos (x + k)) xa :> - sin (xa + k)" +apply (rule lemma_DERIV_subst) +apply (rule_tac f = cos and g = "%x. x + k" in DERIV_chain2) +apply (best intro!: DERIV_intros intro: DERIV_chain2)+ +apply (simp (no_asm)) +done + +lemma isCont_cos [simp]: "isCont cos x" +by (rule DERIV_cos [THEN DERIV_isCont]) + +lemma isCont_sin [simp]: "isCont sin x" +by (rule DERIV_sin [THEN DERIV_isCont]) + +lemma isCont_exp [simp]: "isCont exp x" +by (rule DERIV_exp [THEN DERIV_isCont]) + +lemma sin_zero_abs_cos_one: "sin x = 0 ==> abs(cos x) = 1" +by (auto simp add: sin_zero_iff even_mult_two_ex) + +lemma exp_eq_one_iff [simp]: "(exp x = 1) = (x = 0)" +apply auto +apply (drule_tac f = ln in arg_cong, auto) +done + +lemma cos_one_sin_zero: "cos x = 1 ==> sin x = 0" +by (cut_tac x = x in sin_cos_squared_add3, auto) + + +lemma real_root_less_mono: "[| 0 \ x; x < y |] ==> root(Suc n) x < root(Suc n) y" +apply (frule order_le_less_trans, assumption) +apply (frule_tac n1 = n in real_root_pow_pos2 [THEN ssubst]) +apply (rotate_tac 1, assumption) +apply (frule_tac n1 = n in real_root_pow_pos [THEN ssubst]) +apply (rotate_tac 3, assumption) +apply (drule_tac y = "root (Suc n) y ^ Suc n" in order_less_imp_le) +apply (frule_tac n = n in real_root_pos_pos_le) +apply (frule_tac n = n in real_root_pos_pos) +apply (drule_tac x = "root (Suc n) x" and y = "root (Suc n) y" in realpow_increasing) +apply (assumption, assumption) +apply (drule_tac x = "root (Suc n) x" in order_le_imp_less_or_eq) +apply auto +apply (drule_tac f = "%x. x ^ (Suc n) " in arg_cong) +apply (auto simp add: real_root_pow_pos2 simp del: realpow_Suc) +done + +lemma real_root_le_mono: "[| 0 \ x; x \ y |] ==> root(Suc n) x \ root(Suc n) y" +apply (drule_tac y = y in order_le_imp_less_or_eq) +apply (auto dest: real_root_less_mono intro: order_less_imp_le) +done + +lemma real_root_less_iff [simp]: "[| 0 \ x; 0 \ y |] ==> (root(Suc n) x < root(Suc n) y) = (x < y)" +apply (auto intro: real_root_less_mono) +apply (rule ccontr, drule linorder_not_less [THEN iffD1]) +apply (drule_tac x = y and n = n in real_root_le_mono, auto) +done + +lemma real_root_le_iff [simp]: "[| 0 \ x; 0 \ y |] ==> (root(Suc n) x \ root(Suc n) y) = (x \ y)" +apply (auto intro: real_root_le_mono) +apply (simp (no_asm) add: linorder_not_less [symmetric]) +apply auto +apply (drule_tac x = y and n = n in real_root_less_mono, auto) +done + +lemma real_root_eq_iff [simp]: "[| 0 \ x; 0 \ y |] ==> (root(Suc n) x = root(Suc n) y) = (x = y)" +apply (auto intro!: order_antisym) +apply (rule_tac n1 = n in real_root_le_iff [THEN iffD1]) +apply (rule_tac [4] n1 = n in real_root_le_iff [THEN iffD1], auto) +done + +lemma real_root_pos_unique: "[| 0 \ x; 0 \ y; y ^ (Suc n) = x |] ==> root (Suc n) x = y" +by (auto dest: real_root_pos2 simp del: realpow_Suc) + +lemma real_root_mult: "[| 0 \ x; 0 \ y |] + ==> root(Suc n) (x * y) = root(Suc n) x * root(Suc n) y" +apply (rule real_root_pos_unique) +apply (auto intro!: real_root_pos_pos_le simp add: power_mult_distrib zero_le_mult_iff real_root_pow_pos2 simp del: realpow_Suc) +done + +lemma real_root_inverse: "0 \ x ==> (root(Suc n) (inverse x) = inverse(root(Suc n) x))" +apply (rule real_root_pos_unique) +apply (auto intro: real_root_pos_pos_le simp add: power_inverse [symmetric] real_root_pow_pos2 simp del: realpow_Suc) +done + +lemma real_root_divide: + "[| 0 \ x; 0 \ y |] + ==> (root(Suc n) (x / y) = root(Suc n) x / root(Suc n) y)" +apply (unfold real_divide_def) +apply (auto simp add: real_root_mult real_root_inverse) +done + +lemma real_sqrt_less_mono: "[| 0 \ x; x < y |] ==> sqrt(x) < sqrt(y)" +apply (unfold sqrt_def) +apply (auto intro: real_root_less_mono) +done + +lemma real_sqrt_le_mono: "[| 0 \ x; x \ y |] ==> sqrt(x) \ sqrt(y)" +apply (unfold sqrt_def) +apply (auto intro: real_root_le_mono) +done + +lemma real_sqrt_less_iff [simp]: "[| 0 \ x; 0 \ y |] ==> (sqrt(x) < sqrt(y)) = (x < y)" +by (unfold sqrt_def, auto) + +lemma real_sqrt_le_iff [simp]: "[| 0 \ x; 0 \ y |] ==> (sqrt(x) \ sqrt(y)) = (x \ y)" +by (unfold sqrt_def, auto) + +lemma real_sqrt_eq_iff [simp]: "[| 0 \ x; 0 \ y |] ==> (sqrt(x) = sqrt(y)) = (x = y)" +by (unfold sqrt_def, auto) + +lemma real_sqrt_sos_less_one_iff [simp]: "(sqrt(x\ + y\) < 1) = (x\ + y\ < 1)" +apply (rule real_sqrt_one [THEN subst], safe) +apply (rule_tac [2] real_sqrt_less_mono) +apply (drule real_sqrt_less_iff [THEN [2] rev_iffD1], auto) +done + +lemma real_sqrt_sos_eq_one_iff [simp]: "(sqrt(x\ + y\) = 1) = (x\ + y\ = 1)" +apply (rule real_sqrt_one [THEN subst], safe) +apply (drule real_sqrt_eq_iff [THEN [2] rev_iffD1], auto) +done + +lemma real_divide_square_eq [simp]: "(((r::real) * a) / (r * r)) = a / r" +apply (unfold real_divide_def) +apply (case_tac "r=0") +apply (auto simp add: inverse_mult_distrib mult_ac) +done + + +subsection{*Theorems About Sqrt, Transcendental Functions for Complex*} + +lemma lemma_real_divide_sqrt: + "0 < x ==> 0 \ x/(sqrt (x * x + y * y))" +apply (unfold real_divide_def) +apply (rule real_mult_order [THEN order_less_imp_le], assumption) +apply (subgoal_tac "0 < inverse (sqrt (x\ + y\))") + apply (simp add: numeral_2_eq_2) +apply (simp add: real_sqrt_sum_squares_ge1 [THEN [2] order_less_le_trans]) +done + +lemma lemma_real_divide_sqrt_ge_minus_one: + "0 < x ==> -1 \ x/(sqrt (x * x + y * y))" +apply (rule real_le_trans) +apply (rule_tac [2] lemma_real_divide_sqrt, auto) +done + +lemma real_sqrt_sum_squares_gt_zero1: "x < 0 ==> 0 < sqrt (x * x + y * y)" +apply (rule real_sqrt_gt_zero) +apply (subgoal_tac "0 < x*x & 0 \ y*y", arith) +apply (auto simp add: zero_less_mult_iff) +done + +lemma real_sqrt_sum_squares_gt_zero2: "0 < x ==> 0 < sqrt (x * x + y * y)" +apply (rule real_sqrt_gt_zero) +apply (subgoal_tac "0 < x*x & 0 \ y*y", arith) +apply (auto simp add: zero_less_mult_iff) +done + +lemma real_sqrt_sum_squares_gt_zero3: "x \ 0 ==> 0 < sqrt(x\ + y\)" +apply (cut_tac x = x and y = 0 in linorder_less_linear) +apply (auto intro: real_sqrt_sum_squares_gt_zero2 real_sqrt_sum_squares_gt_zero1 simp add: numeral_2_eq_2) +done + +lemma real_sqrt_sum_squares_gt_zero3a: "y \ 0 ==> 0 < sqrt(x\ + y\)" +apply (drule_tac y = x in real_sqrt_sum_squares_gt_zero3) +apply (auto simp add: real_add_commute) +done + +lemma real_sqrt_sum_squares_eq_cancel [simp]: "sqrt(x\ + y\) = x ==> y = 0" +by (drule_tac f = "%x. x\" in arg_cong, auto) + +lemma real_sqrt_sum_squares_eq_cancel2 [simp]: "sqrt(x\ + y\) = y ==> x = 0" +apply (rule_tac x = y in real_sqrt_sum_squares_eq_cancel) +apply (simp add: real_add_commute) +done + +lemma lemma_real_divide_sqrt_le_one: "x < 0 ==> x/(sqrt (x * x + y * y)) \ 1" +by (insert lemma_real_divide_sqrt_ge_minus_one [of "-x" y], simp) + +lemma lemma_real_divide_sqrt_ge_minus_one2: + "x < 0 ==> -1 \ x/(sqrt (x * x + y * y))" +apply (case_tac "y = 0", auto) +apply (frule abs_minus_eqI2) +apply (auto simp add: inverse_minus_eq) +apply (rule order_less_imp_le) +apply (rule_tac z1 = "sqrt (x * x + y * y) " in real_mult_less_iff1 [THEN iffD1]) +apply (frule_tac [2] y2 = y in + real_sqrt_sum_squares_gt_zero1 [THEN real_not_refl2, THEN not_sym]) +apply (auto intro: real_sqrt_sum_squares_gt_zero1 simp add: mult_ac) +apply (cut_tac x = "-x" and y = y in real_sqrt_sum_squares_ge1) +apply (drule order_le_less [THEN iffD1], safe) +apply (simp add: numeral_2_eq_2) +apply (drule sym [THEN real_sqrt_sum_squares_eq_cancel], simp) +done + +lemma lemma_real_divide_sqrt_le_one2: "0 < x ==> x/(sqrt (x * x + y * y)) \ 1" +by (cut_tac x = "-x" and y = y in lemma_real_divide_sqrt_ge_minus_one2, auto) + + +lemma cos_x_y_ge_minus_one: "-1 \ x / sqrt (x * x + y * y)" +apply (cut_tac x = x and y = 0 in linorder_less_linear, safe) +apply (rule lemma_real_divide_sqrt_ge_minus_one2) +apply (rule_tac [3] lemma_real_divide_sqrt_ge_minus_one, auto) +done + +lemma cos_x_y_ge_minus_one1a [simp]: "-1 \ y / sqrt (x * x + y * y)" +apply (cut_tac x = y and y = x in cos_x_y_ge_minus_one) +apply (simp add: real_add_commute) +done + +lemma cos_x_y_le_one [simp]: "x / sqrt (x * x + y * y) \ 1" +apply (cut_tac x = x and y = 0 in linorder_less_linear, safe) +apply (rule lemma_real_divide_sqrt_le_one) +apply (rule_tac [3] lemma_real_divide_sqrt_le_one2, auto) +done + +lemma cos_x_y_le_one2 [simp]: "y / sqrt (x * x + y * y) \ 1" +apply (cut_tac x = y and y = x in cos_x_y_le_one) +apply (simp add: real_add_commute) +done + +declare cos_arcos [OF cos_x_y_ge_minus_one cos_x_y_le_one, simp] +declare arcos_bounded [OF cos_x_y_ge_minus_one cos_x_y_le_one, simp] + +declare cos_arcos [OF cos_x_y_ge_minus_one1a cos_x_y_le_one2, simp] +declare arcos_bounded [OF cos_x_y_ge_minus_one1a cos_x_y_le_one2, simp] + +lemma cos_abs_x_y_ge_minus_one [simp]: + "-1 \ \x\ / sqrt (x * x + y * y)" +apply (cut_tac x = x and y = 0 in linorder_less_linear) +apply (auto simp add: abs_minus_eqI2 abs_eqI2) +apply (drule lemma_real_divide_sqrt_ge_minus_one, force) +done + +lemma cos_abs_x_y_le_one [simp]: "\x\ / sqrt (x * x + y * y) \ 1" +apply (cut_tac x = x and y = 0 in linorder_less_linear) +apply (auto simp add: abs_minus_eqI2 abs_eqI2) +apply (drule lemma_real_divide_sqrt_ge_minus_one2, force) +done + +declare cos_arcos [OF cos_abs_x_y_ge_minus_one cos_abs_x_y_le_one, simp] +declare arcos_bounded [OF cos_abs_x_y_ge_minus_one cos_abs_x_y_le_one, simp] + +lemma minus_pi_less_zero: "-pi < 0" +apply (simp (no_asm)) +done +declare minus_pi_less_zero [simp] +declare minus_pi_less_zero [THEN order_less_imp_le, simp] + +lemma arcos_ge_minus_pi: "[| -1 \ y; y \ 1 |] ==> -pi \ arcos y" +apply (rule real_le_trans) +apply (rule_tac [2] arcos_lbound, auto) +done + +declare arcos_ge_minus_pi [OF cos_x_y_ge_minus_one cos_x_y_le_one, simp] + +(* How tedious! *) +lemma lemma_divide_rearrange: + "[| x + (y::real) \ 0; 1 - z = x/(x + y) |] ==> z = y/(x + y)" +apply (rule_tac c1 = "x + y" in real_mult_right_cancel [THEN iffD1]) +apply (frule_tac [2] c1 = "x + y" in real_mult_right_cancel [THEN iffD2]) +prefer 2 apply assumption +apply (rotate_tac [2] 2) +apply (drule_tac [2] mult_assoc [THEN subst]) +apply (rotate_tac [2] 2) +apply (frule_tac [2] left_inverse [THEN subst]) +prefer 2 apply assumption +apply (erule_tac [2] V = " (1 - z) * (x + y) = x / (x + y) * (x + y) " in thin_rl) +apply (erule_tac [2] V = "1 - z = x / (x + y) " in thin_rl) +apply (auto simp add: mult_assoc) +apply (auto simp add: right_distrib left_diff_distrib) +done + +lemma lemma_cos_sin_eq: + "[| 0 < x * x + y * y; + 1 - (sin xa)\ = (x / sqrt (x * x + y * y)) ^ 2 |] + ==> (sin xa)\ = (y / sqrt (x * x + y * y)) ^ 2" +by (auto intro: lemma_divide_rearrange + simp add: realpow_divide power2_eq_square [symmetric]) + + +lemma lemma_sin_cos_eq: + "[| 0 < x * x + y * y; + 1 - (cos xa)\ = (y / sqrt (x * x + y * y)) ^ 2 |] + ==> (cos xa)\ = (x / sqrt (x * x + y * y)) ^ 2" +apply (auto simp add: realpow_divide power2_eq_square [symmetric]) +apply (rule add_commute [THEN subst]) +apply (rule lemma_divide_rearrange, simp) +apply (simp add: add_commute) +done + +lemma sin_x_y_disj: + "[| x \ 0; + cos xa = x / sqrt (x * x + y * y) + |] ==> sin xa = y / sqrt (x * x + y * y) | + sin xa = - y / sqrt (x * x + y * y)" +apply (drule_tac f = "%x. x\" in arg_cong) +apply (frule_tac y = y in real_sum_square_gt_zero) +apply (simp add: cos_squared_eq) +apply (subgoal_tac "(sin xa)\ = (y / sqrt (x * x + y * y)) ^ 2") +apply (rule_tac [2] lemma_cos_sin_eq) +apply (auto simp add: realpow_two_disj numeral_2_eq_2 simp del: realpow_Suc) +done + +lemma lemma_cos_not_eq_zero: "x \ 0 ==> x / sqrt (x * x + y * y) \ 0" +apply (unfold real_divide_def) +apply (frule_tac y3 = y in real_sqrt_sum_squares_gt_zero3 [THEN real_not_refl2, THEN not_sym, THEN nonzero_imp_inverse_nonzero]) +apply (auto simp add: power2_eq_square) +done + +lemma cos_x_y_disj: "[| x \ 0; + sin xa = y / sqrt (x * x + y * y) + |] ==> cos xa = x / sqrt (x * x + y * y) | + cos xa = - x / sqrt (x * x + y * y)" +apply (drule_tac f = "%x. x\" in arg_cong) +apply (frule_tac y = y in real_sum_square_gt_zero) +apply (simp add: sin_squared_eq del: realpow_Suc) +apply (subgoal_tac "(cos xa)\ = (x / sqrt (x * x + y * y)) ^ 2") +apply (rule_tac [2] lemma_sin_cos_eq) +apply (auto simp add: realpow_two_disj numeral_2_eq_2 simp del: realpow_Suc) +done + +lemma real_sqrt_divide_less_zero: "0 < y ==> - y / sqrt (x * x + y * y) < 0" +apply (case_tac "x = 0") +apply (auto simp add: abs_eqI2) +apply (drule_tac y = y in real_sqrt_sum_squares_gt_zero3) +apply (auto simp add: zero_less_mult_iff real_divide_def power2_eq_square) +done + +lemma polar_ex1: "[| x \ 0; 0 < y |] ==> \r a. x = r * cos a & y = r * sin a" +apply (rule_tac x = "sqrt (x\ + y\) " in exI) +apply (rule_tac x = "arcos (x / sqrt (x * x + y * y))" in exI) +apply auto +apply (drule_tac y2 = y in real_sqrt_sum_squares_gt_zero3 [THEN real_not_refl2, THEN not_sym]) +apply (auto simp add: power2_eq_square) +apply (unfold arcos_def) +apply (cut_tac x1 = x and y1 = y + in cos_total [OF cos_x_y_ge_minus_one cos_x_y_le_one]) +apply (rule someI2_ex, blast) +apply (erule_tac V = "EX! xa. 0 \ xa & xa \ pi & cos xa = x / sqrt (x * x + y * y) " in thin_rl) +apply (frule sin_x_y_disj, blast) +apply (drule_tac y2 = y in real_sqrt_sum_squares_gt_zero3 [THEN real_not_refl2, THEN not_sym]) +apply (auto simp add: power2_eq_square) +apply (drule sin_ge_zero, assumption) +apply (drule_tac x = x in real_sqrt_divide_less_zero, auto) +done + +lemma real_sum_squares_cancel2a: "x * x = -(y * y) ==> y = (0::real)" +by (auto intro: real_sum_squares_cancel) + +lemma polar_ex2: "[| x \ 0; y < 0 |] ==> \r a. x = r * cos a & y = r * sin a" +apply (cut_tac x = 0 and y = x in linorder_less_linear, auto) +apply (rule_tac x = "sqrt (x\ + y\) " in exI) +apply (rule_tac x = "arcsin (y / sqrt (x * x + y * y))" in exI) +apply (auto dest: real_sum_squares_cancel2a simp add: power2_eq_square) +apply (unfold arcsin_def) +apply (cut_tac x1 = x and y1 = y + in sin_total [OF cos_x_y_ge_minus_one1a cos_x_y_le_one2]) +apply (rule someI2_ex, blast) +apply (erule_tac V = "EX! xa. - (pi/2) \ xa & xa \ pi/2 & sin xa = y / sqrt (x * x + y * y) " in thin_rl) +apply (cut_tac x=x and y=y in cos_x_y_disj, simp, blast, auto) +apply (drule cos_ge_zero, force) +apply (drule_tac x = y in real_sqrt_divide_less_zero) +apply (auto simp add: real_add_commute) +apply (insert polar_ex1 [of x "-y"], simp, clarify) +apply (rule_tac x = r in exI) +apply (rule_tac x = "-a" in exI, simp) +done + +lemma polar_Ex: "\r a. x = r * cos a & y = r * sin a" +apply (case_tac "x = 0", auto) +apply (rule_tac x = y in exI) +apply (rule_tac x = "pi/2" in exI, auto) +apply (cut_tac x = 0 and y = y in linorder_less_linear, auto) +apply (rule_tac [2] x = x in exI) +apply (rule_tac [2] x = 0 in exI, auto) +apply (blast intro: polar_ex1 polar_ex2)+ +done + +lemma real_sqrt_ge_abs1 [simp]: "\x\ \ sqrt (x\ + y\)" +apply (rule_tac n = 1 in realpow_increasing) +apply (auto simp add: numeral_2_eq_2 [symmetric]) +done + +lemma real_sqrt_ge_abs2 [simp]: "\y\ \ sqrt (x\ + y\)" +apply (rule real_add_commute [THEN subst]) +apply (rule real_sqrt_ge_abs1) +done +declare real_sqrt_ge_abs1 [simp] real_sqrt_ge_abs2 [simp] + +lemma real_sqrt_two_gt_zero [simp]: "0 < sqrt 2" +by (auto intro: real_sqrt_gt_zero) + +lemma real_sqrt_two_ge_zero [simp]: "0 \ sqrt 2" +by (auto intro: real_sqrt_ge_zero) + +lemma real_sqrt_two_gt_one [simp]: "1 < sqrt 2" +apply (rule order_less_le_trans [of _ "7/5"], simp) +apply (rule_tac n = 1 in realpow_increasing) + prefer 3 apply (simp add: numeral_2_eq_2 [symmetric] del: realpow_Suc) +apply (simp_all add: numeral_2_eq_2) +done + +lemma lemma_real_divide_sqrt_less: "0 < u ==> u / sqrt 2 < u" +apply (rule_tac z1 = "inverse u" in real_mult_less_iff1 [THEN iffD1], auto) +apply (rule_tac z1 = "sqrt 2" in real_mult_less_iff1 [THEN iffD1], auto) +done + +lemma four_x_squared: + fixes x::real + shows "4 * x\ = (2 * x)\" +by (simp add: power2_eq_square) + + +text{*Needed for the infinitely close relation over the nonstandard + complex numbers*} +lemma lemma_sqrt_hcomplex_capprox: + "[| 0 < u; x < u/2; y < u/2; 0 \ x; 0 \ y |] ==> sqrt (x\ + y\) < u" +apply (rule_tac y = "u/sqrt 2" in order_le_less_trans) +apply (erule_tac [2] lemma_real_divide_sqrt_less) +apply (rule_tac n = 1 in realpow_increasing) +apply (auto simp add: real_0_le_divide_iff realpow_divide numeral_2_eq_2 [symmetric] + simp del: realpow_Suc) +apply (rule_tac t = "u\" in real_sum_of_halves [THEN subst]) +apply (rule add_mono) +apply (auto simp add: four_x_squared simp del: realpow_Suc intro: power_mono) +done + +declare real_sqrt_sum_squares_ge_zero [THEN abs_eqI1, simp] + + +subsection{*A Few Theorems Involving Ln, Derivatives, etc.*} + +lemma lemma_DERIV_ln: + "DERIV ln z :> l ==> DERIV (%x. exp (ln x)) z :> exp (ln z) * l" +by (erule DERIV_fun_exp) + +lemma STAR_exp_ln: "0 < z ==> ( *f* (%x. exp (ln x))) z = z" +apply (rule_tac z = z in eq_Abs_hypreal) +apply (auto simp add: starfun hypreal_zero_def hypreal_less) +done + +lemma hypreal_add_Infinitesimal_gt_zero: "[|e : Infinitesimal; 0 < x |] ==> 0 < hypreal_of_real x + e" +apply (rule_tac c1 = "-e" in add_less_cancel_right [THEN iffD1]) +apply (auto intro: Infinitesimal_less_SReal) +done + +lemma NSDERIV_exp_ln_one: "0 < z ==> NSDERIV (%x. exp (ln x)) z :> 1" +apply (unfold nsderiv_def NSLIM_def, auto) +apply (rule ccontr) +apply (subgoal_tac "0 < hypreal_of_real z + h") +apply (drule STAR_exp_ln) +apply (rule_tac [2] hypreal_add_Infinitesimal_gt_zero) +apply (subgoal_tac "h/h = 1") +apply (auto simp add: exp_ln_iff [symmetric] simp del: exp_ln_iff) +done + +lemma DERIV_exp_ln_one: "0 < z ==> DERIV (%x. exp (ln x)) z :> 1" +by (auto intro: NSDERIV_exp_ln_one simp add: NSDERIV_DERIV_iff [symmetric]) + +lemma lemma_DERIV_ln2: "[| 0 < z; DERIV ln z :> l |] ==> exp (ln z) * l = 1" +apply (rule DERIV_unique) +apply (rule lemma_DERIV_ln) +apply (rule_tac [2] DERIV_exp_ln_one, auto) +done + +lemma lemma_DERIV_ln3: "[| 0 < z; DERIV ln z :> l |] ==> l = 1/(exp (ln z))" +apply (rule_tac c1 = "exp (ln z) " in real_mult_left_cancel [THEN iffD1]) +apply (auto intro: lemma_DERIV_ln2) +done + +lemma lemma_DERIV_ln4: "[| 0 < z; DERIV ln z :> l |] ==> l = 1/z" +apply (rule_tac t = z in exp_ln_iff [THEN iffD2, THEN subst]) +apply (auto intro: lemma_DERIV_ln3) +done + +(* need to rename second isCont_inverse *) + +lemma isCont_inv_fun: "[| 0 < d; \z. \z - x\ \ d --> g(f(z)) = z; + \z. \z - x\ \ d --> isCont f z |] + ==> isCont g (f x)" +apply (simp (no_asm) add: isCont_iff LIM_def) +apply safe +apply (drule_tac ?d1.0 = r in real_lbound_gt_zero) +apply (assumption, safe) +apply (subgoal_tac "\z. \z - x\ \ e --> (g (f z) = z) ") +prefer 2 apply force +apply (subgoal_tac "\z. \z - x\ \ e --> isCont f z") +prefer 2 apply force +apply (drule_tac d = e in isCont_inj_range) +prefer 2 apply (assumption, assumption, safe) +apply (rule_tac x = ea in exI, auto) +apply (rotate_tac 4) +apply (drule_tac x = "f (x) + xa" in spec) +apply auto +apply (drule sym, auto, arith) +done + +lemma isCont_inv_fun_inv: + "[| 0 < d; + \z. \z - x\ \ d --> g(f(z)) = z; + \z. \z - x\ \ d --> isCont f z |] + ==> \e. 0 < e & + (\y. 0 < abs(y - f(x)) & abs(y - f(x)) < e --> f(g(y)) = y)" +apply (drule isCont_inj_range) +prefer 2 apply (assumption, assumption, auto) +apply (rule_tac x = e in exI, auto) +apply (rotate_tac 2) +apply (drule_tac x = y in spec, auto) +done + + +text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*} +lemma LIM_fun_gt_zero: "[| f -- c --> l; 0 < l |] + ==> \r. 0 < r & (\x. x \ c & \c - x\ < r --> 0 < f x)" +apply (auto simp add: LIM_def) +apply (drule_tac x = "l/2" in spec, safe, force) +apply (rule_tac x = s in exI) +apply (auto simp only: abs_interval_iff) +done + +lemma LIM_fun_less_zero: "[| f -- c --> l; l < 0 |] + ==> \r. 0 < r & (\x. x \ c & \c - x\ < r --> f x < 0)" +apply (auto simp add: LIM_def) +apply (drule_tac x = "-l/2" in spec, safe, force) +apply (rule_tac x = s in exI) +apply (auto simp only: abs_interval_iff) +done + + +lemma LIM_fun_not_zero: + "[| f -- c --> l; l \ 0 |] + ==> \r. 0 < r & (\x. x \ c & \c - x\ < r --> f x \ 0)" +apply (cut_tac x = l and y = 0 in linorder_less_linear, auto) +apply (drule LIM_fun_less_zero) +apply (drule_tac [3] LIM_fun_gt_zero, auto) +apply (rule_tac [!] x = r in exI, auto) +done + +ML +{* +val inverse_unique = thm "inverse_unique"; +val real_root_zero = thm "real_root_zero"; +val real_root_pow_pos = thm "real_root_pow_pos"; +val real_root_pow_pos2 = thm "real_root_pow_pos2"; +val real_root_pos = thm "real_root_pos"; +val real_root_pos2 = thm "real_root_pos2"; +val real_root_pos_pos = thm "real_root_pos_pos"; +val real_root_pos_pos_le = thm "real_root_pos_pos_le"; +val real_root_one = thm "real_root_one"; +val root_2_eq = thm "root_2_eq"; +val real_sqrt_zero = thm "real_sqrt_zero"; +val real_sqrt_one = thm "real_sqrt_one"; +val real_sqrt_pow2_iff = thm "real_sqrt_pow2_iff"; +val real_sqrt_pow2 = thm "real_sqrt_pow2"; +val real_sqrt_abs_abs = thm "real_sqrt_abs_abs"; +val real_pow_sqrt_eq_sqrt_pow = thm "real_pow_sqrt_eq_sqrt_pow"; +val real_pow_sqrt_eq_sqrt_abs_pow2 = thm "real_pow_sqrt_eq_sqrt_abs_pow2"; +val real_sqrt_pow_abs = thm "real_sqrt_pow_abs"; +val not_real_square_gt_zero = thm "not_real_square_gt_zero"; +val real_mult_self_eq_zero_iff = thm "real_mult_self_eq_zero_iff"; +val real_sqrt_gt_zero = thm "real_sqrt_gt_zero"; +val real_sqrt_ge_zero = thm "real_sqrt_ge_zero"; +val sqrt_eqI = thm "sqrt_eqI"; +val real_sqrt_mult_distrib = thm "real_sqrt_mult_distrib"; +val real_sqrt_mult_distrib2 = thm "real_sqrt_mult_distrib2"; +val real_sqrt_sum_squares_ge_zero = thm "real_sqrt_sum_squares_ge_zero"; +val real_sqrt_sum_squares_mult_ge_zero = thm "real_sqrt_sum_squares_mult_ge_zero"; +val real_sqrt_sum_squares_mult_squared_eq = thm "real_sqrt_sum_squares_mult_squared_eq"; +val real_sqrt_abs = thm "real_sqrt_abs"; +val real_sqrt_abs2 = thm "real_sqrt_abs2"; +val real_sqrt_pow2_gt_zero = thm "real_sqrt_pow2_gt_zero"; +val real_sqrt_not_eq_zero = thm "real_sqrt_not_eq_zero"; +val real_inv_sqrt_pow2 = thm "real_inv_sqrt_pow2"; +val real_sqrt_eq_zero_cancel = thm "real_sqrt_eq_zero_cancel"; +val real_sqrt_eq_zero_cancel_iff = thm "real_sqrt_eq_zero_cancel_iff"; +val real_sqrt_sum_squares_ge1 = thm "real_sqrt_sum_squares_ge1"; +val real_sqrt_sum_squares_ge2 = thm "real_sqrt_sum_squares_ge2"; +val real_sqrt_ge_one = thm "real_sqrt_ge_one"; +val summable_exp = thm "summable_exp"; +val summable_sin = thm "summable_sin"; +val summable_cos = thm "summable_cos"; +val exp_converges = thm "exp_converges"; +val sin_converges = thm "sin_converges"; +val cos_converges = thm "cos_converges"; +val powser_insidea = thm "powser_insidea"; +val powser_inside = thm "powser_inside"; +val diffs_minus = thm "diffs_minus"; +val diffs_equiv = thm "diffs_equiv"; +val less_add_one = thm "less_add_one"; +val termdiffs_aux = thm "termdiffs_aux"; +val termdiffs = thm "termdiffs"; +val exp_fdiffs = thm "exp_fdiffs"; +val sin_fdiffs = thm "sin_fdiffs"; +val sin_fdiffs2 = thm "sin_fdiffs2"; +val cos_fdiffs = thm "cos_fdiffs"; +val cos_fdiffs2 = thm "cos_fdiffs2"; +val DERIV_exp = thm "DERIV_exp"; +val DERIV_sin = thm "DERIV_sin"; +val DERIV_cos = thm "DERIV_cos"; +val exp_zero = thm "exp_zero"; +val exp_ge_add_one_self = thm "exp_ge_add_one_self"; +val exp_gt_one = thm "exp_gt_one"; +val DERIV_exp_add_const = thm "DERIV_exp_add_const"; +val DERIV_exp_minus = thm "DERIV_exp_minus"; +val DERIV_exp_exp_zero = thm "DERIV_exp_exp_zero"; +val exp_add_mult_minus = thm "exp_add_mult_minus"; +val exp_mult_minus = thm "exp_mult_minus"; +val exp_mult_minus2 = thm "exp_mult_minus2"; +val exp_minus = thm "exp_minus"; +val exp_add = thm "exp_add"; +val exp_ge_zero = thm "exp_ge_zero"; +val exp_not_eq_zero = thm "exp_not_eq_zero"; +val exp_gt_zero = thm "exp_gt_zero"; +val inv_exp_gt_zero = thm "inv_exp_gt_zero"; +val abs_exp_cancel = thm "abs_exp_cancel"; +val exp_real_of_nat_mult = thm "exp_real_of_nat_mult"; +val exp_diff = thm "exp_diff"; +val exp_less_mono = thm "exp_less_mono"; +val exp_less_cancel = thm "exp_less_cancel"; +val exp_less_cancel_iff = thm "exp_less_cancel_iff"; +val exp_le_cancel_iff = thm "exp_le_cancel_iff"; +val exp_inj_iff = thm "exp_inj_iff"; +val exp_total = thm "exp_total"; +val ln_exp = thm "ln_exp"; +val exp_ln_iff = thm "exp_ln_iff"; +val ln_mult = thm "ln_mult"; +val ln_inj_iff = thm "ln_inj_iff"; +val ln_one = thm "ln_one"; +val ln_inverse = thm "ln_inverse"; +val ln_div = thm "ln_div"; +val ln_less_cancel_iff = thm "ln_less_cancel_iff"; +val ln_le_cancel_iff = thm "ln_le_cancel_iff"; +val ln_realpow = thm "ln_realpow"; +val ln_add_one_self_le_self = thm "ln_add_one_self_le_self"; +val ln_less_self = thm "ln_less_self"; +val ln_ge_zero = thm "ln_ge_zero"; +val ln_gt_zero = thm "ln_gt_zero"; +val ln_not_eq_zero = thm "ln_not_eq_zero"; +val ln_less_zero = thm "ln_less_zero"; +val exp_ln_eq = thm "exp_ln_eq"; +val sin_zero = thm "sin_zero"; +val cos_zero = thm "cos_zero"; +val DERIV_sin_sin_mult = thm "DERIV_sin_sin_mult"; +val DERIV_sin_sin_mult2 = thm "DERIV_sin_sin_mult2"; +val DERIV_sin_realpow2 = thm "DERIV_sin_realpow2"; +val DERIV_sin_realpow2a = thm "DERIV_sin_realpow2a"; +val DERIV_cos_cos_mult = thm "DERIV_cos_cos_mult"; +val DERIV_cos_cos_mult2 = thm "DERIV_cos_cos_mult2"; +val DERIV_cos_realpow2 = thm "DERIV_cos_realpow2"; +val DERIV_cos_realpow2a = thm "DERIV_cos_realpow2a"; +val DERIV_cos_realpow2b = thm "DERIV_cos_realpow2b"; +val DERIV_cos_cos_mult3 = thm "DERIV_cos_cos_mult3"; +val DERIV_sin_circle_all = thm "DERIV_sin_circle_all"; +val DERIV_sin_circle_all_zero = thm "DERIV_sin_circle_all_zero"; +val sin_cos_squared_add = thm "sin_cos_squared_add"; +val sin_cos_squared_add2 = thm "sin_cos_squared_add2"; +val sin_cos_squared_add3 = thm "sin_cos_squared_add3"; +val sin_squared_eq = thm "sin_squared_eq"; +val cos_squared_eq = thm "cos_squared_eq"; +val real_gt_one_ge_zero_add_less = thm "real_gt_one_ge_zero_add_less"; +val abs_sin_le_one = thm "abs_sin_le_one"; +val sin_ge_minus_one = thm "sin_ge_minus_one"; +val sin_le_one = thm "sin_le_one"; +val abs_cos_le_one = thm "abs_cos_le_one"; +val cos_ge_minus_one = thm "cos_ge_minus_one"; +val cos_le_one = thm "cos_le_one"; +val DERIV_fun_pow = thm "DERIV_fun_pow"; +val DERIV_fun_exp = thm "DERIV_fun_exp"; +val DERIV_fun_sin = thm "DERIV_fun_sin"; +val DERIV_fun_cos = thm "DERIV_fun_cos"; +val DERIV_intros = thms "DERIV_intros"; +val sin_cos_add = thm "sin_cos_add"; +val sin_add = thm "sin_add"; +val cos_add = thm "cos_add"; +val sin_cos_minus = thm "sin_cos_minus"; +val sin_minus = thm "sin_minus"; +val cos_minus = thm "cos_minus"; +val sin_diff = thm "sin_diff"; +val sin_diff2 = thm "sin_diff2"; +val cos_diff = thm "cos_diff"; +val cos_diff2 = thm "cos_diff2"; +val sin_double = thm "sin_double"; +val cos_double = thm "cos_double"; +val sin_paired = thm "sin_paired"; +val sin_gt_zero = thm "sin_gt_zero"; +val sin_gt_zero1 = thm "sin_gt_zero1"; +val cos_double_less_one = thm "cos_double_less_one"; +val cos_paired = thm "cos_paired"; +val cos_two_less_zero = thm "cos_two_less_zero"; +val cos_is_zero = thm "cos_is_zero"; +val pi_half = thm "pi_half"; +val cos_pi_half = thm "cos_pi_half"; +val pi_half_gt_zero = thm "pi_half_gt_zero"; +val pi_half_less_two = thm "pi_half_less_two"; +val pi_gt_zero = thm "pi_gt_zero"; +val pi_neq_zero = thm "pi_neq_zero"; +val pi_not_less_zero = thm "pi_not_less_zero"; +val pi_ge_zero = thm "pi_ge_zero"; +val minus_pi_half_less_zero = thm "minus_pi_half_less_zero"; +val sin_pi_half = thm "sin_pi_half"; +val cos_pi = thm "cos_pi"; +val sin_pi = thm "sin_pi"; +val sin_cos_eq = thm "sin_cos_eq"; +val minus_sin_cos_eq = thm "minus_sin_cos_eq"; +val cos_sin_eq = thm "cos_sin_eq"; +val sin_periodic_pi = thm "sin_periodic_pi"; +val sin_periodic_pi2 = thm "sin_periodic_pi2"; +val cos_periodic_pi = thm "cos_periodic_pi"; +val sin_periodic = thm "sin_periodic"; +val cos_periodic = thm "cos_periodic"; +val cos_npi = thm "cos_npi"; +val sin_npi = thm "sin_npi"; +val sin_npi2 = thm "sin_npi2"; +val cos_two_pi = thm "cos_two_pi"; +val sin_two_pi = thm "sin_two_pi"; +val sin_gt_zero2 = thm "sin_gt_zero2"; +val sin_less_zero = thm "sin_less_zero"; +val pi_less_4 = thm "pi_less_4"; +val cos_gt_zero = thm "cos_gt_zero"; +val cos_gt_zero_pi = thm "cos_gt_zero_pi"; +val cos_ge_zero = thm "cos_ge_zero"; +val sin_gt_zero_pi = thm "sin_gt_zero_pi"; +val sin_ge_zero = thm "sin_ge_zero"; +val cos_total = thm "cos_total"; +val sin_total = thm "sin_total"; +val reals_Archimedean4 = thm "reals_Archimedean4"; +val cos_zero_lemma = thm "cos_zero_lemma"; +val sin_zero_lemma = thm "sin_zero_lemma"; +val cos_zero_iff = thm "cos_zero_iff"; +val sin_zero_iff = thm "sin_zero_iff"; +val tan_zero = thm "tan_zero"; +val tan_pi = thm "tan_pi"; +val tan_npi = thm "tan_npi"; +val tan_minus = thm "tan_minus"; +val tan_periodic = thm "tan_periodic"; +val add_tan_eq = thm "add_tan_eq"; +val tan_add = thm "tan_add"; +val tan_double = thm "tan_double"; +val tan_gt_zero = thm "tan_gt_zero"; +val tan_less_zero = thm "tan_less_zero"; +val DERIV_tan = thm "DERIV_tan"; +val LIM_cos_div_sin = thm "LIM_cos_div_sin"; +val tan_total_pos = thm "tan_total_pos"; +val tan_total = thm "tan_total"; +val arcsin_pi = thm "arcsin_pi"; +val arcsin = thm "arcsin"; +val sin_arcsin = thm "sin_arcsin"; +val arcsin_bounded = thm "arcsin_bounded"; +val arcsin_lbound = thm "arcsin_lbound"; +val arcsin_ubound = thm "arcsin_ubound"; +val arcsin_lt_bounded = thm "arcsin_lt_bounded"; +val arcsin_sin = thm "arcsin_sin"; +val arcos = thm "arcos"; +val cos_arcos = thm "cos_arcos"; +val arcos_bounded = thm "arcos_bounded"; +val arcos_lbound = thm "arcos_lbound"; +val arcos_ubound = thm "arcos_ubound"; +val arcos_lt_bounded = thm "arcos_lt_bounded"; +val arcos_cos = thm "arcos_cos"; +val arcos_cos2 = thm "arcos_cos2"; +val arctan = thm "arctan"; +val tan_arctan = thm "tan_arctan"; +val arctan_bounded = thm "arctan_bounded"; +val arctan_lbound = thm "arctan_lbound"; +val arctan_ubound = thm "arctan_ubound"; +val arctan_tan = thm "arctan_tan"; +val arctan_zero_zero = thm "arctan_zero_zero"; +val cos_arctan_not_zero = thm "cos_arctan_not_zero"; +val tan_sec = thm "tan_sec"; +val DERIV_sin_add = thm "DERIV_sin_add"; +val sin_cos_npi = thm "sin_cos_npi"; +val sin_cos_npi2 = thm "sin_cos_npi2"; +val cos_2npi = thm "cos_2npi"; +val cos_3over2_pi = thm "cos_3over2_pi"; +val sin_2npi = thm "sin_2npi"; +val sin_3over2_pi = thm "sin_3over2_pi"; +val cos_pi_eq_zero = thm "cos_pi_eq_zero"; +val DERIV_cos_add = thm "DERIV_cos_add"; +val isCont_cos = thm "isCont_cos"; +val isCont_sin = thm "isCont_sin"; +val isCont_exp = thm "isCont_exp"; +val sin_zero_abs_cos_one = thm "sin_zero_abs_cos_one"; +val exp_eq_one_iff = thm "exp_eq_one_iff"; +val cos_one_sin_zero = thm "cos_one_sin_zero"; +val real_root_less_mono = thm "real_root_less_mono"; +val real_root_le_mono = thm "real_root_le_mono"; +val real_root_less_iff = thm "real_root_less_iff"; +val real_root_le_iff = thm "real_root_le_iff"; +val real_root_eq_iff = thm "real_root_eq_iff"; +val real_root_pos_unique = thm "real_root_pos_unique"; +val real_root_mult = thm "real_root_mult"; +val real_root_inverse = thm "real_root_inverse"; +val real_root_divide = thm "real_root_divide"; +val real_sqrt_less_mono = thm "real_sqrt_less_mono"; +val real_sqrt_le_mono = thm "real_sqrt_le_mono"; +val real_sqrt_less_iff = thm "real_sqrt_less_iff"; +val real_sqrt_le_iff = thm "real_sqrt_le_iff"; +val real_sqrt_eq_iff = thm "real_sqrt_eq_iff"; +val real_sqrt_sos_less_one_iff = thm "real_sqrt_sos_less_one_iff"; +val real_sqrt_sos_eq_one_iff = thm "real_sqrt_sos_eq_one_iff"; +val real_divide_square_eq = thm "real_divide_square_eq"; +val real_sqrt_sum_squares_gt_zero1 = thm "real_sqrt_sum_squares_gt_zero1"; +val real_sqrt_sum_squares_gt_zero2 = thm "real_sqrt_sum_squares_gt_zero2"; +val real_sqrt_sum_squares_gt_zero3 = thm "real_sqrt_sum_squares_gt_zero3"; +val real_sqrt_sum_squares_gt_zero3a = thm "real_sqrt_sum_squares_gt_zero3a"; +val real_sqrt_sum_squares_eq_cancel = thm "real_sqrt_sum_squares_eq_cancel"; +val real_sqrt_sum_squares_eq_cancel2 = thm "real_sqrt_sum_squares_eq_cancel2"; +val cos_x_y_ge_minus_one = thm "cos_x_y_ge_minus_one"; +val cos_x_y_ge_minus_one1a = thm "cos_x_y_ge_minus_one1a"; +val cos_x_y_le_one = thm "cos_x_y_le_one"; +val cos_x_y_le_one2 = thm "cos_x_y_le_one2"; +val cos_abs_x_y_ge_minus_one = thm "cos_abs_x_y_ge_minus_one"; +val cos_abs_x_y_le_one = thm "cos_abs_x_y_le_one"; +val minus_pi_less_zero = thm "minus_pi_less_zero"; +val arcos_ge_minus_pi = thm "arcos_ge_minus_pi"; +val sin_x_y_disj = thm "sin_x_y_disj"; +val cos_x_y_disj = thm "cos_x_y_disj"; +val real_sqrt_divide_less_zero = thm "real_sqrt_divide_less_zero"; +val polar_ex1 = thm "polar_ex1"; +val real_sum_squares_cancel2a = thm "real_sum_squares_cancel2a"; +val polar_ex2 = thm "polar_ex2"; +val polar_Ex = thm "polar_Ex"; +val real_sqrt_ge_abs1 = thm "real_sqrt_ge_abs1"; +val real_sqrt_ge_abs2 = thm "real_sqrt_ge_abs2"; +val real_sqrt_two_gt_zero = thm "real_sqrt_two_gt_zero"; +val real_sqrt_two_ge_zero = thm "real_sqrt_two_ge_zero"; +val real_sqrt_two_gt_one = thm "real_sqrt_two_gt_one"; +val STAR_exp_ln = thm "STAR_exp_ln"; +val hypreal_add_Infinitesimal_gt_zero = thm "hypreal_add_Infinitesimal_gt_zero"; +val NSDERIV_exp_ln_one = thm "NSDERIV_exp_ln_one"; +val DERIV_exp_ln_one = thm "DERIV_exp_ln_one"; +val isCont_inv_fun = thm "isCont_inv_fun"; +val isCont_inv_fun_inv = thm "isCont_inv_fun_inv"; +val LIM_fun_gt_zero = thm "LIM_fun_gt_zero"; +val LIM_fun_less_zero = thm "LIM_fun_less_zero"; +val LIM_fun_not_zero = thm "LIM_fun_not_zero"; +*} end diff -r 4b3d280ef06a -r 89840837108e src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Jul 26 15:48:50 2004 +0200 +++ b/src/HOL/IsaMakefile Mon Jul 26 17:34:52 2004 +0200 @@ -154,8 +154,7 @@ Hyperreal/Lim.thy Hyperreal/Log.thy\ Hyperreal/MacLaurin_lemmas.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\ - Hyperreal/Star.thy Hyperreal/Transcendental.ML\ + Hyperreal/SEQ.ML Hyperreal/SEQ.thy Hyperreal/Series.thy Hyperreal/Star.thy \ Hyperreal/Transcendental.thy Hyperreal/fuf.ML Hyperreal/hypreal_arith.ML \ Complex/Complex_Main.thy Complex/CLim.thy Complex/CSeries.thy\ Complex/CStar.thy Complex/Complex.thy Complex/ComplexBin.thy\ diff -r 4b3d280ef06a -r 89840837108e src/HOL/Real/RealDef.thy --- a/src/HOL/Real/RealDef.thy Mon Jul 26 15:48:50 2004 +0200 +++ b/src/HOL/Real/RealDef.thy Mon Jul 26 17:34:52 2004 +0200 @@ -798,6 +798,7 @@ subsubsection{*Density of the Reals*} +(*????FIXME: rename d1, d2 to x, y*) lemma real_lbound_gt_zero: "[| (0::real) < d1; 0 < d2 |] ==> \e. 0 < e & e < d1 & e < d2" apply (rule_tac x = " (min d1 d2) /2" in exI) diff -r 4b3d280ef06a -r 89840837108e src/HOL/Ring_and_Field.thy --- a/src/HOL/Ring_and_Field.thy Mon Jul 26 15:48:50 2004 +0200 +++ b/src/HOL/Ring_and_Field.thy Mon Jul 26 17:34:52 2004 +0200 @@ -729,6 +729,15 @@ thus ?thesis by simp qed +lemma inverse_unique: + assumes ab: "a*b = 1" + shows "inverse a = (b::'a::field)" +proof - + have "a \ 0" using ab by auto + moreover have "inverse a * (a * b) = inverse a" by (simp add: ab) + ultimately show ?thesis by (simp add: mult_assoc [symmetric]) +qed + lemma nonzero_inverse_mult_distrib: assumes anz: "a \ 0" and bnz: "b \ 0"