new theories from Jacques Fleuriot
authorpaulson
Thu, 15 Nov 2001 16:12:49 +0100
changeset 12196 a3be6b3a9c0b
parent 12195 ed2893765a08
child 12197 d9320fb0a570
new theories from Jacques Fleuriot
src/HOL/GroupTheory/Exponent.ML
src/HOL/GroupTheory/Sylow.ML
src/HOL/Hyperreal/EvenOdd.ML
src/HOL/Hyperreal/EvenOdd.thy
src/HOL/Hyperreal/ExtraThms2.ML
src/HOL/Hyperreal/ExtraThms2.thy
src/HOL/Hyperreal/Fact.ML
src/HOL/Hyperreal/Fact.thy
src/HOL/Hyperreal/Hyperreal.thy
src/HOL/Hyperreal/NthRoot.ML
src/HOL/Hyperreal/NthRoot.thy
src/HOL/Hyperreal/Transcendental.ML
src/HOL/Hyperreal/Transcendental.thy
src/HOL/Integ/IntPower.ML
src/HOL/Integ/nat_bin.ML
src/HOL/Integ/nat_simprocs.ML
src/HOL/IsaMakefile
src/HOL/Power.ML
src/HOL/Real/RealAbs.ML
src/HOL/Real/RealPow.ML
src/HOL/ex/NatSum.thy
--- a/src/HOL/GroupTheory/Exponent.ML	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/GroupTheory/Exponent.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -27,7 +27,6 @@
 qed "prime_iff";
 
 Goal "p\\<in>prime ==> 0 < p^a";
-by (rtac zero_less_power 1);
 by (force_tac (claset(), simpset() addsimps [prime_iff]) 1);
 qed "zero_less_prime_power";
 
@@ -443,9 +442,8 @@
 Goal "0 < m ==> exponent p (((p^a) * m) choose p^a) = exponent p m";
 by (case_tac "p \\<in> prime" 1);
 by (Asm_simp_tac 2);
-by (forw_inst_tac [("a","a")] zero_less_prime_power 1);
 by (subgoal_tac "0 < p^a * m & p^a <= p^a * m" 1);
-by (Asm_full_simp_tac 2);
+by (force_tac (claset(), simpset() addsimps [prime_iff]) 2);
 (*A similar trick to the one used in p_not_div_choose_lemma:
   insert an equation; use exponent_mult_add on the LHS; on the RHS, first
   transform the binomial coefficient, then use exponent_mult_add.*)
@@ -461,6 +459,6 @@
 by (Asm_full_simp_tac 1);
 by (asm_simp_tac (simpset() addsimps [zero_less_binomial_iff]) 1);
 by (arith_tac 1);
