Documentation

Mathlib.Data.Nat.Choose.Basic

Binomial coefficients #

This file defines binomial coefficients and proves simple lemmas (i.e. those not requiring more imports).

Main definition and results #

Tags #

binomial coefficient, combination, multicombination, stars and bars

def Nat.choose :

choose n k is the number of k-element subsets in an n-element set. Also known as binomial coefficients.

Equations
Instances For
    @[simp]
    theorem Nat.choose_zero_right (n : ) :
    @[simp]
    theorem Nat.choose_succ_succ' (n : ) (k : ) :
    Nat.choose (n + 1) (k + 1) = Nat.choose n k + Nat.choose n (k + 1)
    theorem Nat.choose_eq_zero_of_lt {n : } {k : } :
    n < kNat.choose n k = 0
    @[simp]
    theorem Nat.choose_self (n : ) :
    @[simp]
    @[simp]
    theorem Nat.choose_one_right (n : ) :
    theorem Nat.triangle_succ (n : ) :
    (n + 1) * (n + 1 - 1) / 2 = n * (n - 1) / 2 + n
    theorem Nat.choose_two_right (n : ) :
    Nat.choose n 2 = n * (n - 1) / 2

    choose n 2 is the n-th triangle number.

    theorem Nat.choose_pos {n : } {k : } :
    k n0 < Nat.choose n k
    theorem Nat.choose_eq_zero_iff {n : } {k : } :
    Nat.choose n k = 0 n < k
    theorem Nat.choose_mul {n : } {k : } {s : } (hkn : k n) (hsk : s k) :
    Nat.choose n k * Nat.choose k s = Nat.choose n s * Nat.choose (n - s) (k - s)
    @[simp]
    theorem Nat.choose_symm {n : } {k : } (hk : k n) :
    Nat.choose n (n - k) = Nat.choose n k
    theorem Nat.choose_symm_of_eq_add {n : } {a : } {b : } (h : n = a + b) :
    theorem Nat.choose_symm_add {a : } {b : } :
    Nat.choose (a + b) a = Nat.choose (a + b) b
    theorem Nat.choose_symm_half (m : ) :
    Nat.choose (2 * m + 1) (m + 1) = Nat.choose (2 * m + 1) m
    theorem Nat.choose_succ_right_eq (n : ) (k : ) :
    Nat.choose n (k + 1) * (k + 1) = Nat.choose n k * (n - k)
    @[simp]
    theorem Nat.choose_succ_self_right (n : ) :
    Nat.choose (n + 1) n = n + 1
    theorem Nat.choose_mul_succ_eq (n : ) (k : ) :
    Nat.choose n k * (n + 1) = Nat.choose (n + 1) k * (n + 1 - k)
    def Nat.fast_choose (n : ) (k : ) :

    A faster implementation of choose, to be used during bytecode evaluation and in compiled code.

    Equations
    Instances For

      Inequalities #

      theorem Nat.choose_le_succ_of_lt_half_left {r : } {n : } (h : r < n / 2) :

      Show that Nat.choose is increasing for small values of the right argument.

      theorem Nat.choose_le_middle (r : ) (n : ) :

      choose n r is maximised when r is n/2.

      Inequalities about increasing the first argument #

      theorem Nat.choose_le_add (a : ) (b : ) (c : ) :
      theorem Nat.choose_le_choose {a : } {b : } (c : ) (h : a b) :
      theorem Nat.choose_mono (b : ) :
      Monotone fun (a : ) => Nat.choose a b

      Multichoose #

      Whereas choose n k is the number of subsets of cardinality k from a type of cardinality n, multichoose n k is the number of multisets of cardinality k from a type of cardinality n.

      Alternatively, whereas choose n k counts the number of combinations, i.e. ways to select k items (up to permutation) from n items without replacement, multichoose n k counts the number of multicombinations, i.e. ways to select k items (up to permutation) from n items with replacement.

      Note that multichoose is not the multinomial coefficient, although it can be computed in terms of multinomial coefficients. For details see https://mathworld.wolfram.com/Multichoose.html

      TODO: Prove that choose (-n) k = (-1)^k * multichoose n k, where choose is the generalized binomial coefficient.

      def Nat.multichoose :

      multichoose n k is the number of multisets of cardinality k from a type of cardinality n.

      Equations
      Instances For
        @[simp]
        theorem Nat.multichoose_succ_succ (n : ) (k : ) :
        Nat.multichoose (n + 1) (k + 1) = Nat.multichoose n (k + 1) + Nat.multichoose (n + 1) k
        @[simp]
        @[simp]
        theorem Nat.multichoose_two (k : ) :
        theorem Nat.multichoose_eq (n : ) (k : ) :