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