-by (asm_simp_tac (simpset() delsimps bad_Sucs
+by (asm_full_simp_tac (simpset() delsimps bad_Sucs
                          addsimps [exponent_mult_add, const_p_fac_right]) 1);
 qed "const_p_fac";  
--- a/src/HOL/GroupTheory/Sylow.ML	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/GroupTheory/Sylow.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -73,7 +73,8 @@
 
 Goal "\\<exists>x. x\\<in>M1";
 by (rtac (card_nonempty RS NotEmpty_ExEl) 1);
-by (simp_tac (simpset() addsimps [card_M1, zero_less_prime_power, prime_p]) 1);
+by (cut_facts_tac [prime_p RS prime_imp_one_less] 1);
+by (asm_simp_tac (simpset() addsimps [card_M1]) 1);
 qed "exists_x_in_M1";
 
 Goal "M1 <= carrier G";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/EvenOdd.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -0,0 +1,390 @@
+(*  Title       : Even.ML
+    Author      : Jacques D. Fleuriot  
+    Copyright   : 1999  University of Edinburgh
+    Description : Even numbers defined
+*)
+
+Goal "even(Suc(Suc n)) = (even(n))";
+by (Auto_tac);
+qed "even_Suc_Suc";
+Addsimps [even_Suc_Suc];
+
+Goal "(even(n)) = (~odd(n))";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "even_not_odd";
+
+Goal "(odd(n)) = (~even(n))";
+by (simp_tac (simpset() addsimps [even_not_odd]) 1);
+qed "odd_not_even";
+
+Goal "even(2*n)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "even_mult_two";
+Addsimps [even_mult_two];
+
+Goal "even(n*2)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "even_mult_two'";
+Addsimps [even_mult_two'];
+
+Goal "even(n + n)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "even_sum_self";
+Addsimps [even_sum_self];
+
+Goal "~odd(2*n)";
+by (induct_tac "n" 1);
+by Auto_tac;
+qed "not_odd_self_mult2";
+Addsimps [not_odd_self_mult2];
+
+Goal "~odd(n + n)";
+by (induct_tac "n" 1);
+by Auto_tac;
+qed "not_odd_sum_self";
+Addsimps [not_odd_sum_self];
+
+Goal "~even(Suc(n + n))";
+by (induct_tac "n" 1);
+by Auto_tac;
+qed "not_even_Suc_sum_self";
+Addsimps [not_even_Suc_sum_self];
+
+Goal "odd(Suc(2*n))";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "odd_mult_two_add_one";
+Addsimps [odd_mult_two_add_one];
+
+Goal "odd(Suc(n + n))";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "odd_sum_Suc_self";
+Addsimps [odd_sum_Suc_self];
+
+Goal "even(Suc n) = odd(n)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "even_Suc_odd_iff";
+
+Goal "odd(Suc n) = even(n)";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "odd_Suc_even_iff";
+
+Goal "even n | odd n";
+by (simp_tac (simpset() addsimps [even_not_odd]) 1);
+qed "even_odd_disj";
+
+(* could be proved automatically before: spoiled by numeral arith! *)
+Goal "EX m. (n = 2*m | n = Suc(2*m))";
+by (induct_tac "n" 1 THEN Auto_tac);
+by (res_inst_tac [("x","Suc m")] exI 1 THEN Auto_tac);
+qed "nat_mult_two_Suc_disj";
+
+Goal "even(n) = (EX m. n = 2*m)";
+by (cut_inst_tac [("n","n")] nat_mult_two_Suc_disj 1);
+by (Auto_tac);
+qed "even_mult_two_ex";
+
+Goal "odd(n) = (EX m. n = Suc (2*m))";
+by (cut_inst_tac [("n","n")] nat_mult_two_Suc_disj 1);
+by (Auto_tac);
+qed "odd_Suc_mult_two_ex";
+
+Goal "even(n) ==> even(m*n)";
+by (auto_tac (claset(),
+              simpset() addsimps [add_mult_distrib2, even_mult_two_ex]));
+qed "even_mult_even";
+Addsimps [even_mult_even];
+
+Goal "(m*2) div 2 = (m::nat)";
+by (rtac div_mult_self_is_m 1);
+by (Simp_tac 1);
+qed "div_mult_self_two_is_m";
+
+Goal "(m + m) div 2 = (m::nat)";
+by (asm_full_simp_tac (simpset() addsimps [m_add_m_eq2]) 1); 
+qed "div_add_self_two_is_m";
+Addsimps [div_add_self_two_is_m];
+
+Goal "((m + m) + 2) div 2 = Suc m";
+by (rtac (div_add_self2 RS ssubst) 1);
+by (Auto_tac);
+qed "div_add_self_two";
+
+Goal "Suc(Suc(m*2)) div 2 = Suc m";
+by (cut_inst_tac [("m","m")] div_add_self_two 1);
+by (auto_tac (claset(),simpset() addsimps [m_add_m_eq2, mult_commute]));
+qed "div_mult_self_Suc_Suc";
+Addsimps [div_mult_self_Suc_Suc];
+
+Goal "Suc(Suc(2*m)) div 2 = Suc m";
+by (auto_tac (claset(),simpset() addsimps [mult_commute]));
+qed "div_mult_self_Suc_Suc2";
+Addsimps [div_mult_self_Suc_Suc2];
+
+Goal "Suc(Suc(m + m)) div 2 = Suc m";
+by (cut_inst_tac [("m","m")] div_mult_self_Suc_Suc 1);
+by (auto_tac (claset(),simpset() addsimps [m_add_m_eq2, mult_commute]));
+qed "div_add_self_Suc_Suc";
+Addsimps [div_add_self_Suc_Suc];
+
+Goal "~ even n ==> (Suc n) div 2 = Suc((n - 1) div 2)";
+by (auto_tac (claset(),simpset() addsimps [odd_not_even RS sym,
+    odd_Suc_mult_two_ex]));
+qed "not_even_imp_Suc_Suc_diff_one_eq";
+Addsimps [not_even_imp_Suc_Suc_diff_one_eq];
+
+Goal "even(m + n) = (even m = even n)";
+by (induct_tac "m" 1);
+by Auto_tac;
+qed "even_add";
+AddIffs [even_add];
+
+Goal "even(m * n) = (even m | even n)";
+by (induct_tac "m" 1);
+by Auto_tac;
+qed "even_mult";
+
+Goal "even (m ^ n) = (even m & n ~= 0)";
+by (induct_tac "n" 1);
+by (auto_tac (claset(),simpset() addsimps [even_mult]));
+qed "even_pow";
+AddIffs [even_pow];
+
+Goal "odd(m + n) = (odd m ~= odd n)";
+by (induct_tac "m" 1);
+by Auto_tac;
+qed "odd_add";
+AddIffs [odd_add];
+
+Goal "odd(m * n) = (odd m & odd n)";
+by (induct_tac "m" 1);
+by Auto_tac;
+qed "odd_mult";
+AddIffs [odd_mult];
+
+Goal "odd (m ^ n) = (odd m | n = 0)";
+by (induct_tac "n" 1);
+by Auto_tac;
+qed "odd_pow";
+
+Goal "0 < n --> ~even (n + n - 1)";
+by (induct_tac "n" 1);
+by Auto_tac;
+qed_spec_mp "not_even_2n_minus_1";
+Addsimps [not_even_2n_minus_1];
+
+Goal "Suc (2 * m) mod 2 = 1";
+by (induct_tac "m" 1);
+by (auto_tac (claset(),simpset() addsimps [mod_Suc]));
+qed "mod_two_Suc_2m";
+Addsimps [mod_two_Suc_2m];
+
+Goal "(Suc (Suc (2 * m)) div 2) = Suc m";
+by (rtac (div_if RS ssubst) 1 THEN Auto_tac);
+qed "div_two_Suc_Suc_2m";
+Addsimps [div_two_Suc_Suc_2m];
+
+Goal "even n ==> 2 * ((n + 1) div 2) = n";
+by (auto_tac (claset(),simpset() addsimps [mult_div_cancel,
+    even_mult_two_ex]));
+qed "lemma_even_div";
+Addsimps [lemma_even_div];
+
+Goal "~even n ==> 2 * ((n + 1) div 2) = Suc n";
+by (auto_tac (claset(),simpset() addsimps [even_not_odd,
+    odd_Suc_mult_two_ex]));
+qed "lemma_not_even_div";
+Addsimps [lemma_not_even_div];
+
+Goal "Suc n div 2 <= Suc (Suc n) div 2";
+by (auto_tac (claset(),simpset() addsimps [div_le_mono]));
+qed "Suc_n_le_Suc_Suc_n_div_2";
+Addsimps [Suc_n_le_Suc_Suc_n_div_2];
+
+Goal "(0::nat) < n --> 0 < (n + 1) div 2";
+by (induct_tac "n" 1);
+by Auto_tac;
+by (rtac (Suc_n_le_Suc_Suc_n_div_2 RSN (2,less_le_trans)) 1);
+by Auto_tac;
+qed_spec_mp "Suc_n_div_2_gt_zero";
+Addsimps [Suc_n_div_2_gt_zero];
+
+Goal "0 < n & even n --> 1 < n";
+by (induct_tac "n" 1);
+by Auto_tac;
+qed_spec_mp "even_gt_zero_gt_one";
+bind_thm ("even_gt_zero_gt_one",conjI RS even_gt_zero_gt_one);
+
+(* more general *)
+Goal "n div 2 <= (Suc n) div 2";
+by (auto_tac (claset(),simpset() addsimps [div_le_mono]));
+qed "le_Suc_n_div_2";
+Addsimps [le_Suc_n_div_2];
+
+Goal "(1::nat) < n --> 0 < n div 2";
+by (induct_tac "n" 1);
+by Auto_tac;
+by (dtac (CLAIM "[|0 < n; ~ 1 < n |] ==> n = (1::nat)") 1);
+by (rtac (le_Suc_n_div_2 RSN (2,less_le_trans)) 3);
+by Auto_tac;
+qed_spec_mp "div_2_gt_zero";
+Addsimps [div_2_gt_zero];
+
+Goal "even n ==> (n + 1) div 2 = n div 2";
+by (rtac (CLAIM "2 * x = 2 * y ==> x = (y::nat)") 1);
+by (rtac (lemma_even_div RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex]));
+qed "lemma_even_div2";
+Addsimps [lemma_even_div2];
+
+Goal "~even n ==> (n + 1) div 2 = Suc (n div 2)";
+by (rtac (CLAIM "2 * x = 2 * y ==> x = (y::nat)") 1);
+by (rtac (lemma_not_even_div RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps [even_not_odd,
+    odd_Suc_mult_two_ex])); 
+by (cut_inst_tac [("m","Suc(2*m)"),("n","2")] mod_div_equality 1); 
+by Auto_tac; 
+qed "lemma_not_even_div2";
+Addsimps [lemma_not_even_div2];
+
+Goal "(Suc n) div 2 = 0 ==> even n";
+by (rtac ccontr 1);
+by (dtac lemma_not_even_div2 1 THEN Auto_tac);
+qed "Suc_n_div_two_zero_even";
+
+Goal "0 < n ==> even n = (~ even(n - 1))";
+by (case_tac "n" 1);
+by Auto_tac;
+qed "even_num_iff";
+
+Goal "0 < n ==> odd n = (~ odd(n - 1))";
+by (case_tac "n" 1);
+by Auto_tac;
+qed "odd_num_iff";
+
+(* Some mod and div stuff *)
+
+Goal "n ~= (0::nat) ==> (m = m div n * n + m mod n) & m mod n < n";
+by (auto_tac (claset() addIs [mod_less_divisor],simpset()
+    addsimps [mod_div_equality]));
+qed "mod_div_eq_less";
+
+Goal "(k*n + m) mod n = m mod (n::nat)";
+by (simp_tac (simpset() addsimps mult_ac @ add_ac) 1);
+qed "mod_mult_self3";
+Addsimps [mod_mult_self3];
+
+Goal "Suc (k*n + m) mod n = Suc m mod n";
+by (rtac (CLAIM "Suc (m + n) = (m + Suc n)" RS ssubst) 1);
+by (rtac mod_mult_self3 1);
+qed "mod_mult_self4";
+Addsimps [mod_mult_self4];
+
+Goal "Suc m mod n = Suc (m mod n) mod n";
+by (cut_inst_tac [("d'","Suc (m mod n) mod n")] (CLAIM "EX d. d' = d") 1);
+by (etac exE 1);
+by (Asm_simp_tac 1);
+by (res_inst_tac [("t","m"),("n1","n")] (mod_div_equality RS subst) 1);
+by Auto_tac;
+qed "mod_Suc_eq_Suc_mod";
+
+Goal "even n = (even (n mod 4))";
+by (cut_inst_tac [("d'","(even (n mod 4))")] (CLAIM "EX d. d' = d") 1);
+by (etac exE 1);
+by (Asm_simp_tac 1);
+by (res_inst_tac [("t","n"),("n1","4")] (mod_div_equality RS subst) 1);
+by (auto_tac (claset(),simpset() addsimps [even_mult,even_num_iff]));
+qed "even_even_mod_4_iff";
+
+Goal "odd n = (odd (n mod 4))";
+by (auto_tac (claset(),simpset() addsimps [odd_not_even,
+    even_even_mod_4_iff RS sym]));
+qed "odd_odd_mod_4_iff";
+
+Goal "odd n ==> ((-1) ^ n) = (-1::real)";
+by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
+qed "odd_realpow_minus_one";
+Addsimps [odd_realpow_minus_one];
+
+Goal "even(4*n)";
+by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex]));
+qed "even_4n";
+Addsimps [even_4n];
+
+Goal "n mod 4 = 0 ==> even(n div 2)";
+by Auto_tac;
+qed "lemma_even_div_2";
+
+Goal "n mod 4 = 0 ==> even(n)";
+by Auto_tac;
+qed "lemma_mod_4_even_n";
+
+Goal "n mod 4 = 1 ==> odd(n)";
+by (res_inst_tac [("t","n"),("n1","4")] (mod_div_equality RS subst) 1);
+by (auto_tac (claset(),simpset() addsimps [odd_num_iff]));
+qed "lemma_mod_4_odd_n";
+
+Goal "n mod 4 = 2 ==> even(n)";
+by (res_inst_tac [("t","n"),("n1","4")] (mod_div_equality RS subst) 1);
+by (auto_tac (claset(),simpset() addsimps [even_num_iff]));
+qed "lemma_mod_4_even_n2";
+
+Goal "n mod 4 = 3 ==> odd(n)";
+by (res_inst_tac [("t","n"),("n1","4")] (mod_div_equality RS subst) 1);
+by (auto_tac (claset(),simpset() addsimps [odd_num_iff]));
+qed "lemma_mod_4_odd_n2";
+
+Goal "even n ==> ((-1) ^ n) = (1::real)";
+by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex]));
+qed "even_realpow_minus_one";
+Addsimps [even_realpow_minus_one];
+
+Goal "((4 * n) + 2) div 2 = (2::nat) * n + 1";
+by (rtac (CLAIM_SIMP "4*n + (2::nat) = 2 * (2*n + 1)" 
+          [add_mult_distrib2] RS ssubst) 1);
+by Auto_tac;
+qed "lemma_4n_add_2_div_2_eq";
+Addsimps [lemma_4n_add_2_div_2_eq];
+
+Goal "(Suc(Suc(4 * n))) div 2 = (2::nat) * n + 1";
+by Auto_tac;
+qed "lemma_Suc_Suc_4n_div_2_eq";
+Addsimps [lemma_Suc_Suc_4n_div_2_eq];
+
+Goal "(Suc(Suc(n * 4))) div 2 = (2::nat) * n + 1";
+by (cut_inst_tac [("n","n")] lemma_Suc_Suc_4n_div_2_eq 1);
+by (auto_tac (claset(),simpset() addsimps [mult_commute] 
+    delsimps [lemma_Suc_Suc_4n_div_2_eq]));
+qed "lemma_Suc_Suc_4n_div_2_eq2";
+Addsimps [lemma_Suc_Suc_4n_div_2_eq2];
+
+Goal "n mod 4 = 3 ==> odd((n - 1) div 2)";
+by (res_inst_tac [("t","n"),("n1","4")] (mod_div_equality RS subst) 1);
+by (asm_full_simp_tac (simpset() addsimps [odd_num_iff]) 1);
+val lemma_odd_mod_4_div_2 = result();
+
+Goal "(4 * n) div 2 = (2::nat) * n";
+by Auto_tac;
+qed "lemma_4n_div_2_eq";
+Addsimps [lemma_4n_div_2_eq];
+
+Goal "(n  * 4) div 2 = (2::nat) * n";
+by Auto_tac;
+qed "lemma_4n_div_2_eq2";
+Addsimps [lemma_4n_div_2_eq2];
+
+Goal "n mod 4 = 1 ==> even ((n - 1) div 2)";
+by (res_inst_tac [("t","n"),("n1","4")] (mod_div_equality RS subst) 1);
+by (dtac ssubst 1 THEN assume_tac 2);
+by (rtac ((CLAIM "(n::nat) + 1 - 1 = n") RS ssubst) 1);
+by Auto_tac;
+val lemma_even_mod_4_div_2 = result();
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/EvenOdd.thy	Thu Nov 15 16:12:49 2001 +0100
@@ -0,0 +1,20 @@
+(*  Title       : EvenOdd.thy
+    Author      : Jacques D. Fleuriot  
+    Copyright   : 1999  University of Edinburgh
+    Description : Even and odd numbers defined 
+*)
+
+EvenOdd = NthRoot +  
+
+consts even :: "nat => bool"
+primrec 
+   even_0   "even 0 = True"
+   even_Suc "even (Suc n) = (~ (even n))"    
+
+consts odd :: "nat => bool"
+primrec 
+   odd_0   "odd 0 = False"
+   odd_Suc "odd (Suc n) = (~ (odd n))"    
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/ExtraThms2.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -0,0 +1,563 @@
+(*lcp: lemmas needed now because 2 is a binary numeral!*)
+
+Goal "m+m = m*(2::nat)";
+by Auto_tac; 
+qed "m_add_m_eq2";
+
+(*lcp: needed for binary 2 MOVE UP???*)
+
+Goal "(0::real) <= x^2";
+by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2]) 1); 
+qed "zero_le_x_squared";
+Addsimps [zero_le_x_squared];
+
+fun multl_by_tac x i = 
+       let val cancel_thm = 
+           CLAIM "[| (0::real)<z; z*x<z*y |] ==> x<y" 
+       in
+           res_inst_tac [("z",x)] cancel_thm i 
+       end;
+
+fun multr_by_tac x i = 
+       let val cancel_thm = 
+           CLAIM "[| (0::real)<z; x*z<y*z |] ==> x<y" 
+       in
+           res_inst_tac [("z",x)] cancel_thm i 
+       end;
+
+(* unused? *)
+Goal "ALL x y. x < y --> (f::real=>real) x < f y ==> inj f";
+by (rtac injI 1 THEN rtac ccontr 1);
+by (dtac (ARITH_PROVE "x ~= y ==> x < y | y < (x::real)") 1);
+by (Step_tac 1);
+by (auto_tac (claset() addSDs [spec],simpset()));
+qed "real_monofun_inj";
+
+(* HyperDef *)
+
+Goal "0 = Abs_hypreal (hyprel `` {%n. 0})";
+by (simp_tac (simpset() addsimps [hypreal_zero_def RS meta_eq_to_obj_eq RS sym]) 1);
+qed "hypreal_zero_num";
+
+Goal "1 = Abs_hypreal (hyprel `` {%n. 1})";
+by (simp_tac (simpset() addsimps [hypreal_one_def RS meta_eq_to_obj_eq RS sym]) 1);
+qed "hypreal_one_num";
+
+(* RealOrd *)
+
+Goalw [real_of_posnat_def] "0 < real_of_posnat n";
+by (rtac (real_gt_zero_preal_Ex RS iffD2) 1);
+by (Blast_tac 1);
+qed "real_of_posnat_gt_zero";
+
+Addsimps [real_of_posnat_gt_zero];
+
+bind_thm ("real_inv_real_of_posnat_gt_zero",
+          real_of_posnat_gt_zero RS real_inverse_gt_0);
+Addsimps [real_inv_real_of_posnat_gt_zero];
+
+bind_thm ("real_of_posnat_ge_zero",real_of_posnat_gt_zero RS order_less_imp_le);
+Addsimps [real_of_posnat_ge_zero];
+
+Goal "real_of_posnat n ~= 0";
+by (rtac (real_of_posnat_gt_zero RS real_not_refl2 RS not_sym) 1);
+qed "real_of_posnat_not_eq_zero";
+Addsimps[real_of_posnat_not_eq_zero];
+
+Addsimps [real_of_posnat_not_eq_zero RS real_mult_inv_left];
+Addsimps [real_of_posnat_not_eq_zero RS real_mult_inv_right];
+
+Goal "1 <= real_of_posnat n";
+by (simp_tac (simpset() addsimps [real_of_posnat_one RS sym]) 1);
+by (induct_tac "n" 1);
+by (auto_tac (claset(),
+	      simpset () addsimps [real_of_posnat_Suc,real_of_posnat_one,
+				   order_less_imp_le]));
+qed "real_of_posnat_ge_one";
+Addsimps [real_of_posnat_ge_one];
+
+Goal "inverse(real_of_posnat n) ~= 0";
+by (rtac ((real_of_posnat_gt_zero RS 
+    real_not_refl2 RS not_sym) RS real_inverse_not_zero) 1);
+qed "real_of_posnat_real_inv_not_zero";
+Addsimps [real_of_posnat_real_inv_not_zero];
+
+Goal "inverse(real_of_posnat x) = inverse(real_of_posnat y) ==> x = y";
+by (rtac (inj_real_of_posnat RS injD) 1);
+by (res_inst_tac [("n2","x")] 
+    (real_of_posnat_real_inv_not_zero RS real_mult_left_cancel RS iffD1) 1);
+by (asm_full_simp_tac (simpset() addsimps [(real_of_posnat_gt_zero RS 
+    real_not_refl2 RS not_sym) RS real_mult_inv_left]) 1);
+qed "real_of_posnat_real_inv_inj";
+
+Goal "r < r + inverse(real_of_posnat n)";
+by Auto_tac;
+qed "real_add_inv_real_of_posnat_less";
+Addsimps [real_add_inv_real_of_posnat_less];
+
+Goal "r <= r + inverse(real_of_posnat n)";
+by Auto_tac;
+qed "real_add_inv_real_of_posnat_le";
+Addsimps [real_add_inv_real_of_posnat_le];
+
+Goal "0 < r ==> r*(1 + -inverse(real_of_posnat n)) < r";
+by (simp_tac (simpset() addsimps [real_add_mult_distrib2]) 1);
+by (res_inst_tac [("C","-r")] real_less_add_left_cancel 1);
+by (auto_tac (claset() addIs [real_mult_order],simpset() 
+    addsimps [real_add_assoc RS sym,real_minus_zero_less_iff2]));
+qed "real_mult_less_self";
+
+Goal "(EX n. inverse(real_of_posnat n) < r) = (EX n. 1 < r * real_of_posnat n)";
+by (Step_tac 1);
+by (dres_inst_tac [("n1","n")] (real_of_posnat_gt_zero 
+                       RS real_mult_less_mono1) 1);
+by (dres_inst_tac [("n2","n")] (real_of_posnat_gt_zero RS 
+        real_inverse_gt_0 RS real_mult_less_mono1) 2);
+by (auto_tac (claset(),
+	      simpset() addsimps [(real_of_posnat_gt_zero RS 
+    real_not_refl2 RS not_sym),real_mult_assoc]));
+qed "real_of_posnat_inv_Ex_iff";
+
+Goal "(inverse(real_of_posnat n) < r) = (1 < r * real_of_posnat n)";
+by (Step_tac 1);
+by (dres_inst_tac [("n1","n")] (real_of_posnat_gt_zero 
+                       RS real_mult_less_mono1) 1);
+by (dres_inst_tac [("n2","n")] (real_of_posnat_gt_zero RS 
+        real_inverse_gt_0 RS real_mult_less_mono1) 2);
+by (auto_tac (claset(),simpset() addsimps [real_mult_assoc]));
+qed "real_of_posnat_inv_iff";
+
+Goal "[| (0::real) <=z; x<y |] ==> z*x<=z*y";
+by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_le_less_mono1]) 1);
+qed "real_mult_le_less_mono2";
+
+Goal "[| (0::real) <=z; x<=y |] ==> z*x<=z*y";
+by (dres_inst_tac [("x","x")] real_le_imp_less_or_eq 1);
+by (auto_tac (claset() addIs [real_mult_le_less_mono2], simpset()));
+qed "real_mult_le_le_mono1";
+
+Goal "[| (0::real)<=z; x<=y |] ==> x*z<=y*z";
+by (dtac (real_mult_le_le_mono1) 1);
+by (auto_tac (claset(),simpset() addsimps [real_mult_commute]));
+qed "real_mult_le_le_mono2";
+
+Goal "(inverse(real_of_posnat n) <= r) = (1 <= r * real_of_posnat n)";
+by (Step_tac 1);
+by (dres_inst_tac [("n2","n")] (real_of_posnat_gt_zero RS 
+    order_less_imp_le RS real_mult_le_le_mono1) 1);
+by (dres_inst_tac [("n3","n")] (real_of_posnat_gt_zero RS 
+        real_inverse_gt_0 RS order_less_imp_le RS 
+        real_mult_le_le_mono1) 2);
+by (auto_tac (claset(),simpset() addsimps real_mult_ac));
+qed "real_of_posnat_inv_le_iff";
+
+Goalw [real_of_posnat_def] 
+      "(real_of_posnat n < real_of_posnat m) = (n < m)";
+by Auto_tac;
+qed "real_of_posnat_less_iff";
+Addsimps [real_of_posnat_less_iff];
+
+Goal "(real_of_posnat n <= real_of_posnat m) = (n <= m)";
+by (auto_tac (claset() addDs [inj_real_of_posnat RS injD],
+    simpset() addsimps [real_le_less,le_eq_less_or_eq]));
+qed "real_of_posnat_le_iff";
+Addsimps [real_of_posnat_le_iff];
+
+Goal "[| (0::real)<z; x*z<y*z |] ==> x<y";
+by Auto_tac;
+qed "real_mult_less_cancel3";
+
+Goal "[| (0::real)<z; z*x<z*y |] ==> x<y";
+by Auto_tac;
+qed "real_mult_less_cancel4";
+
+Goal "0 < u  ==> (u < inverse (real_of_posnat n)) = (real_of_posnat n < inverse(u))";
+by (Step_tac 1);
+by (res_inst_tac [("n2","n")] ((real_of_posnat_gt_zero RS 
+    real_inverse_gt_0) RS real_mult_less_cancel3) 1);
+by (res_inst_tac [("x1","u")] ( real_inverse_gt_0
+   RS real_mult_less_cancel3) 2);
+by (auto_tac (claset(),
+	      simpset() addsimps [real_not_refl2 RS not_sym]));
+by (res_inst_tac [("z","u")] real_mult_less_cancel4 1);
+by (res_inst_tac [("n1","n")] (real_of_posnat_gt_zero RS 
+    real_mult_less_cancel4) 3);
+by (auto_tac (claset(),
+	      simpset() addsimps [real_not_refl2 RS not_sym,
+              real_mult_assoc RS sym]));
+qed "real_of_posnat_less_inv_iff";
+
+Goal "0 < u ==> (u = inverse(real_of_posnat n)) = (real_of_posnat n = inverse u)";
+by Auto_tac;
+qed "real_of_posnat_inv_eq_iff";
+
+Goal "0 <= 1 + -inverse(real_of_posnat n)";
+by (res_inst_tac [("C","inverse(real_of_posnat n)")] real_le_add_right_cancel 1);
+by (simp_tac (simpset() addsimps [real_add_assoc,
+    real_of_posnat_inv_le_iff]) 1);
+qed "real_add_one_minus_inv_ge_zero";
+
+Goal "0 < r ==> 0 <= r*(1 + -inverse(real_of_posnat n))";
+by (dtac (real_add_one_minus_inv_ge_zero RS real_mult_le_less_mono1) 1);
+by (Auto_tac);
+qed "real_mult_add_one_minus_ge_zero";
+
+Goal "x*y = (1::real) ==> y = inverse x";
+by (case_tac "x ~= 0" 1);
+by (res_inst_tac [("c1","x")] (real_mult_left_cancel RS iffD1) 1);
+by Auto_tac;
+qed "real_inverse_unique";
+
+Goal "[| (0::real) < x; x < 1 |] ==> 1 < inverse x";
+by (auto_tac (claset() addDs [real_inverse_less_swap],simpset()));
+qed "real_inverse_gt_one";
+
+Goal "(0 < real (n::nat)) = (0 < n)";
+by (rtac (real_of_nat_less_iff RS subst) 1);
+by Auto_tac;
+qed "real_of_nat_gt_zero_cancel_iff";
+Addsimps [real_of_nat_gt_zero_cancel_iff];
+
+Goal "(real (n::nat) <= 0) = (n = 0)";
+by (rtac ((real_of_nat_zero) RS subst) 1);
+by (rtac (real_of_nat_le_iff RS ssubst) 1);
+by Auto_tac;
+qed "real_of_nat_le_zero_cancel_iff";
+Addsimps [real_of_nat_le_zero_cancel_iff];
+
+Goal "~ real (n::nat) < 0";
+by (simp_tac (simpset() addsimps [symmetric real_le_def,
+    real_of_nat_ge_zero]) 1);
+qed "not_real_of_nat_less_zero";
+Addsimps [not_real_of_nat_less_zero];
+
+Goalw [real_le_def,le_def] 
+      "(0 <= real (n::nat)) = (0 <= n)";
+by (Simp_tac 1);
+qed "real_of_nat_ge_zero_cancel_iff";
+Addsimps [real_of_nat_ge_zero_cancel_iff];
+
+Goal "real n = (if n=0 then 0 else 1 + real ((n::nat) - 1))";
+by (case_tac "n" 1);
+by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc]));
+qed "real_of_nat_num_if";
+
+Goal "4 * real n = real (4 * (n::nat))";
+by Auto_tac;
+qed "real_of_nat_mult_num_4_eq";
+
+(*REDUNDANT
+    Goal "x * x = -(y * y) ==> x = (0::real)";
+    by (auto_tac (claset() addIs [
+	real_sum_squares_cancel],simpset()));
+    qed "real_sum_squares_cancel1a";
+
+    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 * x = -(y * y) ==> x = (0::real) & y=0";
+by (auto_tac (claset() addIs [real_sum_squares_cancel],simpset()));
+qed "real_sum_squares_cancel_a";
+
+Goal "x*x - (1::real) = (x + 1)*(x - 1)";
+by (auto_tac (claset(),simpset() addsimps [real_add_mult_distrib,
+    real_add_mult_distrib2,real_diff_def]));
+qed "real_squared_diff_one_factored";
+
+Goal "(x*x = (1::real)) = (x = 1 | x = - 1)";
+by Auto_tac;
+by (dtac (CLAIM "x = (y::real) ==> x - y = 0") 1);
+by (auto_tac (claset(),simpset() addsimps [real_squared_diff_one_factored]));
+qed "real_mult_is_one";
+AddIffs [real_mult_is_one];
+
+Goal "(x + y/2 <= (y::real)) = (x <= y /2)";
+by Auto_tac;
+qed "real_le_add_half_cancel";
+Addsimps [real_le_add_half_cancel];
+
+Goal "(x::real) - x/2 = x/2";
+by Auto_tac;
+qed "real_minus_half_eq";
+Addsimps [real_minus_half_eq];
+
+Goal "[|(0::real) < x;0 < x1; x1 * y < x * u |] ==> inverse x * y < inverse x1 * u";
+by (multl_by_tac "x" 1);
+by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym]));
+by (simp_tac (simpset() addsimps real_mult_ac) 1);
+by (multr_by_tac "x1" 1);
+by (auto_tac (claset(),simpset() addsimps real_mult_ac));
+qed "real_mult_inverse_cancel";
+
+Goal "[|(0::real) < x;0 < x1; x1 * y < x * u |] ==> y * inverse x < u * inverse x1";
+by (auto_tac (claset() addDs [real_mult_inverse_cancel],simpset() addsimps real_mult_ac));
+qed "real_mult_inverse_cancel2";
+
+Goal "0 < inverse (real (Suc n))";
+by Auto_tac;
+qed "inverse_real_of_nat_gt_zero";
+Addsimps [ inverse_real_of_nat_gt_zero];
+
+Goal "0 <= inverse (real (Suc n))";
+by Auto_tac;
+qed "inverse_real_of_nat_ge_zero";
+Addsimps [ inverse_real_of_nat_ge_zero];
+
+Goal "x ~= 0 ==> x * x + y * y ~= (0::real)";
+by (rtac (CLAIM "!!x. ((x = y ==> False)) ==> x ~= y") 1);
+by (dtac (real_sum_squares_cancel) 1);
+by (Asm_full_simp_tac 1);
+qed "real_sum_squares_not_zero";
+
+Goal "y ~= 0 ==> x * x + y * y ~= (0::real)";
+by (rtac (CLAIM "!!x. ((x = y ==> False)) ==> x ~= y") 1);
+by (dtac (real_sum_squares_cancel2) 1);
+by (Asm_full_simp_tac 1);
+qed "real_sum_squares_not_zero2";
+
+(* RealAbs *)
+
+(* nice theorem *)
+Goal "abs x * abs x = x * (x::real)";
+by (cut_inst_tac [("R1.0","x"),("R2.0","0")] real_linear 1);
+by (auto_tac (claset(),simpset() addsimps [abs_eqI2,abs_minus_eqI2]));
+qed "abs_mult_abs";
+Addsimps [abs_mult_abs];
+
+(* RealPow *)
+Goalw [real_divide_def]
+    "(x/y) ^ n = ((x::real) ^ n/ y ^ n)";
+by (auto_tac (claset(),simpset() addsimps [realpow_mult,
+    realpow_inverse]));
+qed "realpow_divide";
+
+Goal "isCont (%x. x ^ n) x";
+by (rtac (DERIV_pow RS DERIV_isCont) 1);
+qed "isCont_realpow";
+Addsimps [isCont_realpow];
+
+Goal "(0::real) <= r --> 0 <= r ^ n";
+by (induct_tac "n" 1);
+by (auto_tac (claset(),simpset() addsimps [real_0_le_mult_iff]));
+qed_spec_mp "realpow_ge_zero2";
+
+Goal "(0::real) <= x & x <= y --> x ^ n <= y ^ n";
+by (induct_tac "n" 1);
+by (auto_tac (claset() addSIs [real_mult_le_mono],
+    simpset() addsimps [realpow_ge_zero2]));
+qed_spec_mp "realpow_le2";
+
+Goal "(1::real) < r ==> 1 < r ^ (Suc n)";
+by (forw_inst_tac [("n","n")] realpow_ge_one 1);
+by (dtac real_le_imp_less_or_eq 1 THEN Step_tac 1);
+by (forward_tac [(real_zero_less_one RS real_less_trans)] 1);
+by (dres_inst_tac [("y","r ^ n")] (real_mult_less_mono2) 1);
+by (assume_tac 1);
+by (auto_tac (claset() addDs [real_less_trans],simpset()));
+qed "realpow_Suc_gt_one";
+
+Goal "(x ^ 2 + y ^ 2 = (0::real)) = (x = 0 & y = 0)";
+by (auto_tac (claset() addIs [real_sum_squares_cancel, real_sum_squares_cancel2], 
+              simpset() addsimps [numeral_2_eq_2]));
+qed "realpow_two_sum_zero_iff";
+Addsimps [realpow_two_sum_zero_iff];
+
+Goal "(0::real) <= u ^ 2 + v ^ 2";
+by (rtac (real_le_add_order) 1);
+by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
+qed "realpow_two_le_add_order";
+Addsimps [realpow_two_le_add_order];
+
+Goal "(0::real) <= u ^ 2 + v ^ 2 + w ^ 2";
+by (REPEAT(rtac (real_le_add_order) 1));
+by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
+qed "realpow_two_le_add_order2";
+Addsimps [realpow_two_le_add_order2];
+
+Goal "(0::real) <= x*x + y*y";
+by (cut_inst_tac [("u","x"),("v","y")] realpow_two_le_add_order 1);
+by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
+qed "real_mult_self_sum_ge_zero";
+Addsimps [real_mult_self_sum_ge_zero];
+Addsimps [real_mult_self_sum_ge_zero RS abs_eqI1];
+
+Goal "x ~= 0 ==> (0::real) < x * x + y * y";
+by (cut_inst_tac [("x","x"),("y","y")] real_mult_self_sum_ge_zero 1);
+by (dtac real_le_imp_less_or_eq 1);
+by (dres_inst_tac [("y","y")] real_sum_squares_not_zero 1);
+by Auto_tac;
+qed "real_sum_square_gt_zero";
+
+Goal "y ~= 0 ==> (0::real) < x * x + y * y";
+by (rtac (real_add_commute RS subst) 1);
+by (etac real_sum_square_gt_zero 1);
+qed "real_sum_square_gt_zero2";
+
+Goal "-(u * u) <= (x * (x::real))";
+by (res_inst_tac [("j","0")] real_le_trans 1);
+by Auto_tac;
+qed "real_minus_mult_self_le";
+Addsimps [real_minus_mult_self_le];
+
+Goal "-(u ^ 2) <= (x::real) ^ 2";
+by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
+qed "realpow_square_minus_le";
+Addsimps [realpow_square_minus_le];
+
+Goal "(m::real) ^ n = (if n=0 then 1 else m * m ^ (n - 1))";
+by (case_tac "n" 1);
+by Auto_tac;
+qed "realpow_num_eq_if";
+
+Goal "0 < (2::real) ^ (4*d)";
+by (induct_tac "d" 1);
+by (auto_tac (claset(),simpset() addsimps [realpow_num_eq_if]));
+qed "real_num_zero_less_two_pow";
+Addsimps [real_num_zero_less_two_pow];
+
+Goal "x * (4::real)   < y ==> x * (2 ^ 8) < y * (2 ^ 6)";
+by (subgoal_tac "(2::real) ^ 8 = 4 * (2 ^ 6)" 1);
+by (asm_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1);
+by (auto_tac (claset(),simpset() addsimps [realpow_num_eq_if]));
+qed "lemma_realpow_num_two_mono";
+
+Goal "2 ^ 2 = (4::real)";
+by (simp_tac (simpset() addsimps [realpow_num_eq_if]) 1);
+val lemma_realpow_4 = result();
+Addsimps [lemma_realpow_4];
+
+Goal "2 ^ 4 = (16::real)";
+by (simp_tac (simpset() addsimps [realpow_num_eq_if]) 1);
+val lemma_realpow_16 = result();
+Addsimps [lemma_realpow_16];
+
+(* HyperOrd *)
+
+Goal "[| (0::hypreal) < x; y < 0 |] ==> y*x < 0";
+by (auto_tac (claset(),simpset() addsimps [hypreal_mult_commute,
+    hypreal_mult_less_zero]));
+qed "hypreal_mult_less_zero2";
+
+(* HyperPow *)
+Goal "(0::hypreal) <= x * x";
+by (auto_tac (claset(),simpset() addsimps 
+    [hypreal_0_le_mult_iff]));
+qed "hypreal_mult_self_ge_zero";
+Addsimps [hypreal_mult_self_ge_zero];
+
+(* deleted from distribution but I prefer to have it *)
+Goal "[|(0::hypreal) <= x; 0 <= y; x ^ Suc n = y ^ Suc n |] ==> x = y";
+by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
+by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
+by (auto_tac (claset(),simpset() addsimps 
+    [hrealpow,hypreal_mult,hypreal_le,hypreal_zero_num]));
+by (ultra_tac (claset() addIs [realpow_Suc_cancel_eq],
+    simpset()) 1);
+qed "hrealpow_Suc_cancel_eq";
+
+(* NSA.ML: next two were there before? Not in distrib though *)
+
+Goal "[| x + y : HInfinite; y: HFinite |] ==> x : HInfinite";
+by (rtac ccontr 1);
+by (dtac (HFinite_HInfinite_iff RS iffD2) 1);
+by (auto_tac (claset() addDs [HFinite_add],simpset() 
+    addsimps [HInfinite_HFinite_iff]));
+qed "HInfinite_HFinite_add_cancel";
+
+Goal "[| x : HInfinite; y : HFinite |] ==> x + y : HInfinite";
+by (res_inst_tac [("y","-y")] HInfinite_HFinite_add_cancel 1);
+by (auto_tac (claset(),simpset() addsimps [hypreal_add_assoc,
+    HFinite_minus_iff]));
+qed "HInfinite_HFinite_add";
+
+Goal "[| x : HInfinite; x <= y; 0 <= x |] ==> y : HInfinite";
+by (auto_tac (claset() addIs [HFinite_bounded],simpset() 
+    addsimps [HInfinite_HFinite_iff]));
+qed "HInfinite_ge_HInfinite";
+
+Goal "[| x : Infinitesimal; x ~= 0 |] ==> inverse x : HInfinite";
+by (rtac ccontr 1 THEN dtac (HFinite_HInfinite_iff RS iffD2) 1);
+by (auto_tac (claset() addDs [Infinitesimal_HFinite_mult2],simpset()));
+qed "Infinitesimal_inverse_HInfinite";
+
+Goal "n : HNatInfinite ==> inverse (hypreal_of_hypnat n) : Infinitesimal";
+by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
+by (auto_tac (claset(),simpset() addsimps [hypreal_of_hypnat,hypreal_inverse,
+    HNatInfinite_FreeUltrafilterNat_iff,Infinitesimal_FreeUltrafilterNat_iff2]));
+by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
+by Auto_tac;
+by (dres_inst_tac [("x","m + 1")] spec 1);
+by (Ultra_tac 1);
+by (subgoal_tac "abs(inverse (real (Y x))) = inverse(real (Y x))" 1);
+by (auto_tac (claset() addSIs [abs_eqI2],simpset()));
+by (rtac real_inverse_less_swap 1);
+by Auto_tac;
+qed "HNatInfinite_inverse_Infinitesimal";
+Addsimps [HNatInfinite_inverse_Infinitesimal];
+
+Goal "n : HNatInfinite ==> inverse (hypreal_of_hypnat n) ~= 0";
+by (auto_tac (claset() addSIs [hypreal_inverse_not_zero],simpset()));
+qed "HNatInfinite_inverse_not_zero";
+Addsimps [HNatInfinite_inverse_not_zero];
+
+Goal "N : HNatInfinite \
+\     ==> (*fNat* (%x. inverse (real x))) N : Infinitesimal";
+by (res_inst_tac [("f1","inverse")]  (starfun_stafunNat_o2 RS subst) 1);
+by (subgoal_tac "hypreal_of_hypnat N ~= 0" 1);
+by (auto_tac (claset(),simpset() addsimps [starfunNat_real_of_nat]));
+qed "starfunNat_inverse_real_of_nat_Infinitesimal";
+Addsimps [starfunNat_inverse_real_of_nat_Infinitesimal];
+
+Goal "[| x : HInfinite; y : HFinite - Infinitesimal |] \
+\     ==> x * y : HInfinite";
+by (rtac ccontr 1 THEN dtac (HFinite_HInfinite_iff RS iffD2) 1); 
+by (ftac HFinite_Infinitesimal_not_zero 1);
+by (dtac HFinite_not_Infinitesimal_inverse 1);
+by (Step_tac 1 THEN dtac HFinite_mult 1);
+by (auto_tac (claset(),simpset() addsimps [hypreal_mult_assoc,
+    HFinite_HInfinite_iff]));
+qed "HInfinite_HFinite_not_Infinitesimal_mult";
+
+Goal "[| x : HInfinite; y : HFinite - Infinitesimal |] \
+\     ==> y * x : HInfinite";
+by (auto_tac (claset(),simpset() addsimps [hypreal_mult_commute,
+    HInfinite_HFinite_not_Infinitesimal_mult]));
+qed "HInfinite_HFinite_not_Infinitesimal_mult2";
+
+Goal "[| x : HInfinite; 0 < x; y : Reals |] ==> y < x";
+by (auto_tac (claset() addSDs [bspec],simpset() addsimps 
+    [HInfinite_def,hrabs_def,order_less_imp_le]));
+qed "HInfinite_gt_SReal";
+
+Goal "[| x : HInfinite; 0 < x |] ==> 1 < x";
+by (auto_tac (claset() addIs [HInfinite_gt_SReal],simpset()));
+qed "HInfinite_gt_zero_gt_one";
+
+(* not added at proof?? *)
+Addsimps [HInfinite_omega];
+
+(* Add in HyperDef.ML? *)
+Goalw [omega_def] "0 < omega";
+by (auto_tac (claset(),simpset() addsimps [hypreal_less,hypreal_zero_num]));
+qed "hypreal_omega_gt_zero";
+Addsimps [hypreal_omega_gt_zero];
+
+Goal "1 ~: HInfinite";
+by (simp_tac (simpset() addsimps [HInfinite_HFinite_iff]) 1);
+qed "not_HInfinite_one";
+Addsimps [not_HInfinite_one];
+
+
+(* RComplete.ML *)
+
+Goal "0 < x ==> ALL y. EX (n::nat). y < real n * x";
+by (Step_tac 1);
+by (cut_inst_tac [("x","y*inverse(x)")] reals_Archimedean2 1);
+by (Step_tac 1);
+by (forw_inst_tac [("x","y * inverse x")] (real_mult_less_mono1) 1);
+by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,real_of_nat_def]));
+qed "reals_Archimedean3";
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/ExtraThms2.thy	Thu Nov 15 16:12:49 2001 +0100
@@ -0,0 +1,1 @@
+ExtraThms2 = HSeries 
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Fact.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -0,0 +1,90 @@
+(*  Title       : Fact.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Factorial function
+*)
+
+Goal "0 < fact n";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "fact_gt_zero";
+Addsimps [fact_gt_zero];
+
+Goal "fact n ~= 0";
+by (Simp_tac 1);
+qed "fact_not_eq_zero";
+Addsimps [fact_not_eq_zero];
+
+Goal "real (fact n) ~= 0";
+by Auto_tac; 
+qed "real_of_nat_fact_not_zero";
+Addsimps [real_of_nat_fact_not_zero];
+
+Goal "0 < real(fact n)";
+by Auto_tac; 
+qed "real_of_nat_fact_gt_zero";
+Addsimps [real_of_nat_fact_gt_zero];
+
+Goal "0 <= real(fact n)";
+by (Simp_tac 1);
+qed "real_of_nat_fact_ge_zero";
+Addsimps [real_of_nat_fact_ge_zero];
+
+Goal "1 <= fact n";
+by (induct_tac "n" 1);
+by (Auto_tac);
+qed "fact_ge_one";
+Addsimps [fact_ge_one];
+
+Goal "m <= n ==> fact m <= fact n";
+by (dtac le_imp_less_or_eq 1);
+by (auto_tac (claset() addSDs [less_imp_Suc_add],simpset()));
+by (induct_tac "k" 1);
+by (Auto_tac);
+qed "fact_mono";
+
+Goal "[| 0 < m; m < n |] ==> fact m < fact n";
+by (dres_inst_tac [("m","m")] less_imp_Suc_add 1);
+by Auto_tac;
+by (induct_tac "k" 1);
+by (Auto_tac);
+qed "fact_less_mono";
+
+Goal "0 < inverse (real (fact n))";
+by (auto_tac (claset(),simpset() addsimps [real_inverse_gt_0]));
+qed "inv_real_of_nat_fact_gt_zero";
+Addsimps [inv_real_of_nat_fact_gt_zero];
+
+Goal "0 <= inverse (real (fact n))";
+by (auto_tac (claset() addIs [order_less_imp_le],simpset()));
+qed "inv_real_of_nat_fact_ge_zero";
+Addsimps [inv_real_of_nat_fact_ge_zero];
+
+Goal "ALL m. ma < Suc m --> fact (Suc m - ma) = (Suc m - ma) * fact (m - ma)";
+by (induct_tac "ma" 1);
+by Auto_tac;
+by (dres_inst_tac [("x","m - 1")] spec 1);
+by Auto_tac;
+qed_spec_mp "fact_diff_Suc";
+
+Goal "fact 0 = 1";
+by Auto_tac;
+qed "fact_num0";
+Addsimps [fact_num0];
+
+Goal "fact m = (if m=0 then 1 else m * fact (m - 1))";
+by (case_tac "m" 1);
+by Auto_tac;
+qed "fact_num_eq_if";
+
+Goal "fact (m + n) = (if (m + n = 0) then 1 else (m + n) * (fact (m + n - 1)))";
+by (case_tac "m+n" 1);
+by Auto_tac;
+qed "fact_add_num_eq_if";
+
+Goal "fact (m + n) = (if (m = 0) then fact n \
+\     else (m + n) * (fact ((m - 1) + n)))";
+by (case_tac "m" 1);
+by Auto_tac;
+qed "fact_add_num_eq_if2";
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Fact.thy	Thu Nov 15 16:12:49 2001 +0100
@@ -0,0 +1,14 @@
+(*  Title       : Fact.thy 
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Factorial function
+*)
+
+Fact = NatStar + 
+
+consts fact :: nat => nat 
+primrec 
+   fact_0     "fact 0 = 1"
+   fact_Suc   "fact (Suc n) = (Suc n) * fact n" 
+
+end
\ No newline at end of file
--- a/src/HOL/Hyperreal/Hyperreal.thy	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/Hyperreal/Hyperreal.thy	Thu Nov 15 16:12:49 2001 +0100
@@ -1,4 +1,4 @@
 
