| author | wenzelm | 
| Thu, 08 Mar 2012 19:56:57 +0100 | |
| changeset 46837 | 5bdd68f380b3 | 
| parent 40077 | c8a9eaaa2f59 | 
| child 54703 | 499f92dc6e45 | 
| permissions | -rw-r--r-- | 
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
23477diff
changeset | 1 | (* Title: HOL/ex/NatSum.thy | 
| 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
23477diff
changeset | 2 | Author: Tobias Nipkow | 
| 11024 | 3 | *) | 
| 4 | ||
| 5 | header {* Summing natural numbers *}
 | |
| 6 | ||
| 21256 | 7 | theory NatSum imports Main Parity begin | 
| 11024 | 8 | |
| 11786 | 9 | text {*
 | 
| 10 | Summing natural numbers, squares, cubes, etc. | |
| 11 | ||
| 12 | Thanks to Sloane's On-Line Encyclopedia of Integer Sequences, | |
| 13 |   \url{http://www.research.att.com/~njas/sequences/}.
 | |
| 14 | *} | |
| 15 | ||
| 15561 | 16 | lemmas [simp] = | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23431diff
changeset | 17 | ring_distribs | 
| 15561 | 18 |   diff_mult_distrib diff_mult_distrib2 --{*for type nat*}
 | 
| 11024 | 19 | |
| 20 | text {*
 | |
| 12023 | 21 |   \medskip The sum of the first @{text n} odd numbers equals @{text n}
 | 
| 11024 | 22 | squared. | 
| 23 | *} | |
| 24 | ||
| 16593 | 25 | lemma sum_of_odds: "(\<Sum>i=0..<n. Suc (i + i)) = n * n" | 
| 16993 | 26 | by (induct n) auto | 
| 11024 | 27 | |
| 28 | ||
| 29 | text {*
 | |
| 30 |   \medskip The sum of the first @{text n} odd squares.
 | |
| 31 | *} | |
| 32 | ||
| 33 | lemma sum_of_odd_squares: | |
| 15561 | 34 | "3 * (\<Sum>i=0..<n. Suc(2*i) * Suc(2*i)) = n * (4 * n * n - 1)" | 
| 16993 | 35 | by (induct n) auto | 
| 11024 | 36 | |
| 37 | ||
| 38 | text {*
 | |
| 12023 | 39 |   \medskip The sum of the first @{text n} odd cubes
 | 
| 11024 | 40 | *} | 
| 41 | ||
| 42 | lemma sum_of_odd_cubes: | |
| 15561 | 43 | "(\<Sum>i=0..<n. Suc (2*i) * Suc (2*i) * Suc (2*i)) = | 
| 11786 | 44 | n * n * (2 * n * n - 1)" | 
| 16993 | 45 | by (induct n) auto | 
| 11024 | 46 | |
| 47 | text {*
 | |
| 12023 | 48 |   \medskip The sum of the first @{text n} positive integers equals
 | 
| 11024 | 49 |   @{text "n (n + 1) / 2"}.*}
 | 
| 50 | ||
| 11586 | 51 | lemma sum_of_naturals: | 
| 15561 | 52 | "2 * (\<Sum>i=0..n. i) = n * Suc n" | 
| 16993 | 53 | by (induct n) auto | 
| 11024 | 54 | |
| 11586 | 55 | lemma sum_of_squares: | 
| 15561 | 56 | "6 * (\<Sum>i=0..n. i * i) = n * Suc n * Suc (2 * n)" | 
| 16993 | 57 | by (induct n) auto | 
| 11024 | 58 | |
| 11586 | 59 | lemma sum_of_cubes: | 
| 15561 | 60 | "4 * (\<Sum>i=0..n. i * i * i) = n * n * Suc n * Suc n" | 
| 16993 | 61 | by (induct n) auto | 
| 11024 | 62 | |
| 21144 | 63 | text{* \medskip A cute identity: *}
 | 
| 64 | ||
| 65 | lemma sum_squared: "(\<Sum>i=0..n. i)^2 = (\<Sum>i=0..n::nat. i^3)" | |
| 66 | proof(induct n) | |
| 67 | case 0 show ?case by simp | |
| 68 | next | |
| 69 | case (Suc n) | |
| 70 | have "(\<Sum>i = 0..Suc n. i)^2 = | |
| 71 | (\<Sum>i = 0..n. i^3) + (2*(\<Sum>i = 0..n. i)*(n+1) + (n+1)^2)" | |
| 72 | (is "_ = ?A + ?B") | |
| 40077 | 73 | using Suc by(simp add:eval_nat_numeral) | 
| 21144 | 74 | also have "?B = (n+1)^3" | 
| 40077 | 75 | using sum_of_naturals by(simp add:eval_nat_numeral) | 
| 21144 | 76 | also have "?A + (n+1)^3 = (\<Sum>i=0..Suc n. i^3)" by simp | 
| 77 | finally show ?case . | |
| 78 | qed | |
| 11024 | 79 | |
| 80 | text {*
 | |
| 15561 | 81 | \medskip Sum of fourth powers: three versions. | 
| 11024 | 82 | *} | 
| 83 | ||
| 84 | lemma sum_of_fourth_powers: | |
| 15561 | 85 | "30 * (\<Sum>i=0..n. i * i * i * i) = | 
| 11786 | 86 | n * Suc n * Suc (2 * n) * (3 * n * n + 3 * n - 1)" | 
| 11024 | 87 | apply (induct n) | 
| 88 | apply simp_all | |
| 12196 | 89 |   apply (case_tac n)  -- {* eliminates the subtraction *} 
 | 
| 90 | apply (simp_all (no_asm_simp)) | |
| 11024 | 91 | done | 
| 92 | ||
| 93 | text {*
 | |
| 16593 | 94 | Two alternative proofs, with a change of variables and much more | 
| 11024 | 95 | subtraction, performed using the integers. *} | 
| 96 | ||
| 97 | lemma int_sum_of_fourth_powers: | |
| 15561 | 98 | "30 * int (\<Sum>i=0..<m. i * i * i * i) = | 
| 99 | int m * (int m - 1) * (int(2 * m) - 1) * | |
| 100 | (int(3 * m * m) - int(3 * m) - 1)" | |
| 16993 | 101 | by (induct m) (simp_all add: int_mult) | 
| 15561 | 102 | |
| 103 | lemma of_nat_sum_of_fourth_powers: | |
| 104 | "30 * of_nat (\<Sum>i=0..<m. i * i * i * i) = | |
| 15114 | 105 | of_nat m * (of_nat m - 1) * (of_nat (2 * m) - 1) * | 
| 106 | (of_nat (3 * m * m) - of_nat (3 * m) - (1::int))" | |
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
21256diff
changeset | 107 | by (induct m) (simp_all add: of_nat_mult) | 
| 11024 | 108 | |
| 109 | ||
| 110 | text {*
 | |
| 12023 | 111 |   \medskip Sums of geometric series: @{text 2}, @{text 3} and the
 | 
| 11786 | 112 | general case. | 
| 113 | *} | |
| 11024 | 114 | |
| 15561 | 115 | lemma sum_of_2_powers: "(\<Sum>i=0..<n. 2^i) = 2^n - (1::nat)" | 
| 16993 | 116 | by (induct n) (auto split: nat_diff_split) | 
| 11024 | 117 | |
| 15561 | 118 | lemma sum_of_3_powers: "2 * (\<Sum>i=0..<n. 3^i) = 3^n - (1::nat)" | 
| 16993 | 119 | by (induct n) auto | 
| 11024 | 120 | |
| 15561 | 121 | lemma sum_of_powers: "0 < k ==> (k - 1) * (\<Sum>i=0..<n. k^i) = k^n - (1::nat)" | 
| 16993 | 122 | by (induct n) auto | 
| 11024 | 123 | |
| 124 | end |