mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-13 15:48:39 +00:00
Lots of speedups (#116)
This commit is contained in:
@@ -16,6 +16,7 @@ open import Numbers.Modulo.Group
|
||||
|
||||
open import Numbers.Integers.Integers
|
||||
open import Numbers.Integers.RingStructure.EuclideanDomain
|
||||
open import Numbers.Integers.RingStructure.Archimedean
|
||||
|
||||
open import Numbers.ClassicalReals.Examples
|
||||
open import Numbers.ClassicalReals.RealField.Lemmas
|
||||
@@ -52,12 +53,15 @@ open import Groups.LectureNotes.Lecture1
|
||||
open import Fields.Fields
|
||||
open import Fields.Orders.Partial.Definition
|
||||
open import Fields.Orders.Total.Definition
|
||||
open import Fields.Orders.Total.Archimedean
|
||||
open import Fields.Orders.LeastUpperBounds.Examples
|
||||
open import Fields.Orders.Lemmas
|
||||
open import Fields.Orders.Limits.Definition
|
||||
open import Fields.FieldOfFractions.Field
|
||||
open import Fields.FieldOfFractions.Lemmas
|
||||
open import Fields.FieldOfFractions.Order
|
||||
open import Fields.FieldOfFractions.Order
|
||||
--open import Fields.FieldOfFractions.Archimedean
|
||||
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
@@ -100,7 +104,7 @@ open import Rings.Orders.Total.BaseExpansion
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.DirectSum
|
||||
open import Setoids.Lists
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Total.Lemmas
|
||||
open import Setoids.Functions.Definition
|
||||
open import Setoids.Functions.Extension
|
||||
open import Setoids.Algebra.Lemmas
|
||||
@@ -159,8 +163,7 @@ open import Modules.PolynomialModule
|
||||
open import Modules.Lemmas
|
||||
open import Modules.DirectSum
|
||||
|
||||
open import Fields.CauchyCompletion.Ring
|
||||
open import Fields.CauchyCompletion.Comparison
|
||||
open import Fields.CauchyCompletion.Archimedean
|
||||
open import Fields.Orders.Limits.Lemmas
|
||||
|
||||
open import ProjectEuler.Problem1
|
||||
|
@@ -9,7 +9,8 @@ open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
@@ -31,7 +32,9 @@ open import Fields.Lemmas F
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Fields.Orders.Total.Lemmas {F = F} (record { oRing = order })
|
||||
open import Rings.Orders.Total.AbsoluteValue order
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Rings.Orders.Total.Cauchy order
|
||||
open import Rings.InitialRing R
|
||||
|
||||
private
|
||||
@@ -39,12 +42,12 @@ private
|
||||
lemm zero a b = refl
|
||||
lemm (succ m) a b = lemm m (Sequence.tail a) (Sequence.tail b)
|
||||
|
||||
_+C_ : CauchyCompletion → CauchyCompletion → CauchyCompletion
|
||||
CauchyCompletion.elts (record { elts = a ; converges = convA } +C record { elts = b ; converges = convB }) = apply _+_ a b
|
||||
CauchyCompletion.converges (record { elts = a ; converges = convA } +C record { elts = b ; converges = convB }) ε 0<e with halve charNot2 ε
|
||||
... | e/2 , e/2Pr with convA e/2 (halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq e/2Pr) 0<e))
|
||||
CauchyCompletion.converges (record { elts = a ; converges = convA } +C record { elts = b ; converges = convB }) ε 0<e | e/2 , e/2Pr | Na , prA with convB e/2 (halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq e/2Pr) 0<e))
|
||||
CauchyCompletion.converges (record { elts = a ; converges = convA } +C record { elts = b ; converges = convB }) ε 0<e | e/2 , e/2Pr | Na , prA | Nb , prB = (Na +N Nb) , t
|
||||
private
|
||||
additionConverges : (a b : CauchyCompletion) → cauchy (apply _+_ (CauchyCompletion.elts a) (CauchyCompletion.elts b))
|
||||
additionConverges record { elts = a ; converges = convA } record { elts = b ; converges = convB } ε 0<e with halve charNot2 ε
|
||||
... | e/2 , e/2Pr with convA e/2 (halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq e/2Pr) 0<e))
|
||||
... | Na , prA with convB e/2 (halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq e/2Pr) 0<e))
|
||||
... | Nb , prB = (Na +N Nb) , t
|
||||
where
|
||||
t : {m n : ℕ} → Na +N Nb <N m → Na +N Nb <N n → abs ((index (apply _+_ a b) m) + inverse (index (apply _+_ a b) n)) < ε
|
||||
t {m} {n} <m <n with prA {m} {n} (inequalityShrinkLeft <m) (inequalityShrinkLeft <n)
|
||||
@@ -53,6 +56,10 @@ CauchyCompletion.converges (record { elts = a ; converges = convA } +C record {
|
||||
... | inl tri rewrite lemm m a b | lemm n a b = SetoidPartialOrder.<WellDefined pOrder (Equivalence.reflexive eq) e/2Pr (SetoidPartialOrder.<Transitive pOrder {_} {(abs ((index a m) + (inverse (index a n)))) + (abs ((index b m) + (inverse (index b n))))} (<WellDefined (absWellDefined _ _ (Equivalence.transitive eq (Equivalence.symmetric eq (+Associative {index a m})) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq {index a m}) (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq (+Associative {index b m})) (+WellDefined (Equivalence.reflexive eq {index b m}) (Equivalence.symmetric eq (invContravariant additiveGroup)))))) (+Associative {index a m})))) (Equivalence.reflexive eq) tri) (ringAddInequalities am-an<e/2 bm-bn<e/2))
|
||||
... | inr tri rewrite lemm m a b | lemm n a b = SetoidPartialOrder.<WellDefined pOrder (Equivalence.reflexive eq) e/2Pr (<WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq tri) (absWellDefined _ _ (Equivalence.transitive eq (Equivalence.symmetric eq (+Associative {index a m})) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq {index a m}) (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq (+Associative {index b m})) (+WellDefined (Equivalence.reflexive eq {index b m}) (Equivalence.symmetric eq (invContravariant additiveGroup)))))) (+Associative {index a m}))))) (Equivalence.reflexive eq) (ringAddInequalities am-an<e/2 bm-bn<e/2))
|
||||
|
||||
_+C_ : CauchyCompletion → CauchyCompletion → CauchyCompletion
|
||||
CauchyCompletion.elts (record { elts = a ; converges = convA } +C record { elts = b ; converges = convB }) = apply _+_ a b
|
||||
CauchyCompletion.converges (a +C b) = additionConverges a b
|
||||
|
||||
inverseDistributes : {r s : A} → inverse (r + s) ∼ inverse r + inverse s
|
||||
inverseDistributes = Equivalence.transitive eq (invContravariant additiveGroup) groupIsAbelian
|
||||
|
||||
|
@@ -10,7 +10,8 @@ open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
@@ -32,6 +33,7 @@ open Field F
|
||||
open import Fields.Lemmas F
|
||||
open import Fields.Orders.Lemmas {F = F} record { oRing = order }
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Rings.Orders.Total.AbsoluteValue order
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Addition order F
|
||||
|
28
Fields/CauchyCompletion/Archimedean.agda
Normal file
28
Fields/CauchyCompletion/Archimedean.agda
Normal file
@@ -0,0 +1,28 @@
|
||||
{-# 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.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Orders.Archimedean
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Fields.Orders.Total.Archimedean
|
||||
|
||||
module Fields.CauchyCompletion.Archimedean {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) (arch : ArchimedeanField {F = F} record { oRing = order }) where
|
||||
|
||||
open import Fields.CauchyCompletion.Group order F
|
||||
open import Fields.CauchyCompletion.Ring order F
|
||||
open import Fields.CauchyCompletion.Comparison order F
|
||||
--open import Fields.CauchyCompletion.PartiallyOrderedRing order F
|
||||
|
||||
--CArchimedean : Archimedean (toGroup CRing CpOrderedRing)
|
||||
--CArchimedean = ?
|
@@ -10,7 +10,8 @@ open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
@@ -32,9 +33,11 @@ open import Fields.Orders.Lemmas {F = F} {pRing} (record { oRing = order })
|
||||
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Rings.Orders.Total.AbsoluteValue order
|
||||
open import Fields.Lemmas F
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Setoid order F
|
||||
open import Fields.CauchyCompletion.Group order F
|
||||
open import Fields.Orders.Total.Lemmas {F = F} (record { oRing = order })
|
||||
|
||||
-- "<C rational", where the r indicates where the rational number goes
|
||||
@@ -48,11 +51,12 @@ record _<Cr_ (a : CauchyCompletion) (b : A) : Set (m ⊔ o) where
|
||||
<CrWellDefinedRight : (a : CauchyCompletion) (b c : A) → b ∼ c → a <Cr b → a <Cr c
|
||||
<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)
|
||||
<CrWellDefinedLemma a b e/2 e 0<e pr a-b<e | inl (inl 0<a-b) = ringAddInequalities (halfLess e/2 e 0<e pr) (<WellDefined identLeft (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) invLeft) identRight)) (orderRespectsAddition 0<a-b b))
|
||||
<CrWellDefinedLemma a b e/2 e 0<e pr a-b<e | inl (inr a-b<0) = <WellDefined (Equivalence.transitive eq (+WellDefined (invContravariant (Ring.additiveGroup R)) groupIsAbelian) (Equivalence.transitive eq (Equivalence.transitive eq +Associative (+WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (invTwice additiveGroup b) invLeft) identRight)) (Equivalence.reflexive eq))) groupIsAbelian)) (Equivalence.transitive eq +Associative (+WellDefined pr (Equivalence.reflexive eq))) (orderRespectsAddition a-b<e (e/2 + a))
|
||||
<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)
|
||||
private
|
||||
<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)
|
||||
<CrWellDefinedLemma a b e/2 e 0<e pr a-b<e | inl (inl 0<a-b) = ringAddInequalities (halfLess e/2 e 0<e pr) (<WellDefined identLeft (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) invLeft) identRight)) (orderRespectsAddition 0<a-b b))
|
||||
<CrWellDefinedLemma a b e/2 e 0<e pr a-b<e | inl (inr a-b<0) = <WellDefined (Equivalence.transitive eq (+WellDefined (invContravariant (Ring.additiveGroup R)) groupIsAbelian) (Equivalence.transitive eq (Equivalence.transitive eq +Associative (+WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (invTwice additiveGroup b) invLeft) identRight)) (Equivalence.reflexive eq))) groupIsAbelian)) (Equivalence.transitive eq +Associative (+WellDefined pr (Equivalence.reflexive eq))) (orderRespectsAddition a-b<e (e/2 + a))
|
||||
<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 record { e = ε ; 0<e = 0<e ; N = N ; property = pr } with halve charNot2 ε
|
||||
@@ -77,11 +81,12 @@ record _r<C_ (a : A) (b : CauchyCompletion) : Set (m ⊔ o) where
|
||||
r<CWellDefinedLeft : (a b : A) (c : CauchyCompletion) → (a ∼ b) → (a r<C c) → b r<C c
|
||||
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)
|
||||
r<CWellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a | inl (inl 0<a-b) = SetoidPartialOrder.<Transitive pOrder (<WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq prE/2) (Equivalence.reflexive eq)) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) invRight)) identRight)))) (Equivalence.reflexive eq) (orderRespectsAddition c+e<a (inverse e/2))) (<WellDefined (Equivalence.transitive eq +Associative (+WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) invLeft) identRight)) (Equivalence.reflexive eq))) (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (invLeft)) identRight))) (orderRespectsAddition pr (b + inverse e/2)))
|
||||
r<CWellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a | inl (inr a-b<0) = SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (halfLess e/2 e 0<e prE/2) c)) c+e<a) (<WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) invLeft)) identRight) identLeft (orderRespectsAddition a-b<0 b))
|
||||
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)
|
||||
private
|
||||
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)
|
||||
r<CWellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a | inl (inl 0<a-b) = SetoidPartialOrder.<Transitive pOrder (<WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq prE/2) (Equivalence.reflexive eq)) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) invRight)) identRight)))) (Equivalence.reflexive eq) (orderRespectsAddition c+e<a (inverse e/2))) (<WellDefined (Equivalence.transitive eq +Associative (+WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) invLeft) identRight)) (Equivalence.reflexive eq))) (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (invLeft)) identRight))) (orderRespectsAddition pr (b + inverse e/2)))
|
||||
r<CWellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a | inl (inr a-b<0) = SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<Transitive pOrder (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (halfLess e/2 e 0<e prE/2) c)) c+e<a) (<WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) invLeft)) identRight) identLeft (orderRespectsAddition a-b<0 b))
|
||||
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 record { e = e ; 0<e = 0<e ; N = N ; pr = pr } with halve charNot2 e
|
||||
@@ -142,6 +147,26 @@ record _<C_ (a : CauchyCompletion) (b : CauchyCompletion) : Set (m ⊔ o) where
|
||||
<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) }
|
||||
|
||||
flip<CR : {a : CauchyCompletion} {b : A} → a <Cr b → inverse b r<C Group.inverse CGroup a
|
||||
_r<C_.e (flip<CR {a} {b} record { e = e ; 0<e = 0<e ; N = N ; property = property }) = e
|
||||
_r<C_.0<e (flip<CR {a} {b} record { e = e ; 0<e = 0<e ; N = N ; property = property }) = 0<e
|
||||
_r<C_.N (flip<CR {a} {b} record { e = e ; 0<e = 0<e ; N = N ; property = property }) = N
|
||||
_r<C_.pr (flip<CR {a} {b} record { e = e ; 0<e = 0<e ; N = N ; property = property }) m N<m with property m N<m
|
||||
... | bl = <WellDefined (transitive (transitive (symmetric +Associative) (+WellDefined reflexive (transitive groupIsAbelian (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invLeft) identRight))))) groupIsAbelian) (transitive +Associative (transitive (+WellDefined invRight reflexive) (transitive identLeft (identityOfIndiscernablesRight _∼_ reflexive (mapAndIndex (CauchyCompletion.elts a) inverse m))))) (orderRespectsAddition bl (inverse b + inverse (index (CauchyCompletion.elts a) m)))
|
||||
|
||||
flipR<C : {a : A} {b : CauchyCompletion} → a r<C b → Group.inverse CGroup b <Cr inverse a
|
||||
_<Cr_.e (flipR<C record { e = e ; 0<e = 0<e ; N = N ; pr = pr }) = e
|
||||
_<Cr_.0<e (flipR<C record { e = e ; 0<e = 0<e ; N = N ; pr = pr }) = 0<e
|
||||
_<Cr_.N (flipR<C record { e = e ; 0<e = 0<e ; N = N ; pr = pr }) = N
|
||||
_<Cr_.property (flipR<C {a} {b} record { e = e ; 0<e = 0<e ; N = N ; pr = pr }) m N<m with pr m N<m
|
||||
... | t = <WellDefined (transitive (+WellDefined groupIsAbelian (+WellDefined reflexive (identityOfIndiscernablesRight _∼_ reflexive (mapAndIndex (CauchyCompletion.elts b) inverse m)))) (transitive (symmetric +Associative) (+WellDefined reflexive (transitive +Associative (transitive (+WellDefined invRight reflexive) identLeft))))) (transitive groupIsAbelian (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invLeft) identRight))) (orderRespectsAddition t (inverse a + inverse (index (CauchyCompletion.elts b) m)))
|
||||
|
||||
flip<C : {a b : CauchyCompletion} → a <C b → Group.inverse CGroup b <C Group.inverse CGroup a
|
||||
flip<C {a} {b} record { i = i ; a<i = a<i ; i<b = i<b } = record { i = inverse i ; a<i = flipR<C i<b ; i<b = flip<CR a<i }
|
||||
|
||||
flip<C' : {a b : CauchyCompletion} → Group.inverse CGroup b <C Group.inverse CGroup a → a <C b
|
||||
flip<C' {a} {b} a<b = <CWellDefined (invTwice CGroup a) (invTwice CGroup b) (flip<C a<b)
|
||||
|
||||
<CIrreflexive : {a : CauchyCompletion} → a <C a → False
|
||||
<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)
|
||||
|
@@ -9,11 +9,12 @@ open import Groups.Definition
|
||||
open import Fields.Fields
|
||||
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 Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
|
||||
module Fields.CauchyCompletion.Definition {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (F : Field R) where
|
||||
|
||||
@@ -27,6 +28,7 @@ open Group additiveGroup
|
||||
open Field F
|
||||
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Rings.Orders.Total.AbsoluteValue order
|
||||
open import Rings.Orders.Total.Cauchy order
|
||||
open import Groups.Lemmas additiveGroup
|
||||
|
||||
@@ -37,11 +39,11 @@ cauchyWellDefined {s1} {s2} prop c e 0<e with c e 0<e
|
||||
record CauchyCompletion : Set (m ⊔ o) where
|
||||
field
|
||||
elts : Sequence A
|
||||
converges : cauchy elts
|
||||
converges : (cauchy elts)
|
||||
|
||||
injection : A → CauchyCompletion
|
||||
CauchyCompletion.elts (injection a) = constSequence a
|
||||
CauchyCompletion.converges (injection a) = λ ε 0<e → 0 , λ {m} {n} _ _ → <WellDefined (symmetric (identityOfIndiscernablesRight _∼_ (absWellDefined (index (constSequence a) m + inverse (index (constSequence a) n)) 0R (t m n)) absZero)) reflexive 0<e
|
||||
CauchyCompletion.converges (injection a) = (λ ε 0<e → 0 , λ {m} {n} _ _ → <WellDefined (symmetric (identityOfIndiscernablesRight _∼_ (absWellDefined (index (constSequence a) m + inverse (index (constSequence a) n)) 0R (t m n)) absZero)) reflexive 0<e)
|
||||
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)))
|
||||
|
@@ -1,4 +1,4 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
|
||||
{-# OPTIONS --allow-unsolved-metas --warning=error --without-K --guardedness #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
@@ -11,13 +11,14 @@ open import Groups.Groups
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
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 _<_} {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
|
||||
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) where
|
||||
|
||||
open Setoid S
|
||||
open PartiallyOrderedRing pRing
|
||||
@@ -27,17 +28,18 @@ open SetoidTotalOrder (TotallyOrderedRing.total order)
|
||||
open Field F
|
||||
open Group (Ring.additiveGroup R)
|
||||
|
||||
open import Rings.Orders.Total.Cauchy 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
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
open import Fields.CauchyCompletion.Group order F charNot2
|
||||
open import Fields.CauchyCompletion.Ring order F charNot2
|
||||
open import Fields.CauchyCompletion.Comparison order F charNot2
|
||||
open import Fields.CauchyCompletion.Multiplication order F
|
||||
open import Fields.CauchyCompletion.Addition order F
|
||||
open import Fields.CauchyCompletion.Setoid order F
|
||||
open import Fields.CauchyCompletion.Group order F
|
||||
open import Fields.CauchyCompletion.Ring order F
|
||||
open import Fields.CauchyCompletion.Comparison order F
|
||||
|
||||
Cnontrivial : (pr : Setoid._∼_ cauchyCompletionSetoid (injection (Ring.0R R)) (injection (Ring.1R R))) → False
|
||||
Cnontrivial pr with pr (Ring.1R R) (0<1 (charNot2ImpliesNontrivial R charNot2))
|
||||
Cnontrivial pr with pr (Ring.1R R) (0<1 nontrivial)
|
||||
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)
|
||||
|
||||
|
@@ -10,7 +10,8 @@ open import Groups.Homomorphisms.Definition
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
@@ -28,13 +29,14 @@ open Group (Ring.additiveGroup R)
|
||||
open Ring R
|
||||
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Rings.Orders.Total.AbsoluteValue order
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Addition order F
|
||||
open import Fields.CauchyCompletion.Setoid order F
|
||||
|
||||
abstract
|
||||
+CCommutative : {a b : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a +C b) (b +C a)
|
||||
+CCommutative {a} {b} ε 0<e = 0 , ans
|
||||
+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
|
||||
@@ -58,13 +60,13 @@ private
|
||||
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}))
|
||||
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}))
|
||||
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)
|
||||
@@ -88,7 +90,7 @@ private
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (a +C injection 0G)) (map inverse (CauchyCompletion.elts a)) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (constSequence 0G) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts a) inverse m) | indexAndConst 0G m = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (Equivalence.transitive eq (+WellDefined (identRight) (Equivalence.reflexive eq)) (invRight))) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) absZero))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
CidentLeft : {a : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (injection 0G +C a) a
|
||||
CidentLeft {a} = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection 0G +C a} {a +C injection 0G} {a} (+CCommutative {injection 0G} {a}) (CidentRight {a})
|
||||
CidentLeft {a} = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection 0G +C a} {a +C injection 0G} {a} (+CCommutative (injection 0G) a) (CidentRight {a})
|
||||
|
||||
CinvRight : {a : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a +C (-C a)) (injection 0G)
|
||||
CinvRight {a} ε 0<e = 0 , ans
|
||||
@@ -103,7 +105,7 @@ Group.inverse CGroup = -C_
|
||||
Group.+Associative CGroup {a} {b} {c} = Cassoc {a} {b} {c}
|
||||
Group.identRight CGroup {a} = CidentRight {a}
|
||||
Group.identLeft CGroup {a} = CidentLeft {a}
|
||||
Group.invLeft CGroup {a} = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {(-C a) +C a} {a +C (-C a)} {injection 0G} (+CCommutative { -C a} {a}) (CinvRight {a})
|
||||
Group.invLeft CGroup {a} = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {(-C a) +C a} {a +C (-C a)} {injection 0G} (+CCommutative (-C a) a) (CinvRight {a})
|
||||
Group.invRight CGroup {a} = CinvRight {a}
|
||||
|
||||
CInjectionGroupHom : GroupHom (Ring.additiveGroup R) CGroup injection
|
||||
|
@@ -10,7 +10,8 @@ open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
@@ -33,6 +34,8 @@ open import Fields.Orders.Total.Lemmas {F = F} record { oRing = order }
|
||||
open import Fields.Lemmas F
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Rings.Orders.Total.Cauchy order
|
||||
open import Rings.Orders.Total.AbsoluteValue order
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Setoid order F
|
||||
open import Fields.CauchyCompletion.Approximation order F
|
||||
@@ -44,11 +47,10 @@ 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+)))
|
||||
|
||||
_*C_ : CauchyCompletion → CauchyCompletion → CauchyCompletion
|
||||
CauchyCompletion.elts (record { elts = a ; converges = aConv } *C record { elts = b ; converges = bConv }) = apply _*_ a b
|
||||
CauchyCompletion.converges (record { elts = a ; converges = aConv } *C record { elts = b ; converges = bConv }) e 0<e with boundModulus record { elts = a ; converges = aConv }
|
||||
... | aBound , (Na , prABound) with boundModulus record { elts = b ; converges = bConv }
|
||||
... | bBound , (Nb , prBBound) = N , ans
|
||||
timesConverges : (a b : CauchyCompletion) → cauchy (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts b))
|
||||
timesConverges record { elts = a ; converges = aConv } record { elts = b ; converges = bConv } e 0<e with boundModulus record { elts = a ; converges = aConv }
|
||||
... | aBound , (Na , prABound) with boundModulus record { elts = b ; converges = bConv }
|
||||
... | bBound , (Nb , prBBound) = N , ans
|
||||
where
|
||||
boundBoth : A
|
||||
boundBoth = aBound + bBound
|
||||
@@ -151,6 +153,10 @@ CauchyCompletion.converges (record { elts = a ; converges = aConv } *C record {
|
||||
... | 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 = <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
|
||||
|
||||
_*C_ : CauchyCompletion → CauchyCompletion → CauchyCompletion
|
||||
CauchyCompletion.elts (record { elts = a ; converges = aConv } *C record { elts = b ; converges = bConv }) = apply _*_ a b
|
||||
CauchyCompletion.converges (a *C b) = timesConverges a b
|
||||
|
||||
*CCommutative : {a b : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a *C b) (b *C a)
|
||||
*CCommutative {a} {b} ε 0<e = 0 , ans
|
||||
where
|
||||
@@ -159,7 +165,7 @@ CauchyCompletion.converges (record { elts = a ; converges = aConv } *C record {
|
||||
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 (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts b)) (map inverse (apply _*_ (CauchyCompletion.elts b) (CauchyCompletion.elts 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
|
||||
|
||||
abstract
|
||||
private
|
||||
|
||||
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
|
||||
@@ -229,11 +235,12 @@ abstract
|
||||
multiplicationPreservedRight : {a b : A} {c : CauchyCompletion} → (a ∼ b) → Setoid._∼_ cauchyCompletionSetoid (c *C injection a) (c *C injection b)
|
||||
multiplicationPreservedRight {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} (multiplicationPreservedLeft {a} {b} {c} a=b) (*CCommutative {injection b} {c}))
|
||||
|
||||
multiplicationPreserved : {a b c d : A} → (a ∼ b) → (c ∼ d) → Setoid._∼_ cauchyCompletionSetoid (injection a *C injection c) (injection b *C injection d)
|
||||
multiplicationPreserved {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} (multiplicationPreservedRight {c} {d} {injection a} c=d) (multiplicationPreservedLeft {a} {b} {injection d} a=b)
|
||||
multiplicationPreserved : {a b c d : A} → (a ∼ b) → (c ∼ d) → Setoid._∼_ cauchyCompletionSetoid (injection a *C injection c) (injection b *C injection d)
|
||||
multiplicationPreserved {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} (multiplicationPreservedRight {c} {d} {injection a} c=d) (multiplicationPreservedLeft {a} {b} {injection d} a=b)
|
||||
|
||||
private
|
||||
multiplicationWellDefinedRight : (a b c : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid b c → Setoid._∼_ cauchyCompletionSetoid (a *C b) (a *C c)
|
||||
multiplicationWellDefinedRight 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} (multiplicationWellDefinedLeft b c a b=c) (*CCommutative {c} {a}))
|
||||
|
||||
multiplicationWellDefined : {a b c d : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid c d → Setoid._∼_ cauchyCompletionSetoid (a *C c) (b *C d)
|
||||
multiplicationWellDefined {a} {b} {c} {d} a=b c=d = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {a *C c} {a *C d} {b *C d} (multiplicationWellDefinedRight a c d c=d) (multiplicationWellDefinedLeft a b d a=b)
|
||||
multiplicationWellDefined : {a b c d : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid c d → Setoid._∼_ cauchyCompletionSetoid (a *C c) (b *C d)
|
||||
multiplicationWellDefined {a} {b} {c} {d} a=b c=d = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {a *C c} {a *C d} {b *C d} (multiplicationWellDefinedRight a c d c=d) (multiplicationWellDefinedLeft a b d a=b)
|
||||
|
@@ -12,7 +12,8 @@ open import Fields.Fields
|
||||
open import Fields.Orders.Total.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
@@ -21,9 +22,19 @@ open import Numbers.Naturals.Order.Lemmas
|
||||
open import Semirings.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Orders.Total.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) where
|
||||
|
||||
private
|
||||
lemma3 : {c d : _} {C : Set c} {S : Setoid {c} {d} C} {_+_ : C → C → C} (G : Group S _+_) → ({x y : C} → Setoid._∼_ S (x + y) (y + x)) → (x y z : C) → Setoid._∼_ S (((Group.inverse G x) + y) + (x + z)) (y + z)
|
||||
lemma3 {S = S} {_+_ = _+_} G ab x y z = transitive +Associative (+WellDefined (transitive (ab {inverse x + y}) (transitive +Associative (transitive (+WellDefined invRight reflexive) identLeft))) reflexive)
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence (Setoid.eq S)
|
||||
open Group G
|
||||
|
||||
open Setoid S
|
||||
open SetoidTotalOrder (TotallyOrderedRing.total order)
|
||||
open SetoidPartialOrder pOrder
|
||||
@@ -32,6 +43,7 @@ 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
|
||||
@@ -41,9 +53,12 @@ open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Addition order F
|
||||
open import Fields.CauchyCompletion.Multiplication order F
|
||||
open import Fields.CauchyCompletion.Approximation order F
|
||||
open import Fields.CauchyCompletion.Group order F
|
||||
open import Fields.CauchyCompletion.Ring order F
|
||||
open import Fields.CauchyCompletion.Comparison order F
|
||||
open import Fields.CauchyCompletion.Setoid order F
|
||||
open import Groups.Homomorphisms.Lemmas CInjectionGroupHom
|
||||
open import Setoids.Orders.Total.Lemmas (TotallyOrderedRing.total order)
|
||||
|
||||
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 }
|
||||
@@ -57,21 +72,8 @@ productPositives' a b interA interB 0<iA 0<iB record { e = interA' ; 0<e = 0<int
|
||||
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.
|
||||
-}
|
||||
<COrderRespectsMultiplication : (a b : CauchyCompletion) → (injection 0R <C a) → (injection 0R <C b) → (injection 0R <C (a *C b))
|
||||
<COrderRespectsMultiplication 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 }
|
||||
|
||||
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) }))
|
||||
@@ -79,83 +81,156 @@ cOrderRespectsAdditionLeft' a b c record { e = e ; 0<e = 0<e ; N = N ; property
|
||||
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 = {!!}
|
||||
private
|
||||
l1 : (a : A) (b : CauchyCompletion) → a r<C b → 0R r<C (b +C injection (inverse a))
|
||||
l1 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 (transitive groupIsAbelian (transitive +Associative (+WellDefined invLeft reflexive))) (ans m) (orderRespectsAddition (pr m N<m) (inverse a)) }
|
||||
where
|
||||
ans : (m : ℕ) → (index (CauchyCompletion.elts b) m + inverse a) ∼ index (CauchyCompletion.elts (b +C injection (inverse a))) m
|
||||
ans m rewrite indexAndApply (CauchyCompletion.elts b) (constSequence (inverse a)) _+_ {m} | indexAndConst (inverse a) m = reflexive
|
||||
|
||||
cOrderRespectsAdditionRight : (a : A) (b : CauchyCompletion) (c : CauchyCompletion) → (a r<C b) → (injection a +C c) <C (b +C c)
|
||||
cOrderRespectsAdditionRight a b a<b = {!!}
|
||||
l1' : (a : A) (b : CauchyCompletion) → 0R r<C (b +C injection a) → (inverse a) r<C b
|
||||
l1' 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 (transitive (+WellDefined identLeft reflexive) groupIsAbelian) (symmetric (ans m)) (orderRespectsAddition (pr m N<m) (inverse a)) }
|
||||
where
|
||||
ans : (m : ℕ) → (index (CauchyCompletion.elts b) m) ∼ (index (CauchyCompletion.elts (b +C injection a)) m + inverse a)
|
||||
ans m rewrite indexAndApply (CauchyCompletion.elts b) (constSequence a) _+_ {m} | indexAndConst a m = transitive (symmetric identRight) (transitive (+WellDefined reflexive (symmetric invRight)) +Associative)
|
||||
|
||||
cOrderRespectsAddition : (a b : CauchyCompletion) → (a <C b) → (c : CauchyCompletion) → (a +C c) <C (b +C c)
|
||||
cOrderRespectsAddition a b a<b c = {!!}
|
||||
l2 : (a : CauchyCompletion) (b : A) → a <Cr b → 0R r<C (injection b +C (-C a))
|
||||
l2 a b record { e = e ; 0<e = 0<e ; N = N ; property = property } = record { e = e ; 0<e = 0<e ; N = N ; pr = λ m N<m → <WellDefined (transitive (transitive (symmetric +Associative) (+WellDefined reflexive invRight)) groupIsAbelian) (ans m) (orderRespectsAddition (property m N<m) (inverse (index (CauchyCompletion.elts a) m))) }
|
||||
where
|
||||
ans : (m : ℕ) → (b + inverse (index (CauchyCompletion.elts a) m)) ∼ index (CauchyCompletion.elts (injection b +C (-C a))) m
|
||||
ans m rewrite indexAndApply (CauchyCompletion.elts (injection b)) (CauchyCompletion.elts (-C a)) _+_ {m} | indexAndConst b m | mapAndIndex (CauchyCompletion.elts a) inverse m = reflexive
|
||||
|
||||
l2' : (a : CauchyCompletion) (b : A) → 0R r<C (injection b +C a) → (-C a) <Cr b
|
||||
l2' a b record { e = e ; 0<e = 0<e ; N = N ; pr = pr } = record { e = e ; 0<e = 0<e ; N = N ; property = λ m N<m → <WellDefined (+WellDefined identLeft reflexive) (ans m) (orderRespectsAddition (pr m N<m) (index (CauchyCompletion.elts (-C a)) m)) }
|
||||
where
|
||||
ans : (m : ℕ) → (index (CauchyCompletion.elts (injection b +C a)) m + index (map inverse (CauchyCompletion.elts a)) m) ∼ b
|
||||
ans m rewrite indexAndApply (CauchyCompletion.elts (injection b)) (CauchyCompletion.elts a) _+_ {m} | indexAndConst b m | equalityCommutative (mapAndIndex (CauchyCompletion.elts a) inverse m) = transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight)
|
||||
|
||||
{-
|
||||
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)
|
||||
minAddLemma : (a b : A) → (0R < a) → (min a b) < (a + b)
|
||||
minAddLemma a b 0<a with totality a b
|
||||
... | inl (inl x) = <WellDefined identLeft groupIsAbelian (orderRespectsAddition (<Transitive 0<a x) a)
|
||||
... | inl (inr x) = <WellDefined identLeft reflexive (orderRespectsAddition 0<a b)
|
||||
... | inr x = <WellDefined identLeft groupIsAbelian (orderRespectsAddition (<WellDefined reflexive x 0<a) a)
|
||||
|
||||
addInequalities : {a b : CauchyCompletion} → 0R r<C a → 0R r<C b → 0R r<C (a +C b)
|
||||
addInequalities {a} {b} record { e = eA ; 0<e = 0<eA ; N = Na ; pr = prA } record { e = eB ; 0<e = 0<eB ; N = Nb ; pr = prB } = record { e = min eA eB ; 0<e = minInequalitiesR 0<eA 0<eB ; N = N ; pr = λ m N<m → <Transitive (<WellDefined (symmetric identLeft) reflexive (minAddLemma eA eB 0<eA)) (<WellDefined (+WellDefined identLeft identLeft) (identityOfIndiscernablesLeft _∼_ reflexive (indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts b) _+_ {m})) (ringAddInequalities (prA m (p2 m N<m)) (prB m (p1 m N<m)))) }
|
||||
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
|
||||
N = succ (TotalOrder.max ℕTotalOrder Na Nb)
|
||||
Nb<N : Nb <N succ (TotalOrder.max ℕTotalOrder Na Nb)
|
||||
Nb<N with TotalOrder.totality ℕTotalOrder Na Nb
|
||||
... | inl (inl x) = le 0 refl
|
||||
... | inl (inr x) = lessTransitive x (le 0 refl)
|
||||
... | inr x = le 0 refl
|
||||
Na<N : Na <N succ (TotalOrder.max ℕTotalOrder Na Nb)
|
||||
Na<N with TotalOrder.totality ℕTotalOrder Na Nb
|
||||
... | inl (inl x) = lessTransitive x (le 0 refl)
|
||||
... | inl (inr x) = le 0 refl
|
||||
... | inr x = le 0 (applyEquality succ x)
|
||||
p1 : (m : ℕ) → (N <N m) → Nb <N m
|
||||
p1 m N<m = lessTransitive Nb<N N<m
|
||||
p2 : (m : ℕ) → (N <N m) → Na <N m
|
||||
p2 m N<m = lessTransitive Na<N N<m
|
||||
|
||||
-}
|
||||
|
||||
invInj : (i : A) → Setoid._∼_ cauchyCompletionSetoid (injection (inverse i) +C injection i) (injection 0R)
|
||||
invInj i = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (injection (inverse i) +C injection i) }} {record { converges = CauchyCompletion.converges ((-C (injection i)) +C injection i) }} {record { converges = CauchyCompletion.converges (injection 0R) }} (Group.+WellDefined CGroup {record { converges = CauchyCompletion.converges (injection (inverse i)) }} {record { converges = CauchyCompletion.converges (injection i)}} {record { converges = CauchyCompletion.converges (-C (injection i)) }} {record { converges = CauchyCompletion.converges (injection i) }} homRespectsInverse (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (injection i) }})) (Group.invLeft CGroup {injection i})
|
||||
|
||||
|
||||
cOrderMove : (a b : A) (c : CauchyCompletion) → (injection a +C c) <Cr b → a r<C (injection b +C (-C c))
|
||||
cOrderMove a b c record { e = e ; 0<e = 0<e ; N = N ; property = property } = record { e = e ; 0<e = 0<e ; N = N ; pr = λ m N<m → <WellDefined (transitive (symmetric +Associative) (transitive (+WellDefined reflexive (transitive (+WellDefined (identityOfIndiscernablesRight _∼_ reflexive (indexAndApply (CauchyCompletion.elts (injection a)) (CauchyCompletion.elts c) _+_ {m})) reflexive) (transitive (symmetric +Associative) (transitive (+WellDefined (identityOfIndiscernablesRight _∼_ reflexive (indexAndConst a m)) invRight) identRight)))) groupIsAbelian)) (transitive (+WellDefined (identityOfIndiscernablesLeft _∼_ reflexive (indexAndConst b m)) (identityOfIndiscernablesRight _∼_ reflexive (mapAndIndex (CauchyCompletion.elts c) inverse m))) (identityOfIndiscernablesLeft _∼_ reflexive (indexAndApply (CauchyCompletion.elts (injection b)) (CauchyCompletion.elts (-C c)) _+_ {m}))) (orderRespectsAddition (property m N<m) (inverse (index (CauchyCompletion.elts c) m))) }
|
||||
|
||||
cOrderMove' : (a b : A) (c : CauchyCompletion) → (injection a +C (-C c)) <Cr b → a r<C (injection b +C c)
|
||||
cOrderMove' a b c pr = r<CWellDefinedRight _ _ _ (Group.+WellDefined CGroup {record { converges = CauchyCompletion.converges (injection b) }} {record { converges = CauchyCompletion.converges (-C (-C c)) }} {record { converges = CauchyCompletion.converges (injection b) }} {record {converges = CauchyCompletion.converges c }} (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (injection b) }}) (invTwice CGroup record { converges = CauchyCompletion.converges c })) (cOrderMove a b (-C c) pr)
|
||||
|
||||
cOrderMove'' : (a : CauchyCompletion) (b c : A) → (a +C (-C injection b)) <Cr c → a <Cr (b + c)
|
||||
cOrderMove'' a b c record { e = e ; 0<e = 0<e ; N = N ; property = property } = record { e = e ; 0<e = 0<e ; N = N ; property = λ m N<m → <WellDefined (transitive (symmetric +Associative) (+WellDefined reflexive (transitive (+WellDefined (identityOfIndiscernablesRight _∼_ reflexive (indexAndApply (CauchyCompletion.elts a) _ _+_ {m})) reflexive) (transitive (transitive (symmetric +Associative) (+WellDefined reflexive (transitive (+WellDefined (transitive (identityOfIndiscernablesLeft _∼_ reflexive (mapAndIndex _ inverse m)) (inverseWellDefined additiveGroup (identityOfIndiscernablesRight _∼_ reflexive (indexAndConst b m)))) reflexive) invLeft))) identRight)))) groupIsAbelian (orderRespectsAddition (property m N<m) b) }
|
||||
|
||||
cOrderRespectsAdditionLeft'' : (a b : A) (c : CauchyCompletion) → (a < b) → (injection a +C c) <C (injection b +C c)
|
||||
cOrderRespectsAdditionLeft'' a b c a<b with halve charNot2 (b + inverse a)
|
||||
... | b-a/2 , prDiff with approximateAbove c b-a/2 (halvePositive' prDiff (moveInequality a<b))
|
||||
... | aboveC , (c<aboveC ,, aboveC-C<e) = record { i = a + aboveC ; a<i = <CRelaxR' (SetoidPartialOrder.<WellDefined <COrder (Ring.groupIsAbelian CRing {record { converges = CauchyCompletion.converges c }} {record { converges = CauchyCompletion.converges (injection a) }}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (injection aboveC +C injection a) }} {record { converges = CauchyCompletion.converges (injection a +C injection aboveC) }} {record { converges = CauchyCompletion.converges (injection (a + aboveC)) }} (Ring.groupIsAbelian CRing {record { converges = CauchyCompletion.converges (injection aboveC) }} {record { converges = CauchyCompletion.converges (injection a) }}) (Equivalence.symmetric (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (injection (a + aboveC)) }} {record { converges = CauchyCompletion.converges (injection a +C injection aboveC) }} (GroupHom.groupHom CInjectionGroupHom))) (cOrderRespectsAdditionLeft' c aboveC a c<aboveC)) ; i<b = cOrderMove' _ _ _ (<CRelaxR' (<CTransitive (<CRelaxR t) (<CInherited u))) }
|
||||
where
|
||||
g : ((injection aboveC +C (-C c)) +C injection a) <C ((injection b-a/2) +C (injection a))
|
||||
g = cOrderRespectsAdditionLeft' _ _ a (<CRelaxR' aboveC-C<e)
|
||||
g' : ((injection aboveC +C (-C c)) +C injection a) <C (injection (b-a/2 + a))
|
||||
g' = <CWellDefined (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges ((injection aboveC +C (-C c)) +C injection a) }}) (Equivalence.symmetric (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (injection (b-a/2 + a)) }} {record { converges = CauchyCompletion.converges (injection b-a/2 +C injection a) }} (GroupHom.groupHom CInjectionGroupHom)) g
|
||||
t : (injection (a + aboveC) +C (-C c)) <Cr (b-a/2 + a)
|
||||
t = <CRelaxR' (<CWellDefined (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (((injection aboveC) +C (-C c)) +C injection a) }} {record { converges = CauchyCompletion.converges (injection a +C ((injection aboveC) +C (-C c))) }} {record { converges = CauchyCompletion.converges (injection (a + aboveC) +C (-C c)) }} (Ring.groupIsAbelian CRing {record { converges = CauchyCompletion.converges (injection aboveC +C (-C c)) }} {record { converges = CauchyCompletion.converges (injection a) }}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (injection a +C ((injection aboveC) +C (-C c))) }} {record { converges = CauchyCompletion.converges (((injection a) +C (injection aboveC)) +C (-C c)) }} {record { converges = CauchyCompletion.converges ((injection (a + aboveC)) +C (-C c)) }} (Group.+Associative CGroup {record { converges = CauchyCompletion.converges (injection a) }} {record { converges = CauchyCompletion.converges (injection aboveC) }} {record { converges = CauchyCompletion.converges (-C c) }}) (Group.+WellDefined CGroup {record { converges = CauchyCompletion.converges (injection a +C injection aboveC) }} {record { converges = CauchyCompletion.converges (-C c) }} {record { converges = CauchyCompletion.converges (injection (a + aboveC)) }} {record { converges = CauchyCompletion.converges (-C c) }} (Equivalence.symmetric (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (injection (a + aboveC)) }} {record { converges = CauchyCompletion.converges (injection a +C injection aboveC) }} (GroupHom.groupHom CInjectionGroupHom)) (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (-C c)}})))) (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (injection (b-a/2 + a)) }}) g')
|
||||
lemm : 0R < b-a/2
|
||||
lemm = halvePositive' prDiff (moveInequality a<b)
|
||||
u : (b-a/2 + a) < b
|
||||
u = <WellDefined (transitive (+WellDefined (symmetric prDiff) reflexive) (transitive groupIsAbelian (transitive +Associative (transitive (+WellDefined (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invLeft) identRight)) reflexive) groupIsAbelian)))) (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invLeft) identRight)) (orderRespectsAddition (<WellDefined identLeft (transitive (transitive +Associative (transitive (+WellDefined (transitive groupIsAbelian (+WellDefined reflexive (symmetric (invTwice additiveGroup b-a/2)))) reflexive) (symmetric +Associative))) (+WellDefined reflexive (symmetric (invContravariant additiveGroup)))) (orderRespectsAddition lemm (b + inverse a))) (a + inverse b-a/2))
|
||||
|
||||
cOrderRespectsAdditionLeft''Flip : (a b : A) (c : CauchyCompletion) → (a < b) → (c +C injection a) <C (c +C injection b)
|
||||
cOrderRespectsAdditionLeft''Flip a b c a<b = <CWellDefined ((Ring.groupIsAbelian CRing {record { converges = CauchyCompletion.converges (injection a) }} {record { converges = CauchyCompletion.converges c }})) (Ring.groupIsAbelian CRing {record { converges = CauchyCompletion.converges (injection b) }} {record { converges = CauchyCompletion.converges c }}) (cOrderRespectsAdditionLeft'' a b c a<b)
|
||||
|
||||
cOrderRespectsAdditionLeft''' : (a b : CauchyCompletion) (c : A) → (a <C b) → (a +C injection c) <C (b +C injection c)
|
||||
cOrderRespectsAdditionLeft''' a b c record { i = i ; a<i = a<i ; i<b = i<b } = <CTransitive (cOrderRespectsAdditionLeft' a i c a<i) (<CWellDefined (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (injection (c + i)) }} {record { converges = CauchyCompletion.converges (injection c +C injection i) }} {record { converges = CauchyCompletion.converges (injection i +C injection c) }} (GroupHom.groupHom CInjectionGroupHom) (Ring.groupIsAbelian CRing {record { converges = CauchyCompletion.converges (injection c) }} {record { converges = CauchyCompletion.converges (injection i) }})) (Ring.groupIsAbelian CRing {record { converges = CauchyCompletion.converges (injection c) }} {record { converges = CauchyCompletion.converges b }}) (flip<C' {record { converges = CauchyCompletion.converges (injection (c + i)) }} {record { converges = CauchyCompletion.converges (injection c +C b) }} (<CWellDefined (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges ((-C b) +C injection (inverse c)) }} {record { converges = CauchyCompletion.converges ((-C b) +C (-C (injection c))) }} {record { converges = CauchyCompletion.converges (-C (injection c +C b)) }} (Group.+WellDefined CGroup {record { converges = CauchyCompletion.converges (-C b) }} {record { converges = CauchyCompletion.converges (injection (inverse c)) }} {record { converges = CauchyCompletion.converges (-C b) }} {record { converges = CauchyCompletion.converges (-C (injection c)) }} (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (-C b) }}) homRespectsInverse) (Equivalence.symmetric (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (-C (injection c +C b)) }} {record { converges = CauchyCompletion.converges ((-C b) +C (-C (injection c))) }} (invContravariant CGroup {record { converges = CauchyCompletion.converges (injection c) }} {record { converges = CauchyCompletion.converges b }}))) homRespectsInverse' (cOrderRespectsAdditionLeft' (-C b) (inverse i) (inverse c) (flipR<C i<b)))))
|
||||
|
||||
{-
|
||||
|
||||
Have 0<a, so 0 < a- < a < a+
|
||||
Have c < a- + c, by cOrderRespectsAdditionLeft''
|
||||
Have a- + c < a+ + c by cOrderRespectsAdditionLeft''
|
||||
so a- + c < K < a+ + c
|
||||
|
||||
|
||||
-}
|
||||
|
||||
--cOrderRespectsAdditionLeft'' : (a b : A) (c : CauchyCompletion) → (a < b) → (injection a +C c) <C (injection b +C c)
|
||||
--cOrderRespectsAdditionLeft''' : (a b : CauchyCompletion) (c : A) → (a <C b) → (a +C injection c) <C (b +C injection c)
|
||||
cOrderRespectsAdditionRightZero : (a : CauchyCompletion) → (0R r<C a) → (c : CauchyCompletion) → c <C (a +C c)
|
||||
cOrderRespectsAdditionRightZero a record { e = e ; 0<e = 0<e ; N = N1 ; pr = pr } c with halve charNot2 e
|
||||
... | e/2 , prE/2 with halve charNot2 e/2
|
||||
... | e/4 , prE/4 with halve charNot2 e/4
|
||||
... | e/8 , prE/8 with approximateAbove c e/4 (halvePositive' prE/4 (halvePositive' prE/2 0<e))
|
||||
... | c' , (c<c' ,, c'-c<e/4) with approximateBelow c e/8 (halvePositive' prE/8 (halvePositive' prE/4 (halvePositive' prE/2 0<e)))
|
||||
... | cB , (record { e = f ; 0<e = 0<f ; N = N2 ; pr = pr2 } ,, c-cB<e/8) = record { i = c' + e/4 ; a<i = <CRelaxR' (<CTransitive (<CRelaxR c<c') (<CInherited (<WellDefined identLeft groupIsAbelian (orderRespectsAddition (halvePositive' prE/4 (halvePositive' prE/2 0<e)) c')))) ; i<b = record { e = e/4 ; 0<e = halvePositive' prE/4 (halvePositive' prE/2 0<e) ; N = N ; pr = λ m N<m → {!!} } }
|
||||
where
|
||||
N = TotalOrder.max ℕTotalOrder N1 N2
|
||||
3e/4<e : ((e/4 + e/4) + e/4) < e
|
||||
3e/4<e = <WellDefined identLeft (transitive (+WellDefined reflexive (+WellDefined prE/4 reflexive)) (transitive groupIsAbelian (transitive (symmetric +Associative) (transitive (+WellDefined reflexive prE/4) prE/2)))) (orderRespectsAddition (halvePositive' prE/4 (halvePositive' prE/2 0<e)) ((e/4 + e/4) + e/4))
|
||||
t : c' r<C (c +C injection e/4)
|
||||
t = r<CWellDefinedRight _ _ _ (Ring.groupIsAbelian CRing {record { converges = CauchyCompletion.converges (injection e/4) }} {record { converges = CauchyCompletion.converges c }}) (cOrderMove' _ _ _ (<CRelaxR' c'-c<e/4))
|
||||
u : (c' + e/2) r<C (c +C injection ((e/4 + e/4) + e/4))
|
||||
u = r<CWellDefinedRight record { converges = CauchyCompletion.converges ((c +C injection e/4) +C (injection e/2))} _ _ {!!} (<CRelaxL' (<CWellDefined (GroupHom.groupHom' CInjectionGroupHom) (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges ((c +C injection e/4) +C injection e/2) }}) (cOrderRespectsAdditionLeft''' _ _ e/2 (<CRelaxL t))))
|
||||
u' : ((c' + e/4) + e/4) r<C (c +C injection ((e/4 + e/4) + e/4))
|
||||
u' = r<CWellDefinedLeft _ _ _ (transitive (+WellDefined reflexive (symmetric prE/4)) +Associative) u
|
||||
v : (c +C injection ((e/4 + e/4) + e/4)) <Cr (cB + (((e/4 + e/4) + e/4) + e/8))
|
||||
v with cOrderMove'' _ _ _ (<CRelaxR' c-cB<e/8)
|
||||
... | c<cB+e/8 = <CRelaxR' (<CWellDefined (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (c +C injection ((e/4 + e/4) + e/4))}}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges ((injection (cB + e/8)) +C (injection ((e/4 + e/4) + e/4))) }} {record { converges = CauchyCompletion.converges (injection ((cB + e/8) + ((e/4 + e/4) + e/4))) }} {record { converges = CauchyCompletion.converges (injection (cB + (((e/4 + e/4) + e/4) + e/8))) }} (GroupHom.groupHom' CInjectionGroupHom) (GroupHom.wellDefined CInjectionGroupHom (transitive (symmetric +Associative) (+WellDefined reflexive groupIsAbelian)))) (cOrderRespectsAdditionLeft''' _ _ ((e/4 + e/4) + e/4) (<CRelaxR c<cB+e/8)))
|
||||
w : (m : ℕ) → (N<m : N <N m) → (c +C injection ((e/4 + e/4) + e/4)) <Cr ((index (CauchyCompletion.elts a) m) + (index (CauchyCompletion.elts c) m))
|
||||
w = {!!}
|
||||
--pr2 : (m₁ : ℕ) → N2 <N m₁ → (f + index (CauchyCompletion.elts c) m₁) < (cB + e/8)
|
||||
req : (m : ℕ) → (N<m : N <N m) → ((c' + e/4) + e/4) < index (apply _+_ (CauchyCompletion.elts a) (CauchyCompletion.elts c)) m
|
||||
req = {!!}
|
||||
{-
|
||||
ans : (m : ℕ) → (N<m : N <N m) → (c' + e/2) r<C (injection (index (CauchyCompletion.elts a) m) +C c)
|
||||
ans m N<m = <CRelaxL' (<CTransitive (<CRelaxL u) (<CTransitive v (w m N<m)))
|
||||
req : (m : ℕ) → (N<m : N <N m) → ((c' + e/4) + e/4) < index (apply _+_ (CauchyCompletion.elts a) (CauchyCompletion.elts c)) m
|
||||
req m N<m with ans m N<m
|
||||
... | record { e = f ; 0<e = 0<f ; N = M ; pr = pr } with pr m {!!}
|
||||
... | t rewrite indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts c) _+_ {m} | indexAndConst (CauchyCompletion.elts a) m = {!!}
|
||||
-}
|
||||
|
||||
cOrderRespectsAdditionLeft : (a : CauchyCompletion) (b : A) (c : CauchyCompletion) → (a <Cr b) → (a +C c) <C (injection b +C c)
|
||||
cOrderRespectsAdditionLeft a b c a<b = <CWellDefined {record { converges = CauchyCompletion.converges (a +C c) }}{record { converges = CauchyCompletion.converges (a +C c) }}{record { converges = CauchyCompletion.converges (((-C a) +C injection b) +C (a +C c)) }}{record { converges = CauchyCompletion.converges (injection b +C c) }} (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (a +C c)}}) (lemma3 CGroup (λ {x} {y} → Ring.groupIsAbelian CRing {x} {y}) a (injection b) c) (cOrderRespectsAdditionRightZero ((-C a) +C injection b) (r<CWellDefinedLeft (inverse b + b) 0R ((-C a) +C injection b) invLeft (<CRelaxL' (<CWellDefined {record { converges = CauchyCompletion.converges (injection (inverse b) +C injection b) }}{record { converges = CauchyCompletion.converges (injection (inverse b + b)) }}{record { converges = CauchyCompletion.converges ((-C a) +C injection b) }}{record { converges = CauchyCompletion.converges ((-C a) +C injection b) }} (GroupHom.groupHom' CInjectionGroupHom) (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges ((-C a) +C injection b) }}) (cOrderRespectsAdditionLeft''' _ _ b (<CRelaxL (flip<CR a<b)))))) (a +C c))
|
||||
|
||||
private
|
||||
lemma2 : (a b : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid (Group.inverse CGroup ((-C a) +C (-C b))) (a +C b)
|
||||
lemma2 a b = tr {record { converges = CauchyCompletion.converges (-C ((-C a) +C (-C b))) }} {record { converges = CauchyCompletion.converges ((-C (-C b)) +C (-C (-C a))) }} {record { converges = CauchyCompletion.converges (a +C b) }} (invContravariant CGroup {record { converges = CauchyCompletion.converges (-C a) }} {record { converges = CauchyCompletion.converges (-C b) }}) (tr {record { converges = CauchyCompletion.converges ((-C (-C b)) +C (-C (-C a))) }} {record { converges = CauchyCompletion.converges (b +C a) }} {record { converges = CauchyCompletion.converges (a +C b) }} (Group.+WellDefined CGroup {record { converges = CauchyCompletion.converges (-C (-C b)) }} {record { converges = CauchyCompletion.converges (-C (-C a)) }} {record { converges = CauchyCompletion.converges b }} {record { converges = CauchyCompletion.converges a }} (invTwice CGroup b) (invTwice CGroup a)) (Ring.groupIsAbelian CRing {record { converges = CauchyCompletion.converges b }} {record { converges = CauchyCompletion.converges a }}))
|
||||
where
|
||||
open Setoid cauchyCompletionSetoid
|
||||
open Equivalence (Setoid.eq cauchyCompletionSetoid) renaming (transitive to tr)
|
||||
|
||||
cOrderRespectsAdditionRight : (a : A) (b : CauchyCompletion) (c : CauchyCompletion) → (a r<C b) → (injection a +C c) <C (b +C c)
|
||||
cOrderRespectsAdditionRight a b c a<b = <CWellDefined (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (Group.inverse CGroup (injection (inverse a) +C (-C c))) }} {record { converges = CauchyCompletion.converges (Group.inverse CGroup ((-C injection a) +C (-C c))) }} {record { converges = CauchyCompletion.converges (injection a +C c) }} (inverseWellDefined CGroup {record { converges = CauchyCompletion.converges (injection (inverse a) +C (-C c)) }} {record { converges = CauchyCompletion.converges ((-C (injection a)) +C (-C c)) }} (Group.+WellDefined CGroup {record { converges = CauchyCompletion.converges (injection (inverse a)) }} {record { converges = CauchyCompletion.converges (-C c) }} {record { converges = CauchyCompletion.converges (-C (injection a)) }} {record { converges = CauchyCompletion.converges (-C c) }} homRespectsInverse (Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {record { converges = CauchyCompletion.converges (-C c) }}))) (lemma2 (injection a) c)) (lemma2 b c) (flip<C (cOrderRespectsAdditionLeft _ _ (Group.inverse CGroup c) (flipR<C a<b)))
|
||||
|
||||
cOrderRespectsAddition : (a b : CauchyCompletion) → (a <C b) → (c : CauchyCompletion) → (a +C c) <C (b +C c)
|
||||
cOrderRespectsAddition a b record { i = i ; a<i = a<i ; i<b = i<b } c = SetoidPartialOrder.<Transitive <COrder (cOrderRespectsAdditionLeft a i c a<i) (cOrderRespectsAdditionRight i b c i<b)
|
||||
|
||||
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 }
|
||||
|
||||
PartiallyOrderedRing.orderRespectsMultiplication CpOrderedRing {a} {b} 0<a 0<b = <COrderRespectsMultiplication a b 0<a 0<b
|
||||
|
@@ -9,7 +9,8 @@ open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
@@ -27,6 +28,7 @@ open Field F
|
||||
open Group (Ring.additiveGroup R)
|
||||
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Rings.Orders.Total.AbsoluteValue order
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Multiplication order F
|
||||
open import Fields.CauchyCompletion.Addition order F
|
||||
@@ -57,7 +59,7 @@ CRing : Ring cauchyCompletionSetoid _+C_ _*C_
|
||||
Ring.additiveGroup CRing = CGroup
|
||||
Ring.*WellDefined CRing {a} {b} {c} {d} r=t s=u = multiplicationWellDefined {a} {c} {b} {d} r=t s=u
|
||||
Ring.1R CRing = injection (Ring.1R R)
|
||||
Ring.groupIsAbelian CRing {a} {b} = +CCommutative {a} {b}
|
||||
Ring.groupIsAbelian CRing {a} {b} = +CCommutative a b
|
||||
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}
|
||||
|
@@ -11,7 +11,8 @@ open import Groups.Lemmas
|
||||
open import Fields.Fields
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
@@ -34,16 +35,18 @@ open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Addition order F
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Rings.Orders.Total.AbsoluteValue order
|
||||
open import Rings.InitialRing R
|
||||
open import Fields.Orders.Total.Lemmas {F = F} (record { oRing = order })
|
||||
|
||||
isZero : CauchyCompletion → Set (m ⊔ o)
|
||||
isZero record { elts = elts ; converges = converges } = ∀ ε → 0R < ε → Sg ℕ (λ N → ∀ {m : ℕ} → (N <N m) → (abs (index elts m)) < ε)
|
||||
|
||||
transitiveLemma : {a b c e/2 : A} → abs (a + inverse b) < e/2 → abs (b + inverse c) < e/2 → (abs (a + inverse c)) < (e/2 + e/2)
|
||||
transitiveLemma {a} {b} {c} {e/2} a-b<e/2 b-c<e/2 with triangleInequality (a + inverse b) (b + inverse c)
|
||||
transitiveLemma {a} {b} {c} {e/2} a-b<e/2 b-c<e/2 | inl x = SetoidPartialOrder.<Transitive pOrder (<WellDefined (absWellDefined _ _ (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq +Associative (Equivalence.transitive eq (+WellDefined invLeft (Equivalence.reflexive eq)) identLeft))))) (Equivalence.reflexive eq) x) (ringAddInequalities a-b<e/2 b-c<e/2)
|
||||
transitiveLemma {a} {b} {c} {e/2} a-b<e/2 b-c<e/2 | inr x = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (Equivalence.symmetric eq ((Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq +Associative (Equivalence.transitive eq (+WellDefined invLeft (Equivalence.reflexive eq)) identLeft))))))) x)) (Equivalence.reflexive eq) (ringAddInequalities a-b<e/2 b-c<e/2)
|
||||
private
|
||||
transitiveLemma : {a b c e/2 : A} → abs (a + inverse b) < e/2 → abs (b + inverse c) < e/2 → (abs (a + inverse c)) < (e/2 + e/2)
|
||||
transitiveLemma {a} {b} {c} {e/2} a-b<e/2 b-c<e/2 with triangleInequality (a + inverse b) (b + inverse c)
|
||||
transitiveLemma {a} {b} {c} {e/2} a-b<e/2 b-c<e/2 | inl x = SetoidPartialOrder.<Transitive pOrder (<WellDefined (absWellDefined _ _ (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq +Associative (Equivalence.transitive eq (+WellDefined invLeft (Equivalence.reflexive eq)) identLeft))))) (Equivalence.reflexive eq) x) (ringAddInequalities a-b<e/2 b-c<e/2)
|
||||
transitiveLemma {a} {b} {c} {e/2} a-b<e/2 b-c<e/2 | inr x = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (Equivalence.symmetric eq ((Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq +Associative (Equivalence.transitive eq (+WellDefined invLeft (Equivalence.reflexive eq)) identLeft))))))) x)) (Equivalence.reflexive eq) (ringAddInequalities a-b<e/2 b-c<e/2)
|
||||
|
||||
cauchyCompletionSetoid : Setoid CauchyCompletion
|
||||
(cauchyCompletionSetoid Setoid.∼ a) b = isZero (a +C (-C b))
|
||||
|
68
Fields/FieldOfFractions/Archimedean.agda
Normal file
68
Fields/FieldOfFractions/Archimedean.agda
Normal file
@@ -0,0 +1,68 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Orders.Archimedean
|
||||
|
||||
module Fields.FieldOfFractions.Archimedean {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) {c : _} {_<_ : Rel {_} {c} A} {pOrder : SetoidPartialOrder S _<_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) (nonempty : (Setoid._∼_ S (Ring.0R R) (Ring.1R R)) → False) where
|
||||
|
||||
open import Groups.Cyclic.Definition
|
||||
open import Fields.FieldOfFractions.Setoid I
|
||||
open import Fields.FieldOfFractions.Group I
|
||||
open import Fields.FieldOfFractions.Addition I
|
||||
open import Fields.FieldOfFractions.Ring I
|
||||
open import Fields.FieldOfFractions.Order I order
|
||||
open import Fields.FieldOfFractions.Field I
|
||||
open import Fields.Orders.Total.Archimedean
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open TotallyOrderedRing order
|
||||
open SetoidTotalOrder total
|
||||
|
||||
private
|
||||
|
||||
denomPower : (n : ℕ) → fieldOfFractionsSet.denom (positiveEltPower fieldOfFractionsGroup record { num = 1R ; denom = 1R ; denomNonzero = IntegralDomain.nontrivial I } n) ∼ 1R
|
||||
denomPower zero = reflexive
|
||||
denomPower (succ n) = transitive identIsIdent (denomPower n)
|
||||
|
||||
denomPlus : {a : A} .(a!=0 : a ∼ 0R → False) (n1 n2 : A) → Setoid._∼_ fieldOfFractionsSetoid (fieldOfFractionsPlus record { num = n1 ; denom = a ; denomNonzero = a!=0 } record { num = n2 ; denom = a ; denomNonzero = a!=0 }) (record { num = n1 + n2 ; denom = a ; denomNonzero = a!=0 })
|
||||
denomPlus {a} a!=0 n1 n2 = transitive *Commutative (transitive (*WellDefined reflexive (transitive (+WellDefined *Commutative reflexive) (symmetric *DistributesOver+))) *Associative)
|
||||
|
||||
d : (a : fieldOfFractionsSet) → fieldOfFractionsSet.denom a ∼ 0R → False
|
||||
d record { num = num ; denom = denom ; denomNonzero = denomNonzero } bad = exFalso (denomNonzero bad)
|
||||
|
||||
simpPower : (n : ℕ) → Setoid._∼_ fieldOfFractionsSetoid (positiveEltPower fieldOfFractionsGroup record { num = Ring.1R R ; denom = Ring.1R R ; denomNonzero = IntegralDomain.nontrivial I} n) record { num = positiveEltPower (Ring.additiveGroup R) (Ring.1R R) n ; denom = Ring.1R R ; denomNonzero = IntegralDomain.nontrivial I }
|
||||
simpPower zero = Equivalence.reflexive (Setoid.eq fieldOfFractionsSetoid) {Group.0G fieldOfFractionsGroup}
|
||||
simpPower (succ n) = Equivalence.transitive (Setoid.eq fieldOfFractionsSetoid) {record { denomNonzero = d (fieldOfFractionsPlus (record { num = 1R ; denom = 1R ; denomNonzero = λ t → nonempty (symmetric t) }) (positiveEltPower fieldOfFractionsGroup _ n)) }} {record { denomNonzero = λ t → nonempty (symmetric (transitive (symmetric identIsIdent) t)) }} {record { denomNonzero = λ t → nonempty (symmetric t) }} (Group.+WellDefined fieldOfFractionsGroup {record { num = Ring.1R R ; denom = Ring.1R R ; denomNonzero = IntegralDomain.nontrivial I }} {positiveEltPower fieldOfFractionsGroup record { num = Ring.1R R ; denom = Ring.1R R ; denomNonzero = IntegralDomain.nontrivial I } n} {record { num = Ring.1R R ; denom = Ring.1R R ; denomNonzero = IntegralDomain.nontrivial I }} {record { num = positiveEltPower (Ring.additiveGroup R) (Ring.1R R) n ; denom = Ring.1R R ; denomNonzero = IntegralDomain.nontrivial I }} (Equivalence.reflexive (Setoid.eq fieldOfFractionsSetoid) {record { num = Ring.1R R ; denom = Ring.1R R ; denomNonzero = IntegralDomain.nontrivial I}}) (simpPower n)) (transitive (transitive (transitive *Commutative (transitive identIsIdent (+WellDefined identIsIdent identIsIdent))) (symmetric identIsIdent)) (*WellDefined (symmetric identIsIdent) reflexive))
|
||||
|
||||
lemma : (n : ℕ) {num denom : A} .(d!=0 : denom ∼ 0G → False) → (num * denom) < positiveEltPower additiveGroup 1R n → fieldOfFractionsComparison (record { num = num ; denom = denom ; denomNonzero = d!=0}) record { num = positiveEltPower additiveGroup (Ring.1R R) n ; denom = 1R ; denomNonzero = IntegralDomain.nontrivial I }
|
||||
lemma n {num} {denom} d!=0 numdenom<n with totality 0G denom
|
||||
... | inl (inl 0<denom) with totality 0G 1R
|
||||
... | inl (inl 0<1) = {!!}
|
||||
... | inl (inr x) = exFalso (1<0False x)
|
||||
... | inr x = exFalso (nonempty x)
|
||||
lemma n {num} {denom} d!=0 numdenom<n | inl (inr denom<0) with totality 0G 1R
|
||||
... | inl (inl 0<1) = {!!}
|
||||
... | inl (inr 1<0) = exFalso (1<0False 1<0)
|
||||
... | inr 0=1 = exFalso (nonempty 0=1)
|
||||
lemma n {num} {denom} d!=0 numdenom<n | inr 0=denom = exFalso (d!=0 (symmetric 0=denom))
|
||||
|
||||
fieldOfFractionsArchimedean : Archimedean (toGroup R pRing) → ArchimedeanField {F = fieldOfFractions} record { oRing = fieldOfFractionsOrderedRing }
|
||||
fieldOfFractionsArchimedean arch (record { num = num ; denom = denom ; denomNonzero = denom!=0 }) with arch (num * denom) {!!} {!!} {!!}
|
||||
... | bl = {!!}
|
||||
--... | N , pr = N , SetoidPartialOrder.<WellDefined fieldOfFractionsOrder {record { denomNonzero = denom!=0 }} {record { denomNonzero = denom!=0 }} {record { denomNonzero = λ t → nonempty (symmetric t) }} {record { denomNonzero = d (positiveEltPower fieldOfFractionsGroup record { num = 1R ; denom = 1R ; denomNonzero = λ t → nonempty (symmetric t) } N) }} (Equivalence.reflexive (Setoid.eq fieldOfFractionsSetoid) { record { denomNonzero = denom!=0 } }) (Equivalence.symmetric (Setoid.eq fieldOfFractionsSetoid) {record { denomNonzero = λ t → d (positiveEltPower fieldOfFractionsGroup record { num = 1R ; denom = 1R ; denomNonzero = λ t → nonempty (symmetric t) } N) t }} {record { denomNonzero = λ t → nonempty (symmetric t) }} (simpPower N)) (lemma N denom!=0 pr)
|
@@ -7,7 +7,6 @@ open import Rings.IntegralDomains.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
|
||||
module Fields.FieldOfFractions.Group {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) where
|
||||
|
||||
open import Fields.FieldOfFractions.Setoid I
|
||||
|
@@ -9,7 +9,8 @@ open import Rings.Orders.Total.Lemmas
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
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
|
||||
@@ -171,12 +172,7 @@ private
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = denomC!=0 }) a<b b<c | inl (inl 0<denomA) with totality (Ring.0R R) denomC
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = denomC!=0 }) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) with totality (Ring.0R R) denomB
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = denomC!=0 }) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inl x) with totality (Ring.0R R) denomC
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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
|
||||
inter : ((numA * denomB) * denomC) < ((numB * denomA) * denomC)
|
||||
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))
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = denomC!=0 }) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inl 0<denomB) | inl (inl _) = ringCanCancelPositive order 0<denomB (SetoidPartialOrder.<Transitive pOrder (SetoidPartialOrder.<WellDefined pOrder (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) reflexive (ringCanMultiplyByPositive pRing 0<denomC a<b)) (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)))
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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 (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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 (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = denomC!=0 }) a<b b<c | inl (inl 0<denomA) | inl (inl 0<denomC) | inl (inr denomB<0) with totality (Ring.0R R) denomC
|
||||
@@ -221,17 +217,11 @@ private
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = denomC!=0 }) a<b b<c | inl (inr denomA<0) | inl (inr denomC<0) with totality (Ring.0R R) denomB
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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 (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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 (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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) (swapLemma) (ringCanMultiplyByNegative pRing denomA<0 b<c)))
|
||||
where
|
||||
have : ((numA * denomC) * denomB) < ((numB * denomA) * denomC)
|
||||
have = SetoidPartialOrder.<WellDefined pOrder (swapLemma) reflexive (ringCanMultiplyByNegative pRing denomC<0 a<b)
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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 (SetoidPartialOrder.<WellDefined pOrder swapLemma reflexive (ringCanMultiplyByNegative pRing denomC<0 a<b)) (SetoidPartialOrder.<WellDefined pOrder (swapLemma) (swapLemma) (ringCanMultiplyByNegative pRing denomA<0 b<c)))
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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 (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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 (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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 (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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) (swapLemma) (ringCanMultiplyByNegative pRing denomA<0 b<c)) have)
|
||||
where
|
||||
have : ((numB * denomA) * denomC) < ((numA * denomC) * denomB)
|
||||
have = SetoidPartialOrder.<WellDefined pOrder reflexive (swapLemma) (ringCanMultiplyByNegative pRing denomC<0 a<b)
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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) (swapLemma) (ringCanMultiplyByNegative pRing denomA<0 b<c)) (SetoidPartialOrder.<WellDefined pOrder reflexive (swapLemma) (ringCanMultiplyByNegative pRing denomC<0 a<b)))
|
||||
<transitive (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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 (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = 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 (record { num = numA ; denom = denomA ; denomNonzero = denomA!=0 }) (record { num = numB ; denom = denomB ; denomNonzero = denomB!=0 }) (record { num = numC ; denom = denomC ; denomNonzero = denomC!=0 }) a<b b<c | inl (inr denomA<0) | inr x = exFalso (denomC!=0 (Equivalence.symmetric (Setoid.eq S) x))
|
||||
@@ -538,11 +528,15 @@ PartiallyOrderedRing.orderRespectsMultiplication (fieldOfFractionsPOrderedRing)
|
||||
fieldOfFractionsOrderedRing : TotallyOrderedRing fieldOfFractionsPOrderedRing
|
||||
TotallyOrderedRing.total fieldOfFractionsOrderedRing = fieldOfFractionsTotalOrder
|
||||
|
||||
private
|
||||
fieldOfFractionsOrderInherited' : {x y : A} → x < y → fieldOfFractionsComparison (embedIntoFieldOfFractions x) (embedIntoFieldOfFractions y)
|
||||
fieldOfFractionsOrderInherited' {x} {y} x<y with totality 0R 1R
|
||||
fieldOfFractionsOrderInherited' {x} {y} x<y | inl (inl 0<1) with totality 0R 1R
|
||||
fieldOfFractionsOrderInherited' {x} {y} x<y | inl (inl 0<1) | inl (inl _) = SetoidPartialOrder.<WellDefined pOrder (symmetric (transitive *Commutative identIsIdent)) (symmetric (transitive *Commutative identIsIdent)) x<y
|
||||
fieldOfFractionsOrderInherited' {x} {y} x<y | inl (inl 0<1) | inl (inr 1<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<1 1<0))
|
||||
fieldOfFractionsOrderInherited' {x} {y} x<y | inl (inl 0<1) | inr 0=1 = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<WellDefined pOrder 0=1 reflexive 0<1))
|
||||
fieldOfFractionsOrderInherited' {x} {y} x<y | inl (inr 1<0) = exFalso (1<0False order 1<0)
|
||||
fieldOfFractionsOrderInherited' {x} {y} x<y | inr 0=1 = exFalso (anyComparisonImpliesNontrivial pRing x<y 0=1)
|
||||
|
||||
fieldOfFractionsOrderInherited : {x y : A} → x < y → fieldOfFractionsComparison (embedIntoFieldOfFractions x) (embedIntoFieldOfFractions y)
|
||||
fieldOfFractionsOrderInherited {x} {y} x<y with totality 0R 1R
|
||||
fieldOfFractionsOrderInherited {x} {y} x<y | inl (inl 0<1) with totality 0R 1R
|
||||
fieldOfFractionsOrderInherited {x} {y} x<y | inl (inl 0<1) | inl (inl _) = SetoidPartialOrder.<WellDefined pOrder (symmetric (transitive *Commutative identIsIdent)) (symmetric (transitive *Commutative identIsIdent)) x<y
|
||||
fieldOfFractionsOrderInherited {x} {y} x<y | inl (inl 0<1) | inl (inr 1<0) = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<Transitive pOrder 0<1 1<0))
|
||||
fieldOfFractionsOrderInherited {x} {y} x<y | inl (inl 0<1) | inr 0=1 = exFalso (SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<WellDefined pOrder 0=1 reflexive 0<1))
|
||||
fieldOfFractionsOrderInherited {x} {y} x<y | inl (inr 1<0) = exFalso (1<0False order 1<0)
|
||||
fieldOfFractionsOrderInherited {x} {y} x<y | inr 0=1 = exFalso (anyComparisonImpliesNontrivial pRing x<y 0=1)
|
||||
fieldOfFractionsOrderInherited {x} {y} x<y = fieldOfFractionsOrderInherited' {x} {y} x<y
|
||||
|
@@ -5,7 +5,7 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Subset
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Fields.Fields
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Orders.Total.Lemmas
|
||||
|
@@ -3,7 +3,8 @@
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Subset
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
|
@@ -8,7 +8,8 @@ open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
@@ -121,6 +122,6 @@ abstract
|
||||
IntegralDomain.nontrivial orderedFieldIsIntDom pr = Field.nontrivial F (Equivalence.symmetric (Setoid.eq S) pr)
|
||||
|
||||
charZero : (n : ℕ) → (0R ∼ (fromN (succ n))) → False
|
||||
charZero n 0=sn = irreflexive (<WellDefined 0=sn reflexive (fromNPreservesOrder nontrivial (succIsPositive n)))
|
||||
charZero n 0=sn = irreflexive (<WellDefined 0=sn reflexive (fromNPreservesOrder (0<1 (Field.nontrivial F)) (succIsPositive n)))
|
||||
charZero' : (n : ℕ) → ((fromN (succ n)) ∼ 0R) → False
|
||||
charZero' n pr = charZero n (symmetric pr)
|
||||
|
@@ -5,7 +5,8 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Subset
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
@@ -26,6 +27,7 @@ open TotallyOrderedField oF
|
||||
open TotallyOrderedRing oRing
|
||||
open PartiallyOrderedRing pRing
|
||||
open import Rings.Orders.Total.Lemmas oRing
|
||||
open import Rings.Orders.Total.AbsoluteValue oRing
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open SetoidTotalOrder total
|
||||
open SetoidPartialOrder pOrder
|
||||
|
@@ -5,7 +5,8 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Subset
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
@@ -27,6 +28,7 @@ open TotallyOrderedField oF
|
||||
open TotallyOrderedRing oRing
|
||||
open PartiallyOrderedRing pRing
|
||||
open import Rings.Orders.Total.Lemmas oRing
|
||||
open import Rings.Orders.Total.AbsoluteValue oRing
|
||||
open import Rings.Orders.Partial.Lemmas pRing
|
||||
open SetoidTotalOrder total
|
||||
open SetoidPartialOrder pOrder
|
||||
|
@@ -3,7 +3,7 @@
|
||||
open import Rings.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Functions
|
||||
open import Fields.Fields
|
||||
|
||||
|
65
Fields/Orders/Total/Archimedean.agda
Normal file
65
Fields/Orders/Total/Archimedean.agda
Normal file
@@ -0,0 +1,65 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Rings.Definition
|
||||
open import Fields.Fields
|
||||
open import Fields.Orders.Total.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.Orders.Total.Archimedean {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A → A → A} {R : Ring S _+_ _*_} {c : _} {_<_ : A → A → Set c} {pOrder : SetoidPartialOrder S _<_} {p : PartiallyOrderedRing R pOrder} {F : Field R} (t : TotallyOrderedField F p) where
|
||||
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open import Groups.Lemmas additiveGroup
|
||||
open import Groups.Cyclic.Definition additiveGroup
|
||||
open import Groups.Cyclic.DefinitionLemmas additiveGroup
|
||||
open import Groups.Orders.Archimedean (toGroup R p)
|
||||
open import Rings.Orders.Partial.Lemmas p
|
||||
open import Rings.InitialRing R
|
||||
open SetoidPartialOrder pOrder
|
||||
open PartiallyOrderedRing p
|
||||
open SetoidTotalOrder (TotallyOrderedRing.total (TotallyOrderedField.oRing t))
|
||||
open import Rings.Orders.Total.Lemmas (TotallyOrderedField.oRing t)
|
||||
open Field F
|
||||
|
||||
ArchimedeanField : Set (a ⊔ c)
|
||||
ArchimedeanField = (x : A) → (0R < x) → Sg ℕ (λ n → x < (fromN n))
|
||||
|
||||
private
|
||||
lemma : (r : A) (N : ℕ) → (positiveEltPower 1R N * r) ∼ positiveEltPower r N
|
||||
lemma r zero = timesZero'
|
||||
lemma r (succ n) = transitive *DistributesOver+' (+WellDefined identIsIdent (lemma r n))
|
||||
|
||||
findBound : (r s : A) (N : ℕ) → (1R < r) → (s < positiveEltPower 1R N) → s < positiveEltPower r N
|
||||
findBound r s zero 1<r s<N = s<N
|
||||
findBound r s (succ N) 1<r s<N = <Transitive s<N (<WellDefined (transitive *Commutative identIsIdent) (lemma r (succ N)) (ringCanMultiplyByPositive' (fromNPreservesOrder (0<1 (anyComparisonImpliesNontrivial 1<r)) (succIsPositive N)) 1<r))
|
||||
|
||||
archFieldToGrp : ArchimedeanField → Archimedean
|
||||
archFieldToGrp a r s 0<r 0<s with totality r s
|
||||
archFieldToGrp a r s 0<r 0<s | inl (inr s<r) = 1 , <WellDefined reflexive (symmetric identRight) s<r
|
||||
archFieldToGrp a r s 0<r 0<s | inr r=s = 2 , <WellDefined (transitive identLeft r=s) (+WellDefined reflexive (symmetric identRight)) (orderRespectsAddition 0<r r)
|
||||
... | inl (inl r<s) with allInvertible r (λ r=0 → irreflexive (<WellDefined reflexive r=0 0<r))
|
||||
... | inv , prInv with a inv (reciprocalPositive r inv 0<r (transitive *Commutative prInv))
|
||||
... | N , invBound with a s 0<s
|
||||
... | Ns , nSBound = (Ns *N N) , <WellDefined reflexive (symmetric (elementPowerMultiplies (nonneg Ns) (nonneg N) r)) (findBound (positiveEltPower r N) s Ns m nSBound)
|
||||
where
|
||||
m : 1R < positiveEltPower r N
|
||||
m = <WellDefined prInv (lemma r N) (ringCanMultiplyByPositive 0<r invBound)
|
||||
|
||||
archToArchField : Archimedean → ArchimedeanField
|
||||
archToArchField arch a 0<a with arch 1R a (0<1 (anyComparisonImpliesNontrivial 0<a)) 0<a
|
||||
... | N , pr = N , pr
|
@@ -4,7 +4,7 @@ open import Rings.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Functions
|
||||
open import Fields.Fields
|
||||
|
||||
|
@@ -4,7 +4,8 @@ open import Rings.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import Fields.Fields
|
||||
open import Fields.Orders.Total.Definition
|
||||
@@ -38,7 +39,7 @@ charNotN : (n : ℕ) → fromN (succ n) ∼ 0R → False
|
||||
charNotN n pr = irreflexive (<WellDefined reflexive pr t)
|
||||
where
|
||||
t : 0R < fromN (succ n)
|
||||
t = fromNPreservesOrder (Field.nontrivial F) (succIsPositive n)
|
||||
t = fromNPreservesOrder (0<1 (Field.nontrivial F)) (succIsPositive n)
|
||||
|
||||
charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False
|
||||
charNot2 pr = charNotN 1 (transitive (transitive +Associative identRight) pr)
|
||||
|
@@ -7,13 +7,13 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Integers
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
|
||||
module Groups.Cyclic.Definition {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} (G : Group S _·_) where
|
||||
|
||||
open Setoid S
|
||||
open Group G
|
||||
open Equivalence eq
|
||||
open import Groups.Lemmas G
|
||||
|
||||
positiveEltPower : (x : A) (i : ℕ) → A
|
||||
positiveEltPower x 0 = Group.0G G
|
||||
@@ -40,9 +40,10 @@ elementPowerWellDefinedZ' i j i=j {g} = identityOfIndiscernablesRight _∼_ refl
|
||||
|
||||
elementPowerWellDefinedG : (g h : A) → (g ∼ h) → {n : ℤ} → (elementPower g n) ∼ (elementPower h n)
|
||||
elementPowerWellDefinedG g h g=h {nonneg n} = positiveEltPowerWellDefinedG g h g=h n
|
||||
elementPowerWellDefinedG g h g=h {negSucc x} = inverseWellDefined G (+WellDefined g=h (positiveEltPowerWellDefinedG g h g=h x))
|
||||
elementPowerWellDefinedG g h g=h {negSucc x} = inverseWellDefined (+WellDefined g=h (positiveEltPowerWellDefinedG g h g=h x))
|
||||
|
||||
record CyclicGroup : Set (m ⊔ n) where
|
||||
field
|
||||
generator : A
|
||||
cyclic : {a : A} → (Sg ℤ (λ i → Setoid._∼_ S (elementPower generator i) a))
|
||||
|
||||
|
@@ -15,15 +15,18 @@ module Groups.FreeProduct.Addition {i : _} {I : Set i} (decidableIndex : (x y :
|
||||
open import Groups.FreeProduct.Definition decidableIndex decidableGroups G
|
||||
open import Groups.FreeProduct.Setoid decidableIndex decidableGroups G
|
||||
|
||||
prependEqual : (i : I) (g : A i) .(nonZero : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → (j : I) → i ≡ j → (w : ReducedSequenceBeginningWith j) → ReducedSequence
|
||||
prependEqual i g nonzero .i refl (ofEmpty .i h nonZero) with decidableGroups i (_+_ i g h) (Group.0G (G i))
|
||||
prependEqual i g nonzero .i refl (ofEmpty .i h nonZero) | inl sumZero = empty
|
||||
prependEqual i g nonzero .i refl (ofEmpty .i h nonZero) | inr sumNotZero = nonempty i (ofEmpty i (_+_ i g h) sumNotZero)
|
||||
prependEqual i g nonzero .i refl (prependLetter .i h nonZero x x₁) with decidableGroups i (_+_ i g h) (Group.0G (G i))
|
||||
prependEqual i g nonzero .i refl (prependLetter .i h nonZero {j} x x₁) | inl sumZero = nonempty j x
|
||||
prependEqual i g nonzero .i refl (prependLetter .i h nonZero x pr) | inr sumNotZero = nonempty i (prependLetter i (_+_ i g h) sumNotZero x pr)
|
||||
|
||||
prepend : (i : I) → (g : A i) → .(nonZero : (Setoid._∼_ (S i) g (Group.0G (G i))) → False) → ReducedSequence → ReducedSequence
|
||||
prepend i g nonzero empty = nonempty i (ofEmpty i g nonzero)
|
||||
prepend i g nonzero (nonempty j x) with decidableIndex i j
|
||||
prepend i g nonzero (nonempty .i (ofEmpty .i h nonZero)) | inl refl with decidableGroups i (_+_ i g h) (Group.0G (G i))
|
||||
prepend i g nonzero (nonempty .i (ofEmpty .i h nonZero)) | inl refl | inl sumZero = empty
|
||||
prepend i g nonzero (nonempty .i (ofEmpty .i h nonZero)) | inl refl | inr sumNotZero = nonempty i (ofEmpty i (_+_ i g h) sumNotZero)
|
||||
prepend i g nonzero (nonempty .i (prependLetter .i h nonZero x x₁)) | inl refl with decidableGroups i (_+_ i g h) (Group.0G (G i))
|
||||
prepend i g nonzero (nonempty .i (prependLetter .i h nonZero {j} x x₁)) | inl refl | inl sumZero = nonempty j x
|
||||
prepend i g nonzero (nonempty .i (prependLetter .i h nonZero x pr)) | inl refl | inr sumNotZero = nonempty i (prependLetter i (_+_ i g h) sumNotZero x pr)
|
||||
prepend i g nonzero (nonempty j w) | inl i=j = prependEqual i g nonzero j i=j w
|
||||
prepend i g nonzero (nonempty j x) | inr i!=j = nonempty i (prependLetter i g nonzero x i!=j)
|
||||
|
||||
plus' : {j : _} → (ReducedSequenceBeginningWith j) → ReducedSequence → ReducedSequence
|
||||
|
@@ -1,5 +1,6 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Setoids.Setoids
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Groups.Definition
|
||||
@@ -12,6 +13,8 @@ record GroupHom {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A →
|
||||
field
|
||||
groupHom : {x y : A} → f (x ·A y) ∼ (f x) ·B (f y)
|
||||
wellDefined : {x y : A} → Setoid._∼_ S x y → f x ∼ f y
|
||||
groupHom' : {x y : A} → (f x) ·B (f y) ∼ f (x ·A y)
|
||||
groupHom' = Equivalence.symmetric eq groupHom
|
||||
|
||||
record InjectiveGroupHom {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_} {underf : A → B} (f : GroupHom G H underf) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
open Setoid S renaming (_∼_ to _∼A_)
|
||||
|
@@ -42,3 +42,9 @@ zeroImpliesInverseZero {x} fx=0 = transitive homRespectsInverse (transitive (inv
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
homRespectsInverse' : {a b : A} → Setoid._∼_ T (f (Group.inverse G a) +B f (Group.inverse G b)) (Group.inverse H (f (b +A a)))
|
||||
homRespectsInverse' {a} {b} = transitive (symmetric (GroupHom.groupHom hom)) (transitive (GroupHom.wellDefined hom (Equivalence.symmetric (Setoid.eq S) (invContravariant G))) (homRespectsInverse))
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
@@ -5,7 +5,6 @@ open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
|
20
Groups/Orders/Archimedean.agda
Normal file
20
Groups/Orders/Archimedean.agda
Normal file
@@ -0,0 +1,20 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Numbers.Naturals.Definition
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Orders.Archimedean {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {G : Group S _+_} {c : _} {_<_ : A → A → Set c} {pOrder : SetoidPartialOrder S _<_} (p : PartiallyOrderedGroup G pOrder) where
|
||||
|
||||
open Setoid S
|
||||
open import Groups.Cyclic.Definition G
|
||||
open Group G
|
||||
|
||||
Archimedean : Set (a ⊔ c)
|
||||
Archimedean = (x y : A) → (0G < x) → (0G < y) → Sg ℕ (λ n → y < (positiveEltPower x n))
|
17
Groups/Orders/Partial/Definition.agda
Normal file
17
Groups/Orders/Partial/Definition.agda
Normal file
@@ -0,0 +1,17 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Definition
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Orders.Partial.Definition {n m : _} {A : Set n} {S : Setoid {n} {m} A} {_+_ : A → A → A} (G : Group S _+_) where
|
||||
|
||||
open Group G
|
||||
open Setoid S
|
||||
|
||||
record PartiallyOrderedGroup {p : _} {_<_ : Rel {_} {p} A} (pOrder : SetoidPartialOrder S _<_) : Set (lsuc n ⊔ m ⊔ p) where
|
||||
field
|
||||
orderRespectsAddition : {a b : A} → (a < b) → (c : A) → (a + c) < (b + c)
|
@@ -1,16 +1,17 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
|
||||
module Lists.Definition {a : _} where
|
||||
module Lists.Definition where
|
||||
|
||||
data List (A : Set a) : Set a where
|
||||
data List {a : _} (A : Set a) : Set a where
|
||||
[] : List A
|
||||
_::_ : (x : A) (xs : List A) → List A
|
||||
infixr 10 _::_
|
||||
{-# BUILTIN LIST List #-}
|
||||
|
||||
[_] : {A : Set a} → (a : A) → List A
|
||||
[_] : {a : _} {A : Set a} → (a : A) → List A
|
||||
[ a ] = a :: []
|
||||
|
||||
_++_ : {A : Set a} → List A → List A → List A
|
||||
_++_ : {a : _} {A : Set a} → List A → List A → List A
|
||||
[] ++ m = m
|
||||
(x :: l) ++ m = x :: (l ++ m)
|
||||
|
@@ -14,6 +14,7 @@ data False : Set where
|
||||
data False' {a : _} : Set a where
|
||||
|
||||
record True : Set where
|
||||
{-# BUILTIN UNIT True #-}
|
||||
record True' {a : _} : Set a where
|
||||
|
||||
infix 10 _||_
|
||||
@@ -24,6 +25,9 @@ data _||_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where
|
||||
data Bool : Set where
|
||||
BoolTrue : Bool
|
||||
BoolFalse : Bool
|
||||
{-# BUILTIN BOOL Bool #-}
|
||||
{-# BUILTIN TRUE BoolTrue #-}
|
||||
{-# BUILTIN FALSE BoolFalse #-}
|
||||
|
||||
infix 15 _&&_
|
||||
record _&&_ {a b} (A : Set a) (B : Set b) : Set (a ⊔ b) where
|
||||
|
@@ -4,7 +4,6 @@ open import Numbers.ClassicalReals.RealField
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Subset
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
@@ -14,6 +13,8 @@ open import Groups.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Naturals.Order.Lemmas
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
|
||||
module Numbers.ClassicalReals.Examples (ℝ : RealField) where
|
||||
|
||||
@@ -45,7 +46,7 @@ squarePositive a | inr 0=a = inr (symmetric 0=a)
|
||||
|
||||
fraction<1 : (n m : ℕ) → n <N m → (0<m : fromN m ∼ 0G → False) → ((fromN n) * underlying (allInvertible (fromN m) 0<m)) < 1R
|
||||
fraction<1 n m n<m 0<m with allInvertible (fromN m) 0<m
|
||||
... | 1/m , prM = <WellDefined reflexive (transitive *Commutative prM) (ringCanMultiplyByPositive (inversePositiveIsPositive prM (fromNPreservesOrder nontrivial {0} {m} (zeroLeast n<m))) (fromNPreservesOrder nontrivial n<m))
|
||||
... | 1/m , prM = <WellDefined reflexive (transitive *Commutative prM) (ringCanMultiplyByPositive (inversePositiveIsPositive prM (fromNPreservesOrder (0<1 nontrivial) {0} {m} (zeroLeast n<m))) (fromNPreservesOrder (0<1 nontrivial) n<m))
|
||||
|
||||
1/2 : A
|
||||
1/2 = underlying (allInvertible (1R + 1R) (orderedImpliesCharNot2 nontrivial))
|
||||
@@ -64,13 +65,13 @@ pr1/2' = transitive (symmetric (transitive *DistributesOver+ (+WellDefined (tran
|
||||
3/2<2 = <WellDefined (transitive (+WellDefined reflexive (transitive (symmetric pr1/2) (transitive *Commutative (*WellDefined (+WellDefined reflexive (symmetric identRight)) reflexive)))) (symmetric *DistributesOver+')) reflexive (orderRespectsAddition (<WellDefined (symmetric identIsIdent) reflexive 1/2<1) 1R)
|
||||
|
||||
2<9/4 : (1R + 1R) < ((fromN 3 * 1/2) * (fromN 3 * 1/2))
|
||||
2<9/4 = <WellDefined reflexive (symmetric *Associative) (halveInequality _ _ 1/2 pr1/2' (<WellDefined reflexive (transitive (symmetric *Associative) *Commutative) (halveInequality _ _ 1/2 pr1/2' (<WellDefined (transitive (fromNPreserves* 4 2) (transitive (*WellDefined (+WellDefined reflexive (+WellDefined reflexive (+WellDefined reflexive identRight))) (+WellDefined reflexive identRight)) (transitive *DistributesOver+ (+WellDefined (transitive *Commutative (transitive identIsIdent +Associative)) (transitive *Commutative (transitive identIsIdent +Associative)))))) (fromNPreserves* 3 3) (fromNPreservesOrder nontrivial (le {8} 0 refl))))))
|
||||
2<9/4 = <WellDefined reflexive (symmetric *Associative) (halveInequality _ _ 1/2 pr1/2' (<WellDefined reflexive (transitive (symmetric *Associative) *Commutative) (halveInequality _ _ 1/2 pr1/2' (<WellDefined (transitive (fromNPreserves* 4 2) (transitive (*WellDefined (+WellDefined reflexive (+WellDefined reflexive (+WellDefined reflexive identRight))) (+WellDefined reflexive identRight)) (transitive *DistributesOver+ (+WellDefined (transitive *Commutative (transitive identIsIdent +Associative)) (transitive *Commutative (transitive identIsIdent +Associative)))))) (fromNPreserves* 3 3) (fromNPreservesOrder (0<1 nontrivial) (le {8} 0 refl))))))
|
||||
|
||||
0<2 : 0R < (1R + 1R)
|
||||
0<2 = <WellDefined identLeft reflexive (ringAddInequalities (0<1 nontrivial) (0<1 nontrivial))
|
||||
|
||||
0<3/2 : 0R < (fromN 3 * 1/2)
|
||||
0<3/2 = orderRespectsMultiplication (fromNPreservesOrder nontrivial (le {0} {3} 2 refl)) (inversePositiveIsPositive pr1/2 0<2)
|
||||
0<3/2 = orderRespectsMultiplication (fromNPreservesOrder (0<1 nontrivial) (le {0} {3} 2 refl)) (inversePositiveIsPositive pr1/2 0<2)
|
||||
|
||||
3/2ub : {r : A} → (r * r) < (1R + 1R) → (r < (fromN 3 * 1/2))
|
||||
3/2ub {r} r^2<2 with totality r (fromN 3 * 1/2)
|
||||
@@ -80,10 +81,10 @@ pr1/2' = transitive (symmetric (transitive *DistributesOver+ (+WellDefined (tran
|
||||
|
||||
2/6<1 : ((1R + 1R) * underlying (allInvertible (fromN 6) (charZero' 5))) < 1R
|
||||
2/6<1 with allInvertible (fromN 6) (charZero' 5)
|
||||
2/6<1 | 1/6 , pr1/6 = <WellDefined reflexive (transitive *Commutative pr1/6) (ringCanMultiplyByPositive (inversePositiveIsPositive pr1/6 (fromNPreservesOrder nontrivial {0} {6} (le 5 refl))) (<WellDefined (+WellDefined reflexive identRight) reflexive (fromNPreservesOrder nontrivial {2} {6} (le 3 refl))))
|
||||
2/6<1 | 1/6 , pr1/6 = <WellDefined reflexive (transitive *Commutative pr1/6) (ringCanMultiplyByPositive (inversePositiveIsPositive pr1/6 (fromNPreservesOrder (0<1 nontrivial) {0} {6} (le 5 refl))) (<WellDefined (+WellDefined reflexive identRight) reflexive (fromNPreservesOrder (0<1 nontrivial) {2} {6} (le 3 refl))))
|
||||
|
||||
2<2*2 : (1R + 1R) < ((1R + 1R) * (1R + 1R))
|
||||
2<2*2 = (<WellDefined (+WellDefined reflexive identRight) (transitive (fromNPreserves* 2 2) (*WellDefined (+WellDefined reflexive identRight) (+WellDefined reflexive identRight))) (fromNPreservesOrder nontrivial {2} {4} (le 1 refl)))
|
||||
2<2*2 = (<WellDefined (+WellDefined reflexive identRight) (transitive (fromNPreserves* 2 2) (*WellDefined (+WellDefined reflexive identRight) (+WellDefined reflexive identRight))) (fromNPreservesOrder (0<1 nontrivial) {2} {4} (le 1 refl)))
|
||||
|
||||
square<2Means<2 : (u : A) → (u * u) < (1R + 1R) → u < (1R + 1R)
|
||||
square<2Means<2 u u^2<2 with totality u (1R + 1R)
|
||||
@@ -137,7 +138,7 @@ sqrt2 = sqrt2' , sqrt2IsSqrt2
|
||||
crudeBound : isInInterval (sqrt2' * sqrt2') record { minBound = 0R ; maxBound = 1R + 1R }
|
||||
crudeBound with squarePositive sqrt2'
|
||||
crudeBound | inl (_ ,, snd) = snd ,, sup^2<2
|
||||
crudeBound | inr sqrt2=0 with upperBound 1R (<WellDefined (transitive identRight (symmetric identIsIdent)) (+WellDefined reflexive identRight) (fromNPreservesOrder nontrivial {1} {2} (le 0 refl)))
|
||||
crudeBound | inr sqrt2=0 with upperBound 1R (<WellDefined (transitive identRight (symmetric identIsIdent)) (+WellDefined reflexive identRight) (fromNPreservesOrder (0<1 nontrivial) {1} {2} (le 0 refl)))
|
||||
crudeBound | inr sqrt2=0 | inl 1<sqrt2 = exFalso (irreflexive (<Transitive 1<sqrt2 (<WellDefined (symmetric sqrt2=0) reflexive (0<1 nontrivial))))
|
||||
crudeBound | inr sqrt2=0 | inr 1=sqrt2 = exFalso (nontrivial (transitive (symmetric sqrt2=0) (symmetric 1=sqrt2)))
|
||||
crudeBound' : isInInterval (inverse (sqrt2' * sqrt2')) record { minBound = inverse (1R + 1R) ; maxBound = inverse 0R }
|
||||
@@ -147,7 +148,7 @@ sqrt2 = sqrt2' , sqrt2IsSqrt2
|
||||
numeratorBound' : isInInterval (fromN 2 + inverse (sqrt2' * sqrt2')) record { minBound = 0R ; maxBound = 1R + 1R }
|
||||
numeratorBound' = intervalWellDefined' groupIsAbelian numeratorBound
|
||||
tBound : isInInterval t record { minBound = 0R ; maxBound = (1R + 1R) * (underlying (allInvertible (fromN 6) (charZero' 5))) }
|
||||
tBound = intervalWellDefined (transitive *Commutative timesZero ,, reflexive) (intervalConstantProductContains (inversePositiveIsPositive pr' (fromNPreservesOrder nontrivial {0} {6} (le 5 refl))) numeratorBound')
|
||||
tBound = intervalWellDefined (transitive *Commutative timesZero ,, reflexive) (intervalConstantProductContains (inversePositiveIsPositive pr' (fromNPreservesOrder (0<1 nontrivial) {0} {6} (le 5 refl))) numeratorBound')
|
||||
u : ((((fromN 5) * ((fromN 2) + inverse (sqrt2' * sqrt2'))) * underlying (allInvertible (fromN 6) (charZero' 5))) + (sqrt2' * sqrt2')) < ((fromN 2 + inverse (sqrt2' * sqrt2')) + (sqrt2' * sqrt2'))
|
||||
u = orderRespectsAddition (<WellDefined (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) identIsIdent (ringCanMultiplyByPositive {(fromN 5) * underlying (allInvertible (fromN 6) (charZero' 5))} {1R} {fromN 2 + inverse (sqrt2' * sqrt2')} (moveInequality (<WellDefined reflexive (transitive (symmetric identRight) (symmetric +Associative)) sup^2<2)) (fraction<1 5 6 (le zero refl) λ pr → charZero' 5 pr))) (sqrt2' * sqrt2')
|
||||
tBound' : (((fromN 5) * t) + (sqrt2' * sqrt2')) < (1R + 1R)
|
||||
@@ -171,9 +172,9 @@ sqrt2 = sqrt2' , sqrt2IsSqrt2
|
||||
where
|
||||
abstract
|
||||
1<sqrt2 : 1R < sqrt2'
|
||||
1<sqrt2 with upperBound 1R (<WellDefined (transitive identRight (symmetric identIsIdent)) (+WellDefined reflexive identRight) (fromNPreservesOrder nontrivial {1} {2} (le zero refl)))
|
||||
1<sqrt2 with upperBound 1R (<WellDefined (transitive identRight (symmetric identIsIdent)) (+WellDefined reflexive identRight) (fromNPreservesOrder (0<1 nontrivial) {1} {2} (le zero refl)))
|
||||
1<sqrt2 | inl x = x
|
||||
1<sqrt2 | inr x = exFalso (irreflexive (<Transitive 2<sup^2 (<WellDefined (transitive identRight (transitive (symmetric identIsIdent) (*WellDefined x x))) (+WellDefined reflexive identRight) (fromNPreservesOrder nontrivial {1} {2} (le 0 refl)))))
|
||||
1<sqrt2 | inr x = exFalso (irreflexive (<Transitive 2<sup^2 (<WellDefined (transitive identRight (transitive (symmetric identIsIdent) (*WellDefined x x))) (+WellDefined reflexive identRight) (fromNPreservesOrder (0<1 nontrivial) {1} {2} (le 0 refl)))))
|
||||
0<sqrt2 : 0R < sqrt2'
|
||||
0<sqrt2 = <Transitive (0<1 nontrivial) 1<sqrt2
|
||||
sqrt2<2 : sqrt2' < (1R + 1R)
|
||||
@@ -189,10 +190,10 @@ sqrt2 = sqrt2' , sqrt2IsSqrt2
|
||||
... | 1/4 , pr1/4 = transitive (transitive (transitive (transitive (symmetric ringMinusExtracts) (transitive (transitive (*WellDefined reflexive (transitive (symmetric ringMinusExtracts') (transitive *Commutative (*WellDefined reflexive (transitive invContravariant (transitive groupIsAbelian (+WellDefined reflexive invInv))))))) *Associative) *Commutative)) (*WellDefined reflexive (transitive *Commutative pr1/4))) *Commutative) identIsIdent
|
||||
t<sqrt2 : t < sqrt2'
|
||||
t<sqrt2 with allInvertible (fromN 4) (charZero' 3)
|
||||
... | 1/4 , pr4 = <WellDefined reflexive (transitive (symmetric *Associative) (transitive (*WellDefined reflexive (transitive *Commutative pr4)) (transitive *Commutative identIsIdent))) (ringCanMultiplyByPositive (inversePositiveIsPositive pr4 (fromNPreservesOrder nontrivial {0} {4} (le 3 refl))) (<Transitive (orderRespectsAddition (ringMultiplyPositives 0<sqrt2 0<sqrt2 sqrt2<2 sqrt2<2) (inverse (1R + (1R + 0G)))) (<WellDefined {fromN 2} 2=2*2-2 reflexive (<Transitive {fromN 2} {sqrt2' * fromN 2} (<WellDefined identIsIdent reflexive (ringCanMultiplyByPositive (fromNPreservesOrder nontrivial {0} (le 1 refl)) 1<sqrt2)) (<WellDefined *Commutative *Commutative (ringCanMultiplyByPositive 0<sqrt2 (fromNPreservesOrder nontrivial {2} (le 1 refl))))))))
|
||||
... | 1/4 , pr4 = <WellDefined reflexive (transitive (symmetric *Associative) (transitive (*WellDefined reflexive (transitive *Commutative pr4)) (transitive *Commutative identIsIdent))) (ringCanMultiplyByPositive (inversePositiveIsPositive pr4 (fromNPreservesOrder (0<1 nontrivial) {0} {4} (le 3 refl))) (<Transitive (orderRespectsAddition (ringMultiplyPositives 0<sqrt2 0<sqrt2 sqrt2<2 sqrt2<2) (inverse (1R + (1R + 0G)))) (<WellDefined {fromN 2} 2=2*2-2 reflexive (<Transitive {fromN 2} {sqrt2' * fromN 2} (<WellDefined identIsIdent reflexive (ringCanMultiplyByPositive (fromNPreservesOrder (0<1 nontrivial) {0} (le 1 refl)) 1<sqrt2)) (<WellDefined *Commutative *Commutative (ringCanMultiplyByPositive 0<sqrt2 (fromNPreservesOrder (0<1 nontrivial) {2} (le 1 refl))))))))
|
||||
0<t : 0R < t
|
||||
0<t with allInvertible (fromN 4) (charZero' 3)
|
||||
... | 1/4 , pr4 = orderRespectsMultiplication (<WellDefined reflexive (+WellDefined reflexive (inverseWellDefined (+WellDefined reflexive (symmetric identRight)))) (moveInequality 2<sup^2)) (inversePositiveIsPositive pr4 (fromNPreservesOrder nontrivial {0} {4} (le 3 refl)))
|
||||
... | 1/4 , pr4 = orderRespectsMultiplication (<WellDefined reflexive (+WellDefined reflexive (inverseWellDefined (+WellDefined reflexive (symmetric identRight)))) (moveInequality 2<sup^2)) (inversePositiveIsPositive pr4 (fromNPreservesOrder (0<1 nontrivial) {0} {4} (le 3 refl)))
|
||||
anotherUpperBound : A
|
||||
anotherUpperBound = (sqrt2' + inverse t) * (sqrt2' + inverse t)
|
||||
abstract
|
||||
|
@@ -5,7 +5,7 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Subset
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Fields.Fields
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Orders.Total.Lemmas
|
||||
|
@@ -6,7 +6,7 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Subset
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Fields.Fields
|
||||
open import Rings.Orders.Total.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
|
@@ -8,6 +8,9 @@ module Numbers.Integers.Definition where
|
||||
data ℤ : Set where
|
||||
nonneg : ℕ → ℤ
|
||||
negSucc : ℕ → ℤ
|
||||
{-# BUILTIN INTEGER ℤ #-}
|
||||
{-# BUILTIN INTEGERPOS nonneg #-}
|
||||
{-# BUILTIN INTEGERNEGSUC negSucc #-}
|
||||
|
||||
nonnegInjective : {a b : ℕ} → nonneg a ≡ nonneg b → a ≡ b
|
||||
nonnegInjective refl = refl
|
||||
|
@@ -7,9 +7,9 @@ open import Numbers.Integers.RingStructure.Ring
|
||||
open import Semirings.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Setoids.Orders
|
||||
open import Orders.Total.Definition
|
||||
open import Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
|
||||
module Numbers.Integers.Order where
|
||||
|
||||
|
30
Numbers/Integers/RingStructure/Archimedean.agda
Normal file
30
Numbers/Integers/RingStructure/Archimedean.agda
Normal file
@@ -0,0 +1,30 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Semirings.Definition
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Numbers.Integers.RingStructure.Ring
|
||||
open import Groups.Orders.Archimedean
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Numbers.Integers.Order
|
||||
|
||||
module Numbers.Integers.RingStructure.Archimedean where
|
||||
|
||||
open import Groups.Cyclic.Definition ℤGroup
|
||||
open import Semirings.Solver ℕSemiring multiplicationNIsCommutative
|
||||
|
||||
private
|
||||
lemma : (x y : ℕ) → positiveEltPower (nonneg x) y ≡ nonneg (x *N y)
|
||||
lemma x zero rewrite Semiring.productZeroRight ℕSemiring x = refl
|
||||
lemma x (succ y) rewrite lemma x y | multiplicationNIsCommutative x (succ y) | multiplicationNIsCommutative y x = equalityCommutative (+Zinherits x (x *N y))
|
||||
|
||||
ℤArchimedean : Archimedean (toGroup ℤRing ℤPOrderedRing)
|
||||
ℤArchimedean (nonneg (succ a)) (nonneg (succ b)) 0<a 0<b = succ (succ b) , t
|
||||
where
|
||||
v : b +N (a +N (a +N a *N b)) ≡ a +N (a +N (b +N a *N b))
|
||||
v rewrite Semiring.+Associative ℕSemiring a a (a *N b) | Semiring.+Associative ℕSemiring b (a +N a) (a *N b) | Semiring.+Associative ℕSemiring a b (a *N b) | Semiring.+Associative ℕSemiring a (a +N b) (a *N b) | Semiring.commutative ℕSemiring b (a +N a) | Semiring.+Associative ℕSemiring a a b = refl
|
||||
u : succ ((a +N (a +N a *N b)) +N b) ≡ a +N succ (a +N (b +N a *N b))
|
||||
u = from (succ (plus (plus (const a) (plus (const a) (times (const a) (const b)))) (const b))) to (plus (const a) (succ (plus (const a) (plus (const b) (times (const a) (const b)))))) by applyEquality succ v
|
||||
t : nonneg (succ b) <Z nonneg (succ a) +Z (nonneg (succ a) +Z positiveEltPower (nonneg (succ a)) b)
|
||||
t rewrite lemma (succ a) b = lessInherits (succPreservesInequality (le (a +N (a +N a *N b)) u))
|
@@ -4,7 +4,7 @@ open import Numbers.ClassicalReals.RealField
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Subset
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Definition
|
||||
|
@@ -4,7 +4,7 @@ open import Numbers.ClassicalReals.RealField
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Subset
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Definition
|
||||
|
@@ -9,6 +9,7 @@ infix 15 _+N_
|
||||
_+N_ : ℕ → ℕ → ℕ
|
||||
zero +N y = y
|
||||
succ x +N y = succ (x +N y)
|
||||
{-# BUILTIN NATPLUS _+N_ #-}
|
||||
|
||||
addZeroRight : (x : ℕ) → (x +N zero) ≡ x
|
||||
addZeroRight zero = refl
|
||||
|
@@ -10,6 +10,7 @@ infix 25 _*N_
|
||||
_*N_ : ℕ → ℕ → ℕ
|
||||
zero *N y = zero
|
||||
(succ x) *N y = y +N (x *N y)
|
||||
{-# BUILTIN NATTIMES _*N_ #-}
|
||||
|
||||
productZeroIsZeroLeft : (a : ℕ) → (zero *N a ≡ zero)
|
||||
productZeroIsZeroLeft a = refl
|
||||
|
@@ -12,7 +12,8 @@ open import Rings.Orders.Total.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Fields.Fields
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
|
@@ -16,13 +16,13 @@ open import Rings.Orders.Total.Definition
|
||||
open import Fields.Fields
|
||||
open import Numbers.Primes.PrimeNumbers
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Rationals.Definition
|
||||
open import Semirings.Definition
|
||||
open import Orders.Total.Definition
|
||||
open import Orders.WellFounded.Induction
|
||||
open import Setoids.Orders.Total.Definition
|
||||
|
||||
module Numbers.Rationals.Lemmas where
|
||||
|
||||
@@ -30,6 +30,7 @@ open import Semirings.Lemmas ℕSemiring
|
||||
|
||||
open PartiallyOrderedRing ℤPOrderedRing
|
||||
open import Rings.Orders.Total.Lemmas ℤOrderedRing
|
||||
open import Rings.Orders.Total.AbsoluteValue ℤOrderedRing
|
||||
open SetoidTotalOrder (totalOrderToSetoidTotalOrder ℤOrder)
|
||||
|
||||
evenOrOdd : (a : ℕ) → (Sg ℕ (λ i → i *N 2 ≡ a)) || (Sg ℕ (λ i → succ (i *N 2) ≡ a))
|
||||
|
@@ -1,6 +1,6 @@
|
||||
{-# OPTIONS --warning=error --safe --guardedness --without-K #-}
|
||||
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Rings.Definition
|
||||
|
@@ -1,6 +1,5 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
|
||||
open import LogicalFormulae
|
||||
|
||||
open import Orders.Total.Definition
|
||||
|
@@ -5,7 +5,6 @@ open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
|
@@ -25,10 +25,10 @@ open Setoid S
|
||||
open Equivalence eq
|
||||
open import Rings.Lemmas R
|
||||
open import Groups.Lemmas additiveGroup
|
||||
open import Groups.Cyclic.Definition additiveGroup
|
||||
|
||||
fromN : ℕ → A
|
||||
fromN zero = 0R
|
||||
fromN (succ n) = 1R + fromN n
|
||||
fromN = positiveEltPower 1R
|
||||
|
||||
fromNPreserves+ : (a b : ℕ) → fromN (a +N b) ∼ (fromN a) + (fromN b)
|
||||
fromNPreserves+ zero b = symmetric identLeft
|
||||
|
@@ -7,7 +7,7 @@ open import Rings.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
|
@@ -1,8 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Definition
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Functions
|
||||
open import Rings.Definition
|
||||
|
||||
@@ -13,9 +13,13 @@ module Rings.Orders.Partial.Definition {n m : _} {A : Set n} {S : Setoid {n} {m}
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open Setoid S
|
||||
open import Groups.Orders.Partial.Definition
|
||||
|
||||
record PartiallyOrderedRing {p : _} {_<_ : Rel {_} {p} A} (pOrder : SetoidPartialOrder S _<_) : Set (lsuc n ⊔ m ⊔ p) where
|
||||
field
|
||||
orderRespectsAddition : {a b : A} → (a < b) → (c : A) → (a + c) < (b + c)
|
||||
orderRespectsMultiplication : {a b : A} → (0R < a) → (0R < b) → (0R < (a * b))
|
||||
open SetoidPartialOrder pOrder
|
||||
|
||||
toGroup : {p : _} {_<_ : Rel {_} {p} A} {pOrder : SetoidPartialOrder S _<_} → PartiallyOrderedRing pOrder → PartiallyOrderedGroup additiveGroup pOrder
|
||||
PartiallyOrderedGroup.orderRespectsAddition (toGroup p) = PartiallyOrderedRing.orderRespectsAddition p
|
||||
|
@@ -3,7 +3,7 @@
|
||||
open import LogicalFormulae
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Definition
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
@@ -91,3 +91,17 @@ abstract
|
||||
|
||||
negativeInequality : {a : A} → a < 0G → 0G < inverse a
|
||||
negativeInequality {a} a<0 = <WellDefined invRight identLeft (orderRespectsAddition a<0 (inverse a))
|
||||
|
||||
negativeInequality' : {a : A} → 0G < a → inverse a < 0G
|
||||
negativeInequality' {a} 0<a = <WellDefined identLeft invRight (orderRespectsAddition 0<a (inverse a))
|
||||
|
||||
open import Rings.InitialRing R
|
||||
|
||||
fromNIncreasing : (0R < 1R) → (n : ℕ) → (fromN n) < (fromN (succ n))
|
||||
fromNIncreasing 0<1 zero = <WellDefined reflexive (symmetric identRight) 0<1
|
||||
fromNIncreasing 0<1 (succ n) = <WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (fromNIncreasing 0<1 n) 1R)
|
||||
|
||||
fromNPreservesOrder : (0R < 1R) → {a b : ℕ} → (a <N b) → (fromN a) < (fromN b)
|
||||
fromNPreservesOrder 0<1 {zero} {succ zero} a<b = fromNIncreasing 0<1 0
|
||||
fromNPreservesOrder 0<1 {zero} {succ (succ b)} a<b = <Transitive (fromNPreservesOrder 0<1 (succIsPositive b)) (fromNIncreasing 0<1 (succ b))
|
||||
fromNPreservesOrder 0<1 {succ a} {succ b} a<b = <WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (fromNPreservesOrder 0<1 (canRemoveSuccFrom<N a<b)) 1R)
|
||||
|
217
Rings/Orders/Total/AbsoluteValue.agda
Normal file
217
Rings/Orders/Total/AbsoluteValue.agda
Normal file
@@ -0,0 +1,217 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Definition
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
open import Orders.Total.Definition
|
||||
open import Rings.IntegralDomains.Definition
|
||||
|
||||
module Rings.Orders.Total.AbsoluteValue {n m p : _} {A : Set n} {S : Setoid {n} {m} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} {_<_ : Rel {_} {p} A} {pOrder : SetoidPartialOrder S _<_} {pOrderRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pOrderRing) where
|
||||
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open Setoid S
|
||||
open SetoidPartialOrder pOrder
|
||||
open TotallyOrderedRing order
|
||||
open SetoidTotalOrder total
|
||||
open PartiallyOrderedRing pOrderRing
|
||||
open import Rings.Lemmas R
|
||||
open import Rings.Orders.Partial.Lemmas pOrderRing
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
|
||||
abs : A → A
|
||||
abs a with totality 0R a
|
||||
abs a | inl (inl 0<a) = a
|
||||
abs a | inl (inr a<0) = inverse a
|
||||
abs a | inr 0=a = a
|
||||
|
||||
absWellDefined : (a b : A) → a ∼ b → abs a ∼ abs b
|
||||
absWellDefined a b a=b with totality 0R a
|
||||
absWellDefined a b a=b | inl (inl 0<a) with totality 0R b
|
||||
absWellDefined a b a=b | inl (inl 0<a) | inl (inl 0<b) = a=b
|
||||
absWellDefined a b a=b | inl (inl 0<a) | inl (inr b<0) = exFalso (irreflexive {0G} (<Transitive 0<a (<WellDefined (Equivalence.symmetric eq a=b) (Equivalence.reflexive eq) b<0)))
|
||||
absWellDefined a b a=b | inl (inl 0<a) | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq a=b (Equivalence.symmetric eq 0=b)) 0<a))
|
||||
absWellDefined a b a=b | inl (inr a<0) with totality 0R b
|
||||
absWellDefined a b a=b | inl (inr a<0) | inl (inl 0<b) = exFalso (irreflexive {0G} (<Transitive 0<b (<WellDefined a=b (Equivalence.reflexive eq) a<0)))
|
||||
absWellDefined a b a=b | inl (inr a<0) | inl (inr b<0) = inverseWellDefined additiveGroup a=b
|
||||
absWellDefined a b a=b | inl (inr a<0) | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq a=b (Equivalence.symmetric eq 0=b)) (Equivalence.reflexive eq) a<0))
|
||||
absWellDefined a b a=b | inr 0=a with totality 0R b
|
||||
absWellDefined a b a=b | inr 0=a | inl (inl 0<b) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (Equivalence.transitive eq 0=a a=b)) 0<b))
|
||||
absWellDefined a b a=b | inr 0=a | inl (inr b<0) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq 0=a a=b)) (Equivalence.reflexive eq) b<0))
|
||||
absWellDefined a b a=b | inr 0=a | inr 0=b = a=b
|
||||
|
||||
triangleInequality : (a b : A) → ((abs (a + b)) < ((abs a) + (abs b))) || (abs (a + b) ∼ (abs a) + (abs b))
|
||||
triangleInequality a b with totality 0R (a + b)
|
||||
triangleInequality a b | inl (inl 0<a+b) with totality 0R a
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inl 0<a) with totality 0R b
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inl 0<a) | inl (inl 0<b) = inr (Equivalence.reflexive eq)
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inl 0<a) | inl (inr b<0) = inl (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder b<0 (lemm2 b b<0)) a))
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inl 0<a) | inr 0=b = inr (Equivalence.reflexive eq)
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inr a<0) with totality 0R b
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inr a<0) | inl (inl 0<b) = inl (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder a<0 (lemm2 a a<0)) b)
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inr a<0) | inl (inr b<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<a+b (<WellDefined (Equivalence.reflexive eq) identLeft (ringAddInequalities a<0 b<0))))
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inr a<0) | inr 0=b = inl (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder a<0 (lemm2 a a<0)) b)
|
||||
triangleInequality a b | inl (inl 0<a+b) | inr 0=a with totality 0R b
|
||||
triangleInequality a b | inl (inl 0<a+b) | inr 0=a | inl (inl 0<b) = inr (Equivalence.reflexive eq)
|
||||
triangleInequality a b | inl (inl 0<a+b) | inr 0=a | inl (inr b<0) = inl (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder b<0 (lemm2 b b<0)) a))
|
||||
triangleInequality a b | inl (inl 0<a+b) | inr 0=a | inr 0=b = inr (Equivalence.reflexive eq)
|
||||
triangleInequality a b | inl (inr a+b<0) with totality 0G a
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inl 0<a) with totality 0G b
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inl 0<a) | inl (inl 0<b) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (ringAddInequalities 0<a 0<b)) a+b<0))
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inl 0<a) | inl (inr b<0) = inl (<WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq (invContravariant additiveGroup)) (inverseWellDefined additiveGroup groupIsAbelian)) (Equivalence.reflexive eq) (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder (lemm2' _ 0<a) 0<a) (inverse b)))
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inl 0<a) | inr 0=b = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<a (<WellDefined (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) identRight) (Equivalence.reflexive eq) a+b<0)))
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) with totality 0G b
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) | inl (inl 0<b) = inl (<WellDefined (Equivalence.symmetric eq (invContravariant additiveGroup)) groupIsAbelian (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder (lemm2' _ 0<b) 0<b) (inverse a)))
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) | inl (inr b<0) = inr (Equivalence.transitive eq (invContravariant additiveGroup) groupIsAbelian)
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) | inr 0=b = inr (Equivalence.transitive eq (invContravariant additiveGroup) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (+WellDefined (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=b)) (invIdent additiveGroup)) (Equivalence.reflexive eq)) identLeft) (Equivalence.symmetric eq identRight)) (+WellDefined (Equivalence.reflexive eq) 0=b)))
|
||||
triangleInequality a b | inl (inr a+b<0) | inr 0=a with totality 0G b
|
||||
triangleInequality a b | inl (inr a+b<0) | inr 0=a | inl (inl 0<b) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<b (<WellDefined (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) identLeft) (Equivalence.reflexive eq) a+b<0)))
|
||||
triangleInequality a b | inl (inr a+b<0) | inr 0=a | inl (inr b<0) = inr (Equivalence.transitive eq (invContravariant additiveGroup) (Equivalence.transitive eq groupIsAbelian (+WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq (inverseWellDefined additiveGroup 0=a)) (invIdent additiveGroup)) 0=a) (Equivalence.reflexive eq))))
|
||||
triangleInequality a b | inl (inr a+b<0) | inr 0=a | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.symmetric eq 0=b)) identLeft) (Equivalence.reflexive eq) a+b<0))
|
||||
triangleInequality a b | inr 0=a+b with totality 0G a
|
||||
triangleInequality a b | inr 0=a+b | inl (inl 0<a) with totality 0G b
|
||||
triangleInequality a b | inr 0=a+b | inl (inl 0<a) | inl (inl 0<b) = exFalso (irreflexive {0G} (<WellDefined identLeft (Equivalence.symmetric eq 0=a+b) (ringAddInequalities 0<a 0<b)))
|
||||
triangleInequality a b | inr 0=a+b | inl (inl 0<a) | inl (inr b<0) = inl (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder b<0 (lemm2 _ b<0)) a))
|
||||
triangleInequality a b | inr 0=a+b | inl (inl 0<a) | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (lemm3 _ _ (Equivalence.transitive eq 0=a+b groupIsAbelian) 0=b)) 0<a))
|
||||
triangleInequality a b | inr 0=a+b | inl (inr a<0) with totality 0G b
|
||||
triangleInequality a b | inr 0=a+b | inl (inr a<0) | inl (inl 0<b) = inl (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder a<0 (lemm2 _ a<0)) b)
|
||||
triangleInequality a b | inr 0=a+b | inl (inr a<0) | inl (inr b<0) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq 0=a+b) identLeft (ringAddInequalities a<0 b<0)))
|
||||
triangleInequality a b | inr 0=a+b | inl (inr a<0) | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (lemm3 _ _ (Equivalence.transitive eq 0=a+b groupIsAbelian) 0=b)) (Equivalence.reflexive eq) a<0))
|
||||
triangleInequality a b | inr 0=a+b | inr 0=a with totality 0G b
|
||||
triangleInequality a b | inr 0=a+b | inr 0=a | inl (inl 0<b) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (lemm3 a b 0=a+b 0=a)) 0<b))
|
||||
triangleInequality a b | inr 0=a+b | inr 0=a | inl (inr b<0) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (lemm3 a b 0=a+b 0=a)) (Equivalence.reflexive eq) b<0))
|
||||
triangleInequality a b | inr 0=a+b | inr 0=a | inr 0=b = inr (Equivalence.reflexive eq)
|
||||
|
||||
absZero : abs (Ring.0R R) ≡ Ring.0R R
|
||||
absZero with totality (Ring.0R R) (Ring.0R R)
|
||||
absZero | inl (inl x) = exFalso (irreflexive x)
|
||||
absZero | inl (inr x) = exFalso (irreflexive x)
|
||||
absZero | inr x = refl
|
||||
|
||||
absNegation : (a : A) → (abs a) ∼ (abs (inverse a))
|
||||
absNegation a with totality 0R a
|
||||
absNegation a | inl (inl 0<a) with totality 0G (inverse a)
|
||||
absNegation a | inl (inl 0<a) | inl (inl 0<-a) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<-a (lemm2' a 0<a)))
|
||||
absNegation a | inl (inl 0<a) | inl (inr -a<0) = Equivalence.symmetric eq (invTwice additiveGroup a)
|
||||
absNegation a | inl (inl 0<a) | inr 0=-a = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq (invTwice additiveGroup a)) (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=-a))) (invIdent additiveGroup)) 0<a))
|
||||
absNegation a | inl (inr a<0) with totality 0G (inverse a)
|
||||
absNegation a | inl (inr a<0) | inl (inl 0<-a) = Equivalence.reflexive eq
|
||||
absNegation a | inl (inr a<0) | inl (inr -a<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder (<WellDefined (Equivalence.reflexive eq) (invTwice additiveGroup a) (lemm2 (inverse a) -a<0)) a<0))
|
||||
absNegation a | inl (inr a<0) | inr 0=-a = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq (Equivalence.transitive eq (inverseWellDefined additiveGroup 0=-a) (invTwice additiveGroup a))) (invIdent additiveGroup)) (Equivalence.reflexive eq) a<0))
|
||||
absNegation a | inr 0=a with totality 0G (inverse a)
|
||||
absNegation a | inr 0=a | inl (inl 0<-a) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=a)) (invIdent additiveGroup)) 0<-a))
|
||||
absNegation a | inr 0=a | inl (inr -a<0) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=a)) (invIdent additiveGroup)) (Equivalence.reflexive eq) -a<0))
|
||||
absNegation a | inr 0=a | inr 0=-a = Equivalence.transitive eq (Equivalence.symmetric eq 0=a) 0=-a
|
||||
|
||||
absRespectsTimes : (a b : A) → abs (a * b) ∼ (abs a) * (abs b)
|
||||
absRespectsTimes a b with totality 0R a
|
||||
absRespectsTimes a b | inl (inl 0<a) with totality 0R b
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) | inl (inl 0<ab) = Equivalence.reflexive eq
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) | inl (inr ab<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder (orderRespectsMultiplication 0<a 0<b) ab<0))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=ab) (orderRespectsMultiplication 0<a 0<b)))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inl (inl 0<ab) with <WellDefined (Equivalence.reflexive eq) ringMinusExtracts (orderRespectsMultiplication 0<a (lemm2 b b<0))
|
||||
... | bl = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<ab (<WellDefined (invTwice additiveGroup _) (Equivalence.reflexive eq) (lemm2' _ bl))))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq 0=ab) (Equivalence.reflexive eq) (posTimesNeg a b 0<a b<0)))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) 0<ab))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inl (inr ab<0) = exFalso ((irreflexive {0G} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) (Equivalence.reflexive eq) ab<0)))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inr 0=ab = Equivalence.reflexive eq
|
||||
absRespectsTimes a b | inl (inr a<0) with totality 0R b
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inl (inl 0<ab) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<ab (<WellDefined *Commutative (Equivalence.reflexive eq) (posTimesNeg b a 0<b a<0))))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts'
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq 0=ab *Commutative)) (Equivalence.reflexive eq) (posTimesNeg b a 0<b a<0)))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inl (inl 0<ab) = Equivalence.symmetric eq twoNegativesTimes
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inl (inr ab<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder (negTimesPos a b a<0 b<0) ab<0))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inr 0=ab = exFalso (exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=ab) (negTimesPos a b a<0 b<0))))
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) 0<ab))
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inl (inr ab<0) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) (Equivalence.reflexive eq) ab<0))
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inr 0=ab = Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (Equivalence.transitive eq (Equivalence.transitive eq timesZero (Equivalence.symmetric eq timesZero)) (*WellDefined (Equivalence.reflexive eq) 0=b))
|
||||
absRespectsTimes a b | inr 0=a with totality 0R b
|
||||
absRespectsTimes a b | inr 0=a | inl (inl 0<b) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inr 0=a | inl (inl 0<b) | inl (inl 0<ab) = Equivalence.reflexive eq
|
||||
absRespectsTimes a b | inr 0=a | inl (inl 0<b) | inl (inr ab<0) = Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.transitive eq (*WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) (Equivalence.transitive eq *Commutative timesZero))) (invIdent additiveGroup)) (Equivalence.transitive eq (Equivalence.symmetric eq timesZero) *Commutative)) (*WellDefined 0=a (Equivalence.reflexive eq))
|
||||
absRespectsTimes a b | inr 0=a | inl (inl 0<b) | inr 0=ab = Equivalence.reflexive eq
|
||||
absRespectsTimes a b | inr 0=a | inl (inr b<0) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inr 0=a | inl (inr b<0) | inl (inl 0<ab) = Equivalence.transitive eq (Equivalence.transitive eq (*WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) *Commutative) (Equivalence.transitive eq timesZero (Equivalence.transitive eq (Equivalence.symmetric eq (Equivalence.transitive eq *Commutative timesZero)) (*WellDefined 0=a (Equivalence.reflexive eq))))
|
||||
absRespectsTimes a b | inr 0=a | inl (inr b<0) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts
|
||||
absRespectsTimes a b | inr 0=a | inl (inr b<0) | inr 0=ab = Equivalence.transitive eq (Equivalence.transitive eq (*WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) *Commutative) (Equivalence.transitive eq timesZero (Equivalence.transitive eq (Equivalence.symmetric eq (Equivalence.transitive eq *Commutative timesZero)) (*WellDefined 0=a (Equivalence.reflexive eq))))
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b with totality 0R (a * b)
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) 0<ab))
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b | inl (inr ab<0) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) (Equivalence.reflexive eq) ab<0))
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b | inr 0=ab = Equivalence.reflexive eq
|
||||
|
||||
absNonnegative : {a : A} → (abs a < 0R) → False
|
||||
absNonnegative {a} pr with totality 0R a
|
||||
absNonnegative {a} pr | inl (inl x) = irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder x pr)
|
||||
absNonnegative {a} pr | inl (inr x) = irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder (<WellDefined (Equivalence.reflexive eq) (invTwice additiveGroup a) (lemm2 (inverse a) pr)) x)
|
||||
absNonnegative {a} pr | inr x = irreflexive {0G} (<WellDefined (Equivalence.symmetric eq x) (Equivalence.reflexive eq) pr)
|
||||
|
||||
absZeroImpliesZero : {a : A} → abs a ∼ 0R → a ∼ 0R
|
||||
absZeroImpliesZero {a} a=0 with totality 0R a
|
||||
absZeroImpliesZero {a} a=0 | inl (inl 0<a) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) a=0 0<a))
|
||||
absZeroImpliesZero {a} a=0 | inl (inr a<0) = Equivalence.symmetric eq (lemm3 (inverse a) a (Equivalence.symmetric eq invLeft) (Equivalence.symmetric eq a=0))
|
||||
absZeroImpliesZero {a} a=0 | inr 0=a = a=0
|
||||
|
||||
addingAbsCannotShrink : {a b : A} → 0G < b → 0G < ((abs a) + b)
|
||||
addingAbsCannotShrink {a} {b} 0<b with totality 0G a
|
||||
addingAbsCannotShrink {a} {b} 0<b | inl (inl x) = <WellDefined identLeft (Equivalence.reflexive eq) (ringAddInequalities x 0<b)
|
||||
addingAbsCannotShrink {a} {b} 0<b | inl (inr x) = <WellDefined identLeft (Equivalence.reflexive eq) (ringAddInequalities (lemm2 a x) 0<b)
|
||||
addingAbsCannotShrink {a} {b} 0<b | inr x = <WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (Equivalence.symmetric eq identLeft) (+WellDefined x (Equivalence.reflexive eq))) 0<b
|
||||
|
||||
greaterZeroImpliesEqualAbs : {a : A} → 0R < a → a ∼ abs a
|
||||
greaterZeroImpliesEqualAbs {a} 0<a with totality 0R a
|
||||
... | inl (inl _) = Equivalence.reflexive eq
|
||||
... | inl (inr a<0) = exFalso (irreflexive (SetoidPartialOrder.<Transitive pOrder a<0 0<a))
|
||||
... | inr 0=a = exFalso (irreflexive (<WellDefined 0=a (Equivalence.reflexive eq) 0<a))
|
||||
|
||||
lessZeroImpliesEqualNegAbs : {a : A} → a < 0R → abs a ∼ inverse a
|
||||
lessZeroImpliesEqualNegAbs {a} a<0 with totality 0R a
|
||||
... | inl (inr _) = Equivalence.reflexive eq
|
||||
... | inl (inl 0<a) = exFalso (irreflexive (SetoidPartialOrder.<Transitive pOrder a<0 0<a))
|
||||
... | inr 0=a = exFalso (irreflexive (<WellDefined (Equivalence.reflexive eq) 0=a a<0))
|
||||
|
||||
absZeroIsZero : abs 0R ∼ 0R
|
||||
absZeroIsZero with totality 0R 0R
|
||||
... | inl (inl bad) = exFalso (irreflexive bad)
|
||||
... | inl (inr bad) = exFalso (irreflexive bad)
|
||||
... | inr _ = Equivalence.reflexive eq
|
||||
|
||||
greaterThanAbsImpliesGreaterThan0 : {a b : A} → (abs a) < b → 0R < b
|
||||
greaterThanAbsImpliesGreaterThan0 {a} {b} a<b with totality 0R a
|
||||
greaterThanAbsImpliesGreaterThan0 {a} {b} a<b | inl (inl 0<a) = SetoidPartialOrder.<Transitive pOrder 0<a a<b
|
||||
greaterThanAbsImpliesGreaterThan0 {a} {b} a<b | inl (inr a<0) = SetoidPartialOrder.<Transitive pOrder (lemm2 _ a<0) a<b
|
||||
greaterThanAbsImpliesGreaterThan0 {a} {b} a<b | inr 0=a = <WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq) a<b
|
||||
|
||||
abs1Is1 : abs 1R ∼ 1R
|
||||
abs1Is1 with totality 0R 1R
|
||||
abs1Is1 | inl (inl 0<1) = Equivalence.reflexive eq
|
||||
abs1Is1 | inl (inr 1<0) = exFalso (1<0False 1<0)
|
||||
abs1Is1 | inr 0=1 = Equivalence.reflexive eq
|
||||
|
||||
absBoundedImpliesBounded : {a b : A} → abs a < b → a < b
|
||||
absBoundedImpliesBounded {a} {b} a<b with totality 0G a
|
||||
absBoundedImpliesBounded {a} {b} a<b | inl (inl x) = a<b
|
||||
absBoundedImpliesBounded {a} {b} a<b | inl (inr x) = SetoidPartialOrder.<Transitive pOrder x (SetoidPartialOrder.<Transitive pOrder (lemm2 a x) a<b)
|
||||
absBoundedImpliesBounded {a} {b} a<b | inr x = a<b
|
||||
|
||||
a-bPos : {a b : A} → ((a ∼ b) → False) → 0R < abs (a + inverse b)
|
||||
a-bPos {a} {b} a!=b with totality 0R (a + inverse b)
|
||||
a-bPos {a} {b} a!=b | inl (inl x) = x
|
||||
a-bPos {a} {b} a!=b | inl (inr x) = lemm2 _ x
|
||||
a-bPos {a} {b} a!=b | inr x = exFalso (a!=b (transferToRight additiveGroup (Equivalence.symmetric eq x)))
|
@@ -4,7 +4,8 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Definition
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
@@ -48,21 +49,21 @@ FloorIs.prAbove (addOneToFloor record { prBelow = x ; prAbove = prAbove }) = <We
|
||||
|
||||
private
|
||||
0<n : {x y : A} → (x < y) → 0R < fromN n
|
||||
0<n x<y = fromNPreservesOrder (anyComparisonImpliesNontrivial x<y) (lessTransitive (succIsPositive 0) 1<n)
|
||||
0<n x<y = fromNPreservesOrder (0<1 (anyComparisonImpliesNontrivial x<y)) (lessTransitive (succIsPositive 0) 1<n)
|
||||
|
||||
0<aFromN : {a : A} → (0<a : 0R < a) → 0R < (a * fromN n)
|
||||
0<aFromN 0<a = orderRespectsMultiplication 0<a (0<n 0<a)
|
||||
|
||||
floorWellDefinedLemma : {a : A} {n1 n2 : ℕ} → FloorIs a n1 → FloorIs a n2 → n1 <N n2 → False
|
||||
floorWellDefinedLemma {a} {n1} {n2} record { prAbove = prAbove1 } record { prBelow = inl prBelow } n1<n2 with TotalOrder.totality ℕTotalOrder (succ n1) n2
|
||||
... | inl (inl n1+1<n2) = irreflexive (<Transitive prAbove1 (<Transitive (fromNPreservesOrder (anyComparisonImpliesNontrivial prBelow) n1+1<n2) prBelow))
|
||||
... | inl (inl n1+1<n2) = irreflexive (<Transitive prAbove1 (<Transitive (fromNPreservesOrder (0<1 (anyComparisonImpliesNontrivial prBelow)) n1+1<n2) prBelow))
|
||||
... | inl (inr n2<n1+1) = noIntegersBetweenXAndSuccX n1 n1<n2 n2<n1+1
|
||||
... | inr refl = irreflexive (<Transitive prAbove1 prBelow)
|
||||
floorWellDefinedLemma {a} {n1} {n2} record { prBelow = (inl x) ; prAbove = prAbove1 } record { prBelow = (inr eq) ; prAbove = _ } n1<n2 with TotalOrder.totality ℕTotalOrder (succ n1) n2
|
||||
... | inl (inl n1+1<n2) = irreflexive (<Transitive prAbove1 (<WellDefined reflexive eq (fromNPreservesOrder (anyComparisonImpliesNontrivial prAbove1) n1+1<n2)))
|
||||
... | inl (inl n1+1<n2) = irreflexive (<Transitive prAbove1 (<WellDefined reflexive eq (fromNPreservesOrder (0<1 (anyComparisonImpliesNontrivial prAbove1)) n1+1<n2)))
|
||||
... | inl (inr n2<n1+1) = noIntegersBetweenXAndSuccX n1 n1<n2 n2<n1+1
|
||||
... | inr refl = irreflexive (<WellDefined reflexive eq prAbove1)
|
||||
floorWellDefinedLemma {a} {n1} {n2} record { prBelow = (inr x) ; prAbove = prAbove1 } record { prBelow = (inr eq) ; prAbove = _ } n1<n2 = lessIrreflexive {n1} (fromNPreservesOrder' (anyComparisonImpliesNontrivial prAbove1) (<WellDefined reflexive (transitive eq (symmetric x)) (fromNPreservesOrder (anyComparisonImpliesNontrivial prAbove1) n1<n2)))
|
||||
floorWellDefinedLemma {a} {n1} {n2} record { prBelow = (inr x) ; prAbove = prAbove1 } record { prBelow = (inr eq) ; prAbove = _ } n1<n2 = lessIrreflexive {n1} (fromNPreservesOrder' (anyComparisonImpliesNontrivial prAbove1) (<WellDefined reflexive (transitive eq (symmetric x)) (fromNPreservesOrder (0<1 (anyComparisonImpliesNontrivial prAbove1)) n1<n2)))
|
||||
|
||||
floorWellDefined : {a : A} {n1 n2 : ℕ} → FloorIs a n1 → FloorIs a n2 → n1 ≡ n2
|
||||
floorWellDefined {a} {n1} {n2} record { prBelow = prBelow1 ; prAbove = prAbove1 } record { prBelow = prBelow ; prAbove = prAbove } with TotalOrder.totality ℕTotalOrder n1 n2
|
||||
@@ -80,7 +81,7 @@ private
|
||||
computeFloor' {zero} (succ k) pr a 0<a a<f = exFalso (irreflexive (<Transitive 0<a a<f))
|
||||
computeFloor' {succ k} n pr a 0<a a<f with totality 1R a
|
||||
... | inl (inr a<1) = 0 , (record { prAbove = <WellDefined reflexive (symmetric identRight) a<1 ; prBelow = inl 0<a })
|
||||
... | inr 1=a = 1 , (record { prAbove = <WellDefined (transitive identRight 1=a) reflexive (fromNPreservesOrder (anyComparisonImpliesNontrivial 0<a) {1} (le 0 refl)) ; prBelow = inr (transitive identRight 1=a) })
|
||||
... | inr 1=a = 1 , (record { prAbove = <WellDefined (transitive identRight 1=a) reflexive (fromNPreservesOrder (0<1 (anyComparisonImpliesNontrivial 0<a)) {1} (le 0 refl)) ; prBelow = inr (transitive identRight 1=a) })
|
||||
... | inl (inl 1<a) with computeFloor' {k} (succ n) (transitivity (transitivity (Semiring.commutative ℕSemiring k (succ n)) (applyEquality succ (Semiring.commutative ℕSemiring n k))) pr) (a + inverse 1R) (moveInequality 1<a) (<WellDefined reflexive (transitive groupIsAbelian (transitive +Associative (transitive (+WellDefined invLeft reflexive) identLeft))) (orderRespectsAddition a<f (inverse 1R)))
|
||||
... | N , snd = succ N , floorWellDefined' (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invLeft) identRight)) (addOneToFloor snd)
|
||||
|
||||
@@ -179,7 +180,7 @@ approximationComesFromBelow 1/n 1/nPr a 0<a a<1 (succ m) with computeFloor' 0 (S
|
||||
... | x , record { prBelow = inr x=an ; prAbove = an<x+1 } = inr (transitive (partialSums'Translate 0G _ _ m) (transitive (+WellDefined (transitive (*WellDefined x=an reflexive) (transitive (symmetric *Associative) (*WellDefined reflexive (transitive *Commutative 1/nPr)))) reflexive) (transitive (+WellDefined (transitive *Commutative identIsIdent) (transitive (partialSums'WellDefined' 0G _ (constSequence 0G) (λ m → transitive (multiplyZeroSequence (powerSeq 1/n (1/n * 1/n)) m) (identityOfIndiscernablesRight _∼_ reflexive (equalityCommutative (indexAndConst 0G m)))) m) (transitive (partialSumsConst _ _ m) (transitive identLeft timesZero')))) identRight)))
|
||||
... | x , record { prBelow = inl x<an ; prAbove = an<x+1 } with approximationComesFromBelow 1/n 1/nPr ((a * fromN n) + inverse (fromN x)) (moveInequality x<an) (<WellDefined reflexive (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight)) (orderRespectsAddition an<x+1 (inverse (fromN x)))) m
|
||||
... | inl inter<an = inl (<WellDefined (symmetric (partialSums'Translate 0G (fromN x * 1/n) _ m)) reflexive (<WellDefined (+WellDefined reflexive (symmetric (partialSums'WellDefined 0G (fromN n * 0G) (symmetric timesZero) _ m))) reflexive (ringCanCancelPositive (0<n an<x+1) (<WellDefined (symmetric (transitive *DistributesOver+' (+WellDefined (transitive (symmetric *Associative) (transitive (*WellDefined reflexive 1/nPr) (transitive *Commutative identIsIdent))) (transitive *Commutative (transitive (lemma2 (fromN n) (fromN n * 0G) 1/n (1/n * 1/n) _ m) (transitive (partialSums'WellDefined _ 0G (transitive (*WellDefined reflexive timesZero) timesZero) _ m) (partialSums'WellDefined' 0G _ _ (λ m → applyWellDefined _ _ _ dot (λ {x} {y} {s} → dotWellDefined {x} {y} {s}) (λ m → powerSeqWellDefined {_} {1/n} {_} {1/n} reflexive (transitive (symmetric *Associative) (transitive (transitive *Commutative (*WellDefined 1/nPr reflexive)) identIsIdent)) m) m) m))))))) reflexive (<WellDefined groupIsAbelian (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invLeft) identRight)) (orderRespectsAddition inter<an (fromN x)))))))
|
||||
... | inr inter=an = inr (transitive (partialSums'Translate 0G (fromN x * 1/n) _ m) (cancelIntDom' (isIntDom λ t → anyComparisonImpliesNontrivial a<1 (symmetric t)) ans λ t → irreflexive (<WellDefined reflexive t (fromNPreservesOrder (anyComparisonImpliesNontrivial a<1) (lessTransitive (succIsPositive 0) 1<n)))))
|
||||
... | inr inter=an = inr (transitive (partialSums'Translate 0G (fromN x * 1/n) _ m) (cancelIntDom' (isIntDom λ t → anyComparisonImpliesNontrivial a<1 (symmetric t)) ans λ t → irreflexive (<WellDefined reflexive t (fromNPreservesOrder (0<1 (anyComparisonImpliesNontrivial a<1)) (lessTransitive (succIsPositive 0) 1<n)))))
|
||||
where
|
||||
ans : (((fromN x * 1/n) + index (partialSums' 0G (apply dot (powerSeq 1/n (1/n * 1/n)) (baseNExpansion ((a * fromN n) + inverse (fromN x)) (moveInequality x<an) (<WellDefined reflexive (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight)) (orderRespectsAddition an<x+1 (inverse (fromN x))))))) m) * fromN n) ∼ a * fromN n
|
||||
ans = transitive (transitive *DistributesOver+' (+WellDefined (transitive (symmetric *Associative) (transitive (*WellDefined reflexive 1/nPr) (transitive *Commutative identIsIdent))) *Commutative)) (transitive (+WellDefined reflexive (transitive (lemma2 (fromN n) 0G 1/n (1/n * 1/n) _ m) (transitive (partialSums'WellDefined _ _ timesZero _ m) (partialSums'WellDefined' _ _ _ (λ m → applyWellDefined _ _ _ dot (λ {x} {y} {s} → dotWellDefined {x} {y} {s}) (λ m → powerSeqWellDefined reflexive (transitive (symmetric *Associative) (transitive (*WellDefined reflexive 1/nPr) (transitive *Commutative identIsIdent))) m) m) m)))) (transitive (+WellDefined reflexive inter=an) (transitive groupIsAbelian (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invLeft) identRight)))))
|
||||
@@ -194,7 +195,7 @@ powPos (succ m) 0<a = orderRespectsMultiplication 0<a (powPos m 0<a)
|
||||
|
||||
approximationIsNear : (1/n : A) → (1/nPr : 1/n * (fromN n) ∼ 1R) → (a : A) → (0<a : 0R < a) → (a<1 : a < 1R) → (m : ℕ) → a < ((index (approximations 1/n 1/nPr (baseNExpansion a 0<a a<1)) m) + pow m 1/n)
|
||||
approximationIsNear 1/n 1/npr a 0<a a<1 zero with computeFloor' 0 (Semiring.sumZeroRight ℕSemiring n) (a * fromN n) (0<aFromN 0<a) (<WellDefined reflexive identIsIdent (ringCanMultiplyByPositive (0<n 0<a) a<1))
|
||||
... | x , record { prBelow = inl prBelow ; prAbove = prAbove } = <WellDefined reflexive (transitive (symmetric identLeft) +Associative) (ringCanCancelPositive (0<n prAbove) (<Transitive prAbove (<WellDefined reflexive (transitive (+WellDefined (transitive (*WellDefined reflexive (symmetric 1/npr)) *Associative) (symmetric identIsIdent)) (symmetric *DistributesOver+')) (<WellDefined reflexive (transitive groupIsAbelian (+WellDefined (transitive (symmetric identIsIdent) *Commutative) reflexive)) (orderRespectsAddition (<WellDefined identRight reflexive (fromNPreservesOrder (anyComparisonImpliesNontrivial prAbove) 1<n)) (fromN x))))))
|
||||
... | x , record { prBelow = inl prBelow ; prAbove = prAbove } = <WellDefined reflexive (transitive (symmetric identLeft) +Associative) (ringCanCancelPositive (0<n prAbove) (<Transitive prAbove (<WellDefined reflexive (transitive (+WellDefined (transitive (*WellDefined reflexive (symmetric 1/npr)) *Associative) (symmetric identIsIdent)) (symmetric *DistributesOver+')) (<WellDefined reflexive (transitive groupIsAbelian (+WellDefined (transitive (symmetric identIsIdent) *Commutative) reflexive)) (orderRespectsAddition (<WellDefined identRight reflexive (fromNPreservesOrder (0<1 (anyComparisonImpliesNontrivial prAbove)) 1<n)) (fromN x))))))
|
||||
... | x , record { prBelow = inr prBelow ; prAbove = prAbove } = <WellDefined reflexive (symmetric (transitive (symmetric +Associative) identLeft)) (ringCanCancelPositive (0<n prAbove) (<WellDefined prBelow (symmetric (transitive *DistributesOver+' (+WellDefined (transitive (symmetric *Associative) (transitive *Commutative (transitive (*WellDefined 1/npr reflexive) identIsIdent))) identIsIdent))) (<WellDefined identLeft groupIsAbelian (orderRespectsAddition (0<n prAbove) (fromN x)))))
|
||||
approximationIsNear 1/n 1/npr a 0<a a<1 (succ m) with computeFloor' 0 (Semiring.sumZeroRight ℕSemiring n) (a * fromN n) (0<aFromN 0<a) (<WellDefined reflexive identIsIdent (ringCanMultiplyByPositive (0<n 0<a) a<1))
|
||||
... | x , record { prBelow = inl prBelow ; prAbove = prAbove } = <WellDefined reflexive (+WellDefined (symmetric (partialSums'Translate 0G (fromN x * 1/n) _ m)) reflexive) (ringCanCancelPositive (0<n prAbove) (<WellDefined reflexive (transitive (+WellDefined (transitive (+WellDefined (transitive (transitive (symmetric identIsIdent) (transitive *Commutative (*WellDefined reflexive (symmetric 1/npr))) ) *Associative) *Commutative) (symmetric *DistributesOver+')) (transitive (transitive (transitive (symmetric identIsIdent) (*WellDefined (transitive (symmetric 1/npr) *Commutative) reflexive)) (symmetric *Associative)) *Commutative)) (symmetric *DistributesOver+')) (<WellDefined reflexive (+WellDefined (+WellDefined reflexive (transitive (partialSums'WellDefined' _ _ _ (λ m → applyWellDefined _ _ _ dot (λ {_} {_} {s} → dotWellDefined {_} {_} {s}) (λ m → powerSeqWellDefined reflexive (transitive (transitive (transitive (symmetric identIsIdent) *Commutative) (*WellDefined reflexive (symmetric 1/npr))) *Associative) m) m) m) (symmetric (lemma2 (fromN n) 0G 1/n (1/n * 1/n) _ m)))) reflexive) (<WellDefined reflexive (+WellDefined (+WellDefined reflexive (partialSums'WellDefined _ _ (symmetric timesZero) _ m)) reflexive) (<WellDefined reflexive +Associative u)))))
|
||||
|
@@ -8,7 +8,8 @@ open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
|
@@ -5,10 +5,9 @@ open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
@@ -16,15 +15,8 @@ open import Numbers.Naturals.Order
|
||||
|
||||
module Rings.Orders.Total.Cauchy {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) where
|
||||
|
||||
open Setoid S
|
||||
open SetoidTotalOrder (TotallyOrderedRing.total order)
|
||||
open SetoidPartialOrder pOrder
|
||||
open Equivalence eq
|
||||
open TotallyOrderedRing order
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
open import Rings.Orders.Total.AbsoluteValue order
|
||||
|
||||
cauchy : Sequence A → Set (m ⊔ o)
|
||||
cauchy s = ∀ (ε : A) → (0R < ε) → Sg ℕ (λ N → ∀ {m n : ℕ} → (N <N m) → (N <N n) → abs ((index s m) -R (index s n)) < ε)
|
||||
cauchy s = ∀ (ε : A) → (Ring.0R R < ε) → Sg ℕ (λ N → ∀ {m n : ℕ} → (N <N m) → (N <N n) → abs (Ring._-R_ R (index s m) (index s n)) < ε)
|
||||
|
@@ -1,7 +1,8 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Definition
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Rings.Definition
|
||||
|
@@ -3,7 +3,8 @@
|
||||
open import LogicalFormulae
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Definition
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
@@ -27,27 +28,7 @@ open PartiallyOrderedRing pOrderRing
|
||||
open import Rings.Lemmas R
|
||||
open import Rings.Orders.Partial.Lemmas pOrderRing
|
||||
|
||||
abs : A → A
|
||||
abs a with totality 0R a
|
||||
abs a | inl (inl 0<a) = a
|
||||
abs a | inl (inr a<0) = inverse a
|
||||
abs a | inr 0=a = a
|
||||
|
||||
abstract
|
||||
absWellDefined : (a b : A) → a ∼ b → abs a ∼ abs b
|
||||
absWellDefined a b a=b with totality 0R a
|
||||
absWellDefined a b a=b | inl (inl 0<a) with totality 0R b
|
||||
absWellDefined a b a=b | inl (inl 0<a) | inl (inl 0<b) = a=b
|
||||
absWellDefined a b a=b | inl (inl 0<a) | inl (inr b<0) = exFalso (irreflexive {0G} (<Transitive 0<a (<WellDefined (Equivalence.symmetric eq a=b) (Equivalence.reflexive eq) b<0)))
|
||||
absWellDefined a b a=b | inl (inl 0<a) | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq a=b (Equivalence.symmetric eq 0=b)) 0<a))
|
||||
absWellDefined a b a=b | inl (inr a<0) with totality 0R b
|
||||
absWellDefined a b a=b | inl (inr a<0) | inl (inl 0<b) = exFalso (irreflexive {0G} (<Transitive 0<b (<WellDefined a=b (Equivalence.reflexive eq) a<0)))
|
||||
absWellDefined a b a=b | inl (inr a<0) | inl (inr b<0) = inverseWellDefined additiveGroup a=b
|
||||
absWellDefined a b a=b | inl (inr a<0) | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq a=b (Equivalence.symmetric eq 0=b)) (Equivalence.reflexive eq) a<0))
|
||||
absWellDefined a b a=b | inr 0=a with totality 0R b
|
||||
absWellDefined a b a=b | inr 0=a | inl (inl 0<b) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (Equivalence.transitive eq 0=a a=b)) 0<b))
|
||||
absWellDefined a b a=b | inr 0=a | inl (inr b<0) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq 0=a a=b)) (Equivalence.reflexive eq) b<0))
|
||||
absWellDefined a b a=b | inr 0=a | inr 0=b = a=b
|
||||
|
||||
lemm2 : (a : A) → a < 0G → 0G < inverse a
|
||||
lemm2 a a<0 with totality 0R (inverse a)
|
||||
@@ -67,48 +48,6 @@ abstract
|
||||
t : a + 0G ∼ 0G
|
||||
t = Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) 0=-a) (invRight {a})
|
||||
|
||||
triangleInequality : (a b : A) → ((abs (a + b)) < ((abs a) + (abs b))) || (abs (a + b) ∼ (abs a) + (abs b))
|
||||
triangleInequality a b with totality 0R (a + b)
|
||||
triangleInequality a b | inl (inl 0<a+b) with totality 0R a
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inl 0<a) with totality 0R b
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inl 0<a) | inl (inl 0<b) = inr (Equivalence.reflexive eq)
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inl 0<a) | inl (inr b<0) = inl (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder b<0 (lemm2 b b<0)) a))
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inl 0<a) | inr 0=b = inr (Equivalence.reflexive eq)
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inr a<0) with totality 0R b
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inr a<0) | inl (inl 0<b) = inl (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder a<0 (lemm2 a a<0)) b)
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inr a<0) | inl (inr b<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<a+b (<WellDefined (Equivalence.reflexive eq) identLeft (ringAddInequalities a<0 b<0))))
|
||||
triangleInequality a b | inl (inl 0<a+b) | inl (inr a<0) | inr 0=b = inl (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder a<0 (lemm2 a a<0)) b)
|
||||
triangleInequality a b | inl (inl 0<a+b) | inr 0=a with totality 0R b
|
||||
triangleInequality a b | inl (inl 0<a+b) | inr 0=a | inl (inl 0<b) = inr (Equivalence.reflexive eq)
|
||||
triangleInequality a b | inl (inl 0<a+b) | inr 0=a | inl (inr b<0) = inl (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder b<0 (lemm2 b b<0)) a))
|
||||
triangleInequality a b | inl (inl 0<a+b) | inr 0=a | inr 0=b = inr (Equivalence.reflexive eq)
|
||||
triangleInequality a b | inl (inr a+b<0) with totality 0G a
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inl 0<a) with totality 0G b
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inl 0<a) | inl (inl 0<b) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (ringAddInequalities 0<a 0<b)) a+b<0))
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inl 0<a) | inl (inr b<0) = inl (<WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq (invContravariant additiveGroup)) (inverseWellDefined additiveGroup groupIsAbelian)) (Equivalence.reflexive eq) (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder (lemm2' _ 0<a) 0<a) (inverse b)))
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inl 0<a) | inr 0=b = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<a (<WellDefined (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) identRight) (Equivalence.reflexive eq) a+b<0)))
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) with totality 0G b
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) | inl (inl 0<b) = inl (<WellDefined (Equivalence.symmetric eq (invContravariant additiveGroup)) groupIsAbelian (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder (lemm2' _ 0<b) 0<b) (inverse a)))
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) | inl (inr b<0) = inr (Equivalence.transitive eq (invContravariant additiveGroup) groupIsAbelian)
|
||||
triangleInequality a b | inl (inr a+b<0) | inl (inr a<0) | inr 0=b = inr (Equivalence.transitive eq (invContravariant additiveGroup) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (+WellDefined (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=b)) (invIdent additiveGroup)) (Equivalence.reflexive eq)) identLeft) (Equivalence.symmetric eq identRight)) (+WellDefined (Equivalence.reflexive eq) 0=b)))
|
||||
triangleInequality a b | inl (inr a+b<0) | inr 0=a with totality 0G b
|
||||
triangleInequality a b | inl (inr a+b<0) | inr 0=a | inl (inl 0<b) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<b (<WellDefined (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) identLeft) (Equivalence.reflexive eq) a+b<0)))
|
||||
triangleInequality a b | inl (inr a+b<0) | inr 0=a | inl (inr b<0) = inr (Equivalence.transitive eq (invContravariant additiveGroup) (Equivalence.transitive eq groupIsAbelian (+WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq (inverseWellDefined additiveGroup 0=a)) (invIdent additiveGroup)) 0=a) (Equivalence.reflexive eq))))
|
||||
triangleInequality a b | inl (inr a+b<0) | inr 0=a | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.symmetric eq 0=b)) identLeft) (Equivalence.reflexive eq) a+b<0))
|
||||
triangleInequality a b | inr 0=a+b with totality 0G a
|
||||
triangleInequality a b | inr 0=a+b | inl (inl 0<a) with totality 0G b
|
||||
triangleInequality a b | inr 0=a+b | inl (inl 0<a) | inl (inl 0<b) = exFalso (irreflexive {0G} (<WellDefined identLeft (Equivalence.symmetric eq 0=a+b) (ringAddInequalities 0<a 0<b)))
|
||||
triangleInequality a b | inr 0=a+b | inl (inl 0<a) | inl (inr b<0) = inl (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder b<0 (lemm2 _ b<0)) a))
|
||||
triangleInequality a b | inr 0=a+b | inl (inl 0<a) | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (lemm3 _ _ (Equivalence.transitive eq 0=a+b groupIsAbelian) 0=b)) 0<a))
|
||||
triangleInequality a b | inr 0=a+b | inl (inr a<0) with totality 0G b
|
||||
triangleInequality a b | inr 0=a+b | inl (inr a<0) | inl (inl 0<b) = inl (orderRespectsAddition (SetoidPartialOrder.<Transitive pOrder a<0 (lemm2 _ a<0)) b)
|
||||
triangleInequality a b | inr 0=a+b | inl (inr a<0) | inl (inr b<0) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq 0=a+b) identLeft (ringAddInequalities a<0 b<0)))
|
||||
triangleInequality a b | inr 0=a+b | inl (inr a<0) | inr 0=b = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (lemm3 _ _ (Equivalence.transitive eq 0=a+b groupIsAbelian) 0=b)) (Equivalence.reflexive eq) a<0))
|
||||
triangleInequality a b | inr 0=a+b | inr 0=a with totality 0G b
|
||||
triangleInequality a b | inr 0=a+b | inr 0=a | inl (inl 0<b) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (lemm3 a b 0=a+b 0=a)) 0<b))
|
||||
triangleInequality a b | inr 0=a+b | inr 0=a | inl (inr b<0) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (lemm3 a b 0=a+b 0=a)) (Equivalence.reflexive eq) b<0))
|
||||
triangleInequality a b | inr 0=a+b | inr 0=a | inr 0=b = inr (Equivalence.reflexive eq)
|
||||
|
||||
ringMinusFlipsOrder : {x : A} → (Ring.0R R) < x → (Group.inverse (Ring.additiveGroup R) x) < (Ring.0R R)
|
||||
ringMinusFlipsOrder {x = x} 0<x with totality (Ring.0R R) (Group.inverse (Ring.additiveGroup R) x)
|
||||
ringMinusFlipsOrder {x} 0<x | inl (inl 0<inv) = exFalso (SetoidPartialOrder.irreflexive pOrder bad')
|
||||
@@ -200,27 +139,6 @@ abstract
|
||||
r : y < x
|
||||
r = SetoidPartialOrder.<WellDefined pOrder (Equivalence.transitive eq (symmetric (Group.+Associative additiveGroup)) (Equivalence.transitive eq (Group.+WellDefined additiveGroup reflexive (invLeft)) identRight)) (Group.identLeft additiveGroup) (orderRespectsAddition q x)
|
||||
|
||||
absZero : abs (Ring.0R R) ≡ Ring.0R R
|
||||
absZero with totality (Ring.0R R) (Ring.0R R)
|
||||
absZero | inl (inl x) = exFalso (irreflexive x)
|
||||
absZero | inl (inr x) = exFalso (irreflexive x)
|
||||
absZero | inr x = refl
|
||||
|
||||
absNegation : (a : A) → (abs a) ∼ (abs (inverse a))
|
||||
absNegation a with totality 0R a
|
||||
absNegation a | inl (inl 0<a) with totality 0G (inverse a)
|
||||
absNegation a | inl (inl 0<a) | inl (inl 0<-a) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<-a (lemm2' a 0<a)))
|
||||
absNegation a | inl (inl 0<a) | inl (inr -a<0) = Equivalence.symmetric eq (invTwice additiveGroup a)
|
||||
absNegation a | inl (inl 0<a) | inr 0=-a = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq (invTwice additiveGroup a)) (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=-a))) (invIdent additiveGroup)) 0<a))
|
||||
absNegation a | inl (inr a<0) with totality 0G (inverse a)
|
||||
absNegation a | inl (inr a<0) | inl (inl 0<-a) = Equivalence.reflexive eq
|
||||
absNegation a | inl (inr a<0) | inl (inr -a<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder (<WellDefined (Equivalence.reflexive eq) (invTwice additiveGroup a) (lemm2 (inverse a) -a<0)) a<0))
|
||||
absNegation a | inl (inr a<0) | inr 0=-a = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq (Equivalence.transitive eq (inverseWellDefined additiveGroup 0=-a) (invTwice additiveGroup a))) (invIdent additiveGroup)) (Equivalence.reflexive eq) a<0))
|
||||
absNegation a | inr 0=a with totality 0G (inverse a)
|
||||
absNegation a | inr 0=a | inl (inl 0<-a) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=a)) (invIdent additiveGroup)) 0<-a))
|
||||
absNegation a | inr 0=a | inl (inr -a<0) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=a)) (invIdent additiveGroup)) (Equivalence.reflexive eq) -a<0))
|
||||
absNegation a | inr 0=a | inr 0=-a = Equivalence.transitive eq (Equivalence.symmetric eq 0=a) 0=-a
|
||||
|
||||
posTimesNeg : (a b : A) → (0G < a) → (b < 0G) → (a * b) < 0G
|
||||
posTimesNeg a b 0<a b<0 with orderRespectsMultiplication 0<a (lemm2 _ b<0)
|
||||
... | bl = <WellDefined (invTwice additiveGroup _) (Equivalence.reflexive eq) (lemm2' _ (<WellDefined (Equivalence.reflexive eq) ringMinusExtracts bl))
|
||||
@@ -229,67 +147,6 @@ abstract
|
||||
negTimesPos a b a<0 b<0 with orderRespectsMultiplication (lemm2 _ a<0) (lemm2 _ b<0)
|
||||
... | bl = <WellDefined (Equivalence.reflexive eq) twoNegativesTimes bl
|
||||
|
||||
absRespectsTimes : (a b : A) → abs (a * b) ∼ (abs a) * (abs b)
|
||||
absRespectsTimes a b with totality 0R a
|
||||
absRespectsTimes a b | inl (inl 0<a) with totality 0R b
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) | inl (inl 0<ab) = Equivalence.reflexive eq
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) | inl (inr ab<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder (orderRespectsMultiplication 0<a 0<b) ab<0))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inl 0<b) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=ab) (orderRespectsMultiplication 0<a 0<b)))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inl (inl 0<ab) with <WellDefined (Equivalence.reflexive eq) ringMinusExtracts (orderRespectsMultiplication 0<a (lemm2 b b<0))
|
||||
... | bl = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<ab (<WellDefined (invTwice additiveGroup _) (Equivalence.reflexive eq) (lemm2' _ bl))))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq 0=ab) (Equivalence.reflexive eq) (posTimesNeg a b 0<a b<0)))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) 0<ab))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inl (inr ab<0) = exFalso ((irreflexive {0G} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) (Equivalence.reflexive eq) ab<0)))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inr 0=ab = Equivalence.reflexive eq
|
||||
absRespectsTimes a b | inl (inr a<0) with totality 0R b
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inl (inl 0<ab) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder 0<ab (<WellDefined *Commutative (Equivalence.reflexive eq) (posTimesNeg b a 0<b a<0))))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts'
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq 0=ab *Commutative)) (Equivalence.reflexive eq) (posTimesNeg b a 0<b a<0)))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inl (inl 0<ab) = Equivalence.symmetric eq twoNegativesTimes
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inl (inr ab<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder (negTimesPos a b a<0 b<0) ab<0))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inr 0=ab = exFalso (exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=ab) (negTimesPos a b a<0 b<0))))
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) 0<ab))
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inl (inr ab<0) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) (Equivalence.reflexive eq) ab<0))
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inr 0=ab = Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (Equivalence.transitive eq (Equivalence.transitive eq timesZero (Equivalence.symmetric eq timesZero)) (*WellDefined (Equivalence.reflexive eq) 0=b))
|
||||
absRespectsTimes a b | inr 0=a with totality 0R b
|
||||
absRespectsTimes a b | inr 0=a | inl (inl 0<b) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inr 0=a | inl (inl 0<b) | inl (inl 0<ab) = Equivalence.reflexive eq
|
||||
absRespectsTimes a b | inr 0=a | inl (inl 0<b) | inl (inr ab<0) = Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.transitive eq (*WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) (Equivalence.transitive eq *Commutative timesZero))) (invIdent additiveGroup)) (Equivalence.transitive eq (Equivalence.symmetric eq timesZero) *Commutative)) (*WellDefined 0=a (Equivalence.reflexive eq))
|
||||
absRespectsTimes a b | inr 0=a | inl (inl 0<b) | inr 0=ab = Equivalence.reflexive eq
|
||||
absRespectsTimes a b | inr 0=a | inl (inr b<0) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inr 0=a | inl (inr b<0) | inl (inl 0<ab) = Equivalence.transitive eq (Equivalence.transitive eq (*WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) *Commutative) (Equivalence.transitive eq timesZero (Equivalence.transitive eq (Equivalence.symmetric eq (Equivalence.transitive eq *Commutative timesZero)) (*WellDefined 0=a (Equivalence.reflexive eq))))
|
||||
absRespectsTimes a b | inr 0=a | inl (inr b<0) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts
|
||||
absRespectsTimes a b | inr 0=a | inl (inr b<0) | inr 0=ab = Equivalence.transitive eq (Equivalence.transitive eq (*WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq)) *Commutative) (Equivalence.transitive eq timesZero (Equivalence.transitive eq (Equivalence.symmetric eq (Equivalence.transitive eq *Commutative timesZero)) (*WellDefined 0=a (Equivalence.reflexive eq))))
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b with totality 0R (a * b)
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) 0<ab))
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b | inl (inr ab<0) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) (Equivalence.reflexive eq) ab<0))
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b | inr 0=ab = Equivalence.reflexive eq
|
||||
|
||||
absNonnegative : {a : A} → (abs a < 0R) → False
|
||||
absNonnegative {a} pr with totality 0R a
|
||||
absNonnegative {a} pr | inl (inl x) = irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder x pr)
|
||||
absNonnegative {a} pr | inl (inr x) = irreflexive {0G} (SetoidPartialOrder.<Transitive pOrder (<WellDefined (Equivalence.reflexive eq) (invTwice additiveGroup a) (lemm2 (inverse a) pr)) x)
|
||||
absNonnegative {a} pr | inr x = irreflexive {0G} (<WellDefined (Equivalence.symmetric eq x) (Equivalence.reflexive eq) pr)
|
||||
|
||||
a-bPos : {a b : A} → ((a ∼ b) → False) → 0R < abs (a + inverse b)
|
||||
a-bPos {a} {b} a!=b with totality 0R (a + inverse b)
|
||||
a-bPos {a} {b} a!=b | inl (inl x) = x
|
||||
a-bPos {a} {b} a!=b | inl (inr x) = lemm2 _ x
|
||||
a-bPos {a} {b} a!=b | inr x = exFalso (a!=b (transferToRight additiveGroup (Equivalence.symmetric eq x)))
|
||||
|
||||
absZeroImpliesZero : {a : A} → abs a ∼ 0R → a ∼ 0R
|
||||
absZeroImpliesZero {a} a=0 with totality 0R a
|
||||
absZeroImpliesZero {a} a=0 | inl (inl 0<a) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) a=0 0<a))
|
||||
absZeroImpliesZero {a} a=0 | inl (inr a<0) = Equivalence.symmetric eq (lemm3 (inverse a) a (Equivalence.symmetric eq invLeft) (Equivalence.symmetric eq a=0))
|
||||
absZeroImpliesZero {a} a=0 | inr 0=a = a=0
|
||||
|
||||
halvePositive : (a : A) → 0R < (a + a) → 0R < a
|
||||
halvePositive a 0<2a with totality 0R a
|
||||
halvePositive a 0<2a | inl (inl x) = x
|
||||
@@ -305,71 +162,21 @@ abstract
|
||||
0<1 0!=1 | inl (inr x) = <WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq twoNegativesTimes identIsIdent) (orderRespectsMultiplication (lemm2 1R x) (lemm2 1R x))
|
||||
0<1 0!=1 | inr x = exFalso (0!=1 x)
|
||||
|
||||
addingAbsCannotShrink : {a b : A} → 0G < b → 0G < ((abs a) + b)
|
||||
addingAbsCannotShrink {a} {b} 0<b with totality 0G a
|
||||
addingAbsCannotShrink {a} {b} 0<b | inl (inl x) = <WellDefined identLeft (Equivalence.reflexive eq) (ringAddInequalities x 0<b)
|
||||
addingAbsCannotShrink {a} {b} 0<b | inl (inr x) = <WellDefined identLeft (Equivalence.reflexive eq) (ringAddInequalities (lemm2 a x) 0<b)
|
||||
addingAbsCannotShrink {a} {b} 0<b | inr x = <WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (Equivalence.symmetric eq identLeft) (+WellDefined x (Equivalence.reflexive eq))) 0<b
|
||||
|
||||
1<0False : (1R < 0R) → False
|
||||
1<0False 1<0 with orderRespectsMultiplication (lemm2 _ 1<0) (lemm2 _ 1<0)
|
||||
... | bl = exFalso (irreflexive (SetoidPartialOrder.<Transitive pOrder 1<0 (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (twoNegativesTimes) identIsIdent) bl)))
|
||||
|
||||
greaterZeroImpliesEqualAbs : {a : A} → 0R < a → a ∼ abs a
|
||||
greaterZeroImpliesEqualAbs {a} 0<a with totality 0R a
|
||||
... | inl (inl _) = Equivalence.reflexive eq
|
||||
... | inl (inr a<0) = exFalso (irreflexive (SetoidPartialOrder.<Transitive pOrder a<0 0<a))
|
||||
... | inr 0=a = exFalso (irreflexive (<WellDefined 0=a (Equivalence.reflexive eq) 0<a))
|
||||
|
||||
lessZeroImpliesEqualNegAbs : {a : A} → a < 0R → abs a ∼ inverse a
|
||||
lessZeroImpliesEqualNegAbs {a} a<0 with totality 0R a
|
||||
... | inl (inr _) = Equivalence.reflexive eq
|
||||
... | inl (inl 0<a) = exFalso (irreflexive (SetoidPartialOrder.<Transitive pOrder a<0 0<a))
|
||||
... | inr 0=a = exFalso (irreflexive (<WellDefined (Equivalence.reflexive eq) 0=a a<0))
|
||||
|
||||
absZeroIsZero : abs 0R ∼ 0R
|
||||
absZeroIsZero with totality 0R 0R
|
||||
... | inl (inl bad) = exFalso (irreflexive bad)
|
||||
... | inl (inr bad) = exFalso (irreflexive bad)
|
||||
... | inr _ = Equivalence.reflexive eq
|
||||
|
||||
greaterThanAbsImpliesGreaterThan0 : {a b : A} → (abs a) < b → 0R < b
|
||||
greaterThanAbsImpliesGreaterThan0 {a} {b} a<b with totality 0R a
|
||||
greaterThanAbsImpliesGreaterThan0 {a} {b} a<b | inl (inl 0<a) = SetoidPartialOrder.<Transitive pOrder 0<a a<b
|
||||
greaterThanAbsImpliesGreaterThan0 {a} {b} a<b | inl (inr a<0) = SetoidPartialOrder.<Transitive pOrder (lemm2 _ a<0) a<b
|
||||
greaterThanAbsImpliesGreaterThan0 {a} {b} a<b | inr 0=a = <WellDefined (Equivalence.symmetric eq 0=a) (Equivalence.reflexive eq) a<b
|
||||
|
||||
abs1Is1 : abs 1R ∼ 1R
|
||||
abs1Is1 with totality 0R 1R
|
||||
abs1Is1 | inl (inl 0<1) = Equivalence.reflexive eq
|
||||
abs1Is1 | inl (inr 1<0) = exFalso (1<0False 1<0)
|
||||
abs1Is1 | inr 0=1 = Equivalence.reflexive eq
|
||||
|
||||
absBoundedImpliesBounded : {a b : A} → abs a < b → a < b
|
||||
absBoundedImpliesBounded {a} {b} a<b with totality 0G a
|
||||
absBoundedImpliesBounded {a} {b} a<b | inl (inl x) = a<b
|
||||
absBoundedImpliesBounded {a} {b} a<b | inl (inr x) = SetoidPartialOrder.<Transitive pOrder x (SetoidPartialOrder.<Transitive pOrder (lemm2 a x) a<b)
|
||||
absBoundedImpliesBounded {a} {b} a<b | inr x = a<b
|
||||
|
||||
orderedImpliesCharNot2 : (0R ∼ 1R → False) → 1R + 1R ∼ 0R → False
|
||||
orderedImpliesCharNot2 0!=1 x = irreflexive (<WellDefined (identRight {0R}) x (ringAddInequalities (0<1 0!=1) (0<1 0!=1)))
|
||||
|
||||
open import Rings.InitialRing R
|
||||
open Equivalence eq
|
||||
|
||||
fromNIncreasing : (0R ∼ 1R → False) → (n : ℕ) → (fromN n) < (fromN (succ n))
|
||||
fromNIncreasing 0!=1 zero = <WellDefined reflexive (symmetric identRight) (0<1 0!=1)
|
||||
fromNIncreasing 0!=1 (succ n) = <WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (fromNIncreasing 0!=1 n) 1R)
|
||||
|
||||
fromNPreservesOrder : (0R ∼ 1R → False) → {a b : ℕ} → (a <N b) → (fromN a) < (fromN b)
|
||||
fromNPreservesOrder 0!=1 {zero} {succ zero} a<b = fromNIncreasing 0!=1 0
|
||||
fromNPreservesOrder 0!=1 {zero} {succ (succ b)} a<b = <Transitive (fromNPreservesOrder 0!=1 (succIsPositive b)) (fromNIncreasing 0!=1 (succ b))
|
||||
fromNPreservesOrder 0!=1 {succ a} {succ b} a<b = <WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (fromNPreservesOrder 0!=1 (canRemoveSuccFrom<N a<b)) 1R)
|
||||
|
||||
fromNPreservesOrder' : (0R ∼ 1R → False) → {a b : ℕ} → (fromN a) < (fromN b) → a <N b
|
||||
fromNPreservesOrder' nontrivial {a} {b} a<b with TotalOrder.totality ℕTotalOrder a b
|
||||
... | inl (inl x) = x
|
||||
... | inl (inr x) = exFalso (irreflexive (<Transitive a<b (fromNPreservesOrder nontrivial x)))
|
||||
... | inl (inr x) = exFalso (irreflexive (<Transitive a<b (fromNPreservesOrder (0<1 nontrivial) x)))
|
||||
... | inr x = exFalso (irreflexive (<WellDefined (fromNWellDefined x) reflexive a<b))
|
||||
|
||||
reciprocalPositive : (a b : A) → .(0<a : 0R < a) → (a * b ∼ 1R) → 0R < b
|
||||
|
@@ -52,8 +52,9 @@ simplePlus (times a b) (succ c) = succ (plus (times a b) c)
|
||||
simplePlus (times a b) (plus c d) = plus (times a b) (plus c d)
|
||||
simplePlus (times a b) (times c d) = plus (times a b) (times c d)
|
||||
|
||||
lemma : (a b c : A) → a + (b + c) ≡ b + (a + c)
|
||||
lemma a b c rewrite commutative a (b + c) | equalityCommutative (+Associative b c a) = applyEquality (b +_) (commutative _ _)
|
||||
private
|
||||
lemma : (a b c : A) → a + (b + c) ≡ b + (a + c)
|
||||
lemma a b c rewrite commutative a (b + c) | equalityCommutative (+Associative b c a) = applyEquality (b +_) (commutative _ _)
|
||||
|
||||
simplePlusIsPlus : (a b : Expr) → eval (simplePlus a b) ≡ eval (plus a b)
|
||||
simplePlusIsPlus (const x) (const x₁) = refl
|
||||
|
19
Setoids/Orders/Orders.agda
Normal file
19
Setoids/Orders/Orders.agda
Normal file
@@ -0,0 +1,19 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Orders.Total.Definition
|
||||
open import Orders.Partial.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Setoids.Orders where
|
||||
|
||||
partialOrderToSetoidPartialOrder : {a b : _} {A : Set a} (P : PartialOrder {a} A {b}) → SetoidPartialOrder (reflSetoid A) (PartialOrder._<_ P)
|
||||
SetoidPartialOrder.<WellDefined (partialOrderToSetoidPartialOrder P) a=b c=d a<c rewrite a=b | c=d = a<c
|
||||
SetoidPartialOrder.irreflexive (partialOrderToSetoidPartialOrder P) = PartialOrder.irreflexive P
|
||||
SetoidPartialOrder.<Transitive (partialOrderToSetoidPartialOrder P) = PartialOrder.<Transitive P
|
||||
|
||||
totalOrderToSetoidTotalOrder : {a b : _} {A : Set a} (T : TotalOrder {a} A {b}) → SetoidTotalOrder (partialOrderToSetoidPartialOrder (TotalOrder.order T))
|
||||
SetoidTotalOrder.totality (totalOrderToSetoidTotalOrder T) = TotalOrder.totality T
|
@@ -8,7 +8,7 @@ open import Functions
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Setoids.Orders where
|
||||
module Setoids.Orders.Partial.Definition where
|
||||
|
||||
record SetoidPartialOrder {a b c : _} {A : Set a} (S : Setoid {a} {b} A) (_<_ : Rel {a} {c} A) : Set (a ⊔ b ⊔ c) where
|
||||
open Setoid S
|
||||
@@ -19,27 +19,7 @@ record SetoidPartialOrder {a b c : _} {A : Set a} (S : Setoid {a} {b} A) (_<_ :
|
||||
_<=_ : Rel {a} {b ⊔ c} A
|
||||
a <= b = (a < b) || (a ∼ b)
|
||||
|
||||
record SetoidTotalOrder {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_<_ : Rel {a} {c} A} (P : SetoidPartialOrder S _<_) : Set (a ⊔ b ⊔ c) where
|
||||
open Setoid S
|
||||
field
|
||||
totality : (a b : A) → ((a < b) || (b < a)) || (a ∼ b)
|
||||
partial : SetoidPartialOrder S _<_
|
||||
partial = P
|
||||
min : A → A → A
|
||||
min a b with totality a b
|
||||
min a b | inl (inl a<b) = a
|
||||
min a b | inl (inr b<a) = b
|
||||
min a b | inr a=b = a
|
||||
max : A → A → A
|
||||
max a b with totality a b
|
||||
max a b | inl (inl a<b) = b
|
||||
max a b | inl (inr b<a) = a
|
||||
max a b | inr a=b = b
|
||||
|
||||
partialOrderToSetoidPartialOrder : {a b : _} {A : Set a} (P : PartialOrder {a} A {b}) → SetoidPartialOrder (reflSetoid A) (PartialOrder._<_ P)
|
||||
SetoidPartialOrder.<WellDefined (partialOrderToSetoidPartialOrder P) a=b c=d a<c rewrite a=b | c=d = a<c
|
||||
SetoidPartialOrder.irreflexive (partialOrderToSetoidPartialOrder P) = PartialOrder.irreflexive P
|
||||
SetoidPartialOrder.<Transitive (partialOrderToSetoidPartialOrder P) = PartialOrder.<Transitive P
|
||||
|
||||
totalOrderToSetoidTotalOrder : {a b : _} {A : Set a} (T : TotalOrder {a} A {b}) → SetoidTotalOrder (partialOrderToSetoidPartialOrder (TotalOrder.order T))
|
||||
SetoidTotalOrder.totality (totalOrderToSetoidTotalOrder T) = TotalOrder.totality T
|
33
Setoids/Orders/Total/Definition.agda
Normal file
33
Setoids/Orders/Total/Definition.agda
Normal file
@@ -0,0 +1,33 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Orders.Total.Definition
|
||||
open import Orders.Partial.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Setoids.Orders.Total.Definition where
|
||||
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
|
||||
record SetoidTotalOrder {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_<_ : Rel {a} {c} A} (P : SetoidPartialOrder S _<_) : Set (a ⊔ b ⊔ c) where
|
||||
open Setoid S
|
||||
field
|
||||
totality : (a b : A) → ((a < b) || (b < a)) || (a ∼ b)
|
||||
partial : SetoidPartialOrder S _<_
|
||||
partial = P
|
||||
min : A → A → A
|
||||
min a b with totality a b
|
||||
min a b | inl (inl a<b) = a
|
||||
min a b | inl (inr b<a) = b
|
||||
min a b | inr a=b = a
|
||||
max : A → A → A
|
||||
max a b with totality a b
|
||||
max a b | inl (inl a<b) = b
|
||||
max a b | inl (inr b<a) = a
|
||||
max a b | inr a=b = b
|
||||
|
||||
totalOrderToSetoidTotalOrder : {a b : _} {A : Set a} (T : TotalOrder {a} A {b}) → SetoidTotalOrder (partialOrderToSetoidPartialOrder (TotalOrder.order T))
|
||||
SetoidTotalOrder.totality (totalOrderToSetoidTotalOrder T) = TotalOrder.totality T
|
67
Setoids/Orders/Total/Lemmas.agda
Normal file
67
Setoids/Orders/Total/Lemmas.agda
Normal file
@@ -0,0 +1,67 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Orders.Total.Definition
|
||||
open import Orders.Partial.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders.Partial.Definition
|
||||
open import Setoids.Orders.Total.Definition
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Setoids.Orders.Total.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {c : _} {_<_ : A → A → Set c} {P : SetoidPartialOrder S _<_} (T : SetoidTotalOrder P) where
|
||||
|
||||
open SetoidTotalOrder T
|
||||
open SetoidPartialOrder P
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
maxInequalitiesR : {a b c : A} → (a < b) → (a < c) → (a < max b c)
|
||||
maxInequalitiesR {a} {b} {c} a<b a<c with totality b c
|
||||
... | inl (inl x) = a<c
|
||||
... | inl (inr x) = a<b
|
||||
... | inr x = a<c
|
||||
|
||||
minInequalitiesR : {a b c : A} → (a < b) → (a < c) → (a < min b c)
|
||||
minInequalitiesR {a} {b} {c} a<b a<c with totality b c
|
||||
... | inl (inl x) = a<b
|
||||
... | inl (inr x) = a<c
|
||||
... | inr x = a<b
|
||||
|
||||
maxInequalitiesL : {a b c : A} → (a < c) → (b < c) → (max a b < c)
|
||||
maxInequalitiesL {a} {b} {c} a<b a<c with totality a b
|
||||
... | inl (inl x) = a<c
|
||||
... | inl (inr x) = a<b
|
||||
... | inr x = a<c
|
||||
|
||||
minInequalitiesL : {a b c : A} → (a < c) → (b < c) → (min a b < c)
|
||||
minInequalitiesL {a} {b} {c} a<b a<c with totality a b
|
||||
... | inl (inl x) = a<b
|
||||
... | inl (inr x) = a<c
|
||||
... | inr x = a<b
|
||||
|
||||
minLessL : (a b : A) → min a b <= a
|
||||
minLessL a b with totality a b
|
||||
... | inl (inl x) = inr reflexive
|
||||
... | inl (inr x) = inl x
|
||||
... | inr x = inr reflexive
|
||||
|
||||
minLessR : (a b : A) → min a b <= b
|
||||
minLessR a b with totality a b
|
||||
... | inl (inl x) = inl x
|
||||
... | inl (inr x) = inr reflexive
|
||||
... | inr x = inr x
|
||||
|
||||
maxGreaterL : (a b : A) → a <= max a b
|
||||
maxGreaterL a b with totality a b
|
||||
... | inl (inl x) = inl x
|
||||
... | inl (inr x) = inr reflexive
|
||||
... | inr x = inr x
|
||||
|
||||
maxGreaterR : (a b : A) → b <= max a b
|
||||
maxGreaterR a b with totality a b
|
||||
... | inl (inl x) = inr reflexive
|
||||
... | inl (inr x) = inl x
|
||||
... | inr x = inr reflexive
|
Reference in New Issue
Block a user