-theory Hyperreal = HSeries:
+theory Hyperreal = Transcendental:
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/NthRoot.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -0,0 +1,170 @@
+(*  Title       : NthRoot.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Existence of nth root. Adapted from
+                   http://www.math.unl.edu/~webnotes
+*)
+
+(*----------------------------------------------------------------------
+                         Existence of nth root
+   Various lemmas needed for this result. We follow the proof
+   given by John Lindsay Orr (jorr@math.unl.edu) in his Analysis
+   Webnotes available on the www at http://www.math.unl.edu/~webnotes
+   Lemmas about sequences of reals are used to reach the result.
+ ---------------------------------------------------------------------*)
+
+Goal "[| (0::real) < a; 0 < n |] \
+\     ==> EX s. s : {x. x ^ n <= a & 0 < x}";
+by (case_tac "1 <= a" 1);
+by (res_inst_tac [("x","1")] exI 1);
+by (dtac not_real_leE 2);
+by (dtac (less_not_refl2 RS  not0_implies_Suc) 2);
+by (auto_tac (claset() addSIs [realpow_Suc_le_self],
+    simpset() addsimps [real_zero_less_one]));
+qed "lemma_nth_realpow_non_empty";
+
+Goal "[| (0::real) < a; 0 < n |] \
+\     ==> EX u. isUb (UNIV::real set) {x. x ^ n <= a & 0 < x} u";
+by (case_tac "1 <= a" 1);
+by (res_inst_tac [("x","a")] exI 1);
+by (dtac not_real_leE 2);
+by (res_inst_tac [("x","1")] exI 2);
+by (ALLGOALS(rtac (setleI RS isUbI)));
+by (Auto_tac);
+by (ALLGOALS(rtac ccontr));
+by (ALLGOALS(dtac not_real_leE));
+by (dtac realpow_ge_self2 1 THEN assume_tac 1);
+by (dres_inst_tac [("n","n")] (conjI 
+    RSN (2,conjI RS realpow_less)) 1);
+by (REPEAT(assume_tac 1));
+by (dtac real_le_trans 1 THEN assume_tac 1);
+by (dres_inst_tac [("y","y ^ n")] order_less_le_trans 1);
+by (assume_tac 1 THEN etac real_less_irrefl 1);
+by (dres_inst_tac [("n","n")] ((real_zero_less_one) RS (conjI 
+    RSN (2,conjI RS realpow_less))) 1);
+by (Auto_tac);
+qed "lemma_nth_realpow_isUb_ex";
+
+Goal "[| (0::real) < a; 0 < n |] \
+\    ==> EX u. isLub (UNIV::real set) {x. x ^ n <= a & 0 < x} u";
+by (blast_tac (claset() addIs [lemma_nth_realpow_isUb_ex,
+    lemma_nth_realpow_non_empty,reals_complete]) 1);
+qed "nth_realpow_isLub_ex";
+ 
+(*---------------------------------------------
+          First Half -- lemmas first
+ --------------------------------------------*)
+Goal "isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u \
+\          ==> u + inverse(real_of_posnat k) ~: {x. x ^ n <= a & 0 < x}";
+by (Step_tac 1 THEN dtac isLubD2 1 THEN Blast_tac 1);
+by (asm_full_simp_tac (simpset() addsimps [real_le_def]) 1);
+val lemma_nth_realpow_seq = result();
+
+Goal "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u; \
+\        0 < a; 0 < n |] ==> 0 < u";
+by (dtac lemma_nth_realpow_non_empty 1 THEN Auto_tac);
+by (dres_inst_tac [("y","s")] (isLub_isUb RS isUbD) 1);
+by (auto_tac (claset() addIs [order_less_le_trans],simpset()));
+val lemma_nth_realpow_isLub_gt_zero = result();
+
+Goal "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u; \
+\        0 < a; 0 < n |] ==> ALL k. a <= (u + inverse(real_of_posnat k)) ^ n";
+by (Step_tac 1);
+by (forward_tac [lemma_nth_realpow_seq] 1 THEN Step_tac 1);
+by (auto_tac (claset() addEs [real_less_asym],
+    simpset() addsimps [real_le_def]));
+by (fold_tac [real_le_def]);
+by (rtac real_less_trans 1);
+by (auto_tac (claset() addIs [real_inv_real_of_posnat_gt_zero,
+    lemma_nth_realpow_isLub_gt_zero],simpset()));
+val lemma_nth_realpow_isLub_ge = result();
+
+(*-----------------------
+   First result we want
+ ----------------------*)
+Goal "[| (0::real) < a; 0 < n; \
+\    isLub (UNIV::real set) \
+\    {x. x ^ n <= a & 0 < x} u |] ==> a <= u ^ n";
+by (forward_tac [lemma_nth_realpow_isLub_ge] 1 THEN Step_tac 1);
+by (rtac (LIMSEQ_inverse_real_of_posnat_add RS LIMSEQ_pow 
+    RS LIMSEQ_le_const) 1);
+by (auto_tac (claset(),simpset() addsimps [real_of_nat_def,
+    real_of_posnat_Suc]));
+qed "realpow_nth_ge";
+
+(*---------------------------------------------
+          Second Half
+ --------------------------------------------*)
+
+Goal "[| isLub (UNIV::real set) S u; x < u |] \
+\          ==> ~ isUb (UNIV::real set) S x";
+by (Step_tac 1);
+by (dtac isLub_le_isUb 1);
+by (assume_tac 1);
+by (dtac order_less_le_trans 1);
+by (auto_tac (claset(),simpset() 
+   addsimps [real_less_not_refl]));
+qed "less_isLub_not_isUb";
+
+Goal "~ isUb (UNIV::real set) S u \
+\              ==> EX x: S. u < x";
+by (rtac ccontr 1 THEN etac swap 1);
+by (rtac (setleI RS isUbI) 1);
+by (auto_tac (claset(),simpset() addsimps [real_le_def]));
+qed "not_isUb_less_ex";
+
+Goal "[| isLub (UNIV::real set) {x. x ^ n <= a & (0::real) < x} u; \
+\      0 < a; 0 < n |] ==> ALL k. (u*(1 + -inverse(real_of_posnat k))) ^ n <= a";
+by (Step_tac 1);
+by (forward_tac [less_isLub_not_isUb RS not_isUb_less_ex] 1);
+by (res_inst_tac [("n","k")] real_mult_less_self 1);
+by (blast_tac (claset() addIs [lemma_nth_realpow_isLub_gt_zero]) 1);
+by (Step_tac 1);
+by (dres_inst_tac [("n","k")] (lemma_nth_realpow_isLub_gt_zero RS 
+          real_mult_add_one_minus_ge_zero) 1);
+by (dtac (conjI RS realpow_le2) 3);
+by (rtac (CLAIM "x < y ==> (x::real) <= y") 3);
+by (auto_tac (claset() addIs [real_le_trans],simpset()));
+val lemma_nth_realpow_isLub_le = result();
+
+(*-----------------------
+   Second result we want
+ ----------------------*)
+Goal "[| (0::real) < a; 0 < n; \
+\    isLub (UNIV::real set) \
+\    {x. x ^ n <= a & 0 < x} u |] ==> u ^ n <= a";
+by (forward_tac [lemma_nth_realpow_isLub_le] 1 THEN Step_tac 1);
+by (rtac (LIMSEQ_inverse_real_of_posnat_add_minus_mult RS LIMSEQ_pow 
+    RS LIMSEQ_le_const2) 1);
+by (auto_tac (claset(),simpset() addsimps [real_of_nat_def,real_of_posnat_Suc]));
+qed "realpow_nth_le";
+
+(*----------- The theorem at last! -----------*)
+Goal "[| (0::real) < a; 0 < n |] ==> EX r. r ^ n = a";
+by (forward_tac [nth_realpow_isLub_ex] 1 THEN Auto_tac);
+by (auto_tac (claset() addIs [realpow_nth_le,
+    realpow_nth_ge,real_le_anti_sym],simpset()));
+qed "realpow_nth";
+
+(* positive only *)
+Goal "[| (0::real) < a; 0 < n |] ==> EX r. 0 < r & r ^ n = a";
+by (forward_tac [nth_realpow_isLub_ex] 1 THEN Auto_tac);
+by (auto_tac (claset() addIs [realpow_nth_le,
+    realpow_nth_ge,real_le_anti_sym,
+    lemma_nth_realpow_isLub_gt_zero],simpset()));
+qed "realpow_pos_nth";
+
+Goal "(0::real) < a  ==> EX r. 0 < r & r ^ Suc n = a";
+by (blast_tac (claset() addIs [realpow_pos_nth]) 1);
+qed "realpow_pos_nth2";
+
+(* uniqueness of nth positive root *)
+Goal "[| (0::real) < a; 0 < n |] ==> EX! r. 0 < r & r ^ n = a";
+by (auto_tac (claset() addSIs [realpow_pos_nth],simpset()));
+by (cut_inst_tac [("R1.0","r"),("R2.0","y")] real_linear 1);
+by (Auto_tac);
+by (dres_inst_tac [("x","r")] (conjI RS realpow_less) 1);
+by (dres_inst_tac [("x","y")] (conjI RS realpow_less) 3);
+by (auto_tac (claset(),simpset() addsimps [real_less_not_refl]));
+qed "realpow_pos_nth_unique";
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/NthRoot.thy	Thu Nov 15 16:12:49 2001 +0100
@@ -0,0 +1,8 @@
+(*  Title       : NthRoot.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998  University of Cambridge
+    Description : Existence of nth root. Adapted from
+                   http://www.math.unl.edu/~webnotes
+*)
+
+NthRoot = SEQ + ExtraThms2
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Transcendental.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -0,0 +1,2515 @@
+(*  Title       : Transcendental.ML
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998,1999  University of Cambridge
+                  1999 University of Edinburgh
+    Description : Power Series
+*)
+
+Goalw [root_def] "root (Suc n) 0 = 0";
+by (safe_tac (claset() addSIs [some_equality,realpow_zero] 
+    addSEs [realpow_zero_zero]));
+qed "real_root_zero";
+Addsimps [real_root_zero];
+
+Goalw [root_def]
+     "0 < 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")] realpow_gt_zero 2);
+by (auto_tac (claset(),simpset() addsimps [real_0_less_mult_iff]));
+by (res_inst_tac [("R1.0","u"),("R2.0","x")] real_linear_less2 1);
+by (dres_inst_tac [("n3","n"),("x","u")] 
+    (zero_less_Suc RSN  (3,conjI RSN (2,conjI RS realpow_less))) 1);
+by (dres_inst_tac [("n3","n"),("x","x")] 
+     (zero_less_Suc RSN (3,conjI RSN (2,conjI RS realpow_less))) 4);
+by (auto_tac (claset(),simpset() addsimps [real_less_not_refl])); 
+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 [realpow_gt_zero],simpset() addsimps [real_0_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 [("R1.0","u"),("R2.0","1")] real_linear_less2 1);
+by (dres_inst_tac [("n","n")] realpow_Suc_less_one 1);
+by (dres_inst_tac [("n","n")] realpow_Suc_gt_one 4);
+by (auto_tac (claset(),simpset() addsimps [real_less_not_refl]));
+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 [numeral_0_eq_0, numeral_1_eq_1]
+	                addsimps [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);
+by (stac numeral_2_eq_2 1); 
+by (asm_full_simp_tac (simpset()delsimps [realpow_Suc]) 1); 
+qed "real_pow_sqrt_eq_sqrt_abs_pow";
+
+Goal "0 <= x ==> sqrt(x) ^ 2 = sqrt(abs(x) ^ 2)";
+by (asm_full_simp_tac (simpset() addsimps [real_pow_sqrt_eq_sqrt_abs_pow]) 1); 
+by (stac numeral_2_eq_2 1); 
+by (asm_full_simp_tac (simpset()delsimps [realpow_Suc]) 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 [("R1.0","x"),("R2.0","0")] real_linear 1);
+by Auto_tac;
+by (ftac (real_mult_order) 2);
+by (asm_full_simp_tac (simpset() addsimps [real_0_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 [real_less_asym],
+    simpset() addsimps real_mult_ac@[realpow_mult RS sym,realpow_two_disj,
+    realpow_gt_zero, 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 [real_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 [real_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 [real_0_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,
+    real_0_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 (rtac (real_pow_sqrt_eq_sqrt_pow RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2, abs_mult,abs_ge_zero]));
+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 (forward_tac [real_sqrt_pow2_gt_zero] 1);
+by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2, real_less_not_refl]));
+qed "real_sqrt_not_eq_zero";
+
+Goal "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x";
+by (forward_tac [real_sqrt_not_eq_zero] 1);
+by (cut_inst_tac [("n1","2"),("r1","sqrt x")] (realpow_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_facts_tac [real_zero_less_one RS real_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 Step_tac 1);
+by (res_inst_tac [("N","n"),("c","r")] ratio_test 1);
+by (auto_tac (claset(),simpset() addsimps [abs_mult,real_mult_assoc RS sym]
+    delsimps [fact_Suc]));
+by (rtac real_mult_le_le_mono2 1);
+by (res_inst_tac [("w1","abs x")] (real_mult_commute RS ssubst) 2);
+by (rtac (fact_Suc RS ssubst) 2);
+by (rtac (real_of_nat_mult RS ssubst) 2);
+by (auto_tac (claset(),simpset() addsimps [abs_mult,real_inverse_distrib,
+    abs_ge_zero]));
+by (auto_tac (claset(), simpset() addsimps 
+     [real_mult_assoc RS sym, abs_ge_zero, abs_eqI2,
+      real_inverse_gt_0]));
+by (rtac (CLAIM "x < (y::real) ==> x <= y") 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,
+    real_mult_assoc,abs_inverse]));
+by (rtac real_less_trans 1);
+by (auto_tac (claset(),simpset() addsimps real_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 [realpow_abs,
+    abs_ge_zero,abs_mult,real_0_le_mult_iff]));
+by (auto_tac (claset() addIs [real_mult_le_le_mono2],
+    simpset() addsimps [real_inverse_gt_0,abs_eqI2,abs_ge_zero]));
+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 [realpow_abs,abs_ge_zero,abs_mult,
+    real_0_le_mult_iff]));
+by (auto_tac (claset() addSIs [real_mult_le_le_mono2],
+    simpset() addsimps [real_inverse_gt_0,abs_eqI2,abs_ge_zero]));
+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);
+val lemma_STAR_sin = result();
+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);
+val lemma_STAR_cos = result();
+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);
+val lemma_STAR_cos1 = result();
+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);
+val lemma_STAR_cos2 = result();
+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 (rtac (lemma_realpow_diff RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps real_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 (auto_tac (claset(),simpset() delsimps [sumr_Suc]));
+by (rtac (sumr_Suc RS ssubst) 1);
+by (dtac sym 1);
+by (auto_tac (claset(),simpset() addsimps [lemma_realpow_diff_sumr,
+    real_add_mult_distrib2,real_diff_def] @ 
+    real_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,
+    realpow_add RS sym] delsimps [sumr_Suc]));
+by (res_inst_tac [("c1","x - y")] (real_mult_left_cancel RS iffD1) 1);
+by (rtac (real_minus_minus RS subst) 2);
+by (rtac (real_minus_mult_eq1 RS ssubst) 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|.     *)
+(* ------------------------------------------------------------------------ *)
+
+Goalw [real_divide_def] "1/(x::real) = inverse x";
+by (Simp_tac 1);
+qed "real_divide_eq_inverse";
+
+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) <z; x*z<=y*z |] ==> x<=y" [real_mult_le_cancel1]) 1);
+by (auto_tac (claset(),
+    simpset() addsimps [real_mult_assoc,realpow_abs RS sym]));
+by (dres_inst_tac [("x","0")] spec 2 THEN Force_tac 2); 
+by (auto_tac (claset(),simpset() addsimps [abs_ge_zero,
+    abs_mult,realpow_abs RS sym] @ real_mult_ac));
+by (res_inst_tac [("x2","z ^ n")] (abs_ge_zero RS real_le_imp_less_or_eq
+    RS disjE) 1 THEN dtac sym 2);
+by (auto_tac (claset() addSIs [real_mult_le_le_mono2],
+    simpset() addsimps [real_mult_assoc RS sym,
+    realpow_abs RS sym,summable_def]));
+by (res_inst_tac [("x","K * inverse(1 - (abs(z) * inverse(abs x)))")] exI 1);
+by (auto_tac (claset() addSIs [sums_mult],simpset() addsimps [real_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 [realpow_abs]));
+by (subgoal_tac "x ~= 0" 1);
+by (subgoal_tac "x ~= 0" 3);
+by (auto_tac (claset(),simpset() addsimps 
+    [abs_inverse RS sym,realpow_not_zero,abs_mult 
+     RS sym,realpow_inverse,realpow_mult RS sym]));
+by (auto_tac (claset() addSIs [geometric_sums],simpset() addsimps
+    [realpow_abs RS sym,real_divide_eq_inverse RS sym]));
+by (res_inst_tac [("z","abs(x)")] (CLAIM_SIMP 
+    "[|(0::real)<z; x*z<y*z |] ==> x<y" [real_mult_less_cancel1]) 1);
+by (auto_tac (claset(),simpset() addsimps [abs_ge_zero,
+    abs_mult RS sym,real_mult_assoc]));
+qed "powser_insidea";
+
+Goal "[| summable (%n. f(n) * (x ^ n)); abs(z) < abs(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 [realpow_abs,abs_mult 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 [real_add_mult_distrib2,
+    real_diff_def,realpow_add RS sym] 
+    @ real_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 [real_diff_mult_distrib2]
+    @ real_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,
+    real_diff_mult_distrib2 RS sym,real_mult_assoc] 
+    delsimps [realpow_Suc,sumr_Suc]));
+by (rtac (real_mult_left_cancel RS iffD2) 1);
+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,
+    real_add_mult_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] 
+    @ real_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 (rtac (lemma_termdiff2 RS ssubst) 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_ge_zero,abs_mult]));
+by (case_tac "n" 1 THEN Auto_tac);
+by (dtac less_add_one 1);
+by (auto_tac (claset(),simpset() addsimps [realpow_add,real_add_assoc RS sym,
+    CLAIM_SIMP "(a * b) * c = a * (c * (b::real))" real_mult_ac] 
+    delsimps [sumr_Suc]));
+by (auto_tac (claset() addSIs [real_mult_le_mono],simpset() addsimps 
+    [abs_ge_zero] delsimps [sumr_Suc])); 
+by (auto_tac (claset() addSIs [realpow_le2],simpset() addsimps 
+    [realpow_abs RS sym,abs_ge_zero] 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")] realpow_ge_zero2 2);
+by (auto_tac (claset(),simpset() addsimps 
+    [real_of_nat_le_iff RS sym] 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 [real_mult_le_mono],simpset() addsimps [abs_mult,
+    realpow_add,abs_ge_zero]));
+by (auto_tac (claset() addSIs [realpow_le2,realpow_ge_zero],simpset() addsimps 
+    [realpow_abs RS sym,abs_ge_zero]));
+by (ALLGOALS(arith_tac));
+qed "lemma_termdiff3";
+
+Addsimps [abs_ge_zero];
+
+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*inverse 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) <z; x*z<=y*z |] ==> x<=y" [real_mult_le_cancel1]) 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 [real_inverse_gt_0,
+    real_0_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 [("R2.0","K * e")] real_less_trans 2);
+by (res_inst_tac [("z","inverse K")] (CLAIM_SIMP 
+    "[|(0::real) <z; z*x<z*y |] ==> x<y" [real_mult_less_cancel1]) 3);
+by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 4);
+by Auto_tac;
+by (auto_tac (claset(),simpset() addsimps real_mult_ac));
+qed "lemma_termdiff4";
+
+Goal "[| 0 < k; \
+\           summable f; \
+\           ALL h. 0 < abs(h) & abs(h) < k --> \
+\                   (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 real_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 [real_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] @ real_add_ac @ real_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 [real_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,real_diff_mult_distrib2 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] @ real_add_ac @ real_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,
+     real_add_mult_distrib2] @ real_add_ac @ real_mult_ac) 1);
+(* 46 *)
+by (dtac real_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] @ real_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 real_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);
+(* 75 *)
+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"
+     real_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 real_mult_le_le_mono1 1);
+by (rtac (real_add_commute RS subst) 2);
+by (simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 2);
+by (rtac lemma_termdiff3 2);
+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 (rtac (fact_Suc RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (rtac (real_inverse_distrib RS ssubst) 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 (rtac (fact_Suc RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (rtac (even_Suc RS ssubst) 1);
+by (rtac (real_inverse_distrib RS ssubst) 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 (rtac (fact_Suc RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (rtac (even_Suc RS ssubst) 1);
+by (rtac (real_inverse_distrib RS ssubst) 1);
+by Auto_tac;
+qed "sin_fdiffs2";
+
+(* thms in EvenOdd needed *)
+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 (rtac (fact_Suc RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (rtac (even_Suc RS ssubst) 1);
+by (rtac (real_inverse_distrib RS ssubst) 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_not_even RS sym,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 (rtac (fact_Suc RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (rtac (even_Suc RS ssubst) 1);
+by (rtac (real_inverse_distrib RS ssubst) 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_not_even RS sym,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]));
+val lemma_exp_ext = result();
+
+Goalw [exp_def] "DERIV exp x :> exp(x)";
+by (rtac (lemma_exp_ext RS ssubst) 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]));
+val lemma_sin_ext = result();
+
+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]));
+val lemma_cos_ext = result();
+
+Goalw [cos_def] "DERIV sin x :> cos(x)";
+by (rtac (lemma_sin_ext RS ssubst) 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 (rtac (lemma_cos_ext RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps [lemma_sin_minus,
+    cos_fdiffs2 RS sym,real_minus_mult_eq1]));
+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,realpow_ge_zero,real_0_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 (real_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 (real_minus_mult_eq1 RS subst) 1);
+by (rtac (real_minus_mult_eq2 RS ssubst) 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 real_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 (simpset() delsimps [exp_add_mult_minus] 
+    addsimps real_mult_ac) 1);
+qed "exp_add";
+
+Goal "0 <= exp x";
+by (res_inst_tac [("t","x")] (real_sum_of_halves RS subst) 1);
+by (rtac (exp_add RS ssubst) 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))" [real_le_less]]) 1);
+qed "exp_gt_zero";
+Addsimps [exp_gt_zero];
+
+Goal "0 < inverse(exp x)";
+by (auto_tac (claset() addIs [real_inverse_gt_0],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,
+    real_add_mult_distrib2,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 real_leI, 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];
+
+Goalw [real_le_def] "(exp(x) <= exp(y)) = (x <= y)";
+by (Auto_tac);
+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 [real_le_diff_eq]));
+by (dtac (CLAIM_SIMP "x <= y ==> (0::real) <= y - x" [real_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 [("R1.0","1"),("R2.0","y")] real_linear_less2 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)],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 [("x1","ln x")] (real_add_left_cancel RS iffD1) 1);
+by (auto_tac (claset(),simpset() addsimps [real_inverse_gt_0,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 [real_inverse_gt_0,
+    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];
+
+Goalw [real_le_def] "[| 0 < x; 0 < y|] ==> (ln x <= ln y) = (x <= y)";
+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";
+
+Addsimps [hypreal_less_not_refl];
+
+(* ------------------------------------------------------------------------ *)
+(* 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]));
+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 real_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);
+val lemma_DERIV_subst = result();
+
+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 (rtac (real_add_commute RS ssubst) 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 [("x1","(cos(x) ^ 2)")] (real_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 [("x1","(sin(x) ^ 2)")] (real_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";
+
+Goalw [real_le_def] "abs(sin x) <= 1";
+by (rtac notI 1);
+by (dtac realpow_two_gt_one 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];
+
+Goalw [real_le_def] "abs(cos x) <= 1";
+by (rtac notI 1);
+by (dtac realpow_two_gt_one 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,
+    real_add_mult_distrib,real_add_mult_distrib2] @ 
+    real_mult_ac @ real_add_ac));
+val lemma_DERIV_sin_cos_add = result();
+
+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,
+    real_add_mult_distrib,real_add_mult_distrib2]
+    @ real_mult_ac @ real_add_ac));
+val lemma_DERIV_sin_cos_minus = result();
+
+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);
+val lemma_real_of_nat_six_mult = result();
+
+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 (res_inst_tac [("z","real (Suc (Suc (Suc (Suc 2))))")]
+    (CLAIM "[|(0::real)<z; z*x<z*y |] ==> x<y") 1);
+by (asm_simp_tac (simpset() addsimps [numeral_2_eq_2 RS sym, real_mult_assoc RS sym]) 2);
+by (stac (CLAIM "6 = 2 * (3::real)") 2);
+by (rtac real_mult_less_mono 2);  
+by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc] delsimps [fact_Suc]));
+by (rtac (fact_Suc RS ssubst) 1);
+by (rtac (fact_Suc RS ssubst) 1);
+by (rtac (fact_Suc RS ssubst) 1);
+by (rtac (fact_Suc RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (simp_tac (simpset() addsimps [real_divide_def,
+    real_inverse_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 [real_mult_assoc] 
+    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 (auto_tac (claset(),simpset() addsimps [real_mult_assoc,realpow_gt_zero] 
+    delsimps [fact_Suc]));
+by (rtac real_mult_less_mono 1);
+by (ALLGOALS(Asm_simp_tac));
+by (TRYALL(rtac real_less_trans));
+by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc] delsimps [fact_Suc]));
+by (res_inst_tac [("y","0")] order_less_le_trans 1);
+by (ALLGOALS(Asm_simp_tac));
+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 (auto_tac (claset(),simpset() addsimps [cos_squared_eq,
+    real_minus_add_distrib RS sym, real_minus_zero_less_iff2,sin_gt_zero1,
+    real_add_order,realpow_gt_zero,cos_double] delsimps 
+    [realpow_Suc,real_minus_add_distrib]));
+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 [realpow_gt_zero];
+
+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 [("R2.0",
+    "sumr 0 (Suc (Suc (Suc 0))) (%n. -((- 1) ^ n /(real (fact(2 * n))) \
+\               * 2 ^ (2 * n)))")] real_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 (rtac (fact_Suc RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (rtac (real_of_nat_mult RS ssubst) 1);
+by (rtac real_mult_less_mono 1);
+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 [("R1.0","xa"),("R2.0","y")] real_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 Auto_tac;
+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,real_add_mult_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,real_add_mult_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 (rtac real_less_trans 2 THEN assume_tac 2);
+by Auto_tac;
+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 "[| 0 < x; x < pi/2 |] ==> 0 < cos x";
+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 (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 [("R1.0","x"),("R2.0","0")] real_linear_less2 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 [real_le_less,
+    cos_gt_zero_pi]));
+by Auto_tac;
+qed "cos_ge_zero";
+
+Goal "[| 0 < x; x < pi  |] ==> 0 < sin x";
+by (rtac (sin_cos_eq RS ssubst) 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 [real_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 [("R1.0","xa"),("R2.0","y")] real_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. ~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,
+    real_add_mult_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,
+    real_add_mult_distrib]) 1);
+by Auto_tac;
+qed "cos_zero_lemma";
+
+Goal "[| 0 <= x; sin x = 0 |] ==> \
+\     EX n. 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_not_even RS sym,
+    odd_Suc_mult_two_ex,real_of_nat_Suc,
+    real_add_mult_distrib,real_mult_assoc RS sym]));
+qed "sin_zero_lemma";
+
+(* also spoilt by numeral arithmetic *)
+Goal "(cos x = 0) = \
+\     ((EX n. ~even n & (x = real n * (pi/2))) |   \
+\      (EX n. ~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_not_even RS sym,
+    odd_Suc_mult_two_ex,real_of_nat_Suc,real_add_mult_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. even n & (x = real n * (pi/2))) |   \
+\      (EX n. 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 [real_minus_mult_eq1]) 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() addsimps [real_inverse_distrib RS sym] 
+    @ real_mult_ac));
+by (res_inst_tac [("c1","cos x * cos y")] (real_mult_right_cancel RS subst) 1);
+by (auto_tac (claset(), simpset() addsimps [real_mult_assoc, 
+    real_mult_not_zero,real_diff_mult_distrib,cos_add]));
+val lemma_tan_add1 = result();
+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 [real_mult_assoc,
+    real_mult_not_zero,real_add_mult_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 [real_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,
+    real_inverse_gt_0],simpset()));
+by (rtac (CLAIM "-x < y ==> -y < (x::real)") 1 THEN Auto_tac);
+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 [("z1","1")] ((real_mult_0) RS subst) 1);
+by (rtac LIM_mult2 1); 
+by (rtac ((real_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 (real_inverse_less_iff 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 [real_inverse_distrib,abs_mult]));
+by (dres_inst_tac [("x","cos e")] (real_inverse_gt_0) 1);
+by (dres_inst_tac [("x","inverse (cos e)")] abs_eqI2 1);
+by (auto_tac (claset() addSDs [abs_eqI2],simpset() addsimps real_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()));
+by (rtac (CLAIM "-x < y ==> -y < (x::real)") 1 THEN Auto_tac);
+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 [("R1.0","xa"),("R2.0","y")] real_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 (Auto_tac);
+qed "arctan_bounded";
+
+Goal "- (pi/2) < arctan y";
+by (Auto_tac);
+qed "arctan_lbound";
+
+Goal "arctan y < pi/2";
+by (Auto_tac);
+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,
+    real_add_mult_distrib,real_le_def,
+    real_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 (realpow_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
+    [realpow_mult,real_add_mult_distrib,realpow_divide,
+     tan_def,real_mult_assoc,realpow_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,real_add_mult_distrib,real_add_mult_distrib2]) 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,real_add_mult_distrib,real_add_mult_distrib2]) 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 [real_add_mult_distrib2,
+    sin_add,real_add_mult_distrib] @ real_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,
+    real_add_mult_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,realpow_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 (rtac (real_add_mult_distrib RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps [cos_add] @ real_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 (rtac (real_add_mult_distrib RS ssubst) 1);
+by (auto_tac (claset(),simpset() addsimps [sin_add] @real_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,real_add_mult_distrib2,real_add_mult_distrib,
+    real_minus_mult_eq2]) 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,real_add_mult_distrib,real_add_mult_distrib2]) 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,real_add_mult_distrib,real_add_mult_distrib2]) 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";
+
+(*-------------------------------------------------------------------------------*)
+(* Add continuity checker in backup of theory?                                   *)
+(*-------------------------------------------------------------------------------*)
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hyperreal/Transcendental.thy	Thu Nov 15 16:12:49 2001 +0100
@@ -0,0 +1,49 @@
+(*  Title       : Transcendental.thy
+    Author      : Jacques D. Fleuriot
+    Copyright   : 1998,1999 University of Cambridge
+                  1999 University of Edinburgh
+    Description : Power Series, transcendental functions etc.
+*)
+
+Transcendental = NthRoot + Fact + HSeries + EvenOdd + Lim + 
+
+constdefs
+    root :: [nat,real] => real
+    "root n x == (@u. ((0::real) < x --> 0 < u) & (u ^ n = x))"
+
+    sqrt :: real => real
+    "sqrt x == root 2 x"
+
+    exp :: real => real
+    "exp x == suminf(%n. inverse(real (fact n)) * (x ^ n))"
+
+    sin :: real => real
+    "sin x == suminf(%n. (if even(n) then 0 else
+             ((- 1) ^ ((n - Suc 0) div 2))/(real (fact n))) * x ^ n)"
+ 
+    diffs :: (nat => real) => nat => real
+    "diffs c == (%n. real (Suc n) * c(Suc n))"
+
+    cos :: real => real
+    "cos x == suminf(%n. (if even(n) then ((- 1) ^ (n div 2))/(real (fact n)) 
+                          else 0) * x ^ n)"
+  
+    ln :: real => real
+    "ln x == (@u. exp u = x)"
+
+    pi :: real
+    "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)"
+
+    arcos :: real => real
+    "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)"
+  
+end 
--- a/src/HOL/Integ/IntPower.ML	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/Integ/IntPower.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -39,3 +39,23 @@
 by (Asm_simp_tac 1);
 qed "zpower_zpower";
 
