Skip to content
Snippets Groups Projects
Commit 2b0d874d authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Show that Pigeon hole principle on nat is a special case of the finite version.

parent 296f00b9
No related branches found
No related tags found
No related merge requests found
......@@ -427,3 +427,15 @@ Proof.
cut (Inj eq eq f); [intros ?%inj_card; lia|].
intros x1 x2 ?. apply dec_stable. naive_solver.
Qed.
Lemma nat_pigeon_hole (f : nat nat) (n1 n2 : nat) :
n2 < n1
( i, i < n1 f i < n2)
i1 i2, i1 < i2 f i1 = f i2.
Proof.
intros Hn Hf. pose (f' (i : fin n1) := nat_to_fin (Hf _ (fin_to_nat_lt i))).
destruct (finite_pigeon_hole f') as (i1&i2&Hi&Hf'); [by rewrite !fin_card|].
apply (not_inj (f:=fin_to_nat)) in Hi. apply (f_equal fin_to_nat) in Hf'.
unfold f' in Hf'. rewrite !fin_to_nat_to_fin in Hf'.
destruct (decide (i1 < i2)); [exists i1, i2|exists i2, i1]; lia.
Qed.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment