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