+
+(*** Logical equivalences for inequalities ***)
+
+Goal "(x^n = 0) = (x = (0::int) & 0<n)";
+by (induct_tac "n" 1);
+by Auto_tac; 
+qed "zpower_eq_0_iff";
+Addsimps [zpower_eq_0_iff];
+
+Goal "(0 < (abs x)^n) = (x ~= (0::int) | n=0)";
+by (induct_tac "n" 1);
+by (auto_tac (claset(), simpset() addsimps [int_0_less_mult_iff]));  
+qed "zero_less_zpower_abs_iff";
+Addsimps [zero_less_zpower_abs_iff];
+
+Goal "(0::int) <= (abs x)^n";
+by (induct_tac "n" 1);
+by (auto_tac (claset(), simpset() addsimps [int_0_le_mult_iff]));  
+qed "zero_le_zpower_abs";
+Addsimps [zero_le_zpower_abs];
--- a/src/HOL/Integ/nat_bin.ML	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/Integ/nat_bin.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -27,7 +27,7 @@
 by (simp_tac (simpset() addsimps [numeral_1_eq_1]) 1); 
 qed "numeral_1_eq_Suc_0";
 
-Goalw [nat_number_of_def, One_nat_def] "2 = Suc 1";
+Goalw [nat_number_of_def] "2 = Suc (Suc 0)";
 by (rtac nat_2 1); 
 qed "numeral_2_eq_2";
 
