author | wenzelm |
Fri, 29 Mar 2019 12:24:34 +0100 | |
changeset 70010 | 499896e3a7b0 |
parent 68224 | 1f7308050349 |
permissions | -rw-r--r-- |
64974 | 1 |
(* Title: HOL/ex/NatSum.thy |
2 |
Author: Tobias Nipkow |
|
11024 | 3 |
*) |
4 |
||
61343 | 5 |
section \<open>Summing natural numbers\<close> |
11024 | 6 |
|
64974 | 7 |
theory NatSum |
8 |
imports Main |
|
9 |
begin |
|
11024 | 10 |
|
61343 | 11 |
text \<open> |
11786 | 12 |
Summing natural numbers, squares, cubes, etc. |
13 |
||
14 |
Thanks to Sloane's On-Line Encyclopedia of Integer Sequences, |
|
68224 | 15 |
\<^url>\<open>https://oeis.org\<close>. |
61343 | 16 |
\<close> |
11786 | 17 |
|
15561 | 18 |
lemmas [simp] = |
23477
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents:
23431
diff
changeset
|
19 |
ring_distribs |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
64974
diff
changeset
|
20 |
diff_mult_distrib diff_mult_distrib2 \<comment> \<open>for type nat\<close> |
11024 | 21 |
|
64974 | 22 |
|
23 |
text \<open>\<^medskip> The sum of the first \<open>n\<close> odd numbers equals \<open>n\<close> squared.\<close> |
|
11024 | 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 |
||
64974 | 29 |
text \<open>\<^medskip> The sum of the first \<open>n\<close> odd squares.\<close> |
11024 | 30 |
|
31 |
lemma sum_of_odd_squares: |
|
15561 | 32 |
"3 * (\<Sum>i=0..<n. Suc(2*i) * Suc(2*i)) = n * (4 * n * n - 1)" |
16993 | 33 |
by (induct n) auto |
11024 | 34 |
|
35 |
||
64974 | 36 |
text \<open>\<^medskip> The sum of the first \<open>n\<close> odd cubes.\<close> |
11024 | 37 |
|
38 |
lemma sum_of_odd_cubes: |
|
15561 | 39 |
"(\<Sum>i=0..<n. Suc (2*i) * Suc (2*i) * Suc (2*i)) = |
11786 | 40 |
n * n * (2 * n * n - 1)" |
16993 | 41 |
by (induct n) auto |
11024 | 42 |
|
64974 | 43 |
|
44 |
text \<open>\<^medskip> The sum of the first \<open>n\<close> positive integers equals \<open>n (n + 1) / 2\<close>.\<close> |
|
11024 | 45 |
|
64974 | 46 |
lemma sum_of_naturals: "2 * (\<Sum>i=0..n. i) = n * Suc n" |
47 |
by (induct n) auto |
|
48 |
||
49 |
lemma sum_of_squares: "6 * (\<Sum>i=0..n. i * i) = n * Suc n * Suc (2 * n)" |
|
16993 | 50 |
by (induct n) auto |
11024 | 51 |
|
64974 | 52 |
lemma sum_of_cubes: "4 * (\<Sum>i=0..n. i * i * i) = n * n * Suc n * Suc n" |
16993 | 53 |
by (induct n) auto |
11024 | 54 |
|
64974 | 55 |
|
56 |
text \<open>\<^medskip> A cute identity:\<close> |
|
11024 | 57 |
|
64974 | 58 |
lemma sum_squared: "(\<Sum>i=0..n. i)^2 = (\<Sum>i=0..n. i^3)" for n :: nat |
59 |
proof (induct n) |
|
60 |
case 0 |
|
61 |
show ?case by simp |
|
21144 | 62 |
next |
63 |
case (Suc n) |
|
64 |
have "(\<Sum>i = 0..Suc n. i)^2 = |
|
65 |
(\<Sum>i = 0..n. i^3) + (2*(\<Sum>i = 0..n. i)*(n+1) + (n+1)^2)" |
|
66 |
(is "_ = ?A + ?B") |
|
64974 | 67 |
using Suc by (simp add: eval_nat_numeral) |
21144 | 68 |
also have "?B = (n+1)^3" |
64974 | 69 |
using sum_of_naturals by (simp add: eval_nat_numeral) |
21144 | 70 |
also have "?A + (n+1)^3 = (\<Sum>i=0..Suc n. i^3)" by simp |
71 |
finally show ?case . |
|
72 |
qed |
|
11024 | 73 |
|
64974 | 74 |
|
75 |
text \<open>\<^medskip> Sum of fourth powers: three versions.\<close> |
|
11024 | 76 |
|
77 |
lemma sum_of_fourth_powers: |
|
15561 | 78 |
"30 * (\<Sum>i=0..n. i * i * i * i) = |
11786 | 79 |
n * Suc n * Suc (2 * n) * (3 * n * n + 3 * n - 1)" |
64974 | 80 |
proof (induct n) |
81 |
case 0 |
|
82 |
show ?case by simp |
|
83 |
next |
|
84 |
case (Suc n) |
|
85 |
then show ?case |
|
86 |
by (cases n) \<comment> \<open>eliminates the subtraction\<close> |
|
87 |
simp_all |
|
88 |
qed |
|
11024 | 89 |
|
61343 | 90 |
text \<open> |
16593 | 91 |
Two alternative proofs, with a change of variables and much more |
64974 | 92 |
subtraction, performed using the integers. |
93 |
\<close> |
|
11024 | 94 |
|
95 |
lemma int_sum_of_fourth_powers: |
|
15561 | 96 |
"30 * int (\<Sum>i=0..<m. i * i * i * i) = |
97 |
int m * (int m - 1) * (int(2 * m) - 1) * |
|
98 |
(int(3 * m * m) - int(3 * m) - 1)" |
|
64974 | 99 |
by (induct m) simp_all |
15561 | 100 |
|
101 |
lemma of_nat_sum_of_fourth_powers: |
|
102 |
"30 * of_nat (\<Sum>i=0..<m. i * i * i * i) = |
|
15114 | 103 |
of_nat m * (of_nat m - 1) * (of_nat (2 * m) - 1) * |
104 |
(of_nat (3 * m * m) - of_nat (3 * m) - (1::int))" |
|
64974 | 105 |
by (induct m) simp_all |
11024 | 106 |
|
107 |
||
64974 | 108 |
text \<open>\<^medskip> Sums of geometric series: \<open>2\<close>, \<open>3\<close> and the general case.\<close> |
11024 | 109 |
|
15561 | 110 |
lemma sum_of_2_powers: "(\<Sum>i=0..<n. 2^i) = 2^n - (1::nat)" |
64974 | 111 |
by (induct n) auto |
11024 | 112 |
|
15561 | 113 |
lemma sum_of_3_powers: "2 * (\<Sum>i=0..<n. 3^i) = 3^n - (1::nat)" |
16993 | 114 |
by (induct n) auto |
11024 | 115 |
|
64974 | 116 |
lemma sum_of_powers: "0 < k \<Longrightarrow> (k - 1) * (\<Sum>i=0..<n. k^i) = k^n - 1" |
117 |
for k :: nat |
|
16993 | 118 |
by (induct n) auto |
11024 | 119 |
|
120 |
end |