Skip to content

Commit 289bb19

Browse files
A proof of the pigeonhole principle (#387)
1 parent f342904 commit 289bb19

File tree

2 files changed

+15
-1
lines changed

2 files changed

+15
-1
lines changed

CHANGELOG.md

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -308,6 +308,8 @@ Other minor additions
308308
* Added new proof to `Data.Fin.Properties`:
309309
```agda
310310
toℕ-fromℕ≤″ : toℕ (fromℕ≤″ m m<n) ≡ m
311+
312+
pigeonhole : m < n → (f : Fin n → Fin m) → ∃₂ λ i j → i ≢ j × f i ≡ f j
311313
```
312314

313315
* Added new proofs to `Data.List.Any.Properties`:

src/Data/Fin/Properties.agda

Lines changed: 13 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -20,7 +20,7 @@ open import Data.Nat as ℕ using (ℕ; zero; suc; s≤s; z≤n; _∸_)
2020
)
2121
import Data.Nat.Properties as ℕₚ
2222
open import Data.Unit using (tt)
23-
open import Data.Product using (∃; ∄; _×_; _,_; map; proj₁)
23+
open import Data.Product using (∃; ∃₂; ∄; _×_; _,_; map; proj₁)
2424
open import Function using (_∘_; id)
2525
open import Function.Injection using (_↣_)
2626
open import Relation.Binary as B hiding (Decidable)
@@ -496,6 +496,18 @@ all? P? with decFinSubset U? (λ {f} _ → P? f)
496496
¬ ( i P i) (∃ λ i ¬ P i)
497497
¬∀⟶∃¬ n P P? ¬P = map id proj₁ (¬∀⟶∃¬-smallest n P P? ¬P)
498498

499+
-- The pigeonhole principle.
500+
501+
pigeonhole : {m n} m ℕ.< n (f : Fin n Fin m)
502+
∃₂ λ i j i ≢ j × f i ≡ f j
503+
pigeonhole (s≤s z≤n) f = contradiction (f zero) λ()
504+
pigeonhole (s≤s (s≤s m≤n)) f with any? (λ k f zero ≟ f (suc k))
505+
... | yes (j , f₀≡fⱼ) = zero , suc j , (λ()) , f₀≡fⱼ
506+
... | no f₀≢fₖ with pigeonhole (s≤s m≤n) (λ j punchOut (f₀≢fₖ ∘ (j ,_ )))
507+
... | (i , j , i≢j , fᵢ≡fⱼ) =
508+
suc i , suc j , i≢j ∘ suc-injective ,
509+
punchOut-injective (f₀≢fₖ ∘ (i ,_)) _ fᵢ≡fⱼ
510+
499511
------------------------------------------------------------------------
500512
-- Categorical
501513

0 commit comments

Comments
 (0)