--- a/src/HOL/Integ/nat_simprocs.ML	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/Integ/nat_simprocs.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -227,16 +227,18 @@
     numeral in a coefficient position.
 **)
 
+fun ignore_Sucs (Const ("Suc", _) $ t) = ignore_Sucs t
+  | ignore_Sucs t = t;
+
 fun is_numeral (Const("Numeral.number_of", _) $ w) = true
   | is_numeral _ = false;
 
 fun prod_has_numeral t = exists is_numeral (dest_prod t);
 
 fun restricted_dest_Sucs_sum t = 
-  let val ts = dest_Sucs_sum t
-  in  if exists prod_has_numeral ts then ts
-      else raise TERM("Nat_Numeral_Simprocs.restricted_dest_Sucs_sum", ts)
-  end;
+    if exists prod_has_numeral (dest_sum (ignore_Sucs t))
+    then dest_Sucs_sum t
+    else raise TERM("Nat_Numeral_Simprocs.restricted_dest_Sucs_sum", [t]);
 
 
 (*** Applying CancelNumeralsFun ***)
--- a/src/HOL/IsaMakefile	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/IsaMakefile	Thu Nov 15 16:12:49 2001 +0100
@@ -131,20 +131,23 @@
 
 HOL-Real-Hyperreal: HOL-Real $(LOG)/HOL-Real-Hyperreal.gz
 
