| 
16732
 | 
     1  | 
(*  Title:      HOL/Binomial.thy
  | 
| 
 | 
     2  | 
    ID:         $Id$
  | 
| 
 | 
     3  | 
    Author:     Lawrence C Paulson
  | 
| 
 | 
     4  | 
    Copyright   1997  University of Cambridge
  | 
| 
 | 
     5  | 
*)
  | 
| 
 | 
     6  | 
  | 
| 
 | 
     7  | 
header{*Binomial Coefficients*}
 | 
| 
 | 
     8  | 
  | 
| 
 | 
     9  | 
theory Binomial
  | 
| 
17508
 | 
    10  | 
imports GCD
  | 
| 
16732
 | 
    11  | 
begin
  | 
| 
 | 
    12  | 
  | 
| 
 | 
    13  | 
text{*This development is based on the work of Andy Gordon and
 | 
| 
 | 
    14  | 
Florian Kammueller*}
  | 
| 
 | 
    15  | 
  | 
| 
 | 
    16  | 
consts
  | 
| 
 | 
    17  | 
  binomial :: "nat \<Rightarrow> nat \<Rightarrow> nat"      (infixl "choose" 65)
  | 
| 
 | 
    18  | 
  | 
| 
 | 
    19  | 
primrec
  | 
| 
 | 
    20  | 
  binomial_0:   "(0     choose k) = (if k = 0 then 1 else 0)"
  | 
| 
 | 
    21  | 
  | 
| 
 | 
    22  | 
  binomial_Suc: "(Suc n choose k) =
  | 
| 
 | 
    23  | 
                 (if k = 0 then 1 else (n choose (k - 1)) + (n choose k))"
  | 
| 
 | 
    24  | 
  | 
| 
 | 
    25  | 
lemma binomial_n_0 [simp]: "(n choose 0) = 1"
  | 
| 
17508
 | 
    26  | 
by (cases n) simp_all
  | 
| 
16732
 | 
    27  | 
  | 
| 
 | 
    28  | 
lemma binomial_0_Suc [simp]: "(0 choose Suc k) = 0"
  | 
| 
 | 
    29  | 
by simp
  | 
| 
 | 
    30  | 
  | 
| 
 | 
    31  | 
lemma binomial_Suc_Suc [simp]:
  | 
| 
 | 
    32  | 
     "(Suc n choose Suc k) = (n choose k) + (n choose Suc k)"
  | 
| 
 | 
    33  | 
by simp
  | 
| 
 | 
    34  | 
  | 
| 
 | 
    35  | 
lemma binomial_eq_0 [rule_format]: "\<forall>k. n < k --> (n choose k) = 0"
  | 
| 
 | 
    36  | 
apply (induct "n", auto)
  | 
| 
 | 
    37  | 
apply (erule allE)
  | 
| 
 | 
    38  | 
apply (erule mp, arith)
  | 
| 
 | 
    39  | 
done
  | 
| 
 | 
    40  | 
  | 
| 
 | 
    41  | 
declare binomial_0 [simp del] binomial_Suc [simp del]
  | 
| 
 | 
    42  | 
  | 
| 
 | 
    43  | 
lemma binomial_n_n [simp]: "(n choose n) = 1"
  | 
| 
 | 
    44  | 
apply (induct "n")
  | 
| 
 | 
    45  | 
apply (simp_all add: binomial_eq_0)
  | 
| 
 | 
    46  | 
done
  | 
| 
 | 
    47  | 
  | 
| 
 | 
    48  | 
lemma binomial_Suc_n [simp]: "(Suc n choose n) = Suc n"
  | 
| 
 | 
    49  | 
by (induct "n", simp_all)
  | 
| 
 | 
    50  | 
  | 
| 
 | 
    51  | 
lemma binomial_1 [simp]: "(n choose Suc 0) = n"
  | 
| 
 | 
    52  | 
by (induct "n", simp_all)
  | 
| 
 | 
    53  | 
  | 
| 
 | 
    54  | 
lemma zero_less_binomial [rule_format]: "k \<le> n --> 0 < (n choose k)"
  | 
| 
 | 
    55  | 
by (rule_tac m = n and n = k in diff_induct, simp_all)
  | 
