Z is a Euclidean domain (#86)

This commit is contained in:
Patrick Stevens
2019-12-07 13:00:18 +00:00
committed by GitHub
parent cfd9787bb8
commit e192f0e1f1
38 changed files with 1018 additions and 486 deletions

View File

@@ -0,0 +1,107 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Multiplication
open import Numbers.Integers.Order
open import Numbers.Integers.RingStructure.Ring
open import Numbers.Integers.RingStructure.IntegralDomain
open import Semirings.Definition
open import Groups.Definition
open import Rings.Definition
open import Setoids.Setoids
open import Rings.EuclideanDomains.Definition
open import Orders
open import Numbers.Naturals.EuclideanAlgorithm
module Numbers.Integers.RingStructure.EuclideanDomain where
private
norm :
norm (nonneg x) = x
norm (negSucc x) = succ x
norm' : {a : } (a!=0 : (a nonneg 0) False)
norm' {a} _ = norm a
multiplyIncreasesSuccLemma : (a b : ) succ (succ a) <N (succ (succ a)) *N (succ (succ b))
multiplyIncreasesSuccLemma a b with multiplyIncreases (succ (succ b)) (succ (succ a)) (le b (applyEquality succ (Semiring.commutative Semiring b 1))) (le (succ a) (applyEquality (λ i succ (succ i)) (Semiring.sumZeroRight Semiring a)))
... | bl rewrite multiplicationNIsCommutative (succ (succ b)) (succ (succ a)) = bl
multiplyIncreasesSucc : (a b : ) succ a ≤N (succ a) *N (succ b)
multiplyIncreasesSucc zero zero = inr refl
multiplyIncreasesSucc zero (succ b) = inl (le (b +N 0) (Semiring.commutative Semiring (succ (b +N 0)) 1))
multiplyIncreasesSucc (succ a) zero = inr (applyEquality (λ i succ (succ i)) (equalityCommutative (productWithOneRight a)))
multiplyIncreasesSucc (succ a) (succ b) = inl (multiplyIncreasesSuccLemma a b)
multiplyIncreasesSuccLemma' : (a b : ) succ (succ a) <N succ ((succ (succ a)) *N succ b) +N succ a
multiplyIncreasesSuccLemma' a b = succPreservesInequality (le (b +N succ (b +N a *N succ b)) refl)
multiplyIncreasesSucc' : (a b : ) succ a ≤N succ ((b +N a *N b) +N a)
multiplyIncreasesSucc' zero zero = inr refl
multiplyIncreasesSucc' zero (succ b) = inl (le b (applyEquality succ (transitivity (Semiring.commutative Semiring b 1) (transitivity (equalityCommutative (Semiring.sumZeroRight Semiring (succ b))) (equalityCommutative (Semiring.sumZeroRight Semiring _))))))
multiplyIncreasesSucc' (succ a) zero rewrite multiplicationNIsCommutative a zero = inr refl
multiplyIncreasesSucc' (succ a) (succ b) = inl (multiplyIncreasesSuccLemma' a b)
normSize : {a b : } (a!=0 : (a nonneg 0) False) (b!=0 : (b nonneg 0) False) (c : ) b (a *Z c) (norm a) ≤N (norm b)
normSize {nonneg a} {nonneg b} a!=0 b!=0 (nonneg zero) b=ac rewrite nonnegInjective b=ac | multiplicationNIsCommutative a 0 = exFalso (b!=0 refl)
normSize {nonneg a} {nonneg b} a!=0 b!=0 (nonneg (succ 0)) b=ac rewrite nonnegInjective b=ac | multiplicationNIsCommutative a 1 = inr (equalityCommutative (Semiring.sumZeroRight Semiring a))
normSize {nonneg 0} {nonneg b} a!=0 b!=0 (nonneg (succ (succ c))) b=ac = exFalso (a!=0 refl)
normSize {nonneg (succ a)} {nonneg (succ b)} a!=0 b!=0 (nonneg (succ (succ c))) b=ac rewrite nonnegInjective b=ac = multiplyIncreasesSucc a (succ c)
normSize {nonneg 0} {nonneg b} a!=0 b!=0 (negSucc c) bad = exFalso (a!=0 refl)
normSize {nonneg (succ a)} {nonneg b} a!=0 b!=0 (negSucc c) ()
normSize {nonneg a} {negSucc b} a!=0 b!=0 (nonneg c) ()
normSize {nonneg (succ a)} {negSucc b} a!=0 b!=0 (negSucc c) pr rewrite negSuccInjective pr = multiplyIncreasesSucc' a c
normSize {negSucc a} {nonneg zero} _ b!=0 c pr = exFalso (b!=0 refl)
normSize {negSucc a} {nonneg (succ b)} _ _ (nonneg zero) ()
normSize {negSucc a} {nonneg (succ b)} _ _ (nonneg (succ x)) ()
normSize {negSucc a} {nonneg (succ b)} _ _ (negSucc c) pr rewrite nonnegInjective pr = multiplyIncreasesSucc a c
normSize {negSucc a} {negSucc b} _ _ (nonneg (succ c)) pr rewrite negSuccInjective pr | multiplicationNIsCommutative c a | Semiring.commutative Semiring (a +N a *N c) c | Semiring.+Associative Semiring c a (a *N c) | Semiring.commutative Semiring c a | equalityCommutative (Semiring.+Associative Semiring a c (a *N c)) | Semiring.commutative Semiring a (c +N a *N c) = multiplyIncreasesSucc' a c
divAlg : {a b : } (a!=0 : (a nonneg 0) False) (b!=0 : (b nonneg 0) False) DivisionAlgorithmResult Ring norm' a!=0 b!=0
divAlg {nonneg zero} {b} a!=0 b!=0 = exFalso (a!=0 refl)
divAlg {nonneg (succ a)} {nonneg zero} a!=0 b!=0 = exFalso (b!=0 refl)
divAlg {nonneg (succ a)} {nonneg (succ b)} a!=0 b!=0 with divisionAlg (succ b) (succ a)
divAlg {nonneg (succ a)} {nonneg (succ b)} a!=0 b!=0 | record { quot = quot ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = record { quotient = nonneg quot ; rem = nonneg rem ; remSmall = u ; divAlg = transitivity (applyEquality nonneg (equalityCommutative pr)) t }
where
t : nonneg ((quot +N b *N quot) +N rem) nonneg (quot *N succ b) +Z nonneg rem
t rewrite multiplicationNIsCommutative (succ b) quot = +Zinherits (quot *N succ b) rem
u : (nonneg rem nonneg 0) || Sg ((nonneg rem nonneg 0) False) (λ pr rem <N succ b)
u with TotalOrder.totality Order (nonneg 0) (nonneg rem)
u | inl (inl 0<rem) = inr ((λ p TotalOrder.irreflexive Order (TotalOrder.<WellDefined Order refl p 0<rem)) , remIsSmall)
u | inr 0=rem rewrite 0=rem = inl refl
divAlg {nonneg (succ a)} {negSucc b} _ _ with divisionAlg (succ b) (succ a)
divAlg {nonneg (succ a)} {negSucc b} _ _ | record { quot = zero ; rem = rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl quotSmall } rewrite multiplicationNIsCommutative b 0 = record { quotient = nonneg 0 ; rem = nonneg rem ; remSmall = inr (p , remIsSmall) ; divAlg = applyEquality nonneg (equalityCommutative pr) }
where
p : (nonneg rem nonneg 0) False
p pr2 rewrite pr = naughtE (equalityCommutative (nonnegInjective pr2))
divAlg {nonneg (succ a)} {negSucc b} _ _ | record { quot = succ quot ; rem = zero ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl quotSmall } = record { quotient = negSucc quot ; rem = nonneg zero ; remSmall = inl refl ; divAlg = applyEquality nonneg (transitivity (equalityCommutative pr) (applyEquality (λ i i +N 0) (multiplicationNIsCommutative (succ b) (succ quot)))) }
divAlg {nonneg (succ a)} {negSucc b} _ _ | record { quot = succ quot ; rem = succ rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = inl quotSmall } = record { quotient = negSucc quot ; rem = nonneg (succ rem) ; remSmall = inr ((λ ()) , remIsSmall) ; divAlg = applyEquality nonneg (transitivity (equalityCommutative pr) (applyEquality (λ i i +N succ rem) (multiplicationNIsCommutative (succ b) (succ quot)))) }
divAlg {negSucc a} {nonneg zero} _ b!=0 = exFalso (b!=0 refl)
divAlg {negSucc a} {nonneg (succ b)} _ _ with divisionAlg (succ b) (succ a)
divAlg {negSucc a} {nonneg (succ b)} _ _ | record { quot = zero ; rem = zero ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = exFalso (naughtE (transitivity (equalityCommutative (transitivity (Semiring.sumZeroRight Semiring _) (multiplicationNIsCommutative b 0))) pr))
divAlg {negSucc a} {nonneg (succ b)} _ _ | record { quot = succ quot ; rem = zero ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = record { quotient = negSucc quot ; rem = nonneg 0 ; remSmall = inl refl ; divAlg = applyEquality negSucc (succInjective (transitivity (equalityCommutative pr) t)) }
where
t : succ (quot +N b *N succ quot) +N 0 succ ((quot +N b *N quot) +N b)
t rewrite Semiring.sumZeroRight Semiring (succ (quot +N b *N succ quot)) | Semiring.commutative Semiring (quot +N b *N quot) b | Semiring.+Associative Semiring b quot (b *N quot) | Semiring.commutative Semiring b quot | equalityCommutative (Semiring.+Associative Semiring quot b (b *N quot)) | multiplicationNIsCommutative b (succ quot) | multiplicationNIsCommutative quot b = refl
divAlg {negSucc a} {nonneg (succ b)} _ _ | record { quot = zero ; rem = succ rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } rewrite multiplicationNIsCommutative b 0 | equalityCommutative (succInjective pr) = record { quotient = nonneg zero ; rem = negSucc rem ; remSmall = inr (((λ ()) , remIsSmall)) ; divAlg = refl }
divAlg {negSucc a} {nonneg (succ b)} _ _ | record { quot = succ quot ; rem = succ rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = record { quotient = negSucc quot ; rem = negSucc rem ; remSmall = inr ((λ ()) , remIsSmall) ; divAlg = applyEquality negSucc (succInjective (transitivity (equalityCommutative pr) t)) }
where
t : succ b *N succ quot +N succ rem succ (succ (succ b *N quot +N b) +N rem)
t rewrite Semiring.commutative Semiring ((quot +N b *N quot) +N b) rem | Semiring.commutative Semiring (succ rem) ((quot +N b *N quot) +N b) | multiplicationNIsCommutative b (succ quot) | equalityCommutative (Semiring.+Associative Semiring quot (b *N quot) b) | Semiring.commutative Semiring b (quot *N b) | multiplicationNIsCommutative b quot = refl
divAlg {negSucc a} {negSucc b} _ _ with divisionAlg (succ b) (succ a)
divAlg {negSucc a} {negSucc b} _ _ | record { quot = zero ; rem = zero ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } = exFalso (naughtE (transitivity (equalityCommutative (transitivity (Semiring.sumZeroRight Semiring (b *N 0)) (multiplicationNIsCommutative b 0))) pr))
divAlg {negSucc a} {negSucc b} _ _ | record { quot = zero ; rem = succ rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } rewrite multiplicationNIsCommutative b 0 | equalityCommutative (succInjective pr) = record { quotient = nonneg 0 ; rem = negSucc rem ; remSmall = inr ((λ ()) , remIsSmall) ; divAlg = refl }
divAlg {negSucc a} {negSucc b} _ _ | record { quot = succ quot ; rem = zero ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } rewrite Semiring.sumZeroRight Semiring (quot +N b *N succ quot) | Semiring.commutative Semiring b (succ quot) = record { quotient = nonneg (succ quot) ; rem = nonneg 0 ; remSmall = inl refl ; divAlg = applyEquality negSucc (succInjective (transitivity (equalityCommutative pr) t)) }
where
t : succ (quot +N b *N succ quot) succ ((b +N quot *N b) +N quot)
t rewrite Semiring.commutative Semiring quot (b *N succ quot) | multiplicationNIsCommutative b (succ quot) = refl
divAlg {negSucc a} {negSucc b} _ _ | record { quot = succ quot ; rem = succ rem ; pr = pr ; remIsSmall = inl remIsSmall ; quotSmall = quotSmall } rewrite multiplicationNIsCommutative b (succ quot) | Semiring.commutative Semiring quot (b +N quot *N b) | Semiring.commutative Semiring ((b +N quot *N b) +N quot) (succ rem) | Semiring.commutative Semiring rem ((b +N quot *N b) +N quot) = record { quotient = nonneg (succ quot) ; rem = negSucc rem ; remSmall = inr (((λ ()) , remIsSmall)) ; divAlg = applyEquality negSucc (equalityCommutative (succInjective pr)) }
EuclideanDomain : EuclideanDomain Ring
EuclideanDomain.isIntegralDomain EuclideanDomain = IntDom
EuclideanDomain.norm EuclideanDomain = norm'
EuclideanDomain.normSize EuclideanDomain = normSize
EuclideanDomain.divisionAlg EuclideanDomain {a} {b} a!=0 b!=0 = divAlg {a} {b} a!=0 b!=0