-$(LOG)/HOL-Real-Hyperreal.gz: $(OUT)/HOL-Real \
-  Hyperreal/Filter.ML Hyperreal/Filter.thy Hyperreal/HRealAbs.ML \
-  Hyperreal/HRealAbs.thy Hyperreal/HSeries.ML Hyperreal/HSeries.thy \
-  Hyperreal/HyperArith.thy Hyperreal/HyperArith0.ML \
-  Hyperreal/HyperArith0.thy Hyperreal/HyperBin.ML \
-  Hyperreal/HyperBin.thy Hyperreal/HyperDef.ML Hyperreal/HyperDef.thy \
-  Hyperreal/HyperNat.ML Hyperreal/HyperNat.thy Hyperreal/HyperOrd.ML \
-  Hyperreal/HyperOrd.thy Hyperreal/HyperPow.ML Hyperreal/HyperPow.thy \
-  Hyperreal/Hyperreal.thy Hyperreal/Lim.ML Hyperreal/Lim.thy \
-  Hyperreal/NSA.ML Hyperreal/NSA.thy Hyperreal/NatStar.ML \
-  Hyperreal/NatStar.thy Hyperreal/ROOT.ML Hyperreal/SEQ.ML \
-  Hyperreal/SEQ.thy Hyperreal/Series.ML Hyperreal/Series.thy \
-  Hyperreal/Star.ML Hyperreal/Star.thy Hyperreal/Zorn.ML \
-  Hyperreal/Zorn.thy Hyperreal/fuf.ML Hyperreal/hypreal_arith.ML \
+$(LOG)/HOL-Real-Hyperreal.gz: $(OUT)/HOL-Real Hyperreal/ROOT.ML\
+  Hyperreal/EvenOdd.ML Hyperreal/EvenOdd.thy Hyperreal/ExtraThms2.ML\
+  Hyperreal/ExtraThms2.thy Hyperreal/Fact.ML Hyperreal/Fact.thy\
+  Hyperreal/Filter.ML Hyperreal/Filter.thy Hyperreal/HRealAbs.ML\
+  Hyperreal/HRealAbs.thy Hyperreal/HSeries.ML Hyperreal/HSeries.thy\
+  Hyperreal/HyperArith0.ML Hyperreal/HyperArith0.thy Hyperreal/HyperArith.thy\
+  Hyperreal/HyperBin.ML Hyperreal/HyperBin.thy Hyperreal/HyperDef.ML\
+  Hyperreal/HyperDef.thy Hyperreal/HyperNat.ML Hyperreal/HyperNat.thy\
+  Hyperreal/HyperOrd.ML Hyperreal/HyperOrd.thy Hyperreal/HyperPow.ML\
+  Hyperreal/HyperPow.thy Hyperreal/Hyperreal.thy Hyperreal/Lim.ML\
+  Hyperreal/Lim.thy Hyperreal/NatStar.ML Hyperreal/NatStar.thy\
+  Hyperreal/NSA.ML Hyperreal/NSA.thy\
+  Hyperreal/NthRoot.ML Hyperreal/NthRoot.thy\
+  Hyperreal/SEQ.ML Hyperreal/SEQ.thy Hyperreal/Series.ML Hyperreal/Series.thy\
+  Hyperreal/Star.ML Hyperreal/Star.thy Hyperreal/Transcendental.ML\
+  Hyperreal/Transcendental.thy Hyperreal/Zorn.ML Hyperreal/Zorn.thy\
+  Hyperreal/fuf.ML Hyperreal/hypreal_arith.ML \
   Hyperreal/hypreal_arith0.ML
 	@$(ISATOOL) usedir $(OUT)/HOL-Real Hyperreal
 
