mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-08 21:38:40 +00:00
Split out Integers and remove use of K (#39)
This commit is contained in:
@@ -5,6 +5,8 @@
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.BinaryNaturals.Definition
|
||||
|
||||
open import Numbers.Integers.Integers
|
||||
|
||||
open import Lists.Lists
|
||||
|
||||
open import Groups.Groups
|
||||
|
@@ -3,7 +3,6 @@
|
||||
-- This file contains everything that is --safe, but uses K.
|
||||
|
||||
open import PrimeNumbers
|
||||
open import Numbers.Integers
|
||||
open import Numbers.Rationals
|
||||
open import Numbers.RationalsLemmas
|
||||
open import Numbers.BinaryNaturals.Multiplication -- TODO there's no reason for this to need K
|
||||
|
@@ -5,7 +5,7 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers
|
||||
open import Numbers.Integers.Integers
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
|
@@ -3,7 +3,7 @@
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Orders
|
||||
open import Numbers.Integers
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Sets.FinSet
|
||||
|
@@ -5,13 +5,14 @@ open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Rationals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Rings.Definition
|
||||
open import IntegersModN
|
||||
open import Semirings.Definition
|
||||
|
||||
module Groups.LectureNotes.Lecture1 where
|
||||
|
||||
@@ -27,13 +28,19 @@ module Groups.LectureNotes.Lecture1 where
|
||||
integersMinusNotGroup record { wellDefined = wellDefined ; identity = identity ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multAssoc {nonneg 3} {nonneg 2} {nonneg 1}
|
||||
integersMinusNotGroup record { wellDefined = wellDefined ; identity = identity ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | ()
|
||||
|
||||
negSuccInjective : {a b : ℕ} → (negSucc a ≡ negSucc b) → a ≡ b
|
||||
negSuccInjective {a} {.a} refl = refl
|
||||
|
||||
nonnegInjective : {a b : ℕ} → (nonneg a ≡ nonneg b) → a ≡ b
|
||||
nonnegInjective {a} {.a} refl = refl
|
||||
|
||||
integersTimesNotGroup : Group (reflSetoid ℤ) (_*Z_) → False
|
||||
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg zero) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with multIdentLeft {negSucc 1}
|
||||
... | ()
|
||||
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ zero)) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with invLeft {nonneg zero}
|
||||
... | bl with inverse (nonneg zero)
|
||||
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ zero)) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | nonneg zero
|
||||
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ zero)) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | nonneg (succ x)
|
||||
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ zero)) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | p | nonneg (succ x) = naughtE (nonnegInjective (transitivity (applyEquality nonneg (equalityCommutative (Semiring.productZeroRight ℕSemiring x))) p))
|
||||
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ zero)) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } | () | negSucc x
|
||||
integersTimesNotGroup record { wellDefined = wellDefined ; identity = (nonneg (succ (succ x))) ; inverse = inverse ; multAssoc = multAssoc ; multIdentRight = multIdentRight ; multIdentLeft = multIdentLeft ; invLeft = invLeft ; invRight = invRight } with succInjective (negSuccInjective (multIdentLeft {negSucc 1}))
|
||||
... | ()
|
||||
|
File diff suppressed because it is too large
Load Diff
95
Numbers/Integers/Addition.agda
Normal file
95
Numbers/Integers/Addition.agda
Normal file
@@ -0,0 +1,95 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers.Definition
|
||||
open import Semirings.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Setoids.Setoids
|
||||
|
||||
module Numbers.Integers.Addition where
|
||||
|
||||
infix 15 _+Z_
|
||||
_+Z_ : ℤ → ℤ → ℤ
|
||||
nonneg zero +Z b = b
|
||||
nonneg (succ x) +Z nonneg y = nonneg (succ x +N y)
|
||||
nonneg (succ x) +Z negSucc zero = nonneg x
|
||||
nonneg (succ x) +Z negSucc (succ y) = nonneg x +Z negSucc y
|
||||
negSucc x +Z nonneg zero = negSucc x
|
||||
negSucc zero +Z nonneg (succ y) = nonneg y
|
||||
negSucc (succ x) +Z nonneg (succ y) = negSucc x +Z nonneg y
|
||||
negSucc x +Z negSucc y = negSucc (succ x +N y)
|
||||
|
||||
+Zinherits : (a b : ℕ) → nonneg (a +N b) ≡ (nonneg a) +Z (nonneg b)
|
||||
+Zinherits zero b = refl
|
||||
+Zinherits (succ a) b = refl
|
||||
|
||||
+ZCommutative : (a b : ℤ) → a +Z b ≡ b +Z a
|
||||
+ZCommutative (nonneg zero) (nonneg zero) = refl
|
||||
+ZCommutative (nonneg zero) (nonneg (succ y)) = applyEquality (λ i → nonneg (succ i)) (Semiring.commutative ℕSemiring 0 y)
|
||||
+ZCommutative (nonneg (succ x)) (nonneg zero) = applyEquality (λ i → nonneg (succ i)) (Semiring.commutative ℕSemiring x 0)
|
||||
+ZCommutative (nonneg (succ x)) (nonneg (succ y)) = applyEquality (λ i → nonneg (succ i)) (transitivity (Semiring.commutative ℕSemiring x (succ y)) (transitivity (applyEquality succ (Semiring.commutative ℕSemiring y x)) (Semiring.commutative ℕSemiring (succ x) y)))
|
||||
+ZCommutative (nonneg zero) (negSucc y) = refl
|
||||
+ZCommutative (nonneg (succ x)) (negSucc zero) = refl
|
||||
+ZCommutative (nonneg (succ x)) (negSucc (succ y)) = +ZCommutative (nonneg x) (negSucc y)
|
||||
+ZCommutative (negSucc x) (nonneg zero) = refl
|
||||
+ZCommutative (negSucc zero) (nonneg (succ y)) = refl
|
||||
+ZCommutative (negSucc (succ x)) (nonneg (succ y)) = +ZCommutative (negSucc x) (nonneg y)
|
||||
+ZCommutative (negSucc x) (negSucc y) = applyEquality (λ i → negSucc (succ i)) (Semiring.commutative ℕSemiring x y)
|
||||
|
||||
+ZAssociative : (a b c : ℤ) → a +Z (b +Z c) ≡ (a +Z b) +Z c
|
||||
+ZAssociative (nonneg zero) (nonneg b) (nonneg c) = refl
|
||||
+ZAssociative (nonneg (succ a)) (nonneg zero) (nonneg c) rewrite Semiring.commutative ℕSemiring a 0 = refl
|
||||
+ZAssociative (nonneg (succ a)) (nonneg (succ b)) (nonneg c) = applyEquality (λ i → nonneg (succ i)) (Semiring.+Associative ℕSemiring a (succ b) c)
|
||||
+ZAssociative (nonneg zero) (nonneg b) (negSucc c) = refl
|
||||
+ZAssociative (nonneg (succ a)) (nonneg zero) (negSucc c) rewrite Semiring.sumZeroRight ℕSemiring a = refl
|
||||
+ZAssociative (nonneg (succ a)) (nonneg (succ b)) (negSucc zero) = applyEquality nonneg (transitivity (applyEquality succ (Semiring.commutative ℕSemiring a b)) (Semiring.commutative ℕSemiring (succ b) a))
|
||||
+ZAssociative (nonneg (succ a)) (nonneg (succ b)) (negSucc (succ c)) = transitivity (+ZAssociative (nonneg (succ a)) (nonneg b) (negSucc c)) (applyEquality (λ i → nonneg i +Z negSucc c) (transitivity (applyEquality succ (Semiring.commutative ℕSemiring a b)) (Semiring.commutative ℕSemiring (succ b) a)))
|
||||
+ZAssociative (nonneg zero) (negSucc b) c = refl
|
||||
+ZAssociative (nonneg (succ a)) (negSucc zero) (nonneg zero) = +ZCommutative (nonneg 0) (nonneg a)
|
||||
+ZAssociative (nonneg (succ a)) (negSucc zero) (nonneg (succ c)) = transitivity (applyEquality nonneg (transitivity (applyEquality succ (Semiring.commutative ℕSemiring a c)) (Semiring.commutative ℕSemiring (succ c) a))) (+Zinherits a (succ c))
|
||||
+ZAssociative (nonneg (succ a)) (negSucc zero) (negSucc c) = refl
|
||||
+ZAssociative (nonneg (succ a)) (negSucc (succ b)) (nonneg zero) = +ZCommutative (nonneg 0) _
|
||||
+ZAssociative (nonneg (succ a)) (negSucc (succ b)) (nonneg (succ c)) = transitivity (applyEquality (nonneg (succ a) +Z_) (+ZCommutative (negSucc b) (nonneg c))) (transitivity (+ZAssociative (nonneg (succ a)) (nonneg c) (negSucc b)) (transitivity (transitivity (transitivity (transitivity (applyEquality (λ i → nonneg (succ i) +Z negSucc b) (Semiring.commutative ℕSemiring a c)) (+ZCommutative (nonneg (succ (c +N a))) (negSucc b))) (applyEquality (negSucc b +Z_) (+ZCommutative (nonneg (succ c)) (nonneg a)))) (+ZAssociative (negSucc b) (nonneg a) (nonneg (succ c)))) (applyEquality (_+Z nonneg (succ c)) (+ZCommutative (negSucc b) (nonneg a)))))
|
||||
+ZAssociative (nonneg (succ a)) (negSucc (succ b)) (negSucc c) = +ZAssociative (nonneg a) (negSucc b) (negSucc c)
|
||||
+ZAssociative (negSucc a) (nonneg zero) c = refl
|
||||
+ZAssociative (negSucc zero) (nonneg (succ b)) (nonneg c) = +Zinherits b c
|
||||
+ZAssociative (negSucc zero) (nonneg (succ b)) (negSucc zero) = +ZCommutative (negSucc 0) (nonneg b)
|
||||
+ZAssociative (negSucc zero) (nonneg (succ b)) (negSucc (succ c)) = transitivity (+ZCommutative (negSucc 0) (nonneg b +Z negSucc c)) (transitivity (equalityCommutative (+ZAssociative (nonneg b) (negSucc c) (negSucc 0))) (applyEquality (λ i → nonneg b +Z negSucc (succ i)) (Semiring.sumZeroRight ℕSemiring c)))
|
||||
+ZAssociative (negSucc (succ a)) (nonneg (succ b)) (nonneg zero) = transitivity (applyEquality (λ i → negSucc a +Z (nonneg i)) (Semiring.sumZeroRight ℕSemiring b)) (+ZCommutative (nonneg 0) (negSucc a +Z nonneg b))
|
||||
+ZAssociative (negSucc (succ a)) (nonneg (succ b)) (nonneg (succ c)) = transitivity (applyEquality (negSucc a +Z_) (+Zinherits b (succ c))) (+ZAssociative (negSucc a) (nonneg b) (nonneg (succ c)))
|
||||
+ZAssociative (negSucc (succ a)) (nonneg (succ b)) (negSucc zero) = transitivity (+ZCommutative (negSucc (succ a)) (nonneg b)) (transitivity (transitivity (applyEquality (λ i → nonneg b +Z negSucc (succ i)) (equalityCommutative (Semiring.sumZeroRight ℕSemiring a))) (+ZAssociative (nonneg b) (negSucc a) (negSucc 0))) (applyEquality (_+Z negSucc 0) (+ZCommutative (nonneg b) (negSucc a))))
|
||||
+ZAssociative (negSucc (succ a)) (nonneg (succ b)) (negSucc (succ c)) = transitivity (+ZCommutative (negSucc (succ a)) (nonneg b +Z negSucc c)) (transitivity (transitivity (equalityCommutative (+ZAssociative (nonneg b) (negSucc c) (negSucc (succ a)))) (transitivity (applyEquality (λ i → nonneg b +Z negSucc i) (Semiring.commutative ℕSemiring (succ c) (succ a))) (+ZAssociative (nonneg b) (negSucc a) (negSucc (succ c))))) (applyEquality (_+Z negSucc (succ c)) (+ZCommutative (nonneg b) (negSucc a))))
|
||||
+ZAssociative (negSucc a) (negSucc b) (nonneg zero) = refl
|
||||
+ZAssociative (negSucc a) (negSucc zero) (nonneg (succ c)) = applyEquality (λ i → negSucc i +Z nonneg c) (Semiring.commutative ℕSemiring 0 a)
|
||||
+ZAssociative (negSucc a) (negSucc (succ b)) (nonneg (succ c)) = transitivity (+ZAssociative (negSucc a) (negSucc b) (nonneg c)) (applyEquality (λ i → negSucc i +Z nonneg c) (transitivity (applyEquality succ (Semiring.commutative ℕSemiring a b)) (Semiring.commutative ℕSemiring (succ b) a)))
|
||||
+ZAssociative (negSucc a) (negSucc b) (negSucc c) = applyEquality (λ i → negSucc (succ i)) (transitivity (Semiring.+Associative ℕSemiring a (succ b) c) (applyEquality (_+N c) (transitivity (Semiring.commutative ℕSemiring a (succ b)) (applyEquality succ (Semiring.commutative ℕSemiring b a)))))
|
||||
|
||||
additiveInverseExists : (a : ℕ) → (negSucc a +Z nonneg (succ a)) ≡ nonneg 0
|
||||
additiveInverseExists zero = refl
|
||||
additiveInverseExists (succ a) = additiveInverseExists a
|
||||
|
||||
additiveInverse : (a : ℤ) → ℤ
|
||||
additiveInverse (nonneg zero) = nonneg 0
|
||||
additiveInverse (nonneg (succ x)) = negSucc x
|
||||
additiveInverse (negSucc x) = nonneg (succ x)
|
||||
|
||||
ℤGroup : Group (reflSetoid ℤ) (_+Z_)
|
||||
Group.wellDefined ℤGroup refl refl = refl
|
||||
Group.identity ℤGroup = nonneg 0
|
||||
Group.inverse ℤGroup = additiveInverse
|
||||
Group.multAssoc ℤGroup {a} {b} {c} = +ZAssociative a b c
|
||||
Group.multIdentRight ℤGroup {nonneg zero} = refl
|
||||
Group.multIdentRight ℤGroup {nonneg (succ x)} = applyEquality (λ i → nonneg (succ i)) (Semiring.commutative ℕSemiring x 0)
|
||||
Group.multIdentRight ℤGroup {negSucc x} = refl
|
||||
Group.multIdentLeft ℤGroup = refl
|
||||
Group.invLeft ℤGroup {nonneg zero} = refl
|
||||
Group.invLeft ℤGroup {nonneg (succ x)} = additiveInverseExists x
|
||||
Group.invLeft ℤGroup {negSucc x} = transitivity (+ZCommutative (nonneg (succ x)) (negSucc x)) (additiveInverseExists x)
|
||||
Group.invRight ℤGroup {nonneg zero} = refl
|
||||
Group.invRight ℤGroup {nonneg (succ x)} = transitivity (+ZCommutative (nonneg (succ x)) (negSucc x)) (additiveInverseExists x)
|
||||
Group.invRight ℤGroup {negSucc x} = additiveInverseExists x
|
||||
|
||||
ℤAbGrp : AbelianGroup ℤGroup
|
||||
ℤAbGrp = record { commutative = λ {a} {b} → +ZCommutative a b }
|
35
Numbers/Integers/Definition.agda
Normal file
35
Numbers/Integers/Definition.agda
Normal file
@@ -0,0 +1,35 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Definition
|
||||
|
||||
module Numbers.Integers.Definition where
|
||||
|
||||
data ℤ : Set where
|
||||
nonneg : ℕ → ℤ
|
||||
negSucc : ℕ → ℤ
|
||||
|
||||
data ℤSimple : Set where
|
||||
negativeSucc : (a : ℕ) → ℤSimple
|
||||
positiveSucc : (a : ℕ) → ℤSimple
|
||||
zZero : ℤSimple
|
||||
|
||||
convertZ : ℤ → ℤSimple
|
||||
convertZ (nonneg zero) = zZero
|
||||
convertZ (nonneg (succ x)) = positiveSucc x
|
||||
convertZ (negSucc x) = negativeSucc x
|
||||
|
||||
convertZ' : ℤSimple → ℤ
|
||||
convertZ' (negativeSucc a) = negSucc a
|
||||
convertZ' (positiveSucc a) = nonneg (succ a)
|
||||
convertZ' zZero = nonneg 0
|
||||
|
||||
zIsZ : (a : ℤ) → convertZ' (convertZ a) ≡ a
|
||||
zIsZ (nonneg zero) = refl
|
||||
zIsZ (nonneg (succ x)) = refl
|
||||
zIsZ (negSucc x) = refl
|
||||
|
||||
zIsZ' : (a : ℤSimple) → convertZ (convertZ' a) ≡ a
|
||||
zIsZ' (negativeSucc a) = refl
|
||||
zIsZ' (positiveSucc a) = refl
|
||||
zIsZ' zZero = refl
|
17
Numbers/Integers/Integers.agda
Normal file
17
Numbers/Integers/Integers.agda
Normal file
@@ -0,0 +1,17 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Numbers.Integers.Definition
|
||||
open import Numbers.Integers.Addition
|
||||
open import Numbers.Integers.Multiplication
|
||||
open import Numbers.Integers.Order
|
||||
open import Groups.Definition
|
||||
|
||||
module Numbers.Integers.Integers where
|
||||
|
||||
open Numbers.Integers.Definition using (ℤ ; nonneg ; negSucc) public
|
||||
open Numbers.Integers.Addition using (_+Z_ ; ℤGroup ; ℤAbGrp) public
|
||||
open Numbers.Integers.Multiplication using (_*Z_ ; ℤIntDom ; ℤRing) public
|
||||
open Numbers.Integers.Order using (_<Z_ ; ℤOrderedRing) public
|
||||
|
||||
_-Z_ : ℤ → ℤ → ℤ
|
||||
a -Z b = a +Z (Group.inverse ℤGroup b)
|
199
Numbers/Integers/Multiplication.agda
Normal file
199
Numbers/Integers/Multiplication.agda
Normal file
@@ -0,0 +1,199 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Naturals.Multiplication
|
||||
open import Numbers.Integers.Definition
|
||||
open import Numbers.Integers.Addition
|
||||
open import Semirings.Definition
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Rings.IntegralDomains
|
||||
|
||||
module Numbers.Integers.Multiplication where
|
||||
|
||||
infix 25 _*Z_
|
||||
_*Z_ : ℤ → ℤ → ℤ
|
||||
nonneg x *Z nonneg y = nonneg (x *N y)
|
||||
nonneg zero *Z negSucc y = nonneg zero
|
||||
nonneg (succ x) *Z negSucc y = negSucc ((succ x) *N y +N x)
|
||||
negSucc x *Z nonneg zero = nonneg zero
|
||||
negSucc x *Z nonneg (succ y) = negSucc ((succ y) *N x +N y)
|
||||
negSucc x *Z negSucc y = nonneg ((succ x) *N (succ y))
|
||||
|
||||
*ZInherits : (a b : ℕ) → nonneg (a *N b) ≡ (nonneg a) *Z (nonneg b)
|
||||
*ZInherits a b = refl
|
||||
|
||||
*ZCommutative : (a b : ℤ) → a *Z b ≡ b *Z a
|
||||
*ZCommutative (nonneg x) (nonneg y) = applyEquality nonneg (multiplicationNIsCommutative x y)
|
||||
*ZCommutative (nonneg zero) (negSucc y) = refl
|
||||
*ZCommutative (nonneg (succ x)) (negSucc y) = refl
|
||||
*ZCommutative (negSucc x) (nonneg zero) = refl
|
||||
*ZCommutative (negSucc x) (nonneg (succ y)) = refl
|
||||
*ZCommutative (negSucc x) (negSucc y) = applyEquality nonneg (multiplicationNIsCommutative (succ x) (succ y))
|
||||
|
||||
*ZleftIdent : (a : ℤ) → (nonneg 1) *Z a ≡ a
|
||||
*ZleftIdent (nonneg x) = applyEquality nonneg (Semiring.commutative ℕSemiring x 0)
|
||||
*ZleftIdent (negSucc x) = applyEquality negSucc (transitivity (Semiring.commutative ℕSemiring (x +N 0) 0) (Semiring.commutative ℕSemiring x 0))
|
||||
|
||||
*ZrightIdent : (a : ℤ) → a *Z (nonneg 1) ≡ a
|
||||
*ZrightIdent (nonneg x) = applyEquality nonneg (Semiring.productOneRight ℕSemiring x)
|
||||
*ZrightIdent (negSucc x) = applyEquality negSucc (transitivity (Semiring.commutative ℕSemiring (x +N 0) 0) (Semiring.commutative ℕSemiring x 0))
|
||||
|
||||
*ZZeroLeft : (a : ℤ) → nonneg 0 *Z a ≡ nonneg 0
|
||||
*ZZeroLeft (nonneg x) = refl
|
||||
*ZZeroLeft (negSucc x) = refl
|
||||
|
||||
*ZZeroRight : (a : ℤ) → a *Z nonneg 0 ≡ nonneg 0
|
||||
*ZZeroRight (nonneg x) = applyEquality nonneg (Semiring.productZeroRight ℕSemiring x)
|
||||
*ZZeroRight (negSucc x) = refl
|
||||
|
||||
*ZAssociative : (a b c : ℤ) → a *Z (b *Z c) ≡ (a *Z b) *Z c
|
||||
*ZAssociative (nonneg zero) b c = transitivity (*ZZeroLeft (b *Z c)) (transitivity (equalityCommutative (*ZZeroLeft c)) (applyEquality (_*Z c) (equalityCommutative (*ZZeroLeft b))))
|
||||
*ZAssociative (nonneg (succ a)) (nonneg zero) c rewrite Semiring.productZeroRight ℕSemiring a | *ZZeroLeft c | Semiring.productZeroRight ℕSemiring a = refl
|
||||
*ZAssociative (nonneg (succ a)) (nonneg (succ b)) (nonneg zero) rewrite Semiring.productZeroRight ℕSemiring b | Semiring.productZeroRight ℕSemiring (b +N a *N succ b) | Semiring.productZeroRight ℕSemiring a = refl
|
||||
*ZAssociative (nonneg (succ a)) (nonneg (succ b)) (nonneg (succ c)) = applyEquality nonneg (Semiring.*Associative ℕSemiring (succ a) (succ b) (succ c))
|
||||
*ZAssociative (nonneg (succ a)) (nonneg (succ b)) (negSucc x) rewrite productDistributes' b (a *N succ b) x | Semiring.+Associative ℕSemiring x (b *N x) ((a *N succ b) *N x) | equalityCommutative (Semiring.+Associative ℕSemiring ((x +N b *N x) +N b) (a *N ((x +N b *N x) +N b)) a) | equalityCommutative (Semiring.+Associative ℕSemiring (x +N b *N x) ((a *N succ b) *N x) (b +N a *N succ b)) | equalityCommutative (Semiring.+Associative ℕSemiring (x +N b *N x) b (a *N ((x +N b *N x) +N b) +N a)) = applyEquality (λ i → negSucc ((x +N b *N x) +N i)) ans
|
||||
where
|
||||
ans3 : (succ b +N (succ b *N x)) *N a ≡ succ x *N (a *N succ b)
|
||||
ans3 rewrite multiplicationNIsCommutative (succ b) x = transitivity (equalityCommutative (Semiring.*Associative ℕSemiring (succ x) (succ b) a)) (applyEquality (succ x *N_) (multiplicationNIsCommutative (succ b) a))
|
||||
ans2 : a *N ((x +N b *N x) +N b) +N a ≡ (a *N succ b) +N (a *N succ b) *N x
|
||||
ans2 rewrite multiplicationNIsCommutative (a *N succ b) x | Semiring.commutative ℕSemiring (succ b *N x) b | multiplicationNIsCommutative a (b +N (succ b *N x)) | Semiring.commutative ℕSemiring ((b +N (succ b *N x)) *N a) a = ans3
|
||||
ans : b +N (a *N ((x +N b *N x) +N b) +N a) ≡ (a *N succ b) *N x +N (b +N a *N succ b)
|
||||
ans rewrite Semiring.commutative ℕSemiring ((a *N succ b) *N x) (b +N a *N succ b) | equalityCommutative (Semiring.+Associative ℕSemiring b (a *N succ b) ((a *N succ b) *N x)) = applyEquality (b +N_) ans2
|
||||
*ZAssociative (nonneg (succ a)) (negSucc b) (nonneg zero) = applyEquality nonneg (Semiring.productZeroRight ℕSemiring a)
|
||||
*ZAssociative (nonneg (succ a)) (negSucc b) (nonneg (succ c)) = applyEquality negSucc ans
|
||||
where
|
||||
ans4 : (b +N b *N c) +N (a *N b +N (a *N b) *N c) ≡ (b +N a *N b) +N (b *N c +N (a *N b) *N c)
|
||||
ans4 rewrite Semiring.+Associative ℕSemiring (b +N b *N c) (a *N b) ((a *N b) *N c) | Semiring.+Associative ℕSemiring (b +N a *N b) (b *N c) ((a *N b) *N c) = applyEquality (_+N (a *N b) *N c) (transitivity (transitivity (applyEquality (_+N a *N b) (Semiring.commutative ℕSemiring b (b *N c))) (equalityCommutative (Semiring.+Associative ℕSemiring (b *N c) b (a *N b)))) (Semiring.commutative ℕSemiring (b *N c) (b +N a *N b)))
|
||||
ans3 : ((b +N b *N c) +N (a *N b +N (a *N b) *N c)) +N ((c +N a *N c) +N a) ≡ ((b +N a *N b) +N (b *N c +N (a *N b) *N c)) +N ((a +N a *N c) +N c)
|
||||
ans3 rewrite Semiring.commutative ℕSemiring (c +N a *N c) a | Semiring.commutative ℕSemiring c (a *N c) | Semiring.+Associative ℕSemiring a (a *N c) c = applyEquality (λ i → i +N ((a +N a *N c) +N c)) {(b +N b *N c) +N (a *N b +N (a *N b) *N c)} {(b +N a *N b) +N (b *N c +N (a *N b) *N c)} ans4
|
||||
ans2 : (((b +N c *N b) +N (a *N b +N a *N (c *N b))) +N (c +N a *N c)) +N a ≡ (((b +N a *N b) +N (c *N b +N c *N (a *N b))) +N (a +N c *N a)) +N c
|
||||
ans2 rewrite multiplicationNIsCommutative c b | multiplicationNIsCommutative c (a *N b) | Semiring.*Associative ℕSemiring a b c | multiplicationNIsCommutative c a = transitivity (equalityCommutative (Semiring.+Associative ℕSemiring ((b +N b *N c) +N (a *N b +N (a *N b) *N c)) (c +N a *N c) a)) (transitivity ans3 (Semiring.+Associative ℕSemiring ((b +N a *N b) +N (b *N c +N (a *N b) *N c)) (a +N a *N c) c))
|
||||
ans : succ a *N ((succ c *N b) +N c) +N a ≡ succ c *N ((succ a *N b) +N a) +N c
|
||||
ans = transitivity (applyEquality (_+N a) (productDistributes (succ a) (succ c *N b) c)) (transitivity (transitivity (applyEquality (λ i → ((((b +N c *N b) +N i) +N (c +N a *N c)) +N a)) (productDistributes a b (c *N b))) (transitivity ans2 (applyEquality (λ i → (((b +N a *N b) +N i) +N (a +N c *N a)) +N c) (equalityCommutative (productDistributes c b (a *N b)))))) (applyEquality (_+N c) (equalityCommutative (productDistributes (succ c) (succ a *N b) a))))
|
||||
*ZAssociative (nonneg (succ a)) (negSucc b) (negSucc x) = applyEquality nonneg ans
|
||||
where
|
||||
ans : succ a *N (succ b *N succ x) ≡ succ ((succ a *N b) +N a) *N succ x
|
||||
ans rewrite Semiring.commutative ℕSemiring (succ a *N b) a | multiplicationNIsCommutative (succ a) b = transitivity (Semiring.*Associative ℕSemiring (succ a) (succ b) (succ x)) (applyEquality (_*N succ x) {succ a *N succ b} {succ b *N succ a} (multiplicationNIsCommutative (succ a) (succ b)))
|
||||
*ZAssociative (negSucc a) (nonneg zero) c rewrite *ZZeroLeft c = refl
|
||||
*ZAssociative (negSucc a) (nonneg (succ b)) (nonneg zero) rewrite multiplicationNIsCommutative b 0 = refl
|
||||
*ZAssociative (negSucc a) (nonneg (succ b)) (nonneg (succ c)) = applyEquality negSucc ans
|
||||
where
|
||||
ans2 : (d : ℕ) → ((succ b *N d) *N a) +N b *N d ≡ d *N ((succ b *N a) +N b)
|
||||
ans2 d = transitivity (transitivity (applyEquality (((d +N b *N d) *N a) +N_) (multiplicationNIsCommutative b d)) (applyEquality (_+N d *N b) (transitivity (applyEquality (_*N a) {d +N b *N d} {d *N succ b} (multiplicationNIsCommutative (succ b) d)) (equalityCommutative (Semiring.*Associative ℕSemiring d (succ b) a))))) (equalityCommutative (Semiring.+DistributesOver* ℕSemiring d (succ b *N a) b))
|
||||
ans : (succ (c +N b *N succ c) *N a) +N (c +N b *N succ c) ≡ (succ c *N ((succ b *N a) +N b)) +N c
|
||||
ans rewrite Semiring.commutative ℕSemiring c (b *N succ c) | Semiring.+Associative ℕSemiring (succ (b *N succ c +N c) *N a) (b *N succ c) c | Semiring.commutative ℕSemiring (b *N succ c) c = applyEquality (_+N c) (ans2 (succ c))
|
||||
*ZAssociative (negSucc a) (nonneg (succ b)) (negSucc c) = applyEquality nonneg ans
|
||||
where
|
||||
ans : succ a *N ((succ ((succ b) *N c)) +N b) ≡ (succ (succ b *N a) +N b) *N succ c
|
||||
ans rewrite Semiring.commutative ℕSemiring (succ b *N a) b | multiplicationNIsCommutative (succ b) a = transitivity (applyEquality (succ a *N_) (applyEquality succ (transitivity (equalityCommutative (Semiring.+Associative ℕSemiring c (b *N c) b)) (applyEquality (c +N_) (transitivity (transitivity (Semiring.commutative ℕSemiring (b *N c) b) (applyEquality (b +N_) (multiplicationNIsCommutative b c))) (multiplicationNIsCommutative (succ c) b)))))) (Semiring.*Associative ℕSemiring (succ a) (succ b) (succ c))
|
||||
*ZAssociative (negSucc a) (negSucc b) (nonneg zero) rewrite Semiring.productZeroRight ℕSemiring (b +N a *N succ b) = refl
|
||||
*ZAssociative (negSucc a) (negSucc b) (nonneg (succ x)) = applyEquality nonneg (transitivity (applyEquality (succ a *N_) ans) (Semiring.*Associative ℕSemiring (succ a) (succ b) (succ x)))
|
||||
where
|
||||
ans : succ ((succ x *N b) +N x) ≡ (succ b *N succ x)
|
||||
ans rewrite Semiring.commutative ℕSemiring (succ x *N b) x | multiplicationNIsCommutative (succ x) b = refl
|
||||
*ZAssociative (negSucc a) (negSucc b) (negSucc x) = applyEquality negSucc t
|
||||
where
|
||||
l1 : (a b c d e f g : ℕ) → (a +N b) +N d ≡ e +N (f +N g) → (a +N b) +N (c +N d) ≡ (c +N e) +N (f +N g)
|
||||
l1 a b c d e f g pr rewrite Semiring.commutative ℕSemiring c d | Semiring.+Associative ℕSemiring (a +N b) d c | pr | Semiring.commutative ℕSemiring (e +N (f +N g)) c | Semiring.+Associative ℕSemiring c e (f +N g) = refl
|
||||
v : (((succ b) *N succ x) *N a) ≡ (succ x) *N (a *N succ b)
|
||||
v = transitivity (applyEquality (_*N a) (multiplicationNIsCommutative (succ b) (succ x))) (transitivity (equalityCommutative (Semiring.*Associative ℕSemiring (succ x) (succ b) a)) (applyEquality (succ x *N_) (multiplicationNIsCommutative (succ b) a)))
|
||||
u : (a +N (x +N b *N succ x) *N a) +N b *N succ x ≡ (b +N a *N succ b) *N x +N (b +N a *N succ b)
|
||||
u = transitivity (transitivity (transitivity (Semiring.commutative ℕSemiring (a +N (x +N b *N succ x) *N a) (b *N succ x)) (transitivity (applyEquality (_+N (a +N (x +N b *N succ x) *N a)) (multiplicationNIsCommutative b (succ x))) (transitivity (applyEquality ((b +N x *N b) +N_) v) (equalityCommutative (productDistributes (succ x) b (a *N succ b)))))) (applyEquality ((b +N a *N succ b) +N_) (multiplicationNIsCommutative x (b +N a *N succ b)))) (Semiring.commutative ℕSemiring (b +N a *N succ b) ((b +N a *N succ b) *N x))
|
||||
t : (a +N (x +N b *N succ x) *N a) +N (x +N b *N succ x) ≡ (x +N (b +N a *N succ b) *N x) +N (b +N a *N succ b)
|
||||
t = l1 a ((x +N b *N succ x) *N a) x (b *N succ x) ((b +N (a *N succ b)) *N x) b (a *N succ b) u
|
||||
|
||||
multLemma : (a b : ℕ) → a *N succ b ≡ a *N b +N a
|
||||
multLemma a b rewrite multiplicationNIsCommutative a (succ b) | Semiring.commutative ℕSemiring a (b *N a) | multiplicationNIsCommutative a b = refl
|
||||
|
||||
negSucc+Nonneg : (a b : ℕ) → negSucc a ≡ negSucc (a +N b) +Z nonneg b
|
||||
negSucc+Nonneg zero zero = refl
|
||||
negSucc+Nonneg zero (succ b) = negSucc+Nonneg zero b
|
||||
negSucc+Nonneg (succ a) zero rewrite Semiring.commutative ℕSemiring a 0 = refl
|
||||
negSucc+Nonneg (succ a) (succ b) rewrite Semiring.commutative ℕSemiring a (succ b) | Semiring.commutative ℕSemiring b a = negSucc+Nonneg (succ a) b
|
||||
|
||||
distLemma : (a b : ℕ) → negSucc a *Z nonneg b ≡ negSucc ((a +N b *N a) +N b) +Z nonneg (succ a)
|
||||
distLemma a b rewrite Semiring.commutative ℕSemiring (a +N b *N a) b | Semiring.commutative ℕSemiring a (b *N a) | Semiring.+Associative ℕSemiring b (b *N a) a = transitivity u (transitivity (equalityCommutative (+ZAssociative (negSucc ((b +N b *N a) +N a)) (nonneg a) (nonneg 1))) (applyEquality ((negSucc ((b +N b *N a) +N a)) +Z_) {nonneg a +Z nonneg 1} {nonneg (succ a)} (+ZCommutative (nonneg a) (nonneg 1))))
|
||||
where
|
||||
v : (a b : ℕ) → negSucc a *Z nonneg b ≡ (negSucc (b +N b *N a)) +Z nonneg 1
|
||||
v zero zero = refl
|
||||
v zero (succ b) rewrite Semiring.productZeroRight ℕSemiring b = applyEquality negSucc (equalityCommutative (Semiring.sumZeroRight ℕSemiring b))
|
||||
v (succ a) zero = refl
|
||||
v (succ a) (succ b) = applyEquality negSucc (Semiring.commutative ℕSemiring (succ (a +N b *N succ a)) b)
|
||||
u : negSucc a *Z nonneg b ≡ (negSucc ((b +N b *N a) +N a) +Z nonneg a) +Z nonneg 1
|
||||
u = transitivity (v a b) (applyEquality (_+Z nonneg 1) (negSucc+Nonneg (b +N b *N a) a))
|
||||
|
||||
*ZDistributesOver+Z : (a b c : ℤ) → a *Z (b +Z c) ≡ (a *Z b) +Z (a *Z c)
|
||||
*ZDistributesOver+Z (nonneg zero) b c rewrite *ZZeroLeft (b +Z c) | *ZZeroLeft b | *ZZeroLeft c = refl
|
||||
*ZDistributesOver+Z (nonneg (succ a)) (nonneg zero) (nonneg c) rewrite Semiring.productZeroRight ℕSemiring a = refl
|
||||
*ZDistributesOver+Z (nonneg (succ a)) (nonneg (succ b)) (nonneg c) = applyEquality nonneg (productDistributes (succ a) (succ b) c)
|
||||
*ZDistributesOver+Z (nonneg (succ a)) (nonneg zero) (negSucc c) rewrite Semiring.productZeroRight ℕSemiring a = refl
|
||||
*ZDistributesOver+Z (nonneg (succ a)) (nonneg (succ b)) (negSucc zero) rewrite Semiring.productZeroRight ℕSemiring a = t a
|
||||
where
|
||||
t : (a : ℕ) → nonneg (succ a *N b) ≡ nonneg (succ a *N succ b) +Z negSucc a
|
||||
t zero = refl
|
||||
t (succ a) = transitivity (+Zinherits b (b +N a *N b)) (transitivity (applyEquality (nonneg b +Z_) (t a)) (transitivity (+ZAssociative (nonneg b) (nonneg (succ (b +N a *N succ b))) (negSucc a)) (applyEquality (_+Z negSucc a) {nonneg b +Z nonneg (succ (b +N a *N succ b))} {nonneg (b +N succ (b +N a *N succ b))} (equalityCommutative (+Zinherits b _)))))
|
||||
*ZDistributesOver+Z (nonneg (succ a)) (nonneg (succ x)) (negSucc (succ b)) = transitivity (applyEquality ((nonneg (succ a)) *Z_) (+ZCommutative (nonneg x) (negSucc b))) (transitivity (transitivity (*ZDistributesOver+Z (nonneg (succ a)) (negSucc b) (nonneg x)) t) (+ZCommutative (negSucc ((b +N a *N succ b) +N a)) (nonneg (x +N a *N succ x))))
|
||||
where
|
||||
t : negSucc ((b +N a *N b) +N a) +Z nonneg (x +N a *N x) ≡ negSucc ((b +N a *N succ b) +N a) +Z nonneg (x +N a *N succ x)
|
||||
t rewrite multLemma a x | multLemma a b | Semiring.+Associative ℕSemiring x (a *N x) a | Semiring.+Associative ℕSemiring b (a *N b) a | Semiring.commutative ℕSemiring (x +N a *N x) a | +Zinherits a (x +N a *N x) | +ZAssociative (negSucc (((b +N a *N b) +N a) +N a)) (nonneg a) (nonneg (x +N a *N x)) = applyEquality (_+Z nonneg (x +N a *N x)) {negSucc ((b +N a *N b) +N a)} {negSucc (((b +N a *N b) +N a) +N a) +Z nonneg a} (negSucc+Nonneg ((b +N a *N b) +N a) a)
|
||||
*ZDistributesOver+Z (nonneg (succ a)) (negSucc b) (nonneg zero) rewrite Semiring.productZeroRight ℕSemiring a = refl
|
||||
*ZDistributesOver+Z (nonneg (succ a)) (negSucc zero) (nonneg (succ x)) rewrite Semiring.productZeroRight ℕSemiring a = t a x
|
||||
where
|
||||
t : (a x : ℕ) → nonneg (succ a *N x) ≡ negSucc a +Z nonneg (succ a *N succ x)
|
||||
t zero x = refl
|
||||
t (succ a) x with t a x
|
||||
... | bl rewrite +Zinherits x (a *N x) | +Zinherits x (x +N a *N x) = transitivity (transitivity (applyEquality (nonneg x +Z_) (+Zinherits x (a *N x))) (+ZCommutative (nonneg x) (nonneg x +Z nonneg (a *N x)))) (transitivity (applyEquality (_+Z nonneg x) bl) (transitivity (equalityCommutative (+ZAssociative (negSucc a) (nonneg (succ x +N a *N succ x)) (nonneg x))) (applyEquality (λ i → negSucc a +Z nonneg i) {succ (x +N a *N succ x) +N x} {x +N succ (x +N a *N succ x)} (Semiring.commutative ℕSemiring (succ (x +N a *N succ x)) x))))
|
||||
*ZDistributesOver+Z (nonneg (succ a)) (negSucc (succ b)) (nonneg (succ x)) = transitivity (*ZDistributesOver+Z (nonneg (succ a)) (negSucc b) (nonneg x)) t
|
||||
where
|
||||
t : negSucc ((b +N a *N b) +N a) +Z nonneg (x +N a *N x) ≡ negSucc ((b +N a *N succ b) +N a) +Z nonneg (x +N a *N succ x)
|
||||
t rewrite multLemma a x | multLemma a b | Semiring.+Associative ℕSemiring x (a *N x) a | Semiring.+Associative ℕSemiring b (a *N b) a | Semiring.commutative ℕSemiring (x +N a *N x) a | +Zinherits a (x +N a *N x) | +ZAssociative (negSucc (((b +N a *N b) +N a) +N a)) (nonneg a) (nonneg (x +N a *N x)) = applyEquality (_+Z nonneg (x +N a *N x)) {negSucc ((b +N a *N b) +N a)} {negSucc (((b +N a *N b) +N a) +N a) +Z nonneg a} (negSucc+Nonneg ((b +N a *N b) +N a) a)
|
||||
*ZDistributesOver+Z (nonneg (succ a)) (negSucc b) (negSucc x) = applyEquality negSucc (transitivity (applyEquality (_+N a) (transitivity (productDistributes (succ a) (succ b) x) (applyEquality (_+N (x +N a *N x)) {succ b +N a *N succ b} {succ (b +N a *N b) +N a} (applyEquality succ u)))) (equalityCommutative (Semiring.+Associative ℕSemiring (succ ((b +N a *N b) +N a)) (succ a *N x) a)))
|
||||
where
|
||||
u : b +N a *N succ b ≡ (succ a *N b) +N a
|
||||
u rewrite multiplicationNIsCommutative a (succ b) | Semiring.commutative ℕSemiring a (b *N a) | multiplicationNIsCommutative b a = Semiring.+Associative ℕSemiring b (a *N b) a
|
||||
*ZDistributesOver+Z (negSucc a) (nonneg zero) (nonneg x) = refl
|
||||
*ZDistributesOver+Z (negSucc a) (nonneg zero) (negSucc x) = refl
|
||||
*ZDistributesOver+Z (negSucc a) (nonneg (succ b)) (nonneg zero) rewrite Semiring.commutative ℕSemiring b 0 = refl
|
||||
*ZDistributesOver+Z (negSucc a) (nonneg (succ b)) (nonneg (succ c)) = applyEquality negSucc t
|
||||
where
|
||||
v : a *N b +N a *N succ c ≡ c *N a +N (a +N b *N a)
|
||||
v rewrite multiplicationNIsCommutative b a | Semiring.+Associative ℕSemiring (c *N a) a (a *N b) | Semiring.commutative ℕSemiring (c *N a +N a) (a *N b) | Semiring.commutative ℕSemiring (c *N a) a = applyEquality (λ i → a *N b +N i) (transitivity (multLemma a c) (transitivity (Semiring.commutative ℕSemiring (a *N c) a) (applyEquality (a +N_) (multiplicationNIsCommutative a c))))
|
||||
u : (a +N (b +N succ c) *N a) +N b ≡ (a +N c *N a) +N ((a +N b *N a) +N b)
|
||||
u rewrite Semiring.+Associative ℕSemiring (a +N c *N a) (a +N b *N a) b | (equalityCommutative (Semiring.+Associative ℕSemiring a (c *N a) (a +N b *N a))) = applyEquality (λ i → (a +N i) +N b) (transitivity (multiplicationNIsCommutative (b +N succ c) a) (transitivity (productDistributes a b (succ c)) v))
|
||||
t : (a +N (b +N succ c) *N a) +N (b +N succ c) ≡ succ (((a +N b *N a) +N b) +N ((a +N c *N a) +N c))
|
||||
t rewrite Semiring.+Associative ℕSemiring ((a +N (b +N succ c) *N a)) b (succ c) | Semiring.commutative ℕSemiring ((a +N b *N a) +N b) ((a +N c *N a) +N c) | Semiring.commutative ℕSemiring (a +N c *N a) c | equalityCommutative (Semiring.+Associative ℕSemiring (succ c) (a +N c *N a) ((a +N b *N a) +N b)) | Semiring.commutative ℕSemiring (succ c) ((a +N c *N a) +N ((a +N b *N a) +N b)) = applyEquality (_+N succ c) u
|
||||
*ZDistributesOver+Z (negSucc a) (nonneg (succ b)) (negSucc zero) rewrite Semiring.productOneRight ℕSemiring a = distLemma a b
|
||||
*ZDistributesOver+Z (negSucc a) (nonneg (succ b)) (negSucc (succ c)) = transitivity (*ZDistributesOver+Z (negSucc a) (nonneg b) (negSucc c)) (transitivity (applyEquality (_+Z nonneg (succ (c +N a *N succ c))) (distLemma a b)) (transitivity (equalityCommutative (+ZAssociative (negSucc ((a +N b *N a) +N b)) (nonneg (succ a)) (nonneg (succ (c +N a *N succ c))))) (applyEquality (negSucc ((a +N b *N a) +N b) +Z_) {nonneg (succ (a +N succ (c +N a *N succ c)))} {nonneg (succ (succ (c +N a *N succ (succ c))))} (applyEquality nonneg (t (succ a) (succ c))))))
|
||||
where
|
||||
t : (x y : ℕ) → x +N x *N y ≡ x *N succ y
|
||||
t x y = transitivity (Semiring.commutative ℕSemiring x (x *N y)) (equalityCommutative (multLemma x y))
|
||||
*ZDistributesOver+Z (negSucc a) (negSucc b) (nonneg zero) rewrite Semiring.commutative ℕSemiring (b +N a *N succ b) 0 = refl
|
||||
*ZDistributesOver+Z (negSucc a) (negSucc zero) (nonneg (succ b)) = transitivity (distLemma a b) (transitivity (+ZCommutative (negSucc ((a +N b *N a) +N b)) (nonneg (succ a))) (applyEquality (λ i → nonneg (succ i) +Z negSucc ((a +N b *N a) +N b)) (equalityCommutative (Semiring.productOneRight ℕSemiring a))))
|
||||
*ZDistributesOver+Z (negSucc a) (negSucc (succ c)) (nonneg (succ b)) = transitivity (applyEquality (negSucc a *Z_) (+ZCommutative (negSucc c) (nonneg b))) (transitivity (transitivity (*ZDistributesOver+Z (negSucc a) (nonneg b) (negSucc c)) (transitivity (applyEquality (_+Z nonneg (succ (c +N a *N succ c))) (distLemma a b)) (transitivity (equalityCommutative (+ZAssociative (negSucc ((a +N b *N a) +N b)) (nonneg (succ a)) (nonneg (succ (c +N a *N succ c))))) (applyEquality (negSucc ((a +N b *N a) +N b) +Z_) {nonneg (succ (a +N succ (c +N a *N succ c)))} {nonneg (succ (succ (c +N a *N succ (succ c))))} (applyEquality nonneg (t (succ a) (succ c))))))) (+ZCommutative (negSucc ((a +N b *N a) +N b)) (nonneg (succ (succ (c +N a *N succ (succ c)))))))
|
||||
where
|
||||
t : (x y : ℕ) → x +N x *N y ≡ x *N succ y
|
||||
t x y = transitivity (Semiring.commutative ℕSemiring x (x *N y)) (equalityCommutative (multLemma x y))
|
||||
*ZDistributesOver+Z (negSucc a) (negSucc b) (negSucc c) = applyEquality nonneg (transitivity (applyEquality (λ i → succ a *N (succ i)) (transitivity (applyEquality succ (Semiring.commutative ℕSemiring b c)) (Semiring.commutative ℕSemiring (succ c) b))) (productDistributes (succ a) (succ b) (succ c)))
|
||||
|
||||
ℤRing : Ring (reflSetoid ℤ) _+Z_ _*Z_
|
||||
Ring.additiveGroup ℤRing = ℤGroup
|
||||
Ring.multWellDefined ℤRing refl refl = refl
|
||||
Ring.1R ℤRing = nonneg 1
|
||||
Ring.groupIsAbelian ℤRing {a} {b} = +ZCommutative a b
|
||||
Ring.multAssoc ℤRing {a} {b} {c} = *ZAssociative a b c
|
||||
Ring.multCommutative ℤRing {a} {b} = *ZCommutative a b
|
||||
Ring.multDistributes ℤRing {a} {b} {c} = *ZDistributesOver+Z a b c
|
||||
Ring.multIdentIsIdent ℤRing {a} = *ZleftIdent a
|
||||
|
||||
intDom : (a b : ℤ) → a *Z b ≡ nonneg 0 → (a ≡ nonneg 0) || (b ≡ nonneg 0)
|
||||
intDom (nonneg zero) (nonneg b) pr = inl refl
|
||||
intDom (nonneg (succ a)) (nonneg zero) pr = inr refl
|
||||
intDom (nonneg zero) (negSucc b) pr = inl refl
|
||||
intDom (negSucc a) (nonneg zero) pr = inr refl
|
||||
|
||||
ℤIntDom : IntegralDomain ℤRing
|
||||
IntegralDomain.intDom ℤIntDom {a} {b} = intDom a b
|
||||
IntegralDomain.nontrivial ℤIntDom ()
|
99
Numbers/Integers/Order.agda
Normal file
99
Numbers/Integers/Order.agda
Normal file
@@ -0,0 +1,99 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers.Definition
|
||||
open import Numbers.Integers.Addition
|
||||
open import Numbers.Integers.Multiplication
|
||||
open import Semirings.Definition
|
||||
open import Rings.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Orders
|
||||
|
||||
module Numbers.Integers.Order where
|
||||
|
||||
infix 5 _<Z_
|
||||
record _<Z_ (a : ℤ) (b : ℤ) : Set where
|
||||
constructor le
|
||||
field
|
||||
x : ℕ
|
||||
proof : (nonneg (succ x)) +Z a ≡ b
|
||||
|
||||
lessLemma : (a x : ℕ) → succ x +N a ≡ a → False
|
||||
lessLemma zero x pr = naughtE (equalityCommutative pr)
|
||||
lessLemma (succ a) x pr = lessLemma a x q
|
||||
where
|
||||
q : succ x +N a ≡ a
|
||||
q rewrite Semiring.commutative ℕSemiring a (succ x) | Semiring.commutative ℕSemiring x a | Semiring.commutative ℕSemiring (succ a) x = succInjective pr
|
||||
|
||||
nonnegInjective : {a b : ℕ} → (nonneg a ≡ nonneg b) → (a ≡ b)
|
||||
nonnegInjective {a} {.a} refl = refl
|
||||
|
||||
irreflexive : (x : ℤ) → x <Z x → False
|
||||
irreflexive (nonneg x) record { x = y ; proof = proof } = lessLemma x y (nonnegInjective proof)
|
||||
irreflexive (negSucc a) record { x = x ; proof = proof } = naughtE (equalityCommutative q)
|
||||
where
|
||||
pr' : nonneg (succ x) +Z (negSucc a +Z nonneg (succ a)) ≡ negSucc a +Z nonneg (succ a)
|
||||
pr' rewrite +ZAssociative (nonneg (succ x)) (negSucc a) (nonneg (succ a)) = applyEquality (λ t → t +Z nonneg (succ a)) proof
|
||||
pr'' : nonneg (succ x) +Z nonneg 0 ≡ nonneg 0
|
||||
pr'' rewrite equalityCommutative (additiveInverseExists a) = identityOfIndiscernablesLeft (nonneg (succ x) +Z (negSucc a +Z nonneg (succ a))) (negSucc a +Z nonneg (succ a)) (nonneg (succ (x +N zero))) _≡_ pr' q
|
||||
where
|
||||
q : nonneg (succ x) +Z (negSucc a +Z nonneg (succ a)) ≡ nonneg (succ (x +N 0))
|
||||
q rewrite Semiring.commutative ℕSemiring x 0 | additiveInverseExists a | Semiring.commutative ℕSemiring x 0 = refl
|
||||
pr''' : succ x +N 0 ≡ 0
|
||||
pr''' = nonnegInjective pr''
|
||||
q : succ x ≡ 0
|
||||
q rewrite Semiring.commutative ℕSemiring 0 (succ x) = pr'''
|
||||
|
||||
lessZTransitive : {a b c : ℤ} → (a <Z b) → (b <Z c) → (a <Z c)
|
||||
lessZTransitive {a} {b} {c} (le d1 pr1) (le d2 pr2) rewrite equalityCommutative pr1 = le (d1 +N succ d2) pr
|
||||
where
|
||||
pr : nonneg (succ (d1 +N succ d2)) +Z a ≡ c
|
||||
pr rewrite +ZAssociative (nonneg (succ d2)) (nonneg (succ d1)) a | Semiring.commutative ℕSemiring (succ d2) (succ d1) = pr2
|
||||
|
||||
lessInherits : {a b : ℕ} → (a <N b) → ((nonneg a) <Z (nonneg b))
|
||||
_<Z_.x (lessInherits {a} {b} (le x proof)) = x
|
||||
_<Z_.proof (lessInherits {a} {.(succ (x +N a))} (le x refl)) = refl
|
||||
|
||||
lessInheritsNegsucc : {a b : ℕ} → (a <N b) → ((negSucc b) <Z negSucc a)
|
||||
_<Z_.x (lessInheritsNegsucc {a} {b} (le x proof)) = x
|
||||
_<Z_.proof (lessInheritsNegsucc {a} {b} (le x proof)) rewrite equalityCommutative proof = transitivity (transitivity (+ZCommutative (nonneg x) (negSucc (x +N a))) (applyEquality (λ i → negSucc i +Z nonneg x) (Semiring.commutative ℕSemiring x a))) (equalityCommutative (negSucc+Nonneg a x))
|
||||
|
||||
lessNegsuccNonneg : {a b : ℕ} → (negSucc a <Z nonneg b)
|
||||
_<Z_.x (lessNegsuccNonneg {a} {b}) = a +N b
|
||||
_<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} | 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
|
||||
... | 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
|
||||
|
||||
ℤOrder : TotalOrder ℤ
|
||||
PartialOrder._<_ (TotalOrder.order ℤOrder) = _<Z_
|
||||
PartialOrder.irreflexive (TotalOrder.order ℤOrder) {a} = irreflexive a
|
||||
PartialOrder.transitive (TotalOrder.order ℤOrder) = lessZTransitive
|
||||
TotalOrder.totality ℤOrder a b = lessThanTotalZ {a} {b}
|
||||
|
||||
orderRespectsAddition : (a b : ℤ) → a <Z b → (c : ℤ) → a +Z c <Z b +Z c
|
||||
orderRespectsAddition (nonneg a) (nonneg b) (le x proof) (nonneg c) = le x (transitivity (+ZAssociative (nonneg (succ x)) (nonneg a) (nonneg c)) (applyEquality (_+Z nonneg c) proof))
|
||||
orderRespectsAddition (nonneg a) (nonneg b) (le x proof) (negSucc c) = le x (transitivity (+ZAssociative (nonneg (succ x)) (nonneg a) (negSucc c)) (applyEquality (_+Z negSucc c) proof))
|
||||
orderRespectsAddition (negSucc a) (nonneg b) (le x proof) (nonneg c) = le x (transitivity (+ZAssociative (nonneg (succ x)) (negSucc a) (nonneg c)) (applyEquality (_+Z nonneg c) proof))
|
||||
orderRespectsAddition (negSucc a) (nonneg b) (le x proof) (negSucc c) = le x (transitivity (+ZAssociative (nonneg (succ x)) (negSucc a) (negSucc c)) (applyEquality (_+Z negSucc c) proof))
|
||||
orderRespectsAddition (negSucc a) (negSucc b) (le x proof) (nonneg c) = le x (transitivity (+ZAssociative (nonneg (succ x)) (negSucc a) (nonneg c)) (applyEquality (_+Z nonneg c) proof))
|
||||
orderRespectsAddition (negSucc a) (negSucc b) (le x proof) (negSucc c) = le x (transitivity (+ZAssociative (nonneg (succ x)) (negSucc a) (negSucc c)) (applyEquality (_+Z negSucc c) proof))
|
||||
|
||||
orderRespectsMultiplication : (a b : ℤ) → nonneg 0 <Z a → nonneg 0 <Z b → nonneg 0 <Z a *Z b
|
||||
orderRespectsMultiplication (nonneg (succ a)) (nonneg (succ b)) 0<a 0<b = lessInherits (succIsPositive (b +N a *N succ b))
|
||||
|
||||
ℤOrderedRing : OrderedRing ℤRing (totalOrderToSetoidTotalOrder ℤOrder)
|
||||
OrderedRing.orderRespectsAddition ℤOrderedRing {a} {b} = orderRespectsAddition a b
|
||||
OrderedRing.orderRespectsMultiplication ℤOrderedRing {a} {b} = orderRespectsMultiplication a b
|
@@ -2,7 +2,7 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers
|
||||
open import Numbers.Integers.Integers
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
|
@@ -2,7 +2,7 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Rationals
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
|
@@ -4,7 +4,7 @@ open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Functions
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers
|
||||
open import Numbers.Integers.Integers
|
||||
open import IntegersModN
|
||||
open import Rings.Examples.Proofs
|
||||
open import PrimeNumbers
|
||||
|
@@ -7,7 +7,7 @@ open import Groups.Definition
|
||||
open import Orders
|
||||
open import Rings.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Numbers.Integers
|
||||
open import Numbers.Integers.Integers
|
||||
open import PrimeNumbers
|
||||
open import IntegersModN
|
||||
|
||||
|
Reference in New Issue
Block a user