diff -r e7df213a1918 -r 0115764233e4 src/HOL/ex/NatSum.thy --- a/src/HOL/ex/NatSum.thy Tue Jun 28 16:12:03 2005 +0200 +++ b/src/HOL/ex/NatSum.thy Tue Jun 28 17:56:04 2005 +0200 @@ -18,7 +18,6 @@ *} lemmas [simp] = - lessThan_Suc atMost_Suc setsum_op_ivl_Suc setsum_cl_ivl_Suc left_distrib right_distrib left_diff_distrib right_diff_distrib --{*for true subtraction*} diff_mult_distrib diff_mult_distrib2 --{*for type nat*} @@ -28,10 +27,8 @@ squared. *} -lemma sum_of_odds: "(\i \ {0..i=0..i=0..i=0..i=0..n. i) = n * Suc n" - apply (induct n) - apply auto - done + by (induct n, auto) lemma sum_of_squares: "6 * (\i=0..n. i * i) = n * Suc n * Suc (2 * n)" - apply (induct n) - apply auto - done + by (induct n, auto) lemma sum_of_cubes: "4 * (\i=0..n. i * i * i) = n * n * Suc n * Suc n" - apply (induct n) - apply auto - done + by (induct n, auto) text {* @@ -95,24 +80,20 @@ done text {* - Tow alternative proofs, with a change of variables and much more + Two alternative proofs, with a change of variables and much more subtraction, performed using the integers. *} lemma int_sum_of_fourth_powers: "30 * int (\i=0..i=0..i=0.. (k - 1) * (\i=0..