| 
 | 
    56  | 
  | 
| 
 | 
    57  | 
lemma binomial_eq_0_iff: "(n choose k = 0) = (n<k)"
  | 
| 
 | 
    58  | 
apply (safe intro!: binomial_eq_0)
  | 
| 
 | 
    59  | 
apply (erule contrapos_pp)
  | 
| 
 | 
    60  | 
apply (simp add: zero_less_binomial)
  | 
| 
 | 
    61  | 
done
  | 
| 
 | 
    62  | 
  | 
| 
 | 
    63  | 
lemma zero_less_binomial_iff: "(0 < n choose k) = (k\<le>n)"
  | 
| 
 | 
    64  | 
by (simp add: linorder_not_less [symmetric] binomial_eq_0_iff [symmetric])
  | 
| 
 | 
    65  | 
  | 
| 
 | 
    66  | 
(*Might be more useful if re-oriented*)
  | 
| 
 | 
    67  | 
lemma Suc_times_binomial_eq [rule_format]:
  | 
| 
 | 
    68  | 
     "\<forall>k. k \<le> n --> Suc n * (n choose k) = (Suc n choose Suc k) * Suc k"
  | 
| 
 | 
    69  | 
apply (induct "n")
  | 
| 
 | 
    70  | 
apply (simp add: binomial_0, clarify)
  | 
| 
 | 
    71  | 
apply (case_tac "k")
  | 
| 
 | 
    72  | 
apply (auto simp add: add_mult_distrib add_mult_distrib2 le_Suc_eq
  | 
| 
 | 
    73  | 
                      binomial_eq_0)
  | 
| 
 | 
    74  | 
done
  | 
| 
 | 
    75  | 
  | 
| 
 | 
    76  | 
text{*This is the well-known version, but it's harder to use because of the
 | 
| 
 | 
    77  | 
  need to reason about division.*}
  | 
| 
 | 
    78  | 
lemma binomial_Suc_Suc_eq_times:
  | 
| 
 | 
    79  | 
     "k \<le> n ==> (Suc n choose Suc k) = (Suc n * (n choose k)) div Suc k"
  | 
| 
 | 
    80  | 
by (simp add: Suc_times_binomial_eq div_mult_self_is_m zero_less_Suc
  | 
| 
 | 
    81  | 
        del: mult_Suc mult_Suc_right)
  | 
| 
 | 
    82  | 
  | 
| 
 | 
    83  | 
text{*Another version, with -1 instead of Suc.*}
 | 
| 
 | 
    84  | 
lemma times_binomial_minus1_eq:
  | 
| 
 | 
    85  | 
     "[|k \<le> n;  0<k|] ==> (n choose k) * k = n * ((n - 1) choose (k - 1))"
  | 
| 
 | 
    86  | 
apply (cut_tac n = "n - 1" and k = "k - 1" in Suc_times_binomial_eq)
  | 
| 
 | 
    87  | 
apply (simp split add: nat_diff_split, auto)
  | 
| 
 | 
    88  | 
done
  | 
| 
 | 
    89  | 
  | 
| 
 | 
    90  | 
subsubsection {* Theorems about @{text "choose"} *}
 | 
| 
 | 
    91  | 
  | 
| 
 | 
    92  | 
text {*
 | 
| 
 | 
    93  | 
  \medskip Basic theorem about @{text "choose"}.  By Florian
 | 
| 
 | 
    94  | 
  Kamm\"uller, tidied by LCP.
  | 
| 
 | 
    95  | 
*}
  | 
| 
 | 
    96  | 
  | 
| 
 | 
    97  | 
lemma card_s_0_eq_empty:
  | 
| 
 | 
    98  | 
    "finite A ==> card {B. B \<subseteq> A & card B = 0} = 1"
 | 
| 
 | 
    99  | 
  apply (simp cong add: conj_cong add: finite_subset [THEN card_0_eq])
  | 
| 
 | 
   100  | 
  apply (simp cong add: rev_conj_cong)
  | 
| 
 | 
   101  | 
  done
  | 
| 
 | 
   102  | 
  | 
| 
 | 
   103  | 
lemma choose_deconstruct: "finite M ==> x \<notin> M
  | 
