mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-11 14:48:42 +00:00
Cantor pairing function is bijective (#42)
This commit is contained in:
20
Sets/CantorBijection/CantorBijection.agda
Normal file
20
Sets/CantorBijection/CantorBijection.agda
Normal file
@@ -0,0 +1,20 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
open import WellFoundedInduction
|
||||
open import Sets.CantorBijection.Proofs
|
||||
|
||||
module Sets.CantorBijection.CantorBijection where
|
||||
|
||||
open Sets.CantorBijection.Proofs using (cantorInverse ; cantorInverseLemma) public
|
||||
|
||||
cantorBijection : Bijection cantorInverse
|
||||
Injection.property (Bijection.inj cantorBijection) {x} {y} = cantorInverseInjective x y
|
||||
Surjection.property (Bijection.surj cantorBijection) = cantorInverseSurjective
|
98
Sets/CantorBijection/Order.agda
Normal file
98
Sets/CantorBijection/Order.agda
Normal file
@@ -0,0 +1,98 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
open import WellFoundedInduction
|
||||
|
||||
module Sets.CantorBijection.Order where
|
||||
|
||||
order : Rel (ℕ && ℕ)
|
||||
order (a ,, b) (c ,, d) = ((a +N b) <N (c +N d)) || (((a +N b) ≡ (c +N d)) && (b <N d))
|
||||
|
||||
totalOrder : TotalOrder (ℕ && ℕ)
|
||||
PartialOrder._<_ (TotalOrder.order totalOrder) = order
|
||||
PartialOrder.irreflexive (TotalOrder.order totalOrder) {fst ,, snd} (inl x) = exFalso (TotalOrder.irreflexive ℕTotalOrder x)
|
||||
PartialOrder.irreflexive (TotalOrder.order totalOrder) {fst ,, snd} (inr (_ ,, p)) = exFalso (TotalOrder.irreflexive ℕTotalOrder p)
|
||||
PartialOrder.transitive (TotalOrder.order totalOrder) {x1 ,, x2} {y1 ,, y2} {z1 ,, z2} (inl pr1) (inl pr2) = inl (TotalOrder.transitive ℕTotalOrder pr1 pr2)
|
||||
PartialOrder.transitive (TotalOrder.order totalOrder) {x1 ,, x2} {y1 ,, y2} {z1 ,, z2} (inl pr1) (inr (f1 ,, f2)) = inl (identityOfIndiscernablesRight _ _ _ _<N_ pr1 f1)
|
||||
PartialOrder.transitive (TotalOrder.order totalOrder) {x1 ,, x2} {y1 ,, y2} {z1 ,, z2} (inr (f1 ,, f2)) (inl x) = inl (identityOfIndiscernablesLeft _ _ _ _<N_ x (equalityCommutative f1))
|
||||
PartialOrder.transitive (TotalOrder.order totalOrder) {x1 ,, x2} {y1 ,, y2} {z1 ,, z2} (inr (f1 ,, f2)) (inr (fst ,, snd)) = inr (transitivity f1 fst ,, TotalOrder.transitive ℕTotalOrder f2 snd)
|
||||
TotalOrder.totality totalOrder (a ,, b) (c ,, d) with TotalOrder.totality ℕTotalOrder (a +N b) (c +N d)
|
||||
TotalOrder.totality totalOrder (a ,, b) (c ,, d) | inl (inl x) = inl (inl (inl x))
|
||||
TotalOrder.totality totalOrder (a ,, b) (c ,, d) | inl (inr x) = inl (inr (inl x))
|
||||
TotalOrder.totality totalOrder (a ,, b) (c ,, d) | inr eq with TotalOrder.totality ℕTotalOrder b d
|
||||
TotalOrder.totality totalOrder (a ,, b) (c ,, d) | inr eq | inl (inl x) = inl (inl (inr (eq ,, x)))
|
||||
TotalOrder.totality totalOrder (a ,, b) (c ,, d) | inr eq | inl (inr x) = inl (inr (inr (equalityCommutative eq ,, x)))
|
||||
TotalOrder.totality totalOrder (a ,, b) (c ,, .b) | inr eq | inr refl rewrite canSubtractFromEqualityRight {a} {b} {c} eq = inr refl
|
||||
|
||||
leastElement : {y : ℕ && ℕ} → order y (0 ,, 0) → False
|
||||
leastElement {zero ,, b} (inl ())
|
||||
leastElement {zero ,, b} (inr ())
|
||||
leastElement {succ a ,, b} (inl ())
|
||||
leastElement {succ a ,, b} (inr ())
|
||||
|
||||
mustDescend : (a b t : ℕ) → order (a ,, b) (t ,, zero) → a +N b <N t
|
||||
mustDescend a b zero ord = exFalso (leastElement ord)
|
||||
mustDescend a b (succ t) (inl x) rewrite Semiring.sumZeroRight ℕSemiring t = x
|
||||
|
||||
orderWellfounded : WellFounded order
|
||||
orderWellfounded (a ,, b) = access (go b a)
|
||||
where
|
||||
g0 : (c : ℕ) (y : ℕ && ℕ) → order y (c ,, 0) → Accessible order y
|
||||
-- We want to induct on the second entry of x here, so we decompose so as to put that first.
|
||||
g : (x : ℕ) → ((y : ℕ) → y <N x → (b : ℕ) (z : ℕ && ℕ) → order z (b ,, y) → Accessible order z) → (c : ℕ) (y : ℕ && ℕ) → order y (c ,, x) → Accessible order y
|
||||
|
||||
g0 = rec <NWellfounded (λ z → (x : ℕ && ℕ) (x₁ : order x (z ,, zero)) → Accessible order x) f
|
||||
where
|
||||
p : (a b : ℕ) → a ≡ b → (a ,, zero) ≡ (b ,, zero)
|
||||
p a .a refl = refl
|
||||
|
||||
f : (x : ℕ) (x₁ : (x₂ : ℕ) (x₃ : x₂ <N x) (x₄ : ℕ && ℕ) (x₅ : order x₄ (x₂ ,, zero)) → Accessible order x₄) (x₂ : ℕ && ℕ) (x₃ : order x₂ (x ,, zero)) → Accessible order x₂
|
||||
f zero pr y ord = exFalso (leastElement ord)
|
||||
f (succ m) pr (fst ,, snd) (inl pr2) = h snd fst pr2
|
||||
where
|
||||
go : (x : ℕ) (x₁ : (x₂ : ℕ) (x₃ : x₂ <N x) (x₄ : ℕ) (x₅ : x₄ +N x₂ <N succ (m +N zero)) → Accessible order (x₄ ,, x₂)) (x₂ : ℕ) (x₃ : x₂ +N x <N succ (m +N zero)) → Accessible order (x₂ ,, x)
|
||||
go bound pr2 toProve bounded with TotalOrder.totality ℕTotalOrder (toProve +N bound) (m +N 0)
|
||||
... | inl (inl bl) = pr m (le 0 refl) (toProve ,, bound) (inl bl)
|
||||
... | inl (inr bl) = exFalso (noIntegersBetweenXAndSuccX (m +N 0) bl bounded)
|
||||
go zero pr2 toProve _ | inr bl rewrite Semiring.sumZeroRight ℕSemiring m | Semiring.sumZeroRight ℕSemiring toProve = access λ i i<z → pr m (le 0 refl) i (inl (mustDescend (_&&_.fst i) (_&&_.snd i) (m +N zero) (identityOfIndiscernablesLeft _ _ _ order (identityOfIndiscernablesRight _ _ _ order i<z (p toProve (m +N 0) (transitivity bl (equalityCommutative (Semiring.sumZeroRight ℕSemiring m))))) refl)))
|
||||
go (succ bound) pr2 toProve _ | inr bl = access desc
|
||||
where
|
||||
desc : (i : ℕ && ℕ) → (order i (toProve ,, succ bound)) → Accessible order i
|
||||
desc (i1 ,, i2) (inl ord) = pr m (le zero refl) (i1 ,, i2) (inl (identityOfIndiscernablesRight _ _ _ _<N_ ord bl))
|
||||
desc (i1 ,, i2) (inr (ord1 ,, ord2)) = pr2 i2 ord2 i1 (le 0 (applyEquality succ (transitivity ord1 bl)))
|
||||
|
||||
h : (a : ℕ) (b : ℕ) (pr2 : b +N a <N succ m +N 0) → Accessible order (b ,, a)
|
||||
h = rec <NWellfounded (λ z → (x2 : ℕ) (x1 : x2 +N z <N succ (m +N zero)) → Accessible order (x2 ,, z)) go
|
||||
|
||||
g zero _ = g0
|
||||
g (succ x) pr zero (y1 ,, y2) (inl pr1) with TotalOrder.totality ℕTotalOrder (y1 +N y2) x
|
||||
g (succ x) pr zero (y1 ,, y2) (inl pr1) | inl (inl y1+y2<x) = pr x (le 0 refl) zero (y1 ,, y2) (inl y1+y2<x)
|
||||
g (succ x) pr zero (y1 ,, y2) (inl pr1) | inl (inr x<y1+y2) = exFalso (noIntegersBetweenXAndSuccX x x<y1+y2 pr1)
|
||||
g (succ x) pr zero (y1 ,, y2) (inl pr1) | inr y1+y2=x = access (λ y y<zs → pr x (le 0 refl) zero y (ans y y<zs))
|
||||
where
|
||||
ans : (y : ℕ && ℕ) → order y (y1 ,, y2) → (_&&_.fst y +N _&&_.snd y <N x) || ((_&&_.fst y +N _&&_.snd y ≡ x) && (_&&_.snd y <N x))
|
||||
ans (fst ,, snd) (inl x) rewrite y1+y2=x = inl x
|
||||
ans (zero ,, snd) (inr (a1 ,, a2)) rewrite y1+y2=x | a1 = exFalso (bad x y1 y2 y1+y2=x a2)
|
||||
where
|
||||
bad : (x y1 y2 : ℕ) → y1 +N y2 ≡ x → x <N y2 → False
|
||||
bad x zero y2 y1+y2=x x<y2 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _ _ _ _<N_ x<y2 y1+y2=x)
|
||||
bad x (succ y1) y2 y1+y2=x x<y2 rewrite equalityCommutative y1+y2=x = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.transitive ℕTotalOrder x<y2 (identityOfIndiscernablesRight _ _ _ _<N_ (addingIncreases y2 y1) (Semiring.commutative ℕSemiring y2 (succ y1))))
|
||||
ans (succ fst ,, snd) (inr (a1 ,, a2)) rewrite y1+y2=x = inr (a1 ,, le fst a1)
|
||||
g (succ x) pr zero (zero ,, y2) (inr (fst ,, snd)) = exFalso (TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _ _ _ _<N_ snd fst))
|
||||
g (succ x) pr zero (succ y1 ,, y2) (inr (fst ,, snd)) = pr y2 snd (succ (succ y1)) (succ y1 ,, y2) (inl (le 0 refl))
|
||||
g (succ x) pr (succ c) y (inl z) = pr x (le 0 refl) (succ (succ c)) y (inl (identityOfIndiscernablesRight _ _ _ _<N_ z (applyEquality succ (transitivity (Semiring.commutative ℕSemiring c (succ x)) (applyEquality succ (Semiring.commutative ℕSemiring x c))))))
|
||||
g (succ x) pr (succ c) y (inr (fst ,, snd)) with TotalOrder.totality ℕTotalOrder (_&&_.snd y) x
|
||||
... | inl (inl bl) = pr x (le 0 refl) (succ (succ c)) y (inr (transitivity fst (applyEquality succ (transitivity (Semiring.commutative ℕSemiring c (succ x)) (applyEquality succ (Semiring.commutative ℕSemiring x c)))) ,, bl))
|
||||
... | inl (inr bl) = exFalso (noIntegersBetweenXAndSuccX x bl snd)
|
||||
g (succ x) pr (succ c) (y1 ,, .x) (inr (fst ,, _)) | inr refl with canSubtractFromEqualityRight {y1} {x} {succ (succ c)} (transitivity fst (applyEquality succ (transitivity (Semiring.commutative ℕSemiring c (succ x)) (applyEquality succ (Semiring.commutative ℕSemiring x c)))))
|
||||
g (succ x) pr (succ c) (.(succ (succ c)) ,, .x) (inr (_ ,, _)) | inr refl | refl = pr x (le 0 refl) (succ (succ (succ c))) (succ (succ c) ,, x) (inl (le 0 refl))
|
||||
|
||||
go : (a : ℕ) → (b : ℕ) → (y : ℕ && ℕ) → order y (b ,, a) → Accessible order y
|
||||
go = rec <NWellfounded (λ (a : ℕ) → (b : ℕ) → (y : ℕ && ℕ) → order y (b ,, a) → Accessible order y) g
|
158
Sets/CantorBijection/Proofs.agda
Normal file
158
Sets/CantorBijection/Proofs.agda
Normal file
@@ -0,0 +1,158 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
open import WellFoundedInduction
|
||||
open import Sets.CantorBijection.Order
|
||||
|
||||
module Sets.CantorBijection.Proofs where
|
||||
|
||||
open Sets.CantorBijection.Order using (order ; totalOrder ; orderWellfounded)
|
||||
|
||||
cantorInverseLemma : (ℕ && ℕ) → (ℕ && ℕ)
|
||||
cantorInverseLemma (0 ,, s) = (succ s) ,, 0
|
||||
cantorInverseLemma (succ r ,, 0) = r ,, 1
|
||||
cantorInverseLemma (succ r ,, succ s) = (r ,, succ (succ s))
|
||||
|
||||
cantorInverse : ℕ → (ℕ && ℕ)
|
||||
cantorInverse zero = (0 ,, 0)
|
||||
cantorInverse (succ z) = cantorInverseLemma (cantorInverse z)
|
||||
|
||||
cantorInverseLemmaInjective : (a b : ℕ && ℕ) → cantorInverseLemma a ≡ cantorInverseLemma b → a ≡ b
|
||||
cantorInverseLemmaInjective (zero ,, b) (zero ,, .b) refl = refl
|
||||
cantorInverseLemmaInjective (zero ,, b) (succ c ,, zero) ()
|
||||
cantorInverseLemmaInjective (zero ,, b) (succ c ,, succ d) ()
|
||||
cantorInverseLemmaInjective (succ a ,, zero) (zero ,, d) ()
|
||||
cantorInverseLemmaInjective (succ a ,, succ b) (zero ,, d) ()
|
||||
cantorInverseLemmaInjective (succ a ,, zero) (succ .a ,, zero) refl = refl
|
||||
cantorInverseLemmaInjective (succ a ,, succ b) (succ .a ,, succ .b) refl = refl
|
||||
|
||||
cantorInverseLemmaSurjective : (a : ℕ && ℕ) → (a ≡ (0 ,, 0)) || (Sg (ℕ && ℕ) (λ b → cantorInverseLemma b ≡ a))
|
||||
cantorInverseLemmaSurjective (zero ,, zero) = inl refl
|
||||
cantorInverseLemmaSurjective (zero ,, succ zero) = inr ((1 ,, 0) , refl)
|
||||
cantorInverseLemmaSurjective (zero ,, succ (succ b)) = inr ((1 ,, succ b) , refl)
|
||||
cantorInverseLemmaSurjective (succ a ,, zero) = inr ((0 ,, a) , refl)
|
||||
cantorInverseLemmaSurjective (succ a ,, succ zero) = inr ((succ (succ a) ,, 0) , refl)
|
||||
cantorInverseLemmaSurjective (succ a ,, succ (succ b)) = inr ((succ (succ a) ,, succ b) , refl)
|
||||
|
||||
cantorInverseLemmaNotZero : (a : ℕ && ℕ) → (cantorInverseLemma a ≡ (0 ,, 0)) → False
|
||||
cantorInverseLemmaNotZero (succ a ,, zero) ()
|
||||
cantorInverseLemmaNotZero (succ a ,, succ b) ()
|
||||
|
||||
cantorInverseLemmaIncreases : (x : ℕ && ℕ) → order x (cantorInverseLemma x)
|
||||
cantorInverseLemmaIncreases (zero ,, b) = inl (identityOfIndiscernablesRight _ _ _ _<N_ (a<SuccA b) (applyEquality succ (equalityCommutative (Semiring.sumZeroRight ℕSemiring b))))
|
||||
cantorInverseLemmaIncreases (succ a ,, zero) = inr (transitivity (applyEquality succ (Semiring.sumZeroRight ℕSemiring a)) (Semiring.commutative ℕSemiring 1 a) ,, (le 0 refl))
|
||||
cantorInverseLemmaIncreases (succ a ,, succ b) = inr (transitivity (applyEquality succ (Semiring.commutative ℕSemiring a (succ b))) (Semiring.commutative ℕSemiring (succ (succ b)) a) ,, (le 0 refl))
|
||||
|
||||
cantorInverseOrderPreserving : (x y : ℕ) → (x <N y) → order (cantorInverse x) (cantorInverse y)
|
||||
cantorInverseOrderPreserving zero (succ y) x<y with TotalOrder.totality totalOrder (0 ,, 0) (cantorInverseLemma (cantorInverse y))
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inl (inl bl) = bl
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inl (inr bl) = exFalso (leastElement bl)
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inr x = exFalso (leastElement {cantorInverse y} (identityOfIndiscernablesRight _ _ _ order (cantorInverseLemmaIncreases (cantorInverse y)) (equalityCommutative x)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y with cantorInverseOrderPreserving x y (subtractOneFromInequality x<y)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr with cantorInverse x
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | x1 ,, x2 with cantorInverse y
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | zero ,, succ y2 = inl (succPreservesInequality (succIsPositive (y2 +N 0)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero with TotalOrder.totality ℕTotalOrder 1 (y1 +N 1)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inl (inl pr1) = inl pr1
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inl (inr pr1) rewrite Semiring.commutative ℕSemiring y1 1 = exFalso (zeroNeverGreater (subtractOneFromInequality pr1))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inr pr1 = inr (pr1 ,, le 0 refl)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring y1 (succ (succ y2)) = inl (succPreservesInequality (succIsPositive (y2 +N y1)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, zero | zero ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.sumZeroRight ℕSemiring y2 | Semiring.sumZeroRight ℕSemiring x1 = inl (TotalOrder.transitive ℕTotalOrder pr (a<SuccA _))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, zero | succ y1 ,, zero rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring x1 | Semiring.sumZeroRight ℕSemiring y1 = inl pr
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, zero | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.sumZeroRight ℕSemiring x1 = inl (identityOfIndiscernablesRight _ _ _ _<N_ pr (transitivity (applyEquality succ (Semiring.commutative ℕSemiring y1 (succ y2))) (Semiring.commutative ℕSemiring (succ (succ y2)) y1)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, succ x2 | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring x2 | Semiring.sumZeroRight ℕSemiring y2 = inl (succPreservesInequality pr)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, succ x2 | succ y1 ,, zero rewrite Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring y1 | Semiring.sumZeroRight ℕSemiring x2 = ans
|
||||
where
|
||||
ans : (succ (succ x2) <N succ y1) || (succ (succ x2) ≡ succ y1) && (zero <N 1)
|
||||
ans with TotalOrder.totality ℕTotalOrder (succ (succ x2)) (succ y1)
|
||||
ans | inl (inl x) = inl x
|
||||
ans | inl (inr x) = exFalso (noIntegersBetweenXAndSuccX (succ x2) pr x)
|
||||
ans | inr x = inr (x ,, le zero refl)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, succ x2 | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring y1 | Semiring.sumZeroRight ℕSemiring x2 = ans
|
||||
where
|
||||
ans : (succ (succ x2) <N y1 +N succ (succ y2)) || (succ (succ x2) ≡ y1 +N succ (succ y2)) && (zero <N succ (succ y2))
|
||||
ans with TotalOrder.totality ℕTotalOrder (succ (succ x2)) (y1 +N succ (succ y2))
|
||||
ans | inl (inl x) = inl x
|
||||
ans | inl (inr x) = exFalso (noIntegersBetweenXAndSuccX (succ x2) pr (identityOfIndiscernablesLeft _ _ _ _<N_ x (transitivity (Semiring.commutative ℕSemiring y1 (succ (succ y2))) (applyEquality succ (Semiring.commutative ℕSemiring (succ y2) y1)))))
|
||||
ans | inr x = inr (x ,, succIsPositive (succ y2))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, succ x2 | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring y2 = inl (TotalOrder.transitive ℕTotalOrder (identityOfIndiscernablesLeft _ _ _ _<N_ pr (transitivity (applyEquality succ (Semiring.commutative ℕSemiring x1 (succ x2))) (Semiring.commutative ℕSemiring (succ (succ x2)) x1))) (a<SuccA (succ y2)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, succ x2 | succ y1 ,, zero rewrite Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring y1 | Semiring.commutative ℕSemiring x1 (succ x2) | Semiring.commutative ℕSemiring x1 (succ (succ x2)) = inl pr
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, succ x2 | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 (succ x2) | Semiring.commutative ℕSemiring y1 (succ y2) | Semiring.commutative ℕSemiring x1 (succ (succ x2)) | Semiring.commutative ℕSemiring y1 (succ (succ y2)) = inl pr
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) with cantorInverse x
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | x1 ,, x2 with cantorInverse y
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, ()) | zero ,, zero | zero ,, zero
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (() ,, snd) | zero ,, zero | zero ,, succ y2
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (refl ,, snd) | zero ,, succ .y2 | zero ,, succ y2 = exFalso (TotalOrder.irreflexive ℕTotalOrder snd)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (refl ,, snd) | zero ,, succ .(y1 +N succ y2) | succ y1 ,, succ y2 = exFalso (TotalOrder.irreflexive ℕTotalOrder (TotalOrder.transitive ℕTotalOrder snd (identityOfIndiscernablesRight _ _ _ _<N_ (addingIncreases (succ y2) y1) (Semiring.commutative ℕSemiring (succ y2) (succ y1)))))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, zero | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring x1 | succInjective fst | Semiring.commutative ℕSemiring y2 1 | Semiring.sumZeroRight ℕSemiring y2 = inl (le zero refl)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, zero | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.sumZeroRight ℕSemiring x1 = inr (transitivity fst (transitivity (applyEquality succ (Semiring.commutative ℕSemiring y1 (succ y2))) (Semiring.commutative ℕSemiring (succ (succ y2)) y1)) ,, succPreservesInequality snd)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, succ x2 | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring y2 | Semiring.commutative ℕSemiring x1 (succ x2) | Semiring.commutative ℕSemiring (succ (succ x2)) x1 | fst = inl (succPreservesInequality (le zero refl))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, succ x2 | succ y1 ,, succ y2 = inr (transitivity (transitivity (Semiring.commutative ℕSemiring x1 (succ (succ x2))) (applyEquality succ (Semiring.commutative ℕSemiring (succ x2) x1))) (transitivity fst (transitivity (applyEquality succ (Semiring.commutative ℕSemiring y1 (succ y2))) (Semiring.commutative ℕSemiring (succ (succ y2)) y1))) ,, succPreservesInequality snd)
|
||||
|
||||
cantorInverseInjective : (a b : ℕ) → cantorInverse a ≡ cantorInverse b → a ≡ b
|
||||
cantorInverseInjective zero zero pr = refl
|
||||
cantorInverseInjective zero (succ b) pr = exFalso (cantorInverseLemmaNotZero _ (equalityCommutative pr))
|
||||
cantorInverseInjective (succ a) zero pr = exFalso (cantorInverseLemmaNotZero _ pr)
|
||||
cantorInverseInjective (succ a) (succ b) pr = applyEquality succ (cantorInverseInjective a b (cantorInverseLemmaInjective (cantorInverse a) (cantorInverse b) pr))
|
||||
|
||||
-- Some unnecessary things on the way to the final proof
|
||||
cantorInverseDiscrete : (a : ℕ) → (c : ℕ && ℕ) → order (cantorInverse a) c → order c (cantorInverse (succ a)) → False
|
||||
cantorInverseDiscrete zero (zero ,, zero) (inl ()) c<sa
|
||||
cantorInverseDiscrete zero (zero ,, zero) (inr ()) c<sa
|
||||
cantorInverseDiscrete zero (zero ,, succ c) a<c (inl x) = zeroNeverGreater (subtractOneFromInequality x)
|
||||
cantorInverseDiscrete zero (succ b ,, zero) a<c (inl x) = zeroNeverGreater (subtractOneFromInequality x)
|
||||
cantorInverseDiscrete zero (succ b ,, succ c) a<c (inl x) = zeroNeverGreater (subtractOneFromInequality x)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) a<c c<sa with cantorInverse a
|
||||
cantorInverseDiscrete (succ a) (zero ,, zero) (inl ()) (inl x) | zero ,, zero
|
||||
cantorInverseDiscrete (succ a) (zero ,, zero) (inr ()) (inl x) | zero ,, zero
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ zero) a<c (inl x) | zero ,, zero = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ (succ c)) a<c (inl x) | zero ,, zero = zeroNeverGreater (subtractOneFromInequality x)
|
||||
cantorInverseDiscrete (succ a) (succ b ,, c) a<c (inl x) | zero ,, zero = zeroNeverGreater (subtractOneFromInequality x)
|
||||
cantorInverseDiscrete (succ a) (b ,, zero) (inl x) (inr (fst ,, snd)) | zero ,, zero = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _ _ _ _<N_ x fst)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) a<c (inr (fst ,, snd)) | zero ,, zero = zeroNeverGreater (subtractOneFromInequality snd)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl y) (inl x) | zero ,, succ snd rewrite Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.transitive ℕTotalOrder x y)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) (inr (fst ,, _)) (inl x) | zero ,, succ snd rewrite Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd | fst = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, zero) (inl x) (inr (fst ,, snd₁)) | zero ,, succ snd rewrite fst | Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) (inl x) (inr (fst ,, snd1)) | zero ,, succ snd = zeroNeverGreater (subtractOneFromInequality snd1)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ zero) (inr (fst ,, _)) (inr (fst₁ ,, bad)) | zero ,, succ snd = TotalOrder.irreflexive ℕTotalOrder bad
|
||||
cantorInverseDiscrete (succ a) (b ,, succ (succ c)) (inr (fst ,, _)) (inr (fst₁ ,, bad)) | zero ,, succ snd = zeroNeverGreater (subtractOneFromInequality bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x1) (inl x) | succ zero ,, zero = noIntegersBetweenXAndSuccX 1 x1 x
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ zero) (inr (fst ,, snd)) (inl x) | succ zero ,, zero = TotalOrder.irreflexive ℕTotalOrder snd
|
||||
cantorInverseDiscrete (succ a) (succ b ,, succ c) (inr (fst ,, snd)) (inl x) | succ zero ,, zero rewrite Semiring.commutative ℕSemiring b (succ c) = bad fst
|
||||
where
|
||||
bad : {a : ℕ} → 1 ≡ succ (succ a) → False
|
||||
bad ()
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl y) (inl x) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.transitive ℕTotalOrder y x)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (bad ,, _)) (inl x) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _ _ _ _<N_ x bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inr (bad ,, snd)) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _ _ _ _<N_ x bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (_ ,, bad)) (inr (fst₁ ,, snd)) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = noIntegersBetweenXAndSuccX 1 bad snd
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inl y) | succ zero ,, succ snd rewrite Semiring.sumZeroRight ℕSemiring snd = noIntegersBetweenXAndSuccX (succ (succ snd)) x y
|
||||
cantorInverseDiscrete (succ a) (zero ,, c) (inr (fst ,, snd1)) (inl y) | succ zero ,, succ snd rewrite Semiring.sumZeroRight ℕSemiring snd = noIntegersBetweenXAndSuccX (succ (succ snd)) snd1 y
|
||||
cantorInverseDiscrete (succ a) (succ b ,, c) (inr (fst ,, bad)) (inl y) | succ zero ,, succ snd rewrite Semiring.sumZeroRight ℕSemiring snd | fst = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.transitive ℕTotalOrder bad (identityOfIndiscernablesRight _ _ _ _<N_ (addingIncreases c b) (Semiring.commutative ℕSemiring c (succ b))))
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inl y) | succ (succ fst) ,, succ snd rewrite Semiring.commutative ℕSemiring fst (succ (succ (succ snd))) | Semiring.commutative ℕSemiring (succ (succ snd)) fst = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.transitive ℕTotalOrder y x)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inr (y ,, z)) | succ (succ fst) ,, succ snd rewrite y | Semiring.commutative ℕSemiring fst (succ (succ (succ snd))) | Semiring.commutative ℕSemiring (succ (succ snd)) fst = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (x ,, y)) (inl z) | succ (succ fst) ,, succ snd rewrite equalityCommutative x | Semiring.commutative ℕSemiring fst (succ (succ (succ snd))) | Semiring.commutative ℕSemiring (succ (succ snd)) fst = TotalOrder.irreflexive ℕTotalOrder z
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (x ,, y)) (inr (m ,, n)) | succ (succ fst) ,, succ snd = noIntegersBetweenXAndSuccX (succ (succ snd)) y n
|
||||
|
||||
boundedInversesExist : (a : ℕ && ℕ) (s : ℕ) → order a (cantorInverse s) → Sg ℕ (λ i → cantorInverse i ≡ a)
|
||||
boundedInversesExist a zero a<s = exFalso (leastElement a<s)
|
||||
boundedInversesExist a (succ s) a<s with TotalOrder.totality totalOrder a (cantorInverse s)
|
||||
boundedInversesExist a (succ s) _ | inl (inl a<s) = boundedInversesExist a s a<s
|
||||
boundedInversesExist a (succ s) a<s | inl (inr s<a) = exFalso (cantorInverseDiscrete s a s<a a<s)
|
||||
boundedInversesExist a (succ s) a<s | inr a=s = s , equalityCommutative a=s
|
||||
|
||||
cantorInverseSurjective : (x : ℕ && ℕ) → Sg ℕ (λ i → (cantorInverse i) ≡ x)
|
||||
cantorInverseSurjective = rec orderWellfounded (λ z → Sg ℕ (λ z₁ → (cantorInverse z₁) ≡ z)) go
|
||||
where
|
||||
go : (a : ℕ && ℕ) → (pr : (x : ℕ && ℕ) (x₁ : order x a) → Sg ℕ (λ z → (cantorInverse z) ≡ x)) → Sg ℕ (λ i → (cantorInverse i) ≡ a)
|
||||
go a pr with cantorInverseLemmaSurjective a
|
||||
go .(0 ,, 0) pr | inl refl = 0 , refl
|
||||
go a ind | inr (decr , proof) with ind decr (identityOfIndiscernablesRight _ _ _ order (cantorInverseLemmaIncreases decr) proof)
|
||||
go a ind | inr (decr , proof) | boundForDecr , boundIsBound = succ boundForDecr , transitivity (applyEquality cantorInverseLemma boundIsBound) proof
|
@@ -7,6 +7,9 @@ open import Functions
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
open import WellFoundedInduction
|
||||
open import Sets.CantorBijection.CantorBijection
|
||||
|
||||
module Sets.Cardinality where
|
||||
record CountablyInfiniteSet {a : _} (A : Set a) : Set a where
|
||||
@@ -61,121 +64,6 @@ module Sets.Cardinality where
|
||||
q : succ ((a +N a *N succ a) +N 0) ≡ succ n
|
||||
q rewrite Semiring.commutative ℕSemiring (a +N a *N succ a) 0 = applyEquality succ p
|
||||
|
||||
{-
|
||||
triangle : (a : ℕ) → ℕ
|
||||
triangle zero = 0
|
||||
triangle (succ a) = succ a +N triangle a
|
||||
|
||||
cantorPairing' : (a : ℕ) → (b : ℕ) → (s : ℕ) → (s ≡ a +N b) → ℕ
|
||||
cantorPairing' a b zero pr = 0
|
||||
cantorPairing' zero zero (succ s) ()
|
||||
cantorPairing' zero (succ b) (succ s) pr = succ (cantorPairing' 1 b (succ s) pr)
|
||||
cantorPairing' (succ a) zero (succ s) pr = succ (cantorPairing' 0 a s (succInjective (identityOfIndiscernablesRight _ _ _ _≡_ pr (applyEquality (λ i → succ i) (identityOfIndiscernablesLeft _ _ _ _≡_ (Semiring.commutative ℕSemiring a 0) refl)))))
|
||||
cantorPairing' (succ a) (succ b) (succ zero) pr = naughtE f
|
||||
where
|
||||
f : 0 ≡ succ b +N a
|
||||
f rewrite Semiring.commutative ℕSemiring (succ b) a = succInjective pr
|
||||
cantorPairing' (succ a) (succ b) (succ (succ s)) pr = succ (cantorPairing' (succ (succ a)) b (succ (succ s)) (identityOfIndiscernablesRight _ _ _ _≡_ pr (applyEquality succ (succExtracts a b))))
|
||||
|
||||
cantorPairing'Zero : (x y s : ℕ) → (pr : s ≡ x +N y) → cantorPairing' x y s pr ≡ 0 → (x ≡ 0) && (y ≡ 0)
|
||||
cantorPairing'Zero x y zero pr pr' = sumZeroImpliesSummandsZero (equalityCommutative pr)
|
||||
cantorPairing'Zero zero zero (succ s) () pr'
|
||||
cantorPairing'Zero zero (succ y) (succ s) pr ()
|
||||
cantorPairing'Zero (succ x) zero (succ s) pr ()
|
||||
cantorPairing'Zero (succ x) (succ y) (succ zero) pr pr' = naughtE (identityOfIndiscernablesRight _ _ _ _≡_ (succInjective pr) (Semiring.commutative ℕSemiring x (succ y)))
|
||||
cantorPairing'Zero (succ x) (succ y) (succ (succ s)) pr ()
|
||||
|
||||
cantorPairingInj' : (s1 s2 : ℕ) → (x1 x2 y1 y2 : ℕ) → (pr1 : s1 ≡ x1 +N y1) → (pr2 : s2 ≡ x2 +N y2) → (cantorPairing' x1 y1 s1 pr1) ≡ (cantorPairing' x2 y2 s2 pr2) → (x1 ≡ x2) && (y1 ≡ y2)
|
||||
cantorPairingInj' zero zero x1 x2 y1 y2 pr1 pr2 injF = transitivity (_&&_.fst x1y1) (equalityCommutative (_&&_.fst x2y2)) ,, transitivity (_&&_.snd x1y1) (equalityCommutative (_&&_.snd x2y2))
|
||||
where
|
||||
x2y2 : (x2 ≡ 0) && (y2 ≡ 0)
|
||||
x2y2 = sumZeroImpliesSummandsZero (equalityCommutative pr2)
|
||||
x1y1 : (x1 ≡ 0) && (y1 ≡ 0)
|
||||
x1y1 = sumZeroImpliesSummandsZero (equalityCommutative pr1)
|
||||
cantorPairingInj' zero (succ s2) x1 x2 y1 y2 pr1 pr2 injF = naughtE (equalityCommutative bad2)
|
||||
where
|
||||
x2y2 : (x2 ≡ 0) && (y2 ≡ 0)
|
||||
x2y2 = cantorPairing'Zero x2 y2 (succ s2) pr2 (equalityCommutative injF)
|
||||
bad : succ s2 ≡ y2
|
||||
bad = identityOfIndiscernablesRight _ _ _ _≡_ pr2 ((applyEquality (λ i → i +N y2) (_&&_.fst x2y2)))
|
||||
bad2 : succ s2 ≡ 0
|
||||
bad2 = transitivity bad (_&&_.snd x2y2)
|
||||
cantorPairingInj' (succ s1) zero x1 x2 y1 y2 pr1 pr2 injF = naughtE (equalityCommutative bad2)
|
||||
where
|
||||
x1y1 : (x1 ≡ 0) && (y1 ≡ 0)
|
||||
x1y1 = cantorPairing'Zero x1 y1 (succ s1) pr1 injF
|
||||
bad : succ s1 ≡ y1
|
||||
bad = identityOfIndiscernablesRight _ _ _ _≡_ pr1 ((applyEquality (λ i → i +N y1) (_&&_.fst x1y1)))
|
||||
bad2 : succ s1 ≡ 0
|
||||
bad2 = transitivity bad (_&&_.snd x1y1)
|
||||
cantorPairingInj' (succ s1) (succ s2) zero x2 zero y2 () pr2 injF
|
||||
cantorPairingInj' (succ s1) (succ s2) zero zero (succ .s1) zero refl () injF
|
||||
cantorPairingInj' (succ zero) (succ zero) zero zero (succ .0) (succ .0) refl refl injF = refl ,, refl
|
||||
cantorPairingInj' (succ zero) (succ (succ s2)) zero zero (succ .0) (succ .(succ s2)) refl refl injF = naughtE (equalityCommutative bad)
|
||||
where
|
||||
bad : 2 ≡ 0
|
||||
bad = _&&_.fst (cantorPairing'Zero 2 s2 (succ (succ s2)) refl (equalityCommutative (succInjective (succInjective injF))))
|
||||
cantorPairingInj' (succ (succ s1)) (succ zero) zero zero (succ .(succ s1)) (succ .0) refl refl injF = naughtE (equalityCommutative bad)
|
||||
where
|
||||
bad : 2 ≡ 0
|
||||
bad = _&&_.fst (cantorPairing'Zero 2 s1 (succ (succ s1)) refl (succInjective (succInjective injF)))
|
||||
cantorPairingInj' (succ (succ s1)) (succ (succ s2)) zero zero (succ .(succ s1)) (succ .(succ s2)) refl refl injF = refl ,, (applyEquality (λ i → succ (succ i)) rec')
|
||||
where
|
||||
rec : cantorPairing' 2 s1 (succ (succ s1)) refl ≡ cantorPairing' 2 s2 (succ (succ s2)) refl
|
||||
rec = succInjective (succInjective injF)
|
||||
rec' : (s1 ≡ s2)
|
||||
rec' = _&&_.snd (cantorPairingInj' (succ (succ s1)) (succ (succ s2)) 2 2 s1 s2 refl refl rec)
|
||||
cantorPairingInj' (succ s1) (succ s2) zero (succ x2) (succ .s1) zero refl pr2 injF = naughtE (equalityCommutative rec')
|
||||
where
|
||||
rec : cantorPairing' 1 s1 (succ s1) refl ≡ cantorPairing' 0 x2 s2 _
|
||||
rec = succInjective injF
|
||||
rec' : (1 ≡ zero)
|
||||
rec' = _&&_.fst (cantorPairingInj' (succ s1) s2 1 zero s1 x2 refl _ rec)
|
||||
cantorPairingInj' (succ s1) (succ zero) zero (succ x2) (succ .s1) (succ y2) refl pr2 injF = naughtE (identityOfIndiscernablesRight _ _ _ _≡_ (succInjective pr2) (Semiring.commutative ℕSemiring x2 (succ y2)))
|
||||
cantorPairingInj' (succ s1) (succ (succ s2)) zero (succ x2) (succ .s1) (succ y2) refl pr2 injF = naughtE (succInjective rec')
|
||||
where
|
||||
rec : cantorPairing' 1 s1 (succ s1) refl ≡ cantorPairing' (succ (succ x2)) y2 (succ (succ s2)) _
|
||||
rec = succInjective injF
|
||||
rec' : (1 ≡ succ (succ x2))
|
||||
rec' = _&&_.fst (cantorPairingInj' (succ s1) (succ (succ s2)) 1 (succ (succ x2)) s1 y2 refl _ rec)
|
||||
cantorPairingInj' (succ s1) (succ s2) (succ x1) zero y1 zero pr1 () injF
|
||||
cantorPairingInj' (succ s1) (succ s2) (succ zero) zero .s1 (succ .s2) refl refl injF = {!!}
|
||||
cantorPairingInj' (succ s1) (succ s2) (succ (succ x1)) zero zero (succ .s2) pr1 refl injF = {!!}
|
||||
where
|
||||
false : {!!}
|
||||
false = {!!}
|
||||
cantorPairingInj' (succ s1) (succ s2) (succ (succ x1)) zero (succ y1) (succ .s2) pr1 refl injF = {!!}
|
||||
cantorPairingInj' (succ s1) (succ s2) (succ x1) (succ x2) y1 y2 pr1 pr2 injF = {!!}
|
||||
|
||||
cantorPairingInj : Injection cantorPairing
|
||||
Injection.property cantorPairingInj {zero ,, zero} {zero ,, zero} pr = refl
|
||||
Injection.property cantorPairingInj {zero ,, zero} {zero ,, succ snd} pr = exFalso {!!}
|
||||
Injection.property cantorPairingInj {zero ,, zero} {succ fst ,, snd} pr = {!!}
|
||||
Injection.property cantorPairingInj {zero ,, succ snd} {y} pr = {!!}
|
||||
Injection.property cantorPairingInj {succ fst ,, snd} {y} pr = {!!}
|
||||
|
||||
cantorInverse : ℕ → (ℕ && ℕ)
|
||||
cantorInverse zero = (0 ,, 0)
|
||||
cantorInverse (succ z) = f (cantorInverse z)
|
||||
where
|
||||
f : (ℕ && ℕ) → (ℕ && ℕ)
|
||||
f (0 ,, s) = (succ s) ,, 0
|
||||
f (succ r ,, 0) = r ,, 1
|
||||
f (succ r ,, succ s) = (r ,, succ (succ s))
|
||||
|
||||
cantorInvertible : Invertible cantorPairing
|
||||
Invertible.inverse cantorInvertible = cantorInverse
|
||||
Invertible.isLeft cantorInvertible zero = refl
|
||||
Invertible.isLeft cantorInvertible (succ b) with inspect (cantorInverse b)
|
||||
Invertible.isLeft cantorInvertible (succ b) | (zero ,, zero) with≡ x = {!!}
|
||||
Invertible.isLeft cantorInvertible (succ b) | (zero ,, succ snd) with≡ x = {!!}
|
||||
Invertible.isLeft cantorInvertible (succ b) | (succ fst ,, snd) with≡ x = {!!}
|
||||
Invertible.isRight cantorInvertible (a ,, b) = {!!}
|
||||
|
||||
cantorBijection : Bijection cantorPairing
|
||||
cantorBijection = invertibleImpliesBijection cantorInvertible
|
||||
-}
|
||||
|
||||
pairUnionIsCountable : {a b : _} {A : Set a} {B : Set b} → (X : CountablyInfiniteSet A) → (Y : CountablyInfiniteSet B) → CountablyInfiniteSet (A || B)
|
||||
CountablyInfiniteSet.counting (pairUnionIsCountable X Y) (inl r) = 2 *N (CountablyInfiniteSet.counting X r)
|
||||
CountablyInfiniteSet.counting (pairUnionIsCountable X Y) (inr s) = succ (2 *N (CountablyInfiniteSet.counting Y s))
|
||||
|
Reference in New Issue
Block a user