Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 233eff0

Browse files
committed
feat(data/fintype/card_embedding): the birthday problem (#7363)
Co-authored-by: Eric <37984851+ericrbg@users.noreply.github.com>
1 parent 1924742 commit 233eff0

File tree

9 files changed

+292
-7
lines changed

9 files changed

+292
-7
lines changed
Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,22 @@
1+
/-
2+
Copyright (c) 2021 Eric Rodriguez. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Eric Rodriguez
5+
-/
6+
import data.fintype.card_embedding
7+
8+
/-!
9+
# Birthday Problem
10+
11+
This file proves Theorem 93 from the [100 Theorems List](https://www.cs.ru.nl/~freek/100/).
12+
13+
As opposed to the standard probabilistic statement, we instead state the birthday problem
14+
in terms of injective functions. The general result about `fintype.card (α ↪ β)` which this proof
15+
uses is `fintype.card_embedding`.
16+
-/
17+
18+
local notation `‖` x `‖` := fintype.card x
19+
20+
theorem birthday :
21+
2 * ‖fin 23 ↪ fin 365‖ < ‖fin 23 → fin 365‖ ∧ 2 * ‖fin 22 ↪ fin 365‖ > ‖fin 22 → fin 365‖ :=
22+
by split; norm_num [nat.desc_fac]

docs/100.yaml

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -343,6 +343,10 @@
343343
title : Pick’s Theorem
344344
93:
345345
title : The Birthday Problem
346+
decl : fintype.card_embedding
347+
links :
348+
mathlib archive: https://github.com/leanprover-community/mathlib/blob/master/archive/100-theorems-list/93_birthday_problem.lean
349+
author : Eric Rodriguez
346350
94:
347351
title : The Law of Cosines
348352
decl : euclidean_geometry.dist_sq_eq_dist_sq_add_dist_sq_sub_two_mul_dist_mul_dist_mul_cos_angle

src/data/equiv/basic.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1418,6 +1418,14 @@ def subtype_prod_equiv_prod {α : Type u} {β : Type v} {p : α → Prop} {q :
14181418
λ ⟨⟨_, _⟩, ⟨_, _⟩⟩, rfl,
14191419
λ ⟨⟨_, _⟩, ⟨_, _⟩⟩, rfl⟩
14201420

1421+
/-- A subtype of a `prod` is equivalent to a sigma type whose fibers are subtypes. -/
1422+
def subtype_prod_equiv_sigma_subtype {α β : Type*} (p : α → β → Prop) :
1423+
{x : α × β // p x.1 x.2} ≃ Σ a, {b : β // p a b} :=
1424+
{ to_fun := λ x, ⟨x.1.1, x.1.2, x.prop⟩,
1425+
inv_fun := λ x, ⟨⟨x.1, x.2⟩, x.2.prop⟩,
1426+
left_inv := λ x, by ext; refl,
1427+
right_inv := λ ⟨a, b, pab⟩, rfl }
1428+
14211429
end
14221430

14231431
section subtype_equiv_codomain

src/data/equiv/embedding.lean

Lines changed: 85 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,85 @@
1+
/-
2+
Copyright (c) 2021 Eric Rodriguez. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Eric Rodriguez
5+
-/
6+
import logic.embedding
7+
import data.set.lattice
8+
9+
/-!
10+
# Equivalences on embeddings
11+
12+
This file shows some advanced equivalences on embeddings, useful for constructing larger
13+
embeddings from smaller ones.
14+
-/
15+
16+
open function.embedding
17+
18+
namespace equiv
19+
20+
/-- Embeddings from a sum type are equivalent to two separate embeddings with disjoint ranges. -/
21+
def sum_embedding_equiv_prod_embedding_disjoint {α β γ : Type*} :
22+
((α ⊕ β) ↪ γ) ≃ {f : (α ↪ γ) × (β ↪ γ) // disjoint (set.range f.1) (set.range f.2)} :=
23+
{ to_fun := λ f, ⟨(inl.trans f, inr.trans f),
24+
begin
25+
rintros _ ⟨⟨a, h⟩, ⟨b, rfl⟩⟩,
26+
simp only [trans_apply, inl_apply, inr_apply] at h,
27+
have : sum.inl a = sum.inr b := f.injective h,
28+
simp only at this,
29+
assumption
30+
end⟩,
31+
inv_fun := λ ⟨⟨f, g⟩, disj⟩,
32+
⟨λ x, match x with
33+
| (sum.inl a) := f a
34+
| (sum.inr b) := g b
35+
end,
36+
begin
37+
rintros (a₁|b₁) (a₂|b₂) f_eq;
38+
simp only [equiv.coe_fn_symm_mk, sum.elim_inl, sum.elim_inr] at f_eq,
39+
{ rw f.injective f_eq },
40+
{ simp! only at f_eq, exfalso, exact disj ⟨⟨a₁, by simp⟩, ⟨b₂, by simp [f_eq]⟩⟩ },
41+
{ simp! only at f_eq, exfalso, exact disj ⟨⟨a₂, by simp⟩, ⟨b₁, by simp [f_eq]⟩⟩ },
42+
{ rw g.injective f_eq }
43+
end⟩,
44+
left_inv := λ f, by { dsimp only, ext, cases x; simp! },
45+
right_inv := λ ⟨⟨f, g⟩, _⟩, by { simp only [prod.mk.inj_iff], split; ext; simp! } }
46+
47+
/-- Embeddings whose range lies within a set are equivalent to embeddings to that set.
48+
This is `function.embedding.cod_restrict` as an equiv. -/
49+
def cod_restrict (α : Type*) {β : Type*} (bs : set β) :
50+
{f : α ↪ β // ∀ a, f a ∈ bs} ≃ (α ↪ bs) :=
51+
{ to_fun := λ f, (f : α ↪ β).cod_restrict bs f.prop,
52+
inv_fun := λ f, ⟨f.trans (function.embedding.subtype _), λ a, (f a).prop⟩,
53+
left_inv := λ x, by ext; refl,
54+
right_inv := λ x, by ext; refl }
55+
56+
/-- Pairs of embeddings with disjoint ranges are equivalent to a dependent sum of embeddings,
57+
in which the second embedding cannot take values in the range of the first. -/
58+
def prod_embedding_disjoint_equiv_sigma_embedding_restricted {α β γ : Type*} :
59+
{f : (α ↪ γ) × (β ↪ γ) // disjoint (set.range f.1) (set.range f.2)} ≃
60+
(Σ f : α ↪ γ, β ↪ ↥((set.range f)ᶜ)) :=
61+
(subtype_prod_equiv_sigma_subtype $
62+
λ (a : α ↪ γ) (b : β ↪ _), disjoint (set.range a) (set.range b)).trans $
63+
equiv.sigma_congr_right $ λ a,
64+
(subtype_equiv_prop begin
65+
ext f,
66+
rw [←set.range_subset_iff, set.subset_compl_iff_disjoint],
67+
exact disjoint.comm.trans disjoint_iff,
68+
end).trans (cod_restrict _ _)
69+
70+
/-- A combination of the above results, allowing us to turn one embedding over a sum type
71+
into two dependent embeddings, the second of which avoids any members of the range
72+
of the first. This is helpful for constructing larger embeddings out of smaller ones. -/
73+
def sum_embedding_equiv_sigma_embedding_restricted {α β γ : Type*} :
74+
((α ⊕ β) ↪ γ) ≃ (Σ f : α ↪ γ, β ↪ ↥((set.range f)ᶜ))
75+
:= equiv.trans sum_embedding_equiv_prod_embedding_disjoint
76+
prod_embedding_disjoint_equiv_sigma_embedding_restricted
77+
78+
/-- Embeddings from a single-member type are equivalent to members of the target type. -/
79+
def unique_embedding_equiv_result {α β : Type*} [unique α] : (α ↪ β) ≃ β :=
80+
{ to_fun := λ f, f (default α),
81+
inv_fun := λ x, ⟨λ _, x, λ _ _ _, subsingleton.elim _ _⟩,
82+
left_inv := λ _, by { ext, simp_rw [function.embedding.coe_fn_mk], congr },
83+
right_inv := λ _, by simp }
84+
85+
end equiv

src/data/fintype/basic.lean

Lines changed: 37 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -938,6 +938,14 @@ lemma equiv_of_fintype_self_embedding_to_embedding {α : Type*} [fintype α] (e
938938
e.equiv_of_fintype_self_embedding.to_embedding = e :=
939939
by { ext, refl, }
940940

941+
/-- If `‖β‖ < ‖α‖` there are no embeddings `α ↪ β`.
942+
This is a formulation of the pigeonhole principle.
943+
944+
Note this cannot be an instance as it needs `h`. -/
945+
@[simp] lemma is_empty_of_card_lt {α β} [fintype α] [fintype β]
946+
(h : fintype.card β < fintype.card α) : is_empty (α ↪ β) :=
947+
⟨λ f, let ⟨x, y, ne, feq⟩ := fintype.exists_ne_map_eq_of_card_lt f h in ne $ f.injective feq⟩
948+
941949
end function.embedding
942950

943951
@[simp]
@@ -1017,6 +1025,10 @@ fintype.of_surjective quotient.mk (λ x, quotient.induction_on x (λ x, ⟨x, rf
10171025
instance finset.fintype [fintype α] : fintype (finset α) :=
10181026
⟨univ.powerset, λ x, finset.mem_powerset.2 (finset.subset_univ _)⟩
10191027

1028+
instance function.embedding.fintype {α β} [fintype α] [fintype β]
1029+
[decidable_eq α] [decidable_eq β] : fintype (α ↪ β) :=
1030+
fintype.of_equiv _ (equiv.subtype_injective_equiv_embedding α β)
1031+
10201032
@[simp] lemma fintype.card_finset [fintype α] :
10211033
fintype.card (finset α) = 2 ^ (fintype.card α) :=
10221034
finset.card_powerset finset.univ
@@ -1402,6 +1414,10 @@ not_fintype α
14021414
@[simp] lemma not_nonempty_fintype {α : Type*} : ¬ nonempty (fintype α) ↔ infinite α :=
14031415
not_nonempty_iff.trans is_empty_fintype
14041416

1417+
/-- A non-infinite type is a fintype. -/
1418+
noncomputable def fintype_of_not_infinite {α : Type*} (h : ¬ infinite α) : fintype α :=
1419+
((not_iff_comm.mp not_nonempty_fintype).mp h).some
1420+
14051421
lemma finset.exists_minimal {α : Type*} [preorder α] (s : finset α) (h : s.nonempty) :
14061422
∃ m ∈ s, ∀ x ∈ s, ¬ (x < m) :=
14071423
begin
@@ -1488,6 +1504,18 @@ begin
14881504
intros x y, contrapose, apply hf,
14891505
end
14901506

1507+
instance function.embedding.is_empty {α β} [infinite α] [fintype β] : is_empty (α ↪ β) :=
1508+
⟨λ f, let ⟨x, y, ne, feq⟩ := fintype.exists_ne_map_eq_of_infinite f in ne $ f.injective feq⟩
1509+
1510+
@[priority 100]
1511+
noncomputable instance function.embedding.fintype'' {α β : Type*} [fintype β] : fintype (α ↪ β) :=
1512+
begin
1513+
by_cases h : infinite α,
1514+
{ resetI, apply_instance },
1515+
{ have := fintype_of_not_infinite h, classical, apply_instance }
1516+
-- the `classical` generates `decidable_eq α/β` instances, and resets instance cache
1517+
end
1518+
14911519
/--
14921520
The strong pigeonhole principle for infinitely many pigeons in
14931521
finitely many pigeonholes. If there are infinitely many pigeons in
@@ -1499,12 +1527,11 @@ See also: `fintype.exists_ne_map_eq_of_infinite`
14991527
lemma fintype.exists_infinite_fiber [infinite α] [fintype β] (f : α → β) :
15001528
∃ y : β, infinite (f ⁻¹' {y}) :=
15011529
begin
1502-
classical, by_contra hf, push_neg at hf,
1503-
haveI h' : ∀ (y : β), fintype (f ⁻¹' {y}) := begin
1504-
intro y, specialize hf y,
1505-
rw [←not_nonempty_fintype, not_not] at hf,
1506-
exact classical.choice hf,
1507-
end,
1530+
classical,
1531+
by_contra hf,
1532+
push_neg at hf,
1533+
1534+
haveI := λ y, fintype_of_not_infinite $ hf y,
15081535
let key : fintype α :=
15091536
{ elems := univ.bUnion (λ (y : β), (f ⁻¹' {y}).to_finset),
15101537
complete := by simp },
@@ -1517,6 +1544,10 @@ assume (hf : surjective f),
15171544
have H : infinite α := infinite.of_surjective f hf,
15181545
by exactI not_fintype α
15191546

1547+
-- the instance generated by `is_empty → subsingleton → fintype is non-computable
1548+
instance function.embedding.fintype' {α β} [infinite α] [fintype β] : fintype (α ↪ β) :=
1549+
{ elems := finset.empty, complete := is_empty_elim }
1550+
15201551
instance nat.infinite : infinite ℕ :=
15211552
⟨λ ⟨s, hs⟩, finset.not_mem_range_self $ s.subset_range_sup_succ (hs _)⟩
15221553

src/data/fintype/card_embedding.lean

Lines changed: 83 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,83 @@
1+
/-
2+
Copyright (c) 2021 Eric Rodriguez. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Eric Rodriguez
5+
-/
6+
import data.fintype.card
7+
import data.equiv.fin
8+
import data.equiv.embedding
9+
10+
/-!
11+
# Birthday Problem
12+
13+
This file establishes the cardinality of `α ↪ β` in full generality.
14+
-/
15+
16+
local notation `|` x `|` := finset.card x
17+
local notation `‖` x `‖` := fintype.card x
18+
19+
open_locale nat
20+
21+
namespace fintype
22+
23+
-- We need the separate `fintype α` instance as it contains data,
24+
-- and may not match definitionally with the instance coming from `unique.fintype`.
25+
lemma card_embedding_of_unique
26+
{α β : Type*} [unique α] [fintype α] [fintype β] [decidable_eq α] [decidable_eq β]:
27+
‖α ↪ β‖ = ‖β‖ := card_congr equiv.unique_embedding_equiv_result
28+
29+
private lemma card_embedding_aux (n : ℕ) (β) [fintype β] [decidable_eq β] (h : n ≤ ‖β‖) :
30+
‖fin n ↪ β‖ = nat.desc_fac (‖β‖ - n) n :=
31+
begin
32+
induction n with n hn,
33+
{ nontriviality (fin 0 ↪ β),
34+
rw [nat.desc_fac_zero, fintype.card_eq_one_iff],
35+
refine ⟨nonempty.some nontrivial.to_nonempty, λ x, function.embedding.ext fin.elim0⟩ },
36+
37+
rw [nat.succ_eq_add_one, ←card_congr (equiv.embedding_congr fin_sum_fin_equiv (equiv.refl β))],
38+
rw card_congr equiv.sum_embedding_equiv_sigma_embedding_restricted,
39+
-- these `rw`s create goals for instances, which it doesn't infer for some reason
40+
all_goals { try { apply_instance } }, -- however, this needs to be done here instead of at the end
41+
-- else, a later `simp`, which depends on the `fintype` instance, won't work.
42+
43+
have : ∀ (f : fin n ↪ β), ‖fin 1 ↪ ↥((set.range f)ᶜ)‖ = ‖β‖ - n,
44+
{ intro f,
45+
rw card_embedding_of_unique,
46+
rw card_of_finset' (finset.map f finset.univ)ᶜ,
47+
{ rw [finset.card_compl, finset.card_map, finset.card_fin] },
48+
{ simp } },
49+
50+
-- putting `card_sigma` in `simp` causes it not to fully simplify
51+
rw card_sigma,
52+
simp only [this, finset.sum_const, finset.card_univ, nsmul_eq_mul, nat.cast_id],
53+
54+
replace h := nat.lt_of_succ_le h,
55+
rw [hn h.le, mul_comm, nat.desc_fac_of_sub h]
56+
end
57+
58+
variables {α β : Type*} [fintype α] [fintype β] [decidable_eq α] [decidable_eq β]
59+
60+
/- Establishes the cardinality of the type of all injections, if any exist. -/
61+
@[simp] theorem card_embedding (h : ‖α‖ ≤ ‖β‖) : ‖α ↪ β‖ = (nat.desc_fac (‖β‖ - ‖α‖) ‖α‖) :=
62+
begin
63+
trunc_cases fintype.trunc_equiv_fin α with eq,
64+
rw fintype.card_congr (equiv.embedding_congr eq (equiv.refl β)),
65+
exact card_embedding_aux _ _ h,
66+
end
67+
68+
/-- If `‖β‖ < ‖α‖` there are no embeddings `α ↪ β`.
69+
This is a formulation of the pigeonhole principle. -/
70+
@[simp] theorem card_embedding_eq_zero (h : ‖β‖ < ‖α‖) : ‖α ↪ β‖ = 0 :=
71+
card_eq_zero_iff.mpr $ function.embedding.is_empty_of_card_lt h
72+
73+
theorem card_embedding_eq_if : ‖α ↪ β‖ = if ‖α‖ ≤ ‖β‖ then nat.desc_fac (‖β‖ - ‖α‖) ‖α‖ else 0 :=
74+
begin
75+
split_ifs with h,
76+
{ exact card_embedding h },
77+
{ exact card_embedding_eq_zero (not_le.mp h) }
78+
end
79+
80+
lemma card_embedding_eq_infinite {α β} [infinite α] [fintype β] : ‖α ↪ β‖ = 0 :=
81+
by rw card_eq_zero_iff; apply_instance
82+
83+
end fintype

src/data/nat/factorial.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -184,6 +184,14 @@ begin
184184
exact (nat.mul_div_cancel' $ factorial_dvd_factorial $ le.intro rfl).symm
185185
end
186186

187+
lemma desc_fac_of_sub {n k : ℕ} (h : n < k) :
188+
(k - n) * (k - n).desc_fac n = (k - (n + 1)).desc_fac (n + 1) :=
189+
begin
190+
set t := k - n.succ with ht,
191+
suffices h' : k - n = t.succ, by rw [←ht, h', succ_desc_fac, desc_fac_succ],
192+
rw [ht, succ_eq_add_one, ←sub_sub_assoc h (succ_pos _), succ_sub_one],
193+
end
194+
187195
end desc_fac
188196

189197
end nat

src/logic/embedding.lean

Lines changed: 43 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -62,6 +62,9 @@ lemma ext_iff {α β} {f g : embedding α β} : (∀ x, f x = g x) ↔ f = g :=
6262
@[simp] theorem coe_fn_mk {α β} (f : α → β) (i) :
6363
(@mk _ _ f i : α → β) = f := rfl
6464

65+
@[simp] lemma mk_coe {α β : Type*} (f : α ↪ β) (inj) : (⟨f, inj⟩ : α ↪ β) = f :=
66+
by { ext, simp }
67+
6568
theorem injective {α β} (f : α ↪ β) : injective f := f.inj'
6669

6770
@[refl, simps {simp_rhs := tt}]
@@ -232,6 +235,46 @@ end function
232235

233236
namespace equiv
234237

238+
open function.embedding
239+
240+
/-- The type of embeddings `α ↪ β` is equivalent to
241+
the subtype of all injective functions `α → β`. -/
242+
def subtype_injective_equiv_embedding (α β : Sort*) :
243+
{f : α → β // function.injective f} ≃ (α ↪ β) :=
244+
{ to_fun := λ f, ⟨f.val, f.property⟩,
245+
inv_fun := λ f, ⟨f, f.injective⟩,
246+
left_inv := λ f, by simp,
247+
right_inv := λ f, by {ext, refl} }
248+
249+
/-- If `α₁ ≃ α₂` and `β₁ ≃ β₂`, then the type of embeddings `α₁ ↪ β₁`
250+
is equivalent to the type of embeddings `α₂ ↪ β₂`. -/
251+
@[congr, simps apply] def embedding_congr {α β γ δ : Sort*}
252+
(h : α ≃ β) (h' : γ ≃ δ) : (α ↪ γ) ≃ (β ↪ δ) :=
253+
{ to_fun := λ f, h.symm.to_embedding.trans $ f.trans $ h'.to_embedding,
254+
inv_fun := λ f, h.to_embedding.trans $ f.trans $ h'.symm.to_embedding,
255+
left_inv := λ x, by {ext, simp},
256+
right_inv := λ x, by {ext, simp} }
257+
258+
@[simp] lemma embedding_congr_refl {α β : Sort*} :
259+
embedding_congr (equiv.refl α) (equiv.refl β) = equiv.refl (α ↪ β) :=
260+
by {ext, refl}
261+
262+
@[simp] lemma embedding_congr_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Sort*}
263+
(e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) :
264+
embedding_congr (e₁.trans e₂) (e₁'.trans e₂') =
265+
(embedding_congr e₁ e₁').trans (embedding_congr e₂ e₂') :=
266+
rfl
267+
268+
@[simp] lemma embedding_congr_symm {α₁ β₁ α₂ β₂ : Sort*} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) :
269+
(embedding_congr e₁ e₂).symm = embedding_congr e₁.symm e₂.symm :=
270+
rfl
271+
272+
lemma embedding_congr_apply_trans {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort*}
273+
(ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) (f : α₁ ↪ β₁) (g : β₁ ↪ γ₁) :
274+
equiv.embedding_congr ea ec (f.trans g) =
275+
(equiv.embedding_congr ea eb f).trans (equiv.embedding_congr eb ec g) :=
276+
by {ext, simp}
277+
235278
@[simp]
236279
lemma refl_to_embedding {α : Type*} : (equiv.refl α).to_embedding = function.embedding.refl α := rfl
237280

src/ring_theory/polynomial/pochhammer.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,8 @@ import data.polynomial.eval
1111
1212
We define and prove some basic relations about
1313
`pochhammer S n : polynomial S = X * (X+1) * ... * (X + n - 1)`
14-
which is also known as the rising factorial.
14+
which is also known as the rising factorial. A version of this definition
15+
that is focused on `nat` can be found in `data.nat.factorial` as `desc_fac`.
1516
1617
## Implementation
1718

0 commit comments

Comments
 (0)