{-# OPTIONS --safe --warning=error --without-K #-} open import LogicalFormulae open import Functions.Definition open import Functions.Lemmas open import Numbers.Naturals.Semiring open import Numbers.Naturals.Order open import Numbers.Naturals.Order.Lemmas open import Sets.Cardinality.Finite.Definition open import Semirings.Definition open import Sets.CantorBijection.CantorBijection open import Orders.Total.Definition module Sets.Cardinality where open import Semirings.Lemmas ℕSemiring record CountablyInfiniteSet {a : _} (A : Set a) : Set a where field counting : A → ℕ countingIsBij : Bijection counting data Countable {a : _} (A : Set a) : Set a where finite : FiniteSet A → Countable A infinite : CountablyInfiniteSet A → Countable A ℕCountable : CountablyInfiniteSet ℕ ℕCountable = record { counting = id ; countingIsBij = invertibleImpliesBijection (record { inverse = id ; isLeft = λ b → refl ; isRight = λ a → refl}) } doubleLemma : (a b : ℕ) → 2 *N a ≡ 2 *N b → a ≡ b doubleLemma a b pr = productCancelsLeft 2 a b (le 1 refl) pr evenCannotBeOdd : (a b : ℕ) → 2 *N a ≡ succ (2 *N b) → False evenCannotBeOdd zero b () evenCannotBeOdd (succ a) zero pr rewrite Semiring.commutative ℕSemiring a 0 | Semiring.commutative ℕSemiring a (succ a) = naughtE (equalityCommutative (succInjective pr)) evenCannotBeOdd (succ a) (succ b) pr = evenCannotBeOdd a b pr'' where pr' : a +N a ≡ (b +N succ b) pr' rewrite Semiring.commutative ℕSemiring a 0 | Semiring.commutative ℕSemiring b 0 | Semiring.commutative ℕSemiring a (succ a) = succInjective (succInjective pr) pr'' : 2 *N a ≡ succ (2 *N b) pr'' rewrite Semiring.commutative ℕSemiring a 0 | Semiring.commutative ℕSemiring b 0 | Semiring.commutative ℕSemiring (succ b) b = pr' aMod2 : (a : ℕ) → Sg ℕ (λ i → (2 *N i ≡ a) || (succ (2 *N i) ≡ a)) aMod2 zero = (0 , inl refl) aMod2 (succ a) with aMod2 a aMod2 (succ a) | b , inl x = b , inr (applyEquality succ x) aMod2 (succ a) | b , inr x = (succ b) , inl pr where pr : succ (b +N succ (b +N 0)) ≡ succ a pr rewrite Semiring.commutative ℕSemiring b (succ (b +N 0)) | Semiring.commutative ℕSemiring (b +N 0) b = applyEquality succ x sqrtFloor : (a : ℕ) → Sg (ℕ && ℕ) (λ pair → ((_&&_.fst pair) *N (_&&_.fst pair) +N (_&&_.snd pair) ≡ a) && ((_&&_.snd pair)