Skip to content

Commit 15fee0a

Browse files
grhkm21Parcly-Taxel
andcommitted
feat(Data/Nat/Choose): Lucas' theorem (#8612)
This PR proves Lucas' theorem, which states that $\binom{n}{k} \equiv \prod_i \binom{\lfloor n / p^i \rfloor \pmod{p}}{\lfloor k / p^i \rfloor \pmod{p}} \equiv \prod_i \binom{n_i}{k_i} \pmod{p}$, where $n_i$ and $k_i$ are the $i$-th base-$p$ digits of $n$ and $k$ respectively. This is proven by looking at Freshman's dream closely + induction Co-authored-by: Jeremy Tan Jie Rui <reddeloostw@gmail.com> Co-authored-by: grhkm21 <83517584+grhkm21@users.noreply.github.com>
1 parent add0ea7 commit 15fee0a

File tree

3 files changed

+111
-0
lines changed

3 files changed

+111
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2277,6 +2277,7 @@ import Mathlib.Data.Nat.Choose.Cast
22772277
import Mathlib.Data.Nat.Choose.Central
22782278
import Mathlib.Data.Nat.Choose.Dvd
22792279
import Mathlib.Data.Nat.Choose.Factorization
2280+
import Mathlib.Data.Nat.Choose.Lucas
22802281
import Mathlib.Data.Nat.Choose.Multinomial
22812282
import Mathlib.Data.Nat.Choose.Sum
22822283
import Mathlib.Data.Nat.Choose.Vandermonde

