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