--- a/src/HOL/Power.ML	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/Power.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -81,8 +81,28 @@
 qed "power_dvd_imp_le";
 
 
+(*** Logical equivalences for inequalities ***)
 
-(*** Binomial Coefficients, following Andy Gordon and Florian Kammueller ***)
+Goal "(x^n = 0) = (x = (0::nat) & 0<n)";
+by (induct_tac "n" 1);
+by Auto_tac; 
+qed "power_eq_0_iff";
+Addsimps [power_eq_0_iff];
+
+Goal "(0 < x^n) = (x ~= (0::nat) | n=0)";
+by (induct_tac "n" 1);
+by Auto_tac; 
+qed "zero_less_power_iff";
+Addsimps [zero_less_power_iff];
+
+Goal "(0::nat) <= x^n";
+by (induct_tac "n" 1);
+by Auto_tac; 
+qed "zero_le_power";
+Addsimps [zero_le_power];
+
+
+(**** Binomial Coefficients, following Andy Gordon and Florian Kammueller ****)
 
 Goal "(n choose 0) = 1";
 by (case_tac "n" 1);
--- a/src/HOL/Real/RealAbs.ML	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/Real/RealAbs.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -65,6 +65,7 @@
 Goalw [real_abs_def] "(0::real)<= abs x";
 by (Simp_tac 1);
 qed "abs_ge_zero";