Mathlib/Algebra/CharP/Basic.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -110,6 +110,11 @@ theorem add_pow_char_pow [CommSemiring R] {p : ℕ} [Fact p.Prime] [CharP R p] {
110110
(x + y) ^ p ^ n = x ^ p ^ n + y ^ p ^ n :=
111111
add_pow_char_pow_of_commute _ _ _ (Commute.all _ _)
112112

113+
theorem add_pow_eq_add_pow_mod_mul_pow_add_pow_div
114+
[CommSemiring R] {p : ℕ} [Fact p.Prime] [CharP R p] {n : ℕ} (x y : R) :
115+
(x + y) ^ n = (x + y) ^ (n % p) * (x ^ p + y ^ p) ^ (n / p) := by
116+
rw [← add_pow_char, ← pow_mul, ← pow_add, Nat.mod_add_div]
117+
113118
theorem sub_pow_char [CommRing R] {p : ℕ} [Fact p.Prime] [CharP R p] (x y : R) :
114119
(x - y) ^ p = x ^ p - y ^ p :=
115120
sub_pow_char_of_commute _ _ _ (Commute.all _ _)
@@ -118,6 +123,11 @@ theorem sub_pow_char_pow [CommRing R] {p : ℕ} [Fact p.Prime] [CharP R p] {n :
118123
(x - y) ^ p ^ n = x ^ p ^ n - y ^ p ^ n :=
119124
sub_pow_char_pow_of_commute _ _ _ (Commute.all _ _)
120125

126+
theorem sub_pow_eq_sub_pow_mod_mul_pow_sub_pow_div
127+
[CommRing R] {p : ℕ} [Fact p.Prime] [CharP R p] {n : ℕ} (x y : R) :
128+
(x - y) ^ n = (x - y) ^ (n % p) * (x ^ p - y ^ p) ^ (n / p) := by
129+
rw [← sub_pow_char, ← pow_mul, ← pow_add, Nat.mod_add_div]
130+
121131
theorem CharP.neg_one_pow_char [Ring R] (p : ℕ) [CharP R p] [Fact p.Prime] :
122132
(-1 : R) ^ p = -1 := by
123133
rw [eq_neg_iff_add_eq_zero]

Mathlib/Data/Nat/Choose/Lucas.lean

Lines changed: 100 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,100 @@
1+
/-
2+
Copyright (c) 2023 Gareth Ma. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Gareth Ma
5+
-/
6+
import Mathlib.Data.ZMod.Basic
7+
import Mathlib.RingTheory.Polynomial.Basic
8+
9+
/-!
10+
# Lucas's theorem
11+
12+
This file contains a proof of [Lucas's theorem](https://en.wikipedia.org/wiki/Lucas's_theorem) about
13+
binomial coefficients, which says that for primes `p`, `n` choose `k` is congruent to product of
14+
`n_i` choose `k_i` modulo `p`, where `n_i` and `k_i` are the base-`p` digits of `n` and `k`,
15+
respectively.
16+
17+
## Main statements
18+
19+
* `lucas_theorem`: the binomial coefficient `n choose k` is congruent to the product of `n_i choose
20+
k_i` modulo `p`, where `n_i` and `k_i` are the base-`p` digits of `n` and `k`, respectively.
21+
-/
22+
23+
open Finset hiding choose
24+
25+
open Nat BigOperators Polynomial
26+
27+
namespace Choose
28+
29+
variable {n k p : ℕ} [Fact p.Prime]
30+
31+
/-- For primes `p`, `choose n k` is congruent to `choose (n % p) (k % p) * choose (n / p) (k / p)`
32+
modulo `p`. Also see `choose_modEq_choose_mod_mul_choose_div_nat` for the version with `MOD`. -/
33+
theorem choose_modEq_choose_mod_mul_choose_div :
34+
choose n k ≡ choose (n % p) (k % p) * choose (n / p) (k / p) [ZMOD p] := by
35+
have decompose : ((X : (ZMod p)[X]) + 1) ^ n = (X + 1) ^ (n % p) * (X ^ p + 1) ^ (n / p) := by
36+
simpa using add_pow_eq_add_pow_mod_mul_pow_add_pow_div _ (X : (ZMod p)[X]) 1
37+
simp only [← ZMod.intCast_eq_intCast_iff, Int.cast_mul, Int.cast_ofNat,
38+
← coeff_X_add_one_pow _ n k, ← eq_intCast (Int.castRingHom (ZMod p)), ← coeff_map,
39+
Polynomial.map_pow, Polynomial.map_add, Polynomial.map_one, map_X, decompose]
40+
simp only [add_pow, one_pow, mul_one, ← pow_mul, sum_mul_sum]
41+
conv_lhs =>
42+
enter [1, 2, k, 2, k']
43+
rw [← mul_assoc, mul_right_comm _ _ (X ^ (p * k')), ← pow_add, mul_assoc, ← cast_mul]
44+
have h_iff : ∀ x ∈ range (n % p + 1) ×ˢ range (n / p + 1),
45+
k = x.1 + p * x.2 ↔ (k % p, k / p) = x := by
46+
intro ⟨x₁, x₂⟩ hx
47+
rw [Prod.mk.injEq]
48+
constructor <;> intro h
49+
· simp only [mem_product, mem_range] at hx
50+
have h' : x₁ < p := lt_of_lt_of_le hx.left $ mod_lt _ Fin.size_pos'
51+
rw [h, add_mul_mod_self_left, add_mul_div_left _ _ Fin.size_pos', eq_comm (b := x₂)]
52+
exact ⟨mod_eq_of_lt h', self_eq_add_left.mpr (div_eq_of_lt h')⟩
53+
· rw [← h.left, ← h.right, mod_add_div]
54+
simp only [finset_sum_coeff, coeff_mul_natCast, coeff_X_pow, ite_mul, zero_mul, ← cast_mul]
55+
rw [← sum_product', sum_congr rfl (fun a ha ↦ if_congr (h_iff a ha) rfl rfl), sum_ite_eq]
56+
split_ifs with h
57+
· simp
58+
· rw [mem_product, mem_range, mem_range, not_and_or, lt_succ, not_le, not_lt] at h
59+
cases h <;> simp [choose_eq_zero_of_lt (by tauto)]
60+
61+
/-- For primes `p`, `choose n k` is congruent to `choose (n % p) (k % p) * choose (n / p) (k / p)`
62+
modulo `p`. Also see `choose_modEq_choose_mod_mul_choose_div` for the version with `ZMOD`. -/
63+
theorem choose_modEq_choose_mod_mul_choose_div_nat :
64+
choose n k ≡ choose (n % p) (k % p) * choose (n / p) (k / p) [MOD p] := by
65+
rw [← Int.natCast_modEq_iff]
66+
exact_mod_cast choose_modEq_choose_mod_mul_choose_div
67+
68+
/-- For primes `p`, `choose n k` is congruent to the product of `choose (⌊n / p ^ i⌋ % p)
69+
(⌊k / p ^ i⌋ % p)` over i < a, multiplied by `choose (⌊n / p ^ a⌋) (⌊k / p ^ a⌋)`, modulo `p`. -/
70+
theorem choose_modEq_choose_mul_prod_range_choose (a : ℕ) :
71+
choose n k ≡ choose (n / p ^ a) (k / p ^ a) *
72+
∏ i in range a, choose (n / p ^ i % p) (k / p ^ i % p) [ZMOD p] :=
73+
match a with
74+
| Nat.zero => by simp
75+
| Nat.succ a => (choose_modEq_choose_mul_prod_range_choose a).trans <| by
76+
rw [prod_range_succ, cast_mul, ← mul_assoc, mul_right_comm]
77+
gcongr
78+
apply choose_modEq_choose_mod_mul_choose_div.trans
79+
simp_rw [pow_succ, Nat.div_div_eq_div_mul, mul_comm]
80+
rfl
81+
82+
/-- **Lucas's Theorem**: For primes `p`, `choose n k` is congruent to the product of
83+
`choose (⌊n / p ^ i⌋ % p) (⌊k / p ^ i⌋ % p)` over `i` modulo `p`. -/
84+
theorem choose_modEq_prod_range_choose {a : ℕ} (ha₁ : n < p ^ a) (ha₂ : k < p ^ a) :
85+
choose n k ≡ ∏ i in range a, choose (n / p ^ i % p) (k / p ^ i % p) [ZMOD p] := by
86+
apply (choose_modEq_choose_mul_prod_range_choose a).trans
87+
simp_rw [Nat.div_eq_of_lt ha₁, Nat.div_eq_of_lt ha₂, choose, cast_one, one_mul, cast_prod,
88+
Int.ModEq.refl]
89+
90+
/-- **Lucas's Theorem**: For primes `p`, `choose n k` is congruent to the product of
91+
`choose (⌊n / p ^ i⌋ % p) (⌊k / p ^ i⌋ % p)` over `i` modulo `p`. -/
92+
theorem choose_modEq_prod_range_choose_nat {a : ℕ} (ha₁ : n < p ^ a) (ha₂ : k < p ^ a) :
93+
choose n k ≡ ∏ i in range a, choose (n / p ^ i % p) (k / p ^ i % p) [MOD p] := by
94+
rw [← Int.natCast_modEq_iff]
95+
exact_mod_cast choose_modEq_prod_range_choose ha₁ ha₂
96+
97+
alias lucas_theorem := choose_modEq_prod_range_choose
98+
alias lucas_theorem_nat := choose_modEq_prod_range_choose_nat
99+
100+
end Choose

0 commit comments

Comments
 (0)