Files
agdaproofs/Sets/CantorBijection/Proofs.agda
2019-09-01 10:23:25 +01:00

159 lines
18 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

{-# 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