+Addsimps [abs_ge_zero];
 
 Goalw [real_abs_def] "abs(abs x)=abs (x::real)";
 by (Simp_tac 1);
--- a/src/HOL/Real/RealPow.ML	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/Real/RealPow.ML	Thu Nov 15 16:12:49 2001 +0100
@@ -3,6 +3,8 @@
     Author      : Jacques D. Fleuriot  
     Copyright   : 1998  University of Cambridge
     Description : Natural Powers of reals theory
+
+FIXME: most of the theorems for Suc (Suc 0) are redundant!
 *)
 
 bind_thm ("realpow_Suc", thm "realpow_Suc");
@@ -349,3 +351,24 @@
 by (blast_tac (claset() addIs [realpow_increasing, order_antisym, 
 			       order_eq_refl, sym]) 1);
 qed_spec_mp "realpow_Suc_cancel_eq";
+
+
+(*** Logical equivalences for inequalities ***)
+
+Goal "(x^n = 0) = (x = (0::real) & 0<n)";
+by (induct_tac "n" 1);
+by Auto_tac; 
+qed "realpow_eq_0_iff";
+Addsimps [realpow_eq_0_iff];
+
+Goal "(0 < (abs x)^n) = (x ~= (0::real) | n=0)";
+by (induct_tac "n" 1);
+by (auto_tac (claset(), simpset() addsimps [real_0_less_mult_iff]));  
+qed "zero_less_realpow_abs_iff";
+Addsimps [zero_less_realpow_abs_iff];
+
+Goal "(0::real) <= (abs x)^n";
+by (induct_tac "n" 1);
+by (auto_tac (claset(), simpset() addsimps [real_0_le_mult_iff]));  
+qed "zero_le_realpow_abs";
+Addsimps [zero_le_realpow_abs];
--- a/src/HOL/ex/NatSum.thy	Thu Nov 15 15:07:16 2001 +0100
+++ b/src/HOL/ex/NatSum.thy	Thu Nov 15 16:12:49 2001 +0100
@@ -37,11 +37,10 @@
 *}
 
 lemma sum_of_odd_squares:
-  "3 * (\<Sum>i \<in> {..n(}. Suc (i + i) * Suc (i + i)) =
+  "3 * (\<Sum>i \<in> {..n(}. Suc (2*i) * Suc (2*i)) =
     n * (4 * n * n - 1)"
   apply (induct n)
-   apply (case_tac [2] n)  -- {* eliminates the subtraction *}
-    apply auto
+   apply (auto split: nat_diff_split) (*eliminate the subtraction*)
   done
 
 
@@ -49,12 +48,13 @@
   \medskip The sum of the first @{text n} odd cubes
 *}
 
+lemma numeral_2_eq_2: "2 = Suc (Suc 0)" by (auto ); 
+
 lemma sum_of_odd_cubes:
-  "(\<Sum>i \<in> {..n(}. Suc (i + i) * Suc (i + i) * Suc (i + i)) =
+  "(\<Sum>i \<in> {..n(}. Suc (2*i) * Suc (2*i) * Suc (2*i)) =
     n * n * (2 * n * n - 1)"
   apply (induct n)
-   apply (case_tac [2] n)  -- {* eliminates the subtraction *}
-    apply auto
+   apply (auto split: nat_diff_split) (*eliminate the subtraction*)
   done
 
 text {*
@@ -88,9 +88,9 @@
   "30 * (\<Sum>i \<in> {..n}. i * i * i * i) =
     n * Suc n * Suc (2 * n) * (3 * n * n + 3 * n - 1)"
   apply (induct n)
-   apply auto
-  apply (case_tac n)  -- {* eliminates the subtraction *}
    apply simp_all
+  apply (case_tac n)  -- {* eliminates the subtraction *} 
+   apply (simp_all (no_asm_simp))
   done
 
 text {*