Theory Binomial_Int
section ‹Integer Binomial Coefficients›
theory Binomial_Int
imports Complex_Main
begin
lemma upper_le_binomial:
assumes "0 < k" and "k < n"
shows "n ≤ n choose k"
proof -
from assms have "1 ≤ n" by simp
define k' where "k' = (if n div 2 ≤ k then k else n - k)"
from assms have 1: "k' ≤ n - 1" and 2: "n div 2 ≤ k'" by (auto simp: k'_def)
from assms(2) have "k ≤ n" by simp
have "n choose k = n choose k'" by (simp add: k'_def binomial_symmetric[OF ‹k ≤ n›])
have "n = n choose 1" by (simp only: choose_one)
also from ‹1 ≤ n› have "… = n choose (n - 1)" by (rule binomial_symmetric)
also from 1 2 have "… ≤ n choose k'" by (rule binomial_antimono) simp
also have "… = n choose k" by (simp add: k'_def binomial_symmetric[OF ‹k ≤ n›])
finally show ?thesis .
qed
text ‹Restore original sort constraints:›
setup ‹Sign.add_const_constraint (@{const_name gbinomial}, SOME @{typ "'a::{semidom_divide,semiring_char_0} ⇒ nat ⇒ 'a"})›
lemma gbinomial_0_left: "0 gchoose k = (if k = 0 then 1 else 0)"
by (cases k) simp_all
lemma gbinomial_eq_0_int:
assumes "n < k"
shows "(int n) gchoose k = 0"
proof -
have "∃a∈{0..<k}. int n - int a = 0"
proof
show "int n - int n = 0" by simp
next
from assms show "n ∈ {0..<k}" by simp
qed
with finite_atLeastLessThan have eq: "prod (λi. int n - int i) {0..<k} = 0" by (rule prod_zero)
show ?thesis by (simp add: gbinomial_prod_rev eq)
qed
corollary gbinomial_eq_0: "0 ≤ a ⟹ a < int k ⟹ a gchoose k = 0"
by (metis nat_eq_iff2 nat_less_iff gbinomial_eq_0_int)
lemma int_binomial: "int (n choose k) = (int n) gchoose k"
proof (cases "k ≤ n")
case True
from refl have eq: "(∏i = 0..<k. int (n - i)) = (∏i = 0..<k. int n - int i)"
proof (rule prod.cong)
fix i
assume "i ∈ {0..<k}"
with True show "int (n - i) = int n - int i" by simp
qed
show ?thesis
by (simp add: gbinomial_binomial[symmetric] gbinomial_prod_rev zdiv_int eq)
next
case False
thus ?thesis by (simp add: gbinomial_eq_0_int)
qed
lemma falling_fact_pochhammer: "prod (λi. a - int i) {0..<k} = (- 1) ^ k * pochhammer (- a) k"
proof -
have eq: "z ^ Suc n * prod f {0..n} = prod (λx. z * f x) {0..n}" for z::int and n f
by (induct n) (simp_all add: ac_simps)
show ?thesis
proof (cases k)
case 0
thus ?thesis by (simp add: pochhammer_minus)
next
case (Suc n)
thus ?thesis
by (simp only: pochhammer_prod atLeastLessThanSuc_atLeastAtMost
prod.atLeast_Suc_atMost_Suc_shift eq flip: power_mult_distrib) (simp add: of_nat_diff)
qed
qed
lemma falling_fact_pochhammer': "prod (λi. a - int i) {0..<k} = pochhammer (a - int k + 1) k"
by (simp add: falling_fact_pochhammer pochhammer_minus')
lemma gbinomial_int_pochhammer: "(a::int) gchoose k = (- 1) ^ k * pochhammer (- a) k div fact k"
by (simp only: gbinomial_prod_rev falling_fact_pochhammer)
lemma gbinomial_int_pochhammer': "a gchoose k = pochhammer (a - int k + 1) k div fact k"
by (simp only: gbinomial_prod_rev falling_fact_pochhammer')
lemma fact_dvd_pochhammer: "fact k dvd pochhammer (a::int) k"
proof -
have dvd: "y ≠ 0 ⟹ ((of_int (x div y))::'a::field_char_0) = of_int x / of_int y ⟹ y dvd x"
for x y :: int
by (smt dvd_triv_left mult.commute nonzero_eq_divide_eq of_int_eq_0_iff of_int_eq_iff of_int_mult)
show ?thesis
proof (cases "0 < a")
case True
moreover define n where "n = nat (a - 1) + k"
ultimately have a: "a = int n - int k + 1" by simp
from fact_nonzero show ?thesis unfolding a
proof (rule dvd)
have "of_int (pochhammer (int n - int k + 1) k div fact k) = (of_int (int n gchoose k)::rat)"
by (simp only: gbinomial_int_pochhammer')
also have "… = of_int (int (n choose k))" by (simp only: int_binomial)
also have "… = of_nat (n choose k)" by simp
also have "… = (of_nat n) gchoose k" by (fact binomial_gbinomial)
also have "… = pochhammer (of_nat n - of_nat k + 1) k / fact k"
by (fact gbinomial_pochhammer')
also have "… = pochhammer (of_int (int n - int k + 1)) k / fact k" by simp
also have "… = (of_int (pochhammer (int n - int k + 1) k)) / (of_int (fact k))"
by (simp only: of_int_fact pochhammer_of_int)
finally show "of_int (pochhammer (int n - int k + 1) k div fact k) =
of_int (pochhammer (int n - int k + 1) k) / rat_of_int (fact k)" .
qed
next
case False
moreover define n where "n = nat (- a)"
ultimately have a: "a = - int n" by simp
from fact_nonzero have "fact k dvd (-1)^k * pochhammer (- int n) k"
proof (rule dvd)
have "of_int ((-1)^k * pochhammer (- int n) k div fact k) = (of_int (int n gchoose k)::rat)"
by (simp only: gbinomial_int_pochhammer)
also have "… = of_int (int (n choose k))" by (simp only: int_binomial)
also have "… = of_nat (n choose k)" by simp
also have "… = (of_nat n) gchoose k" by (fact binomial_gbinomial)
also have "… = (-1)^k * pochhammer (- of_nat n) k / fact k"
by (fact gbinomial_pochhammer)
also have "… = (-1)^k * pochhammer (of_int (- int n)) k / fact k" by simp
also have "… = (-1)^k * (of_int (pochhammer (- int n) k)) / (of_int (fact k))"
by (simp only: of_int_fact pochhammer_of_int)
also have "… = (of_int ((-1)^k * pochhammer (- int n) k)) / (of_int (fact k))" by simp
finally show "of_int ((- 1) ^ k * pochhammer (- int n) k div fact k) =
of_int ((- 1) ^ k * pochhammer (- int n) k) / rat_of_int (fact k)" .
qed
thus ?thesis unfolding a by (metis dvdI dvd_mult_unit_iff' minus_one_mult_self)
qed
qed
lemma gbinomial_int_negated_upper: "(a gchoose k) = (-1) ^ k * ((int k - a - 1) gchoose k)"
by (simp add: gbinomial_int_pochhammer pochhammer_minus algebra_simps fact_dvd_pochhammer div_mult_swap)
lemma gbinomial_int_mult_fact: "fact k * (a gchoose k) = (∏i = 0..<k. a - int i)"
by (simp only: gbinomial_int_pochhammer' fact_dvd_pochhammer dvd_mult_div_cancel falling_fact_pochhammer')
corollary gbinomial_int_mult_fact': "(a gchoose k) * fact k = (∏i = 0..<k. a - int i)"
using gbinomial_int_mult_fact[of k a] by (simp add: ac_simps)
lemma gbinomial_int_binomial:
"a gchoose k = (if 0 ≤ a then int ((nat a) choose k) else (-1::int)^k * int ((k + (nat (- a)) - 1) choose k))"
by (auto simp: int_binomial gbinomial_int_negated_upper[of a] int_ops(6))
corollary gbinomial_nneg: "0 ≤ a ⟹ a gchoose k = int ((nat a) choose k)"
by (simp add: gbinomial_int_binomial)
corollary gbinomial_neg: "a < 0 ⟹ a gchoose k = (-1::int)^k * int ((k + (nat (- a)) - 1) choose k)"
by (simp add: gbinomial_int_binomial)
lemma of_int_gbinomial: "of_int (a gchoose k) = (of_int a :: 'a::field_char_0) gchoose k"
proof -
have of_int_div: "y dvd x ⟹ of_int (x div y) = of_int x / (of_int y :: 'a)" for x y :: int by auto
show ?thesis
by (simp add: gbinomial_int_pochhammer' gbinomial_pochhammer' of_int_div fact_dvd_pochhammer
pochhammer_of_int[symmetric])
qed
lemma uminus_one_gbinomial [simp]: "(- 1::int) gchoose k = (- 1) ^ k"
by (simp add: gbinomial_int_binomial)
lemma gbinomial_int_Suc_Suc: "(x + 1::int) gchoose (Suc k) = (x gchoose k) + (x gchoose (Suc k))"
proof (rule linorder_cases)
assume 1: "x + 1 < 0"
hence 2: "x < 0" by simp
then obtain n where 3: "nat (- x) = Suc n" using not0_implies_Suc by fastforce
hence 4: "nat (- x - 1) = n" by simp
show ?thesis
proof (cases k)
case 0
show ?thesis by (simp add: ‹k = 0›)
next
case (Suc k')
from 1 2 3 4 show ?thesis by (simp add: ‹k = Suc k'› gbinomial_int_binomial int_distrib(2))
qed
next
assume "x + 1 = 0"
hence "x = - 1" by simp
thus ?thesis by simp
next
assume "0 < x + 1"
hence "0 ≤ x + 1" and "0 ≤ x" and "nat (x + 1) = Suc (nat x)" by simp_all
thus ?thesis by (simp add: gbinomial_int_binomial)
qed
corollary plus_Suc_gbinomial:
"(x + (1 + int k)) gchoose (Suc k) = ((x + int k) gchoose k) + ((x + int k) gchoose (Suc k))"
(is "?l = ?r")
proof -
have "?l = (x + int k + 1) gchoose (Suc k)" by (simp only: ac_simps)
also have "… = ?r" by (fact gbinomial_int_Suc_Suc)
finally show ?thesis .
qed
lemma gbinomial_int_n_n [simp]: "(int n) gchoose n = 1"
proof (induct n)
case 0
show ?case by simp
next
case (Suc n)
have "int (Suc n) gchoose Suc n = (int n + 1) gchoose Suc n" by (simp add: add.commute)
also have "… = (int n gchoose n) + (int n gchoose (Suc n))" by (fact gbinomial_int_Suc_Suc)
finally show ?case by (simp add: Suc gbinomial_eq_0)
qed
lemma gbinomial_int_Suc_n [simp]: "(1 + int n) gchoose n = 1 + int n"
proof (induct n)
case 0
show ?case by simp
next
case (Suc n)
have "1 + int (Suc n) gchoose Suc n = (1 + int n) + 1 gchoose Suc n" by simp
also have "… = (1 + int n gchoose n) + (1 + int n gchoose (Suc n))" by (fact gbinomial_int_Suc_Suc)
also have "… = 1 + int n + (int (Suc n) gchoose (Suc n))" by (simp add: Suc)
also have "… = 1 + int (Suc n)" by (simp only: gbinomial_int_n_n)
finally show ?case .
qed
lemma zbinomial_eq_0_iff [simp]: "a gchoose k = 0 ⟷ (0 ≤ a ∧ a < int k)"
proof
assume a: "a gchoose k = 0"
have 1: "b < int k" if "b gchoose k = 0" for b
proof (rule ccontr)
assume "¬ b < int k"
hence "0 ≤ b" and "k ≤ nat b" by simp_all
from this(1) have "int ((nat b) choose k) = b gchoose k" by (simp add: gbinomial_int_binomial)
also have "… = 0" by (fact that)
finally show False using ‹k ≤ nat b› by simp
qed
show "0 ≤ a ∧ a < int k"
proof
show "0 ≤ a"
proof (rule ccontr)
assume "¬ 0 ≤ a"
hence "(-1) ^ k * ((int k - a - 1) gchoose k) = a gchoose k"
by (simp add: gbinomial_int_negated_upper[of a])
also have "… = 0" by (fact a)
finally have "(int k - a - 1) gchoose k = 0" by simp
hence "int k - a - 1 < int k" by (rule 1)
with ‹¬ 0 ≤ a› show False by simp
qed
next
from a show "a < int k" by (rule 1)
qed
qed (auto intro: gbinomial_eq_0)
subsection ‹Sums›
lemma gchoose_rising_sum_nat: "(∑j≤n. int j + int k gchoose k) = (int n + int k + 1) gchoose (Suc k)"
proof -
have "(∑j≤n. int j + int k gchoose k) = int (∑j≤n. k + j choose k)"
by (simp add: int_binomial add.commute)
also have "(∑j≤n. k + j choose k) = (k + n + 1) choose (k + 1)" by (fact choose_rising_sum(1))
also have "int … = (int n + int k + 1) gchoose (Suc k)"
by (simp add: int_binomial ac_simps del: binomial_Suc_Suc)
finally show ?thesis .
qed
lemma gchoose_rising_sum:
assumes "0 ≤ n"
shows "(∑j=0..n. j + int k gchoose k) = (n + int k + 1) gchoose (Suc k)"
proof -
from _ refl have "(∑j=0..n. j + int k gchoose k) = (∑j∈int ` {0..nat n}. j + int k gchoose k)"
proof (rule sum.cong)
from assms show "{0..n} = int ` {0..nat n}" by (simp add: image_int_atLeastAtMost)
qed
also have "… = (∑j≤nat n. int j + int k gchoose k)" by (simp add: sum.reindex atMost_atLeast0)
also have "… = (int (nat n) + int k + 1) gchoose (Suc k)" by (fact gchoose_rising_sum_nat)
also from assms have "… = (n + int k + 1) gchoose (Suc k)" by (simp add: add.assoc add.commute)
finally show ?thesis .
qed
subsection ‹Inequalities›
lemma binomial_mono:
assumes "m ≤ n"
shows "m choose k ≤ n choose k"
proof -
define l where "l = n - m"
with assms have n: "n = m + l" by simp
have "m choose k ≤ (m + l) choose k"
proof (induct l)
case 0
show ?case by simp
next
case *: (Suc l)
show ?case
proof (cases k)
case 0
thus ?thesis by simp
next
case k: (Suc k0)
note *
also have "m + l choose k ≤ m + l choose k + (m + l choose k0)" by simp
also have "… = m + Suc l choose k" by (simp add: k)
finally show ?thesis .
qed
qed
thus ?thesis by (simp only: n)
qed
lemma binomial_plus_le:
assumes "0 < k"
shows "(m choose k) + (n choose k) ≤ (m + n) choose k"
proof -
define k0 where "k0 = k - 1"
with assms have k: "k = Suc k0" by simp
show ?thesis unfolding k
proof (induct n)
case 0
show ?case by simp
next
case (Suc n)
have "m choose Suc k0 + (Suc n choose Suc k0) = m choose Suc k0 + (n choose Suc k0) + (n choose k0)"
by (simp only: binomial_Suc_Suc)
also from Suc have "… ≤ (m + n) choose Suc k0 + ((m + n) choose k0)"
proof (rule add_mono)
have "n ≤ m + n" by simp
thus "n choose k0 ≤ m + n choose k0" by (rule binomial_mono)
qed
also have "… = m + Suc n choose Suc k0" by simp
finally show ?case .
qed
qed
lemma binomial_ineq_1: "2 * ((n + i) choose k) ≤ n choose k + ((n + 2 * i) choose k)"
proof (cases k)
case 0
thus ?thesis by simp
next
case k: (Suc k0)
show ?thesis unfolding k
proof (induct i)
case 0
thus ?case by simp
next
case (Suc i)
have "2 * (n + Suc i choose Suc k0) = 2 * (n + i choose k0) + 2 * (n + i choose Suc k0)"
by simp
also have "… ≤ (n + 2 * i choose k0 + (Suc (n + 2 * i) choose k0)) + (n choose Suc k0 + (n + 2 * i choose Suc k0))"
proof (rule add_mono)
have "n + i choose k0 ≤ n + 2 * i choose k0" by (rule binomial_mono) simp
moreover have "n + 2 * i choose k0 ≤ Suc (n + 2 * i) choose k0" by (rule binomial_mono) simp
ultimately show "2 * (n + i choose k0) ≤ n + 2 * i choose k0 + (Suc (n + 2 * i) choose k0)"
by simp
qed (fact Suc)
also have "… = n choose Suc k0 + (n + 2 * Suc i choose Suc k0)" by simp
finally show ?case .
qed
qed
lemma gbinomial_int_nonneg:
assumes "0 ≤ (x::int)"
shows "0 ≤ x gchoose k"
proof -
have "0 ≤ int (nat x choose k)" by simp
also from assms have "… = x gchoose k" by (simp add: int_binomial)
finally show ?thesis .
qed
lemma gbinomial_int_mono:
assumes "0 ≤ x" and "x ≤ (y::int)"
shows "x gchoose k ≤ y gchoose k"
proof -
from assms have "nat x ≤ nat y" by simp
hence "nat x choose k ≤ nat y choose k" by (rule binomial_mono)
hence "int (nat x choose k) ≤ int (nat y choose k)" by (simp only: zle_int)
hence "int (nat x) gchoose k ≤ int (nat y) gchoose k" by (simp only: int_binomial)
with assms show ?thesis by simp
qed
lemma gbinomial_int_plus_le:
assumes "0 < k" and "0 ≤ x" and "0 ≤ (y::int)"
shows "(x gchoose k) + (y gchoose k) ≤ (x + y) gchoose k"
proof -
from assms(1) have "nat x choose k + (nat y choose k) ≤ nat x + nat y choose k"
by (rule binomial_plus_le)
hence "int (nat x choose k + (nat y choose k)) ≤ int (nat x + nat y choose k)"
by (simp only: zle_int)
hence "int (nat x) gchoose k + (int (nat y) gchoose k) ≤ int (nat x) + int (nat y) gchoose k"
by (simp only: int_plus int_binomial)
with assms(2, 3) show ?thesis by simp
qed
lemma binomial_int_ineq_1:
assumes "0 ≤ x" and "0 ≤ (y::int)"
shows "2 * (x + y gchoose k) ≤ x gchoose k + ((x + 2 * y) gchoose k)"
proof -
from binomial_ineq_1[of "nat x" "nat y" k]
have "int (2 * (nat x + nat y choose k)) ≤ int (nat x choose k + (nat x + 2 * nat y choose k))"
by (simp only: zle_int)
hence "2 * (int (nat x) + int (nat y) gchoose k) ≤ int (nat x) gchoose k + (int (nat x) + 2 * int (nat y) gchoose k)"
by (simp only: int_binomial int_plus int_ops(7)) simp
with assms show ?thesis by simp
qed
corollary binomial_int_ineq_2:
assumes "0 ≤ y" and "y ≤ (x::int)"
shows "2 * (x gchoose k) ≤ x - y gchoose k + (x + y gchoose k)"
proof -
from assms(2) have "0 ≤ x - y" by simp
hence "2 * ((x - y) + y gchoose k) ≤ x - y gchoose k + ((x - y + 2 * y) gchoose k)"
using assms(1) by (rule binomial_int_ineq_1)
thus ?thesis by smt
qed
corollary binomial_int_ineq_3:
assumes "0 ≤ y" and "y ≤ 2 * (x::int)"
shows "2 * (x gchoose k) ≤ y gchoose k + (2 * x - y gchoose k)"
proof (cases "y ≤ x")
case True
hence "0 ≤ x - y" by simp
moreover from assms(1) have "x - y ≤ x" by simp
ultimately have "2 * (x gchoose k) ≤ x - (x - y) gchoose k + (x + (x - y) gchoose k)"
by (rule binomial_int_ineq_2)
thus ?thesis by simp
next
case False
hence "0 ≤ y - x" by simp
moreover from assms(2) have "y - x ≤ x" by simp
ultimately have "2 * (x gchoose k) ≤ x - (y - x) gchoose k + (x + (y - x) gchoose k)"
by (rule binomial_int_ineq_2)
thus ?thesis by simp
qed
subsection ‹Backward Difference Operator›
definition bw_diff :: "('a ⇒ 'a) ⇒ 'a ⇒ 'a::{ab_group_add,one}"
where "bw_diff f x = f x - f (x - 1)"
lemma bw_diff_const [simp]: "bw_diff (λ_. c) = (λ_. 0)"
by (rule ext) (simp add: bw_diff_def)
lemma bw_diff_id [simp]: "bw_diff (λx. x) = (λ_. 1)"
by (rule ext) (simp add: bw_diff_def)
lemma bw_diff_plus [simp]: "bw_diff (λx. f x + g x) = (λx. bw_diff f x + bw_diff g x)"
by (rule ext) (simp add: bw_diff_def)
lemma bw_diff_uminus [simp]: "bw_diff (λx. - f x) = (λx. - bw_diff f x)"
by (rule ext) (simp add: bw_diff_def)
lemma bw_diff_minus [simp]: "bw_diff (λx. f x - g x) = (λx. bw_diff f x - bw_diff g x)"
by (rule ext) (simp add: bw_diff_def)
lemma bw_diff_const_pow: "(bw_diff ^^ k) (λ_. c) = (if k = 0 then λ_. c else (λ_. 0))"
by (induct k, simp_all)
lemma bw_diff_id_pow:
"(bw_diff ^^ k) (λx. x) = (if k = 0 then (λx. x) else if k = 1 then (λ_. 1) else (λ_. 0))"
by (induct k, simp_all)
lemma bw_diff_plus_pow [simp]:
"(bw_diff ^^ k) (λx. f x + g x) = (λx. (bw_diff ^^ k) f x + (bw_diff ^^ k) g x)"
by (induct k, simp_all)
lemma bw_diff_uminus_pow [simp]: "(bw_diff ^^ k) (λx. - f x) = (λx. - (bw_diff ^^ k) f x)"
by (induct k, simp_all)
lemma bw_diff_minus_pow [simp]:
"(bw_diff ^^ k) (λx. f x - g x) = (λx. (bw_diff ^^ k) f x - (bw_diff ^^ k) g x)"
by (induct k, simp_all)
lemma bw_diff_sum_pow [simp]:
"(bw_diff ^^ k) (λx. (∑i∈I. f i x)) = (λx. (∑i∈I. (bw_diff ^^ k) (f i) x))"
by (induct I rule: infinite_finite_induct, simp_all add: bw_diff_const_pow)
lemma bw_diff_gbinomial:
assumes "0 < k"
shows "bw_diff (λx::int. (x + n) gchoose k) = (λx. (x + n - 1) gchoose (k - 1))"
proof (rule ext)
fix x::int
from assms have eq: "Suc (k - Suc 0) = k" by simp
have "x + n gchoose k = (x + n - 1) + 1 gchoose (Suc (k - 1))" by (simp add: eq)
also have "… = (x + n - 1) gchoose (k - 1) + ((x + n - 1) gchoose (Suc (k - 1)))"
by (fact gbinomial_int_Suc_Suc)
finally show "bw_diff (λx. x + n gchoose k) x = x + n - 1 gchoose (k - 1)"
by (simp add: eq bw_diff_def algebra_simps)
qed
lemma bw_diff_gbinomial_pow:
"(bw_diff ^^ l) (λx::int. (x + n) gchoose k) =
(if l ≤ k then (λx. (x + n - int l) gchoose (k - l)) else (λ_. 0))"
proof -
have *: "l0 ≤ k ⟹ (bw_diff ^^ l0) (λx::int. (x + n) gchoose k) = (λx. (x + n - int l0) gchoose (k - l0))"
for l0
proof (induct l0)
case 0
show ?case by simp
next
case (Suc l0)
from Suc.prems have "0 < k - l0" and "l0 ≤ k" by simp_all
from this(2) have eq: "(bw_diff ^^ l0) (λx. x + n gchoose k) = (λx. x + n - int l0 gchoose (k - l0))"
by (rule Suc.hyps)
have "(bw_diff ^^ Suc l0) (λx. x + n gchoose k) = bw_diff (λx. x + (n - int l0) gchoose (k - l0))"
by (simp add: eq algebra_simps)
also from ‹0 < k - l0› have "… = (λx. (x + (n - int l0) - 1) gchoose (k - l0 - 1))"
by (rule bw_diff_gbinomial)
also have "… = (λx. x + n - int (Suc l0) gchoose (k - Suc l0))" by (simp add: algebra_simps)
finally show ?case .
qed
show ?thesis
proof (simp add: * split: if_split, intro impI)
assume "¬ l ≤ k"
hence "(l - k) + k = l" and "l - k ≠ 0" by simp_all
hence "(bw_diff ^^ l) (λx. x + n gchoose k) = (bw_diff ^^ ((l - k) + k)) (λx. x + n gchoose k)"
by (simp only:)
also have "… = (bw_diff ^^ (l - k)) (λ_. 1)" by (simp add: * funpow_add)
also from ‹l - k ≠ 0› have "… = (λ_. 0)" by (simp add: bw_diff_const_pow)
finally show "(bw_diff ^^ l) (λx. x + n gchoose k) = (λ_. 0)" .
qed
qed
end