author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 73477 | 1d8a79aa2a99 |
permissions | -rw-r--r-- |
73477 | 1 |
(* Author: Amine Chaieb |
63487 | 2 |
Author: Florian Haftmann |
3 |
Author: Lukas Bulwahn |
|
4 |
Author: Manuel Eberl |
|
5 |
*) |
|
63071 | 6 |
|
63145 | 7 |
section \<open>Stirling numbers of first and second kind\<close> |
63071 | 8 |
|
9 |
theory Stirling |
|
65552
f533820e7248
theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
wenzelm
parents:
64267
diff
changeset
|
10 |
imports Main |
63071 | 11 |
begin |
12 |
||
63145 | 13 |
subsection \<open>Stirling numbers of the second kind\<close> |
63071 | 14 |
|
15 |
fun Stirling :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
63487 | 16 |
where |
17 |
"Stirling 0 0 = 1" |
|
18 |
| "Stirling 0 (Suc k) = 0" |
|
19 |
| "Stirling (Suc n) 0 = 0" |
|
20 |
| "Stirling (Suc n) (Suc k) = Suc k * Stirling n (Suc k) + Stirling n k" |
|
63071 | 21 |
|
63487 | 22 |
lemma Stirling_1 [simp]: "Stirling (Suc n) (Suc 0) = 1" |
63071 | 23 |
by (induct n) simp_all |
24 |
||
63487 | 25 |
lemma Stirling_less [simp]: "n < k \<Longrightarrow> Stirling n k = 0" |
63071 | 26 |
by (induct n k rule: Stirling.induct) simp_all |
27 |
||
63487 | 28 |
lemma Stirling_same [simp]: "Stirling n n = 1" |
63071 | 29 |
by (induct n) simp_all |
30 |
||
63487 | 31 |
lemma Stirling_2_2: "Stirling (Suc (Suc n)) (Suc (Suc 0)) = 2 ^ Suc n - 1" |
63071 | 32 |
proof (induct n) |
63487 | 33 |
case 0 |
34 |
then show ?case by simp |
|
63071 | 35 |
next |
36 |
case (Suc n) |
|
37 |
have "Stirling (Suc (Suc (Suc n))) (Suc (Suc 0)) = |
|
63487 | 38 |
2 * Stirling (Suc (Suc n)) (Suc (Suc 0)) + Stirling (Suc (Suc n)) (Suc 0)" |
39 |
by simp |
|
63071 | 40 |
also have "\<dots> = 2 * (2 ^ Suc n - 1) + 1" |
41 |
by (simp only: Suc Stirling_1) |
|
42 |
also have "\<dots> = 2 ^ Suc (Suc n) - 1" |
|
43 |
proof - |
|
63487 | 44 |
have "(2::nat) ^ Suc n - 1 > 0" |
45 |
by (induct n) simp_all |
|
46 |
then have "2 * ((2::nat) ^ Suc n - 1) > 0" |
|
47 |
by simp |
|
48 |
then have "2 \<le> 2 * ((2::nat) ^ Suc n)" |
|
49 |
by simp |
|
63071 | 50 |
with add_diff_assoc2 [of 2 "2 * 2 ^ Suc n" 1] |
63487 | 51 |
have "2 * 2 ^ Suc n - 2 + (1::nat) = 2 * 2 ^ Suc n + 1 - 2" . |
52 |
then show ?thesis |
|
53 |
by (simp add: nat_distrib) |
|
63071 | 54 |
qed |
55 |
finally show ?case by simp |
|
56 |
qed |
|
57 |
||
63487 | 58 |
lemma Stirling_2: "Stirling (Suc n) (Suc (Suc 0)) = 2 ^ n - 1" |
63071 | 59 |
using Stirling_2_2 by (cases n) simp_all |
60 |
||
63487 | 61 |
|
63145 | 62 |
subsection \<open>Stirling numbers of the first kind\<close> |
63071 | 63 |
|
64 |
fun stirling :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
63487 | 65 |
where |
66 |
"stirling 0 0 = 1" |
|
67 |
| "stirling 0 (Suc k) = 0" |
|
68 |
| "stirling (Suc n) 0 = 0" |
|
69 |
| "stirling (Suc n) (Suc k) = n * stirling n (Suc k) + stirling n k" |
|
63071 | 70 |
|
71 |
lemma stirling_0 [simp]: "n > 0 \<Longrightarrow> stirling n 0 = 0" |
|
72 |
by (cases n) simp_all |
|
73 |
||
63487 | 74 |
lemma stirling_less [simp]: "n < k \<Longrightarrow> stirling n k = 0" |
63071 | 75 |
by (induct n k rule: stirling.induct) simp_all |
76 |
||
63487 | 77 |
lemma stirling_same [simp]: "stirling n n = 1" |
63071 | 78 |
by (induct n) simp_all |
79 |
||
63487 | 80 |
lemma stirling_Suc_n_1: "stirling (Suc n) (Suc 0) = fact n" |
63071 | 81 |
by (induct n) auto |
82 |
||
63487 | 83 |
lemma stirling_Suc_n_n: "stirling (Suc n) n = Suc n choose 2" |
84 |
by (induct n) (auto simp add: numerals(2)) |
|
63071 | 85 |
|
86 |
lemma stirling_Suc_n_2: |
|
87 |
assumes "n \<ge> Suc 0" |
|
88 |
shows "stirling (Suc n) 2 = (\<Sum>k=1..n. fact n div k)" |
|
63487 | 89 |
using assms |
63071 | 90 |
proof (induct n) |
63487 | 91 |
case 0 |
92 |
then show ?case by simp |
|
63071 | 93 |
next |
94 |
case (Suc n) |
|
95 |
show ?case |
|
96 |
proof (cases n) |
|
63487 | 97 |
case 0 |
98 |
then show ?thesis |
|
99 |
by (simp add: numerals(2)) |
|
63071 | 100 |
next |
101 |
case Suc |
|
63487 | 102 |
then have geq1: "Suc 0 \<le> n" |
103 |
by simp |
|
63071 | 104 |
have "stirling (Suc (Suc n)) 2 = Suc n * stirling (Suc n) 2 + stirling (Suc n) (Suc 0)" |
105 |
by (simp only: stirling.simps(4)[of "Suc n"] numerals(2)) |
|
63487 | 106 |
also have "\<dots> = Suc n * (\<Sum>k=1..n. fact n div k) + fact n" |
63071 | 107 |
using Suc.hyps[OF geq1] |
108 |
by (simp only: stirling_Suc_n_1 of_nat_fact of_nat_add of_nat_mult) |
|
63487 | 109 |
also have "\<dots> = Suc n * (\<Sum>k=1..n. fact n div k) + Suc n * fact n div Suc n" |
64240 | 110 |
by (metis nat.distinct(1) nonzero_mult_div_cancel_left) |
63487 | 111 |
also have "\<dots> = (\<Sum>k=1..n. fact (Suc n) div k) + fact (Suc n) div Suc n" |
64267 | 112 |
by (simp add: sum_distrib_left div_mult_swap dvd_fact) |
63487 | 113 |
also have "\<dots> = (\<Sum>k=1..Suc n. fact (Suc n) div k)" |
114 |
by simp |
|
63071 | 115 |
finally show ?thesis . |
116 |
qed |
|
117 |
qed |
|
118 |
||
119 |
lemma of_nat_stirling_Suc_n_2: |
|
120 |
assumes "n \<ge> Suc 0" |
|
121 |
shows "(of_nat (stirling (Suc n) 2)::'a::field_char_0) = fact n * (\<Sum>k=1..n. (1 / of_nat k))" |
|
63487 | 122 |
using assms |
63071 | 123 |
proof (induct n) |
63487 | 124 |
case 0 |
125 |
then show ?case by simp |
|
63071 | 126 |
next |
127 |
case (Suc n) |
|
128 |
show ?case |
|
129 |
proof (cases n) |
|
63487 | 130 |
case 0 |
131 |
then show ?thesis |
|
132 |
by (auto simp add: numerals(2)) |
|
63071 | 133 |
next |
134 |
case Suc |
|
63487 | 135 |
then have geq1: "Suc 0 \<le> n" |
136 |
by simp |
|
63071 | 137 |
have "(of_nat (stirling (Suc (Suc n)) 2)::'a) = |
63487 | 138 |
of_nat (Suc n * stirling (Suc n) 2 + stirling (Suc n) (Suc 0))" |
63071 | 139 |
by (simp only: stirling.simps(4)[of "Suc n"] numerals(2)) |
63487 | 140 |
also have "\<dots> = of_nat (Suc n) * (fact n * (\<Sum>k = 1..n. 1 / of_nat k)) + fact n" |
63071 | 141 |
using Suc.hyps[OF geq1] |
142 |
by (simp only: stirling_Suc_n_1 of_nat_fact of_nat_add of_nat_mult) |
|
63487 | 143 |
also have "\<dots> = fact (Suc n) * (\<Sum>k = 1..n. 1 / of_nat k) + fact (Suc n) * (1 / of_nat (Suc n))" |
63071 | 144 |
using of_nat_neq_0 by auto |
63487 | 145 |
also have "\<dots> = fact (Suc n) * (\<Sum>k = 1..Suc n. 1 / of_nat k)" |
63071 | 146 |
by (simp add: distrib_left) |
147 |
finally show ?thesis . |
|
148 |
qed |
|
149 |
qed |
|
150 |
||
64267 | 151 |
lemma sum_stirling: "(\<Sum>k\<le>n. stirling n k) = fact n" |
63071 | 152 |
proof (induct n) |
153 |
case 0 |
|
63487 | 154 |
then show ?case by simp |
63071 | 155 |
next |
156 |
case (Suc n) |
|
157 |
have "(\<Sum>k\<le>Suc n. stirling (Suc n) k) = stirling (Suc n) 0 + (\<Sum>k\<le>n. stirling (Suc n) (Suc k))" |
|
70113
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents:
68975
diff
changeset
|
158 |
by (simp only: sum.atMost_Suc_shift) |
63487 | 159 |
also have "\<dots> = (\<Sum>k\<le>n. stirling (Suc n) (Suc k))" |
160 |
by simp |
|
161 |
also have "\<dots> = (\<Sum>k\<le>n. n * stirling n (Suc k) + stirling n k)" |
|
162 |
by simp |
|
63071 | 163 |
also have "\<dots> = n * (\<Sum>k\<le>n. stirling n (Suc k)) + (\<Sum>k\<le>n. stirling n k)" |
64267 | 164 |
by (simp add: sum.distrib sum_distrib_left) |
63071 | 165 |
also have "\<dots> = n * fact n + fact n" |
166 |
proof - |
|
167 |
have "n * (\<Sum>k\<le>n. stirling n (Suc k)) = n * ((\<Sum>k\<le>Suc n. stirling n k) - stirling n 0)" |
|
70113
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents:
68975
diff
changeset
|
168 |
by (metis add_diff_cancel_left' sum.atMost_Suc_shift) |
63487 | 169 |
also have "\<dots> = n * (\<Sum>k\<le>n. stirling n k)" |
170 |
by (cases n) simp_all |
|
171 |
also have "\<dots> = n * fact n" |
|
172 |
using Suc.hyps by simp |
|
63071 | 173 |
finally have "n * (\<Sum>k\<le>n. stirling n (Suc k)) = n * fact n" . |
63487 | 174 |
moreover have "(\<Sum>k\<le>n. stirling n k) = fact n" |
175 |
using Suc.hyps . |
|
63071 | 176 |
ultimately show ?thesis by simp |
177 |
qed |
|
178 |
also have "\<dots> = fact (Suc n)" by simp |
|
179 |
finally show ?case . |
|
180 |
qed |
|
181 |
||
182 |
lemma stirling_pochhammer: |
|
63487 | 183 |
"(\<Sum>k\<le>n. of_nat (stirling n k) * x ^ k) = (pochhammer x n :: 'a::comm_semiring_1)" |
184 |
proof (induct n) |
|
185 |
case 0 |
|
186 |
then show ?case by simp |
|
187 |
next |
|
63071 | 188 |
case (Suc n) |
189 |
have "of_nat (n * stirling n 0) = (0 :: 'a)" by (cases n) simp_all |
|
63487 | 190 |
then have "(\<Sum>k\<le>Suc n. of_nat (stirling (Suc n) k) * x ^ k) = |
191 |
(of_nat (n * stirling n 0) * x ^ 0 + |
|
192 |
(\<Sum>i\<le>n. of_nat (n * stirling n (Suc i)) * (x ^ Suc i))) + |
|
193 |
(\<Sum>i\<le>n. of_nat (stirling n i) * (x ^ Suc i))" |
|
70113
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents:
68975
diff
changeset
|
194 |
by (subst sum.atMost_Suc_shift) (simp add: sum.distrib ring_distribs) |
63071 | 195 |
also have "\<dots> = pochhammer x (Suc n)" |
70113
c8deb8ba6d05
Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents:
68975
diff
changeset
|
196 |
by (subst sum.atMost_Suc_shift [symmetric]) |
68406 | 197 |
(simp add: algebra_simps sum.distrib sum_distrib_left pochhammer_Suc flip: Suc) |
63071 | 198 |
finally show ?case . |
63487 | 199 |
qed |
63071 | 200 |
|
201 |
||
202 |
text \<open>A row of the Stirling number triangle\<close> |
|
203 |
||
63487 | 204 |
definition stirling_row :: "nat \<Rightarrow> nat list" |
205 |
where "stirling_row n = [stirling n k. k \<leftarrow> [0..<Suc n]]" |
|
63071 | 206 |
|
207 |
lemma nth_stirling_row: "k \<le> n \<Longrightarrow> stirling_row n ! k = stirling n k" |
|
208 |
by (simp add: stirling_row_def del: upt_Suc) |
|
209 |
||
210 |
lemma length_stirling_row [simp]: "length (stirling_row n) = Suc n" |
|
211 |
by (simp add: stirling_row_def) |
|
212 |
||
213 |
lemma stirling_row_nonempty [simp]: "stirling_row n \<noteq> []" |
|
214 |
using length_stirling_row[of n] by (auto simp del: length_stirling_row) |
|
215 |
||
216 |
||
217 |
subsubsection \<open>Efficient code\<close> |
|
218 |
||
219 |
text \<open> |
|
63487 | 220 |
Naively using the defining equations of the Stirling numbers of the first |
221 |
kind to compute them leads to exponential run time due to repeated |
|
222 |
computations. We can use memoisation to compute them row by row without |
|
223 |
repeating computations, at the cost of computing a few unneeded values. |
|
63071 | 224 |
|
225 |
As a bonus, this is very efficient for applications where an entire row of |
|
63487 | 226 |
Stirling numbers is needed. |
63071 | 227 |
\<close> |
228 |
||
63487 | 229 |
definition zip_with_prev :: "('a \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'b list" |
66656 | 230 |
where "zip_with_prev f x xs = map2 f (x # xs) xs" |
63071 | 231 |
|
232 |
lemma zip_with_prev_altdef: |
|
233 |
"zip_with_prev f x xs = |
|
63487 | 234 |
(if xs = [] then [] else f x (hd xs) # [f (xs!i) (xs!(i+1)). i \<leftarrow> [0..<length xs - 1]])" |
63071 | 235 |
proof (cases xs) |
63487 | 236 |
case Nil |
237 |
then show ?thesis |
|
238 |
by (simp add: zip_with_prev_def) |
|
239 |
next |
|
63071 | 240 |
case (Cons y ys) |
63487 | 241 |
then have "zip_with_prev f x xs = f x (hd xs) # zip_with_prev f y ys" |
63071 | 242 |
by (simp add: zip_with_prev_def) |
243 |
also have "zip_with_prev f y ys = map (\<lambda>i. f (xs ! i) (xs ! (i + 1))) [0..<length xs - 1]" |
|
244 |
unfolding Cons |
|
63487 | 245 |
by (induct ys arbitrary: y) |
68406 | 246 |
(simp_all add: zip_with_prev_def upt_conv_Cons flip: map_Suc_upt del: upt_Suc) |
63487 | 247 |
finally show ?thesis |
248 |
using Cons by simp |
|
249 |
qed |
|
63071 | 250 |
|
251 |
||
63487 | 252 |
primrec stirling_row_aux |
253 |
where |
|
254 |
"stirling_row_aux n y [] = [1]" |
|
255 |
| "stirling_row_aux n y (x#xs) = (y + n * x) # stirling_row_aux n x xs" |
|
63071 | 256 |
|
257 |
lemma stirling_row_aux_correct: |
|
258 |
"stirling_row_aux n y xs = zip_with_prev (\<lambda>a b. a + n * b) y xs @ [1]" |
|
63487 | 259 |
by (induct xs arbitrary: y) (simp_all add: zip_with_prev_def) |
63071 | 260 |
|
261 |
lemma stirling_row_code [code]: |
|
262 |
"stirling_row 0 = [1]" |
|
263 |
"stirling_row (Suc n) = stirling_row_aux n 0 (stirling_row n)" |
|
63487 | 264 |
proof goal_cases |
265 |
case 1 |
|
266 |
show ?case by (simp add: stirling_row_def) |
|
267 |
next |
|
268 |
case 2 |
|
63071 | 269 |
have "stirling_row (Suc n) = |
63487 | 270 |
0 # [stirling_row n ! i + stirling_row n ! (i+1) * n. i \<leftarrow> [0..<n]] @ [1]" |
68975
5ce4d117cea7
A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents:
68406
diff
changeset
|
271 |
proof (rule nth_equalityI, goal_cases length nth) |
63071 | 272 |
case (nth i) |
63487 | 273 |
from nth have "i \<le> Suc n" |
274 |
by simp |
|
275 |
then consider "i = 0 \<or> i = Suc n" | "i > 0" "i \<le> n" |
|
276 |
by linarith |
|
277 |
then show ?case |
|
63071 | 278 |
proof cases |
63487 | 279 |
case 1 |
280 |
then show ?thesis |
|
281 |
by (auto simp: nth_stirling_row nth_append) |
|
282 |
next |
|
283 |
case 2 |
|
284 |
then show ?thesis |
|
285 |
by (cases i) (simp_all add: nth_append nth_stirling_row) |
|
286 |
qed |
|
287 |
next |
|
288 |
case length |
|
289 |
then show ?case by simp |
|
290 |
qed |
|
63071 | 291 |
also have "0 # [stirling_row n ! i + stirling_row n ! (i+1) * n. i \<leftarrow> [0..<n]] @ [1] = |
63487 | 292 |
zip_with_prev (\<lambda>a b. a + n * b) 0 (stirling_row n) @ [1]" |
63071 | 293 |
by (cases n) (auto simp add: zip_with_prev_altdef stirling_row_def hd_map simp del: upt_Suc) |
294 |
also have "\<dots> = stirling_row_aux n 0 (stirling_row n)" |
|
295 |
by (simp add: stirling_row_aux_correct) |
|
63487 | 296 |
finally show ?case . |
297 |
qed |
|
63071 | 298 |
|
299 |
lemma stirling_code [code]: |
|
63487 | 300 |
"stirling n k = |
301 |
(if k = 0 then (if n = 0 then 1 else 0) |
|
302 |
else if k > n then 0 |
|
303 |
else if k = n then 1 |
|
304 |
else stirling_row n ! k)" |
|
63071 | 305 |
by (simp add: nth_stirling_row) |
306 |
||
307 |
end |