# HG changeset patch # User wenzelm # Date 1485791240 -3600 # Node ID d0e55f85fd8a1b42a61d907b9a3c9a7b5e445289 # Parent ea56dd12deb0d83b58509ab6729aab6ba5ed5d21 misc tuning and modernization; diff -r ea56dd12deb0 -r d0e55f85fd8a src/HOL/ex/NatSum.thy --- a/src/HOL/ex/NatSum.thy Wed Feb 01 13:50:20 2017 +0100 +++ b/src/HOL/ex/NatSum.thy Mon Jan 30 16:47:20 2017 +0100 @@ -1,124 +1,120 @@ -(* Title: HOL/ex/NatSum.thy - Author: Tobias Nipkow +(* Title: HOL/ex/NatSum.thy + Author: Tobias Nipkow *) section \Summing natural numbers\ -theory NatSum imports Main begin +theory NatSum + imports Main +begin text \ Summing natural numbers, squares, cubes, etc. Thanks to Sloane's On-Line Encyclopedia of Integer Sequences, - \<^url>\http://www.research.att.com/~njas/sequences\. + \<^url>\http://oeis.org\. \ lemmas [simp] = ring_distribs diff_mult_distrib diff_mult_distrib2 \\for type nat\ -text \ - \medskip The sum of the first \n\ odd numbers equals \n\ - squared. -\ + +text \\<^medskip> The sum of the first \n\ odd numbers equals \n\ squared.\ lemma sum_of_odds: "(\i=0.. - \medskip The sum of the first \n\ odd squares. -\ +text \\<^medskip> The sum of the first \n\ odd squares.\ lemma sum_of_odd_squares: "3 * (\i=0.. - \medskip The sum of the first \n\ odd cubes -\ +text \\<^medskip> The sum of the first \n\ odd cubes.\ lemma sum_of_odd_cubes: "(\i=0.. - \medskip The sum of the first \n\ positive integers equals - \n (n + 1) / 2\.\ + +text \\<^medskip> The sum of the first \n\ positive integers equals \n (n + 1) / 2\.\ -lemma sum_of_naturals: - "2 * (\i=0..n. i) = n * Suc n" +lemma sum_of_naturals: "2 * (\i=0..n. i) = n * Suc n" + by (induct n) auto + +lemma sum_of_squares: "6 * (\i=0..n. i * i) = n * Suc n * Suc (2 * n)" by (induct n) auto -lemma sum_of_squares: - "6 * (\i=0..n. i * i) = n * Suc n * Suc (2 * n)" +lemma sum_of_cubes: "4 * (\i=0..n. i * i * i) = n * n * Suc n * Suc n" by (induct n) auto -lemma sum_of_cubes: - "4 * (\i=0..n. i * i * i) = n * n * Suc n * Suc n" - by (induct n) auto + +text \\<^medskip> A cute identity:\ -text\\medskip A cute identity:\ - -lemma sum_squared: "(\i=0..n. i)^2 = (\i=0..n::nat. i^3)" -proof(induct n) - case 0 show ?case by simp +lemma sum_squared: "(\i=0..n. i)^2 = (\i=0..n. i^3)" for n :: nat +proof (induct n) + case 0 + show ?case by simp next case (Suc n) have "(\i = 0..Suc n. i)^2 = (\i = 0..n. i^3) + (2*(\i = 0..n. i)*(n+1) + (n+1)^2)" (is "_ = ?A + ?B") - using Suc by(simp add:eval_nat_numeral) + using Suc by (simp add: eval_nat_numeral) also have "?B = (n+1)^3" - using sum_of_naturals by(simp add:eval_nat_numeral) + using sum_of_naturals by (simp add: eval_nat_numeral) also have "?A + (n+1)^3 = (\i=0..Suc n. i^3)" by simp finally show ?case . qed -text \ - \medskip Sum of fourth powers: three versions. -\ + +text \\<^medskip> Sum of fourth powers: three versions.\ lemma sum_of_fourth_powers: "30 * (\i=0..n. i * i * i * i) = n * Suc n * Suc (2 * n) * (3 * n * n + 3 * n - 1)" - apply (induct n) - apply simp_all - apply (case_tac n) \ \eliminates the subtraction\ - apply (simp_all (no_asm_simp)) - done +proof (induct n) + case 0 + show ?case by simp +next + case (Suc n) + then show ?case + by (cases n) \ \eliminates the subtraction\ + simp_all +qed text \ Two alternative proofs, with a change of variables and much more - subtraction, performed using the integers.\ + subtraction, performed using the integers. +\ lemma int_sum_of_fourth_powers: "30 * int (\i=0..i=0.. - \medskip Sums of geometric series: \2\, \3\ and the - general case. -\ +text \\<^medskip> Sums of geometric series: \2\, \3\ and the general case.\ lemma sum_of_2_powers: "(\i=0..i=0.. (k - 1) * (\i=0.. (k - 1) * (\i=0..