mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-12 07:08:40 +00:00
Cleanup finset and modulo (#92)
This commit is contained in:
@@ -8,6 +8,8 @@ open import Numbers.BinaryNaturals.Multiplication
|
||||
open import Numbers.BinaryNaturals.Order
|
||||
open import Numbers.BinaryNaturals.Subtraction
|
||||
|
||||
open import Numbers.Modulo.Group
|
||||
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Integers.RingStructure.EuclideanDomain
|
||||
|
||||
@@ -83,8 +85,12 @@ open import Setoids.Orders
|
||||
open import Setoids.Functions.Definition
|
||||
open import Setoids.Functions.Extension
|
||||
|
||||
open import Sets.Cardinality.Infinite.Examples
|
||||
open import Sets.Cardinality.Infinite.Lemmas
|
||||
open import Sets.Cardinality.Countable.Lemmas
|
||||
open import Sets.Cardinality.Finite.Lemmas
|
||||
open import Sets.Cardinality
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSet.Lemmas
|
||||
|
||||
open import Decidable.Sets
|
||||
open import Decidable.Relations
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --warning=error --safe --guardedness #-}
|
||||
|
||||
-- This file contains everything that is --safe, but uses K.
|
||||
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Numbers.Primes.IntegerFactorisation
|
||||
open import Numbers.Rationals.Definition
|
||||
@@ -13,8 +12,6 @@ open import Logic.PropositionalLogic
|
||||
open import Logic.PropositionalLogicExamples
|
||||
open import Logic.PropositionalAxiomsTautology
|
||||
|
||||
open import Numbers.Modulo.Group
|
||||
|
||||
open import Sets.FinSetWithK
|
||||
|
||||
open import Rings.Examples.Examples
|
||||
|
@@ -11,13 +11,11 @@ module Functions where
|
||||
_∘_ : {a b c : _} {A : Set a} {B : Set b} {C : Set c} → (f : B → C) → (g : A → B) → (A → C)
|
||||
g ∘ f = λ a → g (f a)
|
||||
|
||||
record Injection {a b : _} {A : Set a} {B : Set b} (f : A → B) : Set (a ⊔ b) where
|
||||
field
|
||||
property : {x y : A} → (f x ≡ f y) → x ≡ y
|
||||
Injection : {a b : _} {A : Set a} {B : Set b} (f : A → B) → Set (a ⊔ b)
|
||||
Injection {A = A} f = {x y : A} → (f x ≡ f y) → x ≡ y
|
||||
|
||||
record Surjection {a b : _} {A : Set a} {B : Set b} (f : A → B) : Set (a ⊔ b) where
|
||||
field
|
||||
property : (b : B) → Sg A (λ a → f a ≡ b)
|
||||
Surjection : {a b : _} {A : Set a} {B : Set b} (f : A → B) → Set (a ⊔ b)
|
||||
Surjection {A = A} {B = B} f = (b : B) → Sg A (λ a → f a ≡ b)
|
||||
|
||||
record Bijection {a b : _} {A : Set a} {B : Set b} (f : A → B) : Set (a ⊔ b) where
|
||||
field
|
||||
@@ -31,28 +29,28 @@ module Functions where
|
||||
isRight : (a : A) → inverse (f a) ≡ a
|
||||
|
||||
invertibleImpliesBijection : {a b : _} {A : Set a} {B : Set b} {f : A → B} → Invertible f → Bijection f
|
||||
Injection.property (Bijection.inj (invertibleImpliesBijection {a} {b} {A} {B} {f} record { inverse = inverse ; isLeft = isLeft ; isRight = isRight })) {x} {y} fx=fy = ans
|
||||
Bijection.inj (invertibleImpliesBijection {a} {b} {A} {B} {f} record { inverse = inverse ; isLeft = isLeft ; isRight = isRight }) {x} {y} fx=fy = ans
|
||||
where
|
||||
bl : inverse (f x) ≡ inverse (f y)
|
||||
bl = applyEquality inverse fx=fy
|
||||
ans : x ≡ y
|
||||
ans rewrite equalityCommutative (isRight x) | equalityCommutative (isRight y) = bl
|
||||
Surjection.property (Bijection.surj (invertibleImpliesBijection {a} {b} {A} {B} {f} record { inverse = inverse ; isLeft = isLeft ; isRight = isRight })) y = (inverse y , isLeft y)
|
||||
Bijection.surj (invertibleImpliesBijection {a} {b} {A} {B} {f} record { inverse = inverse ; isLeft = isLeft ; isRight = isRight }) y = (inverse y , isLeft y)
|
||||
|
||||
bijectionImpliesInvertible : {a b : _} {A : Set a} {B : Set b} {f : A → B} → Bijection f → Invertible f
|
||||
Invertible.inverse (bijectionImpliesInvertible record { inj = inj ; surj = record { property = property } }) b = underlying (property b)
|
||||
Invertible.isLeft (bijectionImpliesInvertible {f = f} record { inj = inj ; surj = surj }) b with Surjection.property surj b
|
||||
Invertible.inverse (bijectionImpliesInvertible record { inj = inj ; surj = surj }) b = underlying (surj b)
|
||||
Invertible.isLeft (bijectionImpliesInvertible {f = f} record { inj = inj ; surj = surj }) b with surj b
|
||||
Invertible.isLeft (bijectionImpliesInvertible {f = f} record { inj = inj ; surj = surj }) b | a , prop = prop
|
||||
Invertible.isRight (bijectionImpliesInvertible {f = f} record { inj = inj ; surj = surj }) a with Surjection.property surj (f a)
|
||||
Invertible.isRight (bijectionImpliesInvertible {f = f} record { inj = record { property = property } ; surj = surj }) a | a₁ , b = property b
|
||||
Invertible.isRight (bijectionImpliesInvertible {f = f} record { inj = inj ; surj = surj }) a with surj (f a)
|
||||
Invertible.isRight (bijectionImpliesInvertible {f = f} record { inj = property ; surj = surj }) a | a₁ , b = property b
|
||||
|
||||
injComp : {a b c : _} {A : Set a} {B : Set b} {C : Set c} {f : A → B} {g : B → C} → Injection f → Injection g → Injection (g ∘ f)
|
||||
Injection.property (injComp {f = f} {g} record { property = propF } record { property = propG }) pr = propF (propG pr)
|
||||
injComp {f = f} {g} propF propG pr = propF (propG pr)
|
||||
|
||||
surjComp : {a b c : _} {A : Set a} {B : Set b} {C : Set c} {f : A → B} {g : B → C} → Surjection f → Surjection g → Surjection (g ∘ f)
|
||||
Surjection.property (surjComp {f = f} {g} record { property = propF } record { property = propG }) c with propG c
|
||||
Surjection.property (surjComp {f = f} {g} record { property = propF } record { property = propG }) c | b , pr with propF b
|
||||
Surjection.property (surjComp {f = f} {g} record { property = propF } record { property = propG }) c | b , pr | a , pr2 = a , pr'
|
||||
surjComp {f = f} {g} propF propG c with propG c
|
||||
surjComp {f = f} {g} propF propG c | b , pr with propF b
|
||||
surjComp {f = f} {g} propF propG c | b , pr | a , pr2 = a , pr'
|
||||
where
|
||||
pr' : g (f a) ≡ c
|
||||
pr' rewrite pr2 = pr
|
||||
@@ -62,18 +60,18 @@ module Functions where
|
||||
Bijection.surj (bijectionComp bijF bijG) = surjComp (Bijection.surj bijF) (Bijection.surj bijG)
|
||||
|
||||
compInjRightInj : {a b c : _} {A : Set a} {B : Set b} {C : Set c} {f : A → B} {g : B → C} → Injection (g ∘ f) → Injection f
|
||||
Injection.property (compInjRightInj {f = f} {g} record { property = property }) {x} {y} fx=fy = property (applyEquality g fx=fy)
|
||||
compInjRightInj {f = f} {g} property {x} {y} fx=fy = property (applyEquality g fx=fy)
|
||||
|
||||
compSurjLeftSurj : {a b c : _} {A : Set a} {B : Set b} {C : Set c} {f : A → B} {g : B → C} → Surjection (g ∘ f) → Surjection g
|
||||
Surjection.property (compSurjLeftSurj {f = f} {g} record { property = property }) b with property b
|
||||
Surjection.property (compSurjLeftSurj {f = f} {g} record { property = property }) b | a , b1 = f a , b1
|
||||
compSurjLeftSurj {f = f} {g} property b with property b
|
||||
compSurjLeftSurj {f = f} {g} property b | a , b1 = f a , b1
|
||||
|
||||
injectionPreservedUnderExtensionalEq : {a b : _} {A : Set a} {B : Set b} {f g : A → B} → Injection f → ({x : A} → f x ≡ g x) → Injection g
|
||||
Injection.property (injectionPreservedUnderExtensionalEq {A = A} {B} {f} {g} record { property = property } prop) {x} {y} gx=gy = property (transitivity (prop {x}) (transitivity gx=gy (equalityCommutative (prop {y}))))
|
||||
injectionPreservedUnderExtensionalEq {A = A} {B} {f} {g} property prop {x} {y} gx=gy = property (transitivity (prop {x}) (transitivity gx=gy (equalityCommutative (prop {y}))))
|
||||
|
||||
surjectionPreservedUnderExtensionalEq : {a b : _} {A : Set a} {B : Set b} {f g : A → B} → Surjection f → ({x : A} → f x ≡ g x) → Surjection g
|
||||
Surjection.property (surjectionPreservedUnderExtensionalEq {f = f} {g} surj ext) b with (Surjection.property surj b)
|
||||
Surjection.property (surjectionPreservedUnderExtensionalEq {f = f} {g} surj ext) b | a , pr = a , transitivity (equalityCommutative ext) pr
|
||||
surjectionPreservedUnderExtensionalEq {f = f} {g} surj ext b with surj b
|
||||
surjectionPreservedUnderExtensionalEq {f = f} {g} surj ext b | a , pr = a , transitivity (equalityCommutative ext) pr
|
||||
|
||||
bijectionPreservedUnderExtensionalEq : {a b : _} {A : Set a} {B : Set b} {f g : A → B} → Bijection f → ({x : A} → f x ≡ g x) → Bijection g
|
||||
Bijection.inj (bijectionPreservedUnderExtensionalEq record { inj = inj ; surj = surj } ext) = injectionPreservedUnderExtensionalEq inj ext
|
||||
@@ -88,8 +86,8 @@ module Functions where
|
||||
id a = a
|
||||
|
||||
idIsBijective : {a : _} {A : Set a} → Bijection (id {a} {A})
|
||||
Injection.property (Bijection.inj idIsBijective) pr = pr
|
||||
Surjection.property (Bijection.surj idIsBijective) b = b , refl
|
||||
Bijection.inj idIsBijective pr = pr
|
||||
Bijection.surj idIsBijective b = b , refl
|
||||
|
||||
functionCompositionExtensionallyAssociative : {a b c d : _} {A : Set a} {B : Set b} {C : Set c} {D : Set d} → (f : A → B) → (g : B → C) → (h : C → D) → (x : A) → (h ∘ (g ∘ f)) x ≡ ((h ∘ g) ∘ f) x
|
||||
functionCompositionExtensionallyAssociative f g h x = refl
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Abelian.Definition
|
||||
|
@@ -6,7 +6,6 @@ open import Setoids.DirectSum
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Abelian.Definition
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Homomorphisms.Definition
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
|
@@ -8,7 +8,6 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Integers.Addition
|
||||
open import Sets.FinSet
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Subgroups.Definition
|
||||
|
@@ -8,7 +8,6 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Integers.Addition
|
||||
open import Sets.FinSet
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
|
@@ -5,18 +5,17 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
|
||||
module Groups.Definition where
|
||||
|
||||
record Group {lvl1 lvl2} {A : Set lvl1} (S : Setoid {lvl1} {lvl2} A) (_·_ : A → A → A) : Set (lsuc lvl1 ⊔ lvl2) where
|
||||
open Setoid S
|
||||
field
|
||||
+WellDefined : {m n x y : A} → (m ∼ x) → (n ∼ y) → (m · n) ∼ (x · y)
|
||||
0G : A
|
||||
inverse : A → A
|
||||
+Associative : {a b c : A} → (a · (b · c)) ∼ (a · b) · c
|
||||
identRight : {a : A} → (a · 0G) ∼ a
|
||||
identLeft : {a : A} → (0G · a) ∼ a
|
||||
invLeft : {a : A} → (inverse a) · a ∼ 0G
|
||||
invRight : {a : A} → a · (inverse a) ∼ 0G
|
||||
record Group {lvl1 lvl2} {A : Set lvl1} (S : Setoid {lvl1} {lvl2} A) (_·_ : A → A → A) : Set (lsuc lvl1 ⊔ lvl2) where
|
||||
open Setoid S
|
||||
field
|
||||
+WellDefined : {m n x y : A} → (m ∼ x) → (n ∼ y) → (m · n) ∼ (x · y)
|
||||
0G : A
|
||||
inverse : A → A
|
||||
+Associative : {a b c : A} → (a · (b · c)) ∼ (a · b) · c
|
||||
identRight : {a : A} → (a · 0G) ∼ a
|
||||
identLeft : {a : A} → (0G · a) ∼ a
|
||||
invLeft : {a : A} → (inverse a) · a ∼ 0G
|
||||
invRight : {a : A} → a · (inverse a) ∼ 0G
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
|
@@ -6,7 +6,6 @@ open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers.Integers
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
|
@@ -5,7 +5,7 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Sets.FinSet
|
||||
open import Sets.Cardinality.Finite.Definition
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
|
@@ -5,7 +5,6 @@ open import Numbers.Naturals.Semiring -- for length
|
||||
open import Lists.Lists
|
||||
open import Functions
|
||||
open import Groups.SymmetricGroups.Definition
|
||||
open import Sets.FinSet
|
||||
open import Setoids.Setoids
|
||||
--open import Groups.Actions
|
||||
|
||||
|
@@ -5,7 +5,6 @@ open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Sets.FinSet
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
|
@@ -8,7 +8,7 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.WithK
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSet.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.SymmetricGroups.Definition
|
||||
open import Decidable.Sets
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Homomorphisms.Definition
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Homomorphisms.Definition
|
||||
|
@@ -6,7 +6,6 @@ open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Subset
|
||||
open import LogicalFormulae
|
||||
open import Sets.FinSet
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
|
@@ -5,7 +5,6 @@ open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Sets.FinSet
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Homomorphisms.Definition
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Isomorphisms.Definition
|
||||
|
@@ -8,7 +8,9 @@ open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Rationals.Definition
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSet.Definition
|
||||
open import Sets.FinSet.Lemmas
|
||||
open import Sets.Cardinality.Finite.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
@@ -20,191 +22,191 @@ open import Semirings.Definition
|
||||
|
||||
module Groups.LectureNotes.Lecture1 where
|
||||
|
||||
ℤIsGroup : _
|
||||
ℤIsGroup = ℤGroup
|
||||
ℤIsGroup : _
|
||||
ℤIsGroup = ℤGroup
|
||||
|
||||
ℚIsGroup : _
|
||||
ℚIsGroup = Ring.additiveGroup ℚRing
|
||||
ℚIsGroup : _
|
||||
ℚIsGroup = Ring.additiveGroup ℚRing
|
||||
|
||||
-- TODO: R is a group with +
|
||||
-- TODO: R is a group with +
|
||||
|
||||
integersMinusNotGroup : Group (reflSetoid ℤ) (_-Z_) → False
|
||||
integersMinusNotGroup record { +WellDefined = wellDefined ; 0G = identity ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multAssoc {nonneg 3} {nonneg 2} {nonneg 1}
|
||||
integersMinusNotGroup record { +WellDefined = wellDefined ; 0G = identity ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | ()
|
||||
integersMinusNotGroup : Group (reflSetoid ℤ) (_-Z_) → False
|
||||
integersMinusNotGroup record { +WellDefined = wellDefined ; 0G = identity ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multAssoc {nonneg 3} {nonneg 2} {nonneg 1}
|
||||
integersMinusNotGroup record { +WellDefined = wellDefined ; 0G = identity ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | ()
|
||||
|
||||
negSuccInjective : {a b : ℕ} → (negSucc a ≡ negSucc b) → a ≡ b
|
||||
negSuccInjective {a} {.a} refl = refl
|
||||
negSuccInjective : {a b : ℕ} → (negSucc a ≡ negSucc b) → a ≡ b
|
||||
negSuccInjective {a} {.a} refl = refl
|
||||
|
||||
nonnegInjective : {a b : ℕ} → (nonneg a ≡ nonneg b) → a ≡ b
|
||||
nonnegInjective {a} {.a} refl = refl
|
||||
nonnegInjective : {a b : ℕ} → (nonneg a ≡ nonneg b) → a ≡ b
|
||||
nonnegInjective {a} {.a} refl = refl
|
||||
|
||||
integersTimesNotGroup : Group (reflSetoid ℤ) (_*Z_) → False
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg zero) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {negSucc 1}
|
||||
... | ()
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with invLeft {nonneg zero}
|
||||
... | bl with inverse (nonneg zero)
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | nonneg zero
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | p | nonneg (succ x) = naughtE (nonnegInjective (transitivity (applyEquality nonneg (equalityCommutative (Semiring.productZeroRight ℕSemiring x))) p))
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | negSucc x
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ (succ x))) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with succInjective (negSuccInjective (multIdentLeft {negSucc 1}))
|
||||
... | ()
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (negSucc x) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {nonneg 2}
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (negSucc x) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | ()
|
||||
integersTimesNotGroup : Group (reflSetoid ℤ) (_*Z_) → False
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg zero) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {negSucc 1}
|
||||
... | ()
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with invLeft {nonneg zero}
|
||||
... | bl with inverse (nonneg zero)
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | nonneg zero
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | p | nonneg (succ x) = naughtE (nonnegInjective (transitivity (applyEquality nonneg (equalityCommutative (Semiring.productZeroRight ℕSemiring x))) p))
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ zero)) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | negSucc x
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (nonneg (succ (succ x))) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with succInjective (negSuccInjective (multIdentLeft {negSucc 1}))
|
||||
... | ()
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (negSucc x) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {nonneg 2}
|
||||
integersTimesNotGroup record { +WellDefined = wellDefined ; 0G = (negSucc x) ; inverse = inverse ; +Associative = multAssoc ; identRight = multIdentRight ; identLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | ()
|
||||
|
||||
-- TODO: Q is not a group with *Q
|
||||
-- TODO: Q without 0 is a group with *Q
|
||||
-- TODO: {1, -1} is a group with *
|
||||
-- TODO: Q is not a group with *Q
|
||||
-- TODO: Q without 0 is a group with *Q
|
||||
-- TODO: {1, -1} is a group with *
|
||||
|
||||
ℤnIsGroup : (n : ℕ) → (0<n : 0 <N n) → _
|
||||
ℤnIsGroup n pr = ℤnGroup n pr
|
||||
ℤnIsGroup : (n : ℕ) → (0<n : 0 <N n) → _
|
||||
ℤnIsGroup n pr = ℤnGroup n pr
|
||||
|
||||
-- Groups example 8.9 from lecture 1
|
||||
data Weird : Set where
|
||||
e : Weird
|
||||
a : Weird
|
||||
b : Weird
|
||||
c : Weird
|
||||
-- Groups example 8.9 from lecture 1
|
||||
data Weird : Set where
|
||||
e : Weird
|
||||
a : Weird
|
||||
b : Weird
|
||||
c : Weird
|
||||
|
||||
_+W_ : Weird → Weird → Weird
|
||||
e +W t = t
|
||||
a +W e = a
|
||||
a +W a = e
|
||||
a +W b = c
|
||||
a +W c = b
|
||||
b +W e = b
|
||||
b +W a = c
|
||||
b +W b = e
|
||||
b +W c = a
|
||||
c +W e = c
|
||||
c +W a = b
|
||||
c +W b = a
|
||||
c +W c = e
|
||||
_+W_ : Weird → Weird → Weird
|
||||
e +W t = t
|
||||
a +W e = a
|
||||
a +W a = e
|
||||
a +W b = c
|
||||
a +W c = b
|
||||
b +W e = b
|
||||
b +W a = c
|
||||
b +W b = e
|
||||
b +W c = a
|
||||
c +W e = c
|
||||
c +W a = b
|
||||
c +W b = a
|
||||
c +W c = e
|
||||
|
||||
+WAssoc : {x y z : Weird} → (x +W (y +W z)) ≡ (x +W y) +W z
|
||||
+WAssoc {e} {y} {z} = refl
|
||||
+WAssoc {a} {e} {z} = refl
|
||||
+WAssoc {a} {a} {e} = refl
|
||||
+WAssoc {a} {a} {a} = refl
|
||||
+WAssoc {a} {a} {b} = refl
|
||||
+WAssoc {a} {a} {c} = refl
|
||||
+WAssoc {a} {b} {e} = refl
|
||||
+WAssoc {a} {b} {a} = refl
|
||||
+WAssoc {a} {b} {b} = refl
|
||||
+WAssoc {a} {b} {c} = refl
|
||||
+WAssoc {a} {c} {e} = refl
|
||||
+WAssoc {a} {c} {a} = refl
|
||||
+WAssoc {a} {c} {b} = refl
|
||||
+WAssoc {a} {c} {c} = refl
|
||||
+WAssoc {b} {e} {z} = refl
|
||||
+WAssoc {b} {a} {e} = refl
|
||||
+WAssoc {b} {a} {a} = refl
|
||||
+WAssoc {b} {a} {b} = refl
|
||||
+WAssoc {b} {a} {c} = refl
|
||||
+WAssoc {b} {b} {e} = refl
|
||||
+WAssoc {b} {b} {a} = refl
|
||||
+WAssoc {b} {b} {b} = refl
|
||||
+WAssoc {b} {b} {c} = refl
|
||||
+WAssoc {b} {c} {e} = refl
|
||||
+WAssoc {b} {c} {a} = refl
|
||||
+WAssoc {b} {c} {b} = refl
|
||||
+WAssoc {b} {c} {c} = refl
|
||||
+WAssoc {c} {e} {z} = refl
|
||||
+WAssoc {c} {a} {e} = refl
|
||||
+WAssoc {c} {a} {a} = refl
|
||||
+WAssoc {c} {a} {b} = refl
|
||||
+WAssoc {c} {a} {c} = refl
|
||||
+WAssoc {c} {b} {e} = refl
|
||||
+WAssoc {c} {b} {a} = refl
|
||||
+WAssoc {c} {b} {b} = refl
|
||||
+WAssoc {c} {b} {c} = refl
|
||||
+WAssoc {c} {c} {e} = refl
|
||||
+WAssoc {c} {c} {a} = refl
|
||||
+WAssoc {c} {c} {b} = refl
|
||||
+WAssoc {c} {c} {c} = refl
|
||||
+WAssoc : {x y z : Weird} → (x +W (y +W z)) ≡ (x +W y) +W z
|
||||
+WAssoc {e} {y} {z} = refl
|
||||
+WAssoc {a} {e} {z} = refl
|
||||
+WAssoc {a} {a} {e} = refl
|
||||
+WAssoc {a} {a} {a} = refl
|
||||
+WAssoc {a} {a} {b} = refl
|
||||
+WAssoc {a} {a} {c} = refl
|
||||
+WAssoc {a} {b} {e} = refl
|
||||
+WAssoc {a} {b} {a} = refl
|
||||
+WAssoc {a} {b} {b} = refl
|
||||
+WAssoc {a} {b} {c} = refl
|
||||
+WAssoc {a} {c} {e} = refl
|
||||
+WAssoc {a} {c} {a} = refl
|
||||
+WAssoc {a} {c} {b} = refl
|
||||
+WAssoc {a} {c} {c} = refl
|
||||
+WAssoc {b} {e} {z} = refl
|
||||
+WAssoc {b} {a} {e} = refl
|
||||
+WAssoc {b} {a} {a} = refl
|
||||
+WAssoc {b} {a} {b} = refl
|
||||
+WAssoc {b} {a} {c} = refl
|
||||
+WAssoc {b} {b} {e} = refl
|
||||
+WAssoc {b} {b} {a} = refl
|
||||
+WAssoc {b} {b} {b} = refl
|
||||
+WAssoc {b} {b} {c} = refl
|
||||
+WAssoc {b} {c} {e} = refl
|
||||
+WAssoc {b} {c} {a} = refl
|
||||
+WAssoc {b} {c} {b} = refl
|
||||
+WAssoc {b} {c} {c} = refl
|
||||
+WAssoc {c} {e} {z} = refl
|
||||
+WAssoc {c} {a} {e} = refl
|
||||
+WAssoc {c} {a} {a} = refl
|
||||
+WAssoc {c} {a} {b} = refl
|
||||
+WAssoc {c} {a} {c} = refl
|
||||
+WAssoc {c} {b} {e} = refl
|
||||
+WAssoc {c} {b} {a} = refl
|
||||
+WAssoc {c} {b} {b} = refl
|
||||
+WAssoc {c} {b} {c} = refl
|
||||
+WAssoc {c} {c} {e} = refl
|
||||
+WAssoc {c} {c} {a} = refl
|
||||
+WAssoc {c} {c} {b} = refl
|
||||
+WAssoc {c} {c} {c} = refl
|
||||
|
||||
weirdGroup : Group (reflSetoid Weird) _+W_
|
||||
Group.+WellDefined weirdGroup = reflGroupWellDefined
|
||||
Group.0G weirdGroup = e
|
||||
Group.inverse weirdGroup t = t
|
||||
Group.+Associative weirdGroup {r} {s} {t} = +WAssoc {r} {s} {t}
|
||||
Group.identRight weirdGroup {e} = refl
|
||||
Group.identRight weirdGroup {a} = refl
|
||||
Group.identRight weirdGroup {b} = refl
|
||||
Group.identRight weirdGroup {c} = refl
|
||||
Group.identLeft weirdGroup {e} = refl
|
||||
Group.identLeft weirdGroup {a} = refl
|
||||
Group.identLeft weirdGroup {b} = refl
|
||||
Group.identLeft weirdGroup {c} = refl
|
||||
Group.invLeft weirdGroup {e} = refl
|
||||
Group.invLeft weirdGroup {a} = refl
|
||||
Group.invLeft weirdGroup {b} = refl
|
||||
Group.invLeft weirdGroup {c} = refl
|
||||
Group.invRight weirdGroup {e} = refl
|
||||
Group.invRight weirdGroup {a} = refl
|
||||
Group.invRight weirdGroup {b} = refl
|
||||
Group.invRight weirdGroup {c} = refl
|
||||
weirdGroup : Group (reflSetoid Weird) _+W_
|
||||
Group.+WellDefined weirdGroup = reflGroupWellDefined
|
||||
Group.0G weirdGroup = e
|
||||
Group.inverse weirdGroup t = t
|
||||
Group.+Associative weirdGroup {r} {s} {t} = +WAssoc {r} {s} {t}
|
||||
Group.identRight weirdGroup {e} = refl
|
||||
Group.identRight weirdGroup {a} = refl
|
||||
Group.identRight weirdGroup {b} = refl
|
||||
Group.identRight weirdGroup {c} = refl
|
||||
Group.identLeft weirdGroup {e} = refl
|
||||
Group.identLeft weirdGroup {a} = refl
|
||||
Group.identLeft weirdGroup {b} = refl
|
||||
Group.identLeft weirdGroup {c} = refl
|
||||
Group.invLeft weirdGroup {e} = refl
|
||||
Group.invLeft weirdGroup {a} = refl
|
||||
Group.invLeft weirdGroup {b} = refl
|
||||
Group.invLeft weirdGroup {c} = refl
|
||||
Group.invRight weirdGroup {e} = refl
|
||||
Group.invRight weirdGroup {a} = refl
|
||||
Group.invRight weirdGroup {b} = refl
|
||||
Group.invRight weirdGroup {c} = refl
|
||||
|
||||
weirdAb : AbelianGroup weirdGroup
|
||||
AbelianGroup.commutative weirdAb {e} {e} = refl
|
||||
AbelianGroup.commutative weirdAb {e} {a} = refl
|
||||
AbelianGroup.commutative weirdAb {e} {b} = refl
|
||||
AbelianGroup.commutative weirdAb {e} {c} = refl
|
||||
AbelianGroup.commutative weirdAb {a} {e} = refl
|
||||
AbelianGroup.commutative weirdAb {a} {a} = refl
|
||||
AbelianGroup.commutative weirdAb {a} {b} = refl
|
||||
AbelianGroup.commutative weirdAb {a} {c} = refl
|
||||
AbelianGroup.commutative weirdAb {b} {e} = refl
|
||||
AbelianGroup.commutative weirdAb {b} {a} = refl
|
||||
AbelianGroup.commutative weirdAb {b} {b} = refl
|
||||
AbelianGroup.commutative weirdAb {b} {c} = refl
|
||||
AbelianGroup.commutative weirdAb {c} {e} = refl
|
||||
AbelianGroup.commutative weirdAb {c} {a} = refl
|
||||
AbelianGroup.commutative weirdAb {c} {b} = refl
|
||||
AbelianGroup.commutative weirdAb {c} {c} = refl
|
||||
weirdAb : AbelianGroup weirdGroup
|
||||
AbelianGroup.commutative weirdAb {e} {e} = refl
|
||||
AbelianGroup.commutative weirdAb {e} {a} = refl
|
||||
AbelianGroup.commutative weirdAb {e} {b} = refl
|
||||
AbelianGroup.commutative weirdAb {e} {c} = refl
|
||||
AbelianGroup.commutative weirdAb {a} {e} = refl
|
||||
AbelianGroup.commutative weirdAb {a} {a} = refl
|
||||
AbelianGroup.commutative weirdAb {a} {b} = refl
|
||||
AbelianGroup.commutative weirdAb {a} {c} = refl
|
||||
AbelianGroup.commutative weirdAb {b} {e} = refl
|
||||
AbelianGroup.commutative weirdAb {b} {a} = refl
|
||||
AbelianGroup.commutative weirdAb {b} {b} = refl
|
||||
AbelianGroup.commutative weirdAb {b} {c} = refl
|
||||
AbelianGroup.commutative weirdAb {c} {e} = refl
|
||||
AbelianGroup.commutative weirdAb {c} {a} = refl
|
||||
AbelianGroup.commutative weirdAb {c} {b} = refl
|
||||
AbelianGroup.commutative weirdAb {c} {c} = refl
|
||||
|
||||
weirdProjection : Weird → FinSet 4
|
||||
weirdProjection a = ofNat 0 (le 3 refl)
|
||||
weirdProjection b = ofNat 1 (le 2 refl)
|
||||
weirdProjection c = ofNat 2 (le 1 refl)
|
||||
weirdProjection e = ofNat 3 (le zero refl)
|
||||
weirdProjection : Weird → FinSet 4
|
||||
weirdProjection a = ofNat 0 (le 3 refl)
|
||||
weirdProjection b = ofNat 1 (le 2 refl)
|
||||
weirdProjection c = ofNat 2 (le 1 refl)
|
||||
weirdProjection e = ofNat 3 (le zero refl)
|
||||
|
||||
weirdProjectionSurj : Surjection weirdProjection
|
||||
Surjection.property weirdProjectionSurj fzero = a , refl
|
||||
Surjection.property weirdProjectionSurj (fsucc fzero) = b , refl
|
||||
Surjection.property weirdProjectionSurj (fsucc (fsucc fzero)) = c , refl
|
||||
Surjection.property weirdProjectionSurj (fsucc (fsucc (fsucc fzero))) = e , refl
|
||||
Surjection.property weirdProjectionSurj (fsucc (fsucc (fsucc (fsucc ()))))
|
||||
weirdProjectionSurj : Surjection weirdProjection
|
||||
weirdProjectionSurj fzero = a , refl
|
||||
weirdProjectionSurj (fsucc fzero) = b , refl
|
||||
weirdProjectionSurj (fsucc (fsucc fzero)) = c , refl
|
||||
weirdProjectionSurj (fsucc (fsucc (fsucc fzero))) = e , refl
|
||||
weirdProjectionSurj (fsucc (fsucc (fsucc (fsucc ()))))
|
||||
|
||||
weirdProjectionInj : (x y : Weird) → weirdProjection x ≡ weirdProjection y → Setoid._∼_ (reflSetoid Weird) x y
|
||||
weirdProjectionInj e e fx=fy = refl
|
||||
weirdProjectionInj e a ()
|
||||
weirdProjectionInj e b ()
|
||||
weirdProjectionInj e c ()
|
||||
weirdProjectionInj a e ()
|
||||
weirdProjectionInj a a fx=fy = refl
|
||||
weirdProjectionInj a b ()
|
||||
weirdProjectionInj a c ()
|
||||
weirdProjectionInj b e ()
|
||||
weirdProjectionInj b a ()
|
||||
weirdProjectionInj b b fx=fy = refl
|
||||
weirdProjectionInj b c ()
|
||||
weirdProjectionInj c e ()
|
||||
weirdProjectionInj c a ()
|
||||
weirdProjectionInj c b ()
|
||||
weirdProjectionInj c c fx=fy = refl
|
||||
weirdProjectionInj : (x y : Weird) → weirdProjection x ≡ weirdProjection y → Setoid._∼_ (reflSetoid Weird) x y
|
||||
weirdProjectionInj e e fx=fy = refl
|
||||
weirdProjectionInj e a ()
|
||||
weirdProjectionInj e b ()
|
||||
weirdProjectionInj e c ()
|
||||
weirdProjectionInj a e ()
|
||||
weirdProjectionInj a a fx=fy = refl
|
||||
weirdProjectionInj a b ()
|
||||
weirdProjectionInj a c ()
|
||||
weirdProjectionInj b e ()
|
||||
weirdProjectionInj b a ()
|
||||
weirdProjectionInj b b fx=fy = refl
|
||||
weirdProjectionInj b c ()
|
||||
weirdProjectionInj c e ()
|
||||
weirdProjectionInj c a ()
|
||||
weirdProjectionInj c b ()
|
||||
weirdProjectionInj c c fx=fy = refl
|
||||
|
||||
weirdFinite : FiniteGroup weirdGroup (FinSet 4)
|
||||
SetoidToSet.project (FiniteGroup.toSet weirdFinite) = weirdProjection
|
||||
SetoidToSet.wellDefined (FiniteGroup.toSet weirdFinite) x y = applyEquality weirdProjection
|
||||
SetoidToSet.surj (FiniteGroup.toSet weirdFinite) = weirdProjectionSurj
|
||||
SetoidToSet.inj (FiniteGroup.toSet weirdFinite) = weirdProjectionInj
|
||||
FiniteSet.size (FiniteGroup.finite weirdFinite) = 4
|
||||
FiniteSet.mapping (FiniteGroup.finite weirdFinite) = id
|
||||
FiniteSet.bij (FiniteGroup.finite weirdFinite) = idIsBijective
|
||||
weirdFinite : FiniteGroup weirdGroup (FinSet 4)
|
||||
SetoidToSet.project (FiniteGroup.toSet weirdFinite) = weirdProjection
|
||||
SetoidToSet.wellDefined (FiniteGroup.toSet weirdFinite) x y = applyEquality weirdProjection
|
||||
SetoidToSet.surj (FiniteGroup.toSet weirdFinite) = weirdProjectionSurj
|
||||
SetoidToSet.inj (FiniteGroup.toSet weirdFinite) = weirdProjectionInj
|
||||
FiniteSet.size (FiniteGroup.finite weirdFinite) = 4
|
||||
FiniteSet.mapping (FiniteGroup.finite weirdFinite) = id
|
||||
FiniteSet.bij (FiniteGroup.finite weirdFinite) = idIsBijective
|
||||
|
||||
weirdOrder : groupOrder weirdFinite ≡ 4
|
||||
weirdOrder = refl
|
||||
weirdOrder : groupOrder weirdFinite ≡ 4
|
||||
weirdOrder = refl
|
||||
|
||||
-- TODO: dihedral groups
|
||||
-- TODO: matrix groups on R
|
||||
-- TODO: general linear groups on R
|
||||
-- TODO: dihedral groups
|
||||
-- TODO: matrix groups on R
|
||||
-- TODO: general linear groups on R
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Sets.FinSet
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Sets.FinSet
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
|
||||
|
@@ -6,7 +6,6 @@ open import Sets.EquivalenceRelations
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
|
||||
module Groups.Subgroups.Examples {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} (G : Group S _+_) where
|
||||
|
@@ -5,7 +5,6 @@ open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Sets.FinSet
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
|
@@ -5,7 +5,6 @@ open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Sets.FinSet
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
|
@@ -5,7 +5,6 @@ open import Groups.Definition
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Sets.FinSet
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
@@ -5,7 +5,6 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
|
@@ -1,7 +1,9 @@
|
||||
{-# OPTIONS --warning=error --safe #-}
|
||||
|
||||
open import Functions
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSet.Definition
|
||||
open import Sets.FinSet.Lemmas
|
||||
open import Sets.Cardinality.Finite.Definition
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
@@ -169,11 +171,11 @@ weirdProjection c = ofNat 2 (le 1 refl)
|
||||
weirdProjection e = ofNat 3 (le zero refl)
|
||||
|
||||
weirdProjectionSurj : Surjection weirdProjection
|
||||
Surjection.property weirdProjectionSurj fzero = a , refl
|
||||
Surjection.property weirdProjectionSurj (fsucc fzero) = b , refl
|
||||
Surjection.property weirdProjectionSurj (fsucc (fsucc fzero)) = c , refl
|
||||
Surjection.property weirdProjectionSurj (fsucc (fsucc (fsucc fzero))) = e , refl
|
||||
Surjection.property weirdProjectionSurj (fsucc (fsucc (fsucc (fsucc ()))))
|
||||
weirdProjectionSurj fzero = a , refl
|
||||
weirdProjectionSurj (fsucc fzero) = b , refl
|
||||
weirdProjectionSurj (fsucc (fsucc fzero)) = c , refl
|
||||
weirdProjectionSurj (fsucc (fsucc (fsucc fzero))) = e , refl
|
||||
weirdProjectionSurj (fsucc (fsucc (fsucc (fsucc ()))))
|
||||
|
||||
weirdProjectionInj : (x y : Weird) → weirdProjection x ≡ weirdProjection y → Setoid._∼_ (reflSetoid Weird) x y
|
||||
weirdProjectionInj e e fx=fy = refl
|
||||
|
@@ -48,12 +48,12 @@ twoElementSubset {P = P} {Q} {R} One = implies P Q
|
||||
twoElementSubset {P = P} {Q} {R} Two = implies Q R
|
||||
|
||||
twoElementSubsetInj : {a : _} {A : Set a} {P Q R : Propositions A} → (P ≡ Q → False) → Injection (twoElementSubset {P = P} {Q} {R})
|
||||
Injection.property (twoElementSubsetInj {P = P} {Q} {R} p!=q) {One} {One} refl = refl
|
||||
Injection.property (twoElementSubsetInj {P = P} {Q} {R} p!=q) {One} {Two} pr with impliesInjective pr
|
||||
Injection.property (twoElementSubsetInj {P = P} {Q} {R} p!=q) {One} {Two} pr | fst ,, snd = exFalso (p!=q fst)
|
||||
Injection.property (twoElementSubsetInj {P = P} {Q} {R} p!=q) {Two} {One} pr with impliesInjective pr
|
||||
Injection.property (twoElementSubsetInj {P = P} {Q} {R} p!=q) {Two} {One} pr | fst ,, snd = exFalso (p!=q (equalityCommutative fst))
|
||||
Injection.property (twoElementSubsetInj {P = P} {Q} {R} p!=q) {Two} {Two} refl = refl
|
||||
twoElementSubsetInj {P = P} {Q} {R} p!=q {One} {One} refl = refl
|
||||
twoElementSubsetInj {P = P} {Q} {R} p!=q {One} {Two} pr with impliesInjective pr
|
||||
twoElementSubsetInj {P = P} {Q} {R} p!=q {One} {Two} pr | fst ,, snd = exFalso (p!=q fst)
|
||||
twoElementSubsetInj {P = P} {Q} {R} p!=q {Two} {One} pr with impliesInjective pr
|
||||
twoElementSubsetInj {P = P} {Q} {R} p!=q {Two} {One} pr | fst ,, snd = exFalso (p!=q (equalityCommutative fst))
|
||||
twoElementSubsetInj {P = P} {Q} {R} p!=q {Two} {Two} refl = refl
|
||||
|
||||
badBool : BoolFalse ≡ BoolTrue → False
|
||||
badBool ()
|
||||
@@ -73,12 +73,12 @@ pQQR {P = P} {Q} {R} One = implies P Q
|
||||
pQQR {P = P} {Q} {R} Two = implies Q R
|
||||
|
||||
pQQRSubsetInj : {a : _} {A : Set a} {P Q R : Propositions A} → (P ≡ Q → False) → Injection (pQQR {P = P} {Q} {R})
|
||||
Injection.property (pQQRSubsetInj {P = P} {Q} {R} p!=q) {One} {One} refl = refl
|
||||
Injection.property (pQQRSubsetInj {P = P} {Q} {R} p!=q) {One} {Two} pr with impliesInjective pr
|
||||
Injection.property (pQQRSubsetInj {P = P} {Q} {R} p!=q) {One} {Two} pr | p=q ,, _ = exFalso (p!=q p=q)
|
||||
Injection.property (pQQRSubsetInj {P = P} {Q} {R} p!=q) {Two} {One} pr with impliesInjective pr
|
||||
Injection.property (pQQRSubsetInj {P = P} {Q} {R} p!=q) {Two} {One} pr | q=p ,, _ = exFalso (p!=q (equalityCommutative q=p))
|
||||
Injection.property (pQQRSubsetInj {P = P} {Q} {R} p!=q) {Two} {Two} refl = refl
|
||||
pQQRSubsetInj {P = P} {Q} {R} p!=q {One} {One} refl = refl
|
||||
pQQRSubsetInj {P = P} {Q} {R} p!=q {One} {Two} pr with impliesInjective pr
|
||||
pQQRSubsetInj {P = P} {Q} {R} p!=q {One} {Two} pr | p=q ,, _ = exFalso (p!=q p=q)
|
||||
pQQRSubsetInj {P = P} {Q} {R} p!=q {Two} {One} pr with impliesInjective pr
|
||||
pQQRSubsetInj {P = P} {Q} {R} p!=q {Two} {One} pr | q=p ,, _ = exFalso (p!=q (equalityCommutative q=p))
|
||||
pQQRSubsetInj {P = P} {Q} {R} p!=q {Two} {Two} refl = refl
|
||||
|
||||
syntacticEntailmentExample : {a : _} {A : Set a} {P Q R : Propositions A} → (p!=q : P ≡ Q → False) → Proof propositionalAxioms (record { ofElt = pQQR {P = P} {Q} {R} }) 7
|
||||
syntacticEntailmentExample {P = P} {Q} {R} p!=q = nextStep 6 (nextStep 5 (nextStep 4 (nextStep 3 (nextStep 2 (nextStep 1 (nextStep 0 empty (axiom (Two , record { one = P ; two = Q ; three = R }))) (given Two)) (axiom (One , ((implies Q R) ,, P)))) (modusPonens (record { element = implies (implies Q R) (implies P (implies Q R)) ; position = 0 ; pos<N = le 2 refl ; elementIsAt = refl }) (record { element = implies Q R ; position = 1 ; elementIsAt = refl ; pos<N = le 1 refl }) (implies P (implies Q R)) refl)) (modusPonens (record { element = implies (implies P (implies Q R)) (implies (implies P Q) (implies P R)) ; position = 3 ; pos<N = le 0 refl ; elementIsAt = refl }) (record { element = implies P (implies Q R) ; position = 0 ; pos<N = succIsPositive _ ; elementIsAt = refl }) (implies (implies P Q) (implies P R)) refl)) (given One)) (modusPonens (record { element = implies (implies P Q) (implies P R) ; position = 1 ; pos<N = succPreservesInequality (succIsPositive _) ; elementIsAt = refl }) (record { element = implies P Q ; position = 0 ; pos<N = succIsPositive _ ; elementIsAt = refl }) (implies P R) refl)
|
||||
@@ -91,7 +91,7 @@ pSubset : {a : _} {A : Set a} {P : Propositions A} → True → Propositions A
|
||||
pSubset {P = P} record {} = implies P P
|
||||
|
||||
pSubsetInj : {a : _} {A : Set a} {P : Propositions A} → Injection (pSubset {P = P})
|
||||
Injection.property (pSubsetInj {P = P}) {record {}} {record {}} refl = refl
|
||||
pSubsetInj {P = P} {record {}} {record {}} refl = refl
|
||||
|
||||
syntacticEntailmentPP : {a : _} {A : Set a} {P : Propositions A} → Proof propositionalAxioms record { ofElt = pSubset {P = P} } 5
|
||||
syntacticEntailmentPP {P = P} = nextStep 4 (nextStep 3 (nextStep 2 (nextStep 1 (nextStep 0 empty (axiom (One , (P ,, (implies P P))))) (axiom (Two , record { one = P ; two = implies P P ; three = P }))) (modusPonens (record { element = implies (implies P (implies (implies P P) P)) (implies (implies P (implies P P)) (implies P P)) ; position = 0 ; pos<N = succIsPositive _ ; elementIsAt = refl }) (record { element = implies P (implies (implies P P) P) ; position = 1 ; pos<N = succPreservesInequality (succIsPositive _) ; elementIsAt = refl }) (implies (implies P (implies P P)) (implies P P)) refl)) (axiom (One , (P ,, P)))) (modusPonens (record { element = implies (implies P (implies P P)) (implies P P) ; position = 1 ; pos<N = succPreservesInequality (succIsPositive _) ; elementIsAt = refl }) (record { element = implies P (implies P P) ; position = 0 ; pos<N = succIsPositive _ ; elementIsAt = refl }) (implies P P) refl)
|
||||
|
@@ -66,7 +66,7 @@ identityOfIndiscernablesRight {a = a} {b} {.b} prop prAB refl = prAB
|
||||
equalityCommutative : {a : _} {A : Set a} {x y : A} → (x ≡ y) → (y ≡ x)
|
||||
equalityCommutative refl = refl
|
||||
|
||||
exFalso : {n : _} {A : Set n} → False → A
|
||||
exFalso : {n : _} {A : Set n} → .(x : False) → A
|
||||
exFalso {a} = λ ()
|
||||
|
||||
orIsAssociative : {n : _} {a b c : Set n} → ((a || b) || c) → (a || (b || c))
|
||||
|
@@ -1,4 +1,4 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
-- These are explicitly with-K, because we currently encode an element of Zn as
|
||||
-- a natural together with a proof that it is small.
|
||||
|
||||
@@ -10,95 +10,48 @@ open import Groups.FiniteGroups.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSetWithK
|
||||
open import Sets.FinSet.Definition
|
||||
open import Functions
|
||||
open import Numbers.Naturals.WithK
|
||||
open import Semirings.Definition
|
||||
open import Orders.Total.Definition
|
||||
open import Numbers.Modulo.Definition
|
||||
open import Numbers.Modulo.ModuloFunction
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
|
||||
module Numbers.Modulo.Addition where
|
||||
open TotalOrder ℕTotalOrder
|
||||
|
||||
cancelSumFromInequality : {a b c : ℕ} → a +N b <N a +N c → b <N c
|
||||
cancelSumFromInequality {a} {b} {c} (le x proof) = le x help
|
||||
where
|
||||
help : succ x +N b ≡ c
|
||||
help rewrite Semiring.+Associative ℕSemiring (succ x) a b | Semiring.commutative ℕSemiring x a | equalityCommutative (Semiring.+Associative ℕSemiring (succ a) x b) | Semiring.commutative ℕSemiring a (x +N b) | Semiring.commutative ℕSemiring (succ (x +N b)) a = canSubtractFromEqualityLeft {a} proof
|
||||
_+n_ : {n : ℕ} .(pr : 0 <N n) → ℤn n pr → ℤn n pr → ℤn n pr
|
||||
_+n_ {n} pr a b = record { x = mod n pr (ℤn.x a +N ℤn.x b) ; xLess = mod<N pr _ }
|
||||
|
||||
_+n_ : {n : ℕ} {pr : 0 <N n} → ℤn n pr → ℤn n pr → ℤn n pr
|
||||
_+n_ {0} {le x ()} a b
|
||||
_+n_ {succ n} {pr} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } with totality (a +N b) (succ n)
|
||||
_+n_ {succ n} {pr} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl (a+b<n)) = record { x = a +N b ; xLess = a+b<n }
|
||||
_+n_ {succ n} {pr} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr (n<a+b)) = record { x = subtractionNResult.result sub ; xLess = pr2 }
|
||||
where
|
||||
sub : subtractionNResult (succ n) (a +N b) (inl n<a+b)
|
||||
sub = -N (inl n<a+b)
|
||||
pr1 : a +N b <N succ n +N succ n
|
||||
pr1 = addStrongInequalities aLess bLess
|
||||
pr1' : succ n +N (subtractionNResult.result sub) <N succ n +N succ n
|
||||
pr1' rewrite subtractionNResult.pr sub = pr1
|
||||
pr2 : subtractionNResult.result sub <N succ n
|
||||
pr2 = cancelSumFromInequality pr1'
|
||||
_+n_ {succ n} {pr} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr (a+b=n) = record { x = 0 ; xLess = succIsPositive n }
|
||||
plusZnIdentityRight : {n : ℕ} → .(pr : 0 <N n) → (a : ℤn n pr) → (_+n_ pr a record { x = 0 ; xLess = pr }) ≡ a
|
||||
plusZnIdentityRight 0<n record { x = x ; xLess = xLess } rewrite Semiring.sumZeroRight ℕSemiring x = equalityZn (modIsMod 0<n x (<NProp xLess))
|
||||
|
||||
plusZnIdentityRight : {n : ℕ} → {pr : 0 <N n} → (a : ℤn n pr) → (a +n record { x = 0 ; xLess = pr }) ≡ a
|
||||
plusZnIdentityRight {zero} {()} a
|
||||
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } with totality (a +N 0) (succ x)
|
||||
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inl (inl a<sx) = equalityZn _ _ (Semiring.commutative ℕSemiring a 0)
|
||||
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inl (inr sx<a) = exFalso (f aLess sx<a)
|
||||
where
|
||||
f : (aL : a <N succ x) → (sx<a : succ x <N a +N 0) → False
|
||||
f aL sx<a rewrite Semiring.commutative ℕSemiring a 0 = irreflexive (<Transitive aL sx<a)
|
||||
plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inr a=sx = exFalso (f aLess a=sx)
|
||||
where
|
||||
f : (aL : a <N succ x) → (a=sx : a +N 0 ≡ succ x) → False
|
||||
f aL a=sx rewrite Semiring.commutative ℕSemiring a 0 | a=sx = TotalOrder.irreflexive ℕTotalOrder aL
|
||||
plusZnIdentityLeft : {n : ℕ} → .(pr : 0 <N n) → (a : ℤn n pr) → _+n_ pr (record { x = 0 ; xLess = pr }) a ≡ a
|
||||
plusZnIdentityLeft 0<n record { x = x ; xLess = xLess } = equalityZn (modIsMod 0<n x (<NProp xLess))
|
||||
|
||||
plusZnCommutative : {n : ℕ} → .(pr : 0 <N n) → (a b : ℤn n pr) → _+n_ pr a b ≡ _+n_ pr b a
|
||||
plusZnCommutative {n} 0<n a b = equalityZn (applyEquality (mod n 0<n) (Semiring.commutative ℕSemiring (ℤn.x a) _))
|
||||
|
||||
plusZnIdentityLeft : {n : ℕ} → {pr : 0 <N n} → (a : ℤn n pr) → (record { x = 0 ; xLess = pr }) +n a ≡ a
|
||||
plusZnIdentityLeft {zero} {()}
|
||||
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } with totality x (succ n)
|
||||
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inl (inl x<succn) rewrite <NRefl x<succn xLess = refl
|
||||
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inl (inr succn<x) = exFalso (TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder succn<x xLess))
|
||||
plusZnIdentityLeft {succ n} {pr} record { x = x ; xLess = xLess } | inr x=succn rewrite x=succn = exFalso (TotalOrder.irreflexive ℕTotalOrder xLess)
|
||||
plusZnAssociative' : {n : ℕ} → .(pr : 0 <N n) → {a b c : ℕ} → (a <N n) → (b <N n) → (c <N n) → mod n pr ((mod n pr (a +N b)) +N c) ≡ mod n pr (a +N (mod n pr (b +N c)))
|
||||
plusZnAssociative' 0<n {a} {b} {c} a<n b<n c<n with modAddition 0<n a<n b<n
|
||||
plusZnAssociative' {n} 0<n {a} {b} {c} a<n b<n c<n | inl a+b=mod with modAddition 0<n b<n c<n
|
||||
... | inl b+c=mod rewrite equalityCommutative (a+b=mod) | equalityCommutative (b+c=mod) = applyEquality (mod _ 0<n) (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
... | inr b+c=n+mod rewrite equalityCommutative (a+b=mod) | equalityCommutative (modAnd+n 0<n (a +N mod _ 0<n (b +N c))) | Semiring.+Associative ℕSemiring n a (mod _ 0<n (b +N c)) | Semiring.commutative ℕSemiring n a | equalityCommutative (Semiring.+Associative ℕSemiring a n (mod _ 0<n (b +N c))) | equalityCommutative b+c=n+mod = applyEquality (mod _ 0<n) (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
plusZnAssociative' 0<n {a} {b} {c} a<n b<n c<n | inr a+b=n+mod with modAddition 0<n b<n c<n
|
||||
plusZnAssociative' {n} 0<n {a} {b} {c} a<n b<n c<n | inr a+b=n+mod | inl b+c=mod rewrite equalityCommutative b+c=mod | equalityCommutative (modAnd+n 0<n (mod _ 0<n (a +N b) +N c)) | Semiring.+Associative ℕSemiring n (mod _ 0<n (a +N b)) c | equalityCommutative a+b=n+mod = applyEquality (mod _ 0<n) (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
plusZnAssociative' {n} 0<n {a} {b} {c} a<n b<n c<n | inr a+b=n+mod | inr b+c=n+mod rewrite equalityCommutative (modAnd+n 0<n ((mod _ 0<n (a +N b)) +N c)) | equalityCommutative (modAnd+n 0<n (a +N mod _ 0<n (b +N c))) | Semiring.+Associative ℕSemiring n (mod _ 0<n (a +N b)) c | equalityCommutative (a+b=n+mod) | Semiring.commutative ℕSemiring a (mod n 0<n (b +N c)) | Semiring.+Associative ℕSemiring n (mod n 0<n (b +N c)) a | equalityCommutative b+c=n+mod | Semiring.commutative ℕSemiring (b +N c) a = applyEquality (mod _ 0<n) (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
|
||||
subLemma : {a b c : ℕ} → (pr1 : a <N b) → (pr2 : a <N c) → (pr : b ≡ c) → (subtractionNResult.result (-N (inl pr1))) ≡ (subtractionNResult.result (-N (inl pr2)))
|
||||
subLemma {a} {b} {c} a<b a<c b=c rewrite b=c | <NRefl a<b a<c = refl
|
||||
|
||||
plusZnCommutative : {n : ℕ} → {pr : 0 <N n} → (a b : ℤn n pr) → (a +n b) ≡ b +n a
|
||||
plusZnCommutative {zero} {()} a b
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } with totality (a +N b) (succ n)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) with totality (b +N a) (succ n)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inl (inl b+a<sn) = equalityZn _ _ (Semiring.commutative ℕSemiring a b)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inl (inr sn<b+a) = exFalso (f a+b<sn sn<b+a)
|
||||
plusZnAssociative : {n : ℕ} → .(pr : 0 <N n) → (a b c : ℤn n pr) → _+n_ pr (_+n_ pr a b) c ≡ _+n_ pr a (_+n_ pr b c)
|
||||
plusZnAssociative {succ n} 0<sn record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } = equalityZn t
|
||||
where
|
||||
f : (a +N b) <N succ n → succ n <N b +N a → False
|
||||
f pr1 pr2 rewrite Semiring.commutative ℕSemiring b a = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr2 pr1)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) | inr b+a=sn = exFalso (f a+b<sn b+a=sn)
|
||||
where
|
||||
f : (a +N b) <N succ n → b +N a ≡ succ n → False
|
||||
f pr1 eq rewrite Semiring.commutative ℕSemiring b a | eq = TotalOrder.irreflexive ℕTotalOrder pr1
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) with totality (b +N a) (succ n)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inl (inl b+a<sn) = exFalso (f sn<a+b b+a<sn)
|
||||
where
|
||||
f : succ n <N a +N b → b +N a <N succ n → False
|
||||
f pr1 pr2 rewrite Semiring.commutative ℕSemiring b a = TotalOrder.irreflexive ℕTotalOrder (<Transitive sn<a+b b+a<sn)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inl (inr sn<b+a) = equalityZn _ _ (subLemma sn<a+b sn<b+a (Semiring.commutative ℕSemiring a b))
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inr sn<a+b) | inr sn=b+a = exFalso (f sn<a+b sn=b+a)
|
||||
where
|
||||
f : succ n <N a +N b → b +N a ≡ succ n → False
|
||||
f pr eq rewrite Semiring.commutative ℕSemiring b a | eq = TotalOrder.irreflexive ℕTotalOrder pr
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b with totality (b +N a) (succ n)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inl (inl b+a<sn) = exFalso f
|
||||
where
|
||||
f : False
|
||||
f rewrite Semiring.commutative ℕSemiring b a | sn=a+b = TotalOrder.irreflexive ℕTotalOrder b+a<sn
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inl (inr sn<b+a) = exFalso f
|
||||
where
|
||||
f : False
|
||||
f rewrite Semiring.commutative ℕSemiring b a | sn=a+b = TotalOrder.irreflexive ℕTotalOrder sn<b+a
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inr sn=a+b | inr sn=b+a = equalityZn _ _ refl
|
||||
ma = mod (succ n) 0<sn a
|
||||
mb = mod (succ n) 0<sn b
|
||||
mc = mod (succ n) 0<sn c
|
||||
v : mod (succ n) 0<sn ((mod (succ n) 0<sn (ma +N mb)) +N mc) ≡ mod (succ n) 0<sn (ma +N mod (succ n) 0<sn (mb +N mc))
|
||||
v = plusZnAssociative' 0<sn (mod<N 0<sn a) (mod<N 0<sn b) (mod<N 0<sn c)
|
||||
u : mod (succ n) 0<sn ((mod (succ n) 0<sn (mod (succ n) 0<sn a +N (mod (succ n) 0<sn b))) +N c) ≡ mod (succ n) 0<sn (a +N mod (succ n) 0<sn ((mod (succ n) 0<sn b) +N (mod (succ n) 0<sn c)))
|
||||
u rewrite equalityCommutative (modIsMod 0<sn c (<NProp cLess)) | equalityCommutative (modIsMod 0<sn a (<NProp aLess)) | modIdempotent 0<sn a | modIdempotent 0<sn c = v
|
||||
t : mod (succ n) 0<sn (mod (succ n) 0<sn (a +N b) +N c) ≡ mod (succ n) 0<sn (a +N mod (succ n) 0<sn (b +N c))
|
||||
t rewrite modExtracts {succ n} 0<sn a b | modExtracts {succ n} 0<sn b c = u
|
||||
|
@@ -1,6 +1,4 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
-- These are explicitly with-K, because we currently encode an element of Zn as
|
||||
-- a natural together with a proof that it is small.
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
@@ -8,20 +6,22 @@ open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSetWithK
|
||||
open import Sets.FinSet.Definition
|
||||
open import Functions
|
||||
open import Numbers.Naturals.WithK
|
||||
open import Semirings.Definition
|
||||
open import Numbers.Naturals.EuclideanAlgorithm
|
||||
open import Orders.Total.Definition
|
||||
|
||||
module Numbers.Modulo.Definition where
|
||||
record ℤn (n : ℕ) (pr : 0 <N n) : Set where
|
||||
field
|
||||
x : ℕ
|
||||
xLess : x <N n
|
||||
|
||||
equalityZn : {n : ℕ} {pr : 0 <N n} → (a b : ℤn n pr) → (ℤn.x a ≡ ℤn.x b) → a ≡ b
|
||||
equalityZn record { x = a ; xLess = aLess } record { x = .a ; xLess = bLess } refl rewrite <NRefl aLess bLess = refl
|
||||
record ℤn (n : ℕ) .(pr : 0 <N n) : Set where
|
||||
field
|
||||
x : ℕ
|
||||
.xLess : x <N n
|
||||
|
||||
equalityZn : {n : ℕ} .{pr : 0 <N n} → {a b : ℤn n pr} → (ℤn.x a ≡ ℤn.x b) → a ≡ b
|
||||
equalityZn {a = record { x = a ; xLess = aLess }} {record { x = .a ; xLess = bLess }} refl = refl
|
||||
|
||||
|
@@ -1,6 +1,4 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
-- These are explicitly with-K, because we currently encode an element of Zn as
|
||||
-- a natural together with a proof that it is small.
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
@@ -11,616 +9,53 @@ open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSetWithK
|
||||
open import Sets.FinSet.Definition
|
||||
open import Sets.FinSet.Lemmas
|
||||
open import Sets.Cardinality.Finite.Definition
|
||||
open import Functions
|
||||
open import Numbers.Naturals.WithK
|
||||
open import Semirings.Definition
|
||||
open import Numbers.Modulo.Definition
|
||||
open import Numbers.Modulo.Addition
|
||||
open import Orders.Total.Definition
|
||||
open import Numbers.Naturals.EuclideanAlgorithm
|
||||
open import Numbers.Modulo.ModuloFunction
|
||||
|
||||
module Numbers.Modulo.Group where
|
||||
|
||||
open TotalOrder ℕTotalOrder
|
||||
open Semiring ℕSemiring
|
||||
|
||||
help30 : {a b c n : ℕ} → (c <N n) → (a +N b ≡ n) → (n<b+c : n <N b +N c) → (n <N a +N subtractionNResult.result (-N (inl n<b+c))) → False
|
||||
help30 {a} {b} {c} {n} c<n a+b=n n<b+c x = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr5 c<n)
|
||||
where
|
||||
pr : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
|
||||
pr = identityOfIndiscernablesRight _<N_ (additionPreservesInequality n x) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
pr2 : n +N n <N a +N (b +N c)
|
||||
pr2 = identityOfIndiscernablesRight _<N_ pr (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
pr3 : n +N n <N (a +N b) +N c
|
||||
pr3 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = pr2
|
||||
pr4 : n +N n <N c +N n
|
||||
pr4 rewrite Semiring.commutative ℕSemiring c n = identityOfIndiscernablesRight _<N_ pr3 (applyEquality (_+N c) a+b=n)
|
||||
pr5 : n <N c
|
||||
pr5 = subtractionPreservesInequality n pr4
|
||||
private
|
||||
0<s : {n : ℕ} → 0 <N succ n
|
||||
0<s {n} = le n (applyEquality succ (Semiring.sumZeroRight ℕSemiring n))
|
||||
|
||||
help31 : {a b c n : ℕ} → (a +N b ≡ n) → (n<b+c : n <N b +N c) → (a +N subtractionNResult.result (-N (inl n<b+c))) ≡ c
|
||||
help31 {a} {b} {c} {n} a+b=n n<b+c = canSubtractFromEqualityLeft pr2
|
||||
where
|
||||
pr1 : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N c
|
||||
pr1 rewrite addMinus' (inl n<b+c) | Semiring.+Associative ℕSemiring a b c | a+b=n = refl
|
||||
pr2 : n +N (a +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N c
|
||||
pr2 = identityOfIndiscernablesLeft _≡_ pr1 (lemm a n (subtractionNResult.result (-N (inl n<b+c))))
|
||||
where
|
||||
lemm : (a b c : ℕ) → a +N (b +N c) ≡ b +N (a +N c)
|
||||
lemm a b c rewrite Semiring.+Associative ℕSemiring a b c | Semiring.commutative ℕSemiring a b | equalityCommutative (Semiring.+Associative ℕSemiring b a c) = refl
|
||||
inverseN : {n : ℕ} → .(0<n : 0 <N n) → (x : ℤn n 0<n) → ℤn n 0<n
|
||||
inverseN 0<n record { x = 0 ; xLess = _ } = record { x = 0 ; xLess = 0<n }
|
||||
inverseN 0<n record { x = succ x ; xLess = xLess } with <NProp xLess
|
||||
... | le subtr pr = record { x = succ subtr ; xLess = le x (transitivity (commutative (succ x) (succ subtr)) pr) }
|
||||
|
||||
help7 : {a b c n : ℕ} → b +N c ≡ n → a <N n → (n<a+b : n <N a +N b) → (subtractionNResult.result (-N (inl n<a+b)) +N c ≡ n) → False
|
||||
help7 {a} {b} {c} {n} b+c=n a<n n<a+b x = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ a<n pr5)
|
||||
where
|
||||
pr : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c ≡ n +N n
|
||||
pr = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) x) (Semiring.+Associative ℕSemiring n _ c)
|
||||
pr2 : (a +N b) +N c ≡ n +N n
|
||||
pr2 = identityOfIndiscernablesLeft _≡_ pr (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
pr3 : a +N (b +N c) ≡ n +N n
|
||||
pr3 rewrite Semiring.+Associative ℕSemiring a b c = pr2
|
||||
pr4 : a +N n ≡ n +N n
|
||||
pr4 = identityOfIndiscernablesLeft _≡_ pr3 (applyEquality (a +N_) b+c=n)
|
||||
pr5 : a ≡ n
|
||||
pr5 = canSubtractFromEqualityRight pr4
|
||||
invLeft : {n : ℕ} → .(0<n : 0 <N n) → (x : ℤn n 0<n) → _+n_ 0<n (inverseN 0<n x) x ≡ record { x = 0 ; xLess = 0<n }
|
||||
invLeft {n} 0<n record { x = 0 ; xLess = xLess } = plusZnIdentityLeft 0<n (record { x = 0 ; xLess = 0<n })
|
||||
invLeft {n} 0<n record { x = (succ x) ; xLess = xLess } with <NProp xLess
|
||||
... | le subtr pr rewrite pr = equalityZn (modN 0<n)
|
||||
|
||||
help29 : {a b c n : ℕ} → (c <N n) → (n<b+c : n <N b +N c) → (a +N subtractionNResult.result (-N (inl n<b+c))) ≡ n → a +N b ≡ n → False
|
||||
help29 {a} {b} {c} {n} c<n n<b+c pr a+b=n = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ c<n p4)
|
||||
where
|
||||
p1 : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) ≡ n +N n
|
||||
p1 = identityOfIndiscernablesLeft _≡_ (applyEquality (_+N n) pr) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
p2 : (a +N b) +N c ≡ n +N n
|
||||
p2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = identityOfIndiscernablesLeft _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
p3 : n +N c ≡ n +N n
|
||||
p3 = identityOfIndiscernablesLeft _≡_ p2 (applyEquality (_+N c) a+b=n)
|
||||
p4 : c ≡ n
|
||||
p4 = canSubtractFromEqualityLeft p3
|
||||
|
||||
help28 : {a b c n : ℕ} → (n<a+'b+c : n <N a +N (b +N c)) → (a +N b ≡ n) → subtractionNResult.result (-N (inl n<a+'b+c)) ≡ c
|
||||
help28 {a} {b} {c} {n} pr a+b=n = canSubtractFromEqualityLeft p2
|
||||
where
|
||||
p1 : a +N (b +N c) ≡ n +N c
|
||||
p1 rewrite Semiring.+Associative ℕSemiring a b c | a+b=n = refl
|
||||
p2 : n +N subtractionNResult.result (-N (inl pr)) ≡ n +N c
|
||||
p2 = identityOfIndiscernablesLeft _≡_ p1 (equalityCommutative (addMinus' (inl pr)))
|
||||
|
||||
help27 : {a b c n : ℕ} → (a +N b ≡ succ n) → (a +N (b +N c) <N succ n) → a +N (b +N c) ≡ c
|
||||
help27 {a} {b} {zero} {n} a+b=sn a+b+c<sn rewrite Semiring.commutative ℕSemiring b 0 | a+b=sn = exFalso (TotalOrder.irreflexive ℕTotalOrder a+b+c<sn)
|
||||
help27 {a} {b} {succ c} {n} a+b=sn a+b+c<sn rewrite Semiring.+Associative ℕSemiring a b (succ c) | a+b=sn = exFalso (cannotAddAndEnlarge' a+b+c<sn)
|
||||
|
||||
help26 : {a b c n : ℕ} → (a +N b ≡ n) → (a +N (b +N c) ≡ n) → 0 ≡ c
|
||||
help26 {a} {b} {c} {n} a+b=n a+b+c=n rewrite Semiring.+Associative ℕSemiring a b c | a+b=n | Semiring.commutative ℕSemiring n c = equalityCommutative (canSubtractFromEqualityRight a+b+c=n)
|
||||
|
||||
help25 : {a b c n : ℕ} → (a +N b ≡ n) → (b +N c ≡ n) → (a +N 0 ≡ c)
|
||||
help25 {a} {b} {c} {n} a+b=n b+c=n rewrite Semiring.commutative ℕSemiring a 0 | Semiring.commutative ℕSemiring a b | equalityCommutative a+b=n = equalityCommutative (canSubtractFromEqualityLeft b+c=n)
|
||||
|
||||
help24 : {a n : ℕ} → (a <N n) → (n <N a +N 0) → False
|
||||
help24 {a} {n} a<n n<a+0 rewrite Semiring.commutative ℕSemiring a 0 = TotalOrder.irreflexive ℕTotalOrder (<Transitive a<n n<a+0)
|
||||
|
||||
help23 : {a n : ℕ} → (a <N n) → (a +N 0 ≡ n) → False
|
||||
help23 {a} {n} a<n a+0=n rewrite Semiring.commutative ℕSemiring a 0 | a+0=n = TotalOrder.irreflexive ℕTotalOrder a<n
|
||||
|
||||
help22 : {a b c n : ℕ} → (a +N b ≡ n) → (c ≡ n) → (n<b+c : n <N b +N c) → (n <N a +N subtractionNResult.result (-N (inl n<b+c))) → False
|
||||
help22 {a} {b} {c} {.c} a+b=n refl n<b+c pr = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _<N_ pr4 a+b=n)
|
||||
where
|
||||
pr1 : c +N c <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N c)
|
||||
pr1 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality c pr) (equalityCommutative (Semiring.+Associative ℕSemiring a _ c))
|
||||
pr2 : c +N c <N a +N (b +N c)
|
||||
pr2 = identityOfIndiscernablesRight _<N_ pr1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
pr3 : c +N c <N (a +N b) +N c
|
||||
pr3 = identityOfIndiscernablesRight _<N_ pr2 (Semiring.+Associative ℕSemiring a b c)
|
||||
pr4 : c <N a +N b
|
||||
pr4 = subtractionPreservesInequality c pr3
|
||||
|
||||
help21 : {a b c n : ℕ} → (a +N b ≡ n) → (b +N c ≡ n) → (c ≡ n) → (a <N n) → False
|
||||
help21 {a} {b} {c} {.c} a+b=n b+c=n refl a<n = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ a<n a=c)
|
||||
where
|
||||
b=0 : b ≡ 0
|
||||
a=c : a ≡ c
|
||||
a=c = identityOfIndiscernablesLeft _≡_ a+b=n lemm
|
||||
where
|
||||
lemm : a +N b ≡ a
|
||||
lemm rewrite b=0 | Semiring.commutative ℕSemiring a 0 = refl
|
||||
b=0' : (b c : ℕ) → (b +N c ≡ c) → b ≡ 0
|
||||
b=0' zero c b+c=c = refl
|
||||
b=0' (succ b) zero b+c=c = exFalso (naughtE (equalityCommutative b+c=c))
|
||||
b=0' (succ b) (succ c) b+c=c = b=0' (succ b) c bl2
|
||||
where
|
||||
bl2 : succ b +N c ≡ c
|
||||
bl2 rewrite Semiring.commutative ℕSemiring b c | Semiring.commutative ℕSemiring (succ c) b = succInjective b+c=c
|
||||
b=0 = b=0' b c b+c=n
|
||||
|
||||
help20 : {a b c n : ℕ} → (c ≡ n) → (a +N b ≡ n) → (n<b+c : n <N b +N c) → (a +N subtractionNResult.result (-N (inl n<b+c)) <N n) → False
|
||||
help20 {a} {b} {c} {n} c=n a+b=n n<b+c pr = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ pr5 c=n)
|
||||
where
|
||||
pr1 : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
|
||||
pr1 = identityOfIndiscernablesLeft _<N_ (additionPreservesInequality n pr) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
pr2 : a +N (b +N c) <N n +N n
|
||||
pr2 = identityOfIndiscernablesLeft _<N_ pr1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
pr3 : (a +N b) +N c <N n +N n
|
||||
pr3 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = pr2
|
||||
pr4 : c +N n <N n +N n
|
||||
pr4 = identityOfIndiscernablesLeft _<N_ pr3 (transitivity (applyEquality (_+N c) a+b=n) (Semiring.commutative ℕSemiring n c))
|
||||
pr5 : c <N n
|
||||
pr5 = subtractionPreservesInequality n pr4
|
||||
|
||||
help19 : {a b c n : ℕ} → (b+c<n : b +N c <N n) → (n<a+b : n <N a +N b) → (a <N n) → (subtractionNResult.result (-N (inl n<a+b)) +N c ≡ n) → False
|
||||
help19 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ r q')
|
||||
where
|
||||
p : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c ≡ n +N n
|
||||
p = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_ ) pr) (Semiring.+Associative ℕSemiring n _ c)
|
||||
q : (a +N b) +N c ≡ n +N n
|
||||
q = identityOfIndiscernablesLeft _≡_ p (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
q' : a +N (b +N c) ≡ n +N n
|
||||
q' = identityOfIndiscernablesLeft _≡_ q (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
r : a +N (b +N c) <N n +N n
|
||||
r = addStrongInequalities a<n b+c<n
|
||||
|
||||
help18 : {a b c n : ℕ} → (b+c<n : b +N c <N n) → (n<a+b : n <N a +N b) → (a <N n) → (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) → False
|
||||
help18 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive ℕTotalOrder (<Transitive p4 a<n)
|
||||
where
|
||||
p : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
p = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr) (Semiring.+Associative ℕSemiring n _ c)
|
||||
p' : n +N n <N (a +N b) +N c
|
||||
p' = identityOfIndiscernablesRight _<N_ p (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
p2 : n +N n <N a +N (b +N c)
|
||||
p2 = identityOfIndiscernablesRight _<N_ p' (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
p3 : n +N n <N a +N n
|
||||
p3 = <Transitive p2 (additionPreservesInequalityOnLeft a b+c<n)
|
||||
p4 : n <N a
|
||||
p4 = subtractionPreservesInequality n p3
|
||||
|
||||
help17 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (a +N subtractionNResult.result (-N (inl n<b+c)) <N n) → (subtractionNResult.result (-N (inl n<a+b)) +N c) ≡ n → False
|
||||
help17 {a} {b} {c} {n} n<b+c n<a+b p1 p2 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ pr1'' pr3)
|
||||
where
|
||||
pr1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
|
||||
pr1' = identityOfIndiscernablesLeft _<N_ (additionPreservesInequality n p1) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
pr1'' : a +N (b +N c) <N n +N n
|
||||
pr1'' = identityOfIndiscernablesLeft _<N_ pr1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
pr2' : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c ≡ n +N n
|
||||
pr2' = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) p2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
pr2'' : (a +N b) +N c ≡ n +N n
|
||||
pr2'' = identityOfIndiscernablesLeft _≡_ pr2' (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
pr3 : a +N (b +N c) ≡ n +N n
|
||||
pr3 = identityOfIndiscernablesLeft _≡_ pr2'' (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
|
||||
help16 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (a +N subtractionNResult.result (-N (inl n<b+c))) <N n → (pr : n <N subtractionNResult.result (-N (inl n<a+b)) +N c) → a +N subtractionNResult.result (-N (inl n<b+c)) ≡ subtractionNResult.result (-N (inl pr))
|
||||
help16 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = exFalso (TotalOrder.irreflexive ℕTotalOrder (<Transitive pr3 pr1''))
|
||||
where
|
||||
pr1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) <N n +N n
|
||||
pr1' = identityOfIndiscernablesLeft _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
pr1'' : a +N (b +N c) <N n +N n
|
||||
pr1'' = identityOfIndiscernablesLeft _<N_ pr1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
pr2' : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
pr2' = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
pr2'' : n +N n <N (a +N b) +N c
|
||||
pr2'' = identityOfIndiscernablesRight _<N_ pr2' (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
pr3 : n +N n <N a +N (b +N c)
|
||||
pr3 = identityOfIndiscernablesRight _<N_ pr2'' (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
|
||||
help15 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (n <N a +N subtractionNResult.result (-N (inl n<b+c))) → (subtractionNResult.result (-N (inl n<a+b)) +N c) <N n → False
|
||||
help15 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive ℕTotalOrder (<Transitive p2'' p1')
|
||||
where
|
||||
p1 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c <N n +N n
|
||||
p1 = identityOfIndiscernablesLeft _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
p1' : (a +N b) +N c <N n +N n
|
||||
p1' = identityOfIndiscernablesLeft _<N_ p1 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
p2 : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
|
||||
p2 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
p2' : n +N n <N a +N (b +N c)
|
||||
p2' = identityOfIndiscernablesRight _<N_ p2 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
p2'' : n +N n <N (a +N b) +N c
|
||||
p2'' = identityOfIndiscernablesRight _<N_ p2' (Semiring.+Associative ℕSemiring a b c)
|
||||
|
||||
help14 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (pr1 : n <N a +N subtractionNResult.result (-N (inl n<b+c))) → (pr2 : n <N subtractionNResult.result (-N (inl n<a+b)) +N c) → subtractionNResult.result (-N (inl pr1)) ≡ subtractionNResult.result (-N (inl pr2))
|
||||
help14 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = equivalentSubtraction _ _ _ _ pr1 pr2 (transitivity (Semiring.+Associative ℕSemiring n _ c) (transitivity (applyEquality (_+N c) (addMinus' (inl n<a+b))) (transitivity (equalityCommutative (Semiring.+Associative ℕSemiring a b c)) (equalityCommutative p2))))
|
||||
where
|
||||
p1 : (a +N subtractionNResult.result (-N (inl n<b+c))) +N n ≡ a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
|
||||
p1 = equalityCommutative (Semiring.+Associative ℕSemiring a _ n)
|
||||
p2 : (a +N subtractionNResult.result (-N (inl n<b+c))) +N n ≡ a +N (b +N c)
|
||||
p2 = identityOfIndiscernablesRight _≡_ p1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
|
||||
help13 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (n <N a +N subtractionNResult.result (-N (inl n<b+c))) → (subtractionNResult.result (-N (inl n<a+b)) +N c ≡ n) → False
|
||||
help13 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _<N_ lemm1' lemm3)
|
||||
where
|
||||
lemm1 : n +N n <N a +N (subtractionNResult.result (-N (inl n<b+c)) +N n)
|
||||
lemm1 = identityOfIndiscernablesRight _<N_ (additionPreservesInequality n pr1) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
lemm1' : n +N n <N a +N (b +N c)
|
||||
lemm1' = identityOfIndiscernablesRight _<N_ lemm1 (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
lemm2 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c ≡ n +N n
|
||||
lemm2 = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) pr2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
lemm2' : (a +N b) +N c ≡ n +N n
|
||||
lemm2' = identityOfIndiscernablesLeft _≡_ lemm2 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
lemm3 : a +N (b +N c) ≡ n +N n
|
||||
lemm3 rewrite Semiring.+Associative ℕSemiring a b c = lemm2'
|
||||
|
||||
help12 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (a +N subtractionNResult.result (-N (inl n<b+c))) ≡ n → subtractionNResult.result (-N (inl n<a+b)) +N c <N n → False
|
||||
help12 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive ℕTotalOrder lemm4
|
||||
where
|
||||
pr : {a b c : ℕ} → a +N (b +N c) ≡ b +N (a +N c)
|
||||
pr {a} {b} {c} rewrite Semiring.+Associative ℕSemiring a b c | Semiring.commutative ℕSemiring a b | equalityCommutative (Semiring.+Associative ℕSemiring b a c) = refl
|
||||
lemm1 : (n +N subtractionNResult.result (-N (inl n<a+b))) +N c <N n +N n
|
||||
lemm1 = identityOfIndiscernablesLeft _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
lemm2 : (a +N b) +N c <N n +N n
|
||||
lemm2 = identityOfIndiscernablesLeft _<N_ lemm1 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
lemm1' : a +N (subtractionNResult.result (-N (inl n<b+c)) +N n) ≡ n +N n
|
||||
lemm1' = identityOfIndiscernablesLeft _≡_ (applyEquality (_+N n) pr1) (equalityCommutative (Semiring.+Associative ℕSemiring a _ n))
|
||||
lemm2' : a +N (b +N c) ≡ n +N n
|
||||
lemm2' = identityOfIndiscernablesLeft _≡_ lemm1' (applyEquality (a +N_) (addMinus (inl n<b+c)))
|
||||
lemm3 : (a +N b) +N c ≡ n +N n
|
||||
lemm3 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = lemm2'
|
||||
lemm4 : (a +N b) +N c <N (a +N b) +N c
|
||||
lemm4 = identityOfIndiscernablesRight _<N_ lemm2 (equalityCommutative lemm3)
|
||||
|
||||
help11 : {a b c n : ℕ} → (a <N n) → (b +N c ≡ n) → (n<a+b : n <N a +N b) → (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) → False
|
||||
help11 {a} {b} {c} {n} a<n b+c=n n<a+b pr1 = TotalOrder.irreflexive ℕTotalOrder (<Transitive a<n lemm5)
|
||||
where
|
||||
pr : {a b c : ℕ} → a +N (b +N c) ≡ b +N (a +N c)
|
||||
pr {a} {b} {c} rewrite Semiring.+Associative ℕSemiring a b c | Semiring.commutative ℕSemiring a b | equalityCommutative (Semiring.+Associative ℕSemiring b a c) = refl
|
||||
lemm : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
lemm = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr1) (Semiring.+Associative ℕSemiring n _ c)
|
||||
lemm2 : n +N n <N (a +N b) +N c
|
||||
lemm2 = identityOfIndiscernablesRight _<N_ lemm (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
lemm3 : n +N n <N a +N (b +N c)
|
||||
lemm3 = identityOfIndiscernablesRight _<N_ lemm2 (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
lemm4 : n +N n <N a +N n
|
||||
lemm4 = identityOfIndiscernablesRight _<N_ lemm3 (applyEquality (a +N_) b+c=n)
|
||||
lemm5 : n <N a
|
||||
lemm5 = subtractionPreservesInequality n lemm4
|
||||
|
||||
help10 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → (a +N subtractionNResult.result (-N (inl n<b+c)) ≡ n) → (n <N subtractionNResult.result (-N (inl n<a+b)) +N c) → False
|
||||
help10 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive ℕTotalOrder lemm6
|
||||
where
|
||||
pr : {a b c : ℕ} → a +N (b +N c) ≡ b +N (a +N c)
|
||||
pr {a} {b} {c} rewrite Semiring.+Associative ℕSemiring a b c | Semiring.commutative ℕSemiring a b | equalityCommutative (Semiring.+Associative ℕSemiring b a c) = refl
|
||||
lemm : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N n
|
||||
lemm = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) pr1) (pr {n} {a})
|
||||
lemm2 : a +N (b +N c) ≡ n +N n
|
||||
lemm2 = identityOfIndiscernablesLeft _≡_ lemm (applyEquality (a +N_) (addMinus' (inl n<b+c)))
|
||||
lemm3 : n +N n <N (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
lemm3 = identityOfIndiscernablesRight _<N_ (additionPreservesInequalityOnLeft n pr2) (Semiring.+Associative ℕSemiring n _ c)
|
||||
lemm4 : n +N n <N (a +N b) +N c
|
||||
lemm4 = identityOfIndiscernablesRight _<N_ lemm3 (applyEquality (_+N c) (addMinus' (inl n<a+b)))
|
||||
lemm5 : n +N n <N a +N (b +N c)
|
||||
lemm5 = identityOfIndiscernablesRight _<N_ lemm4 (equalityCommutative (Semiring.+Associative ℕSemiring a b c))
|
||||
lemm6 : a +N (b +N c) <N a +N (b +N c)
|
||||
lemm6 = identityOfIndiscernablesLeft _<N_ lemm5 (equalityCommutative lemm2)
|
||||
|
||||
help9 : {a n : ℕ} → (a +N 0 ≡ n) → (a <N n) → False
|
||||
help9 {a} {n} n=a+0 a<n rewrite Semiring.commutative ℕSemiring a 0 | n=a+0 = TotalOrder.irreflexive ℕTotalOrder a<n
|
||||
|
||||
help8 : {a n : ℕ} → (n <N a +N 0) → (a <N n) → False
|
||||
help8 {a} {n} n<a+0 a<n rewrite Semiring.commutative ℕSemiring a 0 = TotalOrder.irreflexive ℕTotalOrder (<Transitive a<n n<a+0)
|
||||
|
||||
help6 : {a b c n : ℕ} → (b +N c ≡ n) → (n<a+b : n <N a +N b) → (a +N 0 ≡ subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
help6 {a} {b} {c} {n} b+c=n n<a+b rewrite Semiring.commutative ℕSemiring a 0 = canSubtractFromEqualityLeft {n} lem'
|
||||
where
|
||||
lem : n +N a ≡ (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
lem rewrite addMinus' (inl n<a+b) | equalityCommutative (Semiring.+Associative ℕSemiring a b c) | b+c=n = Semiring.commutative ℕSemiring n a
|
||||
lem' : n +N a ≡ n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
lem' = identityOfIndiscernablesRight _≡_ lem (equalityCommutative (Semiring.+Associative ℕSemiring n _ c))
|
||||
|
||||
help5 : {a b c n : ℕ} → (n<b+c : n <N b +N c) → (n<a+b : n <N a +N b) → a +N subtractionNResult.result (-N (inl n<b+c)) ≡ subtractionNResult.result (-N (inl n<a+b)) +N c
|
||||
help5 {a} {b} {c} {n} n<b+c n<a+b = canSubtractFromEqualityLeft {n} lemma''
|
||||
where
|
||||
lemma : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) ≡ (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
lemma rewrite addMinus' (inl n<b+c) | addMinus' (inl n<a+b) = Semiring.+Associative ℕSemiring a b c
|
||||
lemma' : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
lemma' = identityOfIndiscernablesRight _≡_ lemma (equalityCommutative (Semiring.+Associative ℕSemiring n _ c))
|
||||
lemma'' : n +N (a +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
lemma'' = identityOfIndiscernablesLeft _≡_ lemma' (pr {a} {n} {subtractionNResult.result (-N (inl n<b+c))})
|
||||
where
|
||||
pr : {a b c : ℕ} → a +N (b +N c) ≡ b +N (a +N c)
|
||||
pr {a} {b} {c} rewrite Semiring.+Associative ℕSemiring a b c | Semiring.commutative ℕSemiring a b | equalityCommutative (Semiring.+Associative ℕSemiring b a c) = refl
|
||||
|
||||
help4 : {a b c n : ℕ} → (n<a+'b+c : n <N a +N (b +N c)) → (n<a+b : n <N a +N b) → (subtractionNResult.result (-N (inl n<a+'b+c)) ≡ subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
help4 {a} {b} {c} {n} n<a+'b+c n<a+b = canSubtractFromEqualityLeft lemma'
|
||||
where
|
||||
lemma : (n +N subtractionNResult.result (-N (inl n<a+'b+c))) ≡ (n +N subtractionNResult.result (-N (inl n<a+b))) +N c
|
||||
lemma rewrite addMinus' (inl n<a+'b+c) | addMinus' (inl n<a+b) = Semiring.+Associative ℕSemiring a b c
|
||||
lemma' : n +N subtractionNResult.result (-N (inl n<a+'b+c)) ≡ n +N (subtractionNResult.result (-N (inl n<a+b)) +N c)
|
||||
lemma' = identityOfIndiscernablesRight _≡_ lemma (equalityCommutative (Semiring.+Associative ℕSemiring n (subtractionNResult.result (-N (inl n<a+b))) c))
|
||||
|
||||
help3 : {a b c n : ℕ} → (a <N n) → (b <N n) → (c <N n) → (a +N b <N n) → (pr : n <N b +N c) → a +N subtractionNResult.result (-N (inl pr)) ≡ n → False
|
||||
help3 {a} {b} {c} {n} a<n b<n c<n a+b<n n<b+c pr = TotalOrder.irreflexive ℕTotalOrder (<Transitive (inter4 inter3) c<n)
|
||||
where
|
||||
inter : a +N (n +N subtractionNResult.result (-N (inl n<b+c))) ≡ n +N n
|
||||
inter = identityOfIndiscernablesLeft _≡_ (applyEquality (n +N_) pr) (lemma n a (subtractionNResult.result (-N (inl n<b+c))))
|
||||
where
|
||||
lemma : (a b c : ℕ) → a +N (b +N c) ≡ b +N (a +N c)
|
||||
lemma a b c rewrite Semiring.+Associative ℕSemiring a b c | Semiring.+Associative ℕSemiring b a c = applyEquality (_+N c) (Semiring.commutative ℕSemiring a b)
|
||||
inter2 : n +N n ≡ a +N (b +N c)
|
||||
inter2 = equalityCommutative (identityOfIndiscernablesLeft _≡_ inter (applyEquality (a +N_) (addMinus' (inl n<b+c))))
|
||||
inter3 : n +N n <N n +N c
|
||||
inter3 rewrite inter2 | Semiring.+Associative ℕSemiring a b c = additionPreservesInequality c a+b<n
|
||||
inter4 : (n +N n <N n +N c) → n <N c
|
||||
inter4 pr rewrite Semiring.commutative ℕSemiring n c = subtractionPreservesInequality n pr
|
||||
|
||||
help2 : {a b c n : ℕ} → (sn<b+c : succ n <N b +N c) → (sn<a+b+c : succ n <N (a +N b) +N c) → a +N subtractionNResult.result (-N (inl sn<b+c)) ≡ subtractionNResult.result (-N (inl sn<a+b+c))
|
||||
help2 {a} {b} {c} {n} sn<b+c sn<a+b+c = res inter
|
||||
where
|
||||
inter : a +N (subtractionNResult.result (-N (inl sn<b+c)) +N succ n) ≡ subtractionNResult.result (-N (inl sn<a+b+c)) +N succ n
|
||||
inter rewrite addMinus (inl sn<b+c) | addMinus (inl sn<a+b+c) = Semiring.+Associative ℕSemiring a b c
|
||||
res : (a +N (subtractionNResult.result (-N (inl sn<b+c)) +N succ n) ≡ subtractionNResult.result (-N (inl sn<a+b+c)) +N succ n) → a +N subtractionNResult.result (-N (inl sn<b+c)) ≡ subtractionNResult.result (-N (inl sn<a+b+c))
|
||||
res pr = canSubtractFromEqualityRight {_} {succ n} (identityOfIndiscernablesLeft _≡_ pr (Semiring.+Associative ℕSemiring a (subtractionNResult.result (-N (inl sn<b+c))) (succ n)))
|
||||
|
||||
help1 : {a b c n : ℕ} → (sn<b+c : succ n <N b +N c) → (pr1 : succ n <N a +N subtractionNResult.result (-N (inl sn<b+c))) → (a +N b <N succ n) → (a <N succ n) → (b <N succ n) → (c <N succ n) → False
|
||||
help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn with -N (inl sn<b+c)
|
||||
help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn | record { result = b+c-sn ; pr = Prb+c-sn } = ans
|
||||
where
|
||||
b+c-nNonzero : b+c-sn ≡ 0 → False
|
||||
b+c-nNonzero pr rewrite (equalityCommutative Prb+c-sn) | pr | Semiring.commutative ℕSemiring n 0 = TotalOrder.irreflexive ℕTotalOrder sn<b+c
|
||||
2sn<a+b+c' : succ n +N succ n <N succ n +N (a +N b+c-sn)
|
||||
2sn<a+b+c' = additionPreservesInequalityOnLeft (succ n) pr1
|
||||
2sn<a+b+c'' : succ n +N succ n <N a +N (succ n +N b+c-sn)
|
||||
2sn<a+b+c'' rewrite Semiring.+Associative ℕSemiring a (succ n) b+c-sn | Semiring.commutative ℕSemiring a (succ n) | equalityCommutative (Semiring.+Associative ℕSemiring (succ n) a b+c-sn) = 2sn<a+b+c'
|
||||
eep : succ n +N succ n <N a +N (b +N c)
|
||||
eep rewrite equalityCommutative Prb+c-sn = 2sn<a+b+c''
|
||||
eep2 : a +N (b +N c) <N succ n +N c
|
||||
eep2 rewrite Semiring.+Associative ℕSemiring a b c = additionPreservesInequality c a+b<sn
|
||||
eep2' : a +N (b +N c) <N succ n +N succ n
|
||||
eep2' = <Transitive eep2 (additionPreservesInequalityOnLeft (succ n) c<sn)
|
||||
ans : False
|
||||
ans = TotalOrder.irreflexive ℕTotalOrder (<Transitive eep eep2')
|
||||
|
||||
plusZnAssociative : {n : ℕ} → {pr : 0 <N n} → (a b c : ℤn n pr) → a +n (b +n c) ≡ ((a +n b) +n c)
|
||||
plusZnAssociative {zero} {()}
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess} record { x = c ; xLess = cLess } with totality (a +N b) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) with totality ((a +N b) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = equalityZn _ _ (Semiring.+Associative ℕSemiring a b c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = exFalso (false {succ n} a+b+c<sn sn<a+'b+c)
|
||||
where
|
||||
false : {x : ℕ} → (a +N b) +N c <N succ n → succ n <N a +N (b +N c) → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr1 pr2)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inl b+c<sn) | inr sn=a+b+c = exFalso (false a+b+c<sn sn=a+b+c)
|
||||
where
|
||||
false : {x : ℕ} → (a +N b) +N c <N x → (a +N (b +N c)) ≡ x → False
|
||||
false p1 p2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | p2 = TotalOrder.irreflexive ℕTotalOrder p1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) with totality (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inl (inl x) with -N (inl sn<b+c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inl (inl x) | record { result = result ; pr = pr } = exFalso (false a+'b+c<sn pr)
|
||||
where
|
||||
false : a +N (b +N c) <N succ n → succ n +N result ≡ b +N c → False
|
||||
false pr1 pr2 rewrite equalityCommutative pr2 | Semiring.+Associative ℕSemiring a (succ n) result | Semiring.commutative ℕSemiring a (succ n) | equalityCommutative (Semiring.+Associative ℕSemiring (succ n) a result) = cannotAddAndEnlarge' pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inl (inr x) with -N (inl sn<b+c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inl (inr x) | record { result = result ; pr = pr } = exFalso false
|
||||
where
|
||||
lemma : a +N (succ n +N result) ≡ a +N (b +N c)
|
||||
lemma' : a +N (succ n +N result) <N succ n
|
||||
lemma'' : succ n +N (a +N result) <N succ n
|
||||
lemma'' = identityOfIndiscernablesLeft _<N_ lemma' (transitivity (Semiring.+Associative ℕSemiring a (succ n) result) (transitivity (applyEquality (λ t → t +N result) (Semiring.commutative ℕSemiring a (succ n))) (equalityCommutative (Semiring.+Associative ℕSemiring (succ n) a result))))
|
||||
lemma = applyEquality (λ i → a +N i) pr
|
||||
lemma' rewrite lemma = a+'b+c<sn
|
||||
false : False
|
||||
false = cannotAddAndEnlarge' lemma''
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl _) | inr x with -N (inl sn<b+c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inl a+'b+c<sn) | inr x | record { result = result ; pr = pr } = exFalso false
|
||||
where
|
||||
lemma : a +N (succ n +N result) <N succ n
|
||||
lemma rewrite pr = a+'b+c<sn
|
||||
lemma' : succ n +N (a +N result) <N succ n
|
||||
lemma' = identityOfIndiscernablesLeft _<N_ lemma (transitivity (Semiring.+Associative ℕSemiring a (succ n) result) (transitivity (applyEquality (λ t → t +N result) (Semiring.commutative ℕSemiring a (succ n))) (equalityCommutative (Semiring.+Associative ℕSemiring (succ n) a result))))
|
||||
false : False
|
||||
false = cannotAddAndEnlarge' lemma'
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inl (inr x) = equalityZn _ _ (exFalso (false {succ n} a+b+c<sn x))
|
||||
where
|
||||
false : {x : ℕ} → (a +N b) +N c <N succ n → succ n <N a +N (b +N c) → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr1 pr2)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inl (inr sn<b+c) | inr x = exFalso (false a+b+c<sn x)
|
||||
where
|
||||
false : (a +N b) +N c <N succ n → a +N (b +N c) ≡ succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative pr2 | equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inl a+b+c<sn) | inr sn=b+c = exFalso (false a+b+c<sn sn=b+c)
|
||||
where
|
||||
false : (a +N b) +N c <N succ n → b +N c ≡ succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 | Semiring.commutative ℕSemiring a (succ n) = cannotAddAndEnlarge' a+b+c<sn
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = exFalso (false sn<a+b+c a+'b+c<sn)
|
||||
where
|
||||
false : (succ n <N (a +N b) +N c) → a +N (b +N c) <N succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr1 pr2)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = equalityZn _ _ ans
|
||||
where
|
||||
lemma : succ n +N ((a +N b) +N c) ≡ (a +N (b +N c)) +N succ n
|
||||
lemma rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = Semiring.commutative ℕSemiring (succ n) _
|
||||
ans : subtractionNResult.result (-N (inl sn<a+'b+c)) ≡ subtractionNResult.result (-N (inl sn<a+b+c))
|
||||
ans = equivalentSubtraction _ _ _ _ sn<a+'b+c sn<a+b+c lemma
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inl b+c<sn) | inr x = exFalso (false sn<a+b+c x)
|
||||
where
|
||||
false : succ n <N (a +N b) +N c → (a +N (b +N c)) ≡ succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 = TotalOrder.irreflexive ℕTotalOrder sn<a+b+c
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inl a+b+c<sn) = exFalso (false sn<a+b+c a+b+c<sn)
|
||||
where
|
||||
false : succ n <N (a +N b) +N c → (a +N (b +N c)) <N succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr1 pr2)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) with totality (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) | inl (inl x) = equalityZn _ _ (help2 {a} {b} {c} {n} sn<b+c sn<a+b+c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) | inl (inr x) = equalityZn _ _ (exFalso (help1 sn<b+c x a+b<sn aLess bLess cLess))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inl (inr _) | inr x = exFalso (help3 aLess bLess cLess a+b<sn sn<b+c x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inl (inr sn<b+c) | inr a+b+c=sn = exFalso (false sn<a+b+c a+b+c=sn)
|
||||
where
|
||||
false : (succ n <N (a +N b) +N c) → (a +N (b +N c) ≡ succ n) → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 = TotalOrder.irreflexive ℕTotalOrder pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inl x) = exFalso (false sn<a+b+c x)
|
||||
where
|
||||
false : succ n <N (a +N b) +N c → a +N (b +N c) <N succ n → False
|
||||
false p1 p2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) = TotalOrder.irreflexive ℕTotalOrder (<Transitive p1 p2)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) with totality (a +N 0) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inl (inl x) = equalityZn _ _ (ans sn=b+c)
|
||||
where
|
||||
ans : b +N c ≡ succ n → a +N 0 ≡ subtractionNResult.result (-N (inl sn<a+b+c))
|
||||
ans pr with -N (inl sn<a+b+c)
|
||||
ans b+c=sn | record { result = result ; pr = pr1 } rewrite Semiring.commutative ℕSemiring a 0 | equalityCommutative (Semiring.+Associative ℕSemiring a b c) | b+c=sn | Semiring.commutative ℕSemiring (succ n) result = equalityCommutative (canSubtractFromEqualityRight pr1)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inl (inr x) = exFalso (false b a+b<sn x)
|
||||
where
|
||||
false : (b : ℕ) → a +N b <N succ n → succ n <N a +N 0 → False
|
||||
false zero pr1 pr2 rewrite Semiring.commutative ℕSemiring a 0 = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr1 pr2)
|
||||
false (succ b) pr1 pr2 rewrite Semiring.commutative ℕSemiring a 0 = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr2 (<Transitive (le b (Semiring.commutative ℕSemiring (succ b) a)) pr1))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inl (inr _) | inr x = exFalso (false b a+b<sn x)
|
||||
where
|
||||
false : (b : ℕ) → a +N b <N succ n → a +N 0 ≡ succ n → False
|
||||
false zero pr1 pr2 rewrite pr2 = TotalOrder.irreflexive ℕTotalOrder pr1
|
||||
false (succ b) pr1 pr2 rewrite Semiring.commutative ℕSemiring a 0 | pr2 = cannotAddAndEnlarge' pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inl (inr sn<a+b+c) | inr sn=b+c | inr x = exFalso (false sn<a+b+c x)
|
||||
where
|
||||
false : succ n <N (a +N b) +N c → a +N (b +N c) ≡ succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 = TotalOrder.irreflexive ℕTotalOrder pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = exFalso (false sn=a+b+c a+'b+c<sn)
|
||||
where
|
||||
false : (a +N b) +N c ≡ succ n → a +N (b +N c) <N succ n → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr1 = TotalOrder.irreflexive ℕTotalOrder pr2
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = exFalso (false sn=a+b+c sn<a+'b+c)
|
||||
where
|
||||
false : (a +N b) +N c ≡ succ n → succ n <N a +N (b +N c) → False
|
||||
false pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr1 = TotalOrder.irreflexive ℕTotalOrder pr2
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inl b+c<sn) | inr _ = equalityZn _ _ refl
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inl (inr sn<b+c) = exFalso (false a sn=a+b+c sn<b+c)
|
||||
where
|
||||
false : (a : ℕ) → (a +N b) +N c ≡ succ n → succ n <N b +N c → False
|
||||
false zero pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | equalityCommutative pr1 = TotalOrder.irreflexive ℕTotalOrder pr2
|
||||
false (succ a) pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | equalityCommutative pr1 = TotalOrder.irreflexive ℕTotalOrder (<Transitive pr2 (le a refl))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn with totality (a +N 0) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inl (inl a+0<sn) = equalityZn _ _ ans
|
||||
where
|
||||
a=0 : (a : ℕ) → (a +N b) +N c ≡ succ n → b +N c ≡ succ n → a ≡ 0
|
||||
a=0 zero pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 = refl
|
||||
a=0 (succ a) pr1 pr2 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 = canSubtractFromEqualityRight pr1
|
||||
ans : a +N 0 ≡ 0
|
||||
ans rewrite Semiring.commutative ℕSemiring a 0 = a=0 a sn=a+b+c b+c=sn
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inl (inr sn<a+0) = exFalso (false sn<a+0 sn=a+b+c b+c=sn)
|
||||
where
|
||||
false : succ n <N a +N 0 → (a +N b) +N c ≡ succ n → b +N c ≡ succ n → False
|
||||
false pr1 pr2 pr3 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr3 | Semiring.commutative ℕSemiring a 0 = zeroNeverGreater {succ n} (identityOfIndiscernablesRight _<N_ pr1 (a=0 a pr2))
|
||||
where
|
||||
a=0 : (a : ℕ) → (a +N succ n ≡ succ n) → a ≡ 0
|
||||
a=0 zero pr = refl
|
||||
a=0 (succ a) pr = canSubtractFromEqualityRight pr
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inl a+b<sn) | inr sn=a+b+c | inr b+c=sn | inr sn=a+0 = exFalso (false sn=a+b+c b+c=sn sn=a+0)
|
||||
where
|
||||
false : (a +N b) +N c ≡ succ n → b +N c ≡ succ n → a +N 0 ≡ succ n → False
|
||||
false pr1 pr2 pr3 rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | pr2 | equalityCommutative pr3 | Semiring.commutative ℕSemiring a 0 = naughtE (identityOfIndiscernablesLeft _≡_ pr3 (a=0 a pr1))
|
||||
where
|
||||
a=0 : (a : ℕ) → (a +N a ≡ a) → a ≡ 0
|
||||
a=0 zero pr = refl
|
||||
a=0 (succ a) pr = exFalso (naughtE {a} (equalityCommutative (canSubtractFromEqualityRight pr)))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = exFalso (false sn<a+b a+'b+c<sn)
|
||||
where
|
||||
false : succ n <N a +N b → a +N (b +N c) <N succ n → False
|
||||
false pr1 pr2 rewrite Semiring.+Associative ℕSemiring a b c = cannotAddAndEnlarge' (<Transitive pr2 pr1)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) with totality (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) | inl (inl x) = equalityZn _ _ (help4 {a} {b} {c} {succ n} sn<a+'b+c sn<a+b)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) | inl (inr x) = exFalso (help18 {a} {b} {c} {succ n} b+c<sn sn<a+b aLess x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) | inr x = exFalso (help19 {a} {b} {c} {succ n} b+c<sn sn<a+b aLess x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inl b+c<sn) | inr a+'b+c=sn = exFalso (false sn<a+b a+'b+c=sn)
|
||||
where
|
||||
false : (succ n <N a +N b) → a +N (b +N c) ≡ succ n → False
|
||||
false pr1 pr2 rewrite Semiring.+Associative ℕSemiring a b c | equalityCommutative pr2 = cannotAddAndEnlarge' pr1
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) with totality (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inl x) with totality (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inl x) | inl (inl x₁) = equalityZn _ _ (help5 {a} {b} {c} {succ n} sn<b+c sn<a+b)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inl x) | inl (inr x1) = equalityZn _ _ (help16 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inl x) | inr x1 = exFalso (help17 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inr x) with totality (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inr x) | inl (inl x1) = equalityZn _ _ (exFalso (help15 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inr x) | inl (inr x1) = equalityZn _ _ (help14 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inl (inr x) | inr x1 = equalityZn _ _ (exFalso (help13 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inr x with totality (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inr x | inl (inl x1) = equalityZn _ _ (exFalso (help12 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inr x | inl (inr x1) = equalityZn _ _ (exFalso (help10 {a} {b} {c} {succ n} sn<b+c sn<a+b x x1))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inl (inr sn<b+c) | inr x | inr x₁ = equalityZn _ _ refl
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn with totality (a +N 0) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inl (inl a+0<sn) with totality (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inl (inl a+0<sn) | inl (inl x) = equalityZn _ _ (help6 {a} {b} {c} {succ n} b+c=sn sn<a+b)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inl (inl _) | inl (inr x) = equalityZn _ _ (exFalso (help11 {a} {b} {c} {succ n} aLess b+c=sn sn<a+b x))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inl (inl a+0<sn) | inr x = equalityZn _ _ (exFalso (help7 {a} {b} {c} {succ n} b+c=sn aLess sn<a+b x))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inl (inr sn<a+0) = exFalso (help8 {a} {succ n} sn<a+0 aLess)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inl (inr sn<a+b) | inr b+c=sn | inr a+0=sn = exFalso (help9 {a} {succ n} a+0=sn aLess)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b with totality c (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inl b+c<sn) with totality (a +N (b +N c)) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inl b+c<sn) | inl (inl a+'b+c<sn) = equalityZn _ _ (help27 {a} {b} {c} {n} sn=a+b a+'b+c<sn)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inl b+c<sn) | inl (inr sn<a+'b+c) = equalityZn _ _ (help28 {a} {b} {c} {succ n} sn<a+'b+c sn=a+b)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inl b+c<sn) | inr a+'b+c=sn = equalityZn _ _ (help26 {a} {b} {c} {succ n} sn=a+b a+'b+c=sn)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inr sn<b+c) with totality (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inr sn<b+c) | inl (inl x) = equalityZn _ _ (help31 {a} {b} {c} {succ n} sn=a+b sn<b+c)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inr sn<b+c) | inl (inr x) = exFalso (help30 {a} {b} {c} {succ n} cLess sn=a+b sn<b+c x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inl (inr sn<b+c) | inr x = exFalso (help29 {a} {b} {c} {succ n} cLess sn<b+c x sn=a+b)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn with totality (a +N 0) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn | inl (inl a+0<sn) = equalityZn _ _ (help25 {a} {b} {c} {succ n} sn=a+b b+c=sn)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn | inl (inr sn<a+0) = exFalso (help24 {a} {succ n} aLess sn<a+0)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inl c<sn) | inr b+c=sn | inr a+0=sn = exFalso (help23 {a} {succ n} aLess a+0=sn)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inl (inr sn<c) = exFalso (TotalOrder.irreflexive ℕTotalOrder (<Transitive sn<c cLess))
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn with totality (b +N c) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inl b+c<sn) = exFalso (help b+c<sn)
|
||||
where
|
||||
help : (b +N c <N succ n) → False
|
||||
help b+c<sn rewrite equalityCommutative c=sn | Semiring.commutative ℕSemiring b c = cannotAddAndEnlarge' b+c<sn
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) with totality (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) | inl (inl x) = exFalso (help20 {a} {b} {c} {succ n} c=sn sn=a+b sn<b+c x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) | inl (inr x) = exFalso (help22 {a} {b} {c} {succ n} sn=a+b c=sn sn<b+c x)
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inl (inr sn<b+c) | inr x = equalityZn _ _ refl
|
||||
plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } record { x = c ; xLess = cLess } | inr sn=a+b | inr c=sn | inr b+c=sn = exFalso (help21 {a} {b} {c} {succ n} sn=a+b b+c=sn c=sn aLess)
|
||||
|
||||
subLess : {a b : ℕ} → (0<a : 0 <N a) → (a<b : a <N b) → subtractionNResult.result (-N (inl a<b)) <N b
|
||||
subLess {zero} {b} 0<a a<b = exFalso (TotalOrder.irreflexive ℕTotalOrder 0<a)
|
||||
subLess {succ a} {b} _ a<b with -N (inl a<b)
|
||||
... | record { result = b-a ; pr = pr } = record { x = a ; proof = pr }
|
||||
|
||||
inverseZn : {n : ℕ} → {pr : 0 <N n} → (a : ℤn n pr) → Sg (ℤn n pr) (λ i → i +n a ≡ record { x = 0 ; xLess = pr } )
|
||||
inverseZn {zero} {()}
|
||||
inverseZn {succ n} {0<n} record { x = zero ; xLess = zeroLess } = record { x = zero ; xLess = zeroLess } , plusZnIdentityLeft _
|
||||
inverseZn {succ n} {0<n} record { x = (succ a) ; xLess = aLess } = ans , pr
|
||||
where
|
||||
ans = record { x = subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) ; xLess = subLess (succIsPositive a) aLess }
|
||||
pr : ans +n record { x = (succ a) ; xLess = aLess } ≡ record { x = 0 ; xLess = 0<n }
|
||||
pr with totality (subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N (succ a)) (succ n)
|
||||
... | inl (inl x) = exFalso f
|
||||
where
|
||||
h : subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N succ a ≡ succ n
|
||||
h with -N (inl (canRemoveSuccFrom<N aLess))
|
||||
h | record { result = result ; pr = pr } rewrite equalityCommutative pr = Semiring.commutative ℕSemiring result (succ a)
|
||||
f : False
|
||||
f = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ x h)
|
||||
... | inl (inr x) = exFalso f
|
||||
where
|
||||
h : subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N succ a ≡ succ n
|
||||
h = addMinus {succ a} {succ n} (inl aLess)
|
||||
f : False
|
||||
f rewrite h = TotalOrder.irreflexive ℕTotalOrder x
|
||||
... | (inr n-a+sa=sn) = equalityZn _ _ refl
|
||||
|
||||
ℤnGroup : (n : ℕ) → (pr : 0 <N n) → Group (reflSetoid (ℤn n pr)) _+n_
|
||||
ℤnGroup n pr = record { 0G = record { x = 0 ; xLess = pr } ; inverse = λ a → underlying (inverseZn a) ; +Associative = λ {a} {b} {c} → plusZnAssociative a b c ; identRight = λ {a} → plusZnIdentityRight a ; identLeft = λ {a} → plusZnIdentityLeft a ; invLeft = λ {a} → helpInvLeft a ; invRight = λ {a} → helpInvRight a ; +WellDefined = reflGroupWellDefined }
|
||||
where
|
||||
helpInvLeft : (a : ℤn n pr) → underlying (inverseZn a) +n a ≡ record { x = 0 ; xLess = pr }
|
||||
helpInvLeft a with inverseZn a
|
||||
... | -a , pr-a = pr-a
|
||||
helpInvRight : (a : ℤn n pr) → a +n underlying (inverseZn a) ≡ record { x = 0 ; xLess = pr }
|
||||
helpInvRight a rewrite plusZnCommutative a (underlying (inverseZn a)) = helpInvLeft a
|
||||
ℤnGroup : (n : ℕ) → .(pr : 0 <N n) → Group (reflSetoid (ℤn n pr)) (_+n_ pr)
|
||||
Group.+WellDefined (ℤnGroup n 0<n) refl refl = refl
|
||||
Group.0G (ℤnGroup n 0<n) = record { x = 0 ; xLess = 0<n }
|
||||
Group.inverse (ℤnGroup n 0<n) = inverseN 0<n
|
||||
Group.+Associative (ℤnGroup n 0<n) {a} {b} {c} = equalityCommutative (plusZnAssociative 0<n a b c)
|
||||
Group.identRight (ℤnGroup n 0<n) {a} = plusZnIdentityRight 0<n a
|
||||
Group.identLeft (ℤnGroup n 0<n) {a} = plusZnIdentityLeft 0<n a
|
||||
Group.invLeft (ℤnGroup n 0<n) {a} = invLeft 0<n a
|
||||
Group.invRight (ℤnGroup n 0<n) {a} = transitivity (plusZnCommutative 0<n a (inverseN 0<n a)) (invLeft 0<n a)
|
||||
|
||||
ℤnAbGroup : (n : ℕ) → (pr : 0 <N n) → AbelianGroup (ℤnGroup n pr)
|
||||
AbelianGroup.commutative (ℤnAbGroup n pr) {a} {b} = plusZnCommutative a b
|
||||
AbelianGroup.commutative (ℤnAbGroup n pr) {a} {b} = plusZnCommutative pr a b
|
||||
|
||||
ℤnFinite : (n : ℕ) → (pr : 0 <N n) → FiniteGroup (ℤnGroup n pr) (FinSet n)
|
||||
SetoidToSet.project (FiniteGroup.toSet (ℤnFinite n pr)) record { x = x ; xLess = xLess } = ofNat x xLess
|
||||
SetoidToSet.wellDefined (FiniteGroup.toSet (ℤnFinite n pr)) x y x=y rewrite x=y = refl
|
||||
Surjection.property (SetoidToSet.surj (FiniteGroup.toSet (ℤnFinite n pr))) b = record { x = toNat b ; xLess = toNatSmaller b } , ofNatToNat b
|
||||
SetoidToSet.inj (FiniteGroup.toSet (ℤnFinite zero ())) x y x=y
|
||||
SetoidToSet.inj (FiniteGroup.toSet (ℤnFinite (succ n) pr)) record { x = x ; xLess = xLess } record { x = y ; xLess = yLess } x=y = equalityZn _ _ b
|
||||
where
|
||||
b : x ≡ y
|
||||
b = ofNatInjective x y xLess yLess x=y
|
||||
FiniteSet.size (FiniteGroup.finite (ℤnFinite n pr)) = n
|
||||
FiniteSet.mapping (FiniteGroup.finite (ℤnFinite n pr)) = id
|
||||
FiniteSet.bij (FiniteGroup.finite (ℤnFinite n pr)) = idIsBijective
|
||||
SetoidToSet.project (FiniteGroup.toSet (ℤnFinite (succ n) 0<n)) record { x = x ; xLess = xLess } = ofNat x xLess
|
||||
SetoidToSet.wellDefined (FiniteGroup.toSet (ℤnFinite (succ n) 0<n)) x y x=y rewrite x=y = refl
|
||||
SetoidToSet.surj (FiniteGroup.toSet (ℤnFinite (succ n) 0<n)) b = record { x = toNat b ; xLess = toNatSmaller b } , ofNatToNat b
|
||||
SetoidToSet.inj (FiniteGroup.toSet (ℤnFinite (succ n) 0<n)) record { x = x ; xLess = xLess } record { x = y ; xLess = yLess } eq = equalityZn (ofNatInjective x y xLess yLess eq)
|
||||
FiniteGroup.finite (ℤnFinite n pr) = record { size = n ; mapping = id ; bij = idIsBijective }
|
||||
|
101
Numbers/Modulo/ModuloFunction.agda
Normal file
101
Numbers/Modulo/ModuloFunction.agda
Normal file
@@ -0,0 +1,101 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.EuclideanAlgorithm
|
||||
open import Semirings.Definition
|
||||
open import Orders.Total.Definition
|
||||
|
||||
module Numbers.Modulo.ModuloFunction where
|
||||
|
||||
open TotalOrder ℕTotalOrder
|
||||
|
||||
private
|
||||
notBigger : (a : ℕ) {n : ℕ} → succ a +N n ≡ a → False
|
||||
notBigger (succ a) {n} pr = notBigger a {n} (succInjective pr)
|
||||
|
||||
notBigger' : (a : ℕ) {n : ℕ} → succ a +N n ≡ n → False
|
||||
notBigger' (succ a) {succ n} pr rewrite Semiring.commutative ℕSemiring a (succ n) | Semiring.commutative ℕSemiring n a = notBigger' _ (succInjective pr)
|
||||
|
||||
abstract
|
||||
mod : (n : ℕ) → .(pr : 0 <N n) → (a : ℕ) → ℕ
|
||||
mod (succ n) 0<n a = divisionAlgResult.rem (divisionAlg (succ n) a)
|
||||
|
||||
modIsMod : {n : ℕ} → .(pr : 0 <N n) → (a : ℕ) → a <N n → mod n pr a ≡ a
|
||||
modIsMod {succ n} 0<n a a<n with divisionAlg (succ n) a
|
||||
modIsMod {succ n} 0<n a a<n | record { quot = zero ; rem = rem ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } rewrite multiplicationNIsCommutative n 0 = pr
|
||||
modIsMod {succ n} 0<n a (le x proof) | record { quot = succ quot ; rem = rem ; pr = pr ; remIsSmall = inl rem<sn ; quotSmall = quotSmall } = exFalso (notBigger a v)
|
||||
where
|
||||
t : ((succ (a +N x)) *N succ quot) +N rem ≡ a
|
||||
t rewrite Semiring.commutative ℕSemiring a x = transitivity (applyEquality (λ i → (i *N succ quot) +N rem) proof) pr
|
||||
u : (succ quot *N succ (a +N x)) +N rem ≡ a
|
||||
u = transitivity (applyEquality (_+N rem) (multiplicationNIsCommutative (succ quot) _)) t
|
||||
v : succ a +N ((x +N quot *N succ (a +N x)) +N rem) ≡ a
|
||||
v = transitivity (applyEquality succ (transitivity (Semiring.+Associative ℕSemiring a _ rem) (applyEquality (_+N rem) (Semiring.+Associative ℕSemiring a x _)))) u
|
||||
|
||||
mod<N : {n : ℕ} → .(pr : 0 <N n) → (a : ℕ) → (mod n pr a) <N n
|
||||
mod<N {succ n} pr a with divisionAlg (succ n) a
|
||||
mod<N {succ n} 0<n a | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl x ; quotSmall = quotSmall } = x
|
||||
|
||||
modIdempotent : {n : ℕ} → .(pr : 0 <N n) → (a : ℕ) → mod n pr (mod n pr a) ≡ mod n pr a
|
||||
modIdempotent 0<n a = modIsMod 0<n (mod _ 0<n a) (mod<N 0<n a)
|
||||
|
||||
modAnd+n : {n : ℕ} → .(pr : 0 <N n) → (a : ℕ) → mod n pr (n +N a) ≡ mod n pr a
|
||||
modAnd+n {succ n} 0<sn a with divisionAlg (succ n) (succ n +N a)
|
||||
modAnd+n {succ n} 0<sn a | record { quot = 0 ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl q } rewrite multiplicationNIsCommutative n 0 | pr = exFalso (notBigger (succ n) (transitivity (applyEquality succ (transitivity (applyEquality succ (Semiring.+Associative ℕSemiring n a _)) (Semiring.commutative ℕSemiring (succ (n +N a)) _))) (_<N_.proof remIsSmall)))
|
||||
modAnd+n {succ n} 0<sn a | record { quot = succ quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl q } = modIsUnique (record { quot = quot ; rem = rem ; pr = t ; remIsSmall = inl remIsSmall ; quotSmall = inl q }) (divisionAlg (succ n) a)
|
||||
where
|
||||
g : succ n +N ((quot *N succ n) +N rem) ≡ succ n +N a
|
||||
g = transitivity (Semiring.+Associative ℕSemiring (succ n) _ rem) (transitivity (applyEquality (_+N rem) (multiplicationNIsCommutative (succ quot) _)) pr)
|
||||
h : (quot *N succ n) +N rem ≡ a
|
||||
h = canSubtractFromEqualityLeft g
|
||||
t : (succ n *N quot) +N rem ≡ a
|
||||
t = transitivity (applyEquality (_+N rem) (multiplicationNIsCommutative (succ n) quot)) h
|
||||
|
||||
modExtracts : {n : ℕ} → .(pr : 0 <N n) → (a b : ℕ) → mod n pr (a +N b) ≡ mod n pr (mod n pr a +N mod n pr b)
|
||||
modExtracts {succ n} 0<sn a b with divisionAlg (succ n) a
|
||||
modExtracts {succ n} 0<sn a b | record { quot = sn/a ; rem = sn%a ; pr = prA ; remIsSmall = remIsSmallA ; quotSmall = quotSmallA } with divisionAlg (succ n) b
|
||||
modExtracts {succ n} 0<sn a b | record { quot = sn/a ; rem = sn%a ; pr = prA ; remIsSmall = remIsSmallA ; quotSmall = quotSmallA } | record { quot = sn/b ; rem = sn%b ; pr = prB ; remIsSmall = remIsSmallB ; quotSmall = quotSmallB } with divisionAlg (succ n) (a +N b)
|
||||
modExtracts {succ n} 0<sn a b | record { quot = sn/a ; rem = sn%a ; pr = prA ; remIsSmall = remIsSmallA ; quotSmall = quotSmallA } | record { quot = sn/b ; rem = sn%b ; pr = prB ; remIsSmall = remIsSmallB ; quotSmall = quotSmallB } | record { quot = sn/a+b ; rem = sn%a+b ; pr = prA+B ; remIsSmall = remIsSmallA+B ; quotSmall = quotSmallA+B } with divisionAlg (succ n) (sn%a +N sn%b)
|
||||
modExtracts {succ n} 0<sn a b | record { quot = sn/a ; rem = sn%a ; pr = prA ; remIsSmall = inl sn%a<sn ; quotSmall = inl _ } | record { quot = sn/b ; rem = sn%b ; pr = prB ; remIsSmall = inl sn%b<sn ; quotSmall = inl _ } | record { quot = sn/a+b ; rem = sn%a+b ; pr = prA+B ; remIsSmall = inl sn%a+b<sn ; quotSmall = inl _ } | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl _ } = equalityCommutative whoa
|
||||
where
|
||||
t : ((succ n *N sn/a) +N sn%a) +N ((succ n *N sn/b) +N sn%b) ≡ a +N b
|
||||
t = +NWellDefined prA prB
|
||||
t' : (succ n *N (sn/a +N sn/b)) +N (sn%a +N sn%b) ≡ a +N b
|
||||
t' = transitivity (transitivity (Semiring.+Associative ℕSemiring (succ n *N (sn/a +N sn/b)) sn%a sn%b) (transitivity (applyEquality (_+N sn%b) (transitivity (transitivity (applyEquality (_+N sn%a) (transitivity (applyEquality (succ n *N_) (Semiring.commutative ℕSemiring sn/a sn/b)) (Semiring.+DistributesOver* ℕSemiring (succ n) sn/b sn/a))) (equalityCommutative (Semiring.+Associative ℕSemiring (succ n *N sn/b) (succ n *N sn/a) sn%a))) (Semiring.commutative ℕSemiring (succ n *N sn/b) ((succ n *N sn/a) +N sn%a)))) (equalityCommutative (Semiring.+Associative ℕSemiring ((succ n *N sn/a) +N sn%a) (succ n *N sn/b) sn%b)))) t
|
||||
thing : ((succ n) *N quot) +N rem ≡ sn%a +N sn%b
|
||||
thing = pr
|
||||
t'' : (succ n *N ((sn/a +N sn/b) +N quot)) +N rem ≡ a +N b
|
||||
t'' rewrite Semiring.+DistributesOver* ℕSemiring (succ n) (sn/a +N sn/b) quot | equalityCommutative (Semiring.+Associative ℕSemiring (succ n *N (sn/a +N sn/b)) (succ n *N quot) rem) | thing = t'
|
||||
whoa : rem ≡ sn%a+b
|
||||
whoa = modUniqueLemma _ _ remIsSmall sn%a+b<sn (transitivity t'' (equalityCommutative prA+B))
|
||||
|
||||
modAddition : {n : ℕ} → .(pr : 0 <N n) → {a b : ℕ} → (a <N n) → (b <N n) → (a +N b ≡ mod n pr (a +N b)) || (a +N b ≡ n +N mod n pr (a +N b))
|
||||
modAddition {succ n} 0<sn {a} {b} a<n b<n with divisionAlg (succ n) (a +N b)
|
||||
modAddition {succ n} 0<sn {a} {b} a<n b<n | record { quot = zero ; rem = rem ; pr = pr ; remIsSmall = inl rem<sn ; quotSmall = inl _ } = inl (equalityCommutative (transitivity (applyEquality (_+N rem) (multiplicationNIsCommutative 0 n)) pr))
|
||||
modAddition {succ n} 0<sn {a} {b} a<n b<n | record { quot = succ zero ; rem = rem ; pr = pr ; remIsSmall = inl rem<sn ; quotSmall = inl _ } = inr (transitivity (equalityCommutative pr) (applyEquality (λ i → succ i +N rem) (transitivity (multiplicationNIsCommutative n 1) (Semiring.sumZeroRight ℕSemiring n))))
|
||||
modAddition {succ n} 0<sn {a} {b} (le x prA) (le y prB) | record { quot = succ (succ quot) ; rem = 0 ; pr = pr ; remIsSmall = inl rem<sn ; quotSmall = inl _ } rewrite Semiring.sumZeroRight ℕSemiring (succ n *N succ (succ quot)) | multiplicationNIsCommutative (succ n) (succ (succ quot)) = exFalso (notBigger' ((x +N succ y) +N (quot *N succ n)) u)
|
||||
where
|
||||
t : ((succ x +N a) +N succ (y +N b)) +N (quot *N succ n) ≡ a +N b
|
||||
t = transitivity (transitivity (applyEquality (_+N quot *N succ n) (+NWellDefined prA prB)) (equalityCommutative (Semiring.+Associative ℕSemiring (succ n) (succ n) (quot *N succ n)))) pr
|
||||
u : ((succ x +N succ y) +N (quot *N succ n)) +N (a +N b) ≡ a +N b
|
||||
u rewrite Semiring.commutative ℕSemiring ((x +N succ y) +N quot *N succ n) (a +N b) | Semiring.+Associative ℕSemiring (succ a +N b) (x +N succ y) (quot *N succ n) = transitivity (applyEquality (λ i → succ i +N quot *N succ n) (transitivity (Semiring.commutative ℕSemiring (a +N b) _) (transitivity (equalityCommutative (Semiring.+Associative ℕSemiring x (succ y) _)) (transitivity (applyEquality (x +N_) (transitivity (Semiring.+Associative ℕSemiring (succ y) a b) (transitivity (applyEquality (_+N b) (Semiring.commutative ℕSemiring _ a)) (equalityCommutative (Semiring.+Associative ℕSemiring a (succ y) b))))) (Semiring.+Associative ℕSemiring x a _))))) t
|
||||
modAddition {succ n} 0<sn {a} {b} a<n b<n | record { quot = succ (succ quot) ; rem = succ rem ; pr = pr ; remIsSmall = inl rem<sn ; quotSmall = inl _ } = exFalso (irreflexive (<Transitive u t))
|
||||
where
|
||||
t : (succ n *N succ (succ quot) +N succ rem) <N succ n +N succ n
|
||||
t = identityOfIndiscernablesLeft _<N_ (addStrongInequalities a<n b<n) (equalityCommutative pr)
|
||||
lemm : {a c d : ℕ} → (a +N a) <N (a *N succ (succ c)) +N succ d
|
||||
lemm {a} {c} {d} = le (a *N c +N d) f
|
||||
where
|
||||
f : succ ((a *N c +N d) +N (a +N a)) ≡ a *N succ (succ c) +N succ d
|
||||
f rewrite multiplicationNIsCommutative a (succ (succ c)) | Semiring.+Associative ℕSemiring a a (c *N a) | multiplicationNIsCommutative c a | Semiring.commutative ℕSemiring (a *N c +N d) (a +N a) | Semiring.+Associative ℕSemiring (a +N a) (a *N c) d | Semiring.commutative ℕSemiring ((a +N a) +N a *N c) (succ d) = applyEquality succ (Semiring.commutative ℕSemiring _ d)
|
||||
u : (succ n +N succ n) <N (succ n *N succ (succ quot) +N succ rem)
|
||||
u = lemm {succ n}
|
||||
|
||||
modN : {n : ℕ} → .(0<n : 0 <N n) → mod n 0<n n ≡ 0
|
||||
modN {succ n} 0<n = modIsUnique (divisionAlg (succ n) (succ n)) record { quot = 1 ; rem = zero ; pr = ans ; remIsSmall = inl (le n (Semiring.sumZeroRight ℕSemiring (succ n))) ; quotSmall = inl (le n (Semiring.sumZeroRight ℕSemiring (succ n))) }
|
||||
where
|
||||
ans : succ (n *N 1 +N 0) ≡ succ n
|
||||
ans rewrite multiplicationNIsCommutative n 1 | Semiring.sumZeroRight ℕSemiring n | Semiring.sumZeroRight ℕSemiring n = refl
|
@@ -356,3 +356,9 @@ euclid a b = inducted (succ a +N b) a b (a<SuccA (a +N b))
|
||||
go maxSum indHyp = λ a b a+b<maxSum → go' maxSum a b a+b<maxSum indHyp
|
||||
inducted : ∀ x → P x
|
||||
inducted = rec <NWellfounded P go
|
||||
|
||||
divOneImpliesOne : {a : ℕ} → a ∣ 1 → a ≡ 1
|
||||
divOneImpliesOne {zero} a|1 = exFalso (zeroDividesNothing _ a|1)
|
||||
divOneImpliesOne {succ zero} a|1 = refl
|
||||
divOneImpliesOne {succ (succ a)} (divides record { quot = zero ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } refl) rewrite Semiring.sumZeroRight ℕSemiring (a *N zero) | multiplicationNIsCommutative a 0 = exFalso (naughtE pr)
|
||||
divOneImpliesOne {succ (succ a)} (divides record { quot = (succ quot) ; rem = .0 ; pr = pr ; remIsSmall = remIsSmall ; quotSmall = quotSmall } refl) rewrite Semiring.commutative ℕSemiring quot (succ (quot +N a *N succ quot)) = exFalso (naughtE (equalityCommutative (succInjective pr)))
|
||||
|
@@ -16,7 +16,7 @@ open import Orders.Total.Definition
|
||||
|
||||
module Numbers.Naturals.Naturals where
|
||||
|
||||
record subtractionNResult (a b : ℕ) (p : a ≤N b) : Set where
|
||||
record subtractionNResult (a b : ℕ) .(p : a ≤N b) : Set where
|
||||
field
|
||||
result : ℕ
|
||||
pr : a +N result ≡ b
|
||||
@@ -32,11 +32,11 @@ subtractionNWellDefined {a} {.a} {inr refl} {pr2} record { result = result1 ; pr
|
||||
|
||||
-N : {a : ℕ} → {b : ℕ} → (pr : a ≤N b) → subtractionNResult a b pr
|
||||
-N {zero} {b} prAB = record { result = b ; pr = refl }
|
||||
-N {succ a} {zero} (inl (le x ()))
|
||||
-N {succ a} {zero} (inl ())
|
||||
-N {succ a} {zero} (inr ())
|
||||
-N {succ a} {succ b} (inl x) with -N {a} {b} (inl (canRemoveSuccFrom<N x))
|
||||
-N {succ a} {succ b} (inl x) | record { result = result ; pr = pr } = record { result = result ; pr = applyEquality succ pr }
|
||||
-N {succ a} {succ .a} (inr refl) = record { result = 0 ; pr = applyEquality succ (addZeroRight a) }
|
||||
-N {succ a} {succ b} (inr pr) = record { result = 0 ; pr = transitivity (applyEquality succ (addZeroRight a)) pr }
|
||||
|
||||
addOneToWeakInequality : {a b : ℕ} → (a ≤N b) → (succ a ≤N succ b)
|
||||
addOneToWeakInequality {a} {b} (inl ineq) = inl (succPreservesInequality ineq)
|
||||
@@ -51,7 +51,7 @@ addMinus {zero} {succ b} pr = applyEquality succ (addZeroRight b)
|
||||
addMinus {succ a} {zero} (inl (le x ()))
|
||||
addMinus {succ a} {zero} (inr ())
|
||||
addMinus {succ a} {succ b} (inl x) with (-N {succ a} {succ b} (inl x))
|
||||
addMinus {succ a} {succ b} (inl x) | record { result = result ; pr = pr } = transitivity (transitivity (applyEquality (_+N succ a) (transitivity (subtractionNWellDefined {p1 = inl (canRemoveSuccFrom<N x)} (record { result = subtractionNResult.result (-N (inl (canRemoveSuccFrom<N x))) ; pr = transitivity (additionNIsCommutative a _) (addMinus (inl (canRemoveSuccFrom<N x)))}) previous) (equalityCommutative t))) (additionNIsCommutative result (succ a))) pr
|
||||
addMinus {succ a} {succ b} (inl x) | record { result = result ; pr = pr } = transitivity (transitivity (applyEquality (_+N succ a) (transitivity (subtractionNWellDefined {p1 = inl (canRemoveSuccFrom<N x)} {p2 = inl (canRemoveSuccFrom<N x)} (record { result = subtractionNResult.result (-N (inl (canRemoveSuccFrom<N x))) ; pr = transitivity (additionNIsCommutative a _) (addMinus (inl (canRemoveSuccFrom<N x)))}) previous) (equalityCommutative t))) (additionNIsCommutative result (succ a))) pr
|
||||
where
|
||||
pr'' : (a <N b) || (a ≡ b)
|
||||
pr'' = (inl (le (_<N_.x x) (transitivity (equalityCommutative (succExtracts (_<N_.x x) a)) (succInjective (_<N_.proof x)))))
|
||||
@@ -296,9 +296,9 @@ subtractProduct {succ zero} {b} {c} aPos b<c = s'
|
||||
r = subtractionNWellDefined {b +N 0 *N b} {c +N 0 *N c} {(lessCast' (lessCast (inl b<c) (equalityCommutative (addZeroRight b))) (equalityCommutative (addZeroRight c)))} {inl (lessRespectsMultiplicationLeft b c 1 b<c aPos)} (underlying resbc'') (-N {b +N 0 *N b} {c +N 0 *N c} (inl (lessRespectsMultiplicationLeft b c 1 b<c aPos)))
|
||||
g (a , b) = b
|
||||
g' (a , b) = b
|
||||
resbc'' = subtractionCast' (equalityCommutative (addZeroRight c)) (underlying resbc')
|
||||
resbc'' = subtractionCast' {pr = inl (identityOfIndiscernablesLeft _<N_ b<c (equalityCommutative (addZeroRight b)))} (equalityCommutative (addZeroRight c)) (underlying resbc')
|
||||
q = transitivity {_} {_} {subtractionNResult.result resbc} {subtractionNResult.result (underlying resbc')} {subtractionNResult.result (underlying resbc'')} (g resbc') (g' resbc'')
|
||||
resbc' = subtractionCast {b} {c} {b +N 0 *N b} (equalityCommutative (addZeroRight b)) resbc
|
||||
resbc' = subtractionCast {b} {c} {b +N 0 *N b} {inl b<c} (equalityCommutative (addZeroRight b)) resbc
|
||||
|
||||
subtractProduct {succ (succ a)} {b} {c} aPos b<c =
|
||||
let
|
||||
|
@@ -80,3 +80,7 @@ cancelInequalityLeft {a} {succ b} {succ c} pr = succPreservesInequality q'
|
||||
q = canFlipMultiplicationsInIneq {b} {a} {c} {a} p
|
||||
q' : b <N c
|
||||
q' = cancelInequalityLeft {a} {b} {c} q
|
||||
|
||||
<NProp : {a b : ℕ} → .(a <N b) → a <N b
|
||||
<NProp {zero} {succ b} a<b = succIsPositive _
|
||||
<NProp {succ a} {succ b} a<b = succPreservesInequality (<NProp (canRemoveSuccFrom<N a<b))
|
||||
|
@@ -35,3 +35,9 @@ productZeroImpliesOperandZero : {a b : ℕ} → a *N b ≡ 0 → (a ≡ 0) || (b
|
||||
productZeroImpliesOperandZero {zero} {b} pr = inl refl
|
||||
productZeroImpliesOperandZero {succ a} {zero} pr = inr refl
|
||||
productZeroImpliesOperandZero {succ a} {succ b} ()
|
||||
|
||||
*NWellDefined : {a b c d : ℕ} → (a ≡ c) → (b ≡ d) → a *N b ≡ c *N d
|
||||
*NWellDefined refl refl = refl
|
||||
|
||||
+NWellDefined : {a b c d : ℕ} → (a ≡ c) → (b ≡ d) → a +N b ≡ c +N d
|
||||
+NWellDefined refl refl = refl
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Sets.FinSet
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
|
@@ -8,6 +8,7 @@ open import Rings.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Numbers.Modulo.Definition
|
||||
@@ -32,25 +33,25 @@ module Rings.Examples.Proofs where
|
||||
subtractionEquiv (succ a) {b} {c} c<b pr = equivalentSubtraction 0 b (succ a) c (succIsPositive a) c<b (equalityCommutative pr)
|
||||
|
||||
modNExampleSurjective' : (n : ℕ) → (pr : 0 <N n) → Surjection (mod' n pr)
|
||||
Surjection.property (modNExampleSurjective' zero ())
|
||||
Surjection.property (modNExampleSurjective' (succ n) pr) record { x = x ; xLess = xLess } with divisionAlg (succ n) x
|
||||
Surjection.property (modNExampleSurjective' (succ n) p) record { x = x ; xLess = xLess } | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = q } = nonneg x , lhs'
|
||||
modNExampleSurjective' zero ()
|
||||
modNExampleSurjective' (succ n) pr record { x = x ; xLess = xLess } with divisionAlg (succ n) x
|
||||
modNExampleSurjective' (succ n) p record { x = x ; xLess = xLess } | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = q } = nonneg x , lhs'
|
||||
where
|
||||
rs' : rem ≡ x
|
||||
rs' = modIsUnique (record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = q }) (record { quot = 0 ; rem = x ; pr = blah ; remIsSmall = inl xLess ; quotSmall = inl (succIsPositive n) })
|
||||
rs' = modIsUnique (record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = q }) (record { quot = 0 ; rem = x ; pr = blah ; remIsSmall = inl (<NProp xLess) ; quotSmall = inl (succIsPositive n) })
|
||||
where
|
||||
blah : n *N 0 +N x ≡ x
|
||||
blah rewrite multiplicationNIsCommutative n 0 = refl
|
||||
lhs : nToZn' (succ n) p x ≡ record { x = rem ; xLess = remIsSmall }
|
||||
lhs with divisionAlg (succ n) x
|
||||
lhs | record { quot = quot' ; rem = rem' ; pr = pr' ; remIsSmall = inl t ; quotSmall = quotSmall } = equalityZn _ _ (equalityCommutative rs)
|
||||
lhs | record { quot = quot' ; rem = rem' ; pr = pr' ; remIsSmall = inl t ; quotSmall = quotSmall } = equalityZn (equalityCommutative rs)
|
||||
where
|
||||
rs : rem ≡ rem'
|
||||
rs = modIsUnique (record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall; quotSmall = q }) (record { quot = quot' ; rem = rem' ; pr = pr' ; remIsSmall = inl t ; quotSmall = quotSmall })
|
||||
lhs | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
|
||||
lhs' : nToZn' (succ n) p x ≡ record { x = x ; xLess = xLess }
|
||||
lhs' = transitivity lhs (equalityZn _ _ rs')
|
||||
Surjection.property (modNExampleSurjective' (succ n) p) record { x = x ; xLess = xLess } | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
|
||||
lhs' = transitivity lhs (equalityZn rs')
|
||||
modNExampleSurjective' (succ n) p record { x = x ; xLess = xLess } | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inr () ; quotSmall = quotSmall }
|
||||
|
||||
{-
|
||||
modNExampleGroupHom' : (n : ℕ) → (pr : 0 <N n) → GroupHom ℤGroup (ℤnGroup n pr) (mod' n pr)
|
||||
|
@@ -1,7 +1,6 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Sets.FinSet
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
|
@@ -5,14 +5,14 @@ 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 Sets.FinSet.Definition
|
||||
open import Semirings.Definition
|
||||
open import Sets.CantorBijection.Proofs
|
||||
|
||||
module Sets.CantorBijection.CantorBijection where
|
||||
|
||||
open Sets.CantorBijection.Proofs using (cantorInverse ; cantorInverseLemma) public
|
||||
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
|
||||
cantorBijection : Bijection cantorInverse
|
||||
Bijection.inj cantorBijection {x} {y} = cantorInverseInjective x y
|
||||
Bijection.surj cantorBijection = cantorInverseSurjective
|
||||
|
@@ -7,7 +7,7 @@ open import Functions
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.WellFounded
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSet.Definition
|
||||
open import Semirings.Definition
|
||||
open import Orders.Total.Definition
|
||||
open import Orders.Partial.Definition
|
||||
|
@@ -6,7 +6,7 @@ open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSet.Definition
|
||||
open import Semirings.Definition
|
||||
open import Sets.CantorBijection.Order
|
||||
open import Orders.Total.Definition
|
||||
|
@@ -7,7 +7,7 @@ open import Functions
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Sets.FinSet
|
||||
open import Sets.Cardinality.Finite.Definition
|
||||
open import Semirings.Definition
|
||||
open import Sets.CantorBijection.CantorBijection
|
||||
open import Orders.Total.Definition
|
||||
@@ -71,21 +71,21 @@ sqrtFloor (succ n) | (a ,, b) , pr | inr x = (succ a ,, 0) , (q ,, succIsPositiv
|
||||
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))
|
||||
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) {inl x} {inl y} pr rewrite Semiring.commutative ℕSemiring (CountablyInfiniteSet.counting X x) 0 | Semiring.commutative ℕSemiring (CountablyInfiniteSet.counting X y) 0 | doubleIsAddTwo (CountablyInfiniteSet.counting X x) | doubleIsAddTwo (CountablyInfiniteSet.counting X y) = applyEquality inl (Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij X)) inter)
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) {inl x} {inl y} pr rewrite Semiring.commutative ℕSemiring (CountablyInfiniteSet.counting X x) 0 | Semiring.commutative ℕSemiring (CountablyInfiniteSet.counting X y) 0 | doubleIsAddTwo (CountablyInfiniteSet.counting X x) | doubleIsAddTwo (CountablyInfiniteSet.counting X y) = applyEquality inl (Bijection.inj (CountablyInfiniteSet.countingIsBij X) inter)
|
||||
where
|
||||
inter : CountablyInfiniteSet.counting X x ≡ CountablyInfiniteSet.counting X y
|
||||
inter = doubleLemma (CountablyInfiniteSet.counting X x) (CountablyInfiniteSet.counting X y) pr
|
||||
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) {inl x} {inr y} pr = exFalso (evenCannotBeOdd (CountablyInfiniteSet.counting X x) (CountablyInfiniteSet.counting Y y) pr)
|
||||
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) {inr x} {inl y} pr = exFalso (evenCannotBeOdd (CountablyInfiniteSet.counting X y) (CountablyInfiniteSet.counting Y x) (equalityCommutative pr))
|
||||
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) {inr x} {inr y} pr = applyEquality inr (Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij Y)) (doubleLemma (CountablyInfiniteSet.counting Y x) (CountablyInfiniteSet.counting Y y) (succInjective pr) ))
|
||||
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) b with aMod2 b
|
||||
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) b | a , inl x with Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij X)) a
|
||||
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) b | a , inl x | r , pr = inl r , ans
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) {inl x} {inr y} pr = exFalso (evenCannotBeOdd (CountablyInfiniteSet.counting X x) (CountablyInfiniteSet.counting Y y) pr)
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) {inr x} {inl y} pr = exFalso (evenCannotBeOdd (CountablyInfiniteSet.counting X y) (CountablyInfiniteSet.counting Y x) (equalityCommutative pr))
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) {inr x} {inr y} pr = applyEquality inr (Bijection.inj (CountablyInfiniteSet.countingIsBij Y) (doubleLemma (CountablyInfiniteSet.counting Y x) (CountablyInfiniteSet.counting Y y) (succInjective pr) ))
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) b with aMod2 b
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) b | a , inl x with Bijection.surj (CountablyInfiniteSet.countingIsBij X) a
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) b | a , inl x | r , pr = inl r , ans
|
||||
where
|
||||
ans : 2 *N CountablyInfiniteSet.counting X r ≡ b
|
||||
ans rewrite pr = x
|
||||
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) b | a , inr x with Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij Y)) a
|
||||
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y))) b | a , inr x | r , pr = inr r , ans
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) b | a , inr x with Bijection.surj (CountablyInfiniteSet.countingIsBij Y) a
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) b | a , inr x | r , pr = inr r , ans
|
||||
where
|
||||
ans : succ (2 *N CountablyInfiniteSet.counting Y r) ≡ b
|
||||
ans rewrite pr = x
|
||||
@@ -100,17 +100,17 @@ reflPair : {a b : _} {A : Set a} {B : Set b} {x1 x2 : A} {y1 y2 : B} → (x1 ≡
|
||||
reflPair refl refl = refl
|
||||
|
||||
bijectionOfCountableSetIsCountable : {a b : _} {A : Set a} {B : Set b} → (X : CountablyInfiniteSet A) {f : A → B} → (bij : Bijection f) → CountablyInfiniteSet B
|
||||
CountablyInfiniteSet.counting (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj }) b with Surjection.property surj b
|
||||
CountablyInfiniteSet.counting (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj }) b with surj b
|
||||
CountablyInfiniteSet.counting (bijectionOfCountableSetIsCountable record { counting = counting ; countingIsBij = countingIsBij } {f} record { inj = inj ; surj = surj }) b | a , pr = counting a
|
||||
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj }))) {m} {n} m=n with Surjection.property surj m
|
||||
Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj }))) {m} {n} m=n | a , b with Surjection.property surj n
|
||||
... | c , d = transitivity (equalityCommutative b) (transitivity (applyEquality f (Injection.property (Bijection.inj (CountablyInfiniteSet.countingIsBij X)) m=n)) d)
|
||||
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj }))) b with Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij X)) b
|
||||
Surjection.property (Bijection.surj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj }))) b | a , pr = f a , s
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj })) {m} {n} m=n with surj m
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj })) {m} {n} m=n | a , b with surj n
|
||||
... | c , d = transitivity (equalityCommutative b) (transitivity (applyEquality f (Bijection.inj (CountablyInfiniteSet.countingIsBij X) m=n)) d)
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj })) b with Bijection.surj (CountablyInfiniteSet.countingIsBij X) b
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj })) b | a , pr = f a , s
|
||||
where
|
||||
s : CountablyInfiniteSet.counting (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj }) (f a) ≡ b
|
||||
s with Surjection.property surj (f a)
|
||||
s | t , u = transitivity (applyEquality (CountablyInfiniteSet.counting X) (Injection.property inj u)) pr
|
||||
s with surj (f a)
|
||||
s | t , u = transitivity (applyEquality (CountablyInfiniteSet.counting X) (inj u)) pr
|
||||
|
||||
N^2Countable : CountablyInfiniteSet (ℕ && ℕ)
|
||||
CountablyInfiniteSet.counting N^2Countable x = Invertible.inverse (bijectionImpliesInvertible (cantorBijection)) x
|
||||
@@ -123,14 +123,14 @@ tupleLmap : {a b c : _} {A : Set a} {B : Set b} {C : Set c} → (f : A → C)
|
||||
tupleLmap f (a ,, b) = (f a ,, b)
|
||||
|
||||
bijectionOnRightOfProduct : {a b c : _} {A : Set a} {B : Set b} {C : Set c} → {f : B → C} → Bijection f → Bijection (tupleRmap {A = A} f)
|
||||
Injection.property (Bijection.inj (bijectionOnRightOfProduct {A = A} {B} {C} {f} bij)) {fst ,, snd} {fst₁ ,, snd₁} gx=gy rewrite firstEqualityOfPair gx=gy | Injection.property (Bijection.inj bij) (secondEqualityOfPair gx=gy) = refl
|
||||
Surjection.property (Bijection.surj (bijectionOnRightOfProduct {A = A} {B} {C} {f} bij)) (fst ,, snd) with Surjection.property (Bijection.surj bij) snd
|
||||
Surjection.property (Bijection.surj (bijectionOnRightOfProduct {A = A} {B} {C} {f} bij)) (fst ,, snd) | b , pr = (fst ,, b) , reflPair refl pr
|
||||
Bijection.inj (bijectionOnRightOfProduct {A = A} {B} {C} {f} bij) {fst ,, snd} {fst₁ ,, snd₁} gx=gy rewrite firstEqualityOfPair gx=gy | Bijection.inj bij (secondEqualityOfPair gx=gy) = refl
|
||||
Bijection.surj (bijectionOnRightOfProduct {A = A} {B} {C} {f} bij) (fst ,, snd) with Bijection.surj bij snd
|
||||
Bijection.surj (bijectionOnRightOfProduct {A = A} {B} {C} {f} bij) (fst ,, snd) | b , pr = (fst ,, b) , reflPair refl pr
|
||||
|
||||
bijectionOnLeftOfProduct : {a b c : _} {A : Set a} {B : Set b} {C : Set c} → {f : A → C} → Bijection f → Bijection (tupleLmap {B = B} f)
|
||||
Injection.property (Bijection.inj (bijectionOnLeftOfProduct {A = A} {B} {C} {f} bij)) {a ,, b} {c ,, d} gx=gy rewrite secondEqualityOfPair gx=gy | Injection.property (Bijection.inj bij) (firstEqualityOfPair gx=gy) = refl
|
||||
Surjection.property (Bijection.surj (bijectionOnLeftOfProduct {A = A} {B} {C} {f} bij)) (fst ,, snd) with Surjection.property (Bijection.surj bij) fst
|
||||
Surjection.property (Bijection.surj (bijectionOnLeftOfProduct {A = A} {B} {C} {f} bij)) (fst ,, snd) | a , b = (a ,, snd) , reflPair b refl
|
||||
Bijection.inj (bijectionOnLeftOfProduct {A = A} {B} {C} {f} bij) {a ,, b} {c ,, d} gx=gy rewrite secondEqualityOfPair gx=gy | Bijection.inj bij (firstEqualityOfPair gx=gy) = refl
|
||||
Bijection.surj (bijectionOnLeftOfProduct {A = A} {B} {C} {f} bij) (fst ,, snd) with Bijection.surj bij fst
|
||||
Bijection.surj (bijectionOnLeftOfProduct {A = A} {B} {C} {f} bij) (fst ,, snd) | a , b = (a ,, snd) , reflPair b refl
|
||||
|
||||
pairProductIsCountable : {a b : _} {A : Set a} {B : Set b} → (X : CountablyInfiniteSet A) → (Y : CountablyInfiniteSet B) → CountablyInfiniteSet (A && B)
|
||||
pairProductIsCountable {A = A} {B = B} X Y = bijectionOfCountableSetIsCountable N^2Countable {(tupleLmap m) ∘ (tupleRmap n)} bijF
|
||||
|
27
Sets/Cardinality/Countable/Definition.agda
Normal file
27
Sets/Cardinality/Countable/Definition.agda
Normal file
@@ -0,0 +1,27 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
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.Countable.Definition 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
|
||||
|
144
Sets/Cardinality/Countable/Lemmas.agda
Normal file
144
Sets/Cardinality/Countable/Lemmas.agda
Normal file
@@ -0,0 +1,144 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Sets.FinSet.Definition
|
||||
open import Semirings.Definition
|
||||
open import Sets.CantorBijection.CantorBijection
|
||||
open import Orders.Total.Definition
|
||||
open import Sets.Cardinality.Countable.Definition
|
||||
|
||||
module Sets.Cardinality.Countable.Lemmas where
|
||||
|
||||
open import Semirings.Lemmas ℕSemiring
|
||||
|
||||
ℕ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) <N 2 *N (_&&_.fst pair) +N 1))
|
||||
sqrtFloor zero = (0 ,, 0) , (refl ,, le zero refl)
|
||||
sqrtFloor (succ n) with sqrtFloor n
|
||||
sqrtFloor (succ n) | (a ,, b) , pr with TotalOrder.totality ℕTotalOrder b (2 *N a)
|
||||
sqrtFloor (succ n) | (a ,, b) , pr | inl (inl x) = (a ,, succ b) , (p ,, q)
|
||||
where
|
||||
p : a *N a +N succ b ≡ succ n
|
||||
p rewrite Semiring.commutative ℕSemiring (a *N a) (succ b) | Semiring.commutative ℕSemiring b (a *N a) = applyEquality succ (_&&_.fst pr)
|
||||
q : succ b <N (a +N (a +N 0)) +N 1
|
||||
q rewrite Semiring.commutative ℕSemiring (a +N (a +N 0)) (succ 0) | Semiring.commutative ℕSemiring a 0 = succPreservesInequality x
|
||||
sqrtFloor (succ n) | (a ,, b) , (_ ,, pr) | inl (inr x) rewrite Semiring.commutative ℕSemiring (a +N (a +N 0)) (succ 0) = exFalso (noIntegersBetweenXAndSuccX (a +N (a +N 0)) x pr)
|
||||
sqrtFloor (succ n) | (a ,, b) , pr | inr x = (succ a ,, 0) , (q ,, succIsPositive _)
|
||||
where
|
||||
p : a +N a *N succ a ≡ n
|
||||
p rewrite x | Semiring.commutative ℕSemiring a 0 | Semiring.commutative ℕSemiring (a +N a) (succ 0) | Semiring.commutative ℕSemiring (a *N a) (succ a +N a) | multiplicationNIsCommutative a (succ a) | Semiring.commutative ℕSemiring (a *N a) (a +N a) | Semiring.+Associative ℕSemiring a a (a *N a) = _&&_.fst pr
|
||||
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
|
||||
|
||||
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))
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) {inl x} {inl y} pr rewrite Semiring.commutative ℕSemiring (CountablyInfiniteSet.counting X x) 0 | Semiring.commutative ℕSemiring (CountablyInfiniteSet.counting X y) 0 | doubleIsAddTwo (CountablyInfiniteSet.counting X x) | doubleIsAddTwo (CountablyInfiniteSet.counting X y) = applyEquality inl (Bijection.inj (CountablyInfiniteSet.countingIsBij X) inter)
|
||||
where
|
||||
inter : CountablyInfiniteSet.counting X x ≡ CountablyInfiniteSet.counting X y
|
||||
inter = doubleLemma (CountablyInfiniteSet.counting X x) (CountablyInfiniteSet.counting X y) pr
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) {inl x} {inr y} pr = exFalso (evenCannotBeOdd (CountablyInfiniteSet.counting X x) (CountablyInfiniteSet.counting Y y) pr)
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) {inr x} {inl y} pr = exFalso (evenCannotBeOdd (CountablyInfiniteSet.counting X y) (CountablyInfiniteSet.counting Y x) (equalityCommutative pr))
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) {inr x} {inr y} pr = applyEquality inr (Bijection.inj (CountablyInfiniteSet.countingIsBij Y) (doubleLemma (CountablyInfiniteSet.counting Y x) (CountablyInfiniteSet.counting Y y) (succInjective pr) ))
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) b with aMod2 b
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) b | a , inl x with Bijection.surj (CountablyInfiniteSet.countingIsBij X) a
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) b | a , inl x | r , pr = inl r , ans
|
||||
where
|
||||
ans : 2 *N CountablyInfiniteSet.counting X r ≡ b
|
||||
ans rewrite pr = x
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) b | a , inr x with Bijection.surj (CountablyInfiniteSet.countingIsBij Y) a
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (pairUnionIsCountable X Y)) b | a , inr x | r , pr = inr r , ans
|
||||
where
|
||||
ans : succ (2 *N CountablyInfiniteSet.counting Y r) ≡ b
|
||||
ans rewrite pr = x
|
||||
|
||||
firstEqualityOfPair : {a b : _} {A : Set a} {B : Set b} → {x1 x2 : A} → {y1 y2 : B} → (x1 ,, y1) ≡ (x2 ,, y2) → x1 ≡ x2
|
||||
firstEqualityOfPair {x1} {x2} {y1} {y2} refl = refl
|
||||
|
||||
secondEqualityOfPair : {a b : _} {A : Set a} {B : Set b} → {x1 x2 : A} → {y1 y2 : B} → (x1 ,, y1) ≡ (x2 ,, y2) → y1 ≡ y2
|
||||
secondEqualityOfPair {x1} {x2} {y1} {y2} refl = refl
|
||||
|
||||
reflPair : {a b : _} {A : Set a} {B : Set b} {x1 x2 : A} {y1 y2 : B} → (x1 ≡ x2) → (y1 ≡ y2) → (x1 ,, y1) ≡ x2 ,, y2
|
||||
reflPair refl refl = refl
|
||||
|
||||
bijectionOfCountableSetIsCountable : {a b : _} {A : Set a} {B : Set b} → (X : CountablyInfiniteSet A) {f : A → B} → (bij : Bijection f) → CountablyInfiniteSet B
|
||||
CountablyInfiniteSet.counting (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj }) b with surj b
|
||||
CountablyInfiniteSet.counting (bijectionOfCountableSetIsCountable record { counting = counting ; countingIsBij = countingIsBij } {f} record { inj = inj ; surj = surj }) b | a , pr = counting a
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj })) {m} {n} m=n with surj m
|
||||
Bijection.inj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj })) {m} {n} m=n | a , b with surj n
|
||||
... | c , d = transitivity (equalityCommutative b) (transitivity (applyEquality f (Bijection.inj (CountablyInfiniteSet.countingIsBij X) m=n)) d)
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj })) b with Bijection.surj (CountablyInfiniteSet.countingIsBij X) b
|
||||
Bijection.surj (CountablyInfiniteSet.countingIsBij (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj })) b | a , pr = f a , s
|
||||
where
|
||||
s : CountablyInfiniteSet.counting (bijectionOfCountableSetIsCountable X {f} record { inj = inj ; surj = surj }) (f a) ≡ b
|
||||
s with surj (f a)
|
||||
s | t , u = transitivity (applyEquality (CountablyInfiniteSet.counting X) (inj u)) pr
|
||||
|
||||
N^2Countable : CountablyInfiniteSet (ℕ && ℕ)
|
||||
CountablyInfiniteSet.counting N^2Countable x = Invertible.inverse (bijectionImpliesInvertible (cantorBijection)) x
|
||||
CountablyInfiniteSet.countingIsBij N^2Countable = invertibleImpliesBijection (inverseIsInvertible (bijectionImpliesInvertible cantorBijection))
|
||||
|
||||
tupleRmap : {a b c : _} {A : Set a} {B : Set b} {C : Set c} → (f : B → C) → (A && B) → (A && C)
|
||||
tupleRmap f (a ,, b) = (a ,, f b)
|
||||
|
||||
tupleLmap : {a b c : _} {A : Set a} {B : Set b} {C : Set c} → (f : A → C) → (A && B) → (C && B)
|
||||
tupleLmap f (a ,, b) = (f a ,, b)
|
||||
|
||||
bijectionOnRightOfProduct : {a b c : _} {A : Set a} {B : Set b} {C : Set c} → {f : B → C} → Bijection f → Bijection (tupleRmap {A = A} f)
|
||||
Bijection.inj (bijectionOnRightOfProduct {A = A} {B} {C} {f} bij) {fst ,, snd} {fst₁ ,, snd₁} gx=gy rewrite firstEqualityOfPair gx=gy | Bijection.inj bij (secondEqualityOfPair gx=gy) = refl
|
||||
Bijection.surj (bijectionOnRightOfProduct {A = A} {B} {C} {f} bij) (fst ,, snd) with Bijection.surj bij snd
|
||||
Bijection.surj (bijectionOnRightOfProduct {A = A} {B} {C} {f} bij) (fst ,, snd) | b , pr = (fst ,, b) , reflPair refl pr
|
||||
|
||||
bijectionOnLeftOfProduct : {a b c : _} {A : Set a} {B : Set b} {C : Set c} → {f : A → C} → Bijection f → Bijection (tupleLmap {B = B} f)
|
||||
Bijection.inj (bijectionOnLeftOfProduct {A = A} {B} {C} {f} bij) {a ,, b} {c ,, d} gx=gy rewrite secondEqualityOfPair gx=gy | Bijection.inj bij (firstEqualityOfPair gx=gy) = refl
|
||||
Bijection.surj (bijectionOnLeftOfProduct {A = A} {B} {C} {f} bij) (fst ,, snd) with Bijection.surj bij fst
|
||||
Bijection.surj (bijectionOnLeftOfProduct {A = A} {B} {C} {f} bij) (fst ,, snd) | a , b = (a ,, snd) , reflPair b refl
|
||||
|
||||
pairProductIsCountable : {a b : _} {A : Set a} {B : Set b} → (X : CountablyInfiniteSet A) → (Y : CountablyInfiniteSet B) → CountablyInfiniteSet (A && B)
|
||||
pairProductIsCountable {A = A} {B = B} X Y = bijectionOfCountableSetIsCountable N^2Countable {(tupleLmap m) ∘ (tupleRmap n)} bijF
|
||||
where
|
||||
bijM = (bijectionImpliesInvertible (CountablyInfiniteSet.countingIsBij X))
|
||||
bijN = (bijectionImpliesInvertible (CountablyInfiniteSet.countingIsBij Y))
|
||||
m : ℕ → A
|
||||
m = Invertible.inverse bijM
|
||||
n : ℕ → B
|
||||
n = Invertible.inverse bijN
|
||||
bM : Bijection m
|
||||
bM = invertibleImpliesBijection (record { inverse = CountablyInfiniteSet.counting X ; isLeft = Invertible.isRight bijM ; isRight = Invertible.isLeft bijM })
|
||||
bN : Bijection n
|
||||
bN = invertibleImpliesBijection (record { inverse = CountablyInfiniteSet.counting Y ; isLeft = Invertible.isRight bijN ; isRight = Invertible.isLeft bijN })
|
||||
bijF : Bijection ((tupleLmap {A = ℕ} {B = B} {C = A} m) ∘ (tupleRmap {A = ℕ} {B = ℕ} {C = B} n))
|
||||
bijF = bijectionComp {f = tupleRmap {A = ℕ} {B = ℕ} {C = B} n} {g = tupleLmap {A = ℕ} {B = B} {C = A} m} (bijectionOnRightOfProduct (invertibleImpliesBijection (bijectionImpliesInvertible bN))) (bijectionOnLeftOfProduct bM)
|
||||
|
||||
-- injectionToNMeansCountable : {a : _} {A : Set a} {f : A → ℕ} → Injection f → InfiniteSet A → Countable A
|
||||
-- injectionToNMeansCountable {f = f} inj record { isInfinite = isInfinite } = {!!}
|
18
Sets/Cardinality/Finite/Definition.agda
Normal file
18
Sets/Cardinality/Finite/Definition.agda
Normal file
@@ -0,0 +1,18 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.FinSet.Definition
|
||||
|
||||
module Sets.Cardinality.Finite.Definition where
|
||||
|
||||
record FiniteSet {a : _} (A : Set a) : Set a where
|
||||
field
|
||||
size : ℕ
|
||||
mapping : FinSet size → A
|
||||
bij : Bijection mapping
|
||||
|
26
Sets/Cardinality/Finite/Lemmas.agda
Normal file
26
Sets/Cardinality/Finite/Lemmas.agda
Normal file
@@ -0,0 +1,26 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.FinSet.Definition
|
||||
open import Sets.FinSet.Lemmas
|
||||
|
||||
module Sets.Cardinality.Finite.Lemmas where
|
||||
|
||||
finsetInjectIntoℕ : {n : ℕ} → Injection (toNat {n})
|
||||
finsetInjectIntoℕ {zero} {()}
|
||||
finsetInjectIntoℕ {succ n} = ans
|
||||
where
|
||||
ans : {n : ℕ} → {x y : FinSet (succ n)} → (toNat x ≡ toNat y) → x ≡ y
|
||||
ans {zero} {fzero} {fzero} _ = refl
|
||||
ans {zero} {fzero} {fsucc ()}
|
||||
ans {zero} {fsucc ()} {y}
|
||||
ans {succ n} {fzero} {fzero} pr = refl
|
||||
ans {succ n} {fzero} {fsucc y} ()
|
||||
ans {succ n} {fsucc x} {fzero} ()
|
||||
ans {succ n} {fsucc x} {fsucc y} pr with succInjective pr
|
||||
... | pr' = applyEquality fsucc (ans {n} {x} {y} pr')
|
15
Sets/Cardinality/Infinite/Definition.agda
Normal file
15
Sets/Cardinality/Infinite/Definition.agda
Normal file
@@ -0,0 +1,15 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.FinSet.Definition
|
||||
|
||||
module Sets.Cardinality.Infinite.Definition where
|
||||
|
||||
InfiniteSet : {a : _} (A : Set a) → Set a
|
||||
InfiniteSet A = (n : ℕ) → (f : FinSet n → A) → (Bijection f) → False
|
||||
|
32
Sets/Cardinality/Infinite/Examples.agda
Normal file
32
Sets/Cardinality/Infinite/Examples.agda
Normal file
@@ -0,0 +1,32 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.FinSet.Definition
|
||||
open import Sets.FinSet.Lemmas
|
||||
open import Sets.Cardinality.Infinite.Definition
|
||||
open import Sets.Cardinality.Finite.Lemmas
|
||||
|
||||
module Sets.Cardinality.Infinite.Examples where
|
||||
|
||||
ℕIsInfinite : InfiniteSet ℕ
|
||||
ℕIsInfinite n f bij = pigeonhole (le 0 refl) badInj
|
||||
where
|
||||
inv : ℕ → FinSet n
|
||||
inv = Invertible.inverse (bijectionImpliesInvertible bij)
|
||||
invInj : Injection inv
|
||||
invInj = Bijection.inj (invertibleImpliesBijection (inverseIsInvertible (bijectionImpliesInvertible bij)))
|
||||
bumpUp : FinSet n → FinSet (succ n)
|
||||
bumpUp = intoSmaller (le 0 refl)
|
||||
bumpUpInj : Injection bumpUp
|
||||
bumpUpInj = intoSmallerInj (le 0 refl)
|
||||
nextInj : Injection (toNat {succ n})
|
||||
nextInj = finsetInjectIntoℕ {succ n}
|
||||
bad : FinSet (succ n) → FinSet n
|
||||
bad a = (inv (toNat a))
|
||||
badInj : Injection bad
|
||||
badInj = injComp nextInj invInj
|
15
Sets/Cardinality/Infinite/Lemmas.agda
Normal file
15
Sets/Cardinality/Infinite/Lemmas.agda
Normal file
@@ -0,0 +1,15 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.Cardinality.Infinite.Definition
|
||||
open import Sets.FinSet.Definition
|
||||
|
||||
module Sets.Cardinality.Infinite.Lemmas where
|
||||
|
||||
finsetNotInfinite : {n : ℕ} → InfiniteSet (FinSet n) → False
|
||||
finsetNotInfinite {n} isInfinite = isInfinite n id idIsBijective
|
425
Sets/FinSet.agda
425
Sets/FinSet.agda
@@ -1,425 +0,0 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
|
||||
module Sets.FinSet where
|
||||
data FinSet : (n : ℕ) → Set where
|
||||
fzero : {n : ℕ} → FinSet (succ n)
|
||||
fsucc : {n : ℕ} → FinSet n → FinSet (succ n)
|
||||
|
||||
data FinNotEquals : {n : ℕ} → (a b : FinSet (succ n)) → Set where
|
||||
fne2 : (a b : FinSet 2) → ((a ≡ fzero) && (b ≡ fsucc (fzero))) || ((b ≡ fzero) && (a ≡ fsucc (fzero))) → FinNotEquals {1} a b
|
||||
fneN : {n : ℕ} → (a b : FinSet (succ (succ (succ n)))) → (((a ≡ fzero) && (Sg (FinSet (succ (succ n))) (λ c → b ≡ fsucc c))) || ((Sg (FinSet (succ (succ n))) (λ c → a ≡ fsucc c)) && (b ≡ fzero))) || (Sg (FinSet (succ (succ n)) && FinSet (succ (succ n))) (λ t → (a ≡ fsucc (_&&_.fst t)) & (b ≡ fsucc (_&&_.snd t)) & FinNotEquals (_&&_.fst t) (_&&_.snd t))) → FinNotEquals {succ (succ n)} a b
|
||||
|
||||
fsuccInjective : {n : ℕ} → {a b : FinSet n} → fsucc a ≡ fsucc b → a ≡ b
|
||||
fsuccInjective refl = refl
|
||||
|
||||
fneSymmetric : {n : ℕ} → {a b : FinSet (succ n)} → FinNotEquals a b → FinNotEquals b a
|
||||
fneSymmetric {.1} {.a1} {.b1} (fne2 a1 b1 (inl x)) = fne2 b1 a1 (inr x)
|
||||
fneSymmetric {.1} {.a1} {.b1} (fne2 a1 b1 (inr x)) = fne2 b1 a1 (inl x)
|
||||
fneSymmetric {.(succ (succ _))} {.fzero} {b} (fneN .fzero b (inl (inl (refl ,, snd)))) = fneN b fzero (inl (inr (snd ,, refl)))
|
||||
fneSymmetric {.(succ (succ _))} {a} {.fzero} (fneN a .fzero (inl (inr (fst ,, refl)))) = fneN fzero a (inl (inl (refl ,, fst)))
|
||||
fneSymmetric {.(succ (succ _))} {a} {b} (fneN a b (inr ((fst ,, snd) , record { one = o ; two = p ; three = q }))) = fneN b a (inr ((snd ,, fst) , record { one = p ; two = o ; three = fneSymmetric q }))
|
||||
|
||||
underlyingProof : {l m : _} {L : Set l} {pr : L → Set m} → (a : Sg L pr) → pr (underlying a)
|
||||
underlyingProof (a , b) = b
|
||||
|
||||
sgEq : {l m : _} {L : Set l} → {pr : L → Set m} → {a b : Sg L pr} → (underlying a ≡ underlying b) → ({c : L} → (r s : pr c) → r ≡ s) → (a ≡ b)
|
||||
sgEq {l} {m} {L} {prop} {(a , b1)} {(.a , b)} refl pr2 with pr2 {a} b b1
|
||||
sgEq {l} {m} {L} {prop} {(a , b1)} {(.a , .b1)} refl pr2 | refl = refl
|
||||
|
||||
|
||||
finSetNotEquals : {n : ℕ} → {a b : FinSet (succ n)} → (a ≡ b → False) → FinNotEquals {n} a b
|
||||
finSetNotEquals {zero} {fzero} {fzero} pr = exFalso (pr refl)
|
||||
finSetNotEquals {zero} {fzero} {fsucc ()} pr
|
||||
finSetNotEquals {zero} {fsucc ()} {b} pr
|
||||
finSetNotEquals {succ zero} {fzero} {fzero} pr = exFalso (pr refl)
|
||||
finSetNotEquals {succ zero} {fzero} {fsucc fzero} pr = fne2 fzero (fsucc fzero) (inl (refl ,, refl))
|
||||
finSetNotEquals {succ zero} {fzero} {fsucc (fsucc ())} pr
|
||||
finSetNotEquals {succ zero} {fsucc fzero} {fzero} pr = fne2 (fsucc fzero) fzero (inr (refl ,, refl))
|
||||
finSetNotEquals {succ zero} {fsucc fzero} {fsucc fzero} pr = exFalso (pr refl)
|
||||
finSetNotEquals {succ zero} {fsucc fzero} {fsucc (fsucc ())} pr
|
||||
finSetNotEquals {succ zero} {fsucc (fsucc ())}
|
||||
finSetNotEquals {succ (succ n)} {fzero} {fzero} pr = exFalso (pr refl)
|
||||
finSetNotEquals {succ (succ n)} {fzero} {fsucc b} pr = fneN fzero (fsucc b) (inl (inl (refl ,, (b , refl))))
|
||||
finSetNotEquals {succ (succ n)} {fsucc a} {fzero} pr = fneN (fsucc a) fzero (inl (inr ((a , refl) ,, refl)))
|
||||
finSetNotEquals {succ (succ n)} {fsucc a} {fsucc b} pr = fneN (fsucc a) (fsucc b) (inr ans)
|
||||
where
|
||||
q : a ≡ b → False
|
||||
q refl = pr refl
|
||||
t : FinNotEquals {succ n} a b
|
||||
t = finSetNotEquals {succ n} {a} {b} q
|
||||
ans : Sg (FinSet (succ (succ n)) && FinSet (succ (succ n))) (λ x → (fsucc a ≡ fsucc (_&&_.fst x)) & fsucc b ≡ fsucc (_&&_.snd x) & FinNotEquals (_&&_.fst x) (_&&_.snd x))
|
||||
ans with t
|
||||
ans | fne2 .fzero .(fsucc fzero) (inl (refl ,, refl)) = (a ,, b) , record { one = refl ; two = refl ; three = fne2 fzero (fsucc fzero) (inl (refl ,, refl)) }
|
||||
ans | fne2 .(fsucc fzero) .fzero (inr (refl ,, refl)) = (a ,, b) , record { one = refl ; two = refl ; three = fne2 (fsucc fzero) fzero (inr (refl ,, refl)) }
|
||||
ans | fneN a b _ = (a ,, b) , record { one = refl ; two = refl ; three = t }
|
||||
|
||||
fzeroNotFsucc : {n : ℕ} → {a : _} → fzero ≡ fsucc {succ n} a → False
|
||||
fzeroNotFsucc ()
|
||||
|
||||
finSetNotEqualsSame : {n : ℕ} → {a : FinSet (succ n)} → FinNotEquals a a → False
|
||||
finSetNotEqualsSame {.1} {fzero} (fne2 .fzero .fzero (inl (fst ,, ())))
|
||||
finSetNotEqualsSame {.1} {fzero} (fne2 .fzero .fzero (inr (fst ,, ())))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fzero} (fneN .fzero .fzero (inl (inl (refl ,, (a , ())))))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fzero} (fneN .fzero .fzero (inl (inr ((a , ()) ,, refl))))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fzero} (fneN .fzero .fzero (inr ((fst ,, snd) , record { one = () ; two = p ; three = q })))
|
||||
finSetNotEqualsSame {.1} {fsucc a} (fne2 .(fsucc a) .(fsucc a) (inl (() ,, snd)))
|
||||
finSetNotEqualsSame {.1} {fsucc a} (fne2 .(fsucc a) .(fsucc a) (inr (() ,, snd)))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fsucc a} (fneN .(fsucc a) .(fsucc a) (inl (inl (() ,, snd))))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fsucc a} (fneN .(fsucc a) .(fsucc a) (inl (inr (fst ,, ()))))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fsucc a} (fneN .(fsucc a) .(fsucc a) (inr ((.a ,, .a) , record { one = refl ; two = refl ; three = q }))) = finSetNotEqualsSame q
|
||||
|
||||
finNotEqualsFsucc : {n : ℕ} → {a b : FinSet (succ n)} → FinNotEquals (fsucc a) (fsucc b) → FinNotEquals a b
|
||||
finNotEqualsFsucc {.0} {a} {b} (fne2 .(fsucc a) .(fsucc b) (inl (() ,, snd)))
|
||||
finNotEqualsFsucc {.0} {a} {b} (fne2 .(fsucc a) .(fsucc b) (inr (() ,, snd)))
|
||||
finNotEqualsFsucc {n} {a} {b} (fneN .(fsucc a) .(fsucc b) (inl (inl (() ,, snd))))
|
||||
finNotEqualsFsucc {n} {a} {b} (fneN .(fsucc a) .(fsucc b) (inl (inr (fst ,, ()))))
|
||||
finNotEqualsFsucc {n} {a} {b} (fneN .(fsucc a) .(fsucc b) (inr ((fst ,, snd) , prf))) = identityOfIndiscernablesRight FinNotEquals ans (equalityCommutative b=snd)
|
||||
where
|
||||
a=fst : a ≡ fst
|
||||
a=fst = fsuccInjective (_&_&_.one prf)
|
||||
b=snd : b ≡ snd
|
||||
b=snd = fsuccInjective (_&_&_.two prf)
|
||||
ans : FinNotEquals a snd
|
||||
ans = identityOfIndiscernablesLeft FinNotEquals (_&_&_.three prf) (equalityCommutative a=fst)
|
||||
|
||||
finSetNotEquals' : {n : ℕ} → {a b : FinSet (succ n)} → FinNotEquals a b → (a ≡ b → False)
|
||||
finSetNotEquals' {n} {a} {.a} fne refl = finSetNotEqualsSame fne
|
||||
|
||||
finset0Empty : FinSet zero → False
|
||||
finset0Empty ()
|
||||
|
||||
finset1OnlyOne : (a b : FinSet 1) → a ≡ b
|
||||
finset1OnlyOne fzero fzero = refl
|
||||
finset1OnlyOne fzero (fsucc ())
|
||||
finset1OnlyOne (fsucc ()) b
|
||||
|
||||
intoSmaller : {n m : ℕ} → (n <N m) → (FinSet n → FinSet m)
|
||||
intoSmaller {zero} {m} n<m = t
|
||||
where
|
||||
t : FinSet 0 → FinSet m
|
||||
t ()
|
||||
intoSmaller {succ n} {zero} ()
|
||||
intoSmaller {succ n} {succ m} n<m with intoSmaller (canRemoveSuccFrom<N n<m)
|
||||
intoSmaller {succ n} {succ m} n<m | prev = t
|
||||
where
|
||||
t : FinSet (succ n) → FinSet (succ m)
|
||||
t fzero = fzero
|
||||
t (fsucc arg) = fsucc (prev arg)
|
||||
|
||||
intoSmallerInj : {n m : ℕ} → (n<m : n <N m) → Injection (intoSmaller n<m)
|
||||
intoSmallerInj {zero} {zero} (le x ())
|
||||
intoSmallerInj {zero} {succ m} n<m = record { property = inj }
|
||||
where
|
||||
t : FinSet 0 → FinSet (succ m)
|
||||
t ()
|
||||
inj : {x y : FinSet zero} → intoSmaller (succIsPositive m) x ≡ intoSmaller (succIsPositive m) y → x ≡ y
|
||||
inj {()} {y}
|
||||
intoSmallerInj {succ n} {zero} ()
|
||||
intoSmallerInj {succ n} {succ m} n<m with intoSmallerInj (canRemoveSuccFrom<N n<m)
|
||||
intoSmallerInj {succ n} {succ m} n<m | prevInj = inj
|
||||
where
|
||||
inj : Injection (intoSmaller n<m)
|
||||
Injection.property inj {fzero} {fzero} pr = refl
|
||||
Injection.property inj {fzero} {fsucc y} ()
|
||||
Injection.property inj {fsucc x} {fzero} ()
|
||||
Injection.property inj {fsucc x} {fsucc y} pr = applyEquality fsucc (Injection.property prevInj (fsuccInjective pr))
|
||||
|
||||
toNat : {n : ℕ} → (a : FinSet n) → ℕ
|
||||
toNat {.(succ _)} fzero = 0
|
||||
toNat {.(succ _)} (fsucc a) = succ (toNat a)
|
||||
|
||||
toNatSmaller : {n : ℕ} → (a : FinSet n) → toNat a <N n
|
||||
toNatSmaller {zero} ()
|
||||
toNatSmaller {succ n} fzero = succIsPositive n
|
||||
toNatSmaller {succ n} (fsucc a) = succPreservesInequality (toNatSmaller a)
|
||||
|
||||
ofNat : {n : ℕ} → (m : ℕ) → (m <N n) → FinSet n
|
||||
ofNat {zero} zero (le x ())
|
||||
ofNat {succ n} zero m<n = fzero
|
||||
ofNat {zero} (succ m) (le x ())
|
||||
ofNat {succ n} (succ m) m<n = fsucc (ofNat {n} m (canRemoveSuccFrom<N m<n))
|
||||
|
||||
ofNatInjective : {n : ℕ} → (x y : ℕ) → (pr : x <N n) → (pr2 : y <N n) → ofNat x pr ≡ ofNat y pr2 → x ≡ y
|
||||
ofNatInjective {zero} zero zero () y<n pr
|
||||
ofNatInjective {zero} zero (succ y) () y<n pr
|
||||
ofNatInjective {zero} (succ x) zero x<n () pr
|
||||
ofNatInjective {zero} (succ x) (succ y) () y<n pr
|
||||
ofNatInjective {succ n} zero zero x<n y<n pr = refl
|
||||
ofNatInjective {succ n} zero (succ y) x<n y<n ()
|
||||
ofNatInjective {succ n} (succ x) zero x<n y<n ()
|
||||
ofNatInjective {succ n} (succ x) (succ y) x<n y<n pr = applyEquality succ (ofNatInjective x y (canRemoveSuccFrom<N x<n) (canRemoveSuccFrom<N y<n) (fsuccInjective pr))
|
||||
|
||||
toNatOfNat : {n : ℕ} → (a : ℕ) → (a<n : a <N n) → toNat (ofNat a a<n) ≡ a
|
||||
toNatOfNat {zero} zero (le x ())
|
||||
toNatOfNat {zero} (succ a) (le x ())
|
||||
toNatOfNat {succ n} zero a<n = refl
|
||||
toNatOfNat {succ n} (succ a) a<n = applyEquality succ (toNatOfNat a (canRemoveSuccFrom<N a<n))
|
||||
|
||||
intoSmallerLemm : {n m : ℕ} → {n<m : n <N (succ m)} → {m<sm : m <N succ m} → (b : FinSet n) → intoSmaller n<m b ≡ ofNat m m<sm → False
|
||||
intoSmallerLemm {.(succ _)} {m} {n<m} fzero pr with intoSmaller (canRemoveSuccFrom<N n<m)
|
||||
intoSmallerLemm {.(succ _)} {zero} {n<m} fzero refl | bl = zeroNeverGreater (canRemoveSuccFrom<N n<m)
|
||||
intoSmallerLemm {.(succ _)} {succ m} {n<m} fzero () | bl
|
||||
intoSmallerLemm {.(succ _)} {m} {n<m} (fsucc b) pr with inspect (intoSmaller (canRemoveSuccFrom<N n<m))
|
||||
intoSmallerLemm {.(succ _)} {zero} {n<m} (fsucc b) pr | bl with≡ _ = zeroNeverGreater (canRemoveSuccFrom<N n<m)
|
||||
intoSmallerLemm {.(succ _)} {succ m} {n<m} {m<sm} (fsucc b) pr | bl with≡ p = intoSmallerLemm {m<sm = canRemoveSuccFrom<N m<sm} b (fsuccInjective pr)
|
||||
|
||||
intoSmallerNotSurj : {n m : ℕ} → {n<m : n <N m} → Surjection (intoSmaller n<m) → False
|
||||
intoSmallerNotSurj {n} {zero} {le x ()} record { property = property }
|
||||
intoSmallerNotSurj {zero} {succ zero} {n<m} record { property = property } with property fzero
|
||||
... | () , _
|
||||
intoSmallerNotSurj {succ n} {succ zero} {n<m} record { property = property } = zeroNeverGreater (canRemoveSuccFrom<N n<m)
|
||||
intoSmallerNotSurj {0} {succ (succ m)} {n<m} record { property = property } with property fzero
|
||||
... | () , _
|
||||
intoSmallerNotSurj {succ n} {succ (succ m)} {n<m} record { property = property } = problem
|
||||
where
|
||||
notHit : FinSet (succ (succ m))
|
||||
notHit = ofNat (succ m) (le zero refl)
|
||||
hitting : Sg (FinSet (succ n)) (λ i → intoSmaller n<m i ≡ notHit)
|
||||
hitting = property notHit
|
||||
problem : False
|
||||
problem with hitting
|
||||
... | a , pr = intoSmallerLemm {succ n} {succ m} {n<m} {le 0 refl} a pr
|
||||
|
||||
finsetDecidableEquality : {n : ℕ} → (x y : FinSet n) → (x ≡ y) || ((x ≡ y) → False)
|
||||
finsetDecidableEquality fzero fzero = inl refl
|
||||
finsetDecidableEquality fzero (fsucc y) = inr λ ()
|
||||
finsetDecidableEquality (fsucc x) fzero = inr λ ()
|
||||
finsetDecidableEquality (fsucc x) (fsucc y) with finsetDecidableEquality x y
|
||||
finsetDecidableEquality (fsucc x) (fsucc y) | inl pr rewrite pr = inl refl
|
||||
finsetDecidableEquality (fsucc x) (fsucc y) | inr pr = inr (λ f → pr (fsuccInjective f))
|
||||
|
||||
subInjection : {n : ℕ} → {f : FinSet (succ (succ n)) → FinSet (succ (succ n))} → Injection f → (FinSet (succ n) → FinSet (succ n))
|
||||
subInjection {n} {f} inj x with inspect (f (fsucc x))
|
||||
subInjection {f = f} inj x | fzero with≡ _ with inspect (f fzero)
|
||||
subInjection {f = f} inj x | fzero with≡ pr | fzero with≡ pr2 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity pr2 (equalityCommutative pr))))
|
||||
subInjection {f = f} inj x | fzero with≡ _ | fsucc bl with≡ _ = bl
|
||||
subInjection {f = f} inj x | fsucc bl with≡ _ = bl
|
||||
|
||||
subInjIsInjective : {n : ℕ} → {f : FinSet (succ (succ n)) → FinSet (succ (succ n))} → (inj : Injection f) → Injection (subInjection inj)
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr with inspect (f (fsucc x))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fzero with≡ _ with inspect (f (fzero))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fzero with≡ pr1 | fzero with≡ pr2 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity pr2 (equalityCommutative pr1))))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fzero with≡ pr1 | fsucc bl with≡ _ with inspect (f (fsucc y))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fzero with≡ pr1 | fsucc bl with≡ _ | fzero with≡ x₁ with inspect (f (fzero))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fzero with≡ pr1 | fsucc bl with≡ _ | fzero with≡ x1 | fzero with≡ x2 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x2 (equalityCommutative x1))))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} refl | fzero with≡ pr1 | fsucc .bl2 with≡ _ | fzero with≡ x1 | fsucc bl2 with≡ _ = fsuccInjective (Injection.property inj (transitivity pr1 (equalityCommutative x1)))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} refl | fzero with≡ pr1 | fsucc .bl2 with≡ pr2 | fsucc bl2 with≡ pr3 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity pr2 (equalityCommutative pr3))))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fsucc bl with≡ _ with inspect (f (fsucc y))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fsucc bl with≡ _ | fzero with≡ x₁ with inspect (f fzero)
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} pr | fsucc bl with≡ _ | fzero with≡ x1 | fzero with≡ x2 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x2 (equalityCommutative x1))))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} refl | fsucc .bl2 with≡ x1 | fzero with≡ x₁ | fsucc bl2 with≡ x2 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x2 (equalityCommutative x1))))
|
||||
Injection.property (subInjIsInjective {f = f} inj) {x} {y} refl | fsucc .y1 with≡ pr1 | fsucc y1 with≡ pr2 = fsuccInjective (Injection.property inj (transitivity pr1 (equalityCommutative pr2)))
|
||||
|
||||
onepointBij : (f : FinSet 1 → FinSet 1) → Bijection f
|
||||
Injection.property (Bijection.inj (onepointBij f)) {x} {y} _ = finset1OnlyOne x y
|
||||
Surjection.property (Bijection.surj (onepointBij f)) fzero with inspect (f fzero)
|
||||
... | fzero with≡ pr = fzero , pr
|
||||
... | fsucc () with≡ _
|
||||
Surjection.property (Bijection.surj (onepointBij f)) (fsucc ())
|
||||
|
||||
nopointBij : (f : FinSet 0 → FinSet 0) → Bijection f
|
||||
Injection.property (Bijection.inj (nopointBij f)) {()}
|
||||
Surjection.property (Bijection.surj (nopointBij f)) ()
|
||||
|
||||
flip : {n : ℕ} → (f : FinSet (succ n) → FinSet (succ n)) → FinSet (succ n) → FinSet (succ n)
|
||||
flip {n} f fzero = f fzero
|
||||
flip {n} f (fsucc a) with inspect (f fzero)
|
||||
flip {n} f (fsucc a) | fzero with≡ x = fsucc a
|
||||
flip {n} f (fsucc a) | fsucc y with≡ x with finsetDecidableEquality a y
|
||||
flip {n} f (fsucc a) | fsucc y with≡ x | inl a=y = fzero
|
||||
flip {n} f (fsucc a) | fsucc y with≡ x | inr a!=y = fsucc a
|
||||
|
||||
flipSwapsF0 : {n : ℕ} → {f : FinSet (succ n) → FinSet (succ n)} → (flip f (f fzero)) ≡ fzero
|
||||
flipSwapsF0 {zero} {f} with inspect (f fzero)
|
||||
flipSwapsF0 {zero} {f} | fzero with≡ x rewrite x = x
|
||||
flipSwapsF0 {zero} {f} | fsucc () with≡ x
|
||||
flipSwapsF0 {succ n} {f = f} with inspect (f fzero)
|
||||
flipSwapsF0 {succ n} {f = f} | fzero with≡ pr rewrite pr = pr
|
||||
flipSwapsF0 {succ n} {f = f} | fsucc b with≡ pr rewrite pr = ans
|
||||
where
|
||||
ans : flip f (fsucc b) ≡ fzero
|
||||
ans with inspect (f fzero)
|
||||
ans | fzero with≡ x = exFalso (fzeroNotFsucc (transitivity (equalityCommutative x) pr))
|
||||
ans | fsucc y with≡ x with finsetDecidableEquality b y
|
||||
ans | fsucc y with≡ x | inl b=y = refl
|
||||
ans | fsucc y with≡ x | inr b!=y = exFalso (b!=y (fsuccInjective (transitivity (equalityCommutative pr) x)))
|
||||
|
||||
flipSwapsZero : {n : ℕ} → {f : FinSet (succ n) → FinSet (succ n)} → (flip f fzero) ≡ f fzero
|
||||
flipSwapsZero {n} {f} = refl
|
||||
|
||||
flipMaintainsEverythingElse : {n : ℕ} → {f : FinSet (succ n) → FinSet (succ n)} → (x : FinSet n) → ((fsucc x ≡ f fzero) → False) → flip f (fsucc x) ≡ fsucc x
|
||||
flipMaintainsEverythingElse {n} {f} x x!=f0 with inspect (f fzero)
|
||||
flipMaintainsEverythingElse {n} {f} x x!=f0 | fzero with≡ pr = refl
|
||||
flipMaintainsEverythingElse {n} {f} x x!=f0 | fsucc bl with≡ pr with finsetDecidableEquality x bl
|
||||
flipMaintainsEverythingElse {n} {f} x x!=f0 | fsucc bl with≡ pr | inl x=bl = exFalso (x!=f0 (transitivity (applyEquality fsucc x=bl) (equalityCommutative pr)))
|
||||
flipMaintainsEverythingElse {n} {f} x x!=f0 | fsucc bl with≡ pr | inr x!=bl = refl
|
||||
|
||||
bijFlip : {n : ℕ} → {f : FinSet (succ n) → FinSet (succ n)} → Bijection (flip f)
|
||||
bijFlip {zero} {f} = onepointBij (flip f)
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fzero} flx=fly = refl
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fsucc y} flx=fly with inspect (f fzero)
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fsucc y} flx=fly | fzero with≡ x rewrite x = exFalso (fzeroNotFsucc flx=fly)
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fsucc y} flx=fly | fsucc f0 with≡ x with finsetDecidableEquality y f0
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fsucc y} flx=fly | fsucc f0 with≡ x | inl x₁ = exFalso (fzeroNotFsucc (transitivity (equalityCommutative flx=fly) x))
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fzero} {fsucc y} flx=fly | fsucc f0 with≡ x | inr pr = exFalso (pr (fsuccInjective (transitivity (equalityCommutative flx=fly) x)))
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fzero} flx=fly with inspect (f fzero)
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fzero} flx=fly | fzero with≡ pr = transitivity flx=fly pr
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fzero} flx=fly | fsucc f0 with≡ x₁ with finsetDecidableEquality x f0
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fzero} flx=fly | fsucc f0 with≡ pr | inl x₂ = exFalso (fzeroNotFsucc (transitivity flx=fly pr))
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fzero} flx=fly | fsucc f0 with≡ x1 | inr x!=f0 = exFalso (x!=f0 (fsuccInjective (transitivity flx=fly x1)))
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly with inspect (f fzero)
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fzero with≡ x₁ = flx=fly
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fsucc f0 with≡ x₁ with finsetDecidableEquality y f0
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fsucc f0 with≡ x₁ | inl y=f0 with finsetDecidableEquality x f0
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fsucc f0 with≡ x₁ | inl y=f0 | inl x=f0 = applyEquality fsucc (transitivity x=f0 (equalityCommutative y=f0))
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} () | fsucc f0 with≡ x₁ | inl y=f0 | inr x!=f0
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fsucc f0 with≡ x₁ | inr y!=f0 with finsetDecidableEquality x f0
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} () | fsucc f0 with≡ x₁ | inr y!=f0 | inl x=f0
|
||||
Injection.property (Bijection.inj (bijFlip {succ n} {f})) {fsucc x} {fsucc y} flx=fly | fsucc f0 with≡ x₁ | inr y!=f0 | inr x!=f0 = flx=fly
|
||||
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) fzero = f fzero , flipSwapsF0 {f = f}
|
||||
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) (fsucc b) with inspect (f fzero)
|
||||
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) (fsucc b) | fzero with≡ x = fsucc b , flipMaintainsEverythingElse {f = f} b λ pr → fzeroNotFsucc (transitivity (equalityCommutative x) (equalityCommutative pr))
|
||||
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) (fsucc b) | fsucc y with≡ x with finsetDecidableEquality y b
|
||||
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) (fsucc b) | fsucc y with≡ x | inl y=b = fzero , transitivity x (applyEquality fsucc y=b)
|
||||
Surjection.property (Bijection.surj (bijFlip {succ n} {f})) (fsucc b) | fsucc y with≡ x | inr y!=b = fsucc b , flipMaintainsEverythingElse {f = f} b λ pr → y!=b (fsuccInjective (transitivity (equalityCommutative x) (equalityCommutative pr)))
|
||||
|
||||
injectionIsBijectionFinset : {n : ℕ} → {f : FinSet n → FinSet n} → Injection f → Bijection f
|
||||
injectionIsBijectionFinset {zero} {f} inj = record { inj = inj ; surj = record { property = λ () } }
|
||||
injectionIsBijectionFinset {succ zero} {f} inj = record { inj = inj ; surj = record { property = ans } }
|
||||
where
|
||||
ans : (b : FinSet (succ zero)) → Sg (FinSet (succ zero)) (λ a → f a ≡ b)
|
||||
ans fzero with inspect (f fzero)
|
||||
ans fzero | fzero with≡ x = fzero , x
|
||||
ans fzero | fsucc () with≡ x
|
||||
ans (fsucc ())
|
||||
injectionIsBijectionFinset {succ (succ n)} {f} inj with injectionIsBijectionFinset {succ n} (subInjIsInjective inj)
|
||||
... | sb = ans
|
||||
where
|
||||
subSurj : Surjection (subInjection inj)
|
||||
subSurj = Bijection.surj sb
|
||||
|
||||
tweakedF : FinSet (succ (succ n)) → FinSet (succ (succ n))
|
||||
tweakedF fzero = fzero
|
||||
tweakedF (fsucc x) = fsucc (subInjection inj x)
|
||||
|
||||
tweakedBij : Bijection tweakedF
|
||||
Injection.property (Bijection.inj tweakedBij) {fzero} {fzero} f'x=f'y = refl
|
||||
Injection.property (Bijection.inj tweakedBij) {fzero} {fsucc y} ()
|
||||
Injection.property (Bijection.inj tweakedBij) {fsucc x} {fzero} ()
|
||||
Injection.property (Bijection.inj tweakedBij) {fsucc x} {fsucc y} f'x=f'y = applyEquality fsucc (Injection.property (subInjIsInjective inj) (fsuccInjective f'x=f'y))
|
||||
Surjection.property (Bijection.surj tweakedBij) fzero = fzero , refl
|
||||
Surjection.property (Bijection.surj tweakedBij) (fsucc b) with Surjection.property subSurj b
|
||||
... | ans , prop = fsucc ans , applyEquality fsucc prop
|
||||
|
||||
compBij : Bijection (λ i → flip f (tweakedF i))
|
||||
compBij = bijectionComp tweakedBij bijFlip
|
||||
|
||||
undoTweakMakesF : {x : FinSet (succ (succ n))} → flip f (tweakedF x) ≡ f x
|
||||
undoTweakMakesF {fzero} = refl
|
||||
undoTweakMakesF {fsucc x} with inspect (f fzero)
|
||||
undoTweakMakesF {fsucc x} | fzero with≡ x₁ with inspect (f (fsucc x))
|
||||
undoTweakMakesF {fsucc x} | fzero with≡ x₁ | fzero with≡ x₂ with inspect (f fzero)
|
||||
undoTweakMakesF {fsucc x} | fzero with≡ x₁ | fzero with≡ x2 | fzero with≡ x3 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x3 (equalityCommutative x2))))
|
||||
undoTweakMakesF {fsucc x} | fzero with≡ x1 | fzero with≡ x2 | fsucc y with≡ x3 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x1 (equalityCommutative x2))))
|
||||
undoTweakMakesF {fsucc x} | fzero with≡ x1 | fsucc y with≡ x2 = equalityCommutative x2
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ with inspect (f (fsucc x))
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fzero with≡ x₂ with inspect (f fzero)
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fzero with≡ x2 | fzero with≡ x3 = exFalso (fzeroNotFsucc (Injection.property inj (transitivity x3 (equalityCommutative x2))))
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fzero with≡ x₂ | fsucc pr with≡ x₃ with finsetDecidableEquality pr y
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fzero with≡ x2 | fsucc pr with≡ x₃ | inl x₄ = equalityCommutative x2
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x1 | fzero with≡ x₂ | fsucc pr with≡ x3 | inr pr!=y = exFalso (pr!=y (fsuccInjective (transitivity (equalityCommutative x3) x1)))
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fsucc thi with≡ x₂ with finsetDecidableEquality thi y
|
||||
undoTweakMakesF {fsucc x} | fsucc .thi with≡ x1 | fsucc thi with≡ x2 | inl refl = exFalso false
|
||||
where
|
||||
p : f fzero ≡ f (fsucc x)
|
||||
p = transitivity x1 (equalityCommutative x2)
|
||||
q : fzero ≡ fsucc x
|
||||
q = Injection.property inj p
|
||||
false : False
|
||||
false = fzeroNotFsucc q
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fsucc thi with≡ x2 | inr thi!=y = equalityCommutative x2
|
||||
|
||||
ans : Bijection f
|
||||
ans = bijectionPreservedUnderExtensionalEq compBij undoTweakMakesF
|
||||
|
||||
pigeonhole : {n m : ℕ} → (m <N n) → {f : FinSet n → FinSet m} → Injection f → False
|
||||
pigeonhole {zero} {zero} () {f} record { property = property }
|
||||
pigeonhole {zero} {succ m} () {f} record { property = property }
|
||||
pigeonhole {succ n} {zero} m<n {f} record { property = property } = finset0Empty (f fzero)
|
||||
pigeonhole {succ zero} {succ m} m<n {f} record { property = property } with canRemoveSuccFrom<N m<n
|
||||
pigeonhole {succ zero} {succ m} m<n {f} record { property = property } | le x ()
|
||||
pigeonhole {succ (succ n)} {succ zero} m<n {f} record { property = property } = fzeroNotFsucc (property (transitivity f0 (equalityCommutative f1)))
|
||||
where
|
||||
f0 : f (fzero) ≡ fzero
|
||||
f0 with inspect (f fzero)
|
||||
... | fzero with≡ x = x
|
||||
... | fsucc () with≡ _
|
||||
f1 : f (fsucc fzero) ≡ fzero
|
||||
f1 with inspect (f (fsucc fzero))
|
||||
... | fzero with≡ x = x
|
||||
... | fsucc () with≡ _
|
||||
pigeonhole {succ (succ n)} {succ (succ m)} m<n {f} injF = intoSmallerNotSurj {_}{_} {m<n} surj
|
||||
where
|
||||
inj : Injection ((intoSmaller m<n) ∘ f)
|
||||
inj = injComp injF (intoSmallerInj m<n)
|
||||
bij : Bijection ((intoSmaller m<n) ∘ f)
|
||||
bij = injectionIsBijectionFinset inj
|
||||
surj : Surjection (intoSmaller m<n)
|
||||
surj = compSurjLeftSurj (Bijection.surj bij)
|
||||
|
||||
--surjectionIsBijectionFinset : {n : ℕ} → {f : FinSet n → FinSet n} → Surjection f → Bijection f
|
||||
--surjectionIsBijectionFinset {zero} {f} surj = nopointBij f
|
||||
--surjectionIsBijectionFinset {succ zero} {f} surj = onepointBij f
|
||||
--surjectionIsBijectionFinset {succ (succ n)} {f} record { property = property } = {!!}
|
||||
|
||||
record FiniteSet {a : _} (A : Set a) : Set a where
|
||||
field
|
||||
size : ℕ
|
||||
mapping : FinSet size → A
|
||||
bij : Bijection mapping
|
||||
|
||||
record InfiniteSet {a : _} (A : Set a) : Set a where
|
||||
field
|
||||
isInfinite : (n : ℕ) → (f : FinSet n → A) → (Bijection f) → False
|
||||
|
||||
finsetInjectIntoℕ : {n : ℕ} → Injection (toNat {n})
|
||||
Injection.property (finsetInjectIntoℕ {zero}) {()}
|
||||
finsetInjectIntoℕ {succ n} = record { property = ans }
|
||||
where
|
||||
ans : {n : ℕ} → {x y : FinSet (succ n)} → (toNat x ≡ toNat y) → x ≡ y
|
||||
ans {zero} {fzero} {fzero} _ = refl
|
||||
ans {zero} {fzero} {fsucc ()}
|
||||
ans {zero} {fsucc ()} {y}
|
||||
ans {succ n} {fzero} {fzero} pr = refl
|
||||
ans {succ n} {fzero} {fsucc y} ()
|
||||
ans {succ n} {fsucc x} {fzero} ()
|
||||
ans {succ n} {fsucc x} {fsucc y} pr with succInjective pr
|
||||
... | pr' = applyEquality fsucc (ans {n} {x} {y} pr')
|
||||
|
||||
ℕIsInfinite : InfiniteSet ℕ
|
||||
InfiniteSet.isInfinite ℕIsInfinite n f bij = pigeonhole (le 0 refl) badInj
|
||||
where
|
||||
inv : ℕ → FinSet n
|
||||
inv = Invertible.inverse (bijectionImpliesInvertible bij)
|
||||
invInj : Injection inv
|
||||
invInj = Bijection.inj (invertibleImpliesBijection (inverseIsInvertible (bijectionImpliesInvertible bij)))
|
||||
bumpUp : FinSet n → FinSet (succ n)
|
||||
bumpUp = intoSmaller (le 0 refl)
|
||||
bumpUpInj : Injection bumpUp
|
||||
bumpUpInj = intoSmallerInj (le 0 refl)
|
||||
nextInj : Injection (toNat {succ n})
|
||||
nextInj = finsetInjectIntoℕ {succ n}
|
||||
bad : FinSet (succ n) → FinSet n
|
||||
bad a = (inv (toNat a))
|
||||
badInj : Injection bad
|
||||
badInj = injComp nextInj invInj
|
||||
|
||||
finsetNotInfinite : {n : ℕ} → InfiniteSet (FinSet n) → False
|
||||
finsetNotInfinite {n} record { isInfinite = isInfinite } = isInfinite n id idIsBijective
|
18
Sets/FinSet/Definition.agda
Normal file
18
Sets/FinSet/Definition.agda
Normal file
@@ -0,0 +1,18 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Definition
|
||||
|
||||
module Sets.FinSet.Definition where
|
||||
|
||||
data FinSet : (n : ℕ) → Set where
|
||||
fzero : {n : ℕ} → FinSet (succ n)
|
||||
fsucc : {n : ℕ} → FinSet n → FinSet (succ n)
|
||||
|
||||
fsuccInjective : {n : ℕ} → {a b : FinSet n} → fsucc a ≡ fsucc b → a ≡ b
|
||||
fsuccInjective refl = refl
|
||||
|
||||
data FinNotEquals : {n : ℕ} → (a b : FinSet (succ n)) → Set where
|
||||
fne2 : (a b : FinSet 2) → ((a ≡ fzero) && (b ≡ fsucc (fzero))) || ((b ≡ fzero) && (a ≡ fsucc (fzero))) → FinNotEquals {1} a b
|
||||
fneN : {n : ℕ} → (a b : FinSet (succ (succ (succ n)))) → (((a ≡ fzero) && (Sg (FinSet (succ (succ n))) (λ c → b ≡ fsucc c))) || ((Sg (FinSet (succ (succ n))) (λ c → a ≡ fsucc c)) && (b ≡ fzero))) || (Sg (FinSet (succ (succ n)) && FinSet (succ (succ n))) (λ t → (a ≡ fsucc (_&&_.fst t)) & (b ≡ fsucc (_&&_.snd t)) & FinNotEquals (_&&_.fst t) (_&&_.snd t))) → FinNotEquals {succ (succ n)} a b
|
377
Sets/FinSet/Lemmas.agda
Normal file
377
Sets/FinSet/Lemmas.agda
Normal file
@@ -0,0 +1,377 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.FinSet.Definition
|
||||
|
||||
module Sets.FinSet.Lemmas where
|
||||
|
||||
fneSymmetric : {n : ℕ} → {a b : FinSet (succ n)} → FinNotEquals a b → FinNotEquals b a
|
||||
fneSymmetric {.1} {.a1} {.b1} (fne2 a1 b1 (inl x)) = fne2 b1 a1 (inr x)
|
||||
fneSymmetric {.1} {.a1} {.b1} (fne2 a1 b1 (inr x)) = fne2 b1 a1 (inl x)
|
||||
fneSymmetric {.(succ (succ _))} {.fzero} {b} (fneN .fzero b (inl (inl (refl ,, snd)))) = fneN b fzero (inl (inr (snd ,, refl)))
|
||||
fneSymmetric {.(succ (succ _))} {a} {.fzero} (fneN a .fzero (inl (inr (fst ,, refl)))) = fneN fzero a (inl (inl (refl ,, fst)))
|
||||
fneSymmetric {.(succ (succ _))} {a} {b} (fneN a b (inr ((fst ,, snd) , record { one = o ; two = p ; three = q }))) = fneN b a (inr ((snd ,, fst) , record { one = p ; two = o ; three = fneSymmetric q }))
|
||||
|
||||
private
|
||||
underlyingProof : {l m : _} {L : Set l} {pr : L → Set m} → (a : Sg L pr) → pr (underlying a)
|
||||
underlyingProof (a , b) = b
|
||||
|
||||
finSetNotEquals : {n : ℕ} → {a b : FinSet (succ n)} → (a ≡ b → False) → FinNotEquals {n} a b
|
||||
finSetNotEquals {zero} {fzero} {fzero} pr = exFalso (pr refl)
|
||||
finSetNotEquals {zero} {fzero} {fsucc ()} pr
|
||||
finSetNotEquals {zero} {fsucc ()} {b} pr
|
||||
finSetNotEquals {succ zero} {fzero} {fzero} pr = exFalso (pr refl)
|
||||
finSetNotEquals {succ zero} {fzero} {fsucc fzero} pr = fne2 fzero (fsucc fzero) (inl (refl ,, refl))
|
||||
finSetNotEquals {succ zero} {fzero} {fsucc (fsucc ())} pr
|
||||
finSetNotEquals {succ zero} {fsucc fzero} {fzero} pr = fne2 (fsucc fzero) fzero (inr (refl ,, refl))
|
||||
finSetNotEquals {succ zero} {fsucc fzero} {fsucc fzero} pr = exFalso (pr refl)
|
||||
finSetNotEquals {succ zero} {fsucc fzero} {fsucc (fsucc ())} pr
|
||||
finSetNotEquals {succ zero} {fsucc (fsucc ())}
|
||||
finSetNotEquals {succ (succ n)} {fzero} {fzero} pr = exFalso (pr refl)
|
||||
finSetNotEquals {succ (succ n)} {fzero} {fsucc b} pr = fneN fzero (fsucc b) (inl (inl (refl ,, (b , refl))))
|
||||
finSetNotEquals {succ (succ n)} {fsucc a} {fzero} pr = fneN (fsucc a) fzero (inl (inr ((a , refl) ,, refl)))
|
||||
finSetNotEquals {succ (succ n)} {fsucc a} {fsucc b} pr = fneN (fsucc a) (fsucc b) (inr ans)
|
||||
where
|
||||
q : a ≡ b → False
|
||||
q refl = pr refl
|
||||
t : FinNotEquals {succ n} a b
|
||||
t = finSetNotEquals {succ n} {a} {b} q
|
||||
ans : Sg (FinSet (succ (succ n)) && FinSet (succ (succ n))) (λ x → (fsucc a ≡ fsucc (_&&_.fst x)) & fsucc b ≡ fsucc (_&&_.snd x) & FinNotEquals (_&&_.fst x) (_&&_.snd x))
|
||||
ans with t
|
||||
ans | fne2 .fzero .(fsucc fzero) (inl (refl ,, refl)) = (a ,, b) , record { one = refl ; two = refl ; three = fne2 fzero (fsucc fzero) (inl (refl ,, refl)) }
|
||||
ans | fne2 .(fsucc fzero) .fzero (inr (refl ,, refl)) = (a ,, b) , record { one = refl ; two = refl ; three = fne2 (fsucc fzero) fzero (inr (refl ,, refl)) }
|
||||
ans | fneN a b _ = (a ,, b) , record { one = refl ; two = refl ; three = t }
|
||||
|
||||
fzeroNotFsucc : {n : ℕ} → {a : _} → fzero ≡ fsucc {succ n} a → False
|
||||
fzeroNotFsucc ()
|
||||
|
||||
finSetNotEqualsSame : {n : ℕ} → {a : FinSet (succ n)} → FinNotEquals a a → False
|
||||
finSetNotEqualsSame {.1} {fzero} (fne2 .fzero .fzero (inl (fst ,, ())))
|
||||
finSetNotEqualsSame {.1} {fzero} (fne2 .fzero .fzero (inr (fst ,, ())))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fzero} (fneN .fzero .fzero (inl (inl (refl ,, (a , ())))))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fzero} (fneN .fzero .fzero (inl (inr ((a , ()) ,, refl))))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fzero} (fneN .fzero .fzero (inr ((fst ,, snd) , record { one = () ; two = p ; three = q })))
|
||||
finSetNotEqualsSame {.1} {fsucc a} (fne2 .(fsucc a) .(fsucc a) (inl (() ,, snd)))
|
||||
finSetNotEqualsSame {.1} {fsucc a} (fne2 .(fsucc a) .(fsucc a) (inr (() ,, snd)))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fsucc a} (fneN .(fsucc a) .(fsucc a) (inl (inl (() ,, snd))))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fsucc a} (fneN .(fsucc a) .(fsucc a) (inl (inr (fst ,, ()))))
|
||||
finSetNotEqualsSame {.(succ (succ _))} {fsucc a} (fneN .(fsucc a) .(fsucc a) (inr ((.a ,, .a) , record { one = refl ; two = refl ; three = q }))) = finSetNotEqualsSame q
|
||||
|
||||
finNotEqualsFsucc : {n : ℕ} → {a b : FinSet (succ n)} → FinNotEquals (fsucc a) (fsucc b) → FinNotEquals a b
|
||||
finNotEqualsFsucc {.0} {a} {b} (fne2 .(fsucc a) .(fsucc b) (inl (() ,, snd)))
|
||||
finNotEqualsFsucc {.0} {a} {b} (fne2 .(fsucc a) .(fsucc b) (inr (() ,, snd)))
|
||||
finNotEqualsFsucc {n} {a} {b} (fneN .(fsucc a) .(fsucc b) (inl (inl (() ,, snd))))
|
||||
finNotEqualsFsucc {n} {a} {b} (fneN .(fsucc a) .(fsucc b) (inl (inr (fst ,, ()))))
|
||||
finNotEqualsFsucc {n} {a} {b} (fneN .(fsucc a) .(fsucc b) (inr ((fst ,, snd) , prf))) = identityOfIndiscernablesRight FinNotEquals ans (equalityCommutative b=snd)
|
||||
where
|
||||
a=fst : a ≡ fst
|
||||
a=fst = fsuccInjective (_&_&_.one prf)
|
||||
b=snd : b ≡ snd
|
||||
b=snd = fsuccInjective (_&_&_.two prf)
|
||||
ans : FinNotEquals a snd
|
||||
ans = identityOfIndiscernablesLeft FinNotEquals (_&_&_.three prf) (equalityCommutative a=fst)
|
||||
|
||||
finSetNotEquals' : {n : ℕ} → {a b : FinSet (succ n)} → FinNotEquals a b → (a ≡ b → False)
|
||||
finSetNotEquals' {n} {a} {.a} fne refl = finSetNotEqualsSame fne
|
||||
|
||||
finset0Empty : FinSet zero → False
|
||||
finset0Empty ()
|
||||
|
||||
finset1OnlyOne : (a b : FinSet 1) → a ≡ b
|
||||
finset1OnlyOne fzero fzero = refl
|
||||
finset1OnlyOne fzero (fsucc ())
|
||||
finset1OnlyOne (fsucc ()) b
|
||||
|
||||
intoSmaller : {n m : ℕ} → .(n <N m) → (FinSet n → FinSet m)
|
||||
intoSmaller {zero} {m} n<m = t
|
||||
where
|
||||
t : FinSet 0 → FinSet m
|
||||
t ()
|
||||
intoSmaller {succ n} {zero} ()
|
||||
intoSmaller {succ n} {succ m} n<m with intoSmaller (canRemoveSuccFrom<N n<m)
|
||||
intoSmaller {succ n} {succ m} n<m | prev = t
|
||||
where
|
||||
t : FinSet (succ n) → FinSet (succ m)
|
||||
t fzero = fzero
|
||||
t (fsucc arg) = fsucc (prev arg)
|
||||
|
||||
intoSmallerInj : {n m : ℕ} → (n<m : n <N m) → Injection (intoSmaller n<m)
|
||||
intoSmallerInj {zero} {zero} (le x ())
|
||||
intoSmallerInj {zero} {succ m} n<m = inj
|
||||
where
|
||||
t : FinSet 0 → FinSet (succ m)
|
||||
t ()
|
||||
inj : {x y : FinSet zero} → intoSmaller (succIsPositive m) x ≡ intoSmaller (succIsPositive m) y → x ≡ y
|
||||
inj {()} {y}
|
||||
intoSmallerInj {succ n} {zero} ()
|
||||
intoSmallerInj {succ n} {succ m} n<m with intoSmallerInj (canRemoveSuccFrom<N n<m)
|
||||
intoSmallerInj {succ n} {succ m} n<m | prevInj = inj
|
||||
where
|
||||
inj : Injection (intoSmaller n<m)
|
||||
inj {fzero} {fzero} pr = refl
|
||||
inj {fzero} {fsucc y} ()
|
||||
inj {fsucc x} {fzero} ()
|
||||
inj {fsucc x} {fsucc y} pr = applyEquality fsucc (prevInj (fsuccInjective pr))
|
||||
|
||||
toNat : {n : ℕ} → (a : FinSet n) → ℕ
|
||||
toNat {.(succ _)} fzero = 0
|
||||
toNat {.(succ _)} (fsucc a) = succ (toNat a)
|
||||
|
||||
toNatSmaller : {n : ℕ} → (a : FinSet n) → toNat a <N n
|
||||
toNatSmaller {zero} ()
|
||||
toNatSmaller {succ n} fzero = succIsPositive n
|
||||
toNatSmaller {succ n} (fsucc a) = succPreservesInequality (toNatSmaller a)
|
||||
|
||||
ofNat : {n : ℕ} → (m : ℕ) → .(m <N n) → FinSet n
|
||||
ofNat {zero} zero ()
|
||||
ofNat {succ n} zero m<n = fzero
|
||||
ofNat {zero} (succ m) ()
|
||||
ofNat {succ n} (succ m) m<n = fsucc (ofNat {n} m (canRemoveSuccFrom<N m<n))
|
||||
|
||||
ofNatInjective : {n : ℕ} → (x y : ℕ) → .(pr : x <N n) → .(pr2 : y <N n) → ofNat x pr ≡ ofNat y pr2 → x ≡ y
|
||||
ofNatInjective {zero} zero zero () y<n pr
|
||||
ofNatInjective {zero} zero (succ y) () y<n pr
|
||||
ofNatInjective {zero} (succ x) zero x<n () pr
|
||||
ofNatInjective {zero} (succ x) (succ y) () y<n pr
|
||||
ofNatInjective {succ n} zero zero x<n y<n pr = refl
|
||||
ofNatInjective {succ n} zero (succ y) x<n y<n ()
|
||||
ofNatInjective {succ n} (succ x) zero x<n y<n ()
|
||||
ofNatInjective {succ n} (succ x) (succ y) x<n y<n pr = applyEquality succ (ofNatInjective x y (canRemoveSuccFrom<N x<n) (canRemoveSuccFrom<N y<n) (fsuccInjective pr))
|
||||
|
||||
toNatInjective : {n : ℕ} → (x y : FinSet n) → toNat x ≡ toNat y → x ≡ y
|
||||
toNatInjective fzero fzero pr = refl
|
||||
toNatInjective (fsucc x) (fsucc y) pr = applyEquality fsucc (toNatInjective x y (succInjective pr))
|
||||
|
||||
toNatOfNat : {n : ℕ} → (a : ℕ) → (a<n : a <N n) → toNat (ofNat a a<n) ≡ a
|
||||
toNatOfNat {zero} zero (le x ())
|
||||
toNatOfNat {zero} (succ a) (le x ())
|
||||
toNatOfNat {succ n} zero a<n = refl
|
||||
toNatOfNat {succ n} (succ a) a<n = applyEquality succ (toNatOfNat a (canRemoveSuccFrom<N a<n))
|
||||
|
||||
intoSmallerLemm : {n m : ℕ} → {n<m : n <N (succ m)} → {m<sm : m <N succ m} → (b : FinSet n) → intoSmaller n<m b ≡ ofNat m m<sm → False
|
||||
intoSmallerLemm {.(succ _)} {m} {n<m} fzero pr with intoSmaller (canRemoveSuccFrom<N n<m)
|
||||
intoSmallerLemm {.(succ _)} {zero} {n<m} fzero refl | bl = zeroNeverGreater (canRemoveSuccFrom<N n<m)
|
||||
intoSmallerLemm {.(succ _)} {succ m} {n<m} fzero () | bl
|
||||
intoSmallerLemm {.(succ _)} {m} {n<m} (fsucc b) pr with inspect (intoSmaller (canRemoveSuccFrom<N n<m))
|
||||
intoSmallerLemm {.(succ _)} {zero} {n<m} (fsucc b) pr | bl with≡ _ = zeroNeverGreater (canRemoveSuccFrom<N n<m)
|
||||
intoSmallerLemm {.(succ _)} {succ m} {n<m} {m<sm} (fsucc b) pr | bl with≡ p = intoSmallerLemm {n<m = canRemoveSuccFrom<N n<m} {m<sm = canRemoveSuccFrom<N m<sm} b (fsuccInjective pr)
|
||||
|
||||
intoSmallerNotSurj : {n m : ℕ} → {n<m : n <N m} → Surjection (intoSmaller n<m) → False
|
||||
intoSmallerNotSurj {n} {zero} {le x ()} property
|
||||
intoSmallerNotSurj {zero} {succ zero} {n<m} property with property fzero
|
||||
... | () , _
|
||||
intoSmallerNotSurj {succ n} {succ zero} {n<m} property = zeroNeverGreater (canRemoveSuccFrom<N n<m)
|
||||
intoSmallerNotSurj {0} {succ (succ m)} {n<m} property with property fzero
|
||||
... | () , _
|
||||
intoSmallerNotSurj {succ n} {succ (succ m)} {n<m} property = problem
|
||||
where
|
||||
notHit : FinSet (succ (succ m))
|
||||
notHit = ofNat (succ m) (le zero refl)
|
||||
hitting : Sg (FinSet (succ n)) (λ i → intoSmaller n<m i ≡ notHit)
|
||||
hitting = property notHit
|
||||
problem : False
|
||||
problem with hitting
|
||||
... | a , pr = intoSmallerLemm {succ n} {succ m} {n<m} {le 0 refl} a pr
|
||||
|
||||
finsetDecidableEquality : {n : ℕ} → (x y : FinSet n) → (x ≡ y) || ((x ≡ y) → False)
|
||||
finsetDecidableEquality fzero fzero = inl refl
|
||||
finsetDecidableEquality fzero (fsucc y) = inr λ ()
|
||||
finsetDecidableEquality (fsucc x) fzero = inr λ ()
|
||||
finsetDecidableEquality (fsucc x) (fsucc y) with finsetDecidableEquality x y
|
||||
finsetDecidableEquality (fsucc x) (fsucc y) | inl pr rewrite pr = inl refl
|
||||
finsetDecidableEquality (fsucc x) (fsucc y) | inr pr = inr (λ f → pr (fsuccInjective f))
|
||||
|
||||
subInjection : {n : ℕ} → {f : FinSet (succ (succ n)) → FinSet (succ (succ n))} → Injection f → (FinSet (succ n) → FinSet (succ n))
|
||||
subInjection {n} {f} inj x with inspect (f (fsucc x))
|
||||
subInjection {f = f} inj x | fzero with≡ _ with inspect (f fzero)
|
||||
subInjection {f = f} inj x | fzero with≡ pr | fzero with≡ pr2 = exFalso (fzeroNotFsucc (inj (transitivity pr2 (equalityCommutative pr))))
|
||||
subInjection {f = f} inj x | fzero with≡ _ | fsucc bl with≡ _ = bl
|
||||
subInjection {f = f} inj x | fsucc bl with≡ _ = bl
|
||||
|
||||
subInjIsInjective : {n : ℕ} → {f : FinSet (succ (succ n)) → FinSet (succ (succ n))} → (inj : Injection f) → Injection (subInjection inj)
|
||||
subInjIsInjective {f = f} inj {x} {y} pr with inspect (f (fsucc x))
|
||||
subInjIsInjective {f = f} inj {x} {y} pr | fzero with≡ _ with inspect (f (fzero))
|
||||
subInjIsInjective {f = f} inj {x} {y} pr | fzero with≡ pr1 | fzero with≡ pr2 = exFalso (fzeroNotFsucc (inj (transitivity pr2 (equalityCommutative pr1))))
|
||||
subInjIsInjective {f = f} inj {x} {y} pr | fzero with≡ pr1 | fsucc bl with≡ _ with inspect (f (fsucc y))
|
||||
subInjIsInjective {f = f} inj {x} {y} pr | fzero with≡ pr1 | fsucc bl with≡ _ | fzero with≡ x₁ with inspect (f (fzero))
|
||||
subInjIsInjective {f = f} inj {x} {y} pr | fzero with≡ pr1 | fsucc bl with≡ _ | fzero with≡ x1 | fzero with≡ x2 = exFalso (fzeroNotFsucc (inj (transitivity x2 (equalityCommutative x1))))
|
||||
subInjIsInjective {f = f} inj {x} {y} refl | fzero with≡ pr1 | fsucc .bl2 with≡ _ | fzero with≡ x1 | fsucc bl2 with≡ _ = fsuccInjective (inj (transitivity pr1 (equalityCommutative x1)))
|
||||
subInjIsInjective {f = f} inj {x} {y} refl | fzero with≡ pr1 | fsucc .bl2 with≡ pr2 | fsucc bl2 with≡ pr3 = exFalso (fzeroNotFsucc (inj (transitivity pr2 (equalityCommutative pr3))))
|
||||
subInjIsInjective {f = f} inj {x} {y} pr | fsucc bl with≡ _ with inspect (f (fsucc y))
|
||||
subInjIsInjective {f = f} inj {x} {y} pr | fsucc bl with≡ _ | fzero with≡ x₁ with inspect (f fzero)
|
||||
subInjIsInjective {f = f} inj {x} {y} pr | fsucc bl with≡ _ | fzero with≡ x1 | fzero with≡ x2 = exFalso (fzeroNotFsucc (inj (transitivity x2 (equalityCommutative x1))))
|
||||
subInjIsInjective {f = f} inj {x} {y} refl | fsucc .bl2 with≡ x1 | fzero with≡ x₁ | fsucc bl2 with≡ x2 = exFalso (fzeroNotFsucc (inj (transitivity x2 (equalityCommutative x1))))
|
||||
subInjIsInjective {f = f} inj {x} {y} refl | fsucc .y1 with≡ pr1 | fsucc y1 with≡ pr2 = fsuccInjective (inj (transitivity pr1 (equalityCommutative pr2)))
|
||||
|
||||
onepointBij : (f : FinSet 1 → FinSet 1) → Bijection f
|
||||
Bijection.inj (onepointBij f) {x} {y} _ = finset1OnlyOne x y
|
||||
Bijection.surj (onepointBij f) fzero with inspect (f fzero)
|
||||
... | fzero with≡ pr = fzero , pr
|
||||
... | fsucc () with≡ _
|
||||
Bijection.surj (onepointBij f) (fsucc ())
|
||||
|
||||
nopointBij : (f : FinSet 0 → FinSet 0) → Bijection f
|
||||
Bijection.inj (nopointBij f) {()}
|
||||
Bijection.surj (nopointBij f) ()
|
||||
|
||||
private
|
||||
flip : {n : ℕ} → (f : FinSet (succ n) → FinSet (succ n)) → FinSet (succ n) → FinSet (succ n)
|
||||
flip {n} f fzero = f fzero
|
||||
flip {n} f (fsucc a) with inspect (f fzero)
|
||||
flip {n} f (fsucc a) | fzero with≡ x = fsucc a
|
||||
flip {n} f (fsucc a) | fsucc y with≡ x with finsetDecidableEquality a y
|
||||
flip {n} f (fsucc a) | fsucc y with≡ x | inl a=y = fzero
|
||||
flip {n} f (fsucc a) | fsucc y with≡ x | inr a!=y = fsucc a
|
||||
|
||||
flipSwapsF0 : {n : ℕ} → {f : FinSet (succ n) → FinSet (succ n)} → (flip f (f fzero)) ≡ fzero
|
||||
flipSwapsF0 {zero} {f} with inspect (f fzero)
|
||||
flipSwapsF0 {zero} {f} | fzero with≡ x rewrite x = x
|
||||
flipSwapsF0 {zero} {f} | fsucc () with≡ x
|
||||
flipSwapsF0 {succ n} {f = f} with inspect (f fzero)
|
||||
flipSwapsF0 {succ n} {f = f} | fzero with≡ pr rewrite pr = pr
|
||||
flipSwapsF0 {succ n} {f = f} | fsucc b with≡ pr rewrite pr = ans
|
||||
where
|
||||
ans : flip f (fsucc b) ≡ fzero
|
||||
ans with inspect (f fzero)
|
||||
ans | fzero with≡ x = exFalso (fzeroNotFsucc (transitivity (equalityCommutative x) pr))
|
||||
ans | fsucc y with≡ x with finsetDecidableEquality b y
|
||||
ans | fsucc y with≡ x | inl b=y = refl
|
||||
ans | fsucc y with≡ x | inr b!=y = exFalso (b!=y (fsuccInjective (transitivity (equalityCommutative pr) x)))
|
||||
|
||||
flipSwapsZero : {n : ℕ} → {f : FinSet (succ n) → FinSet (succ n)} → (flip f fzero) ≡ f fzero
|
||||
flipSwapsZero {n} {f} = refl
|
||||
|
||||
flipMaintainsEverythingElse : {n : ℕ} → {f : FinSet (succ n) → FinSet (succ n)} → (x : FinSet n) → ((fsucc x ≡ f fzero) → False) → flip f (fsucc x) ≡ fsucc x
|
||||
flipMaintainsEverythingElse {n} {f} x x!=f0 with inspect (f fzero)
|
||||
flipMaintainsEverythingElse {n} {f} x x!=f0 | fzero with≡ pr = refl
|
||||
flipMaintainsEverythingElse {n} {f} x x!=f0 | fsucc bl with≡ pr with finsetDecidableEquality x bl
|
||||
flipMaintainsEverythingElse {n} {f} x x!=f0 | fsucc bl with≡ pr | inl x=bl = exFalso (x!=f0 (transitivity (applyEquality fsucc x=bl) (equalityCommutative pr)))
|
||||
flipMaintainsEverythingElse {n} {f} x x!=f0 | fsucc bl with≡ pr | inr x!=bl = refl
|
||||
|
||||
bijFlip : {n : ℕ} → {f : FinSet (succ n) → FinSet (succ n)} → Bijection (flip f)
|
||||
bijFlip {zero} {f} = onepointBij (flip f)
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fzero} {fzero} flx=fly = refl
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fzero} {fsucc y} flx=fly with inspect (f fzero)
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fzero} {fsucc y} flx=fly | fzero with≡ x rewrite x = exFalso (fzeroNotFsucc flx=fly)
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fzero} {fsucc y} flx=fly | fsucc f0 with≡ x with finsetDecidableEquality y f0
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fzero} {fsucc y} flx=fly | fsucc f0 with≡ x | inl x₁ = exFalso (fzeroNotFsucc (transitivity (equalityCommutative flx=fly) x))
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fzero} {fsucc y} flx=fly | fsucc f0 with≡ x | inr pr = exFalso (pr (fsuccInjective (transitivity (equalityCommutative flx=fly) x)))
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fzero} flx=fly with inspect (f fzero)
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fzero} flx=fly | fzero with≡ pr = transitivity flx=fly pr
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fzero} flx=fly | fsucc f0 with≡ x₁ with finsetDecidableEquality x f0
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fzero} flx=fly | fsucc f0 with≡ pr | inl x₂ = exFalso (fzeroNotFsucc (transitivity flx=fly pr))
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fzero} flx=fly | fsucc f0 with≡ x1 | inr x!=f0 = exFalso (x!=f0 (fsuccInjective (transitivity flx=fly x1)))
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fsucc y} flx=fly with inspect (f fzero)
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fsucc y} flx=fly | fzero with≡ x₁ = flx=fly
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fsucc y} flx=fly | fsucc f0 with≡ x₁ with finsetDecidableEquality y f0
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fsucc y} flx=fly | fsucc f0 with≡ x₁ | inl y=f0 with finsetDecidableEquality x f0
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fsucc y} flx=fly | fsucc f0 with≡ x₁ | inl y=f0 | inl x=f0 = applyEquality fsucc (transitivity x=f0 (equalityCommutative y=f0))
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fsucc y} () | fsucc f0 with≡ x₁ | inl y=f0 | inr x!=f0
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fsucc y} flx=fly | fsucc f0 with≡ x₁ | inr y!=f0 with finsetDecidableEquality x f0
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fsucc y} () | fsucc f0 with≡ x₁ | inr y!=f0 | inl x=f0
|
||||
Bijection.inj (bijFlip {succ n} {f}) {fsucc x} {fsucc y} flx=fly | fsucc f0 with≡ x₁ | inr y!=f0 | inr x!=f0 = flx=fly
|
||||
Bijection.surj (bijFlip {succ n} {f}) fzero = f fzero , flipSwapsF0 {f = f}
|
||||
Bijection.surj (bijFlip {succ n} {f}) (fsucc b) with inspect (f fzero)
|
||||
Bijection.surj (bijFlip {succ n} {f}) (fsucc b) | fzero with≡ x = fsucc b , flipMaintainsEverythingElse {f = f} b λ pr → fzeroNotFsucc (transitivity (equalityCommutative x) (equalityCommutative pr))
|
||||
Bijection.surj (bijFlip {succ n} {f}) (fsucc b) | fsucc y with≡ x with finsetDecidableEquality y b
|
||||
Bijection.surj (bijFlip {succ n} {f}) (fsucc b) | fsucc y with≡ x | inl y=b = fzero , transitivity x (applyEquality fsucc y=b)
|
||||
Bijection.surj (bijFlip {succ n} {f}) (fsucc b) | fsucc y with≡ x | inr y!=b = fsucc b , flipMaintainsEverythingElse {f = f} b λ pr → y!=b (fsuccInjective (transitivity (equalityCommutative x) (equalityCommutative pr)))
|
||||
|
||||
injectionIsBijectionFinset : {n : ℕ} → {f : FinSet n → FinSet n} → Injection f → Bijection f
|
||||
injectionIsBijectionFinset {zero} {f} inj = record { inj = inj ; surj = λ () }
|
||||
injectionIsBijectionFinset {succ zero} {f} inj = record { inj = inj ; surj = ans }
|
||||
where
|
||||
ans : (b : FinSet (succ zero)) → Sg (FinSet (succ zero)) (λ a → f a ≡ b)
|
||||
ans fzero with inspect (f fzero)
|
||||
ans fzero | fzero with≡ x = fzero , x
|
||||
ans fzero | fsucc () with≡ x
|
||||
ans (fsucc ())
|
||||
injectionIsBijectionFinset {succ (succ n)} {f} inj with injectionIsBijectionFinset {succ n} (subInjIsInjective inj)
|
||||
... | sb = ans
|
||||
where
|
||||
subSurj : Surjection (subInjection inj)
|
||||
subSurj = Bijection.surj sb
|
||||
|
||||
tweakedF : FinSet (succ (succ n)) → FinSet (succ (succ n))
|
||||
tweakedF fzero = fzero
|
||||
tweakedF (fsucc x) = fsucc (subInjection inj x)
|
||||
|
||||
tweakedBij : Bijection tweakedF
|
||||
Bijection.inj tweakedBij {fzero} {fzero} f'x=f'y = refl
|
||||
Bijection.inj tweakedBij {fzero} {fsucc y} ()
|
||||
Bijection.inj tweakedBij {fsucc x} {fzero} ()
|
||||
Bijection.inj tweakedBij {fsucc x} {fsucc y} f'x=f'y = applyEquality fsucc ((subInjIsInjective inj) (fsuccInjective f'x=f'y))
|
||||
Bijection.surj tweakedBij fzero = fzero , refl
|
||||
Bijection.surj tweakedBij (fsucc b) with subSurj b
|
||||
... | ans , prop = fsucc ans , applyEquality fsucc prop
|
||||
|
||||
compBij : Bijection (λ i → flip f (tweakedF i))
|
||||
compBij = bijectionComp tweakedBij bijFlip
|
||||
|
||||
undoTweakMakesF : {x : FinSet (succ (succ n))} → flip f (tweakedF x) ≡ f x
|
||||
undoTweakMakesF {fzero} = refl
|
||||
undoTweakMakesF {fsucc x} with inspect (f fzero)
|
||||
undoTweakMakesF {fsucc x} | fzero with≡ x₁ with inspect (f (fsucc x))
|
||||
undoTweakMakesF {fsucc x} | fzero with≡ x₁ | fzero with≡ x₂ with inspect (f fzero)
|
||||
undoTweakMakesF {fsucc x} | fzero with≡ x₁ | fzero with≡ x2 | fzero with≡ x3 = exFalso (fzeroNotFsucc (inj (transitivity x3 (equalityCommutative x2))))
|
||||
undoTweakMakesF {fsucc x} | fzero with≡ x1 | fzero with≡ x2 | fsucc y with≡ x3 = exFalso (fzeroNotFsucc (inj (transitivity x1 (equalityCommutative x2))))
|
||||
undoTweakMakesF {fsucc x} | fzero with≡ x1 | fsucc y with≡ x2 = equalityCommutative x2
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ with inspect (f (fsucc x))
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fzero with≡ x₂ with inspect (f fzero)
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fzero with≡ x2 | fzero with≡ x3 = exFalso (fzeroNotFsucc (inj (transitivity x3 (equalityCommutative x2))))
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fzero with≡ x₂ | fsucc pr with≡ x₃ with finsetDecidableEquality pr y
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fzero with≡ x2 | fsucc pr with≡ x₃ | inl x₄ = equalityCommutative x2
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x1 | fzero with≡ x₂ | fsucc pr with≡ x3 | inr pr!=y = exFalso (pr!=y (fsuccInjective (transitivity (equalityCommutative x3) x1)))
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fsucc thi with≡ x₂ with finsetDecidableEquality thi y
|
||||
undoTweakMakesF {fsucc x} | fsucc .thi with≡ x1 | fsucc thi with≡ x2 | inl refl = exFalso false
|
||||
where
|
||||
p : f fzero ≡ f (fsucc x)
|
||||
p = transitivity x1 (equalityCommutative x2)
|
||||
q : fzero ≡ fsucc x
|
||||
q = inj p
|
||||
false : False
|
||||
false = fzeroNotFsucc q
|
||||
undoTweakMakesF {fsucc x} | fsucc y with≡ x₁ | fsucc thi with≡ x2 | inr thi!=y = equalityCommutative x2
|
||||
|
||||
ans : Bijection f
|
||||
ans = bijectionPreservedUnderExtensionalEq compBij undoTweakMakesF
|
||||
|
||||
pigeonhole : {n m : ℕ} → (m <N n) → {f : FinSet n → FinSet m} → Injection f → False
|
||||
pigeonhole {zero} {zero} () {f} property
|
||||
pigeonhole {zero} {succ m} () {f} property
|
||||
pigeonhole {succ n} {zero} m<n {f} property = finset0Empty (f fzero)
|
||||
pigeonhole {succ zero} {succ m} m<n {f} property with canRemoveSuccFrom<N m<n
|
||||
pigeonhole {succ zero} {succ m} m<n {f} property | le x ()
|
||||
pigeonhole {succ (succ n)} {succ zero} m<n {f} property = fzeroNotFsucc (property (transitivity f0 (equalityCommutative f1)))
|
||||
where
|
||||
f0 : f (fzero) ≡ fzero
|
||||
f0 with inspect (f fzero)
|
||||
... | fzero with≡ x = x
|
||||
... | fsucc () with≡ _
|
||||
f1 : f (fsucc fzero) ≡ fzero
|
||||
f1 with inspect (f (fsucc fzero))
|
||||
... | fzero with≡ x = x
|
||||
... | fsucc () with≡ _
|
||||
pigeonhole {succ (succ n)} {succ (succ m)} m<n {f} injF = intoSmallerNotSurj {_}{_} {m<n} surj
|
||||
where
|
||||
inj : Injection ((intoSmaller m<n) ∘ f)
|
||||
inj = injComp injF (intoSmallerInj m<n)
|
||||
bij : Bijection ((intoSmaller m<n) ∘ f)
|
||||
bij = injectionIsBijectionFinset inj
|
||||
surj : Surjection (intoSmaller m<n)
|
||||
surj = compSurjLeftSurj (Bijection.surj bij)
|
||||
|
||||
--surjectionIsBijectionFinset : {n : ℕ} → {f : FinSet n → FinSet n} → Surjection f → Bijection f
|
||||
--surjectionIsBijectionFinset {zero} {f} surj = nopointBij f
|
||||
--surjectionIsBijectionFinset {succ zero} {f} surj = onepointBij f
|
||||
--surjectionIsBijectionFinset {succ (succ n)} {f} record { property = property } = {!!}
|
||||
|
||||
ofNatToNat : {n : ℕ} → (a : FinSet n) → ofNat (toNat a) (toNatSmaller a) ≡ a
|
||||
ofNatToNat {zero} ()
|
||||
ofNatToNat {succ n} fzero = refl
|
||||
ofNatToNat {succ n} (fsucc a) = applyEquality fsucc (ofNatToNat a)
|
@@ -2,12 +2,18 @@
|
||||
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.FinSet
|
||||
open import Sets.FinSet.Definition
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.WithK
|
||||
open import Sets.FinSet.Lemmas
|
||||
|
||||
module Sets.FinSetWithK where
|
||||
|
||||
private
|
||||
sgEq : {l m : _} {L : Set l} → {pr : L → Set m} → {a b : Sg L pr} → (underlying a ≡ underlying b) → ({c : L} → (r s : pr c) → r ≡ s) → (a ≡ b)
|
||||
sgEq {l} {m} {L} {prop} {(a , b1)} {(.a , b)} refl pr2 with pr2 {a} b b1
|
||||
sgEq {l} {m} {L} {prop} {(a , b1)} {(.a , .b1)} refl pr2 | refl = refl
|
||||
|
||||
finNotEqualsRefl : {n : ℕ} {a b : FinSet (succ n)} → (p1 p2 : FinNotEquals a b) → p1 ≡ p2
|
||||
finNotEqualsRefl {.1} {.fzero} {.(fsucc fzero)} (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) = refl
|
||||
finNotEqualsRefl {.1} {.fzero} {.(fsucc fzero)} (fne2 .fzero .(fsucc fzero) (inl (refl ,, refl))) (fne2 .fzero .(fsucc fzero) (inr (() ,, snd)))
|
||||
@@ -75,11 +81,3 @@ finNotEqualsRefl {.(succ (succ _))} {.(fsucc a)} {.(fsucc b)} (fneN (fsucc a) (f
|
||||
r rewrite equalityCommutative t | equalityCommutative t' | equalityCommutative u | equalityCommutative u' = refl
|
||||
ans : fneN (fsucc a) (fsucc b) (inr ((fst ,, snd1) , snd)) ≡ fneN (fsucc a) (fsucc b) (inr ((fst1 ,, snd2) , b1))
|
||||
ans = applyEquality (λ t → fneN (fsucc a) (fsucc b) (inr t)) (sgEq r equality)
|
||||
|
||||
ofNatToNat : {n : ℕ} → (a : FinSet n) → ofNat (toNat a) (toNatSmaller a) ≡ a
|
||||
ofNatToNat {zero} ()
|
||||
ofNatToNat {succ n} fzero = refl
|
||||
ofNatToNat {succ n} (fsucc a) = applyEquality fsucc indHyp
|
||||
where
|
||||
indHyp : ofNat (toNat a) (canRemoveSuccFrom<N (succPreservesInequality (toNatSmaller a))) ≡ a
|
||||
indHyp rewrite <NRefl (canRemoveSuccFrom<N (succPreservesInequality (toNatSmaller a))) (toNatSmaller a) = ofNatToNat a
|
||||
|
Reference in New Issue
Block a user