| 
 | 
   104  | 
  ==> {s. s <= insert x M & card(s) = Suc k}
 | 
| 
 | 
   105  | 
       = {s. s <= M & card(s) = Suc k} Un
 | 
| 
 | 
   106  | 
         {s. EX t. t <= M & card(t) = k & s = insert x t}"
 | 
| 
 | 
   107  | 
  apply safe
  | 
| 
 | 
   108  | 
   apply (auto intro: finite_subset [THEN card_insert_disjoint])
  | 
| 
 | 
   109  | 
  apply (drule_tac x = "xa - {x}" in spec)
 | 
| 
 | 
   110  | 
  apply (subgoal_tac "x \<notin> xa", auto)
  | 
| 
 | 
   111  | 
  apply (erule rev_mp, subst card_Diff_singleton)
  | 
| 
 | 
   112  | 
  apply (auto intro: finite_subset)
  | 
| 
 | 
   113  | 
  done
  | 
| 
 | 
   114  | 
  | 
| 
 | 
   115  | 
text{*There are as many subsets of @{term A} having cardinality @{term k}
 | 
| 
 | 
   116  | 
 as there are sets obtained from the former by inserting a fixed element
  | 
| 
 | 
   117  | 
 @{term x} into each.*}
 | 
| 
 | 
   118  | 
lemma constr_bij:
  | 
| 
 | 
   119  | 
   "[|finite A; x \<notin> A|] ==>
  | 
| 
 | 
   120  | 
    card {B. EX C. C <= A & card(C) = k & B = insert x C} =
 | 
| 
 | 
   121  | 
    card {B. B <= A & card(B) = k}"
 | 
| 
 | 
   122  | 
  apply (rule_tac f = "%s. s - {x}" and g = "insert x" in card_bij_eq)
 | 
| 
 | 
   123  | 
       apply (auto elim!: equalityE simp add: inj_on_def)
  | 
| 
 | 
   124  | 
    apply (subst Diff_insert0, auto)
  | 
| 
 | 
   125  | 
   txt {* finiteness of the two sets *}
 | 
| 
 | 
   126  | 
   apply (rule_tac [2] B = "Pow (A)" in finite_subset)
  | 
| 
 | 
   127  | 
   apply (rule_tac B = "Pow (insert x A)" in finite_subset)
  | 
| 
 | 
   128  | 
   apply fast+
  | 
| 
 | 
   129  | 
  done
  | 
| 
 | 
   130  | 
  | 
| 
 | 
   131  | 
text {*
 | 
| 
 | 
   132  | 
  Main theorem: combinatorial statement about number of subsets of a set.
  | 
| 
 | 
   133  | 
*}
  | 
| 
 | 
   134  | 
  | 
| 
 | 
   135  | 
lemma n_sub_lemma:
  | 
| 
 | 
   136  | 
  "!!A. finite A ==> card {B. B <= A & card B = k} = (card A choose k)"
 | 
| 
 | 
   137  | 
  apply (induct k)
  | 
| 
 | 
   138  | 
   apply (simp add: card_s_0_eq_empty, atomize)
  | 
| 
 | 
   139  | 
  apply (rotate_tac -1, erule finite_induct)
  | 
| 
 | 
   140  | 
   apply (simp_all (no_asm_simp) cong add: conj_cong
  | 
| 
 | 
   141  | 
     add: card_s_0_eq_empty choose_deconstruct)
  | 
| 
 | 
   142  | 
  apply (subst card_Un_disjoint)
  | 
| 
 | 
   143  | 
     prefer 4 apply (force simp add: constr_bij)
  | 
| 
 | 
   144  | 
    prefer 3 apply force
  | 
| 
 | 
   145  | 
   prefer 2 apply (blast intro: finite_Pow_iff [THEN iffD2]
  | 
| 
 | 
   146  | 
     finite_subset [of _ "Pow (insert x F)", standard])
  | 
| 
 | 
   147  | 
  apply (blast intro: finite_Pow_iff [THEN iffD2, THEN [2] finite_subset])
  | 
| 
 | 
   148  | 
  done
  | 
| 
 | 
   149  | 
  | 
| 
 | 
   150  | 
theorem n_subsets:
  | 
| 
 | 
   151  | 
    "finite A ==> card {B. B <= A & card B = k} = (card A choose k)"
 | 
| 
 | 
   152  | 
  by (simp add: n_sub_lemma)
  | 
| 
 | 
   153  | 
  | 
| 
 | 
   154  | 
  | 
| 
 | 
   155  | 
text{* The binomial theorem (courtesy of Tobias Nipkow): *}
 | 
| 
 | 
   156  | 
  | 
| 
 | 
   157  | 
theorem binomial: "(a+b::nat)^n = (\<Sum>k=0..n. (n choose k) * a^k * b^(n-k))"
  | 
| 
 | 
   158  | 
proof (induct n)
  | 
| 
 | 
   159  | 
  case 0 thus ?case by simp
  | 
| 
 | 
   160  | 
next
  | 
| 
 | 
   161  | 
  case (Suc n)
  | 
| 
 | 
   162  | 
  have decomp: "{0..n+1} = {0} \<union> {n+1} \<union> {1..n}"
 | 
| 
 | 
   163  | 
    by (auto simp add:atLeastAtMost_def atLeast_def atMost_def)
  | 
| 
 | 
   164  | 
  have decomp2: "{0..n} = {0} \<union> {1..n}"
 | 
| 
 | 
   165  | 
    by (auto simp add:atLeastAtMost_def atLeast_def atMost_def)
  | 
| 
 | 
   166  | 
  have "(a+b::nat)^(n+1) = (a+b) * (\<Sum>k=0..n. (n choose k) * a^k * b^(n-k))"
  | 
| 
 | 
   167  | 
    using Suc by simp
  | 
| 
 | 
   168  | 
  also have "\<dots> =  a*(\<Sum>k=0..n. (n choose k) * a^k * b^(n-k)) +
  | 
| 
 | 
   169  | 
                   b*(\<Sum>k=0..n. (n choose k) * a^k * b^(n-k))"
  | 
| 
 | 
   170  | 
    by(rule nat_distrib)
  | 
| 
 | 
   171  | 
  also have "\<dots> = (\<Sum>k=0..n. (n choose k) * a^(k+1) * b^(n-k)) +
  | 
| 
 | 
   172  | 
                  (\<Sum>k=0..n. (n choose k) * a^k * b^(n-k+1))"
  | 
| 
19279
 | 
   173  | 
    by(simp add: setsum_right_distrib mult_ac)
  | 
| 
16732
 | 
   174  | 
  also have "\<dots> = (\<Sum>k=0..n. (n choose k) * a^k * b^(n+1-k)) +
  | 
| 
 | 
   175  | 
                  (\<Sum>k=1..n+1. (n choose (k - 1)) * a^k * b^(n+1-k))"
  | 
| 
 | 
   176  | 
    by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le
  | 
| 
 | 
   177  | 
             del:setsum_cl_ivl_Suc)
  | 
| 
 | 
   178  | 
  also have "\<dots> = a^(n+1) + b^(n+1) +
  | 
| 
 | 
   179  | 
                  (\<Sum>k=1..n. (n choose (k - 1)) * a^k * b^(n+1-k)) +
  | 
| 
 | 
   180  | 
                  (\<Sum>k=1..n. (n choose k) * a^k * b^(n+1-k))"
  | 
| 
 | 
   181  | 
    by(simp add: decomp2)
  | 
| 
 | 
   182  | 
  also have
  | 
| 
 | 
   183  | 
    "\<dots> = a^(n+1) + b^(n+1) + (\<Sum>k=1..n. (n+1 choose k) * a^k * b^(n+1-k))"
  | 
| 
 | 
   184  | 
    by(simp add: nat_distrib setsum_addf binomial.simps)
  | 
| 
 | 
   185  | 
  also have "\<dots> = (\<Sum>k=0..n+1. (n+1 choose k) * a^k * b^(n+1-k))"
  | 
| 
 | 
   186  | 
    using decomp by simp
  | 
| 
 | 
   187  | 
  finally show ?case by simp
  | 
| 
 | 
   188  | 
qed
  | 
| 
 | 
   189  | 
  | 
| 
 | 
   190  | 
end
  |