-
Notifications
You must be signed in to change notification settings - Fork 4
/
CantorPairing.v
96 lines (74 loc) · 2.26 KB
/
CantorPairing.v
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
Require Import Arith Lia.
Definition inj {X Y} (f : X -> Y) :=
forall x x', f x = f x' -> x = x'.
Definition inv {X Y} (g : Y -> X) (f : X -> Y) :=
forall x, g (f x) = x.
Definition bijection X Y :=
{ f : X -> Y & { g & inv g f /\ inv f g }}.
Lemma size_rect {X} σ f :
(forall x, (forall y : X, σ y < σ x -> f y) -> f x) -> forall x, f x.
Proof.
intros H x. apply H.
induction (σ x).
- now intros ??%PeanoNat.Nat.nlt_0_r.
- intros. apply H.
intros. apply IHn. lia.
Defined.
(** We define the Cantor pairing function and show that it is a Bijection. *)
Section Cantor.
Definition next '(x,y) :=
match x with
| O => (S y, O)
| S x' => (x', S y)
end.
Fact n00_next : forall p, (0,0) <> next p.
Proof. intros [[] ]; intuition discriminate. Qed.
Fact inj_next : inj next.
Proof. intros [[] ][[] ]; cbn; congruence. Qed.
Fixpoint decode n :=
match n with
| 0 => (0,0)
| S x => next (decode x)
end.
Lemma inj_decode : inj decode.
Proof.
intros n. induction n; intros []; auto.
- now intros ?%n00_next.
- now intros ?%eq_sym%n00_next.
- intros [= ?%inj_next%IHn]. congruence.
Qed.
(* Shows that the function "next" is almost surjective *)
Fact zero_or_next : forall p, {a | p = next a} + {p = (0,0)}.
Proof.
intros [x []].
- destruct x. now right. left; now exists (0, x).
- left; now exists (S x, n).
Defined.
Fixpoint Σ n := match n with 0 => 0 | S x => n + Σ x end.
Definition code '(x,y) := Σ(x+y)+y.
Lemma code_next p : code(next p) = S(code p).
Proof.
destruct p as [[|x] y]; cbn.
- rewrite <-!plus_n_O, Nat.add_comm. reflexivity.
- rewrite !Nat.add_succ_r. reflexivity.
Qed.
Lemma inv_dc : inv decode code.
Proof.
unfold inv.
apply (size_rect code). intros p rec.
destruct (zero_or_next p) as [[? ->] | ->].
- rewrite code_next. cbn. f_equal. apply rec.
rewrite code_next; auto.
- reflexivity.
Qed.
Fact inv_cd : inv code decode.
Proof.
intros ?. apply inj_decode. now rewrite inv_dc.
Qed.
Corollary Bij_Nat_NatNat : bijection nat (nat * nat).
Proof.
exists decode, code. split. apply inv_cd. apply inv_dc.
Qed.
Fact CantorBound x y n : code (x, y) = n -> y < S n.
Proof. cbn; lia. Qed.
End Cantor.