Files
agdaproofs/Numbers/Integers/Order.agda
2020-01-05 15:06:35 +00:00

100 lines
6.0 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Numbers.Integers.RingStructure.Ring
open import Semirings.Definition
open import Rings.Orders.Partial.Definition
open import Rings.Orders.Total.Definition
open import Setoids.Orders
open import Orders.Total.Definition
open import Orders.Partial.Definition
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
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 _≡_ 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 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 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
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))
POrderedRing : PartiallyOrderedRing Ring (SetoidTotalOrder.partial (totalOrderToSetoidTotalOrder Order))
PartiallyOrderedRing.orderRespectsAddition POrderedRing {a} {b} = orderRespectsAddition a b
PartiallyOrderedRing.orderRespectsMultiplication POrderedRing {a} {b} = orderRespectsMultiplication a b
OrderedRing : TotallyOrderedRing POrderedRing
TotallyOrderedRing.total OrderedRing = totalOrderToSetoidTotalOrder Order