mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-11 06:38:39 +00:00
Reshuffle in preparation to break the dependency on N's implementation (#75)
This commit is contained in:
@@ -15,7 +15,9 @@ open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
|
||||
module Fields.CauchyCompletion.Addition {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
|
||||
|
@@ -16,8 +16,9 @@ open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Semirings.Definition
|
||||
|
||||
module Fields.CauchyCompletion.Approximation {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
|
@@ -16,7 +16,8 @@ open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Semirings.Definition
|
||||
|
||||
module Fields.CauchyCompletion.Comparison {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
|
@@ -14,7 +14,8 @@ open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
|
||||
module Fields.CauchyCompletion.Definition {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) where
|
||||
|
||||
|
@@ -16,7 +16,8 @@ open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
|
||||
module Fields.CauchyCompletion.Group {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
|
||||
|
@@ -15,7 +15,9 @@ open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Semirings.Definition
|
||||
|
||||
module Fields.CauchyCompletion.Multiplication {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
|
@@ -15,7 +15,9 @@ open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
|
||||
module Fields.CauchyCompletion.Ring {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
|
||||
|
@@ -15,7 +15,8 @@ open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Semirings.Definition
|
||||
|
||||
module Fields.CauchyCompletion.Setoid {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
|
@@ -5,7 +5,7 @@ open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Integers.Addition
|
||||
open import Sets.FinSet
|
||||
|
@@ -5,7 +5,7 @@ open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Integers.Addition
|
||||
open import Sets.FinSet
|
||||
|
@@ -4,7 +4,7 @@ open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
@@ -1,7 +1,7 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals -- for length
|
||||
open import Numbers.Naturals.Semiring -- for length
|
||||
open import Lists.Lists
|
||||
open import Functions
|
||||
open import Groups.SymmetricGroups.Definition
|
||||
|
@@ -5,7 +5,8 @@ open import WithK
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.WithK
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
@@ -13,22 +14,23 @@ open import Groups.SymmetricGroups.Definition
|
||||
open import DecidableSet
|
||||
|
||||
module Groups.FreeGroups where
|
||||
data FreeCompletion {a : _} (A : Set a) : Set a where
|
||||
|
||||
data FreeCompletion {a : _} (A : Set a) : Set a where
|
||||
ofLetter : A → FreeCompletion A
|
||||
ofInv : A → FreeCompletion A
|
||||
|
||||
freeInverse : {a : _} {A : Set a} (l : FreeCompletion A) → FreeCompletion A
|
||||
freeInverse (ofLetter x) = ofInv x
|
||||
freeInverse (ofInv x) = ofLetter x
|
||||
freeInverse : {a : _} {A : Set a} (l : FreeCompletion A) → FreeCompletion A
|
||||
freeInverse (ofLetter x) = ofInv x
|
||||
freeInverse (ofInv x) = ofLetter x
|
||||
|
||||
ofLetterInjective : {a : _} {A : Set a} {x y : A} → (ofLetter x ≡ ofLetter y) → x ≡ y
|
||||
ofLetterInjective refl = refl
|
||||
ofLetterInjective : {a : _} {A : Set a} {x y : A} → (ofLetter x ≡ ofLetter y) → x ≡ y
|
||||
ofLetterInjective refl = refl
|
||||
|
||||
ofInvInjective : {a : _} {A : Set a} {x y : A} → (ofInv x ≡ ofInv y) → x ≡ y
|
||||
ofInvInjective refl = refl
|
||||
ofInvInjective : {a : _} {A : Set a} {x y : A} → (ofInv x ≡ ofInv y) → x ≡ y
|
||||
ofInvInjective refl = refl
|
||||
|
||||
decidableFreeCompletion : {a : _} {A : Set a} → DecidableSet A → DecidableSet (FreeCompletion A)
|
||||
decidableFreeCompletion {A = A} record { eq = dec } = record { eq = pr }
|
||||
decidableFreeCompletion : {a : _} {A : Set a} → DecidableSet A → DecidableSet (FreeCompletion A)
|
||||
decidableFreeCompletion {A = A} record { eq = dec } = record { eq = pr }
|
||||
where
|
||||
pr : (a b : FreeCompletion A) → (a ≡ b) || (a ≡ b → False)
|
||||
pr (ofLetter x) (ofLetter y) with dec x y
|
||||
@@ -40,70 +42,70 @@ module Groups.FreeGroups where
|
||||
... | inl refl = inl refl
|
||||
... | inr x!=y = inr λ p → x!=y (ofInvInjective p)
|
||||
|
||||
freeCompletionEqual : {a : _} {A : Set a} (dec : DecidableSet A) (x y : FreeCompletion A) → Bool
|
||||
freeCompletionEqual dec x y with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqual dec x y | inl x₁ = BoolTrue
|
||||
freeCompletionEqual dec x y | inr x₁ = BoolFalse
|
||||
freeCompletionEqual : {a : _} {A : Set a} (dec : DecidableSet A) (x y : FreeCompletion A) → Bool
|
||||
freeCompletionEqual dec x y with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqual dec x y | inl x₁ = BoolTrue
|
||||
freeCompletionEqual dec x y | inr x₁ = BoolFalse
|
||||
|
||||
freeCompletionEqualFalse : {a : _} {A : Set a} (dec : DecidableSet A) {x y : FreeCompletion A} → ((x ≡ y) → False) → (freeCompletionEqual dec x y) ≡ BoolFalse
|
||||
freeCompletionEqualFalse dec {x = x} {y} x!=y with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqualFalse dec {x} {y} x!=y | inl x=y = exFalso (x!=y x=y)
|
||||
freeCompletionEqualFalse dec {x} {y} x!=y | inr _ = refl
|
||||
freeCompletionEqualFalse : {a : _} {A : Set a} (dec : DecidableSet A) {x y : FreeCompletion A} → ((x ≡ y) → False) → (freeCompletionEqual dec x y) ≡ BoolFalse
|
||||
freeCompletionEqualFalse dec {x = x} {y} x!=y with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqualFalse dec {x} {y} x!=y | inl x=y = exFalso (x!=y x=y)
|
||||
freeCompletionEqualFalse dec {x} {y} x!=y | inr _ = refl
|
||||
|
||||
freeCompletionEqualFalse' : {a : _} {A : Set a} (dec : DecidableSet A) {x y : FreeCompletion A} → (freeCompletionEqual dec x y) ≡ BoolFalse → (x ≡ y) → False
|
||||
freeCompletionEqualFalse' dec {x} {y} pr with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqualFalse' dec {x} {y} () | inl x₁
|
||||
freeCompletionEqualFalse' dec {x} {y} pr | inr ans = ans
|
||||
freeCompletionEqualFalse' : {a : _} {A : Set a} (dec : DecidableSet A) {x y : FreeCompletion A} → (freeCompletionEqual dec x y) ≡ BoolFalse → (x ≡ y) → False
|
||||
freeCompletionEqualFalse' dec {x} {y} pr with DecidableSet.eq (decidableFreeCompletion dec) x y
|
||||
freeCompletionEqualFalse' dec {x} {y} () | inl x₁
|
||||
freeCompletionEqualFalse' dec {x} {y} pr | inr ans = ans
|
||||
|
||||
data ReducedWord {a : _} {A : Set a} (decA : DecidableSet A) : Set a
|
||||
wordLength : {a : _} {A : Set a} {decA : DecidableSet A} → ReducedWord decA → ℕ
|
||||
firstLetter : {a : _} {A : Set a} {decA : DecidableSet A} → (w : ReducedWord decA) → (0 <N wordLength w) → FreeCompletion A
|
||||
data ReducedWord {a : _} {A : Set a} (decA : DecidableSet A) : Set a
|
||||
wordLength : {a : _} {A : Set a} {decA : DecidableSet A} → ReducedWord decA → ℕ
|
||||
firstLetter : {a : _} {A : Set a} {decA : DecidableSet A} → (w : ReducedWord decA) → (0 <N wordLength w) → FreeCompletion A
|
||||
|
||||
data PrependIsValid {a : _} {A : Set a} (decA : DecidableSet A) (w : ReducedWord decA) (l : FreeCompletion A) : Set a where
|
||||
data PrependIsValid {a : _} {A : Set a} (decA : DecidableSet A) (w : ReducedWord decA) (l : FreeCompletion A) : Set a where
|
||||
wordEmpty : wordLength w ≡ 0 → PrependIsValid decA w l
|
||||
wordEnding : (pr : 0 <N wordLength w) → (freeCompletionEqual decA l (freeInverse (firstLetter w pr)) ≡ BoolFalse) → PrependIsValid decA w l
|
||||
|
||||
data ReducedWord {a} {A} decA where
|
||||
data ReducedWord {a} {A} decA where
|
||||
empty : ReducedWord decA
|
||||
prependLetter : (letter : FreeCompletion A) → (w : ReducedWord decA) → PrependIsValid decA w letter → ReducedWord decA
|
||||
|
||||
firstLetter {a} {A} empty (le x ())
|
||||
firstLetter {a} {A} (prependLetter letter w x) pr = letter
|
||||
firstLetter {a} {A} empty (le x ())
|
||||
firstLetter {a} {A} (prependLetter letter w x) pr = letter
|
||||
|
||||
wordLength {a} {A} empty = 0
|
||||
wordLength {a} {A} (prependLetter letter w pr) = succ (wordLength w)
|
||||
wordLength {a} {A} empty = 0
|
||||
wordLength {a} {A} (prependLetter letter w pr) = succ (wordLength w)
|
||||
|
||||
data FreeGroupGenerators {a : _} (A : Set a) : Set a where
|
||||
data FreeGroupGenerators {a : _} (A : Set a) : Set a where
|
||||
χ : A → FreeGroupGenerators A
|
||||
|
||||
prependLetterInjective : {a : _} {A : Set a} {decA : DecidableSet A} {x : FreeCompletion A} {w1 w2 : ReducedWord decA} {pr1 : PrependIsValid decA w1 x} {pr2 : PrependIsValid decA w2 x} → prependLetter x w1 pr1 ≡ prependLetter x w2 pr2 → w1 ≡ w2
|
||||
prependLetterInjective {x = x} {w1} {.w1} {pr1} {.pr1} refl = refl
|
||||
prependLetterInjective : {a : _} {A : Set a} {decA : DecidableSet A} {x : FreeCompletion A} {w1 w2 : ReducedWord decA} {pr1 : PrependIsValid decA w1 x} {pr2 : PrependIsValid decA w2 x} → prependLetter x w1 pr1 ≡ prependLetter x w2 pr2 → w1 ≡ w2
|
||||
prependLetterInjective {x = x} {w1} {.w1} {pr1} {.pr1} refl = refl
|
||||
|
||||
prependLetterInjective' : {a : _} {A : Set a} {decA : DecidableSet A} {x y : FreeCompletion A} {w1 w2 : ReducedWord decA} {pr1 : PrependIsValid decA w1 x} {pr2 : PrependIsValid decA w2 y} → prependLetter x w1 pr1 ≡ prependLetter y w2 pr2 → x ≡ y
|
||||
prependLetterInjective' refl = refl
|
||||
prependLetterInjective' : {a : _} {A : Set a} {decA : DecidableSet A} {x y : FreeCompletion A} {w1 w2 : ReducedWord decA} {pr1 : PrependIsValid decA w1 x} {pr2 : PrependIsValid decA w2 y} → prependLetter x w1 pr1 ≡ prependLetter y w2 pr2 → x ≡ y
|
||||
prependLetterInjective' refl = refl
|
||||
|
||||
prependLetterRefl : {a : _} {A : Set a} {decA : DecidableSet A} {x : FreeCompletion A} {w : ReducedWord decA} → {pr1 pr2 : PrependIsValid decA w x} → prependLetter x w pr1 ≡ prependLetter x w pr2
|
||||
prependLetterRefl {a} {A} {decA} {x} {empty} {wordEmpty refl} {wordEmpty refl} = refl
|
||||
prependLetterRefl {a} {A} {decA} {x} {empty} {wordEmpty refl} {wordEnding (le x₁ ()) x₂}
|
||||
prependLetterRefl {a} {A} {decA} {x} {empty} {wordEnding (le x₂ ()) x₁} {pr2}
|
||||
prependLetterRefl {a} {A} {decA} {x} {prependLetter letter w x₁} {wordEmpty ()} {pr2}
|
||||
prependLetterRefl {a} {A} {decA} {x} {prependLetter letter w x₁} {wordEnding pr x₂} {wordEmpty ()}
|
||||
prependLetterRefl {a} {A} {decA} {x} {prependLetter letter w x₁} {wordEnding pr2 r2} {wordEnding pr1 r1} rewrite <NRefl pr1 (succIsPositive (wordLength w)) | <NRefl pr2 (succIsPositive (wordLength w)) | reflRefl r1 r2 = refl
|
||||
prependLetterRefl : {a : _} {A : Set a} {decA : DecidableSet A} {x : FreeCompletion A} {w : ReducedWord decA} → {pr1 pr2 : PrependIsValid decA w x} → prependLetter x w pr1 ≡ prependLetter x w pr2
|
||||
prependLetterRefl {a} {A} {decA} {x} {empty} {wordEmpty refl} {wordEmpty refl} = refl
|
||||
prependLetterRefl {a} {A} {decA} {x} {empty} {wordEmpty refl} {wordEnding (le x₁ ()) x₂}
|
||||
prependLetterRefl {a} {A} {decA} {x} {empty} {wordEnding (le x₂ ()) x₁} {pr2}
|
||||
prependLetterRefl {a} {A} {decA} {x} {prependLetter letter w x₁} {wordEmpty ()} {pr2}
|
||||
prependLetterRefl {a} {A} {decA} {x} {prependLetter letter w x₁} {wordEnding pr x₂} {wordEmpty ()}
|
||||
prependLetterRefl {a} {A} {decA} {x} {prependLetter letter w x₁} {wordEnding pr2 r2} {wordEnding pr1 r1} rewrite <NRefl pr1 (succIsPositive (wordLength w)) | <NRefl pr2 (succIsPositive (wordLength w)) | reflRefl r1 r2 = refl
|
||||
|
||||
badPrepend : {a : _} {A : Set a} {decA : DecidableSet A} {x : A} {w : ReducedWord decA} {pr : PrependIsValid decA w (ofInv x)} → (PrependIsValid decA (prependLetter (ofInv x) w pr) (ofLetter x)) → False
|
||||
badPrepend (wordEmpty ())
|
||||
badPrepend {decA = decA} {x = x} (wordEnding pr bad) with DecidableSet.eq (decidableFreeCompletion decA) (ofLetter x) (ofLetter x)
|
||||
badPrepend {decA = decA} {x} (wordEnding pr ()) | inl x₁
|
||||
badPrepend {decA = decA} {x} (wordEnding pr bad) | inr pr2 = pr2 refl
|
||||
badPrepend : {a : _} {A : Set a} {decA : DecidableSet A} {x : A} {w : ReducedWord decA} {pr : PrependIsValid decA w (ofInv x)} → (PrependIsValid decA (prependLetter (ofInv x) w pr) (ofLetter x)) → False
|
||||
badPrepend (wordEmpty ())
|
||||
badPrepend {decA = decA} {x = x} (wordEnding pr bad) with DecidableSet.eq (decidableFreeCompletion decA) (ofLetter x) (ofLetter x)
|
||||
badPrepend {decA = decA} {x} (wordEnding pr ()) | inl x₁
|
||||
badPrepend {decA = decA} {x} (wordEnding pr bad) | inr pr2 = pr2 refl
|
||||
|
||||
badPrepend' : {a : _} {A : Set a} {decA : DecidableSet A} {x : A} {w : ReducedWord decA} {pr : PrependIsValid decA w (ofLetter x)} → (PrependIsValid decA (prependLetter (ofLetter x) w pr) (ofInv x)) → False
|
||||
badPrepend' (wordEmpty ())
|
||||
badPrepend' {decA = decA} {x = x} (wordEnding pr x₁) with DecidableSet.eq (decidableFreeCompletion decA) (ofInv x) (ofInv x)
|
||||
badPrepend' {decA = decA} {x} (wordEnding pr ()) | inl x₂
|
||||
badPrepend' {decA = decA} {x} (wordEnding pr x₁) | inr pr2 = pr2 refl
|
||||
badPrepend' : {a : _} {A : Set a} {decA : DecidableSet A} {x : A} {w : ReducedWord decA} {pr : PrependIsValid decA w (ofLetter x)} → (PrependIsValid decA (prependLetter (ofLetter x) w pr) (ofInv x)) → False
|
||||
badPrepend' (wordEmpty ())
|
||||
badPrepend' {decA = decA} {x = x} (wordEnding pr x₁) with DecidableSet.eq (decidableFreeCompletion decA) (ofInv x) (ofInv x)
|
||||
badPrepend' {decA = decA} {x} (wordEnding pr ()) | inl x₂
|
||||
badPrepend' {decA = decA} {x} (wordEnding pr x₁) | inr pr2 = pr2 refl
|
||||
|
||||
freeGroupGenerators : {a : _} (A : Set a) (decA : DecidableSet A) (w : FreeGroupGenerators A) → SymmetryGroupElements (reflSetoid (ReducedWord decA))
|
||||
freeGroupGenerators A decA (χ x) = sym {f = f} bij
|
||||
freeGroupGenerators : {a : _} (A : Set a) (decA : DecidableSet A) (w : FreeGroupGenerators A) → SymmetryGroupElements (reflSetoid (ReducedWord decA))
|
||||
freeGroupGenerators A decA (χ x) = sym {f = f} bij
|
||||
where
|
||||
open DecidableSet.DecidableSet decA
|
||||
f : ReducedWord decA → ReducedWord decA
|
||||
|
@@ -4,7 +4,8 @@ open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Rationals.Definition
|
||||
open import Sets.FinSet
|
||||
|
@@ -6,7 +6,7 @@ open import Maybe
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Vectors
|
||||
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
|
||||
module KeyValue.KeyValue where
|
||||
|
@@ -6,204 +6,205 @@ open import Maybe
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Vectors
|
||||
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
|
||||
module KeyValue.LinearStore.Implementation where
|
||||
record ReducedMap {a b c : _} (keys : Set a) (values : Set b) (keyOrder : TotalOrder {_} {c} keys) (min : keys) : Set (a ⊔ b ⊔ c)
|
||||
record ReducedMap {a} {b} {c} keys values keyOrder min where
|
||||
|
||||
record ReducedMap {a b c : _} (keys : Set a) (values : Set b) (keyOrder : TotalOrder {_} {c} keys) (min : keys) : Set (a ⊔ b ⊔ c)
|
||||
record ReducedMap {a} {b} {c} keys values keyOrder min where
|
||||
inductive
|
||||
field
|
||||
firstEntry : values
|
||||
next : Maybe (Sg keys (λ nextKey → (ReducedMap keys values keyOrder nextKey) && (TotalOrder._<_ keyOrder min nextKey)))
|
||||
|
||||
addReducedMap : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → {min : keys} → (k : keys) → (v : values) → (m : ReducedMap {a} {b} keys values keyOrder min) → ReducedMap keys values keyOrder (TotalOrder.min keyOrder min k)
|
||||
addReducedMap {keyOrder = keyOrder} {min} k v m with TotalOrder.totality keyOrder min k
|
||||
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) = record { firstEntry = firstEntry ; next = yes (k , (record { firstEntry = v ; next = no} ,, min<k))}
|
||||
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = minVal ; next = yes (nextKey , (m ,, pr)) } | inl (inl min<k) = record { firstEntry = minVal ; next = yes ((TotalOrder.min keyOrder nextKey k) , (addReducedMap {keyOrder = keyOrder} {_} k v m ,, minFromBoth {order = keyOrder} pr min<k))}
|
||||
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } | inl (inr k<min) = record { firstEntry = v ; next = yes (min , (record { firstEntry = firstEntry ; next = next } ,, k<min)) }
|
||||
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } | inr min=k rewrite min=k = record { firstEntry = v ; next = next }
|
||||
addReducedMap : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → {min : keys} → (k : keys) → (v : values) → (m : ReducedMap {a} {b} keys values keyOrder min) → ReducedMap keys values keyOrder (TotalOrder.min keyOrder min k)
|
||||
addReducedMap {keyOrder = keyOrder} {min} k v m with TotalOrder.totality keyOrder min k
|
||||
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) = record { firstEntry = firstEntry ; next = yes (k , (record { firstEntry = v ; next = no} ,, min<k))}
|
||||
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = minVal ; next = yes (nextKey , (m ,, pr)) } | inl (inl min<k) = record { firstEntry = minVal ; next = yes ((TotalOrder.min keyOrder nextKey k) , (addReducedMap {keyOrder = keyOrder} {_} k v m ,, minFromBoth {order = keyOrder} pr min<k))}
|
||||
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } | inl (inr k<min) = record { firstEntry = v ; next = yes (min , (record { firstEntry = firstEntry ; next = next } ,, k<min)) }
|
||||
addReducedMap {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } | inr min=k rewrite min=k = record { firstEntry = v ; next = next }
|
||||
|
||||
lookupReduced : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → {min : keys} → (m : ReducedMap keys values keyOrder min) → (target : keys) → Maybe values
|
||||
lookupReduced {keyOrder = keyOrder} {min} m k with TotalOrder.totality keyOrder min k
|
||||
lookupReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = no } k | inl (inl min<k) = no
|
||||
lookupReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, _))) } k | inl (inl min<k) = lookupReduced {keyOrder = keyOrder} {newMin} m k
|
||||
lookupReduced {keyOrder = keyOrder} {min} m k | inl (inr k<min) = no
|
||||
lookupReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = next } k | inr min=k = yes firstEntry
|
||||
lookupReduced : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → {min : keys} → (m : ReducedMap keys values keyOrder min) → (target : keys) → Maybe values
|
||||
lookupReduced {keyOrder = keyOrder} {min} m k with TotalOrder.totality keyOrder min k
|
||||
lookupReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = no } k | inl (inl min<k) = no
|
||||
lookupReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, _))) } k | inl (inl min<k) = lookupReduced {keyOrder = keyOrder} {newMin} m k
|
||||
lookupReduced {keyOrder = keyOrder} {min} m k | inl (inr k<min) = no
|
||||
lookupReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = next } k | inr min=k = yes firstEntry
|
||||
|
||||
countReduced : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} → (m : ReducedMap keys values keyOrder min) → ℕ
|
||||
countReduced record { firstEntry = firstEntry ; next = no } = 1
|
||||
countReduced record { firstEntry = firstEntry ; next = (yes (key , (m ,, pr))) } = succ (countReduced m)
|
||||
countReduced : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} → (m : ReducedMap keys values keyOrder min) → ℕ
|
||||
countReduced record { firstEntry = firstEntry ; next = no } = 1
|
||||
countReduced record { firstEntry = firstEntry ; next = (yes (key , (m ,, pr))) } = succ (countReduced m)
|
||||
|
||||
lookupReducedSucceedsAfterAdd : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → {min : keys} → (k : keys) → (v : values) → (m : ReducedMap keys values keyOrder min) → (lookupReduced (addReducedMap k v m) k ≡ yes v)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) with TotalOrder.totality keyOrder k k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) | inl (inl k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) | inl (inr k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) | inr p = refl
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k k<min))
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inr min=k rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) | inl (inl _) = lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} k v m
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k k<min))
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) | inr min=k rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) with TotalOrder.totality keyOrder k k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) | inl (inl k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) | inl (inr k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) | inr p = refl
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k with TotalOrder.totality keyOrder k k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inl (inl k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inl (inr k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inr p with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inr p | inl (inl min<k) rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inr p | inl (inr k<min) rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {.k} k v record { firstEntry = firstEntry ; next = next } | inr refl | inr p | inr s = refl
|
||||
lookupReducedSucceedsAfterAdd : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → {min : keys} → (k : keys) → (v : values) → (m : ReducedMap keys values keyOrder min) → (lookupReduced (addReducedMap k v m) k ≡ yes v)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) with TotalOrder.totality keyOrder k k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) | inl (inl k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) | inl (inr k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inl _) | inr p = refl
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k k<min))
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } | inl (inl min<k) | inr min=k rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) | inl (inl _) = lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} k v m
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k k<min))
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (newMin , (m ,, pr))) } | inl (inl min<k) | inr min=k rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) with TotalOrder.totality keyOrder k k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) | inl (inl k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) | inl (inr k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inl (inr k<min) | inr p = refl
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k with TotalOrder.totality keyOrder k k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inl (inl k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inl (inr k<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inr p with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inr p | inl (inl min<k) rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {min} k v m | inr min=k | inr p | inl (inr k<min) rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
lookupReducedSucceedsAfterAdd {keyOrder = keyOrder} {.k} k v record { firstEntry = firstEntry ; next = next } | inr refl | inr p | inr s = refl
|
||||
|
||||
lookupReducedSucceedsAfterUnrelatedAdd : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → {min : keys} → (unrelatedK : keys) → (unrelatedV : values) → (k : keys) → (v : values) → ((TotalOrder._<_ keyOrder unrelatedK k) || (TotalOrder._<_ keyOrder k unrelatedK)) → (m : ReducedMap keys values keyOrder min) → (lookupReduced m k ≡ yes v) → lookupReduced (addReducedMap unrelatedK unrelatedV m) k ≡ yes v
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds with TotalOrder.totality keyOrder min k'
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inl min<k') with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr record { firstEntry = firstEntry ; next = no } () | inl (inl min<k') | inl (inl min<k)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedSucceeds | inl (inl min<k') | inl (inl min<k) = lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {a} k' v' k v pr fst lookupReducedSucceeds
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inl x) m lookupReducedSucceeds | inl (inl min<k') | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<min (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k' x)))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr k<k') m () | inl (inl min<k') | inl (inr k<min)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min v (inl x) m lookupReducedSucceeds | inl (inl min<k') | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) x min<k'))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min v (inr x) record { firstEntry = v2 ; next = no } p | inl (inl min<k') | inr refl = applyEquality yes (yesInjective p)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min v (inr x) record { firstEntry = .v ; next = (yes (a , b)) } refl | inl (inl min<k') | inr refl = applyEquality yes refl
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) with TotalOrder.totality keyOrder k' k
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) | inl (inl k'<k) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) | inl (inl k'<k) | inl (inl min<k) = lookupReducedSucceeds
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m () | inl (inr k'<min) | inl (inl k'<k) | inl (inr k<min)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) | inl (inl k'<k) | inr refl = lookupReducedSucceeds
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inl x) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) x k<k'))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr x) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr _) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') | inl (inl x) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) x (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<k' k'<min)))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr _) m () | inl (inr k'<min) | inl (inr k<k') | inl (inr x)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr _) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') | inr x rewrite x = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<k' k'<min))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inl x) m lookupReducedSucceeds | inl (inr k'<min) | inr k'=k rewrite k'=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr x) m lookupReducedSucceeds | inl (inr k'<min) | inr k'=k rewrite k'=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inr refl with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inl _) record { firstEntry = firstEntry ; next = no } () | inr refl | inl (inl min<k)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inl _) record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedSucceeds | inr refl | inl (inl min<k) = lookupReducedSucceeds
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inr x) m lookupReducedSucceeds | inr refl | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k x))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inl x) m () | inr refl | inl (inr k<min)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inr x) m () | inr refl | inl (inr k<min)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} .min v' min v (inl x) m lookupReducedSucceeds | inr refl | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} .min v' min v (inr x) m lookupReducedSucceeds | inr refl | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → {min : keys} → (unrelatedK : keys) → (unrelatedV : values) → (k : keys) → (v : values) → ((TotalOrder._<_ keyOrder unrelatedK k) || (TotalOrder._<_ keyOrder k unrelatedK)) → (m : ReducedMap keys values keyOrder min) → (lookupReduced m k ≡ yes v) → lookupReduced (addReducedMap unrelatedK unrelatedV m) k ≡ yes v
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds with TotalOrder.totality keyOrder min k'
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inl min<k') with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr record { firstEntry = firstEntry ; next = no } () | inl (inl min<k') | inl (inl min<k)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedSucceeds | inl (inl min<k') | inl (inl min<k) = lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {a} k' v' k v pr fst lookupReducedSucceeds
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inl x) m lookupReducedSucceeds | inl (inl min<k') | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<min (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k' x)))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr k<k') m () | inl (inl min<k') | inl (inr k<min)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min v (inl x) m lookupReducedSucceeds | inl (inl min<k') | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) x min<k'))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min v (inr x) record { firstEntry = v2 ; next = no } p | inl (inl min<k') | inr refl = applyEquality yes (yesInjective p)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min v (inr x) record { firstEntry = .v ; next = (yes (a , b)) } refl | inl (inl min<k') | inr refl = applyEquality yes refl
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) with TotalOrder.totality keyOrder k' k
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) | inl (inl k'<k) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) | inl (inl k'<k) | inl (inl min<k) = lookupReducedSucceeds
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m () | inl (inr k'<min) | inl (inl k'<k) | inl (inr k<min)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inl (inr k'<min) | inl (inl k'<k) | inr refl = lookupReducedSucceeds
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inl x) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) x k<k'))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr x) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr _) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') | inl (inl x) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) x (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<k' k'<min)))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr _) m () | inl (inr k'<min) | inl (inr k<k') | inl (inr x)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr _) m lookupReducedSucceeds | inl (inr k'<min) | inl (inr k<k') | inr x rewrite x = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<k' k'<min))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inl x) m lookupReducedSucceeds | inl (inr k'<min) | inr k'=k rewrite k'=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v (inr x) m lookupReducedSucceeds | inl (inr k'<min) | inr k'=k rewrite k'=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k v pr m lookupReducedSucceeds | inr refl with TotalOrder.totality keyOrder min k
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inl _) record { firstEntry = firstEntry ; next = no } () | inr refl | inl (inl min<k)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inl _) record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedSucceeds | inr refl | inl (inl min<k) = lookupReducedSucceeds
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inr x) m lookupReducedSucceeds | inr refl | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k x))
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inl x) m () | inr refl | inl (inr k<min)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {k'} k' v' k v (inr x) m () | inr refl | inl (inr k<min)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} .min v' min v (inl x) m lookupReducedSucceeds | inr refl | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedSucceedsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} .min v' min v (inr x) m lookupReducedSucceeds | inr refl | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
|
||||
lookupReducedFailsAfterUnrelatedAdd : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → {min : keys} → (unrelatedK : keys) → (unrelatedV : values) → (k : keys) → ((TotalOrder._<_ keyOrder unrelatedK k) || (TotalOrder._<_ keyOrder k unrelatedK)) → (m : ReducedMap keys values keyOrder min) → (lookupReduced m k ≡ no) → lookupReduced (addReducedMap unrelatedK unrelatedV m) k ≡ no
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails with TotalOrder.totality keyOrder min k'
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inl min<k') with TotalOrder.totality keyOrder min k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) with TotalOrder.totality keyOrder k' k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inl (inl x) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inl (inr x) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inl x) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inr x) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedFails | inl (inl min<k') | inl (inl min<k) = lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} k' v' k pr fst lookupReducedFails
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inl min<k') | inl (inr k<min) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min pr m () | inl (inl min<k') | inr refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inr k'<min) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) with TotalOrder.totality keyOrder k' k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl x₁) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl x₁) record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = lookupReducedFails
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inr x₁) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inr x₁) record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = lookupReducedFails
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inr x) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inr k<k') = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inl x) record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inr x) record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) with TotalOrder.totality keyOrder k' k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<min min<k))
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) | inl (inr k<min') = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inr x) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k'<k x))
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inr k<k') = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inl x) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inr x) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m () | inl (inr k'<min) | inr min=k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr m lookupReducedFails | inr refl with TotalOrder.totality keyOrder min k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inr refl | inl (inl min<k) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedFails | inr refl | inl (inl min<k) = lookupReducedFails
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr m lookupReducedFails | inr refl | inl (inr k<min) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr m () | inr refl | inr min=k
|
||||
lookupReducedFailsAfterUnrelatedAdd : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → {min : keys} → (unrelatedK : keys) → (unrelatedV : values) → (k : keys) → ((TotalOrder._<_ keyOrder unrelatedK k) || (TotalOrder._<_ keyOrder k unrelatedK)) → (m : ReducedMap keys values keyOrder min) → (lookupReduced m k ≡ no) → lookupReduced (addReducedMap unrelatedK unrelatedV m) k ≡ no
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails with TotalOrder.totality keyOrder min k'
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inl min<k') with TotalOrder.totality keyOrder min k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) with TotalOrder.totality keyOrder k' k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inl (inl x) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inl (inr x) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inl x) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inr x) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inl min<k') | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedFails | inl (inl min<k') | inl (inl min<k) = lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} k' v' k pr fst lookupReducedFails
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inl min<k') | inl (inr k<min) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {.min} k' v' min pr m () | inl (inl min<k') | inr refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inr k'<min) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) with TotalOrder.totality keyOrder k' k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl x₁) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl x₁) record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = lookupReducedFails
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inr x₁) record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inr x₁) record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inl _) = lookupReducedFails
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inl (inr x) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inl k'<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inl (inr k<k') = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inl x) record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inr x) record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr k'<min) | inl (inl min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) with TotalOrder.totality keyOrder k' k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) with TotalOrder.totality keyOrder min k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<min min<k))
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) | inl (inr k<min') = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inl _) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k (inr x) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inl k'<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k'<k x))
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inl (inr k<k') = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inl x) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .k v' k (inr x) m lookupReducedFails | inl (inr k'<min) | inl (inr k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} k' v' k pr m () | inl (inr k'<min) | inr min=k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr m lookupReducedFails | inr refl with TotalOrder.totality keyOrder min k
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr record { firstEntry = firstEntry ; next = no } lookupReducedFails | inr refl | inl (inl min<k) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedFails | inr refl | inl (inl min<k) = lookupReducedFails
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr m lookupReducedFails | inr refl | inl (inr k<min) = refl
|
||||
lookupReducedFailsAfterUnrelatedAdd {keyOrder = keyOrder} {min} .min v' k pr m () | inr refl | inr min=k
|
||||
|
||||
countReducedBehavesWhenAddingNotPresent : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} → (k : keys) → (v : values) → (m : ReducedMap keys values keyOrder min) → (lookupReduced {keyOrder = keyOrder} m k ≡ no) → countReduced (addReducedMap k v m) ≡ succ (countReduced m)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v m lookupReducedFails with TotalOrder.totality keyOrder k min
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k k<min))
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) | inl (inr _) = refl
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr min<k) with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inr min<k) | inl (inl _) = refl
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedFails | inl (inr min<k) | inl (inl _) = applyEquality succ (countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} k v (_&&_.fst b) lookupReducedFails)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr min<k) | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k k<min))
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v m lookupReducedFails | inr refl with TotalOrder.totality keyOrder min min
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v m lookupReducedFails | inr refl | inl (inl min<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<min)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v m lookupReducedFails | inr refl | inl (inr min<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<min)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v record { firstEntry = firstEntry ; next = no } () | inr refl | inr p
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v record { firstEntry = firstEntry ; next = (yes x) } () | inr refl | inr p
|
||||
countReducedBehavesWhenAddingNotPresent : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} → (k : keys) → (v : values) → (m : ReducedMap keys values keyOrder min) → (lookupReduced {keyOrder = keyOrder} m k ≡ no) → countReduced (addReducedMap k v m) ≡ succ (countReduced m)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v m lookupReducedFails with TotalOrder.totality keyOrder k min
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k k<min))
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) | inl (inr _) = refl
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inl k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr min<k) with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = no } lookupReducedFails | inl (inr min<k) | inl (inl _) = refl
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedFails | inl (inr min<k) | inl (inl _) = applyEquality succ (countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} k v (_&&_.fst b) lookupReducedFails)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr min<k) | inl (inr k<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) min<k k<min))
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v record { firstEntry = firstEntry ; next = next } lookupReducedFails | inl (inr min<k) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {min} k v m lookupReducedFails | inr refl with TotalOrder.totality keyOrder min min
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v m lookupReducedFails | inr refl | inl (inl min<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<min)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v m lookupReducedFails | inr refl | inl (inr min<min) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<min)
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v record { firstEntry = firstEntry ; next = no } () | inr refl | inr p
|
||||
countReducedBehavesWhenAddingNotPresent {keyOrder = keyOrder} {k} k v record { firstEntry = firstEntry ; next = (yes x) } () | inr refl | inr p
|
||||
|
||||
countReducedBehavesWhenAddingPresent : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} → (k : keys) → (v v' : values) → (m : ReducedMap keys values keyOrder min) → (lookupReduced {keyOrder = keyOrder} m k ≡ yes v') → countReduced (addReducedMap k v m) ≡ countReduced m
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds with TotalOrder.totality keyOrder k min
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inl k<min) with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inl k<min) | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<min min<k))
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m () | inl (inl k<min) | inl (inr _)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inl k<min) | inr q = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (identityOfIndiscernablesLeft (TotalOrder._<_ keyOrder) k<min (equalityCommutative q)))
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inr min<k) with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = no } () | inl (inr min<k) | inl (inl _)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedSucceeds | inl (inr min<k) | inl (inl _) = applyEquality succ (countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} k v v' fst lookupReducedSucceeds)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m () | inl (inr min<k) | inl (inr k<min)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inr min<k) | inr q = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (identityOfIndiscernablesLeft (λ a b → TotalOrder._<_ keyOrder a b) min<k q))
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inr q with TotalOrder.totality keyOrder min min
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inr q | inl (inl x) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inr q | inl (inr x) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = no } lookupReducedSucceeds | inr _ | inr p with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = no } () | inr _ | inr p | inl (inl x)
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = no } () | inr _ | inr p | inl (inr x`)
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {keyOrder = keyOrder} {.k} k v v' record { firstEntry = firstEntry ; next = no } lookupReducedSucceeds | inr q | inr p | inr refl = refl
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedSucceeds | inr q | inr p with TotalOrder.totality keyOrder k k
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedSucceeds | inr q | inr p | inl (inl x) = exFalso (TotalOrder.irreflexive keyOrder x)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedSucceeds | inr q | inr p | inl (inr x) = exFalso (TotalOrder.irreflexive keyOrder x)
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {values = values} {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , (m ,, pr))) } q | inr q1 | inr p | inr x with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {values} {keyOrder} {.k} k v v' record { firstEntry = firstEntry ; next = (yes (a , (m ,, pr))) } q | inr refl | inr p | inr x | inl (inl min<k) = exFalso (TotalOrder.irreflexive keyOrder min<k)
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {values} {keyOrder} {.k} k v v' record { firstEntry = firstEntry ; next = (yes (a , (m ,, pr))) } q | inr q1 | inr p | inr x | inr refl = refl
|
||||
countReducedBehavesWhenAddingPresent : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} {min : keys} → (k : keys) → (v v' : values) → (m : ReducedMap keys values keyOrder min) → (lookupReduced {keyOrder = keyOrder} m k ≡ yes v') → countReduced (addReducedMap k v m) ≡ countReduced m
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds with TotalOrder.totality keyOrder k min
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inl k<min) with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inl k<min) | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<min min<k))
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m () | inl (inl k<min) | inl (inr _)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inl k<min) | inr q = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (identityOfIndiscernablesLeft (TotalOrder._<_ keyOrder) k<min (equalityCommutative q)))
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inr min<k) with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = no } () | inl (inr min<k) | inl (inl _)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } lookupReducedSucceeds | inl (inr min<k) | inl (inl _) = applyEquality succ (countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} k v v' fst lookupReducedSucceeds)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m () | inl (inr min<k) | inl (inr k<min)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inl (inr min<k) | inr q = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (identityOfIndiscernablesLeft (λ a b → TotalOrder._<_ keyOrder a b) min<k q))
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inr q with TotalOrder.totality keyOrder min min
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inr q | inl (inl x) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' m lookupReducedSucceeds | inr q | inl (inr x) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) x)
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = no } lookupReducedSucceeds | inr _ | inr p with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = no } () | inr _ | inr p | inl (inl x)
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = no } () | inr _ | inr p | inl (inr x`)
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {keyOrder = keyOrder} {.k} k v v' record { firstEntry = firstEntry ; next = no } lookupReducedSucceeds | inr q | inr p | inr refl = refl
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedSucceeds | inr q | inr p with TotalOrder.totality keyOrder k k
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedSucceeds | inr q | inr p | inl (inl x) = exFalso (TotalOrder.irreflexive keyOrder x)
|
||||
countReducedBehavesWhenAddingPresent {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , b)) } lookupReducedSucceeds | inr q | inr p | inl (inr x) = exFalso (TotalOrder.irreflexive keyOrder x)
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {values = values} {keyOrder = keyOrder} {min} k v v' record { firstEntry = firstEntry ; next = (yes (a , (m ,, pr))) } q | inr q1 | inr p | inr x with TotalOrder.totality keyOrder min k
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {values} {keyOrder} {.k} k v v' record { firstEntry = firstEntry ; next = (yes (a , (m ,, pr))) } q | inr refl | inr p | inr x | inl (inl min<k) = exFalso (TotalOrder.irreflexive keyOrder min<k)
|
||||
countReducedBehavesWhenAddingPresent {keys = keys} {values} {keyOrder} {.k} k v v' record { firstEntry = firstEntry ; next = (yes (a , (m ,, pr))) } q | inr q1 | inr p | inr x | inr refl = refl
|
||||
|
||||
data Map {a b c : _} (keys : Set a) (values : Set b) (keyOrder : TotalOrder {_} {c} keys) : Set (a ⊔ b ⊔ c) where
|
||||
data Map {a b c : _} (keys : Set a) (values : Set b) (keyOrder : TotalOrder {_} {c} keys) : Set (a ⊔ b ⊔ c) where
|
||||
empty : Map keys values keyOrder
|
||||
nonempty : {min : keys} → ReducedMap keys values keyOrder min → Map keys values keyOrder
|
||||
|
||||
addMap : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → (m : Map keys values keyOrder) → (k : keys) → (v : values) → Map keys values keyOrder
|
||||
addMap empty k v = nonempty {min = k} record { firstEntry = v ; next = no }
|
||||
addMap {keys = keys} {values} {keyOrder} (nonempty x) k v = nonempty (addReducedMap {keys = keys} {values} {keyOrder} k v x)
|
||||
addMap : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → (m : Map keys values keyOrder) → (k : keys) → (v : values) → Map keys values keyOrder
|
||||
addMap empty k v = nonempty {min = k} record { firstEntry = v ; next = no }
|
||||
addMap {keys = keys} {values} {keyOrder} (nonempty x) k v = nonempty (addReducedMap {keys = keys} {values} {keyOrder} k v x)
|
||||
|
||||
lookup : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → (m : Map keys values keyOrder) → (target : keys) → Maybe values
|
||||
lookup {keyOrder = keyOrder} empty t = no
|
||||
lookup {keyOrder = keyOrder} (nonempty x) t = lookupReduced x t
|
||||
lookup : {a b c : _} → {keys : Set a} → {values : Set b} → {keyOrder : TotalOrder {_} {c} keys} → (m : Map keys values keyOrder) → (target : keys) → Maybe values
|
||||
lookup {keyOrder = keyOrder} empty t = no
|
||||
lookup {keyOrder = keyOrder} (nonempty x) t = lookupReduced x t
|
||||
|
||||
count : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} → (m : Map keys values keyOrder) → ℕ
|
||||
count {keyOrder = keyOrder} empty = 0
|
||||
count {keyOrder = keyOrder} (nonempty x) = countReduced x
|
||||
count : {a b c : _} {keys : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keys} → (m : Map keys values keyOrder) → ℕ
|
||||
count {keyOrder = keyOrder} empty = 0
|
||||
count {keyOrder = keyOrder} (nonempty x) = countReduced x
|
||||
|
||||
keysReduced : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} {min : keyDom} → (m : ReducedMap keyDom values keyOrder min) → Vec keyDom (countReduced m)
|
||||
keysReduced {min = min} record { firstEntry = firstEntry ; next = no } = min ,- []
|
||||
keysReduced {min = min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } = min ,- (keysReduced fst)
|
||||
keysReduced : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} {min : keyDom} → (m : ReducedMap keyDom values keyOrder min) → Vec keyDom (countReduced m)
|
||||
keysReduced {min = min} record { firstEntry = firstEntry ; next = no } = min ,- []
|
||||
keysReduced {min = min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } = min ,- (keysReduced fst)
|
||||
|
||||
keys : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} → (m : Map keyDom values keyOrder) → Vec keyDom (count m)
|
||||
keys empty = []
|
||||
keys (nonempty m) = keysReduced m
|
||||
keys : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} → (m : Map keyDom values keyOrder) → Vec keyDom (count m)
|
||||
keys empty = []
|
||||
keys (nonempty m) = keysReduced m
|
||||
|
||||
lookupReducedWhenLess : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} {min : keyDom} → (m : ReducedMap keyDom values keyOrder min) → (k : keyDom) → (TotalOrder._<_ keyOrder k min) → (lookupReduced m k ≡ no)
|
||||
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min with TotalOrder.totality keyOrder min k
|
||||
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<min min<k))
|
||||
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min | inl (inr _) = refl
|
||||
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min | inr min=k rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
lookupReducedWhenLess : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} {min : keyDom} → (m : ReducedMap keyDom values keyOrder min) → (k : keyDom) → (TotalOrder._<_ keyOrder k min) → (lookupReduced m k ≡ no)
|
||||
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min with TotalOrder.totality keyOrder min k
|
||||
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min | inl (inl min<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<min min<k))
|
||||
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min | inl (inr _) = refl
|
||||
lookupReducedWhenLess {keyOrder = keyOrder} {min} m k k<min | inr min=k rewrite min=k = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
|
||||
lookupCertainReduced : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} {min : keyDom} → (m : ReducedMap keyDom values keyOrder min) → (k : keyDom) → (vecContains (keysReduced m) k) → Sg values (λ v → lookupReduced m k ≡ yes v)
|
||||
lookupCertainReduced {keyDom = keyDom} {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = no } k pr = firstEntry , q
|
||||
lookupCertainReduced : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} {min : keyDom} → (m : ReducedMap keyDom values keyOrder min) → (k : keyDom) → (vecContains (keysReduced m) k) → Sg values (λ v → lookupReduced m k ≡ yes v)
|
||||
lookupCertainReduced {keyDom = keyDom} {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = no } k pr = firstEntry , q
|
||||
where
|
||||
t : min ≡ k
|
||||
t = vecSolelyContains k pr
|
||||
@@ -215,13 +216,13 @@ module KeyValue.LinearStore.Implementation where
|
||||
q | inr p | inl (inl min<k) rewrite t = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
q | inr p | inl (inr k<min) rewrite t = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
q | inr p | inr x = refl
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k pr with TotalOrder.totality keyOrder min k
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = zero ; index<m = _ ; isHere = isHere } | inl (inl min<k) rewrite isHere = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inl min<k) = lookupCertainReduced {keyOrder = keyOrder} fst k record { index = index ; index<m = canRemoveSuccFrom<N index<m ; isHere = isHere }
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = zero ; index<m = _ ; isHere = isHere } | inl (inr k<min) rewrite isHere = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) with TotalOrder.totality keyOrder a k
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) | inl (inl a<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<min (PartialOrder.<Transitive (TotalOrder.order keyOrder) snd a<k)))
|
||||
lookupCertainReduced {values = values} {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) | inl (inr k<a) = exFalso h
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k pr with TotalOrder.totality keyOrder min k
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = zero ; index<m = _ ; isHere = isHere } | inl (inl min<k) rewrite isHere = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) min<k)
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inl min<k) = lookupCertainReduced {keyOrder = keyOrder} fst k record { index = index ; index<m = canRemoveSuccFrom<N index<m ; isHere = isHere }
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = zero ; index<m = _ ; isHere = isHere } | inl (inr k<min) rewrite isHere = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) k<min)
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) with TotalOrder.totality keyOrder a k
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) | inl (inl a<k) = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) k<min (PartialOrder.<Transitive (TotalOrder.order keyOrder) snd a<k)))
|
||||
lookupCertainReduced {values = values} {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) | inl (inr k<a) = exFalso h
|
||||
where
|
||||
f : Sg values (λ v → lookupReduced fst k ≡ yes v)
|
||||
f = lookupCertainReduced {keyOrder = keyOrder} fst k record { index = index ; index<m = canRemoveSuccFrom<N index<m ; isHere = isHere }
|
||||
@@ -232,9 +233,9 @@ module KeyValue.LinearStore.Implementation where
|
||||
h : False
|
||||
h with f
|
||||
h | a , b rewrite g = noIsNotYes b
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) snd k<min))
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k pr | inr min=k = firstEntry , refl
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k record { index = (succ index) ; index<m = index<m ; isHere = isHere } | inl (inr k<min) | inr refl = exFalso (PartialOrder.irreflexive (TotalOrder.order keyOrder) (PartialOrder.<Transitive (TotalOrder.order keyOrder) snd k<min))
|
||||
lookupCertainReduced {keyOrder = keyOrder} {min} record { firstEntry = firstEntry ; next = (yes (a , (fst ,, snd))) } k pr | inr min=k = firstEntry , refl
|
||||
|
||||
lookupCertain : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} → (m : Map keyDom values keyOrder) → (k : keyDom) → (vecContains (keys m) k) → Sg values (λ v → lookup m k ≡ yes v)
|
||||
lookupCertain {keyOrder = keyOrder} empty k record { index = index ; index<m = (le x ()) ; isHere = isHere }
|
||||
lookupCertain {keyOrder = keyOrder} (nonempty {min} m) k pr = lookupCertainReduced {keyOrder = keyOrder} {min} m k pr
|
||||
lookupCertain : {a b c : _} {keyDom : Set a} {values : Set b} {keyOrder : TotalOrder {_} {c} keyDom} → (m : Map keyDom values keyOrder) → (k : keyDom) → (vecContains (keys m) k) → Sg values (λ v → lookup m k ≡ yes v)
|
||||
lookupCertain {keyOrder = keyOrder} empty k record { index = index ; index<m = (le x ()) ; isHere = isHere }
|
||||
lookupCertain {keyOrder = keyOrder} (nonempty {min} m) k pr = lookupCertainReduced {keyOrder = keyOrder} {min} m k pr
|
||||
|
@@ -8,7 +8,8 @@ open import Groups.Groups
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Rationals.Definition
|
||||
open import Rings.Definition
|
||||
|
@@ -3,6 +3,8 @@
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Subtraction
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Exponentiation
|
||||
|
@@ -2,7 +2,8 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Naturals -- for length
|
||||
open import Numbers.Naturals.Semiring -- for length
|
||||
open import Numbers.Naturals.Order
|
||||
open import Semirings.Definition
|
||||
|
||||
module Lists.Lists where
|
||||
|
@@ -4,7 +4,8 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Vectors
|
||||
|
||||
module Logic.PropositionalLogic where
|
||||
|
@@ -4,7 +4,7 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Logic.PropositionalLogic
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Vectors
|
||||
|
||||
module Logic.PropositionalLogicExamples where
|
||||
@@ -81,7 +81,7 @@ Injection.property (pQQRSubsetInj {P = P} {Q} {R} p!=q) {Two} {One} pr | q=p ,,
|
||||
Injection.property (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 = succIsPositive _ ; elementIsAt = refl }) (record { element = implies Q R ; position = 1 ; elementIsAt = refl ; pos<N = succPreservesInequality (succIsPositive _) }) (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)
|
||||
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)
|
||||
|
||||
pQQRProof : {a : _} {A : Set a} {P Q R : Propositions A} → (p!=q : P ≡ Q → False) → Proves propositionalAxioms (record { ofElt = pQQR {P = P} {Q} {R} }) (implies P R)
|
||||
pQQRProof p!=q = record { n = 6 ; proof = syntacticEntailmentExample p!=q ; ofStatement = refl }
|
||||
|
@@ -3,10 +3,13 @@
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Groups.Definition
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
|
||||
module Numbers.BinaryNaturals.Addition where
|
||||
|
||||
@@ -39,7 +42,7 @@ _+B_ : BinNat → BinNat → BinNat
|
||||
refine b pr with canonical b
|
||||
refine b pr | x :: bl = ::Inj pr
|
||||
t : NToBinNat (0 +N binNatToN (zero :: b)) ≡ zero :: b
|
||||
t with orderIsTotal 0 (binNatToN b)
|
||||
t with TotalOrder.totality ℕTotalOrder 0 (binNatToN b)
|
||||
t | inl (inl pos) = transitivity (doubleIsBitShift (binNatToN b) pos) (applyEquality (zero ::_) (transitivity (binToBin b) (equalityCommutative (refine b prB))))
|
||||
t | inl (inr ())
|
||||
... | inr eq with binNatToNZero b (equalityCommutative eq)
|
||||
@@ -61,9 +64,9 @@ _+B_ : BinNat → BinNat → BinNat
|
||||
|
||||
+BIsInherited' : (a b : BinNat) → a +Binherit b ≡ canonical (a +B b)
|
||||
|
||||
+BinheritLemma a b prA prB with orderIsTotal 0 (binNatToN a +N binNatToN b)
|
||||
+BinheritLemma a b prA prB with TotalOrder.totality ℕTotalOrder 0 (binNatToN a +N binNatToN b)
|
||||
+BinheritLemma a b prA prB | inl (inl x) rewrite doubleIsBitShift (binNatToN a +N binNatToN b) x = applyEquality (one ::_) (+BIsInherited a b prA prB)
|
||||
+BinheritLemma a b prA prB | inr x with sumZeroImpliesOperandsZero (binNatToN a) (equalityCommutative x)
|
||||
+BinheritLemma a b prA prB | inr x with sumZeroImpliesSummandsZero (equalityCommutative x)
|
||||
+BinheritLemma a b prA prB | inr x | fst ,, snd = ans2
|
||||
where
|
||||
bad : b ≡ []
|
||||
@@ -75,10 +78,10 @@ _+B_ : BinNat → BinNat → BinNat
|
||||
|
||||
+BIsInherited [] b _ prB = +BIsInherited[] b prB
|
||||
+BIsInherited (x :: a) [] prA _ = transitivity (applyEquality NToBinNat (Semiring.commutative ℕSemiring (binNatToN (x :: a)) 0)) (transitivity (binToBin (x :: a)) (equalityCommutative prA))
|
||||
+BIsInherited (zero :: as) (zero :: b) prA prB with orderIsTotal 0 (binNatToN as +N binNatToN b)
|
||||
+BIsInherited (zero :: as) (zero :: b) prA prB with TotalOrder.totality ℕTotalOrder 0 (binNatToN as +N binNatToN b)
|
||||
... | inl (inl 0<) rewrite Semiring.commutative ℕSemiring (binNatToN as) 0 | Semiring.commutative ℕSemiring (binNatToN b) 0 | Semiring.+Associative ℕSemiring (binNatToN as +N binNatToN as) (binNatToN b) (binNatToN b) | equalityCommutative (Semiring.+Associative ℕSemiring (binNatToN as) (binNatToN as) (binNatToN b)) | Semiring.commutative ℕSemiring (binNatToN as) (binNatToN b) | Semiring.+Associative ℕSemiring (binNatToN as) (binNatToN b) (binNatToN as) | equalityCommutative (Semiring.+Associative ℕSemiring (binNatToN as +N binNatToN b) (binNatToN as) (binNatToN b)) | Semiring.commutative ℕSemiring 0 ((binNatToN as +N binNatToN b) +N (binNatToN as +N binNatToN b)) | equalityCommutative (Semiring.+Associative ℕSemiring (binNatToN as +N binNatToN b) (binNatToN as +N binNatToN b) 0) = transitivity (doubleIsBitShift (binNatToN as +N binNatToN b) (identityOfIndiscernablesRight _<N_ 0< (Semiring.commutative ℕSemiring (binNatToN b) _))) (applyEquality (zero ::_) (+BIsInherited as b (canonicalDescends as prA) (canonicalDescends b prB)))
|
||||
+BIsInherited (zero :: as) (zero :: b) prA prB | inl (inr ())
|
||||
... | inr p with sumZeroImpliesOperandsZero (binNatToN as) (equalityCommutative p)
|
||||
... | inr p with sumZeroImpliesSummandsZero {binNatToN as} (equalityCommutative p)
|
||||
+BIsInherited (zero :: as) (zero :: b) prA prB | inr p | as=0 ,, b=0 rewrite as=0 | b=0 = exFalso ans
|
||||
where
|
||||
bad : (b : BinNat) → (pr : b ≡ canonical b) → (pr2 : binNatToN b ≡ 0) → b ≡ []
|
||||
@@ -142,7 +145,7 @@ _+B_ : BinNat → BinNat → BinNat
|
||||
where
|
||||
ans : NToBinNat (2 *N (binNatToN as +N binNatToN bs)) ≡ canonical (zero :: (as +B bs))
|
||||
ans with inspect (binNatToN as +N binNatToN bs)
|
||||
ans | zero with≡ x with sumZeroImpliesOperandsZero (binNatToN as) x
|
||||
ans | zero with≡ x with sumZeroImpliesSummandsZero {binNatToN as} x
|
||||
... | as=0 ,, bs=0 rewrite as=0 | bs=0 = foo
|
||||
where
|
||||
u : canonical (as +Binherit bs) ≡ []
|
||||
@@ -162,7 +165,7 @@ _+B_ : BinNat → BinNat → BinNat
|
||||
where
|
||||
ans2 : incr (NToBinNat (2 *N (binNatToN as +N binNatToN bs))) ≡ one :: canonical (as +B bs)
|
||||
ans2 with inspect (binNatToN as +N binNatToN bs)
|
||||
ans2 | zero with≡ x with sumZeroImpliesOperandsZero (binNatToN as) x
|
||||
ans2 | zero with≡ x with sumZeroImpliesSummandsZero {binNatToN as} x
|
||||
ans2 | zero with≡ x | as=0 ,, bs=0 rewrite as=0 | bs=0 = applyEquality (one ::_) (transitivity t (+BIsInherited' as bs))
|
||||
where
|
||||
t : [] ≡ as +Binherit bs
|
||||
@@ -172,7 +175,7 @@ _+B_ : BinNat → BinNat → BinNat
|
||||
where
|
||||
ans : incr (NToBinNat (2 *N (binNatToN as +N binNatToN bs))) ≡ one :: canonical (as +B bs)
|
||||
ans with inspect (binNatToN as +N binNatToN bs)
|
||||
ans | zero with≡ x with sumZeroImpliesOperandsZero (binNatToN as) x
|
||||
ans | zero with≡ x with sumZeroImpliesSummandsZero {binNatToN as} x
|
||||
... | as=0 ,, bs=0 rewrite as=0 | bs=0 = applyEquality (one ::_) (transitivity t (+BIsInherited' as bs))
|
||||
where
|
||||
t : [] ≡ NToBinNat (binNatToN as +N binNatToN bs)
|
||||
@@ -182,7 +185,7 @@ _+B_ : BinNat → BinNat → BinNat
|
||||
where
|
||||
ans : incr (incr (NToBinNat (2 *N (binNatToN as +N binNatToN bs)))) ≡ canonical (zero :: incr (as +B bs))
|
||||
ans with inspect (binNatToN as +N binNatToN bs)
|
||||
... | zero with≡ x with sumZeroImpliesOperandsZero (binNatToN as) x
|
||||
... | zero with≡ x with sumZeroImpliesSummandsZero {binNatToN as} x
|
||||
ans | zero with≡ x | as=0 ,, bs=0 rewrite as=0 | bs=0 = bar
|
||||
where
|
||||
u' : canonical (as +Binherit bs) ≡ []
|
||||
|
@@ -3,7 +3,8 @@
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Groups.Definition
|
||||
open import Semirings.Definition
|
||||
@@ -98,7 +99,7 @@ module Numbers.BinaryNaturals.Definition where
|
||||
canonicalAscends' {i} a pr = canonicalAscends {i} a (t a pr)
|
||||
where
|
||||
t : (a : BinNat) → (canonical a ≡ [] → False) → 0 <N binNatToN a
|
||||
t a pr with orderIsTotal 0 (binNatToN a)
|
||||
t a pr with TotalOrder.totality ℕTotalOrder 0 (binNatToN a)
|
||||
t a pr | inl (inl x) = x
|
||||
t a pr | inr x = exFalso (pr (binNatToNZero a (equalityCommutative x)))
|
||||
|
||||
|
@@ -3,6 +3,7 @@
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Definition
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
@@ -77,7 +78,7 @@ module Numbers.BinaryNaturals.Multiplication where
|
||||
t : binNatToN (as *B (zero :: bs)) +N binNatToN bs ≡ 0
|
||||
t = transitivity (equalityCommutative (nToN _)) (applyEquality binNatToN x)
|
||||
u : (binNatToN (as *B (zero :: bs)) ≡ 0) && (binNatToN bs ≡ 0)
|
||||
u = sumZeroImpliesOperandsZero _ t
|
||||
u = sumZeroImpliesSummandsZero t
|
||||
v : canonical bs ≡ []
|
||||
v with u
|
||||
... | fst ,, snd = binNatToNZero bs snd
|
||||
|
@@ -4,7 +4,9 @@ open import WellFoundedInduction
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Groups.Definition
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
open import Orders
|
||||
@@ -27,7 +29,7 @@ module Numbers.BinaryNaturals.Order where
|
||||
badCompare'' ()
|
||||
|
||||
_<BInherited_ : BinNat → BinNat → Compare
|
||||
a <BInherited b with orderIsTotal (binNatToN a) (binNatToN b)
|
||||
a <BInherited b with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||
(a <BInherited b) | inl (inl x) = FirstLess
|
||||
(a <BInherited b) | inl (inr x) = FirstGreater
|
||||
(a <BInherited b) | inr x = Equal
|
||||
@@ -283,24 +285,24 @@ module Numbers.BinaryNaturals.Order where
|
||||
chopFirstBit m n {one} FirstGreater = refl
|
||||
|
||||
chopDouble : (a b : BinNat) (i : Bit) → (i :: a) <BInherited (i :: b) ≡ a <BInherited b
|
||||
chopDouble a b i with orderIsTotal (binNatToN a) (binNatToN b)
|
||||
chopDouble a b zero | inl (inl a<b) with orderIsTotal (2 *N binNatToN a) (2 *N binNatToN b)
|
||||
chopDouble a b i with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||
chopDouble a b zero | inl (inl a<b) with TotalOrder.totality ℕTotalOrder (2 *N binNatToN a) (2 *N binNatToN b)
|
||||
chopDouble a b zero | inl (inl a<b) | inl (inl x) = refl
|
||||
chopDouble a b zero | inl (inl a<b) | inl (inr b<a) = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) b<a (lessRespectsMultiplicationLeft (binNatToN a) (binNatToN b) 2 a<b (le 1 refl))))
|
||||
chopDouble a b zero | inl (inl a<b) | inr a=b rewrite productCancelsLeft 2 (binNatToN a) (binNatToN b) (le 1 refl) a=b = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) a<b)
|
||||
chopDouble a b one | inl (inl a<b) with orderIsTotal (2 *N binNatToN a) (2 *N binNatToN b)
|
||||
chopDouble a b one | inl (inl a<b) with TotalOrder.totality ℕTotalOrder (2 *N binNatToN a) (2 *N binNatToN b)
|
||||
chopDouble a b one | inl (inl a<b) | inl (inl 2a<2b) = refl
|
||||
chopDouble a b one | inl (inl a<b) | inl (inr 2b<2a) = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) a<b (cancelInequalityLeft {2} 2b<2a)))
|
||||
chopDouble a b one | inl (inl a<b) | inr 2a=2b rewrite productCancelsLeft 2 (binNatToN a) (binNatToN b) (le 1 refl) 2a=2b = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) a<b)
|
||||
chopDouble a b zero | inl (inr b<a) with orderIsTotal (2 *N binNatToN a) (2 *N binNatToN b)
|
||||
chopDouble a b zero | inl (inr b<a) with TotalOrder.totality ℕTotalOrder (2 *N binNatToN a) (2 *N binNatToN b)
|
||||
chopDouble a b zero | inl (inr b<a) | inl (inl 2a<2b) = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) b<a (cancelInequalityLeft {2} {binNatToN a} {binNatToN b} 2a<2b)))
|
||||
chopDouble a b zero | inl (inr b<a) | inl (inr 2b<2a) = refl
|
||||
chopDouble a b zero | inl (inr b<a) | inr 2a=2b rewrite productCancelsLeft 2 (binNatToN a) (binNatToN b) (le 1 refl) 2a=2b = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) b<a)
|
||||
chopDouble a b one | inl (inr b<a) with orderIsTotal (2 *N binNatToN a) (2 *N binNatToN b)
|
||||
chopDouble a b one | inl (inr b<a) with TotalOrder.totality ℕTotalOrder (2 *N binNatToN a) (2 *N binNatToN b)
|
||||
chopDouble a b one | inl (inr b<a) | inl (inl 2a<2b) = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) b<a (cancelInequalityLeft {2} 2a<2b)))
|
||||
chopDouble a b one | inl (inr b<a) | inl (inr x) = refl
|
||||
chopDouble a b one | inl (inr b<a) | inr 2a=2b rewrite productCancelsLeft 2 (binNatToN a) (binNatToN b) (le 1 refl) 2a=2b = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) b<a)
|
||||
chopDouble a b i | inr x with orderIsTotal (binNatToN (i :: a)) (binNatToN (i :: b))
|
||||
chopDouble a b i | inr x with TotalOrder.totality ℕTotalOrder (binNatToN (i :: a)) (binNatToN (i :: b))
|
||||
chopDouble a b zero | inr a=b | inl (inl a<b) rewrite a=b = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) a<b)
|
||||
chopDouble a b one | inr a=b | inl (inl a<b) rewrite a=b = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) a<b)
|
||||
chopDouble a b zero | inr a=b | inl (inr b<a) rewrite a=b = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) b<a)
|
||||
@@ -311,23 +313,23 @@ module Numbers.BinaryNaturals.Order where
|
||||
succNotLess {succ n} (le x proof) = succNotLess {n} (le x (succInjective (transitivity (applyEquality succ (transitivity (Semiring.commutative ℕSemiring (succ x) (succ n)) (transitivity (applyEquality succ (transitivity (Semiring.commutative ℕSemiring n (succ x)) (applyEquality succ (Semiring.commutative ℕSemiring x n)))) (Semiring.commutative ℕSemiring (succ (succ n)) x)))) proof)))
|
||||
|
||||
<BIsInherited : (a b : BinNat) → a <BInherited b ≡ a <B b
|
||||
<BIsInherited [] b with orderIsTotal 0 (binNatToN b)
|
||||
<BIsInherited [] b with TotalOrder.totality ℕTotalOrder 0 (binNatToN b)
|
||||
<BIsInherited [] b | inl (inl x) with inspect (binNatToN b)
|
||||
<BIsInherited [] b | inl (inl x) | 0 with≡ pr rewrite binNatToNZero b pr | pr = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) x)
|
||||
<BIsInherited [] b | inl (inl x) | (succ bl) with≡ pr rewrite pr = equalityCommutative (zeroLess b λ p → zeroNotSucc bl b p pr)
|
||||
<BIsInherited [] b | inr 0=b rewrite canonicalSecond [] b Equal | binNatToNZero b (equalityCommutative 0=b) = refl
|
||||
<BIsInherited (a :: as) [] with orderIsTotal (binNatToN (a :: as)) 0
|
||||
<BIsInherited (a :: as) [] with TotalOrder.totality ℕTotalOrder (binNatToN (a :: as)) 0
|
||||
<BIsInherited (a :: as) [] | inl (inr x) with inspect (binNatToN (a :: as))
|
||||
<BIsInherited (a :: as) [] | inl (inr x) | zero with≡ pr rewrite binNatToNZero (a :: as) pr | pr = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) x)
|
||||
<BIsInherited (a :: as) [] | inl (inr x) | succ y with≡ pr rewrite pr = equalityCommutative (zeroLess' (a :: as) λ i → zeroNotSucc y (a :: as) i pr)
|
||||
<BIsInherited (a :: as) [] | inr x rewrite canonicalFirst (a :: as) [] Equal | binNatToNZero (a :: as) x = refl
|
||||
<BIsInherited (zero :: a) (zero :: b) = transitivity (chopDouble a b zero) (<BIsInherited a b)
|
||||
<BIsInherited (zero :: a) (one :: b) with orderIsTotal (binNatToN (zero :: a)) (binNatToN (one :: b))
|
||||
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) with orderIsTotal (binNatToN a) (binNatToN b)
|
||||
<BIsInherited (zero :: a) (one :: b) with TotalOrder.totality ℕTotalOrder (binNatToN (zero :: a)) (binNatToN (one :: b))
|
||||
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) | inl (inl a<b) = equalityCommutative (equalToFirstLess FirstLess a b (equalityCommutative indHyp))
|
||||
where
|
||||
t : a <BInherited b ≡ FirstLess
|
||||
t with orderIsTotal (binNatToN a) (binNatToN b)
|
||||
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||
t | inl (inl x) = refl
|
||||
t | inl (inr x) = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) x a<b))
|
||||
t | inr x rewrite x = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) a<b)
|
||||
@@ -335,12 +337,12 @@ module Numbers.BinaryNaturals.Order where
|
||||
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
||||
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) | inl (inr b<a) = exFalso (noIntegersBetweenXAndSuccX {2 *N binNatToN a} (2 *N binNatToN b) (lessRespectsMultiplicationLeft (binNatToN b) (binNatToN a) 2 b<a (le 1 refl)) 2a<2b+1)
|
||||
<BIsInherited (zero :: a) (one :: b) | inl (inl 2a<2b+1) | inr a=b rewrite a=b | canonicalFirst a b FirstLess | canonicalSecond (canonical a) b FirstLess | transitivity (equalityCommutative (binToBin a)) (transitivity (applyEquality NToBinNat a=b) (binToBin b)) = equalityCommutative (lemma1 (canonical b))
|
||||
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) with orderIsTotal (binNatToN a) (binNatToN b)
|
||||
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) | inl (inl a<b) = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) 2b+1<2a (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) (lessRespectsMultiplicationLeft (binNatToN a) (binNatToN b) 2 a<b (le 1 refl)) (le zero refl))))
|
||||
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) | inl (inr b<a) = equalityCommutative (equalToFirstGreater FirstLess a b (equalityCommutative indHyp))
|
||||
where
|
||||
t : a <BInherited b ≡ FirstGreater
|
||||
t with orderIsTotal (binNatToN a) (binNatToN b)
|
||||
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||
t | inl (inl x) = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) x b<a))
|
||||
t | inl (inr x) = refl
|
||||
t | inr x rewrite x = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) b<a)
|
||||
@@ -348,12 +350,12 @@ module Numbers.BinaryNaturals.Order where
|
||||
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
||||
<BIsInherited (zero :: a) (one :: b) | inl (inr 2b+1<2a) | inr a=b rewrite a=b = exFalso (succNotLess 2b+1<2a)
|
||||
<BIsInherited (zero :: a) (one :: b) | inr 2a=2b+1 = exFalso (parity (binNatToN b) (binNatToN a) (equalityCommutative 2a=2b+1))
|
||||
<BIsInherited (one :: a) (zero :: b) with orderIsTotal (binNatToN (one :: a)) (binNatToN (zero :: b))
|
||||
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) with orderIsTotal (binNatToN a) (binNatToN b)
|
||||
<BIsInherited (one :: a) (zero :: b) with TotalOrder.totality ℕTotalOrder (binNatToN (one :: a)) (binNatToN (zero :: b))
|
||||
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) | inl (inl a<b) = equalityCommutative (equalToFirstLess FirstGreater a b (equalityCommutative indHyp))
|
||||
where
|
||||
t : a <BInherited b ≡ FirstLess
|
||||
t with orderIsTotal (binNatToN a) (binNatToN b)
|
||||
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||
t | inl (inr x) = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) x a<b))
|
||||
t | inl (inl x) = refl
|
||||
t | inr x rewrite x = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) a<b)
|
||||
@@ -361,12 +363,12 @@ module Numbers.BinaryNaturals.Order where
|
||||
indHyp = transitivity (equalityCommutative t) (<BIsInherited a b)
|
||||
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) | inl (inr b<a) = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) 2a+1<2b (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) (lessRespectsMultiplicationLeft (binNatToN b) (binNatToN a) 2 b<a (le 1 refl)) (le zero refl))))
|
||||
<BIsInherited (one :: a) (zero :: b) | inl (inl 2a+1<2b) | inr a=b rewrite a=b = exFalso (succNotLess 2a+1<2b)
|
||||
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) with orderIsTotal (binNatToN a) (binNatToN b)
|
||||
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) | inl (inl a<b) = exFalso (noIntegersBetweenXAndSuccX {2 *N binNatToN b} (2 *N binNatToN a) (lessRespectsMultiplicationLeft (binNatToN a) (binNatToN b) 2 a<b (le 1 refl)) 2b<2a+1)
|
||||
<BIsInherited (one :: a) (zero :: b) | inl (inr 2b<2a+1) | inl (inr b<a) = equalityCommutative (equalToFirstGreater FirstGreater a b (equalityCommutative indHyp))
|
||||
where
|
||||
t : a <BInherited b ≡ FirstGreater
|
||||
t with orderIsTotal (binNatToN a) (binNatToN b)
|
||||
t with TotalOrder.totality ℕTotalOrder (binNatToN a) (binNatToN b)
|
||||
t | inl (inl x) = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) (PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) x b<a))
|
||||
t | inl (inr x) = refl
|
||||
t | inr x rewrite x = exFalso (PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) b<a)
|
||||
|
@@ -2,6 +2,7 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Lists.Lists
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
|
@@ -1,7 +1,7 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Definition
|
||||
open import Semirings.Definition
|
||||
open import Groups.Groups
|
||||
|
@@ -1,6 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Integers.Definition
|
||||
open import Numbers.Integers.Addition
|
||||
open import Numbers.Integers.Multiplication
|
||||
|
@@ -1,7 +1,7 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Multiplication
|
||||
open import Numbers.Integers.Definition
|
||||
open import Numbers.Integers.Addition
|
||||
|
@@ -1,7 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Integers.Definition
|
||||
open import Numbers.Integers.Addition
|
||||
open import Numbers.Integers.Multiplication
|
||||
@@ -68,13 +69,13 @@ _<Z_.proof (lessNegsuccNonneg {zero} {b}) = refl
|
||||
_<Z_.proof (lessNegsuccNonneg {succ a} {b}) = _<Z_.proof (lessNegsuccNonneg {a} {b})
|
||||
|
||||
lessThanTotalZ : {a b : ℤ} → ((a <Z b) || (b <Z a)) || (a ≡ b)
|
||||
lessThanTotalZ {nonneg a} {nonneg b} with orderIsTotal a b
|
||||
lessThanTotalZ {nonneg a} {nonneg b} with TotalOrder.totality ℕTotalOrder a b
|
||||
lessThanTotalZ {nonneg a} {nonneg b} | inl (inl a<b) = inl (inl (lessInherits a<b))
|
||||
lessThanTotalZ {nonneg a} {nonneg b} | inl (inr b<a) = inl (inr (lessInherits b<a))
|
||||
lessThanTotalZ {nonneg a} {nonneg b} | inr a=b = inr (applyEquality nonneg a=b)
|
||||
lessThanTotalZ {nonneg a} {negSucc b} = inl (inr (lessNegsuccNonneg {b} {a}))
|
||||
lessThanTotalZ {negSucc a} {nonneg x} = inl (inl (lessNegsuccNonneg {a} {x}))
|
||||
lessThanTotalZ {negSucc a} {negSucc b} with orderIsTotal a b
|
||||
lessThanTotalZ {negSucc a} {negSucc b} with TotalOrder.totality ℕTotalOrder a b
|
||||
... | inl (inl a<b) = inl (inr (lessInheritsNegsucc a<b))
|
||||
... | inl (inr b<a) = inl (inl (lessInheritsNegsucc b<a))
|
||||
lessThanTotalZ {negSucc a} {negSucc .a} | inr refl = inr refl
|
||||
|
@@ -7,8 +7,9 @@ open import Groups.Definition
|
||||
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.Addition -- TODO remove this dependency
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet
|
||||
@@ -20,6 +21,7 @@ open import Orders
|
||||
open import Numbers.Modulo.Definition
|
||||
|
||||
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
|
||||
@@ -29,7 +31,7 @@ cancelSumFromInequality {a} {b} {c} (le x proof) = le x help
|
||||
|
||||
_+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 orderIsTotal (a +N b) (succ n)
|
||||
_+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
|
||||
@@ -45,12 +47,12 @@ _+n_ {succ n} {pr} record { x = a ; xLess = aLess } record { x = b ; xLess = bLe
|
||||
|
||||
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 orderIsTotal (a +N 0) (succ x)
|
||||
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 = orderIsIrreflexive aL sx<a
|
||||
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
|
||||
@@ -59,7 +61,7 @@ plusZnIdentityRight {succ x} {_} record { x = a ; xLess = aLess } | inr a=sx = e
|
||||
|
||||
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 orderIsTotal x (succ n)
|
||||
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)
|
||||
@@ -69,28 +71,28 @@ 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 orderIsTotal (a +N b) (succ n)
|
||||
plusZnCommutative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ; xLess = bLess } | inl (inl a+b<sn) with orderIsTotal (b +N a) (succ n)
|
||||
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)
|
||||
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 (orderIsTransitive pr2 pr1)
|
||||
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 orderIsTotal (b +N a) (succ n)
|
||||
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 (orderIsTransitive sn<a+b b+a<sn)
|
||||
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 orderIsTotal (b +N a) (succ n)
|
||||
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
|
||||
|
@@ -7,8 +7,8 @@ open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Addition -- TODO remove this dependency
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet
|
||||
|
@@ -7,8 +7,10 @@ open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Addition -- TODO remove this dependency
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Sets.FinSet
|
||||
@@ -22,8 +24,10 @@ open import Numbers.Modulo.Addition
|
||||
|
||||
module Numbers.Modulo.Group where
|
||||
|
||||
open TotalOrder ℕTotalOrder
|
||||
|
||||
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 (orderIsTransitive pr5 c<n)
|
||||
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))
|
||||
@@ -92,7 +96,7 @@ 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 (orderIsTransitive a<n n<a+0)
|
||||
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
|
||||
@@ -154,7 +158,7 @@ help19 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive ℕTotalOrder
|
||||
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 (orderIsTransitive p4 a<n)
|
||||
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)
|
||||
@@ -163,7 +167,7 @@ help18 {a} {b} {c} {n} b+c<n n<a+b a<n pr = TotalOrder.irreflexive ℕTotalOrder
|
||||
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 = orderIsTransitive p2 (additionPreservesInequalityOnLeft a b+c<n)
|
||||
p3 = <Transitive p2 (additionPreservesInequalityOnLeft a b+c<n)
|
||||
p4 : n <N a
|
||||
p4 = subtractionPreservesInequality n p3
|
||||
|
||||
@@ -182,7 +186,7 @@ help17 {a} {b} {c} {n} n<b+c n<a+b p1 p2 = TotalOrder.irreflexive ℕTotalOrder
|
||||
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 (orderIsTransitive pr3 pr1''))
|
||||
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))
|
||||
@@ -196,7 +200,7 @@ help16 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = exFalso (TotalOrder.irreflexive ℕ
|
||||
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 (orderIsTransitive p2'' p1')
|
||||
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)
|
||||
@@ -250,7 +254,7 @@ help12 {a} {b} {c} {n} n<b+c n<a+b pr1 pr2 = TotalOrder.irreflexive ℕTotalOrde
|
||||
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 (orderIsTransitive a<n lemm5)
|
||||
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
|
||||
@@ -287,7 +291,7 @@ 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 (orderIsTransitive a<n n<a+0)
|
||||
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'
|
||||
@@ -319,7 +323,7 @@ help4 {a} {b} {c} {n} n<a+'b+c n<a+b = canSubtractFromEqualityLeft lemma'
|
||||
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 (orderIsTransitive (inter4 inter3) c<n)
|
||||
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))))
|
||||
@@ -356,27 +360,27 @@ help1 {a} {b} {c} {n} sn<b+c pr1 a+b<sn a<sn b<sn c<sn | record { result = b+c-s
|
||||
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' = orderIsTransitive eep2 (additionPreservesInequalityOnLeft (succ n) c<sn)
|
||||
eep2' = <Transitive eep2 (additionPreservesInequalityOnLeft (succ n) c<sn)
|
||||
ans : False
|
||||
ans = TotalOrder.irreflexive ℕTotalOrder (orderIsTransitive eep eep2')
|
||||
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 orderIsTotal (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 orderIsTotal ((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 orderIsTotal (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 orderIsTotal (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 } 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 (orderIsTransitive pr1 pr2)
|
||||
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 orderIsTotal (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 orderIsTotal (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) 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
|
||||
@@ -405,7 +409,7 @@ plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ;
|
||||
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 (orderIsTransitive pr1 pr2)
|
||||
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
|
||||
@@ -414,12 +418,12 @@ plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ;
|
||||
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 orderIsTotal (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 orderIsTotal (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) 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 (orderIsTransitive pr1 pr2)
|
||||
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
|
||||
@@ -430,12 +434,12 @@ plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ;
|
||||
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 orderIsTotal (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) 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 (orderIsTransitive 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 orderIsTotal (a +N subtractionNResult.result (-N (inl sn<b+c))) (succ n)
|
||||
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)
|
||||
@@ -443,12 +447,12 @@ plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ;
|
||||
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 orderIsTotal (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 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 (orderIsTransitive 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 orderIsTotal (a +N 0) (succ n)
|
||||
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))
|
||||
@@ -457,8 +461,8 @@ plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ;
|
||||
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 (orderIsTransitive pr1 pr2)
|
||||
false (succ b) pr1 pr2 rewrite Semiring.commutative ℕSemiring a 0 = TotalOrder.irreflexive ℕTotalOrder (orderIsTransitive pr2 (orderIsTransitive (le b (Semiring.commutative ℕSemiring (succ b) a)) pr1))
|
||||
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
|
||||
@@ -468,8 +472,8 @@ plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ;
|
||||
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 orderIsTotal (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 orderIsTotal (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 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
|
||||
@@ -483,8 +487,8 @@ plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ;
|
||||
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 (orderIsTransitive 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 orderIsTotal (a +N 0) (succ n)
|
||||
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
|
||||
@@ -508,13 +512,13 @@ plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ;
|
||||
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 orderIsTotal (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 orderIsTotal (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) 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' (orderIsTransitive 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 orderIsTotal (subtractionNResult.result (-N (inl sn<a+b)) +N c) (succ n)
|
||||
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)
|
||||
@@ -522,47 +526,47 @@ plusZnAssociative {succ n} {_} record { x = a ; xLess = aLess } record { x = b ;
|
||||
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 orderIsTotal (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 orderIsTotal (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) 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 orderIsTotal (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) 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 orderIsTotal (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 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 orderIsTotal (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 orderIsTotal (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 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 orderIsTotal 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 orderIsTotal (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 orderIsTotal (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 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 orderIsTotal (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) 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 orderIsTotal (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 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 (orderIsTransitive 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 orderIsTotal (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 (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 orderIsTotal (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) 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
|
||||
@@ -580,7 +584,7 @@ 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 orderIsTotal (subtractionNResult.result (-N (inl (canRemoveSuccFrom<N aLess))) +N (succ a)) (succ 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
|
||||
|
@@ -14,9 +14,10 @@ addZeroRight : (x : ℕ) → (x +N zero) ≡ x
|
||||
addZeroRight zero = refl
|
||||
addZeroRight (succ x) rewrite addZeroRight x = refl
|
||||
|
||||
succExtracts : (x y : ℕ) → (x +N succ y) ≡ (succ (x +N y))
|
||||
succExtracts zero y = refl
|
||||
succExtracts (succ x) y = applyEquality succ (succExtracts x y)
|
||||
private
|
||||
succExtracts : (x y : ℕ) → (x +N succ y) ≡ (succ (x +N y))
|
||||
succExtracts zero y = refl
|
||||
succExtracts (succ x) y = applyEquality succ (succExtracts x y)
|
||||
|
||||
succCanMove : (x y : ℕ) → (x +N succ y) ≡ (succ x +N y)
|
||||
succCanMove x y = transitivity (succExtracts x y) refl
|
||||
|
@@ -4,8 +4,8 @@ open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import WellFoundedInduction
|
||||
open import Functions
|
||||
open import Orders
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Addition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Multiplication
|
||||
@@ -13,64 +13,46 @@ open import Numbers.Naturals.Exponentiation
|
||||
open import Numbers.Naturals.Subtraction
|
||||
open import Semirings.Definition
|
||||
open import Monoids.Definition
|
||||
open import Orders
|
||||
|
||||
module Numbers.Naturals.Naturals where
|
||||
|
||||
open Numbers.Naturals.Definition using (ℕ ; zero ; succ ; succInjective ; naughtE) public
|
||||
open Numbers.Naturals.Addition using (_+N_ ; canSubtractFromEqualityRight ; canSubtractFromEqualityLeft) public
|
||||
open Numbers.Naturals.Multiplication using (_*N_ ; multiplicationNIsCommutative) public
|
||||
open Numbers.Naturals.Exponentiation using (_^N_) public
|
||||
open Numbers.Naturals.Order using (_<N_ ; le; succPreservesInequality; succIsPositive; addingIncreases; zeroNeverGreater; noIntegersBetweenXAndSuccX; a<SuccA; canRemoveSuccFrom<N) public
|
||||
open Numbers.Naturals.Subtraction using (_-N'_) public
|
||||
|
||||
ℕSemiring : Semiring 0 1 _+N_ _*N_
|
||||
Monoid.associative (Semiring.monoid ℕSemiring) a b c = equalityCommutative (additionNIsAssociative a b c)
|
||||
Monoid.idLeft (Semiring.monoid ℕSemiring) _ = refl
|
||||
Monoid.idRight (Semiring.monoid ℕSemiring) a = additionNIsCommutative a 0
|
||||
Semiring.commutative ℕSemiring = additionNIsCommutative
|
||||
Monoid.associative (Semiring.multMonoid ℕSemiring) = multiplicationNIsAssociative
|
||||
Monoid.idLeft (Semiring.multMonoid ℕSemiring) a = additionNIsCommutative a 0
|
||||
Monoid.idRight (Semiring.multMonoid ℕSemiring) a = transitivity (multiplicationNIsCommutative a 1) (additionNIsCommutative a 0)
|
||||
Semiring.productZeroLeft ℕSemiring _ = refl
|
||||
Semiring.productZeroRight ℕSemiring a = multiplicationNIsCommutative a 0
|
||||
Semiring.+DistributesOver* ℕSemiring = productDistributes
|
||||
|
||||
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
|
||||
|
||||
subtractionNWellDefined : {a b : ℕ} → {p1 p2 : a ≤N b} → (s : subtractionNResult a b p1) → (t : subtractionNResult a b p2) → (subtractionNResult.result s ≡ subtractionNResult.result t)
|
||||
subtractionNWellDefined {a} {b} {inl x} {pr2} record { result = result1 ; pr = pr1 } record { result = result ; pr = pr } = canSubtractFromEqualityLeft {a} (transitivity pr1 (equalityCommutative pr))
|
||||
subtractionNWellDefined {a} {.a} {inr refl} {pr2} record { result = result1 ; pr = pr1 } record { result = result2 ; pr = pr } = transitivity g' (equalityCommutative g)
|
||||
subtractionNWellDefined : {a b : ℕ} → {p1 p2 : a ≤N b} → (s : subtractionNResult a b p1) → (t : subtractionNResult a b p2) → (subtractionNResult.result s ≡ subtractionNResult.result t)
|
||||
subtractionNWellDefined {a} {b} {inl x} {pr2} record { result = result1 ; pr = pr1 } record { result = result ; pr = pr } = canSubtractFromEqualityLeft {a} (transitivity pr1 (equalityCommutative pr))
|
||||
subtractionNWellDefined {a} {.a} {inr refl} {pr2} record { result = result1 ; pr = pr1 } record { result = result2 ; pr = pr } = transitivity g' (equalityCommutative g)
|
||||
where
|
||||
g : result2 ≡ 0
|
||||
g = canSubtractFromEqualityLeft {a} {_} {0} (transitivity pr (equalityCommutative (addZeroRight a)))
|
||||
g' : result1 ≡ 0
|
||||
g' = canSubtractFromEqualityLeft {a} {_} {0} (transitivity pr1 (equalityCommutative (addZeroRight a)))
|
||||
|
||||
-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} (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 : {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} (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) }
|
||||
|
||||
addOneToWeakInequality : {a b : ℕ} → (a ≤N b) → (succ a ≤N succ b)
|
||||
addOneToWeakInequality {a} {b} (inl ineq) = inl (succPreservesInequality ineq)
|
||||
addOneToWeakInequality {a} {.a} (inr refl) = inr refl
|
||||
addOneToWeakInequality : {a b : ℕ} → (a ≤N b) → (succ a ≤N succ b)
|
||||
addOneToWeakInequality {a} {b} (inl ineq) = inl (succPreservesInequality ineq)
|
||||
addOneToWeakInequality {a} {.a} (inr refl) = inr refl
|
||||
|
||||
bumpUpSubtraction : {a b : ℕ} → (p1 : a ≤N b) → (s : subtractionNResult a b p1) → Sg (subtractionNResult (succ a) (succ b) (addOneToWeakInequality p1)) (λ n → subtractionNResult.result n ≡ subtractionNResult.result s)
|
||||
bumpUpSubtraction {a} {b} a<=b record { result = result ; pr = pr } = record { result = result ; pr = applyEquality succ pr } , refl
|
||||
bumpUpSubtraction : {a b : ℕ} → (p1 : a ≤N b) → (s : subtractionNResult a b p1) → Sg (subtractionNResult (succ a) (succ b) (addOneToWeakInequality p1)) (λ n → subtractionNResult.result n ≡ subtractionNResult.result s)
|
||||
bumpUpSubtraction {a} {b} a<=b record { result = result ; pr = pr } = record { result = result ; pr = applyEquality succ pr } , refl
|
||||
|
||||
addMinus : {a : ℕ} → {b : ℕ} → (pr : a ≤N b) → subtractionNResult.result (-N {a} {b} pr) +N a ≡ b
|
||||
addMinus {zero} {zero} p = refl
|
||||
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 : {a : ℕ} → {b : ℕ} → (pr : a ≤N b) → subtractionNResult.result (-N {a} {b} pr) +N a ≡ b
|
||||
addMinus {zero} {zero} p = refl
|
||||
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
|
||||
where
|
||||
pr'' : (a <N b) || (a ≡ b)
|
||||
pr'' = (inl (le (_<N_.x x) (transitivity (equalityCommutative (succExtracts (_<N_.x x) a)) (succInjective (_<N_.proof x)))))
|
||||
@@ -80,212 +62,70 @@ module Numbers.Naturals.Naturals where
|
||||
next = bumpUpSubtraction pr'' previous
|
||||
t : result ≡ subtractionNResult.result (underlying next)
|
||||
t = subtractionNWellDefined {succ a} {succ b} {inl x} {addOneToWeakInequality pr''} (record { result = result ; pr = pr }) (underlying next)
|
||||
addMinus {succ a} {succ .a} (inr refl) = refl
|
||||
addMinus {succ a} {succ .a} (inr refl) = refl
|
||||
|
||||
addMinus' : {a b : ℕ} → (pr : a ≤N b) → a +N subtractionNResult.result (-N {a} {b} pr) ≡ b
|
||||
addMinus' {a} {b} pr rewrite additionNIsCommutative a (subtractionNResult.result (-N {a} {b} pr)) = addMinus {a} {b} pr
|
||||
addMinus' : {a b : ℕ} → (pr : a ≤N b) → a +N subtractionNResult.result (-N {a} {b} pr) ≡ b
|
||||
addMinus' {a} {b} pr rewrite additionNIsCommutative a (subtractionNResult.result (-N {a} {b} pr)) = addMinus {a} {b} pr
|
||||
|
||||
additionPreservesInequality : {a b : ℕ} → (c : ℕ) → a <N b → a +N c <N b +N c
|
||||
additionPreservesInequality {a} {b} zero prAB rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 = prAB
|
||||
additionPreservesInequality {a} {b} (succ c) (le x proof) = le x (transitivity (equalityCommutative (additionNIsAssociative (succ x) a (succ c))) (applyEquality (_+N succ c) proof))
|
||||
additionPreservesInequality : {a b : ℕ} → (c : ℕ) → a <N b → a +N c <N b +N c
|
||||
additionPreservesInequality {a} {b} zero prAB rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 = prAB
|
||||
additionPreservesInequality {a} {b} (succ c) (le x proof) = le x (transitivity (equalityCommutative (additionNIsAssociative (succ x) a (succ c))) (applyEquality (_+N succ c) proof))
|
||||
|
||||
additionPreservesInequalityOnLeft : {a b : ℕ} → (c : ℕ) → a <N b → c +N a <N c +N b
|
||||
additionPreservesInequalityOnLeft {a} {b} c prAB = identityOfIndiscernablesRight (λ a b → a <N b) (identityOfIndiscernablesLeft (λ a b → a <N b) (additionPreservesInequality {a} {b} c prAB) (additionNIsCommutative a c)) (additionNIsCommutative b c)
|
||||
additionPreservesInequalityOnLeft : {a b : ℕ} → (c : ℕ) → a <N b → c +N a <N c +N b
|
||||
additionPreservesInequalityOnLeft {a} {b} c prAB = identityOfIndiscernablesRight (λ a b → a <N b) (identityOfIndiscernablesLeft (λ a b → a <N b) (additionPreservesInequality {a} {b} c prAB) (additionNIsCommutative a c)) (additionNIsCommutative b c)
|
||||
|
||||
subtractionPreservesInequality : {a b : ℕ} → (c : ℕ) → a +N c <N b +N c → a <N b
|
||||
subtractionPreservesInequality {a} {b} zero prABC rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 = prABC
|
||||
subtractionPreservesInequality {a} {b} (succ c) (le x proof) = le x (canSubtractFromEqualityRight {b = succ c} (transitivity (additionNIsAssociative (succ x) a (succ c)) proof))
|
||||
multiplyIncreases : (a : ℕ) → (b : ℕ) → succ zero <N a → zero <N b → b <N a *N b
|
||||
multiplyIncreases zero b (le x ()) prB
|
||||
multiplyIncreases (succ zero) b (le zero ()) prb
|
||||
multiplyIncreases (succ zero) b (le (succ x) ()) prb
|
||||
multiplyIncreases (succ (succ a)) (succ b) prA prb = le (b +N a *N succ b) (applyEquality succ (transitivity (Semiring.commutative ℕSemiring _ (succ b)) (transitivity (applyEquality succ (Semiring.commutative ℕSemiring b _)) (Semiring.commutative ℕSemiring _ b))))
|
||||
|
||||
productNonzeroIsNonzero : {a b : ℕ} → zero <N a → zero <N b → zero <N a *N b
|
||||
productNonzeroIsNonzero {zero} {b} prA prB = prA
|
||||
productNonzeroIsNonzero {succ a} {zero} prA ()
|
||||
productNonzeroIsNonzero {succ a} {succ b} prA prB = logical<NImpliesLe (record {})
|
||||
canTimesOneOnLeft : {a b : ℕ} → (a <N b) → (a *N (succ zero)) <N b
|
||||
canTimesOneOnLeft {a} {b} prAB = identityOfIndiscernablesLeft _<N_ prAB (equalityCommutative (productWithOneRight a))
|
||||
|
||||
multiplyIncreases : (a : ℕ) → (b : ℕ) → succ zero <N a → zero <N b → b <N a *N b
|
||||
multiplyIncreases zero b (le x ()) prB
|
||||
multiplyIncreases (succ zero) b prA prb with leImpliesLogical<N {succ zero} {succ zero} prA
|
||||
multiplyIncreases (succ zero) b prA prb | ()
|
||||
multiplyIncreases (succ (succ a)) zero prA ()
|
||||
multiplyIncreases (succ (succ a)) (succ b) prA prB =
|
||||
identityOfIndiscernablesRight _<N_ k refl
|
||||
where
|
||||
h : zero <N (succ a) *N (succ b)
|
||||
h = productNonzeroIsNonzero {succ a} {succ b} (logical<NImpliesLe (record {})) (logical<NImpliesLe (record {}))
|
||||
i : zero +N succ b <N (succ a) *N (succ b) +N succ b
|
||||
i = additionPreservesInequality {zero} {succ a *N succ b} (succ b) h
|
||||
j : zero +N succ b <N succ b +N (succ a *N succ b)
|
||||
j = identityOfIndiscernablesRight _<N_ i (additionNIsCommutative (succ (b +N a *N succ b)) (succ b))
|
||||
k : succ b <N succ b +N (succ a *N succ b)
|
||||
k = identityOfIndiscernablesLeft (λ a b → a <N b) j refl
|
||||
canTimesOneOnRight : {a b : ℕ} → (a <N b) → a <N (b *N (succ zero))
|
||||
canTimesOneOnRight {a} {b} prAB = identityOfIndiscernablesRight _<N_ prAB (equalityCommutative (productWithOneRight b))
|
||||
|
||||
productCancelsRight : (a b c : ℕ) → (zero <N a) → (b *N a ≡ c *N a) → (b ≡ c)
|
||||
productCancelsRight a zero zero aPos eq = refl
|
||||
productCancelsRight zero zero (succ c) (le x ()) eq
|
||||
productCancelsRight (succ a) zero (succ c) aPos eq = contr
|
||||
where
|
||||
h : zero ≡ succ c *N succ a
|
||||
h = eq
|
||||
contr : zero ≡ succ c
|
||||
contr = exFalso (naughtE h)
|
||||
canSwapAddOnLeftOfInequality : {a b c : ℕ} → (a +N b <N c) → (b +N a <N c)
|
||||
canSwapAddOnLeftOfInequality {a} {b} {c} pr = identityOfIndiscernablesLeft _<N_ pr (additionNIsCommutative a b)
|
||||
|
||||
productCancelsRight zero (succ b) zero (le x ()) eq
|
||||
productCancelsRight (succ a) (succ b) zero aPos eq = contr
|
||||
where
|
||||
h : succ b *N succ a ≡ zero
|
||||
h = eq
|
||||
contr : succ b ≡ zero
|
||||
contr = exFalso (naughtE (equalityCommutative h))
|
||||
canSwapAddOnRightOfInequality : {a b c : ℕ} → (a <N b +N c) → (a <N c +N b)
|
||||
canSwapAddOnRightOfInequality {a} {b} {c} pr = identityOfIndiscernablesRight _<N_ pr (additionNIsCommutative b c)
|
||||
|
||||
productCancelsRight zero (succ b) (succ c) (le x ()) eq
|
||||
bumpDownOnRight : (a c : ℕ) → c *N succ a ≡ c *N a +N c
|
||||
bumpDownOnRight a c = transitivity (multiplicationNIsCommutative c (succ a)) (transitivity refl (transitivity (additionNIsCommutative c (a *N c)) ((addingPreservesEqualityRight c (multiplicationNIsCommutative a c) ))))
|
||||
|
||||
productCancelsRight (succ a) (succ b) (succ c) aPos eq = applyEquality succ (productCancelsRight (succ a) b c aPos l)
|
||||
where
|
||||
i : succ a +N b *N succ a ≡ succ c *N succ a
|
||||
i = eq
|
||||
j : succ c *N succ a ≡ succ a +N c *N succ a
|
||||
j = refl
|
||||
k : succ a +N b *N succ a ≡ succ a +N c *N succ a
|
||||
k = transitivity i j
|
||||
l : b *N succ a ≡ c *N succ a
|
||||
l = canSubtractFromEqualityLeft {succ a} {b *N succ a} {c *N succ a} k
|
||||
succIsNonzero : {a : ℕ} → (succ a ≡ zero) → False
|
||||
succIsNonzero {a} ()
|
||||
|
||||
productCancelsLeft : (a b c : ℕ) → (zero <N a) → (a *N b ≡ a *N c) → (b ≡ c)
|
||||
productCancelsLeft a b c aPos pr = productCancelsRight a b c aPos j
|
||||
where
|
||||
i : b *N a ≡ a *N c
|
||||
i = identityOfIndiscernablesLeft _≡_ pr (multiplicationNIsCommutative a b)
|
||||
j : b *N a ≡ c *N a
|
||||
j = identityOfIndiscernablesRight _≡_ i (multiplicationNIsCommutative a c)
|
||||
lessImpliesNotEqual : {a b : ℕ} → (a <N b) → a ≡ b → False
|
||||
lessImpliesNotEqual {a} {.a} prAB refl = TotalOrder.irreflexive ℕTotalOrder prAB
|
||||
|
||||
productCancelsRight' : (a b c : ℕ) → (b *N a ≡ c *N a) → (a ≡ zero) || (b ≡ c)
|
||||
productCancelsRight' zero b c pr = inl refl
|
||||
productCancelsRight' (succ a) b c pr = inr (productCancelsRight (succ a) b c (succIsPositive a) pr)
|
||||
-NIsDecreasing : {a b : ℕ} → (prAB : succ a <N b) → subtractionNResult.result (-N (inl prAB)) <N b
|
||||
-NIsDecreasing {a} {b} prAB with (-N (inl prAB))
|
||||
-NIsDecreasing {a} {b} (le x proof) | record { result = result ; pr = pr } = record { x = a ; proof = pr }
|
||||
|
||||
productCancelsLeft' : (a b c : ℕ) → (a *N b ≡ a *N c) → (a ≡ zero) || (b ≡ c)
|
||||
productCancelsLeft' zero b c pr = inl refl
|
||||
productCancelsLeft' (succ a) b c pr = inr (productCancelsLeft (succ a) b c (succIsPositive a) pr)
|
||||
equalityN : (a b : ℕ) → Sg Bool (λ truth → if truth then a ≡ b else Unit)
|
||||
equalityN zero zero = ( BoolTrue , refl )
|
||||
equalityN zero (succ b) = ( BoolFalse , record {} )
|
||||
equalityN (succ a) zero = ( BoolFalse , record {} )
|
||||
equalityN (succ a) (succ b) with equalityN a b
|
||||
equalityN (succ a) (succ b) | BoolTrue , val = (BoolTrue , applyEquality succ val)
|
||||
equalityN (succ a) (succ b) | BoolFalse , val = (BoolFalse , record {})
|
||||
|
||||
lessRespectsAddition : {a b : ℕ} (c : ℕ) → (a <N b) → ((a +N c) <N (b +N c))
|
||||
lessRespectsAddition {a} {b} zero prAB rewrite additionNIsCommutative a 0 | additionNIsCommutative b 0 = prAB
|
||||
lessRespectsAddition {a} {b} (succ c) prAB rewrite additionNIsCommutative a (succ c) | additionNIsCommutative b (succ c) | additionNIsCommutative c a | additionNIsCommutative c b = succPreservesInequality (lessRespectsAddition c prAB)
|
||||
sumZeroImpliesSummandsZero : {a b : ℕ} → (a +N b ≡ zero) → ((a ≡ zero) && (b ≡ zero))
|
||||
sumZeroImpliesSummandsZero {zero} {zero} pr = record { fst = refl ; snd = refl }
|
||||
sumZeroImpliesSummandsZero {zero} {(succ b)} pr = record { fst = refl ; snd = pr }
|
||||
sumZeroImpliesSummandsZero {(succ a)} {zero} ()
|
||||
sumZeroImpliesSummandsZero {(succ a)} {(succ b)} ()
|
||||
|
||||
canTimesOneOnLeft : {a b : ℕ} → (a <N b) → (a *N (succ zero)) <N b
|
||||
canTimesOneOnLeft {a} {b} prAB = identityOfIndiscernablesLeft _<N_ prAB (equalityCommutative (productWithOneRight a))
|
||||
productWithNonzeroZero : (a b : ℕ) → (a *N succ b ≡ zero) → a ≡ zero
|
||||
productWithNonzeroZero zero b pr = refl
|
||||
productWithNonzeroZero (succ a) b ()
|
||||
|
||||
canTimesOneOnRight : {a b : ℕ} → (a <N b) → a <N (b *N (succ zero))
|
||||
canTimesOneOnRight {a} {b} prAB = identityOfIndiscernablesRight _<N_ prAB (equalityCommutative (productWithOneRight b))
|
||||
|
||||
canSwapAddOnLeftOfInequality : {a b c : ℕ} → (a +N b <N c) → (b +N a <N c)
|
||||
canSwapAddOnLeftOfInequality {a} {b} {c} pr = identityOfIndiscernablesLeft _<N_ pr (additionNIsCommutative a b)
|
||||
|
||||
canSwapAddOnRightOfInequality : {a b c : ℕ} → (a <N b +N c) → (a <N c +N b)
|
||||
canSwapAddOnRightOfInequality {a} {b} {c} pr = identityOfIndiscernablesRight _<N_ pr (additionNIsCommutative b c)
|
||||
|
||||
naturalsAreNonnegative : (a : ℕ) → (a <NLogical zero) → False
|
||||
naturalsAreNonnegative zero pr = pr
|
||||
naturalsAreNonnegative (succ x) pr = pr
|
||||
|
||||
canFlipMultiplicationsInIneq : {a b c d : ℕ} → (a *N b <N c *N d) → b *N a <N d *N c
|
||||
canFlipMultiplicationsInIneq {a} {b} {c} {d} pr = identityOfIndiscernablesRight _<N_ (identityOfIndiscernablesLeft _<N_ pr (multiplicationNIsCommutative a b)) (multiplicationNIsCommutative c d)
|
||||
|
||||
bumpDownOnRight : (a c : ℕ) → c *N succ a ≡ c *N a +N c
|
||||
bumpDownOnRight a c = transitivity (multiplicationNIsCommutative c (succ a)) (transitivity refl (transitivity (additionNIsCommutative c (a *N c)) ((addingPreservesEqualityRight c (multiplicationNIsCommutative a c) ))))
|
||||
|
||||
lessRespectsMultiplicationLeft : (a b c : ℕ) → (a <N b) → (zero <N c) → (c *N a <N c *N b)
|
||||
lessRespectsMultiplicationLeft zero zero c (le x ()) cPos
|
||||
lessRespectsMultiplicationLeft zero (succ b) zero prAB (le x ())
|
||||
lessRespectsMultiplicationLeft zero (succ b) (succ c) prAB cPos = i
|
||||
where
|
||||
j : zero <N succ c *N succ b
|
||||
j = productNonzeroIsNonzero cPos prAB
|
||||
i : succ c *N zero <N succ c *N succ b
|
||||
i = identityOfIndiscernablesLeft _<N_ j (equalityCommutative (productZeroIsZeroRight c))
|
||||
lessRespectsMultiplicationLeft (succ a) zero c (le x ()) cPos
|
||||
lessRespectsMultiplicationLeft (succ a) (succ b) c prAB cPos = m
|
||||
where
|
||||
h : c *N a <N c *N b
|
||||
h = lessRespectsMultiplicationLeft a b c (logical<NImpliesLe (leImpliesLogical<N prAB)) cPos
|
||||
j : c *N a +N c <N c *N b +N c
|
||||
j = lessRespectsAddition c h
|
||||
i : c *N succ a <N c *N b +N c
|
||||
i = identityOfIndiscernablesLeft _<N_ j (equalityCommutative (bumpDownOnRight a c))
|
||||
m : c *N succ a <N c *N succ b
|
||||
m = identityOfIndiscernablesRight _<N_ i (equalityCommutative (bumpDownOnRight b c))
|
||||
|
||||
lessRespectsMultiplication : (a b c : ℕ) → (a <N b) → (zero <N c) → (a *N c <N b *N c)
|
||||
lessRespectsMultiplication a b c prAB cPos = canFlipMultiplicationsInIneq {c} {a} {c} {b} (lessRespectsMultiplicationLeft a b c prAB cPos)
|
||||
|
||||
succIsNonzero : {a : ℕ} → (succ a ≡ zero) → False
|
||||
succIsNonzero {a} ()
|
||||
|
||||
orderIsTotal : (a b : ℕ) → ((a <N b) || (b <N a)) || (a ≡ b)
|
||||
orderIsTotal zero zero = inr refl
|
||||
orderIsTotal zero (succ b) = inl (inl (logical<NImpliesLe (record {})))
|
||||
orderIsTotal (succ a) zero = inl (inr (logical<NImpliesLe (record {})))
|
||||
orderIsTotal (succ a) (succ b) with orderIsTotal a b
|
||||
orderIsTotal (succ a) (succ b) | inl (inl x) = inl (inl (logical<NImpliesLe (leImpliesLogical<N x)))
|
||||
orderIsTotal (succ a) (succ b) | inl (inr x) = inl (inr (logical<NImpliesLe (leImpliesLogical<N x)))
|
||||
orderIsTotal (succ a) (succ b) | inr x = inr (applyEquality succ x)
|
||||
|
||||
orderIsIrreflexive : {a b : ℕ} → (a <N b) → (b <N a) → False
|
||||
orderIsIrreflexive {zero} {b} prAB (le x ())
|
||||
orderIsIrreflexive {(succ a)} {zero} (le x ()) prBA
|
||||
orderIsIrreflexive {(succ a)} {(succ b)} prAB prBA = orderIsIrreflexive {a} {b} (logical<NImpliesLe (leImpliesLogical<N prAB)) (logical<NImpliesLe (leImpliesLogical<N prBA))
|
||||
|
||||
orderIsTransitive : {a b c : ℕ} → (a <N b) → (b <N c) → (a <N c)
|
||||
orderIsTransitive {a} {.(succ (x +N a))} {.(succ (y +N succ (x +N a)))} (le x refl) (le y refl) = le (y +N succ x) (applyEquality succ (additionNIsAssociative y (succ x) a))
|
||||
|
||||
<NWellfounded : WellFounded _<N_
|
||||
<NWellfounded = λ x → access (go x)
|
||||
where
|
||||
lemma : {a b c : ℕ} → a <N b → b <N succ c → a <N c
|
||||
lemma {a} {b} {c} (le y succYAeqB) (le z zbEqC') = le (y +N z) p
|
||||
where
|
||||
zbEqC : z +N b ≡ c
|
||||
zSuccYAEqC : z +N (succ y +N a) ≡ c
|
||||
zSuccYAEqC' : z +N (succ (y +N a)) ≡ c
|
||||
zSuccYAEqC'' : succ (z +N (y +N a)) ≡ c
|
||||
zSuccYAEqC''' : succ ((z +N y) +N a) ≡ c
|
||||
p : succ ((y +N z) +N a) ≡ c
|
||||
p = identityOfIndiscernablesLeft _≡_ zSuccYAEqC''' (applyEquality (λ n → succ (n +N a)) (additionNIsCommutative z y))
|
||||
zSuccYAEqC''' = identityOfIndiscernablesLeft _≡_ zSuccYAEqC'' (equalityCommutative (applyEquality succ (additionNIsAssociative z y a)))
|
||||
zSuccYAEqC'' = identityOfIndiscernablesLeft _≡_ zSuccYAEqC' (succExtracts z (y +N a))
|
||||
zSuccYAEqC' = identityOfIndiscernablesLeft _≡_ zSuccYAEqC (applyEquality (λ r → z +N r) refl)
|
||||
zbEqC = succInjective zbEqC'
|
||||
zSuccYAEqC = identityOfIndiscernablesLeft _≡_ zbEqC (applyEquality (λ r → z +N r) (equalityCommutative succYAeqB))
|
||||
go : ∀ n m → m <N n → Accessible _<N_ m
|
||||
go zero m (le x ())
|
||||
go (succ n) zero mLessN = access (λ y ())
|
||||
go (succ n) (succ m) smLessSN = access (λ o (oLessSM : o <N succ m) → go n o (lemma oLessSM smLessSN))
|
||||
|
||||
lessImpliesNotEqual : {a b : ℕ} → (a <N b) → a ≡ b → False
|
||||
lessImpliesNotEqual {a} {.a} prAB refl = orderIsIrreflexive prAB prAB
|
||||
|
||||
-NIsDecreasing : {a b : ℕ} → (prAB : succ a <N b) → subtractionNResult.result (-N (inl prAB)) <N b
|
||||
-NIsDecreasing {a} {b} prAB with (-N (inl prAB))
|
||||
-NIsDecreasing {a} {b} (le x proof) | record { result = result ; pr = pr } = record { x = a ; proof = pr }
|
||||
|
||||
equalityN : (a b : ℕ) → Sg Bool (λ truth → if truth then a ≡ b else Unit)
|
||||
equalityN zero zero = ( BoolTrue , refl )
|
||||
equalityN zero (succ b) = ( BoolFalse , record {} )
|
||||
equalityN (succ a) zero = ( BoolFalse , record {} )
|
||||
equalityN (succ a) (succ b) with equalityN a b
|
||||
equalityN (succ a) (succ b) | BoolTrue , val = (BoolTrue , applyEquality succ val)
|
||||
equalityN (succ a) (succ b) | BoolFalse , val = (BoolFalse , record {})
|
||||
|
||||
sumZeroImpliesSummandsZero : {a b : ℕ} → (a +N b ≡ zero) → ((a ≡ zero) && (b ≡ zero))
|
||||
sumZeroImpliesSummandsZero {zero} {zero} pr = record { fst = refl ; snd = refl }
|
||||
sumZeroImpliesSummandsZero {zero} {(succ b)} pr = record { fst = refl ; snd = pr }
|
||||
sumZeroImpliesSummandsZero {(succ a)} {zero} ()
|
||||
sumZeroImpliesSummandsZero {(succ a)} {(succ b)} ()
|
||||
|
||||
productWithNonzeroZero : (a b : ℕ) → (a *N succ b ≡ zero) → a ≡ zero
|
||||
productWithNonzeroZero zero b pr = refl
|
||||
productWithNonzeroZero (succ a) b ()
|
||||
|
||||
productOneImpliesOperandsOne : {a b : ℕ} → (a *N b ≡ 1) → (a ≡ 1) && (b ≡ 1)
|
||||
productOneImpliesOperandsOne {zero} {b} ()
|
||||
productOneImpliesOperandsOne {succ a} {zero} pr = exFalso absurd''
|
||||
productOneImpliesOperandsOne : {a b : ℕ} → (a *N b ≡ 1) → (a ≡ 1) && (b ≡ 1)
|
||||
productOneImpliesOperandsOne {zero} {b} ()
|
||||
productOneImpliesOperandsOne {succ a} {zero} pr = exFalso absurd''
|
||||
where
|
||||
absurd : zero *N (succ a) ≡ 1
|
||||
absurd' : 0 ≡ 1
|
||||
@@ -293,7 +133,7 @@ module Numbers.Naturals.Naturals where
|
||||
absurd'' = succIsNonzero (equalityCommutative absurd')
|
||||
absurd = identityOfIndiscernablesLeft _≡_ pr (productZeroIsZeroRight a)
|
||||
absurd' = absurd
|
||||
productOneImpliesOperandsOne {succ a} {succ b} pr = record { fst = r' ; snd = (applyEquality succ (_&&_.fst q)) }
|
||||
productOneImpliesOperandsOne {succ a} {succ b} pr = record { fst = r' ; snd = (applyEquality succ (_&&_.fst q)) }
|
||||
where
|
||||
p : b +N a *N succ b ≡ zero
|
||||
p = succInjective pr
|
||||
@@ -304,45 +144,13 @@ module Numbers.Naturals.Naturals where
|
||||
r' : succ a ≡ 1
|
||||
r' = applyEquality succ r
|
||||
|
||||
oneTimesPlusZero : (a : ℕ) → a ≡ a *N succ zero +N zero
|
||||
oneTimesPlusZero a = identityOfIndiscernablesRight _≡_ (equalityCommutative (productWithOneRight a)) (equalityCommutative (addZeroRight (a *N succ zero)))
|
||||
oneTimesPlusZero : (a : ℕ) → a ≡ a *N succ zero +N zero
|
||||
oneTimesPlusZero a = identityOfIndiscernablesRight _≡_ (equalityCommutative (productWithOneRight a)) (equalityCommutative (addZeroRight (a *N succ zero)))
|
||||
|
||||
cancelInequalityLeft : {a b c : ℕ} → a *N b <N a *N c → b <N c
|
||||
cancelInequalityLeft {a} {zero} {zero} (le x proof) rewrite (productZeroIsZeroRight a) = exFalso (succIsNonzero {x +N zero} proof)
|
||||
cancelInequalityLeft {a} {zero} {succ c} pr = succIsPositive c
|
||||
cancelInequalityLeft {a} {succ b} {zero} (le x proof) rewrite (productZeroIsZeroRight a) = exFalso (succIsNonzero {x +N a *N succ b} proof)
|
||||
cancelInequalityLeft {a} {succ b} {succ c} pr = succPreservesInequality q'
|
||||
where
|
||||
p' : succ b *N a <N succ c *N a
|
||||
p' = canFlipMultiplicationsInIneq {a} {succ b} {a} {succ c} pr
|
||||
p'' : b *N a +N a <N succ c *N a
|
||||
p'' = identityOfIndiscernablesLeft _<N_ p' (additionNIsCommutative a (b *N a))
|
||||
p''' : b *N a +N a <N c *N a +N a
|
||||
p''' = identityOfIndiscernablesRight _<N_ p'' (additionNIsCommutative a (c *N a))
|
||||
p : b *N a <N c *N a
|
||||
p = subtractionPreservesInequality a p'''
|
||||
q : a *N b <N a *N c
|
||||
q = canFlipMultiplicationsInIneq {b} {a} {c} {a} p
|
||||
q' : b <N c
|
||||
q' = cancelInequalityLeft {a} {b} {c} q
|
||||
|
||||
cannotAddAndEnlarge : (a b : ℕ) → a <N succ (a +N b)
|
||||
cannotAddAndEnlarge a b = le b (applyEquality succ (additionNIsCommutative b a))
|
||||
|
||||
cannotAddAndEnlarge' : {a b : ℕ} → a +N b <N a → False
|
||||
cannotAddAndEnlarge' {a} {zero} pr rewrite addZeroRight a = lessIrreflexive pr
|
||||
cannotAddAndEnlarge' {a} {succ b} pr rewrite (succExtracts a b) = lessIrreflexive {a} (lessTransitive {a} {succ (a +N b)} {a} (cannotAddAndEnlarge a b) pr)
|
||||
|
||||
cannotAddAndEnlarge'' : {a b : ℕ} → a +N succ b ≡ a → False
|
||||
cannotAddAndEnlarge'' {a} {b} pr = naughtE (equalityCommutative inter)
|
||||
where
|
||||
inter : succ b ≡ 0
|
||||
inter rewrite additionNIsCommutative a (succ b) = canSubtractFromEqualityRight pr
|
||||
|
||||
equivalentSubtraction' : (a b c d : ℕ) → (a<c : a <N c) → (d<b : d <N b) → (subtractionNResult.result (-N {a} {c} (inl a<c))) ≡ (subtractionNResult.result (-N {d} {b} (inl d<b))) → a +N b ≡ c +N d
|
||||
equivalentSubtraction' a b c d prac prdb eq with -N (inl prac)
|
||||
equivalentSubtraction' a b c d prac prdb eq | record { result = result ; pr = pr } with -N (inl prdb)
|
||||
equivalentSubtraction' a b c d prac prdb refl | record { result = .result ; pr = pr1 } | record { result = result ; pr = pr } rewrite (equalityCommutative pr) = go
|
||||
equivalentSubtraction' : (a b c d : ℕ) → (a<c : a <N c) → (d<b : d <N b) → (subtractionNResult.result (-N {a} {c} (inl a<c))) ≡ (subtractionNResult.result (-N {d} {b} (inl d<b))) → a +N b ≡ c +N d
|
||||
equivalentSubtraction' a b c d prac prdb eq with -N (inl prac)
|
||||
equivalentSubtraction' a b c d prac prdb eq | record { result = result ; pr = pr } with -N (inl prdb)
|
||||
equivalentSubtraction' a b c d prac prdb refl | record { result = .result ; pr = pr1 } | record { result = result ; pr = pr } rewrite (equalityCommutative pr) = go
|
||||
where
|
||||
go : a +N (d +N result) ≡ c +N d
|
||||
go rewrite (equalityCommutative pr1) = t
|
||||
@@ -350,16 +158,16 @@ module Numbers.Naturals.Naturals where
|
||||
t : a +N (d +N result) ≡ (a +N result) +N d
|
||||
t rewrite (additionNIsAssociative a result d) = applyEquality (λ n → a +N n) (additionNIsCommutative d result)
|
||||
|
||||
lessThanMeansPositiveSubtr : {a b : ℕ} → (a<b : a <N b) → (subtractionNResult.result (-N (inl a<b)) ≡ 0) → False
|
||||
lessThanMeansPositiveSubtr {a} {b} a<b pr with -N (inl a<b)
|
||||
lessThanMeansPositiveSubtr {a} {b} a<b pr | record { result = result ; pr = sub } rewrite pr | addZeroRight a = lessImpliesNotEqual a<b sub
|
||||
lessThanMeansPositiveSubtr : {a b : ℕ} → (a<b : a <N b) → (subtractionNResult.result (-N (inl a<b)) ≡ 0) → False
|
||||
lessThanMeansPositiveSubtr {a} {b} a<b pr with -N (inl a<b)
|
||||
lessThanMeansPositiveSubtr {a} {b} a<b pr | record { result = result ; pr = sub } rewrite pr | addZeroRight a = lessImpliesNotEqual a<b sub
|
||||
|
||||
moveOneSubtraction : {a b c : ℕ} → {a<=b : a ≤N b} → (subtractionNResult.result (-N {a} {b} a<=b)) ≡ c → b ≡ a +N c
|
||||
moveOneSubtraction {a} {b} {zero} {inl a<b} pr rewrite addZeroRight a = exFalso (lessThanMeansPositiveSubtr {a} {b} a<b pr)
|
||||
moveOneSubtraction {a} {b} {succ c} {inl a<b} pr with -N (inl a<b)
|
||||
moveOneSubtraction {a} {b} {succ c} {inl a<b} pr | record { result = result ; pr = sub } rewrite pr | sub = refl
|
||||
moveOneSubtraction {a} {.a} {zero} {inr refl} pr = equalityCommutative (addZeroRight a)
|
||||
moveOneSubtraction {a} {.a} {succ c} {inr refl} pr = identityOfIndiscernablesRight _≡_ (equalityCommutative (addZeroRight a)) (applyEquality (λ t → a +N t) pr')
|
||||
moveOneSubtraction : {a b c : ℕ} → {a<=b : a ≤N b} → (subtractionNResult.result (-N {a} {b} a<=b)) ≡ c → b ≡ a +N c
|
||||
moveOneSubtraction {a} {b} {zero} {inl a<b} pr rewrite addZeroRight a = exFalso (lessThanMeansPositiveSubtr {a} {b} a<b pr)
|
||||
moveOneSubtraction {a} {b} {succ c} {inl a<b} pr with -N (inl a<b)
|
||||
moveOneSubtraction {a} {b} {succ c} {inl a<b} pr | record { result = result ; pr = sub } rewrite pr | sub = refl
|
||||
moveOneSubtraction {a} {.a} {zero} {inr refl} pr = equalityCommutative (addZeroRight a)
|
||||
moveOneSubtraction {a} {.a} {succ c} {inr refl} pr = identityOfIndiscernablesRight _≡_ (equalityCommutative (addZeroRight a)) (applyEquality (λ t → a +N t) pr')
|
||||
where
|
||||
selfSub : (r : ℕ) → subtractionNResult.result (-N {r} {r} (inr refl)) ≡ zero
|
||||
selfSub zero = refl
|
||||
@@ -367,15 +175,15 @@ module Numbers.Naturals.Naturals where
|
||||
pr' : 0 ≡ succ c
|
||||
pr' = transitivity (equalityCommutative (selfSub a)) pr
|
||||
|
||||
moveOneSubtraction' : {a b c : ℕ} → {a<=b : a ≤N b} → (b ≡ a +N c) → subtractionNResult.result (-N {a} {b} a<=b) ≡ c
|
||||
moveOneSubtraction' {a} {b} {c} {inl x} pr with -N (inl x)
|
||||
moveOneSubtraction' {a} {b} {c} {inl x} pr | record { result = result ; pr = pr1 } rewrite pr = canSubtractFromEqualityLeft pr1
|
||||
moveOneSubtraction' {a} {b} {c} {inr x} pr with -N (inr x)
|
||||
moveOneSubtraction' {a} {b} {c} {inr x} pr | record { result = result ; pr = pr1 } rewrite pr = canSubtractFromEqualityLeft pr1
|
||||
moveOneSubtraction' : {a b c : ℕ} → {a<=b : a ≤N b} → (b ≡ a +N c) → subtractionNResult.result (-N {a} {b} a<=b) ≡ c
|
||||
moveOneSubtraction' {a} {b} {c} {inl x} pr with -N (inl x)
|
||||
moveOneSubtraction' {a} {b} {c} {inl x} pr | record { result = result ; pr = pr1 } rewrite pr = canSubtractFromEqualityLeft pr1
|
||||
moveOneSubtraction' {a} {b} {c} {inr x} pr with -N (inr x)
|
||||
moveOneSubtraction' {a} {b} {c} {inr x} pr | record { result = result ; pr = pr1 } rewrite pr = canSubtractFromEqualityLeft pr1
|
||||
|
||||
equivalentSubtraction : (a b c d : ℕ) → (a<c : a <N c) → (d<b : d <N b) → a +N b ≡ c +N d → (subtractionNResult.result (-N {a} {c} (inl a<c))) ≡ (subtractionNResult.result (-N {d} {b} (inl d<b)))
|
||||
equivalentSubtraction zero b c d prac (le x proof) eq with (-N (inl (le x proof)))
|
||||
equivalentSubtraction zero b c d prac (le x proof) eq | record { result = result ; pr = pr } = equalityCommutative p''
|
||||
equivalentSubtraction : (a b c d : ℕ) → (a<c : a <N c) → (d<b : d <N b) → a +N b ≡ c +N d → (subtractionNResult.result (-N {a} {c} (inl a<c))) ≡ (subtractionNResult.result (-N {d} {b} (inl d<b)))
|
||||
equivalentSubtraction zero b c d prac (le x proof) eq with (-N (inl (le x proof)))
|
||||
equivalentSubtraction zero b c d prac (le x proof) eq | record { result = result ; pr = pr } = equalityCommutative p''
|
||||
where
|
||||
p : d +N result ≡ c +N d
|
||||
p = transitivity pr eq
|
||||
@@ -383,45 +191,45 @@ module Numbers.Naturals.Naturals where
|
||||
p' = transitivity p (additionNIsCommutative c d)
|
||||
p'' : result ≡ c
|
||||
p'' = canSubtractFromEqualityLeft {d} {result} {c} p'
|
||||
equivalentSubtraction (succ a) b zero d (le x ()) prdb eq
|
||||
equivalentSubtraction (succ a) b (succ c) d prac prdb eq with (-N (inl (canRemoveSuccFrom<N prac)))
|
||||
equivalentSubtraction (succ a) b (succ c) d prac prdb eq | record { result = c-a ; pr = prc-a } with -N (inl prdb)
|
||||
equivalentSubtraction (succ a) b (succ c) d prac prdb eq | record { result = c-a ; pr = prc-a } | record { result = result ; pr = pr } rewrite equalityCommutative prc-a | equalityCommutative pr | equalityCommutative (additionNIsAssociative a d result) | additionNIsCommutative (a +N c-a) d | equalityCommutative (additionNIsAssociative d a c-a) | additionNIsCommutative a d = equalityCommutative (canSubtractFromEqualityLeft eq)
|
||||
equivalentSubtraction (succ a) b zero d (le x ()) prdb eq
|
||||
equivalentSubtraction (succ a) b (succ c) d prac prdb eq with (-N (inl (canRemoveSuccFrom<N prac)))
|
||||
equivalentSubtraction (succ a) b (succ c) d prac prdb eq | record { result = c-a ; pr = prc-a } with -N (inl prdb)
|
||||
equivalentSubtraction (succ a) b (succ c) d prac prdb eq | record { result = c-a ; pr = prc-a } | record { result = result ; pr = pr } rewrite equalityCommutative prc-a | equalityCommutative pr | equalityCommutative (additionNIsAssociative a d result) | additionNIsCommutative (a +N c-a) d | equalityCommutative (additionNIsAssociative d a c-a) | additionNIsCommutative a d = equalityCommutative (canSubtractFromEqualityLeft eq)
|
||||
|
||||
leLemma : (b c : ℕ) → (b ≤N c) ≡ (b +N zero ≤N c +N zero)
|
||||
leLemma b c rewrite addZeroRight c = q
|
||||
leLemma : (b c : ℕ) → (b ≤N c) ≡ (b +N zero ≤N c +N zero)
|
||||
leLemma b c rewrite addZeroRight c = q
|
||||
where
|
||||
q : (b ≤N c) ≡ (b +N zero ≤N c)
|
||||
q rewrite addZeroRight b = refl
|
||||
|
||||
lessCast : {a b c : ℕ} → (pr : a ≤N b) → (eq : a ≡ c) → c ≤N b
|
||||
lessCast {a} {b} pr eq rewrite eq = pr
|
||||
lessCast : {a b c : ℕ} → (pr : a ≤N b) → (eq : a ≡ c) → c ≤N b
|
||||
lessCast {a} {b} pr eq rewrite eq = pr
|
||||
|
||||
lessCast' : {a b c : ℕ} → (pr : a ≤N b) → (eq : b ≡ c) → a ≤N c
|
||||
lessCast' {a} {b} pr eq rewrite eq = pr
|
||||
lessCast' : {a b c : ℕ} → (pr : a ≤N b) → (eq : b ≡ c) → a ≤N c
|
||||
lessCast' {a} {b} pr eq rewrite eq = pr
|
||||
|
||||
subtractionCast : {a b c : ℕ} → {pr : a ≤N b} → (eq : a ≡ c) → (p : subtractionNResult a b pr) → Sg (subtractionNResult c b (lessCast pr eq)) (λ res → subtractionNResult.result p ≡ subtractionNResult.result res)
|
||||
subtractionCast {a} {b} {c} {a<b} eq subt rewrite eq = (subt , refl)
|
||||
subtractionCast : {a b c : ℕ} → {pr : a ≤N b} → (eq : a ≡ c) → (p : subtractionNResult a b pr) → Sg (subtractionNResult c b (lessCast pr eq)) (λ res → subtractionNResult.result p ≡ subtractionNResult.result res)
|
||||
subtractionCast {a} {b} {c} {a<b} eq subt rewrite eq = (subt , refl)
|
||||
|
||||
subtractionCast' : {a b c : ℕ} → {pr : a ≤N b} → (eq : b ≡ c) → (p : subtractionNResult a b pr) → Sg (subtractionNResult a c (lessCast' pr eq)) (λ res → subtractionNResult.result p ≡ subtractionNResult.result res)
|
||||
subtractionCast' {a} {b} {c} {a<b} eq subt rewrite eq = (subt , refl)
|
||||
subtractionCast' : {a b c : ℕ} → {pr : a ≤N b} → (eq : b ≡ c) → (p : subtractionNResult a b pr) → Sg (subtractionNResult a c (lessCast' pr eq)) (λ res → subtractionNResult.result p ≡ subtractionNResult.result res)
|
||||
subtractionCast' {a} {b} {c} {a<b} eq subt rewrite eq = (subt , refl)
|
||||
|
||||
addToRightWeakInequality : (a : ℕ) → {b c : ℕ} → (pr : b ≤N c) → (b ≤N c +N a)
|
||||
addToRightWeakInequality zero {b} {c} (inl x) rewrite (addZeroRight c) = inl x
|
||||
addToRightWeakInequality (succ a) {b} {c} (inl x) = inl (orderIsTransitive x (addingIncreases c a))
|
||||
addToRightWeakInequality zero {b} {.b} (inr refl) = inr (equalityCommutative (addZeroRight b))
|
||||
addToRightWeakInequality (succ a) {b} {.b} (inr refl) = inl (addingIncreases b a)
|
||||
addToRightWeakInequality : (a : ℕ) → {b c : ℕ} → (pr : b ≤N c) → (b ≤N c +N a)
|
||||
addToRightWeakInequality zero {b} {c} (inl x) rewrite (addZeroRight c) = inl x
|
||||
addToRightWeakInequality (succ a) {b} {c} (inl x) = inl (TotalOrder.<Transitive ℕTotalOrder x (addingIncreases c a))
|
||||
addToRightWeakInequality zero {b} {.b} (inr refl) = inr (equalityCommutative (addZeroRight b))
|
||||
addToRightWeakInequality (succ a) {b} {.b} (inr refl) = inl (addingIncreases b a)
|
||||
|
||||
addAssocLemma : (a b c : ℕ) → (a +N b) +N c ≡ (a +N c) +N b
|
||||
addAssocLemma a b c rewrite (additionNIsAssociative a b c) = g
|
||||
addAssocLemma : (a b c : ℕ) → (a +N b) +N c ≡ (a +N c) +N b
|
||||
addAssocLemma a b c rewrite (additionNIsAssociative a b c) = g
|
||||
where
|
||||
g : a +N (b +N c) ≡ (a +N c) +N b
|
||||
g rewrite (additionNIsAssociative a c b) = applyEquality (λ t → a +N t) (additionNIsCommutative b c)
|
||||
|
||||
addIntoSubtraction : (a : ℕ) → {b c : ℕ} → (pr : b ≤N c) → a +N (subtractionNResult.result (-N {b} {c} pr)) ≡ subtractionNResult.result (-N {b} {c +N a} (addToRightWeakInequality a pr))
|
||||
addIntoSubtraction a {b} {c} pr with (-N {b} {c} pr)
|
||||
addIntoSubtraction a {b} {c} pr | record { result = c-b ; pr = prc-b } with (-N {b} {c +N a} (addToRightWeakInequality a pr))
|
||||
addIntoSubtraction a {b} {c} pr | record { result = c-b ; pr = prc-b } | record { result = c+a-b ; pr = prcab } = equalityCommutative g'''
|
||||
addIntoSubtraction : (a : ℕ) → {b c : ℕ} → (pr : b ≤N c) → a +N (subtractionNResult.result (-N {b} {c} pr)) ≡ subtractionNResult.result (-N {b} {c +N a} (addToRightWeakInequality a pr))
|
||||
addIntoSubtraction a {b} {c} pr with (-N {b} {c} pr)
|
||||
addIntoSubtraction a {b} {c} pr | record { result = c-b ; pr = prc-b } with (-N {b} {c +N a} (addToRightWeakInequality a pr))
|
||||
addIntoSubtraction a {b} {c} pr | record { result = c-b ; pr = prc-b } | record { result = c+a-b ; pr = prcab } = equalityCommutative g'''
|
||||
where
|
||||
g : (b +N c+a-b) +N c-b ≡ c +N (a +N c-b)
|
||||
g rewrite (equalityCommutative (additionNIsAssociative c a c-b)) = applyEquality (λ t → t +N c-b) prcab
|
||||
@@ -432,23 +240,23 @@ module Numbers.Naturals.Naturals where
|
||||
g''' : c+a-b ≡ a +N c-b
|
||||
g''' = canSubtractFromEqualityLeft {c} {c+a-b} {a +N c-b} g''
|
||||
|
||||
addStrongInequalities : {a b c d : ℕ} → (a<b : a <N b) → (c<d : c <N d) → (a +N c <N b +N d)
|
||||
addStrongInequalities {zero} {zero} {c} {d} prab prcd = prcd
|
||||
addStrongInequalities {zero} {succ b} {c} {d} prab prcd rewrite (additionNIsCommutative b d) = orderIsTransitive prcd (cannotAddAndEnlarge d b)
|
||||
addStrongInequalities {succ a} {zero} {c} {d} (le x ()) prcd
|
||||
addStrongInequalities {succ a} {succ b} {c} {d} prab prcd = succPreservesInequality (addStrongInequalities (canRemoveSuccFrom<N prab) prcd)
|
||||
addStrongInequalities : {a b c d : ℕ} → (a<b : a <N b) → (c<d : c <N d) → (a +N c <N b +N d)
|
||||
addStrongInequalities {zero} {zero} {c} {d} prab prcd = prcd
|
||||
addStrongInequalities {zero} {succ b} {c} {d} prab prcd rewrite (additionNIsCommutative b d) = TotalOrder.<Transitive ℕTotalOrder prcd (cannotAddAndEnlarge d b)
|
||||
addStrongInequalities {succ a} {zero} {c} {d} (le x ()) prcd
|
||||
addStrongInequalities {succ a} {succ b} {c} {d} prab prcd = succPreservesInequality (addStrongInequalities (canRemoveSuccFrom<N prab) prcd)
|
||||
|
||||
addWeakInequalities : {a b c d : ℕ} → (a<=b : a ≤N b) → (c<=d : c ≤N d) → (a +N c) ≤N (b +N d)
|
||||
addWeakInequalities {a} {b} {c} {d} (inl x) (inl y) = inl (addStrongInequalities x y)
|
||||
addWeakInequalities {a} {b} {c} {.c} (inl x) (inr refl) = inl (additionPreservesInequality c x)
|
||||
addWeakInequalities {a} {.a} {c} {d} (inr refl) (inl x) = inl (additionPreservesInequalityOnLeft a x)
|
||||
addWeakInequalities {a} {.a} {c} {.c} (inr refl) (inr refl) = inr refl
|
||||
addWeakInequalities : {a b c d : ℕ} → (a<=b : a ≤N b) → (c<=d : c ≤N d) → (a +N c) ≤N (b +N d)
|
||||
addWeakInequalities {a} {b} {c} {d} (inl x) (inl y) = inl (addStrongInequalities x y)
|
||||
addWeakInequalities {a} {b} {c} {.c} (inl x) (inr refl) = inl (additionPreservesInequality c x)
|
||||
addWeakInequalities {a} {.a} {c} {d} (inr refl) (inl x) = inl (additionPreservesInequalityOnLeft a x)
|
||||
addWeakInequalities {a} {.a} {c} {.c} (inr refl) (inr refl) = inr refl
|
||||
|
||||
addSubIntoSub : {a b c d : ℕ} → (a<b : a ≤N b) → (c<d : c ≤N d) → (subtractionNResult.result (-N {a} {b} a<b)) +N (subtractionNResult.result (-N {c} {d} c<d)) ≡ subtractionNResult.result (-N {a +N c} {b +N d} (addWeakInequalities a<b c<d))
|
||||
addSubIntoSub {a}{b}{c}{d} a<b c<d with (-N {a} {b} a<b)
|
||||
addSubIntoSub {a} {b} {c} {d} a<b c<d | record { result = b-a ; pr = prb-a } with (-N {c} {d} c<d)
|
||||
addSubIntoSub {a} {b} {c} {d} a<b c<d | record { result = b-a ; pr = prb-a } | record { result = d-c ; pr = prd-c } with (-N {a +N c} {b +N d} (addWeakInequalities a<b c<d))
|
||||
addSubIntoSub {a} {b} {c} {d} a<b c<d | record { result = b-a ; pr = prb-a } | record { result = d-c ; pr = prd-c } | record { result = b+d-a-c ; pr = pr } = equalityCommutative r
|
||||
addSubIntoSub : {a b c d : ℕ} → (a<b : a ≤N b) → (c<d : c ≤N d) → (subtractionNResult.result (-N {a} {b} a<b)) +N (subtractionNResult.result (-N {c} {d} c<d)) ≡ subtractionNResult.result (-N {a +N c} {b +N d} (addWeakInequalities a<b c<d))
|
||||
addSubIntoSub {a}{b}{c}{d} a<b c<d with (-N {a} {b} a<b)
|
||||
addSubIntoSub {a} {b} {c} {d} a<b c<d | record { result = b-a ; pr = prb-a } with (-N {c} {d} c<d)
|
||||
addSubIntoSub {a} {b} {c} {d} a<b c<d | record { result = b-a ; pr = prb-a } | record { result = d-c ; pr = prd-c } with (-N {a +N c} {b +N d} (addWeakInequalities a<b c<d))
|
||||
addSubIntoSub {a} {b} {c} {d} a<b c<d | record { result = b-a ; pr = prb-a } | record { result = d-c ; pr = prd-c } | record { result = b+d-a-c ; pr = pr } = equalityCommutative r
|
||||
where
|
||||
pr' : (a +N c) +N b+d-a-c ≡ (a +N b-a) +N d
|
||||
pr' rewrite prb-a = pr
|
||||
@@ -470,10 +278,10 @@ module Numbers.Naturals.Naturals where
|
||||
r = canSubtractFromEqualityLeft {c} q''''
|
||||
|
||||
|
||||
subtractProduct : {a b c : ℕ} → (aPos : 0 <N a) → (b<c : b <N c) →
|
||||
subtractProduct : {a b c : ℕ} → (aPos : 0 <N a) → (b<c : b <N c) →
|
||||
(a *N (subtractionNResult.result (-N (inl b<c)))) ≡ subtractionNResult.result (-N {a *N b} {a *N c} (inl (lessRespectsMultiplicationLeft b c a b<c aPos)))
|
||||
subtractProduct {zero} {b} {c} aPos b<c = refl
|
||||
subtractProduct {succ zero} {b} {c} aPos b<c = s'
|
||||
subtractProduct {zero} {b} {c} aPos b<c = refl
|
||||
subtractProduct {succ zero} {b} {c} aPos b<c = s'
|
||||
where
|
||||
resbc = -N {b} {c} (inl b<c)
|
||||
resbc' : Sg (subtractionNResult (b +N 0 *N b) c (lessCast (inl b<c) (equalityCommutative (addZeroRight b)))) ((λ res → subtractionNResult.result resbc ≡ subtractionNResult.result res))
|
||||
@@ -493,7 +301,7 @@ module Numbers.Naturals.Naturals where
|
||||
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
|
||||
|
||||
subtractProduct {succ (succ a)} {b} {c} aPos b<c =
|
||||
subtractProduct {succ (succ a)} {b} {c} aPos b<c =
|
||||
let
|
||||
t : (succ a) *N subtractionNResult.result (-N {b} {c} (inl b<c)) ≡ subtractionNResult.result (-N {(succ a) *N b} {(succ a) *N c} (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a))))
|
||||
t = subtractProduct {succ a} {b} {c} (succIsPositive a) b<c
|
||||
@@ -520,46 +328,23 @@ module Numbers.Naturals.Naturals where
|
||||
u = addSubIntoSub (inl b<c) (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a)))
|
||||
v = subtractionNWellDefined {succ (succ a) *N b} {succ (succ a) *N c} {addWeakInequalities (inl b<c) (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a)))} {inl (lessRespectsMultiplicationLeft b c (succ (succ a)) b<c aPos)} (-N {b +N (succ a *N b)} {c +N (succ a *N c)} (addWeakInequalities (inl b<c) (inl (lessRespectsMultiplicationLeft b c (succ a) b<c (succIsPositive a))))) (-N {(succ (succ a)) *N b} {(succ (succ a)) *N c} (inl (lessRespectsMultiplicationLeft b c (succ (succ a)) b<c aPos)))
|
||||
|
||||
subtractProduct' : {a b c : ℕ} → (aPos : 0 <N a) → (b<c : b <N c) →
|
||||
subtractProduct' : {a b c : ℕ} → (aPos : 0 <N a) → (b<c : b <N c) →
|
||||
(subtractionNResult.result (-N (inl b<c))) *N a ≡ subtractionNResult.result (-N {a *N b} {a *N c} (inl (lessRespectsMultiplicationLeft b c a b<c aPos)))
|
||||
subtractProduct' {a} aPos b<c = identityOfIndiscernablesLeft _≡_ (subtractProduct aPos b<c) (multiplicationNIsCommutative a _)
|
||||
subtractProduct' {a} aPos b<c = identityOfIndiscernablesLeft _≡_ (subtractProduct aPos b<c) (multiplicationNIsCommutative a _)
|
||||
|
||||
equalityDecidable : (a b : ℕ) → (a ≡ b) || ((a ≡ b) → False)
|
||||
equalityDecidable zero zero = inl refl
|
||||
equalityDecidable zero (succ b) = inr naughtE
|
||||
equalityDecidable (succ a) zero = inr λ t → naughtE (equalityCommutative t)
|
||||
equalityDecidable (succ a) (succ b) with equalityDecidable a b
|
||||
equalityDecidable (succ a) (succ b) | inl x = inl (applyEquality succ x)
|
||||
equalityDecidable (succ a) (succ b) | inr x = inr (λ t → x (succInjective t))
|
||||
equalityDecidable : (a b : ℕ) → (a ≡ b) || ((a ≡ b) → False)
|
||||
equalityDecidable zero zero = inl refl
|
||||
equalityDecidable zero (succ b) = inr naughtE
|
||||
equalityDecidable (succ a) zero = inr λ t → naughtE (equalityCommutative t)
|
||||
equalityDecidable (succ a) (succ b) with equalityDecidable a b
|
||||
equalityDecidable (succ a) (succ b) | inl x = inl (applyEquality succ x)
|
||||
equalityDecidable (succ a) (succ b) | inr x = inr (λ t → x (succInjective t))
|
||||
|
||||
cannotAddKeepingEquality : (a b : ℕ) → (a ≡ a +N succ b) → False
|
||||
cannotAddKeepingEquality zero zero pr = naughtE pr
|
||||
cannotAddKeepingEquality (succ a) zero pr = cannotAddKeepingEquality a zero (succInjective pr)
|
||||
cannotAddKeepingEquality zero (succ b) pr = naughtE pr
|
||||
cannotAddKeepingEquality (succ a) (succ b) pr = cannotAddKeepingEquality a (succ b) (succInjective pr)
|
||||
cannotAddKeepingEquality : (a b : ℕ) → (a ≡ a +N succ b) → False
|
||||
cannotAddKeepingEquality zero zero pr = naughtE pr
|
||||
cannotAddKeepingEquality (succ a) zero pr = cannotAddKeepingEquality a zero (succInjective pr)
|
||||
cannotAddKeepingEquality zero (succ b) pr = naughtE pr
|
||||
cannotAddKeepingEquality (succ a) (succ b) pr = cannotAddKeepingEquality a (succ b) (succInjective pr)
|
||||
|
||||
ℕTotalOrder : TotalOrder ℕ
|
||||
PartialOrder._<_ (TotalOrder.order ℕTotalOrder) = _<N_
|
||||
PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) = lessIrreflexive
|
||||
PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) = orderIsTransitive
|
||||
TotalOrder.totality ℕTotalOrder = orderIsTotal
|
||||
|
||||
doubleIsAddTwo : (a : ℕ) → a +N a ≡ 2 *N a
|
||||
doubleIsAddTwo a rewrite additionNIsCommutative a 0 = refl
|
||||
|
||||
productZeroImpliesOperandZero : {a b : ℕ} → a *N b ≡ 0 → (a ≡ 0) || (b ≡ 0)
|
||||
productZeroImpliesOperandZero {zero} {b} pr = inl refl
|
||||
productZeroImpliesOperandZero {succ a} {zero} pr = inr refl
|
||||
productZeroImpliesOperandZero {succ a} {succ b} ()
|
||||
|
||||
sumZeroImpliesOperandsZero : (a : ℕ) {b : ℕ} → a +N b ≡ 0 → (a ≡ 0) && (b ≡ 0)
|
||||
sumZeroImpliesOperandsZero zero {zero} pr = refl ,, refl
|
||||
|
||||
inequalityShrinkRight : {a b c : ℕ} → a +N b <N c → b <N c
|
||||
inequalityShrinkRight {a} {b} {c} (le x proof) = le (x +N a) (transitivity (applyEquality succ (additionNIsAssociative x a b)) proof)
|
||||
|
||||
inequalityShrinkLeft : {a b c : ℕ} → a +N b <N c → a <N c
|
||||
inequalityShrinkLeft {a} {b} {c} (le x proof) = le (x +N b) (transitivity (applyEquality succ (transitivity (additionNIsAssociative x b a) (applyEquality (x +N_) (additionNIsCommutative b a)))) proof)
|
||||
|
||||
<NWellDefined' : {a b c d : ℕ} → a ≡ c → b ≡ d → a <N b → c <N d
|
||||
<NWellDefined' {a} {b} {c} {d} a=c b=d a<b rewrite a=c | b=d = a<b
|
||||
<NWellDefined' : {a b c d : ℕ} → a ≡ c → b ≡ d → a <N b → c <N d
|
||||
<NWellDefined' {a} {b} {c} {d} a=c b=d a<b rewrite a=c | b=d = a<b
|
||||
|
@@ -1,17 +1,21 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Semirings.Definition
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Addition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Orders
|
||||
|
||||
module Numbers.Naturals.Order where
|
||||
open Semiring ℕSemiring
|
||||
|
||||
infix 5 _<NLogical_
|
||||
_<NLogical_ : ℕ → ℕ → Set
|
||||
zero <NLogical zero = False
|
||||
zero <NLogical (succ n) = True
|
||||
(succ n) <NLogical zero = False
|
||||
(succ n) <NLogical (succ m) = n <NLogical m
|
||||
private
|
||||
infix 5 _<NLogical_
|
||||
_<NLogical_ : ℕ → ℕ → Set
|
||||
zero <NLogical zero = False
|
||||
zero <NLogical (succ n) = True
|
||||
(succ n) <NLogical zero = False
|
||||
(succ n) <NLogical (succ m) = n <NLogical m
|
||||
|
||||
infix 5 _<N_
|
||||
record _<N_ (a : ℕ) (b : ℕ) : Set where
|
||||
@@ -25,42 +29,43 @@ _≤N_ : ℕ → ℕ → Set
|
||||
a ≤N b = (a <N b) || (a ≡ b)
|
||||
|
||||
succIsPositive : (a : ℕ) → zero <N succ a
|
||||
succIsPositive a = le a (applyEquality succ (additionNIsCommutative a 0))
|
||||
succIsPositive a = le a (applyEquality succ (Semiring.commutative ℕSemiring a 0))
|
||||
|
||||
aLessSucc : (a : ℕ) → (a <NLogical succ a)
|
||||
aLessSucc zero = record {}
|
||||
aLessSucc (succ a) = aLessSucc a
|
||||
|
||||
succPreservesInequalityLogical : {a b : ℕ} → (a <NLogical b) → (succ a <NLogical succ b)
|
||||
succPreservesInequalityLogical {a} {b} prAB = prAB
|
||||
private
|
||||
succPreservesInequalityLogical : {a b : ℕ} → (a <NLogical b) → (succ a <NLogical succ b)
|
||||
succPreservesInequalityLogical {a} {b} prAB = prAB
|
||||
|
||||
lessTransitiveLogical : {a b c : ℕ} → (a <NLogical b) → (b <NLogical c) → (a <NLogical c)
|
||||
lessTransitiveLogical {a} {zero} {zero} prAB prBC = prAB
|
||||
lessTransitiveLogical {a} {(succ b)} {zero} prAB ()
|
||||
lessTransitiveLogical {zero} {zero} {(succ c)} prAB prBC = record {}
|
||||
lessTransitiveLogical {(succ a)} {zero} {(succ c)} () prBC
|
||||
lessTransitiveLogical {zero} {(succ b)} {(succ c)} prAB prBC = record {}
|
||||
lessTransitiveLogical {(succ a)} {(succ b)} {(succ c)} prAB prBC = lessTransitiveLogical {a} {b} {c} prAB prBC
|
||||
lessTransitiveLogical : {a b c : ℕ} → (a <NLogical b) → (b <NLogical c) → (a <NLogical c)
|
||||
lessTransitiveLogical {a} {zero} {zero} prAB prBC = prAB
|
||||
lessTransitiveLogical {a} {(succ b)} {zero} prAB ()
|
||||
lessTransitiveLogical {zero} {zero} {(succ c)} prAB prBC = record {}
|
||||
lessTransitiveLogical {(succ a)} {zero} {(succ c)} () prBC
|
||||
lessTransitiveLogical {zero} {(succ b)} {(succ c)} prAB prBC = record {}
|
||||
lessTransitiveLogical {(succ a)} {(succ b)} {(succ c)} prAB prBC = lessTransitiveLogical {a} {b} {c} prAB prBC
|
||||
|
||||
aLessXPlusSuccA : (a x : ℕ) → (a <NLogical x +N succ a)
|
||||
aLessXPlusSuccA a zero = aLessSucc a
|
||||
aLessXPlusSuccA zero (succ x) = record {}
|
||||
aLessXPlusSuccA (succ a) (succ x) = lessTransitiveLogical {a} {succ a} {x +N succ (succ a)} (aLessXPlusSuccA a zero) (aLessXPlusSuccA (succ a) x)
|
||||
aLessXPlusSuccA : (a x : ℕ) → (a <NLogical x +N succ a)
|
||||
aLessXPlusSuccA a zero = aLessSucc a
|
||||
aLessXPlusSuccA zero (succ x) = record {}
|
||||
aLessXPlusSuccA (succ a) (succ x) = lessTransitiveLogical {a} {succ a} {x +N succ (succ a)} (aLessXPlusSuccA a zero) (aLessXPlusSuccA (succ a) x)
|
||||
|
||||
leImpliesLogical<N : {a b : ℕ} → (a <N b) → (a <NLogical b)
|
||||
leImpliesLogical<N {zero} {zero} (le x ())
|
||||
leImpliesLogical<N {zero} {(succ b)} (le x proof) = record {}
|
||||
leImpliesLogical<N {(succ a)} {zero} (le x ())
|
||||
leImpliesLogical<N {(succ a)} {(succ .(succ a))} (le zero refl) = aLessSucc a
|
||||
leImpliesLogical<N {(succ a)} {(succ .(succ (x +N succ a)))} (le (succ x) refl) = succPreservesInequalityLogical {a} {succ (x +N succ a)} (lessTransitiveLogical {a} {succ a} {succ (x +N succ a)} (aLessSucc a) (aLessXPlusSuccA a x))
|
||||
leImpliesLogical<N : {a b : ℕ} → (a <N b) → (a <NLogical b)
|
||||
leImpliesLogical<N {zero} {zero} (le x ())
|
||||
leImpliesLogical<N {zero} {(succ b)} (le x proof) = record {}
|
||||
leImpliesLogical<N {(succ a)} {zero} (le x ())
|
||||
leImpliesLogical<N {(succ a)} {(succ .(succ a))} (le zero refl) = aLessSucc a
|
||||
leImpliesLogical<N {(succ a)} {(succ .(succ (x +N succ a)))} (le (succ x) refl) = succPreservesInequalityLogical {a} {succ (x +N succ a)} (lessTransitiveLogical {a} {succ a} {succ (x +N succ a)} (aLessSucc a) (aLessXPlusSuccA a x))
|
||||
|
||||
logical<NImpliesLe : {a b : ℕ} → (a <NLogical b) → (a <N b)
|
||||
logical<NImpliesLe {zero} {zero} ()
|
||||
_<N_.x (logical<NImpliesLe {zero} {succ b} prAB) = b
|
||||
_<N_.proof (logical<NImpliesLe {zero} {succ b} prAB) = applyEquality succ (addZeroRight b)
|
||||
logical<NImpliesLe {(succ a)} {zero} ()
|
||||
logical<NImpliesLe {(succ a)} {(succ b)} prAB with logical<NImpliesLe {a} prAB
|
||||
logical<NImpliesLe {(succ a)} {(succ .(succ (x +N a)))} prAB | le x refl = le x (succCanMove (succ x) a)
|
||||
logical<NImpliesLe : {a b : ℕ} → (a <NLogical b) → (a <N b)
|
||||
logical<NImpliesLe {zero} {zero} ()
|
||||
_<N_.x (logical<NImpliesLe {zero} {succ b} prAB) = b
|
||||
_<N_.proof (logical<NImpliesLe {zero} {succ b} prAB) = applyEquality succ (sumZeroRight b)
|
||||
logical<NImpliesLe {(succ a)} {zero} ()
|
||||
logical<NImpliesLe {(succ a)} {(succ b)} prAB with logical<NImpliesLe {a} prAB
|
||||
logical<NImpliesLe {(succ a)} {(succ .(succ (x +N a)))} prAB | le x refl = le x (applyEquality succ (transitivity (commutative x _) (applyEquality succ (commutative a _))))
|
||||
|
||||
lessTransitive : {a b c : ℕ} → (a <N b) → (b <N c) → (a <N c)
|
||||
lessTransitive {a} {b} {c} prab prbc = logical<NImpliesLe (lessTransitiveLogical {a} {b} {c} (leImpliesLogical<N prab) (leImpliesLogical<N prbc))
|
||||
@@ -73,13 +78,13 @@ succPreservesInequality : {a b : ℕ} → (a <N b) → (succ a <N succ b)
|
||||
succPreservesInequality {a} {b} prAB = logical<NImpliesLe (succPreservesInequalityLogical {a} {b} (leImpliesLogical<N prAB))
|
||||
|
||||
canRemoveSuccFrom<N : {a b : ℕ} → (succ a <N succ b) → (a <N b)
|
||||
canRemoveSuccFrom<N {a} {b} (le x proof) rewrite additionNIsCommutative x (succ a) | additionNIsCommutative a x = le x (succInjective proof)
|
||||
canRemoveSuccFrom<N {a} {b} (le x proof) rewrite commutative x (succ a) | commutative a x = le x (succInjective proof)
|
||||
|
||||
a<SuccA : (a : ℕ) → a <N succ a
|
||||
a<SuccA a = le zero refl
|
||||
|
||||
canAddToOneSideOfInequality : {a b : ℕ} (c : ℕ) → a <N b → a <N (b +N c)
|
||||
canAddToOneSideOfInequality {a} {b} c (le x proof) = le (x +N c) (transitivity (applyEquality succ (additionNIsAssociative x c a)) (transitivity (applyEquality (λ i → (succ x) +N i) (additionNIsCommutative c a)) (transitivity (applyEquality succ (equalityCommutative (additionNIsAssociative x a c))) (applyEquality (_+N c) proof))))
|
||||
canAddToOneSideOfInequality {a} {b} c (le x proof) = le (x +N c) (transitivity (applyEquality succ (equalityCommutative (+Associative x c a))) (transitivity (applyEquality (λ i → (succ x) +N i) (commutative c a)) (transitivity (applyEquality succ (+Associative x a c)) (applyEquality (_+N c) proof))))
|
||||
|
||||
addingIncreases : (a b : ℕ) → a <N a +N succ b
|
||||
addingIncreases zero b = succIsPositive b
|
||||
@@ -91,7 +96,7 @@ zeroNeverGreater {a} (le x ())
|
||||
noIntegersBetweenXAndSuccX : {a : ℕ} (x : ℕ) → (x <N a) → (a <N succ x) → False
|
||||
noIntegersBetweenXAndSuccX {zero} x x<a a<sx = zeroNeverGreater x<a
|
||||
noIntegersBetweenXAndSuccX {succ a} x (le y proof) (le z proof1) with succInjective proof1
|
||||
... | ah rewrite (equalityCommutative proof) | (succExtracts z (y +N x)) | equalityCommutative (additionNIsAssociative (succ z) y x) | additionNIsCommutative (succ (z +N y)) x = lessIrreflexive {x} (le (z +N y) (transitivity (additionNIsCommutative _ x) ah))
|
||||
... | ah rewrite (equalityCommutative proof) | (succExtracts z (y +N x)) | +Associative (succ z) y x | commutative (succ (z +N y)) x = lessIrreflexive {x} (le (z +N y) (transitivity (commutative _ x) ah))
|
||||
|
||||
<NWellDefined : {a b : ℕ} → (p1 : a <N b) → (p2 : a <N b) → _<N_.x p1 ≡ _<N_.x p2
|
||||
<NWellDefined {a} {b} (le x proof) (le y proof1) = equalityCommutative r
|
||||
@@ -100,3 +105,80 @@ noIntegersBetweenXAndSuccX {succ a} x (le y proof) (le z proof1) with succInject
|
||||
q = succInjective {y +N a} {x +N a} (transitivity proof1 (equalityCommutative proof))
|
||||
r : y ≡ x
|
||||
r = canSubtractFromEqualityRight q
|
||||
|
||||
private
|
||||
orderIsTotal : (a b : ℕ) → ((a <N b) || (b <N a)) || (a ≡ b)
|
||||
orderIsTotal zero zero = inr refl
|
||||
orderIsTotal zero (succ b) = inl (inl (logical<NImpliesLe (record {})))
|
||||
orderIsTotal (succ a) zero = inl (inr (logical<NImpliesLe (record {})))
|
||||
orderIsTotal (succ a) (succ b) with orderIsTotal a b
|
||||
orderIsTotal (succ a) (succ b) | inl (inl x) = inl (inl (logical<NImpliesLe (leImpliesLogical<N x)))
|
||||
orderIsTotal (succ a) (succ b) | inl (inr x) = inl (inr (logical<NImpliesLe (leImpliesLogical<N x)))
|
||||
orderIsTotal (succ a) (succ b) | inr x = inr (applyEquality succ x)
|
||||
|
||||
orderIsIrreflexive : {a b : ℕ} → (a <N b) → (b <N a) → False
|
||||
orderIsIrreflexive {zero} {b} prAB (le x ())
|
||||
orderIsIrreflexive {(succ a)} {zero} (le x ()) prBA
|
||||
orderIsIrreflexive {(succ a)} {(succ b)} prAB prBA = orderIsIrreflexive {a} {b} (logical<NImpliesLe (leImpliesLogical<N prAB)) (logical<NImpliesLe (leImpliesLogical<N prBA))
|
||||
|
||||
orderIsTransitive : {a b c : ℕ} → (a <N b) → (b <N c) → (a <N c)
|
||||
orderIsTransitive {a} {.(succ (x +N a))} {.(succ (y +N succ (x +N a)))} (le x refl) (le y refl) = le (y +N succ x) (applyEquality succ (equalityCommutative (+Associative y (succ x) a)))
|
||||
|
||||
ℕTotalOrder : TotalOrder ℕ
|
||||
PartialOrder._<_ (TotalOrder.order ℕTotalOrder) = _<N_
|
||||
PartialOrder.irreflexive (TotalOrder.order ℕTotalOrder) = lessIrreflexive
|
||||
PartialOrder.<Transitive (TotalOrder.order ℕTotalOrder) = orderIsTransitive
|
||||
TotalOrder.totality ℕTotalOrder = orderIsTotal
|
||||
|
||||
cannotAddAndEnlarge : (a b : ℕ) → a <N succ (a +N b)
|
||||
cannotAddAndEnlarge a b = le b (applyEquality succ (commutative b a))
|
||||
|
||||
cannotAddAndEnlarge' : {a b : ℕ} → a +N b <N a → False
|
||||
cannotAddAndEnlarge' {a} {zero} pr rewrite sumZeroRight a = lessIrreflexive pr
|
||||
cannotAddAndEnlarge' {a} {succ b} pr rewrite (succExtracts a b) = lessIrreflexive {a} (lessTransitive {a} {succ (a +N b)} {a} (cannotAddAndEnlarge a b) pr)
|
||||
|
||||
cannotAddAndEnlarge'' : {a b : ℕ} → a +N succ b ≡ a → False
|
||||
cannotAddAndEnlarge'' {a} {b} pr = naughtE (equalityCommutative inter)
|
||||
where
|
||||
inter : succ b ≡ 0
|
||||
inter rewrite commutative a (succ b) = canSubtractFromEqualityRight pr
|
||||
|
||||
naturalsAreNonnegative : (a : ℕ) → (a <N zero) → False
|
||||
naturalsAreNonnegative zero ()
|
||||
naturalsAreNonnegative (succ x) ()
|
||||
|
||||
lessRespectsAddition : {a b : ℕ} (c : ℕ) → (a <N b) → ((a +N c) <N (b +N c))
|
||||
lessRespectsAddition {a} {b} zero prAB rewrite commutative a 0 | commutative b 0 = prAB
|
||||
lessRespectsAddition {a} {b} (succ c) prAB rewrite commutative a (succ c) | commutative b (succ c) | commutative c a | commutative c b = succPreservesInequality (lessRespectsAddition c prAB)
|
||||
|
||||
lessRespectsMultiplicationLeft : (a b c : ℕ) → (a <N b) → (zero <N c) → (c *N a <N c *N b)
|
||||
lessRespectsMultiplicationLeft zero zero c (le x ()) cPos
|
||||
lessRespectsMultiplicationLeft zero (succ b) zero prAB (le x ())
|
||||
lessRespectsMultiplicationLeft zero (succ b) (succ c) prAB cPos = i
|
||||
where
|
||||
productNonzeroIsNonzero : {a b : ℕ} → zero <N a → zero <N b → zero <N a *N b
|
||||
productNonzeroIsNonzero {zero} {b} prA prB = prA
|
||||
productNonzeroIsNonzero {succ a} {zero} prA ()
|
||||
productNonzeroIsNonzero {succ a} {succ b} prA prB = le (b +N a *N succ b) (applyEquality succ (Semiring.sumZeroRight ℕSemiring _))
|
||||
|
||||
j : zero <N succ c *N succ b
|
||||
j = productNonzeroIsNonzero cPos prAB
|
||||
i : succ c *N zero <N succ c *N succ b
|
||||
i = identityOfIndiscernablesLeft _<N_ j (equalityCommutative (productZeroRight c))
|
||||
lessRespectsMultiplicationLeft (succ a) zero c (le x ()) cPos
|
||||
lessRespectsMultiplicationLeft (succ a) (succ b) c prAB cPos = m
|
||||
where
|
||||
h : c *N a <N c *N b
|
||||
h = lessRespectsMultiplicationLeft a b c (canRemoveSuccFrom<N prAB) cPos
|
||||
j : c *N a +N c <N c *N b +N c
|
||||
j = lessRespectsAddition c h
|
||||
i : c *N succ a <N c *N b +N c
|
||||
i = identityOfIndiscernablesLeft _<N_ j (equalityCommutative (transitivity (multiplicationNIsCommutative c _) (transitivity (applyEquality (c +N_) (multiplicationNIsCommutative a _)) (commutative c _))))
|
||||
m : c *N succ a <N c *N succ b
|
||||
m = identityOfIndiscernablesRight _<N_ i (equalityCommutative (transitivity (multiplicationNIsCommutative c _) (transitivity (commutative c _) (applyEquality (_+N c) (multiplicationNIsCommutative b _)))))
|
||||
|
||||
canFlipMultiplicationsInIneq : {a b c d : ℕ} → (a *N b <N c *N d) → b *N a <N d *N c
|
||||
canFlipMultiplicationsInIneq {a} {b} {c} {d} pr = identityOfIndiscernablesRight _<N_ (identityOfIndiscernablesLeft _<N_ pr (multiplicationNIsCommutative a b)) (multiplicationNIsCommutative c d)
|
||||
|
||||
lessRespectsMultiplication : (a b c : ℕ) → (a <N b) → (zero <N c) → (a *N c <N b *N c)
|
||||
lessRespectsMultiplication a b c prAB cPos = canFlipMultiplicationsInIneq {c} {a} {c} {b} (lessRespectsMultiplicationLeft a b c prAB cPos)
|
||||
|
83
Numbers/Naturals/Order/Lemmas.agda
Normal file
83
Numbers/Naturals/Order/Lemmas.agda
Normal file
@@ -0,0 +1,83 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Semirings.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Orders
|
||||
|
||||
module Numbers.Naturals.Order.Lemmas where
|
||||
open Semiring ℕSemiring
|
||||
|
||||
inequalityShrinkRight : {a b c : ℕ} → a +N b <N c → b <N c
|
||||
inequalityShrinkRight {a} {b} {c} (le x proof) = le (x +N a) (transitivity (applyEquality succ (equalityCommutative (Semiring.+Associative ℕSemiring x a b))) proof)
|
||||
|
||||
inequalityShrinkLeft : {a b c : ℕ} → a +N b <N c → a <N c
|
||||
inequalityShrinkLeft {a} {b} {c} (le x proof) = le (x +N b) (transitivity (applyEquality succ (transitivity (equalityCommutative (Semiring.+Associative ℕSemiring x b a)) (applyEquality (x +N_) (Semiring.commutative ℕSemiring b a)))) proof)
|
||||
|
||||
productCancelsRight : (a b c : ℕ) → (zero <N a) → (b *N a ≡ c *N a) → (b ≡ c)
|
||||
productCancelsRight a zero zero aPos eq = refl
|
||||
productCancelsRight zero zero (succ c) (le x ()) eq
|
||||
productCancelsRight (succ a) zero (succ c) aPos eq = contr
|
||||
where
|
||||
h : zero ≡ succ c *N succ a
|
||||
h = eq
|
||||
contr : zero ≡ succ c
|
||||
contr = exFalso (naughtE h)
|
||||
productCancelsRight zero (succ b) zero (le x ()) eq
|
||||
productCancelsRight (succ a) (succ b) zero aPos eq = contr
|
||||
where
|
||||
h : succ b *N succ a ≡ zero
|
||||
h = eq
|
||||
contr : succ b ≡ zero
|
||||
contr = exFalso (naughtE (equalityCommutative h))
|
||||
productCancelsRight zero (succ b) (succ c) (le x ()) eq
|
||||
productCancelsRight (succ a) (succ b) (succ c) aPos eq = applyEquality succ (productCancelsRight (succ a) b c aPos l)
|
||||
where
|
||||
i : succ a +N b *N succ a ≡ succ c *N succ a
|
||||
i = eq
|
||||
j : succ c *N succ a ≡ succ a +N c *N succ a
|
||||
j = refl
|
||||
k : succ a +N b *N succ a ≡ succ a +N c *N succ a
|
||||
k = transitivity i j
|
||||
l : b *N succ a ≡ c *N succ a
|
||||
l = canSubtractFromEqualityLeft {succ a} {b *N succ a} {c *N succ a} k
|
||||
|
||||
productCancelsLeft : (a b c : ℕ) → (zero <N a) → (a *N b ≡ a *N c) → (b ≡ c)
|
||||
productCancelsLeft a b c aPos pr = productCancelsRight a b c aPos j
|
||||
where
|
||||
i : b *N a ≡ a *N c
|
||||
i = identityOfIndiscernablesLeft _≡_ pr (multiplicationNIsCommutative a b)
|
||||
j : b *N a ≡ c *N a
|
||||
j = identityOfIndiscernablesRight _≡_ i (multiplicationNIsCommutative a c)
|
||||
|
||||
productCancelsRight' : (a b c : ℕ) → (b *N a ≡ c *N a) → (a ≡ zero) || (b ≡ c)
|
||||
productCancelsRight' zero b c pr = inl refl
|
||||
productCancelsRight' (succ a) b c pr = inr (productCancelsRight (succ a) b c (succIsPositive a) pr)
|
||||
|
||||
productCancelsLeft' : (a b c : ℕ) → (a *N b ≡ a *N c) → (a ≡ zero) || (b ≡ c)
|
||||
productCancelsLeft' zero b c pr = inl refl
|
||||
productCancelsLeft' (succ a) b c pr = inr (productCancelsLeft (succ a) b c (succIsPositive a) pr)
|
||||
|
||||
subtractionPreservesInequality : {a b : ℕ} → (c : ℕ) → a +N c <N b +N c → a <N b
|
||||
subtractionPreservesInequality {a} {b} zero prABC rewrite commutative a 0 | commutative b 0 = prABC
|
||||
subtractionPreservesInequality {a} {b} (succ c) (le x proof) = le x (canSubtractFromEqualityRight {b = succ c} (transitivity (equalityCommutative (+Associative (succ x) a (succ c))) proof))
|
||||
|
||||
cancelInequalityLeft : {a b c : ℕ} → a *N b <N a *N c → b <N c
|
||||
cancelInequalityLeft {a} {zero} {zero} (le x proof) rewrite (productZeroRight a) = exFalso (naughtE (equalityCommutative proof))
|
||||
cancelInequalityLeft {a} {zero} {succ c} pr = succIsPositive c
|
||||
cancelInequalityLeft {a} {succ b} {zero} (le x proof) rewrite (productZeroRight a) = exFalso (naughtE (equalityCommutative proof))
|
||||
cancelInequalityLeft {a} {succ b} {succ c} pr = succPreservesInequality q'
|
||||
where
|
||||
p' : succ b *N a <N succ c *N a
|
||||
p' = canFlipMultiplicationsInIneq {a} {succ b} {a} {succ c} pr
|
||||
p'' : b *N a +N a <N succ c *N a
|
||||
p'' = identityOfIndiscernablesLeft _<N_ p' (commutative a (b *N a))
|
||||
p''' : b *N a +N a <N c *N a +N a
|
||||
p''' = identityOfIndiscernablesRight _<N_ p'' (commutative a (c *N a))
|
||||
p : b *N a <N c *N a
|
||||
p = subtractionPreservesInequality a p'''
|
||||
q : a *N b <N a *N c
|
||||
q = canFlipMultiplicationsInIneq {b} {a} {c} {a} p
|
||||
q' : b <N c
|
||||
q' = cancelInequalityLeft {a} {b} {c} q
|
36
Numbers/Naturals/Order/WellFounded.agda
Normal file
36
Numbers/Naturals/Order/WellFounded.agda
Normal file
@@ -0,0 +1,36 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import WellFoundedInduction
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
|
||||
module Numbers.Naturals.Order.WellFounded where
|
||||
open Semiring ℕSemiring
|
||||
|
||||
<NWellfounded : WellFounded _<N_
|
||||
<NWellfounded = λ x → access (go x)
|
||||
where
|
||||
lemma : {a b c : ℕ} → a <N b → b <N succ c → a <N c
|
||||
lemma {a} {b} {c} (le y succYAeqB) (le z zbEqC') = le (y +N z) p
|
||||
where
|
||||
zbEqC : z +N b ≡ c
|
||||
zSuccYAEqC : z +N (succ y +N a) ≡ c
|
||||
zSuccYAEqC' : z +N (succ (y +N a)) ≡ c
|
||||
zSuccYAEqC'' : succ (z +N (y +N a)) ≡ c
|
||||
zSuccYAEqC''' : succ ((z +N y) +N a) ≡ c
|
||||
p : succ ((y +N z) +N a) ≡ c
|
||||
p = identityOfIndiscernablesLeft _≡_ zSuccYAEqC''' (applyEquality (λ n → succ (n +N a)) (commutative z y))
|
||||
zSuccYAEqC''' = identityOfIndiscernablesLeft _≡_ zSuccYAEqC'' (applyEquality succ (+Associative z y a))
|
||||
zSuccYAEqC'' = identityOfIndiscernablesLeft _≡_ zSuccYAEqC' (succExtracts z (y +N a))
|
||||
zSuccYAEqC' = identityOfIndiscernablesLeft _≡_ zSuccYAEqC (applyEquality (λ r → z +N r) refl)
|
||||
zbEqC = succInjective zbEqC'
|
||||
zSuccYAEqC = identityOfIndiscernablesLeft _≡_ zbEqC (applyEquality (λ r → z +N r) (equalityCommutative succYAeqB))
|
||||
go : ∀ n m → m <N n → Accessible _<N_ m
|
||||
go zero m (le x ())
|
||||
go (succ n) zero mLessN = access (λ y ())
|
||||
go (succ n) (succ m) smLessSN = access (λ o (oLessSM : o <N succ m) → go n o (lemma oLessSM smLessSN))
|
39
Numbers/Naturals/Semiring.agda
Normal file
39
Numbers/Naturals/Semiring.agda
Normal file
@@ -0,0 +1,39 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import WellFoundedInduction
|
||||
open import Functions
|
||||
open import Orders
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Addition
|
||||
open import Numbers.Naturals.Multiplication
|
||||
open import Semirings.Definition
|
||||
open import Monoids.Definition
|
||||
|
||||
module Numbers.Naturals.Semiring where
|
||||
|
||||
open Numbers.Naturals.Definition using (ℕ ; zero ; succ ; succInjective ; naughtE) public
|
||||
open Numbers.Naturals.Addition using (_+N_ ; canSubtractFromEqualityRight ; canSubtractFromEqualityLeft) public
|
||||
open Numbers.Naturals.Multiplication using (_*N_ ; multiplicationNIsCommutative) public
|
||||
|
||||
ℕSemiring : Semiring 0 1 _+N_ _*N_
|
||||
Monoid.associative (Semiring.monoid ℕSemiring) a b c = equalityCommutative (additionNIsAssociative a b c)
|
||||
Monoid.idLeft (Semiring.monoid ℕSemiring) _ = refl
|
||||
Monoid.idRight (Semiring.monoid ℕSemiring) a = additionNIsCommutative a 0
|
||||
Semiring.commutative ℕSemiring = additionNIsCommutative
|
||||
Monoid.associative (Semiring.multMonoid ℕSemiring) = multiplicationNIsAssociative
|
||||
Monoid.idLeft (Semiring.multMonoid ℕSemiring) a = additionNIsCommutative a 0
|
||||
Monoid.idRight (Semiring.multMonoid ℕSemiring) a = transitivity (multiplicationNIsCommutative a 1) (additionNIsCommutative a 0)
|
||||
Semiring.productZeroLeft ℕSemiring _ = refl
|
||||
Semiring.productZeroRight ℕSemiring a = multiplicationNIsCommutative a 0
|
||||
Semiring.+DistributesOver* ℕSemiring = productDistributes
|
||||
Semiring.+DistributesOver*' ℕSemiring a b c rewrite multiplicationNIsCommutative (a +N b) c | multiplicationNIsCommutative a c | multiplicationNIsCommutative b c = productDistributes c a b
|
||||
|
||||
succExtracts : (x y : ℕ) → (x +N succ y) ≡ (succ (x +N y))
|
||||
succExtracts x y = transitivity (Semiring.commutative ℕSemiring x (succ y)) (applyEquality succ (Semiring.commutative ℕSemiring y x))
|
||||
|
||||
productZeroImpliesOperandZero : {a b : ℕ} → a *N b ≡ 0 → (a ≡ 0) || (b ≡ 0)
|
||||
productZeroImpliesOperandZero {zero} {b} pr = inl refl
|
||||
productZeroImpliesOperandZero {succ a} {zero} pr = inr refl
|
||||
productZeroImpliesOperandZero {succ a} {succ b} ()
|
@@ -5,7 +5,8 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import WellFoundedInduction
|
||||
open import Functions
|
||||
open import Orders
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
|
||||
module Numbers.Naturals.WithK where
|
||||
|
||||
|
@@ -1,19 +1,21 @@
|
||||
{-# OPTIONS --warning=error --safe #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Addition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Numbers.Naturals.Order.WellFounded
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import WellFoundedInduction
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
|
||||
module Numbers.Primes.IntegerFactorisation where
|
||||
open TotalOrder ℕTotalOrder
|
||||
|
||||
-- Represent a factorisation into increasing factors
|
||||
-- Note that 0 cannot be expressed this way.
|
||||
record factorisationNonunit (minFactor : ℕ) (a : ℕ) : Set where
|
||||
-- Represent a factorisation into increasing factors
|
||||
-- Note that 0 cannot be expressed this way.
|
||||
record factorisationNonunit (minFactor : ℕ) (a : ℕ) : Set where
|
||||
inductive
|
||||
field
|
||||
1<a : 1 <N a
|
||||
@@ -26,80 +28,80 @@ module Numbers.Primes.IntegerFactorisation where
|
||||
otherFactorsNumber : ℕ
|
||||
otherFactors : ((divisionAlgResult.quot firstFactorDivision ≡ 1) && (otherFactorsNumber ≡ 0)) || (((1 <N divisionAlgResult.quot firstFactorDivision) && (factorisationNonunit firstFactor (divisionAlgResult.quot firstFactorDivision))))
|
||||
|
||||
lemma : (p : ℕ) → p *N 1 +N 0 ≡ p
|
||||
lemma p rewrite Semiring.sumZeroRight ℕSemiring (p *N 1) | Semiring.productOneRight ℕSemiring p = refl
|
||||
lemma : (p : ℕ) → p *N 1 +N 0 ≡ p
|
||||
lemma p rewrite Semiring.sumZeroRight ℕSemiring (p *N 1) | Semiring.productOneRight ℕSemiring p = refl
|
||||
|
||||
lemma' : {a b : ℕ} → a *N zero +N 0 ≡ b → b ≡ zero
|
||||
lemma' {a} {b} pr rewrite Semiring.sumZeroRight ℕSemiring (a *N zero) | Semiring.productZeroRight ℕSemiring a = equalityCommutative pr
|
||||
lemma' : {a b : ℕ} → a *N zero +N 0 ≡ b → b ≡ zero
|
||||
lemma' {a} {b} pr rewrite Semiring.sumZeroRight ℕSemiring (a *N zero) | Semiring.productZeroRight ℕSemiring a = equalityCommutative pr
|
||||
|
||||
primeFactorisation : {p : ℕ} → (pr : Prime p) → factorisationNonunit 1 p
|
||||
primeFactorisation {p} record { p>1 = p>1 ; pr = pr } = record {1<a = p>1 ; firstFactor = p ; firstFactorNontrivial = p>1 ; firstFactorBiggerMin = inl p>1 ; firstFactorDivision = record { quot = 1 ; rem = 0 ; pr = lemma p ; remIsSmall = zeroIsValidRem p ; quotSmall = inl (TotalOrder.<Transitive ℕTotalOrder (le zero refl) p>1) } ; firstFactorDivides = refl ; firstFactorPrime = record { p>1 = p>1 ; pr = pr} ; otherFactors = inl record { fst = refl ; snd = refl } ; otherFactorsNumber = 0 }
|
||||
primeFactorisation : {p : ℕ} → (pr : Prime p) → factorisationNonunit 1 p
|
||||
primeFactorisation {p} record { p>1 = p>1 ; pr = pr } = record {1<a = p>1 ; firstFactor = p ; firstFactorNontrivial = p>1 ; firstFactorBiggerMin = inl p>1 ; firstFactorDivision = record { quot = 1 ; rem = 0 ; pr = lemma p ; remIsSmall = zeroIsValidRem p ; quotSmall = inl (TotalOrder.<Transitive ℕTotalOrder (le zero refl) p>1) } ; firstFactorDivides = refl ; firstFactorPrime = record { p>1 = p>1 ; pr = pr} ; otherFactors = inl record { fst = refl ; snd = refl } ; otherFactorsNumber = 0 }
|
||||
where
|
||||
proof : (s : ℕ) → s *N 1 +N 0 ≡ s
|
||||
proof s rewrite Semiring.sumZeroRight ℕSemiring (s *N 1) | multiplicationNIsCommutative s 1 | Semiring.sumZeroRight ℕSemiring s = refl
|
||||
|
||||
twoAsFact : factorisationNonunit 2 2
|
||||
factorisationNonunit.1<a twoAsFact = succPreservesInequality (succIsPositive 0)
|
||||
factorisationNonunit.firstFactor twoAsFact = 2
|
||||
factorisationNonunit.firstFactorNontrivial twoAsFact = succPreservesInequality (succIsPositive 0)
|
||||
factorisationNonunit.firstFactorBiggerMin twoAsFact = inr refl
|
||||
factorisationNonunit.firstFactorDivision twoAsFact = record { quot = 1 ; rem = 0 ; remIsSmall = zeroIsValidRem 2 ; pr = refl ; quotSmall = inl (le 1 refl) }
|
||||
factorisationNonunit.firstFactorDivides twoAsFact = refl
|
||||
factorisationNonunit.firstFactorPrime twoAsFact = twoIsPrime
|
||||
factorisationNonunit.otherFactorsNumber twoAsFact = 0
|
||||
factorisationNonunit.otherFactors twoAsFact = inl record { fst = refl ; snd = refl }
|
||||
twoAsFact : factorisationNonunit 2 2
|
||||
factorisationNonunit.1<a twoAsFact = succPreservesInequality (succIsPositive 0)
|
||||
factorisationNonunit.firstFactor twoAsFact = 2
|
||||
factorisationNonunit.firstFactorNontrivial twoAsFact = succPreservesInequality (succIsPositive 0)
|
||||
factorisationNonunit.firstFactorBiggerMin twoAsFact = inr refl
|
||||
factorisationNonunit.firstFactorDivision twoAsFact = record { quot = 1 ; rem = 0 ; remIsSmall = zeroIsValidRem 2 ; pr = refl ; quotSmall = inl (le 1 refl) }
|
||||
factorisationNonunit.firstFactorDivides twoAsFact = refl
|
||||
factorisationNonunit.firstFactorPrime twoAsFact = twoIsPrime
|
||||
factorisationNonunit.otherFactorsNumber twoAsFact = 0
|
||||
factorisationNonunit.otherFactors twoAsFact = inl record { fst = refl ; snd = refl }
|
||||
|
||||
fourFact : factorisationNonunit 1 4
|
||||
factorisationNonunit.1<a fourFact = succPreservesInequality (succIsPositive 2)
|
||||
factorisationNonunit.firstFactor fourFact = 2
|
||||
factorisationNonunit.firstFactorNontrivial fourFact = succPreservesInequality (succIsPositive 0)
|
||||
factorisationNonunit.firstFactorBiggerMin fourFact = inl (succPreservesInequality (succIsPositive 0))
|
||||
factorisationNonunit.firstFactorDivision fourFact = record { quot = 2 ; rem = 0 ; remIsSmall = zeroIsValidRem 2 ; pr = refl ; quotSmall = inl (le 1 refl) }
|
||||
factorisationNonunit.firstFactorDivides fourFact = refl
|
||||
factorisationNonunit.firstFactorPrime fourFact = twoIsPrime
|
||||
factorisationNonunit.otherFactorsNumber fourFact = 1
|
||||
factorisationNonunit.otherFactors fourFact = inr record { fst = succPreservesInequality (succIsPositive 0) ; snd = twoAsFact }
|
||||
fourFact : factorisationNonunit 1 4
|
||||
factorisationNonunit.1<a fourFact = succPreservesInequality (succIsPositive 2)
|
||||
factorisationNonunit.firstFactor fourFact = 2
|
||||
factorisationNonunit.firstFactorNontrivial fourFact = succPreservesInequality (succIsPositive 0)
|
||||
factorisationNonunit.firstFactorBiggerMin fourFact = inl (succPreservesInequality (succIsPositive 0))
|
||||
factorisationNonunit.firstFactorDivision fourFact = record { quot = 2 ; rem = 0 ; remIsSmall = zeroIsValidRem 2 ; pr = refl ; quotSmall = inl (le 1 refl) }
|
||||
factorisationNonunit.firstFactorDivides fourFact = refl
|
||||
factorisationNonunit.firstFactorPrime fourFact = twoIsPrime
|
||||
factorisationNonunit.otherFactorsNumber fourFact = 1
|
||||
factorisationNonunit.otherFactors fourFact = inr record { fst = succPreservesInequality (succIsPositive 0) ; snd = twoAsFact }
|
||||
|
||||
lessLemma : {y : ℕ} → (1 <N y) → (zero <N y)
|
||||
lessLemma {.(succ (x +N 1))} (le x refl) = succIsPositive (x +N 1)
|
||||
lessLemma : {y : ℕ} → (1 <N y) → (zero <N y)
|
||||
lessLemma {.(succ (x +N 1))} (le x refl) = succIsPositive (x +N 1)
|
||||
|
||||
canReduceFactorBound : {a i j : ℕ} → factorisationNonunit i a → j <N i → factorisationNonunit j a
|
||||
canReduceFactorBound {a} {i} {j} record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inl ff<i ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors } j<i = record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inl (lessTransitive j<i ff<i) ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors }
|
||||
canReduceFactorBound {a} {i} {j} record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inr ff=i ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors } j<i = record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inl (identityOfIndiscernablesRight _<N_ j<i ff=i) ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors }
|
||||
canReduceFactorBound : {a i j : ℕ} → factorisationNonunit i a → j <N i → factorisationNonunit j a
|
||||
canReduceFactorBound {a} {i} {j} record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inl ff<i ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors } j<i = record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inl (lessTransitive j<i ff<i) ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors }
|
||||
canReduceFactorBound {a} {i} {j} record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inr ff=i ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors } j<i = record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = inl (identityOfIndiscernablesRight _<N_ j<i ff=i) ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors }
|
||||
|
||||
canReduceFactorBound' : {a i j : ℕ} → factorisationNonunit i a → j ≤N i → factorisationNonunit j a
|
||||
canReduceFactorBound' {a} {i} {j} factA (inl x) = canReduceFactorBound factA x
|
||||
canReduceFactorBound' {a} {i} {.i} factA (inr refl) = factA
|
||||
canReduceFactorBound' : {a i j : ℕ} → factorisationNonunit i a → j ≤N i → factorisationNonunit j a
|
||||
canReduceFactorBound' {a} {i} {j} factA (inl x) = canReduceFactorBound factA x
|
||||
canReduceFactorBound' {a} {i} {.i} factA (inr refl) = factA
|
||||
|
||||
canIncreaseFactorBound : {a i : ℕ} → (fact : factorisationNonunit 1 a) → (∀ x → 1 <N x → x <N i → x ∣ a → False) → (i ≤N factorisationNonunit.firstFactor fact) → factorisationNonunit i a
|
||||
canIncreaseFactorBound {a} {i} record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = firstFactorBiggerMin ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors } noSmaller iSmallEnough = record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = Prime.p>1 firstFactorPrime ; firstFactorBiggerMin = iSmallEnough ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors }
|
||||
canIncreaseFactorBound : {a i : ℕ} → (fact : factorisationNonunit 1 a) → (∀ x → 1 <N x → x <N i → x ∣ a → False) → (i ≤N factorisationNonunit.firstFactor fact) → factorisationNonunit i a
|
||||
canIncreaseFactorBound {a} {i} record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = firstFactorBiggerMin ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors } noSmaller iSmallEnough = record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = Prime.p>1 firstFactorPrime ; firstFactorBiggerMin = iSmallEnough ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = otherFactors }
|
||||
|
||||
-- Get the smallest prime factor of the number
|
||||
everyNumberHasAPrimeFactor : {a : ℕ} → (1 <N a) → Sg ℕ (λ i → ((i ∣ a) && (1 <N i)) && ((Prime i) && (∀ x → x <N i → 1 <N x → x ∣ a → False)))
|
||||
everyNumberHasAPrimeFactor {a} 1<a with compositeOrPrime a 1<a
|
||||
everyNumberHasAPrimeFactor {a} 1<a | inl record { n>1 = n>1 ; divisor = divisor ; dividesN = dividesN ; divisorLessN = divisorLessN ; divisorNot1 = divisorNot1 ; divisorPrime = divisorPrime ; noSmallerDivisors = noSmallerDivisors } = ( divisor , record { fst = record { fst = dividesN ; snd = divisorNot1 } ; snd = record { fst = divisorPrime ; snd = noSmallerDivisors } } )
|
||||
everyNumberHasAPrimeFactor {a} 1<a | inr x = (a , record { fst = record { fst = aDivA a ; snd = 1<a } ; snd = record { fst = x ; snd = λ y y<a 1<y y|a → lessImpliesNotEqual 1<y (equalityCommutative (Prime.pr x y|a y<a (lessLemma 1<y))) }} )
|
||||
-- Get the smallest prime factor of the number
|
||||
everyNumberHasAPrimeFactor : {a : ℕ} → (1 <N a) → Sg ℕ (λ i → ((i ∣ a) && (1 <N i)) && ((Prime i) && (∀ x → x <N i → 1 <N x → x ∣ a → False)))
|
||||
everyNumberHasAPrimeFactor {a} 1<a with compositeOrPrime a 1<a
|
||||
everyNumberHasAPrimeFactor {a} 1<a | inl record { n>1 = n>1 ; divisor = divisor ; dividesN = dividesN ; divisorLessN = divisorLessN ; divisorNot1 = divisorNot1 ; divisorPrime = divisorPrime ; noSmallerDivisors = noSmallerDivisors } = ( divisor , record { fst = record { fst = dividesN ; snd = divisorNot1 } ; snd = record { fst = divisorPrime ; snd = noSmallerDivisors } } )
|
||||
everyNumberHasAPrimeFactor {a} 1<a | inr x = (a , record { fst = record { fst = aDivA a ; snd = 1<a } ; snd = record { fst = x ; snd = λ y y<a 1<y y|a → irreflexive (<WellDefined (equalityCommutative (Prime.pr x y|a y<a (lessLemma 1<y))) refl 1<y) }} )
|
||||
|
||||
lemma2 : {a b c : ℕ} → 1 <N a → 0 <N b → a *N b +N 0 ≡ c → b <N c
|
||||
lemma2 {zero} {b} {c} 1<a _ pr = exFalso (zeroNeverGreater 1<a)
|
||||
lemma2 {succ zero} {b} {c} 1<a _ pr = exFalso (lessIrreflexive 1<a)
|
||||
lemma2 {succ (succ a)} {zero} {zero} 1<a t pr = exFalso (lessIrreflexive t)
|
||||
lemma2 {succ (succ a)} {zero} {succ c} 1<a t pr = succIsPositive c
|
||||
lemma2 {succ (succ a)} {succ b} {c} 1<a t pr = le (b +N (a *N succ b)) go
|
||||
lemma2 : {a b c : ℕ} → 1 <N a → 0 <N b → a *N b +N 0 ≡ c → b <N c
|
||||
lemma2 {zero} {b} {c} 1<a _ pr = exFalso (zeroNeverGreater 1<a)
|
||||
lemma2 {succ zero} {b} {c} 1<a _ pr = exFalso (lessIrreflexive 1<a)
|
||||
lemma2 {succ (succ a)} {zero} {zero} 1<a t pr = exFalso (lessIrreflexive t)
|
||||
lemma2 {succ (succ a)} {zero} {succ c} 1<a t pr = succIsPositive c
|
||||
lemma2 {succ (succ a)} {succ b} {c} 1<a t pr = le (b +N (a *N succ b)) go
|
||||
where
|
||||
assocLemm : (a b c : ℕ) → (a +N b) +N c ≡ (a +N c) +N b
|
||||
assocLemm a b c rewrite equalityCommutative (Semiring.+Associative ℕSemiring a b c) | Semiring.commutative ℕSemiring b c | Semiring.+Associative ℕSemiring a c b = refl
|
||||
go : succ ((b +N a *N succ b) +N succ b) ≡ c
|
||||
go rewrite Semiring.sumZeroRight ℕSemiring (succ (b +N succ (b +N a *N succ b))) | equalityCommutative (assocLemm b (succ b) (a *N succ b)) | Semiring.+Associative ℕSemiring b (succ b) (a *N succ b) = pr
|
||||
|
||||
factorIntegerLemma : (x : ℕ) (indHyp : (y : ℕ) (y<x : y <N x) → ((y <N 2) || (factorisationNonunit 1 y))) → ((x <N 2) || (factorisationNonunit 1 x))
|
||||
factorIntegerLemma zero indHyp = inl (succIsPositive 1)
|
||||
factorIntegerLemma (succ zero) indHyp = inl (succPreservesInequality (succIsPositive 0))
|
||||
factorIntegerLemma (succ (succ x)) indHyp with everyNumberHasAPrimeFactor {succ (succ x)} (succPreservesInequality (succIsPositive x))
|
||||
factorIntegerLemma (succ (succ x)) indHyp | a , record { fst = record { fst = divides record {quot = zero ; rem = .0 ; pr = ssxDivA ; remIsSmall = r } refl ; snd = 1<a }; snd = record { fst = primeA ; snd = smallerFactors } } rewrite Semiring.sumZeroRight ℕSemiring (a *N zero) | multiplicationNIsCommutative a 0 = exFalso (naughtE ssxDivA)
|
||||
factorIntegerLemma (succ (succ x)) indHyp | a , record { fst = record { fst = divides record {quot = succ zero ; rem = .0 ; pr = ssxDivA ; remIsSmall = r } refl ; snd = 1<a }; snd = record { fst = primeA ; snd = smallerFactors } } = inr record { 1<a = succPreservesInequality (succIsPositive x) ; firstFactor = a ; firstFactorNontrivial = Prime.p>1 primeA ; firstFactorBiggerMin = inl (Prime.p>1 primeA) ; firstFactorDivision = record { quot = 1 ; rem = 0 ; pr = ssxDivA ; remIsSmall = r ; quotSmall = inl (TotalOrder.<Transitive ℕTotalOrder (le zero refl) 1<a) } ; firstFactorDivides = refl ; firstFactorPrime = record { p>1 = Prime.p>1 primeA ; pr = Prime.pr primeA } ; otherFactors = inl record { fst = refl ; snd = refl } ; otherFactorsNumber = 0 }
|
||||
factorIntegerLemma : (x : ℕ) (indHyp : (y : ℕ) (y<x : y <N x) → ((y <N 2) || (factorisationNonunit 1 y))) → ((x <N 2) || (factorisationNonunit 1 x))
|
||||
factorIntegerLemma zero indHyp = inl (succIsPositive 1)
|
||||
factorIntegerLemma (succ zero) indHyp = inl (succPreservesInequality (succIsPositive 0))
|
||||
factorIntegerLemma (succ (succ x)) indHyp with everyNumberHasAPrimeFactor {succ (succ x)} (succPreservesInequality (succIsPositive x))
|
||||
factorIntegerLemma (succ (succ x)) indHyp | a , record { fst = record { fst = divides record {quot = zero ; rem = .0 ; pr = ssxDivA ; remIsSmall = r } refl ; snd = 1<a }; snd = record { fst = primeA ; snd = smallerFactors } } rewrite Semiring.sumZeroRight ℕSemiring (a *N zero) | multiplicationNIsCommutative a 0 = exFalso (naughtE ssxDivA)
|
||||
factorIntegerLemma (succ (succ x)) indHyp | a , record { fst = record { fst = divides record {quot = succ zero ; rem = .0 ; pr = ssxDivA ; remIsSmall = r } refl ; snd = 1<a }; snd = record { fst = primeA ; snd = smallerFactors } } = inr record { 1<a = succPreservesInequality (succIsPositive x) ; firstFactor = a ; firstFactorNontrivial = Prime.p>1 primeA ; firstFactorBiggerMin = inl (Prime.p>1 primeA) ; firstFactorDivision = record { quot = 1 ; rem = 0 ; pr = ssxDivA ; remIsSmall = r ; quotSmall = inl (TotalOrder.<Transitive ℕTotalOrder (le zero refl) 1<a) } ; firstFactorDivides = refl ; firstFactorPrime = record { p>1 = Prime.p>1 primeA ; pr = Prime.pr primeA } ; otherFactors = inl record { fst = refl ; snd = refl } ; otherFactorsNumber = 0 }
|
||||
|
||||
factorIntegerLemma (succ (succ x)) indHyp | a , record { fst = record { fst = divides record {quot = succ (succ qu) ; rem = .0 ; pr = ssxDivA ; remIsSmall = remSmall } refl ; snd = 1<a }; snd = record { fst = primeA ; snd = smallerFactors } } = inr record { 1<a = succPreservesInequality (succIsPositive x) ; firstFactor = a ; firstFactorNontrivial = Prime.p>1 primeA ; firstFactorBiggerMin = inl (Prime.p>1 primeA) ; firstFactorDivision = record { quot = succ (succ qu) ; rem = 0 ; pr = ssxDivA ; remIsSmall = remSmall ; quotSmall = inl (TotalOrder.<Transitive ℕTotalOrder (le zero refl) 1<a) } ; firstFactorDivides = refl ; firstFactorPrime = record { p>1 = Prime.p>1 primeA ; pr = Prime.pr primeA } ; otherFactors = inr record {fst = succPreservesInequality (succIsPositive qu) ; snd = factNonunit} ; otherFactorsNumber = succ (factorisationNonunit.otherFactorsNumber indHypRes') }
|
||||
factorIntegerLemma (succ (succ x)) indHyp | a , record { fst = record { fst = divides record {quot = succ (succ qu) ; rem = .0 ; pr = ssxDivA ; remIsSmall = remSmall } refl ; snd = 1<a }; snd = record { fst = primeA ; snd = smallerFactors } } = inr record { 1<a = succPreservesInequality (succIsPositive x) ; firstFactor = a ; firstFactorNontrivial = Prime.p>1 primeA ; firstFactorBiggerMin = inl (Prime.p>1 primeA) ; firstFactorDivision = record { quot = succ (succ qu) ; rem = 0 ; pr = ssxDivA ; remIsSmall = remSmall ; quotSmall = inl (TotalOrder.<Transitive ℕTotalOrder (le zero refl) 1<a) } ; firstFactorDivides = refl ; firstFactorPrime = record { p>1 = Prime.p>1 primeA ; pr = Prime.pr primeA } ; otherFactors = inr record {fst = succPreservesInequality (succIsPositive qu) ; snd = factNonunit} ; otherFactorsNumber = succ (factorisationNonunit.otherFactorsNumber indHypRes') }
|
||||
where
|
||||
indHypRes : ((succ (succ qu)) <N 2) || factorisationNonunit 1 (succ (succ qu))
|
||||
indHypRes = indHyp (succ (succ qu)) (lemma2 {a} {succ (succ qu)} {succ (succ x)} 1<a (succIsPositive (succ qu)) ssxDivA)
|
||||
@@ -110,37 +112,37 @@ module Numbers.Primes.IntegerFactorisation where
|
||||
z|ssx : (z : ℕ) → z ∣ succ (succ qu) → z ∣ succ (succ x)
|
||||
z|ssx z z|ssq = (divisibilityTransitive z|ssq (divides (record { quot = a ; rem = 0 ; pr = identityOfIndiscernablesLeft _≡_ ssxDivA (applyEquality (λ t → t +N 0) (multiplicationNIsCommutative a (succ (succ qu)))) ; remIsSmall = zeroIsValidRem (succ (succ qu)) ; quotSmall = inl (succIsPositive _) }) refl))
|
||||
factNonunit : factorisationNonunit a (succ (succ qu))
|
||||
factNonunit with orderIsTotal a (factorisationNonunit.firstFactor indHypRes')
|
||||
factNonunit with totality a (factorisationNonunit.firstFactor indHypRes')
|
||||
factNonunit | inl (inl a<ff) = canIncreaseFactorBound indHypRes' (λ z 1<z z<a z|ssq → smallerFactors z z<a 1<z (z|ssx z z|ssq)) (inl a<ff)
|
||||
factNonunit | inl (inr ff<a) = exFalso (smallerFactors (factorisationNonunit.firstFactor indHypRes') ff<a (factorisationNonunit.firstFactorNontrivial indHypRes') (z|ssx (factorisationNonunit.firstFactor indHypRes') (divides (factorisationNonunit.firstFactorDivision indHypRes') (factorisationNonunit.firstFactorDivides indHypRes'))))
|
||||
factNonunit | inr ff=a = canIncreaseFactorBound indHypRes' (λ z 1<z z<a z|ssq → smallerFactors z z<a 1<z (divisibilityTransitive z|ssq (divides (record { quot = a ; rem = 0 ; pr = identityOfIndiscernablesLeft _≡_ ssxDivA (applyEquality (λ t → t +N 0) (multiplicationNIsCommutative a (succ (succ qu)))) ; remIsSmall = zeroIsValidRem (succ (succ qu)) ; quotSmall = inl (succIsPositive _) }) refl))) (inr ff=a)
|
||||
|
||||
factorInteger : (a : ℕ) → (1 <N a) → factorisationNonunit 1 a
|
||||
factorInteger a 1<a with (rec <NWellfounded (λ n → (n <N 2) || (factorisationNonunit 1 n))) factorIntegerLemma
|
||||
... | bl with bl a
|
||||
factorInteger a 1<a | bl | inl x = exFalso (noIntegersBetweenXAndSuccX 1 1<a x)
|
||||
factorInteger a 1<a | bl | inr x = x
|
||||
factorInteger : (a : ℕ) → (1 <N a) → factorisationNonunit 1 a
|
||||
factorInteger a 1<a with (rec <NWellfounded (λ n → (n <N 2) || (factorisationNonunit 1 n))) factorIntegerLemma
|
||||
... | bl with bl a
|
||||
factorInteger a 1<a | bl | inl x = exFalso (noIntegersBetweenXAndSuccX 1 1<a x)
|
||||
factorInteger a 1<a | bl | inr x = x
|
||||
|
||||
lessTransLemma : {p i j : ℕ} → p <N i → i ≤N j → p <N j
|
||||
lessTransLemma {p} {i} {j} p<i (inl x) = orderIsTransitive p<i x
|
||||
lessTransLemma {p} {i} {j} p<i (inr x) rewrite x = p<i
|
||||
lessTransLemma : {p i j : ℕ} → p <N i → i ≤N j → p <N j
|
||||
lessTransLemma {p} {i} {j} p<i (inl x) = <Transitive p<i x
|
||||
lessTransLemma {p} {i} {j} p<i (inr x) rewrite x = p<i
|
||||
|
||||
lemma4' : {quot rem b : ℕ} → (quot +N quot) +N rem ≡ succ b → quot <N succ b
|
||||
lemma4' {zero} {rem} {b} pr = succIsPositive b
|
||||
lemma4' {succ quot} {rem} {b} pr rewrite equalityCommutative (Semiring.+Associative ℕSemiring quot (succ quot) rem) = succPreservesInequality (le (quot +N rem) (succInjective (transitivity (applyEquality succ (Semiring.commutative ℕSemiring _ quot)) pr)))
|
||||
lemma4' : {quot rem b : ℕ} → (quot +N quot) +N rem ≡ succ b → quot <N succ b
|
||||
lemma4' {zero} {rem} {b} pr = succIsPositive b
|
||||
lemma4' {succ quot} {rem} {b} pr rewrite equalityCommutative (Semiring.+Associative ℕSemiring quot (succ quot) rem) = succPreservesInequality (le (quot +N rem) (succInjective (transitivity (applyEquality succ (Semiring.commutative ℕSemiring _ quot)) pr)))
|
||||
|
||||
lemma4 : {quot a rem b : ℕ} → (a *N quot +N rem ≡ succ b) → (1 <N a) → (quot <N succ b)
|
||||
lemma4 {quot} {zero} {rem} {b} pr 1<a = exFalso (zeroNeverGreater 1<a)
|
||||
lemma4 {quot} {succ zero} {rem} {b} pr 1<a = exFalso (lessIrreflexive 1<a)
|
||||
lemma4 {quot} {succ (succ zero)} {rem} {b} pr 1<a rewrite Semiring.sumZeroRight ℕSemiring quot = lemma4' pr
|
||||
lemma4 {quot} {succ (succ (succ a))} {rem} {b} pr 1<a = lemma4 {quot} {succ (succ a)} {quot +N rem} {b} p' (succPreservesInequality (succIsPositive a))
|
||||
lemma4 : {quot a rem b : ℕ} → (a *N quot +N rem ≡ succ b) → (1 <N a) → (quot <N succ b)
|
||||
lemma4 {quot} {zero} {rem} {b} pr 1<a = exFalso (zeroNeverGreater 1<a)
|
||||
lemma4 {quot} {succ zero} {rem} {b} pr 1<a = exFalso (lessIrreflexive 1<a)
|
||||
lemma4 {quot} {succ (succ zero)} {rem} {b} pr 1<a rewrite Semiring.sumZeroRight ℕSemiring quot = lemma4' pr
|
||||
lemma4 {quot} {succ (succ (succ a))} {rem} {b} pr 1<a = lemma4 {quot} {succ (succ a)} {quot +N rem} {b} p' (succPreservesInequality (succIsPositive a))
|
||||
where
|
||||
p' : (quot +N (quot +N a *N quot)) +N (quot +N rem) ≡ succ b
|
||||
p' rewrite Semiring.commutative ℕSemiring quot (quot +N (quot +N a *N quot)) | Semiring.+Associative ℕSemiring (quot +N (quot +N a *N quot)) quot rem = pr
|
||||
|
||||
noSmallerFactors : {a i p : ℕ} → (factorisationNonunit i a) → (Prime p) → (p <N i) → (p ∣ a) → False
|
||||
noSmallerFactors {a} {i} {p} factA pPrime p<i p|a with rec <NWellfounded (λ b → (factorisationNonunit i b) → p ∣ b → False)
|
||||
... | indHyp = (indHyp prf) a factA p|a
|
||||
noSmallerFactors : {a i p : ℕ} → (factorisationNonunit i a) → (Prime p) → (p <N i) → (p ∣ a) → False
|
||||
noSmallerFactors {a} {i} {p} factA pPrime p<i p|a with rec <NWellfounded (λ b → (factorisationNonunit i b) → p ∣ b → False)
|
||||
... | indHyp = (indHyp prf) a factA p|a
|
||||
where
|
||||
prf : (x : ℕ) (ind : (y : ℕ) (y<x : y <N x) (factY : factorisationNonunit i y) (p|y : p ∣ y) → False) (factX : factorisationNonunit i x) (p|x : p ∣ x) → False
|
||||
prf x ind record { 1<a = 1<a ; firstFactor = firstFactor ; firstFactorNontrivial = firstFactorNontrivial ; firstFactorBiggerMin = firstFactorBiggerMin ; firstFactorDivision = firstFactorDivision ; firstFactorDivides = firstFactorDivides ; firstFactorPrime = firstFactorPrime ; otherFactors = (inl record { fst = quotFirstfact=1 ; snd = otherFactorsNumber }) } p|x = exFalso bad
|
||||
@@ -187,11 +189,11 @@ module Numbers.Primes.IntegerFactorisation where
|
||||
p|q (inl fls) = exFalso (p|ffIsFalse fls)
|
||||
p|q (inr res) = res
|
||||
|
||||
lemma3 : {a : ℕ} → a ≡ 0 → 1 <N a → False
|
||||
lemma3 {a} a=0 pr rewrite a=0 = zeroNeverGreater pr
|
||||
lemma3 : {a : ℕ} → a ≡ 0 → 1 <N a → False
|
||||
lemma3 {a} a=0 pr rewrite a=0 = zeroNeverGreater pr
|
||||
|
||||
firstFactorUniqueLemma : {i : ℕ} → (a : ℕ) → (1 <N a) → (f1 : factorisationNonunit i a) → (f2 : factorisationNonunit i a) → (factorisationNonunit.firstFactor f1 <N factorisationNonunit.firstFactor f2) → False
|
||||
firstFactorUniqueLemma {i} a 1<a f1 f2 ff1<ff2 = go
|
||||
firstFactorUniqueLemma : {i : ℕ} → (a : ℕ) → (1 <N a) → (f1 : factorisationNonunit i a) → (f2 : factorisationNonunit i a) → (factorisationNonunit.firstFactor f1 <N factorisationNonunit.firstFactor f2) → False
|
||||
firstFactorUniqueLemma {i} a 1<a f1 f2 ff1<ff2 = go
|
||||
where
|
||||
p1 = factorisationNonunit.firstFactor f1
|
||||
rem1 = divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f1)
|
||||
@@ -214,7 +216,7 @@ module Numbers.Primes.IntegerFactorisation where
|
||||
p1|second'' = p1|second' p1|second
|
||||
go : False
|
||||
go with p1|second''
|
||||
go | inl x = lessImpliesNotEqual ff1<ff2 x
|
||||
go | inl x = irreflexive (<WellDefined x refl ff1<ff2)
|
||||
go | inr x with factorisationNonunit.otherFactors f2
|
||||
go | inr x | inl record { fst = rem2=1 ; snd = _ } rewrite rem2=1 = exFalso (oneIsNotPrime res)
|
||||
where
|
||||
@@ -231,15 +233,15 @@ module Numbers.Primes.IntegerFactorisation where
|
||||
res rewrite equalityCommutative 1prime' = (factorisationNonunit.firstFactorPrime f1)
|
||||
go | inr x | inr p1|rem2 | inr factorRem2 = noSmallerFactors (_&&_.snd factorRem2) (factorisationNonunit.firstFactorPrime f1) p1<p2 x
|
||||
|
||||
firstFactorUnique : {i : ℕ} → (a : ℕ) → (1 <N a) → (f1 : factorisationNonunit i a) → (f2 : factorisationNonunit i a) → (factorisationNonunit.firstFactor f1 ≡ factorisationNonunit.firstFactor f2)
|
||||
firstFactorUnique {i} a 1<a f1 f2 with orderIsTotal (factorisationNonunit.firstFactor f1) (factorisationNonunit.firstFactor f2)
|
||||
firstFactorUnique {i} a 1<a f1 f2 | inl (inl f1<f2) = exFalso (firstFactorUniqueLemma a 1<a f1 f2 f1<f2)
|
||||
firstFactorUnique {i} a 1<a f1 f2 | inl (inr f2<f1) = exFalso (firstFactorUniqueLemma a 1<a f2 f1 f2<f1)
|
||||
firstFactorUnique {i} a 1<a f1 f2 | inr x = x
|
||||
firstFactorUnique : {i : ℕ} → (a : ℕ) → (1 <N a) → (f1 : factorisationNonunit i a) → (f2 : factorisationNonunit i a) → (factorisationNonunit.firstFactor f1 ≡ factorisationNonunit.firstFactor f2)
|
||||
firstFactorUnique {i} a 1<a f1 f2 with totality (factorisationNonunit.firstFactor f1) (factorisationNonunit.firstFactor f2)
|
||||
firstFactorUnique {i} a 1<a f1 f2 | inl (inl f1<f2) = exFalso (firstFactorUniqueLemma a 1<a f1 f2 f1<f2)
|
||||
firstFactorUnique {i} a 1<a f1 f2 | inl (inr f2<f1) = exFalso (firstFactorUniqueLemma a 1<a f2 f1 f2<f1)
|
||||
firstFactorUnique {i} a 1<a f1 f2 | inr x = x
|
||||
|
||||
factorListLemma : {i : ℕ} → (a : ℕ) → (1 <N a) → (f1 : factorisationNonunit i a) → (f2 : factorisationNonunit i a) → (divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f2)) ≡ (divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f1))
|
||||
factorListLemma {i} a 1<a f1 f2 with firstFactorUnique {i} a 1<a f1 f2
|
||||
... | firstFactEqual = res
|
||||
factorListLemma : {i : ℕ} → (a : ℕ) → (1 <N a) → (f1 : factorisationNonunit i a) → (f2 : factorisationNonunit i a) → (divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f2)) ≡ (divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f1))
|
||||
factorListLemma {i} a 1<a f1 f2 with firstFactorUnique {i} a 1<a f1 f2
|
||||
... | firstFactEqual = res
|
||||
where
|
||||
div1 : divisionAlgResult (factorisationNonunit.firstFactor f1) a
|
||||
div1 = factorisationNonunit.firstFactorDivision f1
|
||||
@@ -266,7 +268,7 @@ module Numbers.Primes.IntegerFactorisation where
|
||||
res : divisionAlgResult.quot div2 ≡ divisionAlgResult.quot div1
|
||||
res = productCancelsLeft (factorisationNonunit.firstFactor f1) (divisionAlgResult.quot div2) (divisionAlgResult.quot div1) positive pr'
|
||||
|
||||
factorListSameLength : {i : ℕ} → (a : ℕ) → (1 <N a) → (f1 : factorisationNonunit i a) → (f2 : factorisationNonunit i a) → (divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f1) ≡ 1) → divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f2) ≡ 1
|
||||
factorListSameLength {i} a 1<a f1 f2 quot=1 with firstFactorUnique {i} a 1<a f1 f2
|
||||
... | firstFactEqual with factorListLemma {i} a 1<a f1 f2
|
||||
... | eq = transitivity eq quot=1
|
||||
factorListSameLength : {i : ℕ} → (a : ℕ) → (1 <N a) → (f1 : factorisationNonunit i a) → (f2 : factorisationNonunit i a) → (divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f1) ≡ 1) → divisionAlgResult.quot (factorisationNonunit.firstFactorDivision f2) ≡ 1
|
||||
factorListSameLength {i} a 1<a f1 f2 quot=1 with firstFactorUnique {i} a 1<a f1 f2
|
||||
... | firstFactEqual with factorListLemma {i} a 1<a f1 f2
|
||||
... | eq = transitivity eq quot=1
|
||||
|
File diff suppressed because it is too large
Load Diff
@@ -2,7 +2,10 @@
|
||||
|
||||
open import WellFoundedInduction
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.WellFounded
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Integers.Order
|
||||
open import Groups.Groups
|
||||
@@ -22,6 +25,8 @@ open import Orders
|
||||
|
||||
module Numbers.Rationals.Lemmas where
|
||||
|
||||
open import Semirings.Lemmas ℕSemiring
|
||||
|
||||
open PartiallyOrderedRing ℤPOrderedRing
|
||||
open import Rings.Orders.Total.Lemmas ℤOrderedRing
|
||||
open SetoidTotalOrder (totalOrderToSetoidTotalOrder ℤOrder)
|
||||
|
131
Orders.agda
131
Orders.agda
@@ -7,13 +7,15 @@ open import Functions
|
||||
|
||||
module Orders where
|
||||
|
||||
record PartialOrder {a b : _} (carrier : Set a) : Set (a ⊔ lsuc b) where
|
||||
record PartialOrder {a b : _} (carrier : Set a) : Set (a ⊔ lsuc b) where
|
||||
field
|
||||
_<_ : Rel {a} {b} carrier
|
||||
irreflexive : {x : carrier} → (x < x) → False
|
||||
<Transitive : {a b c : carrier} → (a < b) → (b < c) → (a < c)
|
||||
<WellDefined : {r s t u : carrier} → (r ≡ t) → (s ≡ u) → r < s → t < u
|
||||
<WellDefined refl refl r<s = r<s
|
||||
|
||||
record TotalOrder {a b : _} (carrier : Set a) : Set (a ⊔ lsuc b) where
|
||||
record TotalOrder {a b : _} (carrier : Set a) : Set (a ⊔ lsuc b) where
|
||||
field
|
||||
order : PartialOrder {a} {b} carrier
|
||||
_<_ : Rel carrier
|
||||
@@ -35,72 +37,73 @@ module Orders where
|
||||
|
||||
irreflexive = PartialOrder.irreflexive order
|
||||
<Transitive = PartialOrder.<Transitive order
|
||||
<WellDefined = PartialOrder.<WellDefined order
|
||||
|
||||
equivMin : {a b : _} {A : Set a} → (order : TotalOrder {a} {b} A) → {x y : A} → (TotalOrder._<_ order x y) → TotalOrder.min order x y ≡ x
|
||||
equivMin order {x} {y} x<y with TotalOrder.totality order x y
|
||||
equivMin order {x} {y} x<y | inl (inl x₁) = refl
|
||||
equivMin order {x} {y} x<y | inl (inr y<x) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) x<y y<x))
|
||||
equivMin order {x} {y} x<y | inr x=y rewrite x=y = refl
|
||||
equivMin : {a b : _} {A : Set a} → (order : TotalOrder {a} {b} A) → {x y : A} → (TotalOrder._<_ order x y) → TotalOrder.min order x y ≡ x
|
||||
equivMin order {x} {y} x<y with TotalOrder.totality order x y
|
||||
equivMin order {x} {y} x<y | inl (inl x₁) = refl
|
||||
equivMin order {x} {y} x<y | inl (inr y<x) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) x<y y<x))
|
||||
equivMin order {x} {y} x<y | inr x=y rewrite x=y = refl
|
||||
|
||||
equivMin' : {a b : _} {A : Set a} → (order : TotalOrder {a} {b} A) → {x y : A} → (TotalOrder.min order x y ≡ x) → (TotalOrder._<_ order x y) || (x ≡ y)
|
||||
equivMin' order {x} {y} minEq with TotalOrder.totality order x y
|
||||
equivMin' order {x} {y} minEq | inl (inl x<y) = inl x<y
|
||||
equivMin' order {x} {y} minEq | inl (inr y<x) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (identityOfIndiscernablesLeft (TotalOrder._<_ order) y<x minEq))
|
||||
equivMin' order {x} {y} minEq | inr x=y = inr x=y
|
||||
equivMin' : {a b : _} {A : Set a} → (order : TotalOrder {a} {b} A) → {x y : A} → (TotalOrder.min order x y ≡ x) → (TotalOrder._<_ order x y) || (x ≡ y)
|
||||
equivMin' order {x} {y} minEq with TotalOrder.totality order x y
|
||||
equivMin' order {x} {y} minEq | inl (inl x<y) = inl x<y
|
||||
equivMin' order {x} {y} minEq | inl (inr y<x) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (identityOfIndiscernablesLeft (TotalOrder._<_ order) y<x minEq))
|
||||
equivMin' order {x} {y} minEq | inr x=y = inr x=y
|
||||
|
||||
minCommutes : {a b : _} {A : Set a} → (order : TotalOrder {a} {b} A) → (x y : A) → (TotalOrder.min order x y) ≡ (TotalOrder.min order y x)
|
||||
minCommutes order x y with TotalOrder.totality order x y
|
||||
minCommutes order x y | inl (inl x<y) with TotalOrder.totality order y x
|
||||
minCommutes order x y | inl (inl x<y) | inl (inl y<x) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) y<x x<y))
|
||||
minCommutes order x y | inl (inl x<y) | inl (inr x<y') = refl
|
||||
minCommutes order x y | inl (inl x<y) | inr y=x = equalityCommutative y=x
|
||||
minCommutes order x y | inl (inr y<x) with TotalOrder.totality order y x
|
||||
minCommutes order x y | inl (inr y<x) | inl (inl y<x') = refl
|
||||
minCommutes order x y | inl (inr y<x) | inl (inr x<y) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) y<x x<y))
|
||||
minCommutes order x y | inl (inr y<x) | inr y=x = refl
|
||||
minCommutes order x y | inr x=y with TotalOrder.totality order y x
|
||||
minCommutes order x y | inr x=y | inl (inl x₁) = x=y
|
||||
minCommutes order x y | inr x=y | inl (inr x₁) = refl
|
||||
minCommutes order x y | inr x=y | inr x₁ = x=y
|
||||
minCommutes : {a b : _} {A : Set a} → (order : TotalOrder {a} {b} A) → (x y : A) → (TotalOrder.min order x y) ≡ (TotalOrder.min order y x)
|
||||
minCommutes order x y with TotalOrder.totality order x y
|
||||
minCommutes order x y | inl (inl x<y) with TotalOrder.totality order y x
|
||||
minCommutes order x y | inl (inl x<y) | inl (inl y<x) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) y<x x<y))
|
||||
minCommutes order x y | inl (inl x<y) | inl (inr x<y') = refl
|
||||
minCommutes order x y | inl (inl x<y) | inr y=x = equalityCommutative y=x
|
||||
minCommutes order x y | inl (inr y<x) with TotalOrder.totality order y x
|
||||
minCommutes order x y | inl (inr y<x) | inl (inl y<x') = refl
|
||||
minCommutes order x y | inl (inr y<x) | inl (inr x<y) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) y<x x<y))
|
||||
minCommutes order x y | inl (inr y<x) | inr y=x = refl
|
||||
minCommutes order x y | inr x=y with TotalOrder.totality order y x
|
||||
minCommutes order x y | inr x=y | inl (inl x₁) = x=y
|
||||
minCommutes order x y | inr x=y | inl (inr x₁) = refl
|
||||
minCommutes order x y | inr x=y | inr x₁ = x=y
|
||||
|
||||
minIdempotent : {a b : _} {A : Set a} → (order : TotalOrder {a} {b} A) → (x : A) → TotalOrder.min order x x ≡ x
|
||||
minIdempotent order x with TotalOrder.totality order x x
|
||||
minIdempotent order x | inl (inl x₁) = refl
|
||||
minIdempotent order x | inl (inr x₁) = refl
|
||||
minIdempotent order x | inr x₁ = refl
|
||||
minIdempotent : {a b : _} {A : Set a} → (order : TotalOrder {a} {b} A) → (x : A) → TotalOrder.min order x x ≡ x
|
||||
minIdempotent order x with TotalOrder.totality order x x
|
||||
minIdempotent order x | inl (inl x₁) = refl
|
||||
minIdempotent order x | inl (inr x₁) = refl
|
||||
minIdempotent order x | inr x₁ = refl
|
||||
|
||||
swapMin : {a b : _} {A : Set a} {order : TotalOrder {a} {b} A} {x y z : A} → (TotalOrder.min order x (TotalOrder.min order y z)) ≡ TotalOrder.min order y (TotalOrder.min order x z)
|
||||
swapMin {order = order} {x} {y} {z} with TotalOrder.totality order y z
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) with TotalOrder.totality order x z
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inl x<z) = minCommutes order x y
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) with TotalOrder.totality order x y
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inl (inl x<y) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) y<z (PartialOrder.<Transitive (TotalOrder.order order) z<x x<y)))
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inl (inr y<x) = equalityCommutative (equivMin order y<z)
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inr x=y rewrite x=y = equalityCommutative (equivMin order y<z)
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inr x=z = minCommutes order x y
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) with TotalOrder.totality order x z
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inl x<z) rewrite minCommutes order y x = equalityCommutative (equivMin order (PartialOrder.<Transitive (TotalOrder.order order) x<z z<y))
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) with TotalOrder.totality order y z
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inl (inl y<z) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) z<y y<z))
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inl (inr z<y') = refl
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inr y=z = equalityCommutative y=z
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inr x=z rewrite x=z | minCommutes order y z = equalityCommutative (equivMin order z<y)
|
||||
swapMin {order = order} {x} {y} {z} | inr y=z with TotalOrder.totality order x z
|
||||
swapMin {order = order} {x} {y} {z} | inr y=z | inl (inl x<z) = minCommutes order x y
|
||||
swapMin {order = order} {x} {y} {z} | inr y=z | inl (inr z<x) rewrite y=z | minIdempotent order z | minCommutes order x z = equivMin order z<x
|
||||
swapMin {order = order} {x} {y} {z} | inr y=z | inr x=z = minCommutes order x y
|
||||
swapMin : {a b : _} {A : Set a} {order : TotalOrder {a} {b} A} {x y z : A} → (TotalOrder.min order x (TotalOrder.min order y z)) ≡ TotalOrder.min order y (TotalOrder.min order x z)
|
||||
swapMin {order = order} {x} {y} {z} with TotalOrder.totality order y z
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) with TotalOrder.totality order x z
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inl x<z) = minCommutes order x y
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) with TotalOrder.totality order x y
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inl (inl x<y) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) y<z (PartialOrder.<Transitive (TotalOrder.order order) z<x x<y)))
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inl (inr y<x) = equalityCommutative (equivMin order y<z)
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inl (inr z<x) | inr x=y rewrite x=y = equalityCommutative (equivMin order y<z)
|
||||
swapMin {order = order} {x} {y} {z} | inl (inl y<z) | inr x=z = minCommutes order x y
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) with TotalOrder.totality order x z
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inl x<z) rewrite minCommutes order y x = equalityCommutative (equivMin order (PartialOrder.<Transitive (TotalOrder.order order) x<z z<y))
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) with TotalOrder.totality order y z
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inl (inl y<z) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) z<y y<z))
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inl (inr z<y') = refl
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inl (inr z<x) | inr y=z = equalityCommutative y=z
|
||||
swapMin {order = order} {x} {y} {z} | inl (inr z<y) | inr x=z rewrite x=z | minCommutes order y z = equalityCommutative (equivMin order z<y)
|
||||
swapMin {order = order} {x} {y} {z} | inr y=z with TotalOrder.totality order x z
|
||||
swapMin {order = order} {x} {y} {z} | inr y=z | inl (inl x<z) = minCommutes order x y
|
||||
swapMin {order = order} {x} {y} {z} | inr y=z | inl (inr z<x) rewrite y=z | minIdempotent order z | minCommutes order x z = equivMin order z<x
|
||||
swapMin {order = order} {x} {y} {z} | inr y=z | inr x=z = minCommutes order x y
|
||||
|
||||
minMin : {a b : _} {A : Set a} → (order : TotalOrder {a} {b} A) → {x y : A} → (TotalOrder.min order x (TotalOrder.min order x y)) ≡ TotalOrder.min order x y
|
||||
minMin order {x} {y} with TotalOrder.totality order x y
|
||||
minMin order {x} {y} | inl (inl x<y) = minIdempotent order x
|
||||
minMin order {x} {y} | inl (inr y<x) with TotalOrder.totality order x y
|
||||
minMin order {x} {y} | inl (inr y<x) | inl (inl x<y) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) y<x x<y))
|
||||
minMin order {x} {y} | inl (inr y<x) | inl (inr y<x') = refl
|
||||
minMin order {x} {y} | inl (inr y<x) | inr x=y = x=y
|
||||
minMin order {x} {y} | inr x=y = minIdempotent order x
|
||||
minMin : {a b : _} {A : Set a} → (order : TotalOrder {a} {b} A) → {x y : A} → (TotalOrder.min order x (TotalOrder.min order x y)) ≡ TotalOrder.min order x y
|
||||
minMin order {x} {y} with TotalOrder.totality order x y
|
||||
minMin order {x} {y} | inl (inl x<y) = minIdempotent order x
|
||||
minMin order {x} {y} | inl (inr y<x) with TotalOrder.totality order x y
|
||||
minMin order {x} {y} | inl (inr y<x) | inl (inl x<y) = exFalso (PartialOrder.irreflexive (TotalOrder.order order) (PartialOrder.<Transitive (TotalOrder.order order) y<x x<y))
|
||||
minMin order {x} {y} | inl (inr y<x) | inl (inr y<x') = refl
|
||||
minMin order {x} {y} | inl (inr y<x) | inr x=y = x=y
|
||||
minMin order {x} {y} | inr x=y = minIdempotent order x
|
||||
|
||||
minFromBoth : {a b : _} {A : Set a} → {order : TotalOrder {a} {b} A} → {min x y : A} → (TotalOrder._<_ order min x) → (TotalOrder._<_ order min y) → (TotalOrder._<_ order min (TotalOrder.min order x y))
|
||||
minFromBoth {a} {order = order} {x = x} {y} prX prY with TotalOrder.totality order x y
|
||||
minFromBoth {a} prX prY | inl (inl x<y) = prX
|
||||
minFromBoth {a} prX prY | inl (inr y<x) = prY
|
||||
minFromBoth {a} prX prY | inr x=y = prX
|
||||
minFromBoth : {a b : _} {A : Set a} → {order : TotalOrder {a} {b} A} → {min x y : A} → (TotalOrder._<_ order min x) → (TotalOrder._<_ order min y) → (TotalOrder._<_ order min (TotalOrder.min order x y))
|
||||
minFromBoth {a} {order = order} {x = x} {y} prX prY with TotalOrder.totality order x y
|
||||
minFromBoth {a} prX prY | inl (inl x<y) = prX
|
||||
minFromBoth {a} prX prY | inl (inr y<x) = prY
|
||||
minFromBoth {a} prX prY | inr x=y = prX
|
||||
|
@@ -3,7 +3,8 @@
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Modulo.Group
|
||||
open import Numbers.Modulo.Definition
|
||||
|
@@ -6,7 +6,9 @@ open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Orders
|
||||
open import Rings.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Numbers.Modulo.Definition
|
||||
|
@@ -15,12 +15,15 @@ record Semiring {a : _} {A : Set a} (Zero One : A) (_+_ : A → A → A) (_*_ :
|
||||
productZeroLeft : (a : A) → Zero * a ≡ Zero
|
||||
productZeroRight : (a : A) → a * Zero ≡ Zero
|
||||
+DistributesOver* : (a b c : A) → a * (b + c) ≡ (a * b) + (a * c)
|
||||
+DistributesOver*' : (a b c : A) → (a + b) * c ≡ (a * c) + (b * c)
|
||||
+Associative = Monoid.associative monoid
|
||||
*Associative = Monoid.associative multMonoid
|
||||
productOneLeft = Monoid.idLeft multMonoid
|
||||
productOneRight = Monoid.idRight multMonoid
|
||||
sumZeroLeft = Monoid.idLeft monoid
|
||||
sumZeroRight = Monoid.idRight monoid
|
||||
+WellDefined : {a b c d : A} → (a ≡ c) → (b ≡ d) → (a + b) ≡ (c + d)
|
||||
+WellDefined refl refl = refl
|
||||
|
||||
-- (b+c)(a+a) == b(a+a) + c(a+a) == ba+ba+ca+ca == (ba+ca) + (ba+ca)
|
||||
-- (b+c)(a+a) ==? (b+c)a+(b+c)a
|
||||
|
14
Semirings/Lemmas.agda
Normal file
14
Semirings/Lemmas.agda
Normal file
@@ -0,0 +1,14 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Monoids.Definition
|
||||
open import Semirings.Definition
|
||||
|
||||
module Semirings.Lemmas {a : _} {A : Set a} {Zero One : A} {_+_ _*_ : A → A → A} (S : Semiring Zero One _+_ _*_) where
|
||||
|
||||
open Semiring S
|
||||
|
||||
doubleIsAddTwo : (a : A) → (a + a) ≡ (One + One) * a
|
||||
doubleIsAddTwo a = transitivity (equalityCommutative (+WellDefined (productOneLeft a) (productOneLeft a))) (equalityCommutative (+DistributesOver*' One One a))
|
@@ -4,7 +4,9 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.WellFounded
|
||||
open import Sets.FinSet
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
|
@@ -4,7 +4,8 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.FinSet
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
@@ -13,143 +14,143 @@ open import Sets.CantorBijection.Order
|
||||
|
||||
module Sets.CantorBijection.Proofs where
|
||||
|
||||
open Sets.CantorBijection.Order using (order ; totalOrder ; orderWellfounded)
|
||||
open Sets.CantorBijection.Order using (order ; totalOrder ; orderWellfounded)
|
||||
|
||||
cantorInverseLemma : (ℕ && ℕ) → (ℕ && ℕ)
|
||||
cantorInverseLemma (0 ,, s) = (succ s) ,, 0
|
||||
cantorInverseLemma (succ r ,, 0) = r ,, 1
|
||||
cantorInverseLemma (succ r ,, succ s) = (r ,, succ (succ s))
|
||||
cantorInverseLemma : (ℕ && ℕ) → (ℕ && ℕ)
|
||||
cantorInverseLemma (0 ,, s) = (succ s) ,, 0
|
||||
cantorInverseLemma (succ r ,, 0) = r ,, 1
|
||||
cantorInverseLemma (succ r ,, succ s) = (r ,, succ (succ s))
|
||||
|
||||
cantorInverse : ℕ → (ℕ && ℕ)
|
||||
cantorInverse zero = (0 ,, 0)
|
||||
cantorInverse (succ z) = cantorInverseLemma (cantorInverse z)
|
||||
cantorInverse : ℕ → (ℕ && ℕ)
|
||||
cantorInverse zero = (0 ,, 0)
|
||||
cantorInverse (succ z) = cantorInverseLemma (cantorInverse z)
|
||||
|
||||
cantorInverseLemmaInjective : (a b : ℕ && ℕ) → cantorInverseLemma a ≡ cantorInverseLemma b → a ≡ b
|
||||
cantorInverseLemmaInjective (zero ,, b) (zero ,, .b) refl = refl
|
||||
cantorInverseLemmaInjective (zero ,, b) (succ c ,, zero) ()
|
||||
cantorInverseLemmaInjective (zero ,, b) (succ c ,, succ d) ()
|
||||
cantorInverseLemmaInjective (succ a ,, zero) (zero ,, d) ()
|
||||
cantorInverseLemmaInjective (succ a ,, succ b) (zero ,, d) ()
|
||||
cantorInverseLemmaInjective (succ a ,, zero) (succ .a ,, zero) refl = refl
|
||||
cantorInverseLemmaInjective (succ a ,, succ b) (succ .a ,, succ .b) refl = refl
|
||||
cantorInverseLemmaInjective : (a b : ℕ && ℕ) → cantorInverseLemma a ≡ cantorInverseLemma b → a ≡ b
|
||||
cantorInverseLemmaInjective (zero ,, b) (zero ,, .b) refl = refl
|
||||
cantorInverseLemmaInjective (zero ,, b) (succ c ,, zero) ()
|
||||
cantorInverseLemmaInjective (zero ,, b) (succ c ,, succ d) ()
|
||||
cantorInverseLemmaInjective (succ a ,, zero) (zero ,, d) ()
|
||||
cantorInverseLemmaInjective (succ a ,, succ b) (zero ,, d) ()
|
||||
cantorInverseLemmaInjective (succ a ,, zero) (succ .a ,, zero) refl = refl
|
||||
cantorInverseLemmaInjective (succ a ,, succ b) (succ .a ,, succ .b) refl = refl
|
||||
|
||||
cantorInverseLemmaSurjective : (a : ℕ && ℕ) → (a ≡ (0 ,, 0)) || (Sg (ℕ && ℕ) (λ b → cantorInverseLemma b ≡ a))
|
||||
cantorInverseLemmaSurjective (zero ,, zero) = inl refl
|
||||
cantorInverseLemmaSurjective (zero ,, succ zero) = inr ((1 ,, 0) , refl)
|
||||
cantorInverseLemmaSurjective (zero ,, succ (succ b)) = inr ((1 ,, succ b) , refl)
|
||||
cantorInverseLemmaSurjective (succ a ,, zero) = inr ((0 ,, a) , refl)
|
||||
cantorInverseLemmaSurjective (succ a ,, succ zero) = inr ((succ (succ a) ,, 0) , refl)
|
||||
cantorInverseLemmaSurjective (succ a ,, succ (succ b)) = inr ((succ (succ a) ,, succ b) , refl)
|
||||
cantorInverseLemmaSurjective : (a : ℕ && ℕ) → (a ≡ (0 ,, 0)) || (Sg (ℕ && ℕ) (λ b → cantorInverseLemma b ≡ a))
|
||||
cantorInverseLemmaSurjective (zero ,, zero) = inl refl
|
||||
cantorInverseLemmaSurjective (zero ,, succ zero) = inr ((1 ,, 0) , refl)
|
||||
cantorInverseLemmaSurjective (zero ,, succ (succ b)) = inr ((1 ,, succ b) , refl)
|
||||
cantorInverseLemmaSurjective (succ a ,, zero) = inr ((0 ,, a) , refl)
|
||||
cantorInverseLemmaSurjective (succ a ,, succ zero) = inr ((succ (succ a) ,, 0) , refl)
|
||||
cantorInverseLemmaSurjective (succ a ,, succ (succ b)) = inr ((succ (succ a) ,, succ b) , refl)
|
||||
|
||||
cantorInverseLemmaNotZero : (a : ℕ && ℕ) → (cantorInverseLemma a ≡ (0 ,, 0)) → False
|
||||
cantorInverseLemmaNotZero (succ a ,, zero) ()
|
||||
cantorInverseLemmaNotZero (succ a ,, succ b) ()
|
||||
cantorInverseLemmaNotZero : (a : ℕ && ℕ) → (cantorInverseLemma a ≡ (0 ,, 0)) → False
|
||||
cantorInverseLemmaNotZero (succ a ,, zero) ()
|
||||
cantorInverseLemmaNotZero (succ a ,, succ b) ()
|
||||
|
||||
cantorInverseLemmaIncreases : (x : ℕ && ℕ) → order x (cantorInverseLemma x)
|
||||
cantorInverseLemmaIncreases (zero ,, b) = inl (identityOfIndiscernablesRight _<N_ (a<SuccA b) (applyEquality succ (equalityCommutative (Semiring.sumZeroRight ℕSemiring b))))
|
||||
cantorInverseLemmaIncreases (succ a ,, zero) = inr (transitivity (applyEquality succ (Semiring.sumZeroRight ℕSemiring a)) (Semiring.commutative ℕSemiring 1 a) ,, (le 0 refl))
|
||||
cantorInverseLemmaIncreases (succ a ,, succ b) = inr (transitivity (applyEquality succ (Semiring.commutative ℕSemiring a (succ b))) (Semiring.commutative ℕSemiring (succ (succ b)) a) ,, (le 0 refl))
|
||||
cantorInverseLemmaIncreases : (x : ℕ && ℕ) → order x (cantorInverseLemma x)
|
||||
cantorInverseLemmaIncreases (zero ,, b) = inl (identityOfIndiscernablesRight _<N_ (a<SuccA b) (applyEquality succ (equalityCommutative (Semiring.sumZeroRight ℕSemiring b))))
|
||||
cantorInverseLemmaIncreases (succ a ,, zero) = inr (transitivity (applyEquality succ (Semiring.sumZeroRight ℕSemiring a)) (Semiring.commutative ℕSemiring 1 a) ,, (le 0 refl))
|
||||
cantorInverseLemmaIncreases (succ a ,, succ b) = inr (transitivity (applyEquality succ (Semiring.commutative ℕSemiring a (succ b))) (Semiring.commutative ℕSemiring (succ (succ b)) a) ,, (le 0 refl))
|
||||
|
||||
cantorInverseOrderPreserving : (x y : ℕ) → (x <N y) → order (cantorInverse x) (cantorInverse y)
|
||||
cantorInverseOrderPreserving zero (succ y) x<y with TotalOrder.totality totalOrder (0 ,, 0) (cantorInverseLemma (cantorInverse y))
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inl (inl bl) = bl
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inl (inr bl) = exFalso (leastElement bl)
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inr x = exFalso (leastElement {cantorInverse y} (identityOfIndiscernablesRight order (cantorInverseLemmaIncreases (cantorInverse y)) (equalityCommutative x)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y with cantorInverseOrderPreserving x y (canRemoveSuccFrom<N x<y)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr with cantorInverse x
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | x1 ,, x2 with cantorInverse y
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | zero ,, succ y2 = inl (succPreservesInequality (succIsPositive (y2 +N 0)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero with TotalOrder.totality ℕTotalOrder 1 (y1 +N 1)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inl (inl pr1) = inl pr1
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inl (inr pr1) rewrite Semiring.commutative ℕSemiring y1 1 = exFalso (zeroNeverGreater (canRemoveSuccFrom<N pr1))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inr pr1 = inr (pr1 ,, le 0 refl)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring y1 (succ (succ y2)) = inl (succPreservesInequality (succIsPositive (y2 +N y1)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, zero | zero ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.sumZeroRight ℕSemiring y2 | Semiring.sumZeroRight ℕSemiring x1 = inl (TotalOrder.<Transitive ℕTotalOrder pr (a<SuccA _))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, zero | succ y1 ,, zero rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring x1 | Semiring.sumZeroRight ℕSemiring y1 = inl pr
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, zero | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.sumZeroRight ℕSemiring x1 = inl (identityOfIndiscernablesRight _<N_ pr (transitivity (applyEquality succ (Semiring.commutative ℕSemiring y1 (succ y2))) (Semiring.commutative ℕSemiring (succ (succ y2)) y1)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, succ x2 | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring x2 | Semiring.sumZeroRight ℕSemiring y2 = inl (succPreservesInequality pr)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, succ x2 | succ y1 ,, zero rewrite Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring y1 | Semiring.sumZeroRight ℕSemiring x2 = ans
|
||||
cantorInverseOrderPreserving : (x y : ℕ) → (x <N y) → order (cantorInverse x) (cantorInverse y)
|
||||
cantorInverseOrderPreserving zero (succ y) x<y with TotalOrder.totality totalOrder (0 ,, 0) (cantorInverseLemma (cantorInverse y))
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inl (inl bl) = bl
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inl (inr bl) = exFalso (leastElement bl)
|
||||
cantorInverseOrderPreserving zero (succ y) x<y | inr x = exFalso (leastElement {cantorInverse y} (identityOfIndiscernablesRight order (cantorInverseLemmaIncreases (cantorInverse y)) (equalityCommutative x)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y with cantorInverseOrderPreserving x y (canRemoveSuccFrom<N x<y)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr with cantorInverse x
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | x1 ,, x2 with cantorInverse y
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | zero ,, succ y2 = inl (succPreservesInequality (succIsPositive (y2 +N 0)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero with TotalOrder.totality ℕTotalOrder 1 (y1 +N 1)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inl (inl pr1) = inl pr1
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inl (inr pr1) rewrite Semiring.commutative ℕSemiring y1 1 = exFalso (zeroNeverGreater (canRemoveSuccFrom<N pr1))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, zero | inr pr1 = inr (pr1 ,, le 0 refl)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, zero | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring y1 (succ (succ y2)) = inl (succPreservesInequality (succIsPositive (y2 +N y1)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, zero | zero ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.sumZeroRight ℕSemiring y2 | Semiring.sumZeroRight ℕSemiring x1 = inl (TotalOrder.<Transitive ℕTotalOrder pr (a<SuccA _))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, zero | succ y1 ,, zero rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring x1 | Semiring.sumZeroRight ℕSemiring y1 = inl pr
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, zero | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.sumZeroRight ℕSemiring x1 = inl (identityOfIndiscernablesRight _<N_ pr (transitivity (applyEquality succ (Semiring.commutative ℕSemiring y1 (succ y2))) (Semiring.commutative ℕSemiring (succ (succ y2)) y1)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, succ x2 | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring x2 | Semiring.sumZeroRight ℕSemiring y2 = inl (succPreservesInequality pr)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, succ x2 | succ y1 ,, zero rewrite Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring y1 | Semiring.sumZeroRight ℕSemiring x2 = ans
|
||||
where
|
||||
ans : (succ (succ x2) <N succ y1) || (succ (succ x2) ≡ succ y1) && (zero <N 1)
|
||||
ans with TotalOrder.totality ℕTotalOrder (succ (succ x2)) (succ y1)
|
||||
ans | inl (inl x) = inl x
|
||||
ans | inl (inr x) = exFalso (noIntegersBetweenXAndSuccX (succ x2) pr x)
|
||||
ans | inr x = inr (x ,, le zero refl)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, succ x2 | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring y1 | Semiring.sumZeroRight ℕSemiring x2 = ans
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | zero ,, succ x2 | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring y1 | Semiring.sumZeroRight ℕSemiring x2 = ans
|
||||
where
|
||||
ans : (succ (succ x2) <N y1 +N succ (succ y2)) || (succ (succ x2) ≡ y1 +N succ (succ y2)) && (zero <N succ (succ y2))
|
||||
ans with TotalOrder.totality ℕTotalOrder (succ (succ x2)) (y1 +N succ (succ y2))
|
||||
ans | inl (inl x) = inl x
|
||||
ans | inl (inr x) = exFalso (noIntegersBetweenXAndSuccX (succ x2) pr (identityOfIndiscernablesLeft _<N_ x (transitivity (Semiring.commutative ℕSemiring y1 (succ (succ y2))) (applyEquality succ (Semiring.commutative ℕSemiring (succ y2) y1)))))
|
||||
ans | inr x = inr (x ,, succIsPositive (succ y2))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, succ x2 | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring y2 = inl (TotalOrder.<Transitive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ pr (transitivity (applyEquality succ (Semiring.commutative ℕSemiring x1 (succ x2))) (Semiring.commutative ℕSemiring (succ (succ x2)) x1))) (a<SuccA (succ y2)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, succ x2 | succ y1 ,, zero rewrite Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring y1 | Semiring.commutative ℕSemiring x1 (succ x2) | Semiring.commutative ℕSemiring x1 (succ (succ x2)) = inl pr
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, succ x2 | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 (succ x2) | Semiring.commutative ℕSemiring y1 (succ y2) | Semiring.commutative ℕSemiring x1 (succ (succ x2)) | Semiring.commutative ℕSemiring y1 (succ (succ y2)) = inl pr
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) with cantorInverse x
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | x1 ,, x2 with cantorInverse y
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, ()) | zero ,, zero | zero ,, zero
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (() ,, snd) | zero ,, zero | zero ,, succ y2
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (refl ,, snd) | zero ,, succ .y2 | zero ,, succ y2 = exFalso (TotalOrder.irreflexive ℕTotalOrder snd)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (refl ,, snd) | zero ,, succ .(y1 +N succ y2) | succ y1 ,, succ y2 = exFalso (TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder snd (identityOfIndiscernablesRight _<N_ (addingIncreases (succ y2) y1) (Semiring.commutative ℕSemiring (succ y2) (succ y1)))))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, zero | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring x1 | succInjective fst | Semiring.commutative ℕSemiring y2 1 | Semiring.sumZeroRight ℕSemiring y2 = inl (le zero refl)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, zero | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.sumZeroRight ℕSemiring x1 = inr (transitivity fst (transitivity (applyEquality succ (Semiring.commutative ℕSemiring y1 (succ y2))) (Semiring.commutative ℕSemiring (succ (succ y2)) y1)) ,, succPreservesInequality snd)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, succ x2 | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring y2 | Semiring.commutative ℕSemiring x1 (succ x2) | Semiring.commutative ℕSemiring (succ (succ x2)) x1 | fst = inl (succPreservesInequality (le zero refl))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, succ x2 | succ y1 ,, succ y2 = inr (transitivity (transitivity (Semiring.commutative ℕSemiring x1 (succ (succ x2))) (applyEquality succ (Semiring.commutative ℕSemiring (succ x2) x1))) (transitivity fst (transitivity (applyEquality succ (Semiring.commutative ℕSemiring y1 (succ y2))) (Semiring.commutative ℕSemiring (succ (succ y2)) y1))) ,, succPreservesInequality snd)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, succ x2 | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring y2 = inl (TotalOrder.<Transitive ℕTotalOrder (identityOfIndiscernablesLeft _<N_ pr (transitivity (applyEquality succ (Semiring.commutative ℕSemiring x1 (succ x2))) (Semiring.commutative ℕSemiring (succ (succ x2)) x1))) (a<SuccA (succ y2)))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, succ x2 | succ y1 ,, zero rewrite Semiring.commutative ℕSemiring y1 1 | Semiring.sumZeroRight ℕSemiring y1 | Semiring.commutative ℕSemiring x1 (succ x2) | Semiring.commutative ℕSemiring x1 (succ (succ x2)) = inl pr
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inl pr | succ x1 ,, succ x2 | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 (succ x2) | Semiring.commutative ℕSemiring y1 (succ y2) | Semiring.commutative ℕSemiring x1 (succ (succ x2)) | Semiring.commutative ℕSemiring y1 (succ (succ y2)) = inl pr
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) with cantorInverse x
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | x1 ,, x2 with cantorInverse y
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, ()) | zero ,, zero | zero ,, zero
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (() ,, snd) | zero ,, zero | zero ,, succ y2
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (refl ,, snd) | zero ,, succ .y2 | zero ,, succ y2 = exFalso (TotalOrder.irreflexive ℕTotalOrder snd)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (refl ,, snd) | zero ,, succ .(y1 +N succ y2) | succ y1 ,, succ y2 = exFalso (TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder snd (identityOfIndiscernablesRight _<N_ (addingIncreases (succ y2) y1) (Semiring.commutative ℕSemiring (succ y2) (succ y1)))))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, zero | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring x1 | succInjective fst | Semiring.commutative ℕSemiring y2 1 | Semiring.sumZeroRight ℕSemiring y2 = inl (le zero refl)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, zero | succ y1 ,, succ y2 rewrite Semiring.commutative ℕSemiring x1 1 | Semiring.sumZeroRight ℕSemiring x1 = inr (transitivity fst (transitivity (applyEquality succ (Semiring.commutative ℕSemiring y1 (succ y2))) (Semiring.commutative ℕSemiring (succ (succ y2)) y1)) ,, succPreservesInequality snd)
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, succ x2 | zero ,, succ y2 rewrite Semiring.sumZeroRight ℕSemiring y2 | Semiring.commutative ℕSemiring x1 (succ x2) | Semiring.commutative ℕSemiring (succ (succ x2)) x1 | fst = inl (succPreservesInequality (le zero refl))
|
||||
cantorInverseOrderPreserving (succ x) (succ y) x<y | inr (fst ,, snd) | succ x1 ,, succ x2 | succ y1 ,, succ y2 = inr (transitivity (transitivity (Semiring.commutative ℕSemiring x1 (succ (succ x2))) (applyEquality succ (Semiring.commutative ℕSemiring (succ x2) x1))) (transitivity fst (transitivity (applyEquality succ (Semiring.commutative ℕSemiring y1 (succ y2))) (Semiring.commutative ℕSemiring (succ (succ y2)) y1))) ,, succPreservesInequality snd)
|
||||
|
||||
cantorInverseInjective : (a b : ℕ) → cantorInverse a ≡ cantorInverse b → a ≡ b
|
||||
cantorInverseInjective zero zero pr = refl
|
||||
cantorInverseInjective zero (succ b) pr = exFalso (cantorInverseLemmaNotZero _ (equalityCommutative pr))
|
||||
cantorInverseInjective (succ a) zero pr = exFalso (cantorInverseLemmaNotZero _ pr)
|
||||
cantorInverseInjective (succ a) (succ b) pr = applyEquality succ (cantorInverseInjective a b (cantorInverseLemmaInjective (cantorInverse a) (cantorInverse b) pr))
|
||||
cantorInverseInjective : (a b : ℕ) → cantorInverse a ≡ cantorInverse b → a ≡ b
|
||||
cantorInverseInjective zero zero pr = refl
|
||||
cantorInverseInjective zero (succ b) pr = exFalso (cantorInverseLemmaNotZero _ (equalityCommutative pr))
|
||||
cantorInverseInjective (succ a) zero pr = exFalso (cantorInverseLemmaNotZero _ pr)
|
||||
cantorInverseInjective (succ a) (succ b) pr = applyEquality succ (cantorInverseInjective a b (cantorInverseLemmaInjective (cantorInverse a) (cantorInverse b) pr))
|
||||
|
||||
-- Some unnecessary things on the way to the final proof
|
||||
cantorInverseDiscrete : (a : ℕ) → (c : ℕ && ℕ) → order (cantorInverse a) c → order c (cantorInverse (succ a)) → False
|
||||
cantorInverseDiscrete zero (zero ,, zero) (inl ()) c<sa
|
||||
cantorInverseDiscrete zero (zero ,, zero) (inr ()) c<sa
|
||||
cantorInverseDiscrete zero (zero ,, succ c) a<c (inl x) = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete zero (succ b ,, zero) a<c (inl x) = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete zero (succ b ,, succ c) a<c (inl x) = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) a<c c<sa with cantorInverse a
|
||||
cantorInverseDiscrete (succ a) (zero ,, zero) (inl ()) (inl x) | zero ,, zero
|
||||
cantorInverseDiscrete (succ a) (zero ,, zero) (inr ()) (inl x) | zero ,, zero
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ zero) a<c (inl x) | zero ,, zero = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ (succ c)) a<c (inl x) | zero ,, zero = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete (succ a) (succ b ,, c) a<c (inl x) | zero ,, zero = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete (succ a) (b ,, zero) (inl x) (inr (fst ,, snd)) | zero ,, zero = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _<N_ x fst)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) a<c (inr (fst ,, snd)) | zero ,, zero = zeroNeverGreater (canRemoveSuccFrom<N snd)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl y) (inl x) | zero ,, succ snd rewrite Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder x y)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) (inr (fst ,, _)) (inl x) | zero ,, succ snd rewrite Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd | fst = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, zero) (inl x) (inr (fst ,, snd₁)) | zero ,, succ snd rewrite fst | Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) (inl x) (inr (fst ,, snd1)) | zero ,, succ snd = zeroNeverGreater (canRemoveSuccFrom<N snd1)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ zero) (inr (fst ,, _)) (inr (fst₁ ,, bad)) | zero ,, succ snd = TotalOrder.irreflexive ℕTotalOrder bad
|
||||
cantorInverseDiscrete (succ a) (b ,, succ (succ c)) (inr (fst ,, _)) (inr (fst₁ ,, bad)) | zero ,, succ snd = zeroNeverGreater (canRemoveSuccFrom<N bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x1) (inl x) | succ zero ,, zero = noIntegersBetweenXAndSuccX 1 x1 x
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ zero) (inr (fst ,, snd)) (inl x) | succ zero ,, zero = TotalOrder.irreflexive ℕTotalOrder snd
|
||||
cantorInverseDiscrete (succ a) (succ b ,, succ c) (inr (fst ,, snd)) (inl x) | succ zero ,, zero rewrite Semiring.commutative ℕSemiring b (succ c) = bad fst
|
||||
cantorInverseDiscrete : (a : ℕ) → (c : ℕ && ℕ) → order (cantorInverse a) c → order c (cantorInverse (succ a)) → False
|
||||
cantorInverseDiscrete zero (zero ,, zero) (inl ()) c<sa
|
||||
cantorInverseDiscrete zero (zero ,, zero) (inr ()) c<sa
|
||||
cantorInverseDiscrete zero (zero ,, succ c) a<c (inl x) = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete zero (succ b ,, zero) a<c (inl x) = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete zero (succ b ,, succ c) a<c (inl x) = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) a<c c<sa with cantorInverse a
|
||||
cantorInverseDiscrete (succ a) (zero ,, zero) (inl ()) (inl x) | zero ,, zero
|
||||
cantorInverseDiscrete (succ a) (zero ,, zero) (inr ()) (inl x) | zero ,, zero
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ zero) a<c (inl x) | zero ,, zero = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ (succ c)) a<c (inl x) | zero ,, zero = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete (succ a) (succ b ,, c) a<c (inl x) | zero ,, zero = zeroNeverGreater (canRemoveSuccFrom<N x)
|
||||
cantorInverseDiscrete (succ a) (b ,, zero) (inl x) (inr (fst ,, snd)) | zero ,, zero = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _<N_ x fst)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) a<c (inr (fst ,, snd)) | zero ,, zero = zeroNeverGreater (canRemoveSuccFrom<N snd)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl y) (inl x) | zero ,, succ snd rewrite Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder x y)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) (inr (fst ,, _)) (inl x) | zero ,, succ snd rewrite Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd | fst = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, zero) (inl x) (inr (fst ,, snd₁)) | zero ,, succ snd rewrite fst | Semiring.commutative ℕSemiring snd 1 | Semiring.sumZeroRight ℕSemiring snd = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, succ c) (inl x) (inr (fst ,, snd1)) | zero ,, succ snd = zeroNeverGreater (canRemoveSuccFrom<N snd1)
|
||||
cantorInverseDiscrete (succ a) (b ,, succ zero) (inr (fst ,, _)) (inr (fst₁ ,, bad)) | zero ,, succ snd = TotalOrder.irreflexive ℕTotalOrder bad
|
||||
cantorInverseDiscrete (succ a) (b ,, succ (succ c)) (inr (fst ,, _)) (inr (fst₁ ,, bad)) | zero ,, succ snd = zeroNeverGreater (canRemoveSuccFrom<N bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x1) (inl x) | succ zero ,, zero = noIntegersBetweenXAndSuccX 1 x1 x
|
||||
cantorInverseDiscrete (succ a) (zero ,, succ zero) (inr (fst ,, snd)) (inl x) | succ zero ,, zero = TotalOrder.irreflexive ℕTotalOrder snd
|
||||
cantorInverseDiscrete (succ a) (succ b ,, succ c) (inr (fst ,, snd)) (inl x) | succ zero ,, zero rewrite Semiring.commutative ℕSemiring b (succ c) = bad fst
|
||||
where
|
||||
bad : {a : ℕ} → 1 ≡ succ (succ a) → False
|
||||
bad ()
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl y) (inl x) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder y x)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (bad ,, _)) (inl x) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _<N_ x bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inr (bad ,, snd)) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _<N_ x bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (_ ,, bad)) (inr (fst₁ ,, snd)) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = noIntegersBetweenXAndSuccX 1 bad snd
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inl y) | succ zero ,, succ snd rewrite Semiring.sumZeroRight ℕSemiring snd = noIntegersBetweenXAndSuccX (succ (succ snd)) x y
|
||||
cantorInverseDiscrete (succ a) (zero ,, c) (inr (fst ,, snd1)) (inl y) | succ zero ,, succ snd rewrite Semiring.sumZeroRight ℕSemiring snd = noIntegersBetweenXAndSuccX (succ (succ snd)) snd1 y
|
||||
cantorInverseDiscrete (succ a) (succ b ,, c) (inr (fst ,, bad)) (inl y) | succ zero ,, succ snd rewrite Semiring.sumZeroRight ℕSemiring snd | fst = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder bad (identityOfIndiscernablesRight _<N_ (addingIncreases c b) (Semiring.commutative ℕSemiring c (succ b))))
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inl y) | succ (succ fst) ,, succ snd rewrite Semiring.commutative ℕSemiring fst (succ (succ (succ snd))) | Semiring.commutative ℕSemiring (succ (succ snd)) fst = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder y x)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inr (y ,, z)) | succ (succ fst) ,, succ snd rewrite y | Semiring.commutative ℕSemiring fst (succ (succ (succ snd))) | Semiring.commutative ℕSemiring (succ (succ snd)) fst = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (x ,, y)) (inl z) | succ (succ fst) ,, succ snd rewrite equalityCommutative x | Semiring.commutative ℕSemiring fst (succ (succ (succ snd))) | Semiring.commutative ℕSemiring (succ (succ snd)) fst = TotalOrder.irreflexive ℕTotalOrder z
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (x ,, y)) (inr (m ,, n)) | succ (succ fst) ,, succ snd = noIntegersBetweenXAndSuccX (succ (succ snd)) y n
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl y) (inl x) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder y x)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (bad ,, _)) (inl x) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _<N_ x bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inr (bad ,, snd)) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = TotalOrder.irreflexive ℕTotalOrder (identityOfIndiscernablesRight _<N_ x bad)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (_ ,, bad)) (inr (fst₁ ,, snd)) | succ (succ fst) ,, zero rewrite Semiring.commutative ℕSemiring fst 2 | Semiring.commutative ℕSemiring fst 1 = noIntegersBetweenXAndSuccX 1 bad snd
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inl y) | succ zero ,, succ snd rewrite Semiring.sumZeroRight ℕSemiring snd = noIntegersBetweenXAndSuccX (succ (succ snd)) x y
|
||||
cantorInverseDiscrete (succ a) (zero ,, c) (inr (fst ,, snd1)) (inl y) | succ zero ,, succ snd rewrite Semiring.sumZeroRight ℕSemiring snd = noIntegersBetweenXAndSuccX (succ (succ snd)) snd1 y
|
||||
cantorInverseDiscrete (succ a) (succ b ,, c) (inr (fst ,, bad)) (inl y) | succ zero ,, succ snd rewrite Semiring.sumZeroRight ℕSemiring snd | fst = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder bad (identityOfIndiscernablesRight _<N_ (addingIncreases c b) (Semiring.commutative ℕSemiring c (succ b))))
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inl y) | succ (succ fst) ,, succ snd rewrite Semiring.commutative ℕSemiring fst (succ (succ (succ snd))) | Semiring.commutative ℕSemiring (succ (succ snd)) fst = TotalOrder.irreflexive ℕTotalOrder (TotalOrder.<Transitive ℕTotalOrder y x)
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inl x) (inr (y ,, z)) | succ (succ fst) ,, succ snd rewrite y | Semiring.commutative ℕSemiring fst (succ (succ (succ snd))) | Semiring.commutative ℕSemiring (succ (succ snd)) fst = TotalOrder.irreflexive ℕTotalOrder x
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (x ,, y)) (inl z) | succ (succ fst) ,, succ snd rewrite equalityCommutative x | Semiring.commutative ℕSemiring fst (succ (succ (succ snd))) | Semiring.commutative ℕSemiring (succ (succ snd)) fst = TotalOrder.irreflexive ℕTotalOrder z
|
||||
cantorInverseDiscrete (succ a) (b ,, c) (inr (x ,, y)) (inr (m ,, n)) | succ (succ fst) ,, succ snd = noIntegersBetweenXAndSuccX (succ (succ snd)) y n
|
||||
|
||||
boundedInversesExist : (a : ℕ && ℕ) (s : ℕ) → order a (cantorInverse s) → Sg ℕ (λ i → cantorInverse i ≡ a)
|
||||
boundedInversesExist a zero a<s = exFalso (leastElement a<s)
|
||||
boundedInversesExist a (succ s) a<s with TotalOrder.totality totalOrder a (cantorInverse s)
|
||||
boundedInversesExist a (succ s) _ | inl (inl a<s) = boundedInversesExist a s a<s
|
||||
boundedInversesExist a (succ s) a<s | inl (inr s<a) = exFalso (cantorInverseDiscrete s a s<a a<s)
|
||||
boundedInversesExist a (succ s) a<s | inr a=s = s , equalityCommutative a=s
|
||||
boundedInversesExist : (a : ℕ && ℕ) (s : ℕ) → order a (cantorInverse s) → Sg ℕ (λ i → cantorInverse i ≡ a)
|
||||
boundedInversesExist a zero a<s = exFalso (leastElement a<s)
|
||||
boundedInversesExist a (succ s) a<s with TotalOrder.totality totalOrder a (cantorInverse s)
|
||||
boundedInversesExist a (succ s) _ | inl (inl a<s) = boundedInversesExist a s a<s
|
||||
boundedInversesExist a (succ s) a<s | inl (inr s<a) = exFalso (cantorInverseDiscrete s a s<a a<s)
|
||||
boundedInversesExist a (succ s) a<s | inr a=s = s , equalityCommutative a=s
|
||||
|
||||
cantorInverseSurjective : (x : ℕ && ℕ) → Sg ℕ (λ i → (cantorInverse i) ≡ x)
|
||||
cantorInverseSurjective = rec orderWellfounded (λ z → Sg ℕ (λ z₁ → (cantorInverse z₁) ≡ z)) go
|
||||
cantorInverseSurjective : (x : ℕ && ℕ) → Sg ℕ (λ i → (cantorInverse i) ≡ x)
|
||||
cantorInverseSurjective = rec orderWellfounded (λ z → Sg ℕ (λ z₁ → (cantorInverse z₁) ≡ z)) go
|
||||
where
|
||||
go : (a : ℕ && ℕ) → (pr : (x : ℕ && ℕ) (x₁ : order x a) → Sg ℕ (λ z → (cantorInverse z) ≡ x)) → Sg ℕ (λ i → (cantorInverse i) ≡ a)
|
||||
go a pr with cantorInverseLemmaSurjective a
|
||||
|
@@ -4,7 +4,9 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Sets.FinSet
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
@@ -12,124 +14,127 @@ open import WellFoundedInduction
|
||||
open import Sets.CantorBijection.CantorBijection
|
||||
|
||||
module Sets.Cardinality where
|
||||
record CountablyInfiniteSet {a : _} (A : Set a) : Set a 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
|
||||
data Countable {a : _} (A : Set a) : Set a where
|
||||
finite : FiniteSet A → Countable A
|
||||
infinite : CountablyInfiniteSet A → Countable A
|
||||
|
||||
ℕCountable : CountablyInfiniteSet ℕ
|
||||
ℕCountable = record { counting = id ; countingIsBij = invertibleImpliesBijection (record { inverse = id ; isLeft = λ b → refl ; isRight = λ a → refl}) }
|
||||
ℕ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
|
||||
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''
|
||||
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
|
||||
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 orderIsTotal b (2 *N a)
|
||||
sqrtFloor (succ n) | (a ,, b) , pr | inl (inl x) = (a ,, succ b) , (p ,, q)
|
||||
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 _)
|
||||
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))
|
||||
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)
|
||||
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)
|
||||
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
|
||||
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
|
||||
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
|
||||
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
|
||||
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
|
||||
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
|
||||
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
|
||||
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 Surjection.property 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
|
||||
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 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
|
||||
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
|
||||
|
||||
N^2Countable : CountablyInfiniteSet (ℕ && ℕ)
|
||||
CountablyInfiniteSet.counting N^2Countable x = Invertible.inverse (bijectionImpliesInvertible (cantorBijection)) x
|
||||
CountablyInfiniteSet.countingIsBij N^2Countable = invertibleImpliesBijection (inverseIsInvertible (bijectionImpliesInvertible cantorBijection))
|
||||
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)
|
||||
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)
|
||||
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)
|
||||
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
|
||||
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
|
||||
|
||||
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
|
||||
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
|
||||
|
||||
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
|
||||
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))
|
||||
|
@@ -1,6 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Sets.FinSet
|
||||
open import LogicalFormulae
|
||||
|
@@ -1,10 +1,11 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Functions
|
||||
open import Semirings.Definition
|
||||
open import Orders
|
||||
|
||||
data Vec {a : _} (X : Set a) : ℕ -> Set a where
|
||||
[] : Vec X zero
|
||||
@@ -175,7 +176,7 @@ all4<=4 : Vec (4 <= 4) 1
|
||||
all4<=4 = os (os (os (os oz))) ,- []
|
||||
|
||||
<=Is≤N : {n m : ℕ} → (n <= m) → n ≤N m
|
||||
<=Is≤N {n} {m} thm with orderIsTotal n m
|
||||
<=Is≤N {n} {m} thm with TotalOrder.totality ℕTotalOrder n m
|
||||
<=Is≤N {n} {m} thm | inl (inl n<m) = inl n<m
|
||||
<=Is≤N {n} {.n} thm | inr refl = inr refl
|
||||
<=Is≤N {zero} {zero} thm | inl (inr m<n) = inl m<n
|
||||
|
Reference in New Issue
Block a user