mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-08 13:28:39 +00:00
Lots of refactoring towards partially-ordered ring R (#109)
This commit is contained in:
@@ -85,6 +85,7 @@ open import Rings.Divisible.Lemmas
|
||||
open import Rings.Associates.Lemmas
|
||||
open import Rings.InitialRing
|
||||
open import Rings.Homomorphisms.Lemmas
|
||||
open import Rings.UniqueFactorisationDomains.Definition
|
||||
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.DirectSum
|
||||
@@ -125,9 +126,6 @@ open import Monoids.Definition
|
||||
open import Semirings.Definition
|
||||
open import Semirings.Solver
|
||||
|
||||
open import Fields.CauchyCompletion.Group
|
||||
open import Fields.CauchyCompletion.Ring
|
||||
|
||||
open import Categories.Definition
|
||||
open import Categories.Functor.Definition
|
||||
open import Categories.Functor.Lemmas
|
||||
@@ -149,4 +147,7 @@ open import Modules.PolynomialModule
|
||||
open import Modules.Lemmas
|
||||
open import Modules.DirectSum
|
||||
|
||||
open import Fields.CauchyCompletion.Ring
|
||||
open import Fields.CauchyCompletion.Comparison
|
||||
|
||||
module Everything.Safe where
|
||||
|
@@ -32,6 +32,7 @@ open import Fields.CauchyCompletion.Definition order F
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
|
||||
private
|
||||
lemm : (m : ℕ) (a b : Sequence A) → index (apply _+_ a b) m ≡ (index a m) + (index b m)
|
||||
lemm zero a b = refl
|
||||
lemm (succ m) a b = lemm m (Sequence.tail a) (Sequence.tail b)
|
||||
|
@@ -40,8 +40,9 @@ open import Fields.CauchyCompletion.Comparison order F charNot2
|
||||
abstract
|
||||
|
||||
chain : {a b : A} (c : CauchyCompletion) → (a r<C c) → (c <Cr b) → a < b
|
||||
chain {a} {b} c (betweenAC , (0<betweenAC ,, (Nac , prAC))) (betweenCB , (0<betweenCB ,, (Nb , prBC))) = SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<betweenAC a)) (<WellDefined groupIsAbelian (Equivalence.reflexive eq) (SetoidPartialOrder.<Transitive pOrder (prAC (succ Nac +N Nb) (le Nb (applyEquality succ (Semiring.commutative ℕSemiring Nb Nac)))) (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<betweenCB (index (Sequence.tail (CauchyCompletion.elts c)) (Nac +N Nb))))))) (prBC (succ Nac +N Nb) (le Nac refl))
|
||||
chain {a} {b} c record { e = betweenAC ; 0<e = 0<betweenAC ; N = Nac ; pr = prAC } record { e = betweenCB ; 0<e = 0<betweenCB ; N = Nb ; property = prBC } = SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<betweenAC a)) (<WellDefined groupIsAbelian (Equivalence.reflexive eq) (SetoidPartialOrder.<Transitive pOrder (prAC (succ Nac +N Nb) (le Nb (applyEquality succ (Semiring.commutative ℕSemiring Nb Nac)))) (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<betweenCB (index (Sequence.tail (CauchyCompletion.elts c)) (Nac +N Nb))))))) (prBC (succ Nac +N Nb) (le Nac refl))
|
||||
|
||||
private
|
||||
approxLemma : (a : CauchyCompletion) (e e/2 : A) → (0G < e) → (e/2 + e/2 ∼ e) → (m N : ℕ) → abs ((index (CauchyCompletion.elts a) m) + inverse (index (CauchyCompletion.elts a) N)) < e/2 → (e/2 + index (CauchyCompletion.elts a) m) < (index (CauchyCompletion.elts a) N + e)
|
||||
approxLemma a e e/2 0<e prE/2 m N ans with totality 0R ((index (CauchyCompletion.elts a) m) + inverse (index (CauchyCompletion.elts a) N))
|
||||
approxLemma a e e/2 0<e prE/2 m N ans | inl (inl x) with <WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (invLeft))) identRight) groupIsAbelian (orderRespectsAddition ans (index (CauchyCompletion.elts a) N))
|
||||
@@ -51,9 +52,14 @@ abstract
|
||||
approxLemma a e e/2 0<e prE/2 m N ans | inr x with transferToRight additiveGroup (Equivalence.symmetric eq x)
|
||||
... | bl = <WellDefined (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq bl)) groupIsAbelian (orderRespectsAddition (halfLess e/2 e 0<e prE/2) (index (CauchyCompletion.elts a) N))
|
||||
|
||||
approximateAboveCrude : (a : CauchyCompletion) → Sg A (λ b → (a <Cr b))
|
||||
approximateAboveCrude : (a : CauchyCompletion) → A
|
||||
approximateAboveCrude a with CauchyCompletion.converges a 1R (0<1 (charNot2ImpliesNontrivial R charNot2))
|
||||
... | N , conv = ((((index (CauchyCompletion.elts a) (succ N)) + 1R) + 1R) + 1R) , (1R , (0<1 (charNot2ImpliesNontrivial R charNot2) ,, (N , ans)))
|
||||
... | N , conv = (((index (CauchyCompletion.elts a) (succ N)) + 1R) + 1R) + 1R
|
||||
|
||||
abstract
|
||||
approximateAboveCrudeIsAbove : (a : CauchyCompletion) → a <Cr (approximateAboveCrude a)
|
||||
approximateAboveCrudeIsAbove a with CauchyCompletion.converges a 1R (0<1 (charNot2ImpliesNontrivial R charNot2))
|
||||
... | N , conv = record { e = 1R ; 0<e = 0<1 (charNot2ImpliesNontrivial R charNot2) ; N = N ; property = ans }
|
||||
where
|
||||
ans : (m : ℕ) → (N <N m) → (1R + index (CauchyCompletion.elts a) m) < (((index (CauchyCompletion.elts a) (succ N) + 1R) + 1R) + 1R)
|
||||
ans m N<m with conv {m} {succ N} N<m (le 0 refl)
|
||||
@@ -72,6 +78,7 @@ abstract
|
||||
... | 0<e/4 with CauchyCompletion.converges a e/4 0<e/4
|
||||
... | N , cauchyBeyondN = index (CauchyCompletion.elts a) (succ N) + e/2
|
||||
|
||||
abstract
|
||||
rationalApproximatelyAboveIsAbove : (a : CauchyCompletion) (e : A) → (0<e : 0G < e) → a <Cr (rationalApproximatelyAbove a e 0<e)
|
||||
rationalApproximatelyAboveIsAbove a e 0<e with halve charNot2 e
|
||||
... | e/2 , prE/2 with halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prE/2) 0<e)
|
||||
@@ -81,7 +88,7 @@ abstract
|
||||
... | N , cauchyBeyondN with halve charNot2 e/4
|
||||
... | e/8 , prE/8 with halvePositive e/8 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prE/8) 0<e/4)
|
||||
... | 0<e/8 with CauchyCompletion.converges a e/8 0<e/8
|
||||
... | N2 , cauchyBeyondN2 = e/8 , (0<e/8 ,, ((N +N N2) , ans2))
|
||||
... | N2 , cauchyBeyondN2 = record { e = e/8 ; 0<e = 0<e/8 ; N = N +N N2 ; property = ans2 }
|
||||
where
|
||||
ans2 : (m : ℕ) → N +N N2 <N m → (e/8 + index (CauchyCompletion.elts a) m) < (index (CauchyCompletion.elts a) (succ N) + e/2)
|
||||
ans2 m <m with cauchyBeyondN {m} {succ N} (inequalityShrinkLeft <m) (le 0 refl)
|
||||
@@ -105,7 +112,7 @@ abstract
|
||||
... | 0<e/8 with CauchyCompletion.converges a e/8 0<e/8
|
||||
... | N8 , cauchyBeyondN8 with halve charNot2 e/8
|
||||
... | e/16 , prE/16 with halvePositive e/16 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prE/16) 0<e/8)
|
||||
... | 0<e/16 = ((e/2 + e/4) + e/8) , ((e/8 , (0<e/8 ,, ((N +N N8) , ans))) ,, (e/16 , (0<e/16 ,, (0 , t'))))
|
||||
... | 0<e/16 = record { i = (e/2 + e/4) + e/8 ; a<i = record { e = e/8 ; 0<e = 0<e/8 ; N = N +N N8 ; property = ans } ; i<b = record { e = e/16 ; 0<e = 0<e/16 ; N = 0 ; pr = t' } }
|
||||
where
|
||||
t' : (m : ℕ) → (0 <N m) → (((e/2 + e/4) + e/8) + e/16) < index (constSequence e) m
|
||||
t' m 0<m rewrite indexAndConst e m = <WellDefined (Equivalence.reflexive eq) prE/2 (<WellDefined (Equivalence.reflexive eq) (+WellDefined (Equivalence.reflexive eq) prE/4) (<WellDefined +Associative (Equivalence.symmetric eq +Associative) (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (<WellDefined (Equivalence.reflexive eq) prE/8 (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (halfLess e/16 e/8 0<e/8 prE/16) e/8))) (e/2 + e/4)))))
|
||||
@@ -130,7 +137,7 @@ abstract
|
||||
|
||||
approximateBelow : (a : CauchyCompletion) → (ε : A) → (0G < ε) → Sg A (λ b → (b r<C a) && (a +C (-C injection b)) <C (injection ε))
|
||||
approximateBelow a e 0<e with approximateAbove (-C a) e 0<e
|
||||
... | x , ((deltaXAnd-A , (0<deltaXA ,, (NdeltaXA , prDeltaXA))) ,, (rationalNear , ((bound , (0<bound ,, (N , prBound))) ,, (bound2 , (0<bound2 ,, (N2 , prBound2)))))) = inverse x , ((deltaXAnd-A , (0<deltaXA ,, (NdeltaXA , λ m N<m → <WellDefined (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (cancel {m} (CauchyCompletion.elts a))) identRight)))))) (Equivalence.transitive eq +Associative (Equivalence.transitive eq (+WellDefined invRight (Equivalence.reflexive eq)) identLeft)) (orderRespectsAddition (prDeltaXA m N<m) (inverse x + (index (CauchyCompletion.elts a) m)))))) ,, (rationalNear , ((bound , (0<bound ,, (N , pr1))) ,, (bound2 , (0<bound2 ,, (N2 , prBound2))))))
|
||||
... | x , (record { e = deltaXAnd-A ; 0<e = 0<deltaXA ; N = NdeltaXA ; property = prDeltaXA } ,, record { i = rationalNear ; a<i = record { e = bound ; 0<e = 0<bound ; N = N ; property = prBound } ; i<b = record { e = bound2 ; 0<e = 0<bound2 ; N = N2 ; pr = prBound2 } }) = inverse x , (record { e = deltaXAnd-A ; 0<e = 0<deltaXA ; N = NdeltaXA ; pr = λ m N<m → <WellDefined (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (cancel {m} (CauchyCompletion.elts a))) identRight)))))) (Equivalence.transitive eq +Associative (Equivalence.transitive eq (+WellDefined invRight (Equivalence.reflexive eq)) identLeft)) (orderRespectsAddition (prDeltaXA m N<m) (inverse x + (index (CauchyCompletion.elts a) m))) } ,, record { i = rationalNear ; a<i = record { e = bound ; 0<e = 0<bound ; N = N ; property = pr1 } ; i<b = record { e = bound2 ; 0<e = 0<bound2 ; N = N2 ; pr = prBound2 } })
|
||||
where
|
||||
cancel : {m : ℕ} (a : Sequence A) → (index (map inverse a) m) + (index a m) ∼ 0G
|
||||
cancel {m} a rewrite equalityCommutative (mapAndIndex a inverse m) = invLeft
|
||||
@@ -143,7 +150,7 @@ abstract
|
||||
... | below , (below<a ,, a-below<e) with approximateAbove a 1R (0<1 (charNot2ImpliesNontrivial R charNot2))
|
||||
... | above , (a<above ,, above-a<e) with totality 0R below
|
||||
boundModulus a | below , (below<a ,, a-below<e) | above , (a<above ,, above-a<e) | inl (inl 0<below) with totality 0R above
|
||||
boundModulus a | below , ((belowBound , (0<belowBound ,, (Nbelow , prBelow))) ,, a-below<e) | above , ((bound , (0<bound ,, (N , ans))) ,, above-a<e) | inl (inl 0<below) | inl (inl 0<above) = above , ((N +N Nbelow) , λ m N<m → SetoidPartialOrder.<Transitive pOrder (res m N<m) (ans m (inequalityShrinkLeft N<m)))
|
||||
boundModulus a | below , (record { e = belowBound ; 0<e = 0<belowBound ; N = Nbelow ; pr = prBelow } ,, a-below<e) | above , (record { e = bound ; 0<e = 0<bound ; N = N ; property = ans } ,, above-a<e) | inl (inl 0<below) | inl (inl 0<above) = above , ((N +N Nbelow) , λ m N<m → SetoidPartialOrder.<Transitive pOrder (res m N<m) (ans m (inequalityShrinkLeft N<m)))
|
||||
where
|
||||
res : (m : ℕ) → ((N +N Nbelow) <N m) → (abs (index (CauchyCompletion.elts a) m)) < (bound + index (CauchyCompletion.elts a) m)
|
||||
res m N<m with totality 0R (index (CauchyCompletion.elts a) m)
|
||||
@@ -154,42 +161,42 @@ abstract
|
||||
boundModulus a | below , (below<a ,, a-below<e) | above , (a<above ,, above-a<e) | inl (inl 0<below) | inr 0=above = exFalso (irreflexive (SetoidPartialOrder.<Transitive pOrder 0<below (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=above) (chain a below<a a<above))))
|
||||
boundModulus a | below , (below<a ,, a-below<e) | above , (a<above ,, above-a<e) | inl (inr below<0) with totality 0R above
|
||||
boundModulus a | below , (below<a ,, a-below<e) | above , (a<above ,, above-a<e) | inl (inr below<0) | inl (inl 0<above) with totality (inverse below) above
|
||||
boundModulus a | below , ((boundBelow , (0<boundBelow ,, (N , prBoundBelow))) ,, a-below<e) | above , ((boundAbove , (0<boundAbove ,, (Nabove , prBoundAbove))) ,, above-a<e) | inl (inr below<0) | inl (inl 0<above) | inl (inl -bel<ab) = above , ((N +N Nabove) , ans)
|
||||
boundModulus a | below , (record { e = boundBelow ; 0<e = 0<boundBelow ; N = N ; pr = prBoundBelow } ,, a-below<e) | above , (record { e = boundAbove ; 0<e = 0<boundAbove ; N = Nabove ; property = prBoundAbove } ,, above-a<e) | inl (inr below<0) | inl (inl 0<above) | inl (inl -bel<ab) = above , ((N +N Nabove) , ans)
|
||||
where
|
||||
ans : (m : ℕ) → (N +N Nabove <N m) → abs (index (CauchyCompletion.elts a) m) < above
|
||||
ans m N<m with totality 0G (index (CauchyCompletion.elts a) m)
|
||||
ans m N<m | inl (inl 0<am) = SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<boundAbove (index (CauchyCompletion.elts a) m))) (prBoundAbove m (inequalityShrinkRight N<m))
|
||||
ans m N<m | inl (inr am<0) = SetoidPartialOrder.<Transitive pOrder (ringSwapNegatives' (prBoundBelow m (inequalityShrinkLeft N<m))) (SetoidPartialOrder.<Transitive pOrder (ringSwapNegatives' (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<boundBelow below))) -bel<ab)
|
||||
ans m N<m | inr 0=am = <WellDefined 0=am (Equivalence.reflexive eq) 0<above
|
||||
boundModulus a | below , ((boundBelow , (0<boundBelow ,, (N , prBoundBelow))) ,, a-below<e) | above , ((boundAbove , (0<boundAbove ,, (Nabove , prBoundAbove))) ,, above-a<e) | inl (inr below<0) | inl (inl 0<above) | inl (inr ab<-bel) = inverse below , ((N +N Nabove) , ans)
|
||||
boundModulus a | below , (record { e = boundBelow ; 0<e = 0<boundBelow ; N = N ; pr = prBoundBelow } ,, a-below<e) | above , (record { e = boundAbove ; 0<e = 0<boundAbove ; N = Nabove ; property = prBoundAbove } ,, above-a<e) | inl (inr below<0) | inl (inl 0<above) | inl (inr ab<-bel) = inverse below , ((N +N Nabove) , ans)
|
||||
where
|
||||
ans : (m : ℕ) → (N +N Nabove <N m) → abs (index (CauchyCompletion.elts a) m) < (inverse below)
|
||||
ans m N<m with totality 0G (index (CauchyCompletion.elts a) m)
|
||||
ans m N<m | inl (inl 0<am) = SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<boundAbove (index (CauchyCompletion.elts a) m))) (prBoundAbove m (inequalityShrinkRight N<m))) ab<-bel
|
||||
ans m N<m | inl (inr am<0) = ringSwapNegatives' (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<boundBelow below)) (prBoundBelow m (inequalityShrinkLeft N<m)))
|
||||
ans m N<m | inr 0=am = <WellDefined 0=am (Equivalence.reflexive eq) (lemm2 below below<0)
|
||||
boundModulus a | below , ((boundBelow , (0<boundBelow ,, (N , prBoundBelow))) ,, a-below<e) | above , ((boundAbove , (0<boundAbove ,, (Nabove , prBoundAbove))) ,, above-a<e) | inl (inr below<0) | inl (inl 0<above) | inr -bel=ab = above , ((N +N Nabove) , ans)
|
||||
boundModulus a | below , (record { e = boundBelow ; 0<e = 0<boundBelow ; N = N ; pr = prBoundBelow } ,, a-below<e) | above , (record { e = boundAbove ; 0<e = 0<boundAbove ; N = Nabove ; property = prBoundAbove } ,, above-a<e) | inl (inr below<0) | inl (inl 0<above) | inr -bel=ab = above , ((N +N Nabove) , ans)
|
||||
where
|
||||
ans : (m : ℕ) → (N +N Nabove <N m) → abs (index (CauchyCompletion.elts a) m) < above
|
||||
ans m N<m with totality 0G (index (CauchyCompletion.elts a) m)
|
||||
ans m N<m | inl (inl 0<am) = SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<boundAbove (index (CauchyCompletion.elts a) m))) (prBoundAbove m (inequalityShrinkRight N<m))
|
||||
ans m N<m | inl (inr am<0) = <WellDefined (Equivalence.reflexive eq) (-bel=ab) (ringSwapNegatives' (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<boundBelow below)) (prBoundBelow m (inequalityShrinkLeft N<m))))
|
||||
ans m N<m | inr 0=am = <WellDefined 0=am (Equivalence.reflexive eq) 0<above
|
||||
boundModulus a | below , ((boundBelow , (0<boundBelow ,, (N , prBoundBelow))) ,, a-below<e) | above , ((boundAbove , (0<boundAbove ,, (Nabove , prBoundAbove))) ,, above-a<e) | inl (inr below<0) | inl (inr above<0) = inverse below , ((N +N Nabove) , ans)
|
||||
boundModulus a | below , (record { e = boundBelow ; 0<e = 0<boundBelow ; N = N ; pr = prBoundBelow } ,, a-below<e) | above , (record { e = boundAbove ; 0<e = 0<boundAbove ; N = Nabove ; property = prBoundAbove } ,, above-a<e) | inl (inr below<0) | inl (inr above<0) = inverse below , ((N +N Nabove) , ans)
|
||||
where
|
||||
ans : (m : ℕ) → ((N +N Nabove) <N m) → abs (index (CauchyCompletion.elts a) m) < inverse below
|
||||
ans m N<m with totality 0R (index (CauchyCompletion.elts a) m)
|
||||
ans m N<m | inl (inl 0<am) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<am (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<boundAbove (index (CauchyCompletion.elts a) m))) (prBoundAbove m (inequalityShrinkRight N<m))) above<0)))
|
||||
ans m N<m | inl (inr am<0) = ringSwapNegatives' (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<boundBelow below)) (prBoundBelow m (inequalityShrinkLeft N<m)))
|
||||
ans m N<m | inr 0=am = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq 0=am) (Equivalence.reflexive eq) (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<boundAbove (index (CauchyCompletion.elts a) m))) (prBoundAbove m (inequalityShrinkRight N<m))) above<0)))
|
||||
boundModulus a | below , ((boundBelow , (0<boundBelow ,, (N , prBoundBelow))) ,, a-below<e) | above , ((boundAbove , (0<boundAbove ,, (Nabove , prBoundAbove))) ,, above-a<e) | inl (inr below<0) | inr 0=above = inverse below , ((N +N Nabove) , ans)
|
||||
boundModulus a | below , (record { e = boundBelow ; 0<e = 0<boundBelow ; N = N ; pr = prBoundBelow } ,, a-below<e) | above , (record { e = boundAbove ; 0<e = 0<boundAbove ; N = Nabove ; property = prBoundAbove } ,, above-a<e) | inl (inr below<0) | inr 0=above = inverse below , ((N +N Nabove) , ans)
|
||||
where
|
||||
ans : (m : ℕ) → ((N +N Nabove) <N m) → abs (index (CauchyCompletion.elts a) m) < inverse below
|
||||
ans m N<m with totality 0R (index (CauchyCompletion.elts a) m)
|
||||
ans m N<m | inl (inl 0<am) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=above) (SetoidPartialOrder.<Transitive pOrder 0<am (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<boundAbove (index (CauchyCompletion.elts a) m))) (prBoundAbove m (inequalityShrinkRight N<m))))))
|
||||
ans m N<m | inl (inr am<0) = ringSwapNegatives' (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<boundBelow below)) (prBoundBelow m (inequalityShrinkLeft N<m)))
|
||||
ans m N<m | inr 0=am = <WellDefined 0=am (Equivalence.reflexive eq) (lemm2 _ below<0)
|
||||
boundModulus a | below , ((boundBelow , ((boundBelowDiff ,, (Nb , ansBelow)))) ,, a-below<e) | above , ((bound , (0<bound ,, (N , ans))) ,, above-a<e) | inr 0=below = above , ((N +N Nb) , λ m N<m → SetoidPartialOrder.<Transitive pOrder (res m N<m) (ans m (inequalityShrinkLeft N<m)))
|
||||
boundModulus a | below , (record { e = boundBelow ; 0<e = boundBelowDiff ; N = Nb ; pr = ansBelow } ,, a-below<e) | above , (record { e = bound ; 0<e = 0<bound ; N = N ; property = ans } ,, above-a<e) | inr 0=below = above , ((N +N Nb) , λ m N<m → SetoidPartialOrder.<Transitive pOrder (res m N<m) (ans m (inequalityShrinkLeft N<m)))
|
||||
where
|
||||
res : (m : ℕ) → (N +N Nb) <N m → (abs (index (CauchyCompletion.elts a) m)) < (bound + index (CauchyCompletion.elts a) m)
|
||||
res m N<m with totality 0R (index (CauchyCompletion.elts a) m)
|
||||
|
@@ -15,6 +15,7 @@ open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
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
|
||||
@@ -35,18 +36,16 @@ open import Fields.Lemmas F
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
|
||||
shrinkInequality : {a b c : ℕ} → a +N b <N c → a <N c
|
||||
shrinkInequality {a} {b} {c} (le x pr) = le (x +N b) (transitivity (applyEquality succ (transitivity (equalityCommutative (Semiring.+Associative ℕSemiring x _ _)) (applyEquality (x +N_) (Semiring.commutative ℕSemiring b a)))) pr)
|
||||
|
||||
shrinkInequality' : {a b c : ℕ} → a +N b <N c → b <N c
|
||||
shrinkInequality' {a} {b} {c} (le x pr) = le (x +N a) (transitivity (applyEquality succ (equalityCommutative (Semiring.+Associative ℕSemiring x _ _))) pr)
|
||||
|
||||
-- "<C rational", where the r indicates where the rational number goes
|
||||
_<Cr_ : CauchyCompletion → A → Set (m ⊔ o)
|
||||
a <Cr b = Sg A (λ ε → (0G < ε) && Sg ℕ (λ N → ((m : ℕ) → (N<m : N <N m) → (ε + index (CauchyCompletion.elts a) m) < b)))
|
||||
record _<Cr_ (a : CauchyCompletion) (b : A) : Set (m ⊔ o) where
|
||||
field
|
||||
e : A
|
||||
0<e : 0G < e
|
||||
N : ℕ
|
||||
property : ((m : ℕ) → (N<m : N <N m) → (e + index (CauchyCompletion.elts a) m) < b)
|
||||
|
||||
<CrWellDefinedRight : (a : CauchyCompletion) (b c : A) → b ∼ c → a <Cr b → a <Cr c
|
||||
<CrWellDefinedRight a b c b=c (ε , (0<e ,, (N , pr))) = ε , (0<e ,, (N , λ m N<m → <WellDefined (Equivalence.reflexive eq) b=c (pr m N<m)))
|
||||
<CrWellDefinedRight a b c b=c record { e = ε ; 0<e = 0<e ; N = N ; property = pr } = record { e = ε ; 0<e = 0<e ; N = N ; property = λ m N<m → <WellDefined (Equivalence.reflexive eq) b=c (pr m N<m) }
|
||||
|
||||
<CrWellDefinedLemma : (a b e/2 e : A) (0<e : 0G < e) (pr : e/2 + e/2 ∼ e) → abs (a + inverse b) < e/2 → (e/2 + b) < (e + a)
|
||||
<CrWellDefinedLemma a b e/2 e 0<e pr a-b<e with totality 0G (a + inverse b)
|
||||
@@ -55,23 +54,27 @@ a <Cr b = Sg A (λ ε → (0G < ε) && Sg ℕ (λ N → ((m : ℕ) → (N<m : N
|
||||
<CrWellDefinedLemma a b e/2 e 0<e pr a-b<e | inr 0=a-b = <WellDefined (Equivalence.reflexive eq) (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (transferToRight (Ring.additiveGroup R) (Equivalence.symmetric eq 0=a-b)))) (orderRespectsAddition (halfLess e/2 e 0<e pr) b)
|
||||
|
||||
<CrWellDefinedLeft : (a b : CauchyCompletion) (c : A) → Setoid._∼_ cauchyCompletionSetoid a b → a <Cr c → b <Cr c
|
||||
<CrWellDefinedLeft a b c a=b (ε , (0<e ,, (N , pr))) with halve charNot2 ε
|
||||
<CrWellDefinedLeft a b c a=b record { e = ε ; 0<e = 0<e ; N = N ; property = pr } with halve charNot2 ε
|
||||
... | e/2 , prE/2 with halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prE/2) 0<e)
|
||||
... | 0<e/2 with a=b e/2 0<e/2
|
||||
... | N2 , prN2 = e/2 , (0<e/2 ,, ((N +N N2) , ans))
|
||||
... | N2 , prN2 = record { e = e/2 ; 0<e = 0<e/2 ; N = N +N N2 ; property = ans }
|
||||
where
|
||||
ans : (m : ℕ) → (N<m : (N +N N2) <N m) → (e/2 + index (CauchyCompletion.elts b) m) < c
|
||||
ans m N<m with prN2 {m} (shrinkInequality' N<m)
|
||||
... | bl rewrite indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts b)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts b) inverse m) = SetoidPartialOrder.<Transitive pOrder (<CrWellDefinedLemma (index (CauchyCompletion.elts a) m) (index (CauchyCompletion.elts b) m) e/2 ε 0<e prE/2 bl) (pr m (shrinkInequality N<m))
|
||||
ans m N<m with prN2 {m} (inequalityShrinkRight N<m)
|
||||
... | bl rewrite indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts b)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts b) inverse m) = SetoidPartialOrder.<Transitive pOrder (<CrWellDefinedLemma (index (CauchyCompletion.elts a) m) (index (CauchyCompletion.elts b) m) e/2 ε 0<e prE/2 bl) (pr m (inequalityShrinkLeft N<m))
|
||||
|
||||
<CrWellDefined : (a b : CauchyCompletion) (c d : A) → Setoid._∼_ cauchyCompletionSetoid a b → c ∼ d → a <Cr c → b <Cr d
|
||||
<CrWellDefined a b c d a=b c=d a<c = <CrWellDefinedLeft a b d a=b (<CrWellDefinedRight a c d c=d a<c)
|
||||
|
||||
_r<C_ : A → CauchyCompletion → Set (m ⊔ o)
|
||||
a r<C b = Sg A (λ ε → (0G < ε) && Sg ℕ (λ N → ((m : ℕ) → (N<m : N <N m) → (a + ε) < index (CauchyCompletion.elts b) m)))
|
||||
record _r<C_ (a : A) (b : CauchyCompletion) : Set (m ⊔ o) where
|
||||
field
|
||||
e : A
|
||||
0<e : 0G < e
|
||||
N : ℕ
|
||||
pr : (m : ℕ) → (N<m : N <N m) → (a + e) < index (CauchyCompletion.elts b) m
|
||||
|
||||
r<CWellDefinedLeft : (a b : A) (c : CauchyCompletion) → (a ∼ b) → (a r<C c) → b r<C c
|
||||
r<CWellDefinedLeft a b c a=b (e , (0<e ,, (N , pr))) = e , (0<e ,, (N , λ m N<m → <WellDefined (+WellDefined a=b (Equivalence.reflexive eq)) (Equivalence.reflexive eq) (pr m N<m)))
|
||||
r<CWellDefinedLeft a b c a=b record { e = e ; 0<e = 0<e ; N = N ; pr = pr } = record { e = e ; 0<e = 0<e ; N = N ; pr = λ m N<m → <WellDefined (+WellDefined a=b (Equivalence.reflexive eq)) (Equivalence.reflexive eq) (pr m N<m) }
|
||||
|
||||
r<CWellDefinedLemma : (a b c e e/2 : A) (_ : e/2 + e/2 ∼ e) (0<e : 0G < e) (_ : abs (a + inverse b) < e/2) (_ : (c + e) < a) → (c + e/2) < b
|
||||
r<CWellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a with totality 0G (a + inverse b)
|
||||
@@ -80,49 +83,72 @@ r<CWellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a | inl (inr a-b<0) = SetoidPar
|
||||
r<CWellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a | inr 0=a-b = SetoidPartialOrder.<Transitive pOrder (<WellDefined groupIsAbelian (Equivalence.reflexive eq) (orderRespectsAddition (halfLess e/2 e 0<e prE/2) c)) (<WellDefined (groupIsAbelian {c} {e}) (transferToRight additiveGroup (Equivalence.symmetric eq 0=a-b)) c+e<a)
|
||||
|
||||
r<CWellDefinedRight : (a b : CauchyCompletion) (c : A) → Setoid._∼_ cauchyCompletionSetoid a b → c r<C a → c r<C b
|
||||
r<CWellDefinedRight a b c a=b (e , (0<e ,, (N , pr))) with halve charNot2 e
|
||||
r<CWellDefinedRight a b c a=b record { e = e ; 0<e = 0<e ; N = N ; pr = pr } with halve charNot2 e
|
||||
... | e/2 , prE/2 with halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prE/2) 0<e)
|
||||
... | 0<e/2 with a=b e/2 0<e/2
|
||||
... | N2 , prN2 = e/2 , (0<e/2 ,, ((N +N N2) , ans))
|
||||
... | N2 , prN2 = record { e = e/2 ; 0<e = 0<e/2 ; N = N +N N2 ; pr = ans }
|
||||
where
|
||||
ans : (m : ℕ) → (N<m : (N +N N2) <N m) → (c + e/2) < index (CauchyCompletion.elts b) m
|
||||
ans m N<m with prN2 {m} (shrinkInequality' N<m)
|
||||
... | bl rewrite indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts b)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts b) inverse m) = r<CWellDefinedLemma (index (CauchyCompletion.elts a) m) (index (CauchyCompletion.elts b) m) c e e/2 prE/2 0<e bl (pr m (shrinkInequality N<m))
|
||||
ans m N<m with prN2 {m} (inequalityShrinkRight N<m)
|
||||
... | bl rewrite indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts b)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts b) inverse m) = r<CWellDefinedLemma (index (CauchyCompletion.elts a) m) (index (CauchyCompletion.elts b) m) c e e/2 prE/2 0<e bl (pr m (inequalityShrinkLeft N<m))
|
||||
|
||||
_<C_ : CauchyCompletion → CauchyCompletion → Set (m ⊔ o)
|
||||
a <C b = Sg A (λ i → (a <Cr i) && (i r<C b))
|
||||
record _<C_ (a : CauchyCompletion) (b : CauchyCompletion) : Set (m ⊔ o) where
|
||||
field
|
||||
i : A
|
||||
a<i : a <Cr i
|
||||
i<b : i r<C b
|
||||
|
||||
<CWellDefined : {a b c d : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid c d → a <C c → b <C d
|
||||
<CWellDefined {a} {b} {c} {d} a=b c=d (bound , (a<bound ,, bound<b)) = bound , (<CrWellDefinedLeft a b bound a=b a<bound ,, r<CWellDefinedRight c d bound c=d bound<b)
|
||||
<CWellDefined {a} {b} {c} {d} a=b c=d record { i = bound ; a<i = a<bound ; i<b = bound<b } = record { i = bound ; a<i = <CrWellDefinedLeft a b bound a=b a<bound ; i<b = r<CWellDefinedRight c d bound c=d bound<b }
|
||||
|
||||
<CrPreserves : {a b : A} → a < b → injection a <Cr b
|
||||
<CrPreserves {a} {b} a<b with dense charNot2 a<b
|
||||
... | between , (a<between ,, between<b) = (between + inverse a) , (<WellDefined invRight (Equivalence.reflexive eq) (orderRespectsAddition a<between (inverse a)) ,, (0 , λ m _ → <WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq identRight) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (ans a m)) +Associative)) (Equivalence.reflexive eq) between<b))
|
||||
<CInheritedL : {a b : A} → a < b → injection a <Cr b
|
||||
<CInheritedL {a} {b} a<b with dense charNot2 a<b
|
||||
... | between , (a<between ,, between<b) = record { e = between + inverse a ; 0<e = <WellDefined invRight (Equivalence.reflexive eq) (orderRespectsAddition a<between (inverse a)) ; N = 0 ; property = λ m _ → <WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq identRight) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (ans a m)) +Associative)) (Equivalence.reflexive eq) between<b }
|
||||
where
|
||||
ans : (x : A) (m : ℕ) → 0G ∼ inverse x + index (constSequence x) m
|
||||
ans x m rewrite indexAndConst x m = Equivalence.symmetric eq invLeft
|
||||
|
||||
r<CPreserves : {a b : A} → a < b → a r<C injection b
|
||||
r<CPreserves {a} {b} a<b with dense charNot2 a<b
|
||||
... | between , (a<between ,, between<b) = (between + inverse a) , (<WellDefined invRight (Equivalence.reflexive eq) (orderRespectsAddition a<between (inverse a)) ,, (0 , λ m _ → <WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq identRight) (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq invLeft))) +Associative) groupIsAbelian) (identityOfIndiscernablesLeft _∼_ (Equivalence.reflexive eq) (indexAndConst b m)) between<b))
|
||||
<CInheritedR : {a b : A} → a < b → a r<C injection b
|
||||
<CInheritedR {a} {b} a<b with dense charNot2 a<b
|
||||
... | between , (a<between ,, between<b) = record { e = between + inverse a ; 0<e = <WellDefined invRight (Equivalence.reflexive eq) (orderRespectsAddition a<between (inverse a)) ; N = 0 ; pr = λ m _ → <WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq identRight) (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq invLeft))) +Associative) groupIsAbelian) (identityOfIndiscernablesLeft _∼_ (Equivalence.reflexive eq) (indexAndConst b m)) between<b }
|
||||
|
||||
<CInherited : {a b : A} → a < b → (injection a) <C (injection b)
|
||||
<CInherited {a} {b} a<b with dense charNot2 a<b
|
||||
... | between , (a<between ,, between<b) = between , (<CrPreserves a<between ,, r<CPreserves between<b)
|
||||
... | between , (a<between ,, between<b) = record { i = between ; a<i = <CInheritedL a<between ; i<b = <CInheritedR between<b }
|
||||
|
||||
<CInherited' : {a b : A} → (injection a) <C (injection b) → a < b
|
||||
<CInherited' {a} {b} (bound , ((Na , (0<Na ,, (Na2 , prA))) ,, ((Nb , (0<Nb ,, (Nb2 , prB)))))) with prA (succ Na2) (le 0 refl)
|
||||
... | bl with prB (succ Nb2) (le 0 refl)
|
||||
... | bl2 rewrite indexAndConst a Na2 | indexAndConst b Nb2 = SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<Na a)) bl) (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<Nb bound)) bl2)
|
||||
<CCollapsesL : (a b : A) → (injection a) <Cr b → a < b
|
||||
<CCollapsesL a b record { e = e ; 0<e = 0<i ; N = N ; property = pr } with pr (succ N) (le 0 refl)
|
||||
... | bl rewrite indexAndConst a N = <Transitive (<WellDefined identLeft reflexive (orderRespectsAddition 0<i a)) bl
|
||||
|
||||
<CCollapsesR : (a b : A) → a r<C injection b → a < b
|
||||
<CCollapsesR a b record { e = e ; 0<e = 0<e ; N = N ; pr = pr } with pr (succ N) (le 0 refl)
|
||||
... | bl rewrite indexAndConst b N = <Transitive (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<e a)) bl
|
||||
|
||||
<CCollapses : {a b : A} → (injection a) <C (injection b) → a < b
|
||||
<CCollapses {a} {b} record { i = inter ; a<i = pr1 ; i<b = pr2 } = <Transitive (<CCollapsesL a inter pr1) (<CCollapsesR inter b pr2)
|
||||
|
||||
<CRelaxL : {a : A} {b : CauchyCompletion} → a r<C b → injection a <C b
|
||||
<CRelaxL {a} {b} record { e = e ; 0<e = 0<e ; N = N ; pr = pr } with halve charNot2 e
|
||||
... | e/2 , e/2pr = record { i = a + e/2 ; a<i = <CInheritedL (<WellDefined identLeft groupIsAbelian (orderRespectsAddition (halvePositive' e/2pr 0<e) a)) ; i<b = record { e = e/2 ; 0<e = halvePositive' e/2pr 0<e ; N = N ; pr = λ m N<m → <WellDefined (transitive (+WellDefined reflexive (symmetric e/2pr)) +Associative) reflexive (pr m N<m) } }
|
||||
|
||||
<CRelaxL' : {a : A} {b : CauchyCompletion} → injection a <C b → a r<C b
|
||||
<CRelaxL' {a} {b} record { i = i ; a<i = a<i ; i<b = record { e = e ; 0<e = 0<e ; N = N ; pr = pr } } = record { e = e ; 0<e = 0<e ; N = N ; pr = λ m N<m → <Transitive (orderRespectsAddition (<CCollapsesL _ _ a<i) e) (pr m N<m) }
|
||||
|
||||
<CRelaxR : {a : CauchyCompletion} {b : A} → a <Cr b → a <C injection b
|
||||
<CRelaxR {a} {b} record { e = ε ; 0<e = 0<ε ; N = N ; property = property } with halve charNot2 ε
|
||||
... | e/2 , e/2+e/2=e = record { i = b + inverse e/2 ; a<i = record { e = e/2 ; 0<e = halvePositive' e/2+e/2=e 0<ε ; N = N ; property = λ m N<m → <WellDefined (transitive groupIsAbelian (transitive +Associative (+WellDefined (transitive (+WellDefined reflexive (symmetric e/2+e/2=e)) (transitive +Associative (transitive (+WellDefined invLeft reflexive) identLeft))) reflexive))) reflexive (orderRespectsAddition (property m N<m) (inverse e/2)) } ; i<b = <CInheritedR (<WellDefined groupIsAbelian identLeft (orderRespectsAddition (<WellDefined reflexive (invIdent additiveGroup) (ringSwapNegatives (<WellDefined (symmetric (invTwice additiveGroup 0R)) (symmetric (invTwice additiveGroup e/2)) (halvePositive' e/2+e/2=e 0<ε)))) b)) }
|
||||
|
||||
<CRelaxR' : {a : CauchyCompletion} {b : A} → a <C injection b → a <Cr b
|
||||
<CRelaxR' {a} {b} record { i = i ; a<i = record { e = ε ; 0<e = 0<ε ; N = N ; property = property } ; i<b = i<b } = record { e = ε ; 0<e = 0<ε ; N = N ; property = λ m N<m → <Transitive (property m N<m) (<CCollapsesR _ _ i<b) }
|
||||
|
||||
<CIrreflexive : {a : CauchyCompletion} → a <C a → False
|
||||
<CIrreflexive {a} (bound , ((bA , (0<bA ,, (Na , prA))) ,, ((bB , (0<bB ,, (Nb , prB)))))) with prA (succ Na +N Nb) (le Nb (applyEquality succ (Semiring.commutative ℕSemiring Nb Na)))
|
||||
<CIrreflexive {a} record { i = bound ; a<i = record { e = bA ; 0<e = 0<bA ; N = Na ; property = prA } ; i<b = record { e = bB ; 0<e = 0<bB ; N = Nb ; pr = prB } } with prA (succ Na +N Nb) (le Nb (applyEquality succ (Semiring.commutative ℕSemiring Nb Na)))
|
||||
... | bl with prB (succ Na +N Nb) (le Na refl)
|
||||
... | bl2 with <WellDefined (Equivalence.transitive eq +Associative (+WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) groupIsAbelian) (Equivalence.reflexive eq))) groupIsAbelian (ringAddInequalities bl bl2)
|
||||
... | bad = irreflexive (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft +Associative (<WellDefined (Equivalence.reflexive eq) groupIsAbelian (orderRespectsAddition (<WellDefined identLeft (Equivalence.reflexive eq) (ringAddInequalities 0<bA 0<bB)) (index (CauchyCompletion.elts a) (succ Na +N Nb) + bound)))) bad)
|
||||
|
||||
<CTransitive : {a b c : CauchyCompletion} → a <C b → b <C c → a <C c
|
||||
<CTransitive {a} {b} {c} (boundA , ((eL , (0<eL ,, prL)) ,, (eR , (0<eR ,, (Nr , prR))))) (boundB , ((fL , (0<fL ,, (Nfl , prFl))) ,, (fR , (0<fR ,, (N , prFR))))) = boundA , ((eL , (0<eL ,, prL)) ,, (eR , (0<eR ,, (((Nr +N Nfl) +N N) , λ m N<m → SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (prR m (shrinkInequality {Nr} (shrinkInequality N<m))) (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<fL (index (CauchyCompletion.elts b) m)))) (prFl m (shrinkInequality' {Nr} (shrinkInequality N<m)))) (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<fR boundB))) (prFR m (shrinkInequality' N<m))))))
|
||||
<CTransitive {a} {b} {c} record { i = boundA ; a<i = record { e = eL ; 0<e = 0<eL ; N = NL ; property = prL } ; i<b = record { e = eR ; 0<e = 0<eR ; N = Nr ; pr = prR } } record { i = boundB ; a<i = record { e = fL ; 0<e = 0<fL ; N = Nfl ; property = prFl } ; i<b = record { e = fR ; 0<e = 0<fR ; N = N ; pr = prFR } } = record { i = boundA ; a<i = record { e = eL ; 0<e = 0<eL ; N = NL ; property = prL } ; i<b = record { e = eR ; 0<e = 0<eR ; N = (Nr +N Nfl) +N N ; pr = λ m N<m → SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (prR m (inequalityShrinkLeft {Nr} (inequalityShrinkLeft N<m))) (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<fL (index (CauchyCompletion.elts b) m)))) (prFl m (inequalityShrinkRight {Nr} (inequalityShrinkLeft N<m)))) (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<fR boundB))) (prFR m (inequalityShrinkRight N<m)) } }
|
||||
|
||||
<COrder : SetoidPartialOrder cauchyCompletionSetoid _<C_
|
||||
SetoidPartialOrder.<WellDefined <COrder {a} {b} {c} {d} a=b c=d a<c = <CWellDefined {a} {b} {c} {d} a=b c=d a<c
|
||||
|
@@ -42,3 +42,11 @@ CauchyCompletion.converges (injection a) = λ ε 0<e → 0 , λ {m} {n} _ _ →
|
||||
where
|
||||
t : (m n : ℕ) → index (constSequence a) m + inverse (index (constSequence a) n) ∼ 0R
|
||||
t m n = identityOfIndiscernablesLeft _∼_ (identityOfIndiscernablesLeft _∼_ invRight (equalityCommutative (applyEquality (λ i → a + inverse i) (indexAndConst a n)))) (applyEquality (_+ inverse (index (constSequence a) n)) (equalityCommutative (indexAndConst a m)))
|
||||
|
||||
-- Some slightly odd things here relating to equality rather than equivalence. Ultimately this is here so we can say Q → R is a genuine injection, not just a setoid one.
|
||||
private
|
||||
lemma : {x y : CauchyCompletion} → x ≡ y → CauchyCompletion.elts x ≡ CauchyCompletion.elts y
|
||||
lemma {x} {.x} refl = refl
|
||||
|
||||
CInjection' : Injection injection
|
||||
CInjection' pr = headInjective (lemma pr)
|
||||
|
@@ -4,7 +4,8 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Order
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Fields.Fields
|
||||
@@ -13,19 +14,20 @@ 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.Field {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {tOrder : SetoidTotalOrder {_<_ = _<_} pOrder} {R : Ring S _+_ _*_} (order : OrderedRing R tOrder) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
module Fields.CauchyCompletion.Field {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
|
||||
|
||||
open Setoid S
|
||||
open SetoidTotalOrder tOrder
|
||||
open PartiallyOrderedRing pRing
|
||||
open SetoidPartialOrder pOrder
|
||||
open Equivalence eq
|
||||
open OrderedRing order
|
||||
open SetoidTotalOrder (TotallyOrderedRing.total order)
|
||||
open Field F
|
||||
open Group (Ring.additiveGroup R)
|
||||
|
||||
open import Rings.Orders.Lemmas(order)
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Multiplication order F charNot2
|
||||
open import Fields.CauchyCompletion.Addition order F charNot2
|
||||
@@ -35,12 +37,12 @@ open import Fields.CauchyCompletion.Ring order F charNot2
|
||||
open import Fields.CauchyCompletion.Comparison order F charNot2
|
||||
|
||||
Cnontrivial : (pr : Setoid._∼_ cauchyCompletionSetoid (injection (Ring.0R R)) (injection (Ring.1R R))) → False
|
||||
Cnontrivial pr with pr (Ring.1R R) (0<1 (charNot2ImpliesNontrivial charNot2))
|
||||
Cnontrivial pr with pr (Ring.1R R) (0<1 (charNot2ImpliesNontrivial R charNot2))
|
||||
Cnontrivial pr | N , b with b {succ N} (le 0 refl)
|
||||
... | bl rewrite indexAndApply (constSequence 0G) (map inverse (constSequence (Ring.1R R))) _+_ {N} | indexAndConst 0G N | equalityCommutative (mapAndIndex (constSequence (Ring.1R R)) inverse N) | indexAndConst (Ring.1R R) N = irreflexive {Ring.1R R} (<WellDefined (Equivalence.transitive eq (absWellDefined _ _ identLeft) (Equivalence.transitive eq (Equivalence.symmetric eq (absNegation (Ring.1R R))) abs1Is1)) (Equivalence.reflexive eq) bl)
|
||||
|
||||
boundedMap : A → A
|
||||
boundedMap a with SetoidTotalOrder.totality tOrder 0G a
|
||||
boundedMap a with totality 0G a
|
||||
boundedMap a | inl (inl x) = underlying (allInvertible a λ pr → irreflexive (<WellDefined (Equivalence.reflexive eq) pr x))
|
||||
boundedMap a | inl (inr x) = underlying (allInvertible a λ pr → irreflexive (<WellDefined pr (Equivalence.reflexive eq) x))
|
||||
boundedMap a | inr x = Ring.1R R
|
||||
|
@@ -25,12 +25,56 @@ open Equivalence eq
|
||||
open TotallyOrderedRing order
|
||||
open Field F
|
||||
open Group (Ring.additiveGroup R)
|
||||
open Ring R
|
||||
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Addition order F charNot2
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
|
||||
abstract
|
||||
+CCommutative : {a b : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a +C b) (b +C a)
|
||||
+CCommutative {a} {b} ε 0<e = 0 , ans
|
||||
where
|
||||
foo : {x y : A} → (x + y) + inverse (y + x) ∼ 0G
|
||||
foo = Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (inverseWellDefined additiveGroup groupIsAbelian)) invRight
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (a +C b)) (map inverse (CauchyCompletion.elts (b +C a)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (a +C b)) (map inverse (CauchyCompletion.elts (b +C a))) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts b) _+_ {m} | equalityCommutative (mapAndIndex (apply _+_ (CauchyCompletion.elts b) (CauchyCompletion.elts a)) inverse m) | indexAndApply (CauchyCompletion.elts b) (CauchyCompletion.elts a) _+_ {m} = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ foo) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) absZero))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
private
|
||||
abstract
|
||||
additionWellDefinedLeft : (a b c : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid (a +C c) (b +C c)
|
||||
additionWellDefinedLeft record { elts = a ; converges = aConv } record { elts = b ; converges = bConv } record { elts = c ; converges = cConv } a=b ε 0<e with a=b ε 0<e
|
||||
... | Na-b , prA-b = Na-b , ans
|
||||
where
|
||||
ans : {m : ℕ} → Na-b <N m → abs (index (apply _+_ (apply _+_ a c) (map inverse (apply _+_ b c))) m) < ε
|
||||
ans {m} mBig with prA-b {m} mBig
|
||||
... | bl rewrite indexAndApply (apply _+_ a c) (map inverse (apply _+_ b c)) _+_ {m} | indexAndApply a c _+_ {m} | equalityCommutative (mapAndIndex (apply _+_ b c) inverse m) | indexAndApply b c _+_ {m} = <WellDefined (absWellDefined _ _ t) (Equivalence.reflexive eq) bl
|
||||
where
|
||||
t : index (apply _+_ a (map inverse b)) m ∼ ((index a m + index c m) + inverse (index b m + index c m))
|
||||
t rewrite indexAndApply a (map inverse b) _+_ {m} | equalityCommutative (mapAndIndex b inverse m) = Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq identLeft) (+WellDefined (Equivalence.symmetric eq invRight) (Equivalence.reflexive eq))) (Equivalence.symmetric eq +Associative)) (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (invContravariant additiveGroup))))) (+Associative {index a m})
|
||||
|
||||
additionPreservedLeft : {a b : A} {c : CauchyCompletion} → (a ∼ b) → Setoid._∼_ cauchyCompletionSetoid (injection a +C c) (injection b +C c)
|
||||
additionPreservedLeft {a} {b} {c} a=b = additionWellDefinedLeft (injection a) (injection b) c (injectionPreservesSetoid a b a=b)
|
||||
|
||||
additionPreservedRight : {a b : A} {c : CauchyCompletion} → (a ∼ b) → Setoid._∼_ cauchyCompletionSetoid (c +C injection a) (c +C injection b)
|
||||
additionPreservedRight {a} {b} {c} a=b = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {c +C injection a} {injection a +C c} {c +C injection b} (+CCommutative {c} {injection a}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection a +C c} {injection b +C c} {c +C injection b} (additionPreservedLeft {a} {b} {c} a=b) (+CCommutative {injection b} {c}))
|
||||
|
||||
additionPreserved : {a b c d : A} → (a ∼ b) → (c ∼ d) → Setoid._∼_ cauchyCompletionSetoid (injection a +C injection c) (injection b +C injection d)
|
||||
additionPreserved {a} {b} {c} {d} a=b c=d = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection a +C injection c} {injection a +C injection d} {injection b +C injection d} (additionPreservedRight {c} {d} {injection a} c=d) (additionPreservedLeft {a} {b} {injection d} a=b)
|
||||
|
||||
additionWellDefinedRight : (a b c : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid b c → Setoid._∼_ cauchyCompletionSetoid (a +C b) (a +C c)
|
||||
additionWellDefinedRight a b c b=c = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {a +C b} {b +C a} {a +C c} (+CCommutative {a} {b}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {b +C a} {c +C a} {a +C c} (additionWellDefinedLeft b c a b=c) (+CCommutative {c} {a}))
|
||||
|
||||
additionWellDefined : {a b c d : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid c d → Setoid._∼_ cauchyCompletionSetoid (a +C c) (b +C d)
|
||||
additionWellDefined {a} {b} {c} {d} a=b c=d = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {a +C c} {a +C d} {b +C d} (additionWellDefinedRight a c d c=d) (additionWellDefinedLeft a b d a=b)
|
||||
|
||||
additionHom : (x y : A) → Setoid._∼_ cauchyCompletionSetoid (injection (x + y)) (injection x +C injection y)
|
||||
additionHom x y ε 0<e = 0 , ans
|
||||
where
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (injection (x + y))) (map inverse (CauchyCompletion.elts (injection x +C injection y)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (injection (x + y))) (map inverse (CauchyCompletion.elts (injection x +C injection y))) _+_ {m} | equalityCommutative (mapAndIndex (apply _+_ (constSequence x) (constSequence y)) inverse m) | indexAndConst (x + y) m | indexAndApply (constSequence x) (constSequence y) _+_ {m} | indexAndConst x m | indexAndConst y m = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ invRight) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) absZero))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
Cassoc : {a b c : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a +C (b +C c)) ((a +C b) +C c)
|
||||
Cassoc {a} {b} {c} ε 0<e = 0 , ans
|
||||
where
|
||||
|
@@ -39,6 +39,7 @@ open import Fields.CauchyCompletion.Approximation order F charNot2
|
||||
0!=1 : {e : A} → (0G < e) → 0R ∼ 1R → False
|
||||
0!=1 {e} 0<e 0=1 = irreflexive (<WellDefined (Equivalence.reflexive eq) (oneZeroImpliesAllZero R 0=1) 0<e)
|
||||
|
||||
private
|
||||
littleLemma : {a b c d : A} → ((a * b) + inverse (c * d)) ∼ ((a * (b + inverse d)) + (d * (a + inverse c)))
|
||||
littleLemma {a} {b} {c} {d} = Equivalence.transitive eq (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq identLeft) (+WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (+WellDefined (Equivalence.transitive eq (ringMinusExtracts R) (inverseWellDefined additiveGroup *Commutative)) (Equivalence.reflexive eq)) (invLeft {d * a}))) (Equivalence.transitive eq (Equivalence.symmetric eq (ringMinusExtracts' R)) *Commutative))) (Equivalence.symmetric eq +Associative))) (+Associative)) (Equivalence.symmetric eq (+WellDefined (*DistributesOver+) (*DistributesOver+)))
|
||||
|
||||
@@ -99,7 +100,7 @@ CauchyCompletion.converges (record { elts = a ; converges = aConv } *C record {
|
||||
N : ℕ
|
||||
N = (Na +N (Nb +N (reallyNa +N reallyNb)))
|
||||
ans : {m : ℕ} {n : ℕ} → N <N m → N <N n → abs (index (apply _*_ a b) m + inverse (index (apply _*_ a b) n)) < e
|
||||
ans {m} {n} N<m N<n rewrite indexAndApply a b _*_ {m} | indexAndApply a b _*_ {n} = ans'
|
||||
ans {m} {n} N<m N<n rewrite indexAndApply a b _*_ {m} | indexAndApply a b _*_ {n} = ans'''
|
||||
where
|
||||
Na<m : Na <N m
|
||||
Na<m = inequalityShrinkLeft N<m
|
||||
@@ -144,10 +145,10 @@ CauchyCompletion.converges (record { elts = a ; converges = aConv } *C record {
|
||||
ans'' with foo
|
||||
... | inl pr = SetoidPartialOrder.<Transitive pOrder (<WellDefined groupIsAbelian groupIsAbelian (<WellDefined (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq bar')) (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq bar)) (orderRespectsAddition pr 0R))) q
|
||||
... | inr pr = <WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq identRight) (+WellDefined (Equivalence.symmetric eq bar') (Equivalence.symmetric eq (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (absWellDefined _ _ (Equivalence.symmetric eq pr)) absZeroIsZero)) timesZero)))) (Equivalence.reflexive eq) 0<e
|
||||
ans' : abs ((index a m * index b m) + inverse (index a n * index b n)) < e
|
||||
ans' with triangleInequality (index a m * (index b m + inverse (index b n))) (index b n * (index a m + inverse (index a n)))
|
||||
ans''' : abs ((index a m * index b m) + inverse (index a n * index b n)) < e
|
||||
ans''' with triangleInequality (index a m * (index b m + inverse (index b n))) (index b n * (index a m + inverse (index a n)))
|
||||
... | inl less = <WellDefined (Equivalence.symmetric eq (absWellDefined ((index a m * index b m) + (inverse (index a n * index b n))) (((index a m) * (index b m + (inverse (index b n)))) + ((index b n) * (index a m + inverse (index a n)))) littleLemma)) (Equivalence.reflexive eq) (SetoidPartialOrder.<Transitive pOrder less (<WellDefined (+WellDefined (Equivalence.symmetric eq (absRespectsTimes (index a m) _)) (Equivalence.symmetric eq (absRespectsTimes (index b n) _))) (Equivalence.reflexive eq) p))
|
||||
... | inr equal rewrite indexAndApply a b _*_ {m} | indexAndApply a b _*_ {n} = <WellDefined (Equivalence.symmetric eq (absWellDefined ((index a m * index b m) + (inverse (index a n * index b n))) (((index a m) * (index b m + (inverse (index b n)))) + ((index b n) * (index a m + inverse (index a n)))) littleLemma)) (Equivalence.reflexive eq) (<WellDefined (Equivalence.symmetric eq equal) (Equivalence.reflexive eq) ((<WellDefined (+WellDefined (Equivalence.symmetric eq (absRespectsTimes (index a m) _)) (Equivalence.symmetric eq (absRespectsTimes (index b n) _))) (Equivalence.reflexive eq) p)))
|
||||
... | inr equal = <WellDefined {_ + _} {abs _} {e} {e} (symmetric (transitive (transitive (absWellDefined ((index a m * index b m) + (inverse (index a n * index b n))) (((index a m) * (index b m + (inverse (index b n)))) + ((index b n) * (index a m + inverse (index a n)))) littleLemma) equal) (+WellDefined (absRespectsTimes (index a m) _) (absRespectsTimes (index b n) _)))) reflexive p
|
||||
|
||||
*CCommutative : {a b : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a *C b) (b *C a)
|
||||
*CCommutative {a} {b} ε 0<e = 0 , ans
|
||||
@@ -162,6 +163,7 @@ abstract
|
||||
multiplicationWellDefinedLeft' : (0!=1 : 0R ∼ 1R → False) (a b c : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid (a *C c) (b *C c)
|
||||
multiplicationWellDefinedLeft' 0!=1 a b c a=b ε 0<e = N , ans
|
||||
where
|
||||
abstract
|
||||
cBoundAndPr : Sg A (λ b → Sg ℕ (λ N → (m : ℕ) → (N <N m) → (abs (index (CauchyCompletion.elts c) m)) < b))
|
||||
cBoundAndPr = boundModulus c
|
||||
cBound : A
|
||||
@@ -197,10 +199,9 @@ abstract
|
||||
cBounded : (m : ℕ) → (N <N m) → abs (index (CauchyCompletion.elts c) m) < cBound
|
||||
cBounded m N<m = cBoundPr m (inequalityShrinkRight N<m)
|
||||
a-bSmall : (m : ℕ) → N <N m → abs ((index (CauchyCompletion.elts a) m) + inverse (index (CauchyCompletion.elts b) m)) < e/c
|
||||
a-bSmall m N<m with abPr {m} (inequalityShrinkLeft N<m)
|
||||
... | f rewrite indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts b)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts b) inverse m) = f
|
||||
a-bSmall m N<m = <WellDefined (absWellDefined _ _ (transitive (identityOfIndiscernablesLeft _∼_ reflexive (equalityCommutative (indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts b)) _+_ {m}))) (+WellDefined reflexive (identityOfIndiscernablesLeft _∼_ reflexive (mapAndIndex (CauchyCompletion.elts b) inverse m))))) reflexive (abPr {m} (inequalityShrinkLeft N<m))
|
||||
ans : {m : ℕ} → N <N m → abs (index (apply _+_ (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts c)) (map inverse (apply _*_ (CauchyCompletion.elts b) (CauchyCompletion.elts c)))) m) < ε
|
||||
ans {m} N<m rewrite indexAndApply (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts c)) (map inverse (apply _*_ (CauchyCompletion.elts b) (CauchyCompletion.elts c))) _+_ {m} | equalityCommutative (mapAndIndex (apply _*_ (CauchyCompletion.elts b) (CauchyCompletion.elts c)) inverse m) | indexAndApply (CauchyCompletion.elts b) (CauchyCompletion.elts c) _*_ {m} | indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts c) _*_ {m} = <WellDefined (absWellDefined _ _ (+WellDefined (Equivalence.reflexive eq) (ringMinusExtracts' R))) (Equivalence.reflexive eq) (<WellDefined (absWellDefined ((index (CauchyCompletion.elts a) m + inverse (index (CauchyCompletion.elts b) m)) * index (CauchyCompletion.elts c) m) _ (Equivalence.transitive eq (Equivalence.transitive eq *Commutative *DistributesOver+) (+WellDefined *Commutative *Commutative))) (Equivalence.reflexive eq) (<WellDefined (Equivalence.symmetric eq (absRespectsTimes _ _)) (Equivalence.reflexive eq) (<WellDefined (Equivalence.reflexive eq) e/cPr (ans' (index (CauchyCompletion.elts a) m) (index (CauchyCompletion.elts b) m) (index (CauchyCompletion.elts c) m) (a-bSmall m N<m) (cBounded m N<m)))))
|
||||
ans {m} N<m = <WellDefined (absWellDefined _ _ (transitive (+WellDefined (identityOfIndiscernablesLeft _∼_ reflexive (indexAndApply _ _ _*_ {m})) (transitive (transitive (ringMinusExtracts' R) (inverseWellDefined additiveGroup (identityOfIndiscernablesRight _∼_ reflexive (equalityCommutative (indexAndApply _ _ _*_ {m}))))) (identityOfIndiscernablesLeft _∼_ reflexive (equalityCommutative (mapAndIndex _ inverse m))))) (identityOfIndiscernablesLeft _∼_ reflexive (indexAndApply _ _ _+_ {m})))) reflexive (<WellDefined (absWellDefined ((index (CauchyCompletion.elts a) m + inverse (index (CauchyCompletion.elts b) m)) * index (CauchyCompletion.elts c) m) _ (Equivalence.transitive eq (Equivalence.transitive eq *Commutative *DistributesOver+) (+WellDefined *Commutative *Commutative))) (Equivalence.reflexive eq) (<WellDefined (Equivalence.symmetric eq (absRespectsTimes _ _)) (Equivalence.reflexive eq) (<WellDefined (Equivalence.reflexive eq) e/cPr (ans' (index (CauchyCompletion.elts a) m) (index (CauchyCompletion.elts b) m) (index (CauchyCompletion.elts c) m) (a-bSmall m N<m) (cBounded m N<m)))))
|
||||
where
|
||||
ans' : (a b c : A) → abs (a + inverse b) < e/c → abs c < cBound → (abs (a + inverse b) * abs c) < (e/c * cBound)
|
||||
ans' a b c a-b<e/c c<bound with totality 0R c
|
||||
|
160
Fields/CauchyCompletion/PartiallyOrderedRing.agda
Normal file
160
Fields/CauchyCompletion/PartiallyOrderedRing.agda
Normal file
@@ -0,0 +1,160 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Fields.Fields
|
||||
open import Fields.Orders.Total.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Semirings.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
|
||||
module Fields.CauchyCompletion.PartiallyOrderedRing {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
|
||||
|
||||
open Setoid S
|
||||
open SetoidTotalOrder (TotallyOrderedRing.total order)
|
||||
open SetoidPartialOrder pOrder
|
||||
open Equivalence eq
|
||||
open PartiallyOrderedRing pRing
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open Field F
|
||||
open import Fields.Lemmas F
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Fields.Orders.Lemmas {F = F} {pRing} (record { oRing = order })
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Addition order F charNot2
|
||||
open import Fields.CauchyCompletion.Multiplication order F charNot2
|
||||
open import Fields.CauchyCompletion.Approximation order F charNot2
|
||||
open import Fields.CauchyCompletion.Ring order F charNot2
|
||||
open import Fields.CauchyCompletion.Comparison order F charNot2
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
|
||||
productPositives : (a b : A) → (injection 0R) <Cr a → (injection 0R) <Cr b → (injection 0R) <Cr (a * b)
|
||||
productPositives a b record { e = eA ; 0<e = 0<eA ; N = Na ; property = prA } record { e = eB ; 0<e = 0<eB ; N = Nb ; property = prB } = record { e = eA * eB ; 0<e = orderRespectsMultiplication 0<eA 0<eB ; N = Na +N Nb ; property = ans }
|
||||
where
|
||||
ans : (m : ℕ) → Na +N Nb <N m → ((eA * eB) + index (CauchyCompletion.elts (injection 0R)) m) < (a * b)
|
||||
ans m <m = <WellDefined (symmetric (transitive (+WellDefined reflexive (identityOfIndiscernablesRight _∼_ reflexive (indexAndConst 0R m))) identRight)) reflexive (ringMultiplyPositives 0<eA 0<eB (<WellDefined (transitive (+WellDefined reflexive (identityOfIndiscernablesRight _∼_ reflexive (indexAndConst 0R m))) identRight) reflexive (prA m (inequalityShrinkLeft <m))) (<WellDefined (transitive (+WellDefined reflexive (identityOfIndiscernablesRight _∼_ reflexive (indexAndConst 0R m))) identRight) reflexive (prB m (inequalityShrinkRight <m))))
|
||||
|
||||
productPositives' : (a b : CauchyCompletion) (interA interB : A) → (0R < interA) → (0R < interB) → (interA r<C a) → (interB r<C b) → (interA * interB) r<C (a *C b)
|
||||
productPositives' a b interA interB 0<iA 0<iB record { e = interA' ; 0<e = 0<interA' ; N = Na ; pr = prA } record { e = interB' ; 0<e = 0<interB' ; N = Nb ; pr = prB } = record { e = interA' * interB' ; 0<e = orderRespectsMultiplication 0<interA' 0<interB' ; N = Na +N Nb ; pr = ans }
|
||||
where
|
||||
ans : (m : ℕ) → (Na +N Nb <N m) → ((interA * interB) + (interA' * interB')) < index (CauchyCompletion.elts (a *C b)) m
|
||||
ans m <m rewrite indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts b) _*_ {m} = <Transitive (<WellDefined identRight (symmetric *DistributesOver+) (<WellDefined reflexive (+WellDefined *Commutative *Commutative) (<WellDefined reflexive (+WellDefined (symmetric *DistributesOver+) (symmetric *DistributesOver+)) (<WellDefined groupIsAbelian (transitive (transitive groupIsAbelian (transitive (symmetric +Associative) (+WellDefined *Commutative (transitive groupIsAbelian (transitive (+WellDefined reflexive *Commutative) (symmetric +Associative)))))) +Associative) (orderRespectsAddition (<WellDefined identRight reflexive (ringAddInequalities (orderRespectsMultiplication 0<iB 0<interA') (orderRespectsMultiplication 0<interB' 0<iA))) ((interA * interB) + (interA' * interB'))))))) (ringMultiplyPositives (<WellDefined identRight reflexive (ringAddInequalities 0<iA 0<interA')) (<WellDefined identRight reflexive (ringAddInequalities 0<iB 0<interB')) (prA m (inequalityShrinkLeft <m)) (prB m (inequalityShrinkRight <m)))
|
||||
|
||||
|
||||
-- a < a'
|
||||
-- b' < b
|
||||
-- then:
|
||||
-- a +C c < a' + c ~ a' + c' < b' + c' ~ b' + c < b +C c
|
||||
{-
|
||||
Have: a <Cr x r<C b
|
||||
|
||||
* Let e = min(distance from a to witness of a<x, distance from x to witness of x<b)
|
||||
* Approximate a above to within e/2
|
||||
* Approximate b below to within e/2
|
||||
* Approximate c (above or below) to within e/2
|
||||
|
||||
Then a' + c' is an appropriate witness.
|
||||
-}
|
||||
|
||||
cOrderRespectsAdditionLeft' : (a : CauchyCompletion) (b : A) (c : A) → (a <Cr b) → (a +C injection c) <C (injection b +C injection c)
|
||||
cOrderRespectsAdditionLeft' a b c record { e = e ; 0<e = 0<e ; N = N ; property = pr } = <CWellDefined {a +C injection c} {a +C injection c} {injection (b + c)} {(injection b) +C (injection c)} (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {a +C injection c}) (GroupHom.groupHom (RingHom.groupHom CInjectionRingHom)) (<CRelaxR (record { e = e ; 0<e = 0<e ; N = N ; property = λ m N<m → <WellDefined (transitive (symmetric +Associative) (+WellDefined reflexive (ans m))) reflexive (orderRespectsAddition (pr m N<m) c) }))
|
||||
where
|
||||
ans : (m : ℕ) → (index (CauchyCompletion.elts a) m + c) ∼ index (apply _+_ (CauchyCompletion.elts a) (constSequence c)) m
|
||||
ans m rewrite indexAndApply (CauchyCompletion.elts a) (constSequence c) _+_ {m} | indexAndConst c m = reflexive
|
||||
|
||||
cOrderRespectsAdditionLeft : (a : CauchyCompletion) (b : A) (c : CauchyCompletion) → (a <Cr b) → (a +C c) <C (injection b +C c)
|
||||
cOrderRespectsAdditionLeft a b c a<b = {!!}
|
||||
|
||||
cOrderRespectsAdditionRight : (a : A) (b : CauchyCompletion) (c : CauchyCompletion) → (a r<C b) → (injection a +C c) <C (b +C c)
|
||||
cOrderRespectsAdditionRight a b a<b = {!!}
|
||||
|
||||
cOrderRespectsAddition : (a b : CauchyCompletion) → (a <C b) → (c : CauchyCompletion) → (a +C c) <C (b +C c)
|
||||
cOrderRespectsAddition a b a<b c = {!!}
|
||||
|
||||
{-
|
||||
cOrderRespectsAddition : (a b : CauchyCompletion) → (a <C b) → (c : CauchyCompletion) → (a +C c) <C (b +C c)
|
||||
cOrderRespectsAddition a b (r , ((r1 , (0<r1 ,, (N1 , prN1))) ,, (r2 , (0<r2 ,, (N2 , prN2))))) c = (a' + c') , (ans1 ,, ans2)
|
||||
where
|
||||
0<min : 0G < min r1 r2
|
||||
0<min with totality r1 r2
|
||||
0<min | inl (inl r1<r2) = 0<r1
|
||||
0<min | inl (inr r2<r1) = 0<r2
|
||||
0<min | inr r1=r2 = 0<r1
|
||||
e/2All : Sg A (λ i → i + i ∼ min r1 r2)
|
||||
e/2All = halve charNot2 (min r1 r2)
|
||||
e/2 : A
|
||||
e/2 = underlying e/2All
|
||||
prE/2 : e/2 + e/2 ∼ min r1 r2
|
||||
prE/2 with e/2All
|
||||
... | _ , pr = pr
|
||||
0<e/2 : 0G < e/2
|
||||
0<e/2 = halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prE/2) 0<min)
|
||||
a'All : Sg A (λ i → (a <Cr i) && (injection i +C (-C a)) <C (injection e/2))
|
||||
a' : A
|
||||
a<a' : a <Cr a'
|
||||
a'Pr : (injection a' +C (-C a)) <C (injection e/2)
|
||||
b'All : Sg A (λ i → (i r<C b) && (b +C (-C injection i)) <C (injection e/2))
|
||||
b' : A
|
||||
b'<b : b' r<C b
|
||||
b'Pr : (b +C (-C injection b')) <C (injection e/2)
|
||||
|
||||
c'All : Sg A (λ i → (c <Cr i) && (injection i +C (-C c)) <C (injection e/2))
|
||||
c' : A
|
||||
c<c' : c <Cr c'
|
||||
c'Pr : (injection c' +C (-C c)) <C (injection e/2)
|
||||
|
||||
-- Now a' + c' is our intervening rational
|
||||
-- and r1 suffices for the witness to a + c < a' + c'
|
||||
-- and r2 suffices for the witness to b' + c' < b + c
|
||||
-- TODO here
|
||||
|
||||
ans1 : (a +C c) <Cr (a' + c')
|
||||
ans1 = r1 , (0<r1 ,, ((N1 +N N2) , ans))
|
||||
where
|
||||
ans : (m : ℕ) → (N1 +N N2) <N m → (r1 + index (CauchyCompletion.elts (a +C c)) m) < (a' + c')
|
||||
ans m N1+N2<m rewrite indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts c) _+_ {m} = <WellDefined (Equivalence.symmetric eq +Associative) reflexive (SetoidPartialOrder.<Transitive pOrder (orderRespectsAddition (prN1 m (inequalityShrinkLeft N1+N2<m)) (index (CauchyCompletion.elts c) m)) {!!})
|
||||
|
||||
ans2 : (a' + c') r<C (b +C c)
|
||||
ans2 = r2 , (0<r2 ,, {!!})
|
||||
|
||||
a'All = approximateAbove a e/2 0<e/2
|
||||
a' = underlying a'All
|
||||
a<a' with a'All
|
||||
... | (_ , (x ,, _)) = x
|
||||
a'Pr with a'All
|
||||
... | (_ , (_ ,, x)) = x
|
||||
b'All = approximateBelow b e/2 0<e/2
|
||||
b' = underlying b'All
|
||||
b'<b with b'All
|
||||
... | (_ , (x ,, _)) = x
|
||||
b'Pr with b'All
|
||||
... | (_ , (_ ,, x)) = x
|
||||
c'All = approximateAbove c e/2 0<e/2
|
||||
c' = underlying c'All
|
||||
c<c' with c'All
|
||||
... | (_ , (x ,, _)) = x
|
||||
c'Pr with c'All
|
||||
... | (_ , (_ ,, x)) = x
|
||||
|
||||
-}
|
||||
|
||||
CpOrderedRing : PartiallyOrderedRing CRing <COrder
|
||||
PartiallyOrderedRing.orderRespectsAddition CpOrderedRing {a} {b} = cOrderRespectsAddition a b
|
||||
PartiallyOrderedRing.orderRespectsMultiplication CpOrderedRing {a} {b} record { i = interA ; a<i = 0<interA ; i<b = interA<a } record { i = interB ; a<i = 0<interB ; i<b = interB<b } = record { i = interA * interB ; a<i = productPositives interA interB 0<interA 0<interB ; i<b = productPositives' a b interA interB (<CCollapsesL 0R _ 0<interA) (<CCollapsesL 0R _ 0<interB) interA<a interB<b }
|
||||
|
@@ -14,6 +14,7 @@ open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Rings.Homomorphisms.Definition
|
||||
|
||||
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
|
||||
|
||||
@@ -32,6 +33,8 @@ open import Fields.CauchyCompletion.Addition order F charNot2
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
open import Fields.CauchyCompletion.Group order F charNot2
|
||||
|
||||
private
|
||||
abstract
|
||||
c*Assoc : {a b c : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a *C (b *C c)) ((a *C b) *C c)
|
||||
c*Assoc {a} {b} {c} ε 0<e = 0 , ans
|
||||
where
|
||||
@@ -59,3 +62,15 @@ Ring.*Associative CRing {a} {b} {c} = c*Assoc {a} {b} {c}
|
||||
Ring.*Commutative CRing {a} {b} = *CCommutative {a} {b}
|
||||
Ring.*DistributesOver+ CRing {a} {b} {c} = *CDistribute {a} {b} {c}
|
||||
Ring.identIsIdent CRing {a} = c*Ident {a}
|
||||
|
||||
private
|
||||
injectionIsRingHom : (a b : A) → Setoid._∼_ cauchyCompletionSetoid (injection (a * b)) (injection a *C injection b)
|
||||
injectionIsRingHom a b ε 0<e = 0 , ans
|
||||
where
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (injection (a * b))) (map inverse (CauchyCompletion.elts (injection a *C injection b)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (constSequence (a * b)) (map inverse (apply _*_ (constSequence a) (constSequence b))) _+_ {m} | indexAndConst (a * b) m | equalityCommutative (mapAndIndex (apply _*_ (constSequence a) (constSequence b)) inverse m) | indexAndApply (constSequence a) (constSequence b) _*_ {m} | indexAndConst a m | indexAndConst b m = <WellDefined (symmetric (transitive (absWellDefined _ _ invRight) absZeroIsZero)) reflexive 0<e
|
||||
|
||||
CInjectionRingHom : RingHom R CRing injection
|
||||
RingHom.preserves1 CInjectionRingHom = Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {injection (Ring.1R R)}
|
||||
RingHom.ringHom CInjectionRingHom {a} {b} = injectionIsRingHom a b
|
||||
RingHom.groupHom CInjectionRingHom = CInjectionGroupHom
|
||||
|
@@ -91,46 +91,6 @@ injectionPreservesSetoid' a b a=b = transferToRight additiveGroup (absZeroImplie
|
||||
... | inl x = x
|
||||
... | inr x = exFalso (absNonnegative x)
|
||||
|
||||
+CCommutative : {a b : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a +C b) (b +C a)
|
||||
+CCommutative {a} {b} ε 0<e = 0 , ans
|
||||
where
|
||||
foo : {x y : A} → (x + y) + inverse (y + x) ∼ 0G
|
||||
foo = Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (inverseWellDefined additiveGroup groupIsAbelian)) invRight
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (a +C b)) (map inverse (CauchyCompletion.elts (b +C a)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (a +C b)) (map inverse (CauchyCompletion.elts (b +C a))) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts b) _+_ {m} | equalityCommutative (mapAndIndex (apply _+_ (CauchyCompletion.elts b) (CauchyCompletion.elts a)) inverse m) | indexAndApply (CauchyCompletion.elts b) (CauchyCompletion.elts a) _+_ {m} = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ foo) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) absZero))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
additionWellDefinedLeft : (a b c : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid (a +C c) (b +C c)
|
||||
additionWellDefinedLeft record { elts = a ; converges = aConv } record { elts = b ; converges = bConv } record { elts = c ; converges = cConv } a=b ε 0<e with a=b ε 0<e
|
||||
... | Na-b , prA-b = Na-b , ans
|
||||
where
|
||||
ans : {m : ℕ} → Na-b <N m → abs (index (apply _+_ (apply _+_ a c) (map inverse (apply _+_ b c))) m) < ε
|
||||
ans {m} mBig with prA-b {m} mBig
|
||||
... | bl rewrite indexAndApply (apply _+_ a c) (map inverse (apply _+_ b c)) _+_ {m} | indexAndApply a c _+_ {m} | equalityCommutative (mapAndIndex (apply _+_ b c) inverse m) | indexAndApply b c _+_ {m} = <WellDefined (absWellDefined _ _ t) (Equivalence.reflexive eq) bl
|
||||
where
|
||||
t : index (apply _+_ a (map inverse b)) m ∼ ((index a m + index c m) + inverse (index b m + index c m))
|
||||
t rewrite indexAndApply a (map inverse b) _+_ {m} | equalityCommutative (mapAndIndex b inverse m) = Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq identLeft) (+WellDefined (Equivalence.symmetric eq invRight) (Equivalence.reflexive eq))) (Equivalence.symmetric eq +Associative)) (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (invContravariant additiveGroup))))) (+Associative {index a m})
|
||||
|
||||
additionPreservedLeft : {a b : A} {c : CauchyCompletion} → (a ∼ b) → Setoid._∼_ cauchyCompletionSetoid (injection a +C c) (injection b +C c)
|
||||
additionPreservedLeft {a} {b} {c} a=b = additionWellDefinedLeft (injection a) (injection b) c (injectionPreservesSetoid a b a=b)
|
||||
|
||||
additionPreservedRight : {a b : A} {c : CauchyCompletion} → (a ∼ b) → Setoid._∼_ cauchyCompletionSetoid (c +C injection a) (c +C injection b)
|
||||
additionPreservedRight {a} {b} {c} a=b = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {c +C injection a} {injection a +C c} {c +C injection b} (+CCommutative {c} {injection a}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection a +C c} {injection b +C c} {c +C injection b} (additionPreservedLeft {a} {b} {c} a=b) (+CCommutative {injection b} {c}))
|
||||
|
||||
additionPreserved : {a b c d : A} → (a ∼ b) → (c ∼ d) → Setoid._∼_ cauchyCompletionSetoid (injection a +C injection c) (injection b +C injection d)
|
||||
additionPreserved {a} {b} {c} {d} a=b c=d = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection a +C injection c} {injection a +C injection d} {injection b +C injection d} (additionPreservedRight {c} {d} {injection a} c=d) (additionPreservedLeft {a} {b} {injection d} a=b)
|
||||
|
||||
additionWellDefinedRight : (a b c : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid b c → Setoid._∼_ cauchyCompletionSetoid (a +C b) (a +C c)
|
||||
additionWellDefinedRight a b c b=c = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {a +C b} {b +C a} {a +C c} (+CCommutative {a} {b}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {b +C a} {c +C a} {a +C c} (additionWellDefinedLeft b c a b=c) (+CCommutative {c} {a}))
|
||||
|
||||
additionWellDefined : {a b c d : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid c d → Setoid._∼_ cauchyCompletionSetoid (a +C c) (b +C d)
|
||||
additionWellDefined {a} {b} {c} {d} a=b c=d = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {a +C c} {a +C d} {b +C d} (additionWellDefinedRight a c d c=d) (additionWellDefinedLeft a b d a=b)
|
||||
|
||||
additionHom : (x y : A) → Setoid._∼_ cauchyCompletionSetoid (injection (x + y)) (injection x +C injection y)
|
||||
additionHom x y ε 0<e = 0 , ans
|
||||
where
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (injection (x + y))) (map inverse (CauchyCompletion.elts (injection x +C injection y)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (injection (x + y))) (map inverse (CauchyCompletion.elts (injection x +C injection y))) _+_ {m} | equalityCommutative (mapAndIndex (apply _+_ (constSequence x) (constSequence y)) inverse m) | indexAndConst (x + y) m | indexAndApply (constSequence x) (constSequence y) _+_ {m} | indexAndConst x m | indexAndConst y m = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ invRight) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) absZero))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
CInjection : SetoidInjection S cauchyCompletionSetoid injection
|
||||
SetoidInjection.wellDefined CInjection {x} {y} x=y = injectionPreservesSetoid x y x=y
|
||||
SetoidInjection.injective CInjection {x} {y} x=y = injectionPreservesSetoid' x y x=y
|
||||
|
@@ -12,11 +12,11 @@ open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
|
||||
module Fields.FieldOfFractions.Order {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} {_<_ : Rel {_} {c} A} {pOrder : SetoidPartialOrder S _<_} {pRing : PartiallyOrderedRing R pOrder} (I : IntegralDomain R) (order : TotallyOrderedRing pRing) where
|
||||
|
||||
open import Fields.FieldOfFractions.Setoid I
|
||||
open import Fields.FieldOfFractions.Ring I
|
||||
open import Fields.FieldOfFractions.Addition I
|
||||
|
||||
open SetoidTotalOrder (TotallyOrderedRing.total order)
|
||||
open import Rings.Orders.Partial.Lemmas
|
||||
@@ -40,6 +40,9 @@ fieldOfFractionsComparison (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , den
|
||||
where
|
||||
open Equivalence (Setoid.eq S)
|
||||
|
||||
private
|
||||
abstract
|
||||
|
||||
fieldOfFractionsOrderWellDefinedLeft : {x y z : fieldOfFractionsSet} → fieldOfFractionsComparison x y → Setoid._∼_ fieldOfFractionsSetoid x z → fieldOfFractionsComparison z y
|
||||
fieldOfFractionsOrderWellDefinedLeft {(numX ,, (denomX , denomX!=0))} {(numY ,, (denomY , denomY!=0))} {(numZ ,, (denomZ , denomZ!=0))} x<y x=z with totality (Ring.0R R) denomZ
|
||||
fieldOfFractionsOrderWellDefinedLeft {numX ,, (denomX , denomX!=0)} {numY ,, (denomY , denomY!=0)} {numZ ,, (denomZ , denomZ!=0)} x<y x=z | inl (inl 0<denomZ) with totality (Ring.0R R) denomY
|
||||
@@ -143,12 +146,8 @@ fieldOfFractionsOrderWellDefinedLeft {numX ,, (denomX , denomX!=0)} {numY ,, (de
|
||||
open Equivalence (Setoid.eq S)
|
||||
fieldOfFractionsOrderWellDefinedLeft {numX ,, (denomX , denomX!=0)} {numY ,, (denomY , denomY!=0)} {numZ ,, (denomZ , denomZ!=0)} x<y x=z | inl (inr denomZ<0) | inl (inr denomY<0) | inl (inr denomX<0) | inr x = exFalso (denomY!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
fieldOfFractionsOrderWellDefinedLeft {numX ,, (denomX , denomX!=0)} {numY ,, (denomY , denomY!=0)} {numZ ,, (denomZ , denomZ!=0)} x<y x=z | inl (inr denomZ<0) | inl (inr denomY<0) | inr x = exFalso (denomX!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
fieldOfFractionsOrderWellDefinedLeft {numX ,, (denomX , denomX!=0)} {numY ,, (denomY , denomY!=0)} {numZ ,, (denomZ , denomZ!=0)} x<y x=z | inl (inr denomZ<0) | inr x = exFalso (denomY!=0 (symmetric x))
|
||||
where
|
||||
open Equivalence (Setoid.eq S)
|
||||
fieldOfFractionsOrderWellDefinedLeft {numX ,, (denomX , denomX!=0)} {numY ,, (denomY , denomY!=0)} {numZ ,, (denomZ , denomZ!=0)} x<y x=z | inr 0=denomZ = exFalso (denomZ!=0 (symmetric 0=denomZ))
|
||||
where
|
||||
open Equivalence (Setoid.eq S)
|
||||
fieldOfFractionsOrderWellDefinedLeft {numX ,, (denomX , denomX!=0)} {numY ,, (denomY , denomY!=0)} {numZ ,, (denomZ , denomZ!=0)} x<y x=z | inl (inr denomZ<0) | inr x = exFalso (denomY!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
fieldOfFractionsOrderWellDefinedLeft {numX ,, (denomX , denomX!=0)} {numY ,, (denomY , denomY!=0)} {numZ ,, (denomZ , denomZ!=0)} x<y x=z | inr 0=denomZ = exFalso (denomZ!=0 (Equivalence.symmetric (Setoid.eq S) 0=denomZ))
|
||||
|
||||
fieldOfFractionsOrderWellDefinedRight : {x y z : fieldOfFractionsSet} → fieldOfFractionsComparison x y → Setoid._∼_ (fieldOfFractionsSetoid) y z → fieldOfFractionsComparison x z
|
||||
fieldOfFractionsOrderWellDefinedRight {numX ,, (denomX , denomX!=0)} {numY ,, (denomY , denomY!=0)} {numZ ,, (denomZ , denomZ!=0)} x<y y=z with totality (Ring.0R R) denomX
|
||||
@@ -205,23 +204,26 @@ swapLemma {S = S} R = transitive (symmetric *Associative) (transitive (*WellDefi
|
||||
open Ring R
|
||||
open Equivalence (Setoid.eq S)
|
||||
|
||||
fieldOfFractionsOrder : SetoidPartialOrder fieldOfFractionsSetoid fieldOfFractionsComparison
|
||||
SetoidPartialOrder.<WellDefined (fieldOfFractionsOrder) {a} {b} {c} {d} a=b c=d a<c = fieldOfFractionsOrderWellDefinedRight {b} {c} {d} (fieldOfFractionsOrderWellDefinedLeft {a} {c} {b} a<c a=b) c=d
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {aNum ,, (aDenom , aDenom!=0)} pr with totality (Ring.0R R) aDenom
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {aNum ,, (aDenom , aDenom!=0)} pr | inl (inl 0<aDenom) with totality (Ring.0R R) aDenom
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {aNum ,, (aDenom , aDenom!=0)} pr | inl (inl 0<aDenom) | inl (inl _) = SetoidPartialOrder.irreflexive pOrder pr
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {aNum ,, (aDenom , aDenom!=0)} pr | inl (inl 0<aDenom) | inl (inr aDenom<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<aDenom aDenom<0))
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {aNum ,, (aDenom , aDenom!=0)} pr | inl (inl 0<aDenom) | inr x = exFalso (aDenom!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {aNum ,, (aDenom , aDenom!=0)} pr | inl (inr aDenom<0) with totality (Ring.0R R) aDenom
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {aNum ,, (aDenom , aDenom!=0)} pr | inl (inr aDenom<0) | inl (inl 0<aDenom) = SetoidPartialOrder.irreflexive pOrder pr
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {aNum ,, (aDenom , aDenom!=0)} pr | inl (inr aDenom<0) | inl (inr _) = SetoidPartialOrder.irreflexive pOrder pr
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {aNum ,, (aDenom , aDenom!=0)} pr | inl (inr aDenom<0) | inr x = exFalso (aDenom!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {aNum ,, (aDenom , aDenom!=0)} pr | inr x = exFalso (aDenom!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c with totality (Ring.0R R) denomA
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) with totality (Ring.0R R) denomC
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) with totality (Ring.0R R) denomB
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inl x) with totality (Ring.0R R) denomC
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inl 0<denomB) | inl (inl _) = ringCanCancelPositive order 0<denomB p
|
||||
private
|
||||
abstract
|
||||
irreflexive : (a : fieldOfFractionsSet) (pr : fieldOfFractionsComparison a a) → False
|
||||
irreflexive (aNum ,, (aDenom , aDenom!=0)) pr with totality (Ring.0R R) aDenom
|
||||
irreflexive (aNum ,, (aDenom , aDenom!=0)) pr | inl (inl 0<aDenom) with totality (Ring.0R R) aDenom
|
||||
irreflexive (aNum ,, (aDenom , aDenom!=0)) pr | inl (inl 0<aDenom) | inl (inl _) = SetoidPartialOrder.irreflexive pOrder pr
|
||||
irreflexive (aNum ,, (aDenom , aDenom!=0)) pr | inl (inl 0<aDenom) | inl (inr aDenom<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<aDenom aDenom<0))
|
||||
irreflexive (aNum ,, (aDenom , aDenom!=0)) pr | inl (inl 0<aDenom) | inr x = exFalso (aDenom!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
irreflexive (aNum ,, (aDenom , aDenom!=0)) pr | inl (inr aDenom<0) with totality (Ring.0R R) aDenom
|
||||
irreflexive (aNum ,, (aDenom , aDenom!=0)) pr | inl (inr aDenom<0) | inl (inl 0<aDenom) = SetoidPartialOrder.irreflexive pOrder pr
|
||||
irreflexive (aNum ,, (aDenom , aDenom!=0)) pr | inl (inr aDenom<0) | inl (inr _) = SetoidPartialOrder.irreflexive pOrder pr
|
||||
irreflexive (aNum ,, (aDenom , aDenom!=0)) pr | inl (inr aDenom<0) | inr x = exFalso (aDenom!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
irreflexive (aNum ,, (aDenom , aDenom!=0)) pr | inr x = exFalso (aDenom!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
|
||||
<transitive : (a b c : fieldOfFractionsSet) (a<b : fieldOfFractionsComparison a b) (b<c : fieldOfFractionsComparison b c) → fieldOfFractionsComparison a c
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c with totality (Ring.0R R) denomA
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) with totality (Ring.0R R) denomC
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) with totality (Ring.0R R) denomB
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inl x) with totality (Ring.0R R) denomC
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inl 0<denomB) | inl (inl _) = ringCanCancelPositive order 0<denomB p
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
@@ -230,29 +232,29 @@ SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA
|
||||
inter = ringCanMultiplyByPositive pRing 0<denomC a<b
|
||||
p : ((numA * denomC) * denomB) < ((numC * denomA) * denomB)
|
||||
p = SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<WellDefined pOrder (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) reflexive inter) (SetoidPartialOrder.<WellDefined pOrder (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) (ringCanMultiplyByPositive pRing 0<denomA b<c))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inl 0<denomB) | inl (inr denomC<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inl 0<denomB) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inr denomB<0) with totality (Ring.0R R) denomC
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inr denomB<0) | inl (inl _) = ringCanCancelNegative order denomB<0 (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<WellDefined pOrder (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) reflexive (ringCanMultiplyByPositive pRing 0<denomA b<c)) (SetoidPartialOrder.<WellDefined pOrder (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) (ringCanMultiplyByPositive pRing 0<denomC a<b)))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inl 0<denomB) | inl (inr denomC<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inl 0<denomB) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inr denomB<0) with totality (Ring.0R R) denomC
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inr denomB<0) | inl (inl _) = ringCanCancelNegative order denomB<0 (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<WellDefined pOrder (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) reflexive (ringCanMultiplyByPositive pRing 0<denomA b<c)) (SetoidPartialOrder.<WellDefined pOrder (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) (ringCanMultiplyByPositive pRing 0<denomC a<b)))
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence (Setoid.eq S)
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inr denomB<0) | inl (inr denomC<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inr denomB<0) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) with totality (Ring.0R R) denomB
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inl (inl 0<denomB) with totality (Ring.0R R) denomC
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inl (inl 0<denomB) | inl (inl 0<denomC) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inl (inl 0<denomB) | inl (inr _) = ringCanCancelPositive order 0<denomB (SetoidPartialOrder.<Transitive pOrder have (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing denomC<0 a<b)))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inr denomB<0) | inl (inr denomC<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inr denomB<0) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) with totality (Ring.0R R) denomB
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inl (inl 0<denomB) with totality (Ring.0R R) denomC
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inl (inl 0<denomB) | inl (inl 0<denomC) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inl (inl 0<denomB) | inl (inr _) = ringCanCancelPositive order 0<denomB (SetoidPartialOrder.<Transitive pOrder have (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing denomC<0 a<b)))
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence (Setoid.eq S)
|
||||
have : ((numC * denomA) * denomB) < ((numB * denomC) * denomA)
|
||||
have = SetoidPartialOrder.<WellDefined pOrder (swapLemma R) reflexive (ringCanMultiplyByPositive pRing 0<denomA b<c)
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inl (inl 0<denomB) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inl (inr denomB<0) with totality (Ring.0R R) denomC
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inl (inl 0<denomB) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inl (inr denomB<0) with totality (Ring.0R R) denomC
|
||||
... | (inl (inl 0<denomC)) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
... | (inl (inr _)) = ringCanCancelNegative order denomB<0 (SetoidPartialOrder.<Transitive pOrder have (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByPositive pRing 0<denomA b<c)))
|
||||
where
|
||||
@@ -262,91 +264,103 @@ SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA
|
||||
have : ((numA * denomC) * denomB) < ((numB * denomA) * denomC)
|
||||
have = SetoidPartialOrder.<WellDefined pOrder (swapLemma R) reflexive (ringCanMultiplyByNegative pRing denomC<0 a<b)
|
||||
... | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inl 0<denomA) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) with totality (Ring.0R R) denomC
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) with totality (Ring.0R R) denomB
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inl 0<denomB) with totality (Ring.0R R) denomC
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inl 0<denomB) | inl (inl _) = ringCanCancelPositive order 0<denomB (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing denomA<0 b<c)) have)
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inl (inr denomC<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inl 0<denomA) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) with totality (Ring.0R R) denomC
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) with totality (Ring.0R R) denomB
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inl 0<denomB) with totality (Ring.0R R) denomC
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inl 0<denomB) | inl (inl _) = ringCanCancelPositive order 0<denomB (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing denomA<0 b<c)) have)
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence (Setoid.eq S)
|
||||
have : ((numB * denomA) * denomC) < ((numA * denomC) * denomB)
|
||||
have = SetoidPartialOrder.<WellDefined pOrder reflexive (swapLemma R) (ringCanMultiplyByPositive pRing 0<denomC a<b)
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inl 0<denomB) | inl (inr denomC<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inl 0<denomB) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inr denomB<0) with totality (Ring.0R R) denomC
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inr denomB<0) | inl (inl _) = ringCanCancelNegative order denomB<0 (SetoidPartialOrder.<Transitive pOrder have (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing denomA<0 b<c)))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inl 0<denomB) | inl (inr denomC<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inl 0<denomB) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inr denomB<0) with totality (Ring.0R R) denomC
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inr denomB<0) | inl (inl _) = ringCanCancelNegative order denomB<0 (SetoidPartialOrder.<Transitive pOrder have (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing denomA<0 b<c)))
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence (Setoid.eq S)
|
||||
have : ((numA * denomC) * denomB) < ((numB * denomA) * denomC)
|
||||
have = SetoidPartialOrder.<WellDefined pOrder (swapLemma R) reflexive (ringCanMultiplyByPositive pRing 0<denomC a<b)
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inr denomB<0) | inl (inr denomC<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inr denomB<0) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) with totality (Ring.0R R) denomB
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inl 0<denomB) with totality (Ring.0R R) denomC
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inl 0<denomB) | inl (inl 0<denomC) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inl 0<denomB) | inl (inr _) = ringCanCancelPositive order 0<denomB (SetoidPartialOrder.<Transitive pOrder have (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing denomA<0 b<c)))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inr denomB<0) | inl (inr denomC<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inl (inr denomB<0) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inl 0<denomC) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) with totality (Ring.0R R) denomB
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inl 0<denomB) with totality (Ring.0R R) denomC
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inl 0<denomB) | inl (inl 0<denomC) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inl 0<denomB) | inl (inr _) = ringCanCancelPositive order 0<denomB (SetoidPartialOrder.<Transitive pOrder have (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing denomA<0 b<c)))
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence (Setoid.eq S)
|
||||
have : ((numA * denomC) * denomB) < ((numB * denomA) * denomC)
|
||||
have = SetoidPartialOrder.<WellDefined pOrder (swapLemma R) reflexive (ringCanMultiplyByNegative pRing denomC<0 a<b)
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inl 0<denomB) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inr denomB<0) with totality (Ring.0R R) denomC
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inr denomB<0) | inl (inl 0<denomC) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inr denomB<0) | inl (inr _) = ringCanCancelNegative order denomB<0 (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing denomA<0 b<c)) have)
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inl 0<denomB) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inr denomB<0) with totality (Ring.0R R) denomC
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inr denomB<0) | inl (inl 0<denomC) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomC denomC<0))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inr denomB<0) | inl (inr _) = ringCanCancelNegative order denomB<0 (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing denomA<0 b<c)) have)
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence (Setoid.eq S)
|
||||
have : ((numB * denomA) * denomC) < ((numA * denomC) * denomB)
|
||||
have = SetoidPartialOrder.<WellDefined pOrder reflexive (swapLemma R) (ringCanMultiplyByNegative pRing denomC<0 a<b)
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inr denomB<0) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inl (inr denomA<0) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} {numC ,, (denomC , denomC!=0)} a<b b<c | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inl (inr denomB<0) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inl (inr denomA<0) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<transitive (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) (numC ,, (denomC , denomC!=0)) a<b b<c | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
|
||||
fieldOfFractionsOrder : SetoidPartialOrder fieldOfFractionsSetoid fieldOfFractionsComparison
|
||||
SetoidPartialOrder.<WellDefined (fieldOfFractionsOrder) {a} {b} {c} {d} a=b c=d a<c = fieldOfFractionsOrderWellDefinedRight {b} {c} {d} (fieldOfFractionsOrderWellDefinedLeft {a} {c} {b} a<c a=b) c=d
|
||||
SetoidPartialOrder.irreflexive (fieldOfFractionsOrder) {a} pr = irreflexive a pr
|
||||
SetoidPartialOrder.<Transitive (fieldOfFractionsOrder) {a} {b} {c} a<b b<c = <transitive a b c a<b b<c
|
||||
|
||||
private
|
||||
<totality : (a b : fieldOfFractionsSet) → ((fieldOfFractionsComparison a b) || (fieldOfFractionsComparison b a)) || (Setoid._∼_ fieldOfFractionsSetoid a b)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) with totality (Ring.0R R) denomA
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) with totality (Ring.0R R) denomB
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) with totality (Ring.0R R) denomA
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inl (inl _) with totality (numA * denomB) (numB * denomA)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inl (inl _) | inl (inl x) = inl (inl x)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inl (inl _) | inl (inr x) = inl (inr x)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inl (inl _) | inr x = inr (Equivalence.transitive (Setoid.eq S) x (Ring.*Commutative R))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inl (inr denomA<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomA denomA<0))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) with totality (Ring.0R R) denomA
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inl (inl _) with totality (numB * denomA) (numA * denomB)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inl (inl _) | inl (inl x) = inl (inl x)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inl (inl _) | inl (inr x) = inl (inr x)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inl (inl _) | inr x = inr (Equivalence.symmetric (Setoid.eq S) (Equivalence.transitive (Setoid.eq S) (Ring.*Commutative R) x))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inl (inr denomA<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomA denomA<0))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) with totality (Ring.0R R) denomB
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) with totality (Ring.0R R) denomA
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inl (inl 0<denomA) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomA denomA<0))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inl (inr _) with totality (numB * denomA) (numA * denomB)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inl (inr _) | inl (inl x) = inl (inl x)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inl (inr _) | inl (inr x) = inl (inr x)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inl (inr _) | inr x = inr (Equivalence.transitive (Setoid.eq S) (Equivalence.symmetric (Setoid.eq S) x) (Ring.*Commutative R))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) with totality (Ring.0R R) denomA
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inl (inl 0<denomA) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomA denomA<0))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inl (inr _) with totality (numA * denomB) (numB * denomA)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inl (inr _) | inl (inl x) = inl (inl x)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inl (inr _) | inl (inr x) = inl (inr x)
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inl (inr _) | inr x = inr (Equivalence.transitive (Setoid.eq S) x (Ring.*Commutative R))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<totality (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
|
||||
fieldOfFractionsTotalOrder : SetoidTotalOrder fieldOfFractionsOrder
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) with totality (Ring.0R R) denomA
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) with totality (Ring.0R R) denomB
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) with totality (Ring.0R R) denomA
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inl (inl _) with totality (numA * denomB) (numB * denomA)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inl (inl _) | inl (inl x) = inl (inl x)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inl (inl _) | inl (inr x) = inl (inr x)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inl (inl _) | inr x = inr (Equivalence.transitive (Setoid.eq S) x (Ring.*Commutative R))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inl (inr denomA<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomA denomA<0))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inl 0<denomB) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) with totality (Ring.0R R) denomA
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inl (inl _) with totality (numB * denomA) (numA * denomB)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inl (inl _) | inl (inl x) = inl (inl x)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inl (inl _) | inl (inr x) = inl (inr x)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inl (inl _) | inr x = inr (Equivalence.symmetric (Setoid.eq S) (Equivalence.transitive (Setoid.eq S) (Ring.*Commutative R) x))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inl (inr denomA<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomA denomA<0))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inl (inr denomB<0) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inl 0<denomA) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) with totality (Ring.0R R) denomB
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) with totality (Ring.0R R) denomA
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inl (inl 0<denomA) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomA denomA<0))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inl (inr _) with totality (numB * denomA) (numA * denomB)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inl (inr _) | inl (inl x) = inl (inl x)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inl (inr _) | inl (inr x) = inl (inr x)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inl (inr _) | inr x = inr (Equivalence.transitive (Setoid.eq S) (Equivalence.symmetric (Setoid.eq S) x) (Ring.*Commutative R))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inl 0<denomB) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) with totality (Ring.0R R) denomA
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inl (inl 0<denomA) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<denomA denomA<0))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inl (inr _) with totality (numA * denomB) (numB * denomA)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inl (inr _) | inl (inl x) = inl (inl x)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inl (inr _) | inl (inr x) = inl (inr x)
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inl (inr _) | inr x = inr (Equivalence.transitive (Setoid.eq S) x (Ring.*Commutative R))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inl (inr denomB<0) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inl (inr denomA<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
SetoidTotalOrder.totality (fieldOfFractionsTotalOrder) x y = <totality x y
|
||||
|
||||
private
|
||||
abstract
|
||||
|
||||
ineqLemma : {x y : A} → (Ring.0R R) < (x * y) → (Ring.0R R) < x → (Ring.0R R) < y
|
||||
ineqLemma {x} {y} 0<xy 0<x with totality (Ring.0R R) y
|
||||
@@ -404,12 +418,13 @@ ineqLemma''' {x} {y} xy<0 x<0 with totality (Ring.0R R) y
|
||||
open Ring R
|
||||
open Equivalence (Setoid.eq S)
|
||||
|
||||
fieldOfFractionsPOrderedRing : PartiallyOrderedRing fieldOfFractionsRing (SetoidTotalOrder.partial fieldOfFractionsTotalOrder)
|
||||
PartiallyOrderedRing.orderRespectsAddition fieldOfFractionsPOrderedRing {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) with totality (Ring.0R R) (denomA * denomC)
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) with totality (Ring.0R R) (denomB * denomC)
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) with totality (Ring.0R R) denomA
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inl 0<dA) with totality (Ring.0R R) denomB
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inl 0<dA) | inl (inl 0<dB) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByPositive pRing 0<dC (SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (symmetric *DistributesOver+) *Commutative)) (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (symmetric *DistributesOver+) *Commutative)) ans))
|
||||
private
|
||||
<orderRespectsAddition : (a b : fieldOfFractionsSet) (a<b : fieldOfFractionsComparison a b) (c : fieldOfFractionsSet) → fieldOfFractionsComparison (fieldOfFractionsPlus a c) (fieldOfFractionsPlus b c)
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) with totality (Ring.0R R) (denomA * denomC)
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) with totality (Ring.0R R) (denomB * denomC)
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) with totality (Ring.0R R) denomA
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inl 0<dA) with totality (Ring.0R R) denomB
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inl 0<dA) | inl (inl 0<dB) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByPositive pRing 0<dC (SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (symmetric *DistributesOver+) *Commutative)) (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (symmetric *DistributesOver+) *Commutative)) ans))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -423,7 +438,7 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
p = SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByPositive pRing 0<dC a<b)
|
||||
ans : ((((numA * denomC) * denomB) + ((denomA * numC) * denomB))) < ((((numB * denomC) * denomA) + ((denomB * numC) * denomA)))
|
||||
ans = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.+WellDefined additiveGroup reflexive (transitive (*WellDefined *Commutative reflexive) (transitive (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) (*WellDefined *Commutative reflexive)))) (PartiallyOrderedRing.orderRespectsAddition pRing p ((denomA * numC) * denomB))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inl 0<dA) | inl (inr dB<0) = exFalso bad
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inl 0<dA) | inl (inr dB<0) = exFalso bad
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -435,9 +450,9 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
... | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
bad : False
|
||||
bad = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<dAdC (SetoidPartialOrder.<WellDefined pOrder reflexive (transitive *Commutative (Ring.timesZero R)) (ringCanMultiplyByNegative pRing dC<0 0<dA)))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inl 0<dA) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inr dA<0) with totality (Ring.0R R) denomB
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inr dA<0) | inl (inl 0<dB) = exFalso bad
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inl 0<dA) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inr dA<0) with totality (Ring.0R R) denomB
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inr dA<0) | inl (inl 0<dB) = exFalso bad
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -454,7 +469,7 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
dC<0 | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
bad : False
|
||||
bad = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<dC dC<0)
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inr dA<0) | inl (inr dB<0) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByNegative pRing dC<0 (SetoidPartialOrder.<WellDefined pOrder (transitive (symmetric *DistributesOver+) *Commutative) (transitive (symmetric *DistributesOver+) *Commutative) have''))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inr dA<0) | inl (inr dB<0) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByNegative pRing dC<0 (SetoidPartialOrder.<WellDefined pOrder (transitive (symmetric *DistributesOver+) *Commutative) (transitive (symmetric *DistributesOver+) *Commutative) have''))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -467,11 +482,11 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
have' = SetoidPartialOrder.<WellDefined pOrder (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)) (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)) have
|
||||
have'' : ((denomA * (numB * denomC)) + (denomA * (denomB * numC))) < ((denomB * (numA * denomC)) + (denomB * (denomA * numC)))
|
||||
have'' = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.+WellDefined additiveGroup reflexive (transitive *Associative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)))) (PartiallyOrderedRing.orderRespectsAddition pRing have' (denomA * (denomB * numC)))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inr dA<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) with totality (Ring.0R R) denomA
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inl 0<dA) with totality (Ring.0R R) denomB
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inl 0<dA) | inl (inl 0<dB) = exFalso bad
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inl (inr dA<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inl 0<dBdC) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) with totality (Ring.0R R) denomA
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inl 0<dA) with totality (Ring.0R R) denomB
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inl 0<dA) | inl (inl 0<dB) = exFalso bad
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -482,7 +497,7 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
dC<0 = ineqLemma'' dBdC<0 0<dB
|
||||
bad : False
|
||||
bad = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<dC dC<0)
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inl 0<dA) | inl (inr dB<0) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByPositive pRing 0<dC ans)
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inl 0<dA) | inl (inr dB<0) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByPositive pRing 0<dC ans)
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -495,9 +510,9 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
have' = PartiallyOrderedRing.orderRespectsAddition pRing (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) have) _
|
||||
ans : (((numB * denomC) + (denomB * numC)) * denomA) < (((numA * denomC) + (denomA * numC)) * denomB)
|
||||
ans = SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (symmetric *DistributesOver+) *Commutative)) (transitive (Group.+WellDefined additiveGroup *Commutative (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative))) (transitive (symmetric *DistributesOver+) *Commutative)) have'
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inl 0<dA) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inr dA<0) with totality (Ring.0R R) denomB
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inr dA<0) | inl (inl 0<dB) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByNegative pRing dC<0 (SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (transitive (Group.+WellDefined additiveGroup reflexive (transitive *Associative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)))) (symmetric *DistributesOver+)) *Commutative)) (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (symmetric *DistributesOver+) *Commutative)) have))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inl 0<dA) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inr dA<0) with totality (Ring.0R R) denomB
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inr dA<0) | inl (inl 0<dB) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByNegative pRing dC<0 (SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (transitive (Group.+WellDefined additiveGroup reflexive (transitive *Associative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)))) (symmetric *DistributesOver+)) *Commutative)) (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (symmetric *DistributesOver+) *Commutative)) have))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -506,7 +521,7 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
dC<0 = ineqLemma'' dBdC<0 0<dB
|
||||
have : (((numA * denomC) * denomB) + ((denomB * numC) * denomA)) < (((numB * denomC) * denomA) + ((denomB * numC) * denomA))
|
||||
have = PartiallyOrderedRing.orderRespectsAddition pRing (SetoidPartialOrder.<WellDefined pOrder (swapLemma R) (swapLemma R) (ringCanMultiplyByNegative pRing dC<0 a<b)) _
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inr dA<0) | inl (inr dB<0) = exFalso bad
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inr dA<0) | inl (inr dB<0) = exFalso bad
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -517,14 +532,14 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
0<dC = ineqLemma''' dBdC<0 dB<0
|
||||
bad : False
|
||||
bad = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<dC dC<0)
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inr dA<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inr 0=dBdC with IntegralDomain.intDom I (Equivalence.symmetric (Setoid.eq S) 0=dBdC)
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inr 0=dBdC | f = exFalso (denomC!=0 (f denomB!=0))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) with totality (Ring.0R R) (denomB * denomC)
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) with totality (Ring.0R R) denomA
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inl 0<dA) with totality (Ring.0R R) denomB
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inl 0<dA) | inl (inl 0<dB) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<dC dC<0))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inl (inr dA<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inl (inr dBdC<0) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inr 0=dBdC with IntegralDomain.intDom I (Equivalence.symmetric (Setoid.eq S) 0=dBdC)
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inl 0<dAdC) | inr 0=dBdC | f = exFalso (denomC!=0 (f denomB!=0))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) with totality (Ring.0R R) (denomB * denomC)
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) with totality (Ring.0R R) denomA
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inl 0<dA) with totality (Ring.0R R) denomB
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inl 0<dA) | inl (inl 0<dB) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<dC dC<0))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -533,7 +548,7 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
0<dC = ineqLemma 0<dBdC 0<dB
|
||||
dC<0 : denomC < 0R
|
||||
dC<0 = ineqLemma'' dAdC<0 0<dA
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inl 0<dA) | inl (inr dB<0) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByNegative pRing dC<0 (SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)) *Commutative) (transitive (symmetric *DistributesOver+) *Commutative)) (transitive (Group.+WellDefined additiveGroup (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)) (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative))) (transitive (symmetric *DistributesOver+) *Commutative)) have))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inl 0<dA) | inl (inr dB<0) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByNegative pRing dC<0 (SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)) *Commutative) (transitive (symmetric *DistributesOver+) *Commutative)) (transitive (Group.+WellDefined additiveGroup (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)) (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative))) (transitive (symmetric *DistributesOver+) *Commutative)) have))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -542,9 +557,9 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
dC<0 = ineqLemma'' dAdC<0 0<dA
|
||||
have : (((numA * denomB) * denomC) + ((denomA * numC) * denomB)) < (((numB * denomA) * denomC) + ((denomA * numC) * denomB))
|
||||
have = PartiallyOrderedRing.orderRespectsAddition pRing (ringCanMultiplyByNegative pRing dC<0 a<b) _
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inl 0<dA) | inr 0=dB = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) 0=dB))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inr dA<0) with totality (Ring.0R R) denomB
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inr dA<0) | inl (inl 0<dB) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByPositive pRing 0<dC (SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (transitive (Group.+WellDefined additiveGroup (transitive *Commutative (transitive (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) *Commutative)) (transitive *Associative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)))) (symmetric *DistributesOver+)) *Commutative)) (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (transitive (Group.+WellDefined additiveGroup (transitive *Commutative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))) reflexive) (symmetric *DistributesOver+)) *Commutative)) have))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inl 0<dA) | inr 0=dB = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) 0=dB))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inr dA<0) with totality (Ring.0R R) denomB
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inr dA<0) | inl (inl 0<dB) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByPositive pRing 0<dC (SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (transitive (Group.+WellDefined additiveGroup (transitive *Commutative (transitive (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) *Commutative)) (transitive *Associative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)))) (symmetric *DistributesOver+)) *Commutative)) (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (transitive (Group.+WellDefined additiveGroup (transitive *Commutative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))) reflexive) (symmetric *DistributesOver+)) *Commutative)) have))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -553,7 +568,7 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
0<dC = ineqLemma 0<dBdC 0<dB
|
||||
have : (((numB * denomA) * denomC) + ((denomA * numC) * denomB)) < (((numA * denomB) * denomC) + ((denomA * numC) * denomB))
|
||||
have = PartiallyOrderedRing.orderRespectsAddition pRing (ringCanMultiplyByPositive pRing 0<dC a<b) _
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inr dA<0) | inl (inr dB<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<dC dC<0))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inr dA<0) | inl (inr dB<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<dC dC<0))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -562,11 +577,11 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
dC<0 = ineqLemma' 0<dBdC dB<0
|
||||
0<dC : 0R < denomC
|
||||
0<dC = ineqLemma''' dAdC<0 dA<0
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inr dA<0) | inr 0=dB = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) 0=dB))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inr 0=dA = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) 0=dA))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) with totality (Ring.0R R) denomA
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inl 0<dA) with totality (Ring.0R R) denomB
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inl 0<dA) | inl (inl 0<dB) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByNegative pRing dC<0 (SetoidPartialOrder.<WellDefined pOrder (transitive (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (Group.+WellDefined additiveGroup (transitive *Commutative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))) reflexive)) (transitive (symmetric *DistributesOver+) *Commutative)) (transitive (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (Group.+WellDefined additiveGroup (transitive (transitive *Associative (*WellDefined *Commutative reflexive)) *Commutative) (transitive *Associative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))))) (transitive (symmetric *DistributesOver+) *Commutative)) have))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inl (inr dA<0) | inr 0=dB = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) 0=dB))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inl 0<dBdC) | inr 0=dA = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) 0=dA))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) with totality (Ring.0R R) denomA
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inl 0<dA) with totality (Ring.0R R) denomB
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inl 0<dA) | inl (inl 0<dB) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByNegative pRing dC<0 (SetoidPartialOrder.<WellDefined pOrder (transitive (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (Group.+WellDefined additiveGroup (transitive *Commutative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))) reflexive)) (transitive (symmetric *DistributesOver+) *Commutative)) (transitive (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (Group.+WellDefined additiveGroup (transitive (transitive *Associative (*WellDefined *Commutative reflexive)) *Commutative) (transitive *Associative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))))) (transitive (symmetric *DistributesOver+) *Commutative)) have))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -575,7 +590,7 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
dC<0 = ineqLemma'' dAdC<0 0<dA
|
||||
have : (((numB * denomA) * denomC) + ((denomB * numC) * denomA)) < (((numA * denomB) * denomC) + ((denomB * numC) * denomA))
|
||||
have = PartiallyOrderedRing.orderRespectsAddition pRing (ringCanMultiplyByNegative pRing dC<0 a<b) _
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inl 0<dA) | inl (inr dB<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder dC<0 0<dC))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inl 0<dA) | inl (inr dB<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder dC<0 0<dC))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -584,9 +599,9 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
dC<0 = ineqLemma'' dAdC<0 0<dA
|
||||
0<dC : 0R < denomC
|
||||
0<dC = ineqLemma''' dBdC<0 dB<0
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inl 0<dA) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inr dA<0) with totality (Ring.0R R) denomB
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inr dA<0) | inl (inl 0<dB) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<dC dC<0))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inl 0<dA) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inr dA<0) with totality (Ring.0R R) denomB
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inr dA<0) | inl (inl 0<dB) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<dC dC<0))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -595,7 +610,7 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
0<dC = ineqLemma''' dAdC<0 dA<0
|
||||
dC<0 : denomC < 0R
|
||||
dC<0 = ineqLemma'' dBdC<0 0<dB
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inr dA<0) | inl (inr dB<0) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByPositive pRing 0<dC (SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (Group.+WellDefined additiveGroup (transitive *Commutative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))) reflexive) (transitive (symmetric *DistributesOver+) *Commutative))) (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (transitive (Group.+WellDefined additiveGroup (transitive *Commutative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))) (transitive *Commutative (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative)))) (symmetric *DistributesOver+)) *Commutative)) have))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inr dA<0) | inl (inr dB<0) = SetoidPartialOrder.<WellDefined pOrder (symmetric *Associative) (symmetric *Associative) (ringCanMultiplyByPositive pRing 0<dC (SetoidPartialOrder.<WellDefined pOrder (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (Group.+WellDefined additiveGroup (transitive *Commutative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))) reflexive) (transitive (symmetric *DistributesOver+) *Commutative))) (transitive (Group.+WellDefined additiveGroup *Commutative *Commutative) (transitive (transitive (Group.+WellDefined additiveGroup (transitive *Commutative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))) (transitive *Commutative (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative)))) (symmetric *DistributesOver+)) *Commutative)) have))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
@@ -604,12 +619,15 @@ PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA
|
||||
0<dC = ineqLemma''' dAdC<0 dA<0
|
||||
have : (((numA * denomB) * denomC) + ((denomA * numC) * denomB)) < (((numB * denomA) * denomC) + ((denomA * numC) * denomB))
|
||||
have = PartiallyOrderedRing.orderRespectsAddition pRing (ringCanMultiplyByPositive pRing 0<dC a<b) _
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inr dA<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inr 0=dBdC with IntegralDomain.intDom I (Equivalence.symmetric (Setoid.eq S) 0=dBdC)
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inl (inr dA<0) | inr x = exFalso (denomB!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inl (inr dBdC<0) | inr x = exFalso (denomA!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inl (inr dAdC<0) | inr 0=dBdC with IntegralDomain.intDom I (Equivalence.symmetric (Setoid.eq S) 0=dBdC)
|
||||
... | f = exFalso (denomC!=0 (f denomB!=0))
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inr (0=dAdC) with IntegralDomain.intDom I (Equivalence.symmetric (Setoid.eq S) 0=dAdC)
|
||||
PartiallyOrderedRing.orderRespectsAddition (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} a<b (numC ,, (denomC , denomC!=0)) | inr 0=dAdC | f = exFalso (denomC!=0 (f denomA!=0))
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inr (0=dAdC) with IntegralDomain.intDom I (Equivalence.symmetric (Setoid.eq S) 0=dAdC)
|
||||
<orderRespectsAddition (numA ,, (denomA , denomA!=0)) (numB ,, (denomB , denomB!=0)) a<b (numC ,, (denomC , denomC!=0)) | inr 0=dAdC | f = exFalso (denomC!=0 (f denomA!=0))
|
||||
|
||||
fieldOfFractionsPOrderedRing : PartiallyOrderedRing fieldOfFractionsRing (SetoidTotalOrder.partial fieldOfFractionsTotalOrder)
|
||||
PartiallyOrderedRing.orderRespectsAddition fieldOfFractionsPOrderedRing {a} {b} a<b c = <orderRespectsAddition a b a<b c
|
||||
PartiallyOrderedRing.orderRespectsMultiplication (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} t u with totality (Ring.0R R) (Ring.1R R)
|
||||
PartiallyOrderedRing.orderRespectsMultiplication (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} 0<a 0<b | inl (inl 0<1) with totality (Ring.0R R) (denomA * denomB)
|
||||
PartiallyOrderedRing.orderRespectsMultiplication (fieldOfFractionsPOrderedRing) {numA ,, (denomA , denomA!=0)} {numB ,, (denomB , denomB!=0)} 0<a 0<b | inl (inl 0<1) | inl (inl 0<dAdB) with totality (Ring.0R R) denomB
|
||||
|
@@ -32,6 +32,12 @@ record _&&_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where
|
||||
fst : A
|
||||
snd : B
|
||||
|
||||
equalityLeft : {a b : _} {A : Set a} {B : Set b} {x1 x2 : A && B} → (x1 ≡ x2) → _&&_.fst x1 ≡ _&&_.fst x2
|
||||
equalityLeft {x1 = a ,, _} {.a ,, _} refl = refl
|
||||
|
||||
equalityRight : {a b : _} {A : Set a} {B : Set b} {x1 x2 : A && B} → (x1 ≡ x2) → _&&_.snd x1 ≡ _&&_.snd x2
|
||||
equalityRight {x1 = _ ,, a} {_ ,, .a} refl = refl
|
||||
|
||||
infix 15 _&_&_
|
||||
record _&_&_ {a b c} (A : Set a) (B : Set b) (C : Set c) : Set (a ⊔ b ⊔ c) where
|
||||
field
|
||||
@@ -77,9 +83,9 @@ orIsAssociative (inl (inl x)) = inl x
|
||||
orIsAssociative (inl (inr x)) = inr (inl x)
|
||||
orIsAssociative (inr x) = inr (inr x)
|
||||
|
||||
leqConstructive : {n : _} → {p : Set n} → (p || (p → False)) → (((p → False) → False) → p)
|
||||
leqConstructive (inl p) = λ _ → p
|
||||
leqConstructive (inr notP) = λ negnegP → exFalso (negnegP notP)
|
||||
lemConstructive : {n : _} → {p : Set n} → (p || (p → False)) → (((p → False) → False) → p)
|
||||
lemConstructive (inl p) = λ _ → p
|
||||
lemConstructive (inr notP) = λ negnegP → exFalso (negnegP notP)
|
||||
|
||||
lemConverse : {n : _} → {p : Set n} → p → ((p → False) → False)
|
||||
lemConverse p = λ notP → notP p
|
||||
|
@@ -1,4 +1,4 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
@@ -10,7 +10,6 @@ open import Rings.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Fields.Fields
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
@@ -28,6 +27,9 @@ open import Fields.FieldOfFractions.Order ℤIntDom ℤOrderedRing
|
||||
ℚ : Set
|
||||
ℚ = fieldOfFractionsSet
|
||||
|
||||
ℚSetoid : Setoid ℚ
|
||||
ℚSetoid = fieldOfFractionsSetoid
|
||||
|
||||
_+Q_ : ℚ → ℚ → ℚ
|
||||
a +Q b = fieldOfFractionsPlus a b
|
||||
|
||||
@@ -43,6 +45,9 @@ a *Q b = fieldOfFractionsTimes a b
|
||||
injectionQ : ℤ → ℚ
|
||||
injectionQ z = z ,, (nonneg 1 , λ ())
|
||||
|
||||
injectionQInjective : Injection injectionQ
|
||||
injectionQInjective pr = equalityLeft pr
|
||||
|
||||
ℚField : Field ℚRing
|
||||
ℚField = fieldOfFractions
|
||||
|
||||
|
@@ -1,9 +1,11 @@
|
||||
{-# OPTIONS --warning=error --safe --guardedness #-}
|
||||
{-# OPTIONS --warning=error --safe --guardedness --without-K #-}
|
||||
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Rings.Definition
|
||||
open import Numbers.Rationals.Definition
|
||||
open import Functions
|
||||
|
||||
module Numbers.Reals.Definition where
|
||||
|
||||
@@ -23,12 +25,20 @@ _+R_ = _+C_
|
||||
_*R_ : ℝ → ℝ → ℝ
|
||||
_*R_ = _*C_
|
||||
|
||||
ℝSetoid = cauchyCompletionSetoid
|
||||
|
||||
_=R_ : ℝ → ℝ → Set
|
||||
a =R b = Setoid._∼_ cauchyCompletionSetoid a b
|
||||
|
||||
ℝRing : Ring cauchyCompletionSetoid _+R_ _*R_
|
||||
ℝRing = CRing
|
||||
|
||||
injectionR : ℚ → ℝ
|
||||
injectionR = injection
|
||||
|
||||
injectionRInjective : Injection injectionR
|
||||
injectionRInjective = CInjection'
|
||||
|
||||
0R : ℝ
|
||||
0R = injection 0Q
|
||||
|
||||
|
@@ -1,42 +0,0 @@
|
||||
{-# OPTIONS --warning=error --safe #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Functions
|
||||
open import Maybe
|
||||
open import Orders
|
||||
|
||||
open import Numbers.Naturals.Naturals
|
||||
|
||||
module RedBlackTree where
|
||||
record BinaryTreeNode {a : _} (carrier : Set a) (order : TotalOrder carrier) (minValue : Maybe carrier) (maxValue : Maybe carrier) : Set a
|
||||
|
||||
valueExtractor : {a : _} {carrier : Set a} {order : TotalOrder carrier} {minValue : Maybe carrier} {maxValue : Maybe carrier} → BinaryTreeNode {a} carrier order minValue maxValue → carrier
|
||||
|
||||
record BinaryTreeNode {a} carrier order minValue maxValue where
|
||||
inductive
|
||||
field
|
||||
value : carrier
|
||||
min<=val : TotalOrder._<_ order minValue maxValue
|
||||
leftChild : Maybe (Sg (BinaryTreeNode {a} carrier order minValue (yes value)) (λ i → TotalOrder._<_ order (valueExtractor {a} {carrier} {order} i) value))
|
||||
rightChild : Maybe (Sg (BinaryTreeNode {a} carrier order (yes value) maxValue) (λ i → TotalOrder._<_ order value (valueExtractor i)))
|
||||
|
||||
valueExtractor t = BinaryTreeNode.value t
|
||||
|
||||
lookup : {a : _} {carrier : Set a} {order : TotalOrder carrier} {minValue : Maybe carrier} {maxValue : Maybe carrier} → (t : BinaryTreeNode carrier order minValue maxValue) → (k : carrier) → Maybe carrier
|
||||
lookup {a} {carrier} {order} record { value = value ; leftChild = leftChild ; rightChild = rightChild } k with TotalOrder.totality order k value
|
||||
lookup {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = rightChild } k | inl (inl k<val) = no
|
||||
lookup {a} {carrier} {order} record { value = value ; leftChild = (yes (tree , _)) ; rightChild = rightChild } k | inl (inl k<val) = lookup tree k
|
||||
lookup {a} {carrier} {order} record { value = value ; leftChild = leftChild ; rightChild = no } k | inl (inr val<k) = no
|
||||
lookup {a} {carrier} {order} record { value = value ; leftChild = leftChild ; rightChild = (yes (tree , _)) } k | inl (inr val<k) = lookup tree k
|
||||
lookup {a} {carrier} {order} record { value = value ; leftChild = leftChild ; rightChild = rightChild } k | inr k=val = yes value
|
||||
|
||||
addTree : {a : _} {carrier : Set a} {order : TotalOrder carrier} {minValue : Maybe carrier} {maxValue : Maybe carrier} → (t : BinaryTreeNode carrier order minValue maxValue) → (k : carrier) → BinaryTreeNode carrier order (yes (defaultValue k minValue)) (yes (defaultValue k maxValue))
|
||||
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = no } k with TotalOrder.totality order k value
|
||||
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = no } k | inl (inr val<k) = record { value = value ; leftChild = no ; rightChild = yes (record { value = k ; leftChild = no ; rightChild = no} , val<k)}
|
||||
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = no } k | inl (inl k<val) = record { value = value ; leftChild = yes (record { value = k ; leftChild = no ; rightChild = no} , k<val) ; rightChild = no }
|
||||
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = no } k | inr k=val = record { value = k ; leftChild = no ; rightChild = no }
|
||||
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = (yes x) } k with TotalOrder.totality order k value
|
||||
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = (yes child) } k | inl (inl k<val) = {!!}
|
||||
addTree {a} {carrier} {order} record { value = value ; leftChild = no ; rightChild = (yes child) } k | inl (inr val<k) = {!!}
|
||||
addTree {a} {carrier} {order} {minValue} {maxValue} record { value = value ; leftChild = no ; rightChild = (yes (child , pr)) } k | inr k=val = record { value = k ; leftChild = no ; rightChild = yes (typeCast child (applyEquality (λ i → BinaryTreeNode carrier order {!!} {!!}) {!!}) , {!!}) }
|
||||
addTree {a} {carrier} {order} record { value = value ; leftChild = (yes x) ; rightChild = rightChild } k = {!!}
|
@@ -293,6 +293,9 @@ abstract
|
||||
halvePositive a 0<2a | inl (inr a<0) = exFalso (irreflexive {a + a} (SetoidPartialOrder.<Transitive pOrder (<WellDefined (Equivalence.reflexive eq) identRight (ringAddInequalities a<0 a<0)) 0<2a))
|
||||
halvePositive a 0<2a | inr x = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq x) (Equivalence.symmetric eq x)) identRight) 0<2a))
|
||||
|
||||
halvePositive' : {a b : A} → (a + a) ∼ b → 0R < b → 0R < a
|
||||
halvePositive' {a} {b} pr 0<b = halvePositive a (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq pr) 0<b)
|
||||
|
||||
0<1 : (0R ∼ 1R → False) → 0R < 1R
|
||||
0<1 0!=1 with totality 0R 1R
|
||||
0<1 0!=1 | inl (inl x) = x
|
||||
|
@@ -13,6 +13,9 @@ record Sequence {a : _} (A : Set a) : Set a where
|
||||
head : A
|
||||
tail : Sequence A
|
||||
|
||||
headInjective : {a : _} {A : Set a} {s1 s2 : Sequence A} → s1 ≡ s2 → Sequence.head s1 ≡ Sequence.head s2
|
||||
headInjective {s1 = s1} {.s1} refl = refl
|
||||
|
||||
constSequence : {a : _} {A : Set a} (k : A) → Sequence A
|
||||
Sequence.head (constSequence k) = k
|
||||
Sequence.tail (constSequence k) = constSequence k
|
||||
|
@@ -1,4 +1,4 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
|
||||
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
@@ -8,6 +8,11 @@ open import Sets.FinSet.Definition
|
||||
open import Sets.FinSet.Lemmas
|
||||
open import Sets.Cardinality.Infinite.Definition
|
||||
open import Sets.Cardinality.Finite.Lemmas
|
||||
open import Numbers.Reals.Definition
|
||||
open import Numbers.Rationals.Definition
|
||||
open import Numbers.Integers.Definition
|
||||
open import Sets.Cardinality.Infinite.Lemmas
|
||||
open import Setoids.Setoids
|
||||
|
||||
module Sets.Cardinality.Infinite.Examples where
|
||||
|
||||
@@ -28,3 +33,7 @@ module Sets.Cardinality.Infinite.Examples where
|
||||
bad a = (inv (toNat a))
|
||||
badInj : Injection bad
|
||||
badInj = injComp nextInj invInj
|
||||
|
||||
ℝIsInfinite : DedekindInfiniteSet ℝ
|
||||
DedekindInfiniteSet.inj ℝIsInfinite n = injectionR (injectionQ (nonneg n))
|
||||
DedekindInfiniteSet.isInjection ℝIsInfinite {x} {y} pr = nonnegInjective (injectionQInjective (injectionRInjective pr))
|
||||
|
Reference in New Issue
Block a user