Minor perf improvements? (#60)

This commit is contained in:
Patrick Stevens
2019-11-02 08:57:15 +00:00
committed by GitHub
parent a66080b8ae
commit 55995ea801
7 changed files with 695 additions and 676 deletions

View File

@@ -37,6 +37,8 @@ open import Fields.CauchyCompletion.Addition order F charNot2
open import Fields.CauchyCompletion.Setoid order F charNot2
open import Fields.CauchyCompletion.Comparison order F charNot2
abstract
chain : {a b : A} (c : CauchyCompletion) (a r<C c) (c <Cr b) a < b
chain {a} {b} c (betweenAC , (0<betweenAC ,, (Nac , prAC))) (betweenCB , (0<betweenCB ,, (Nb , prBC))) = SetoidPartialOrder.transitive pOrder (SetoidPartialOrder.transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<betweenAC a)) (<WellDefined groupIsAbelian (Equivalence.reflexive eq) (SetoidPartialOrder.transitive pOrder (prAC (succ Nac +N Nb) (le Nb (applyEquality succ (Semiring.commutative Semiring Nb Nac)))) (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<betweenCB (index (Sequence.tail (CauchyCompletion.elts c)) (Nac +N Nb))))))) (prBC (succ Nac +N Nb) (le Nac refl))
@@ -81,7 +83,7 @@ rationalApproximatelyAboveIsAbove a e 0<e with halve charNot2 e
... | 0<e/8 with CauchyCompletion.converges a e/8 0<e/8
... | N2 , cauchyBeyondN2 = e/8 , (0<e/8 ,, ((N +N N2) , ans2))
where
ans2 : (m : ) N +N N2 <N m (e/8 + index (CauchyCompletion.elts a) m) < (index (CauchyCompletion.elts a) (succ N) + e/2) -- TODO not sure this is actually true
ans2 : (m : ) N +N N2 <N m (e/8 + index (CauchyCompletion.elts a) m) < (index (CauchyCompletion.elts a) (succ N) + e/2)
ans2 m <m with cauchyBeyondN {m} {succ N} (inequalityShrinkLeft <m) (le 0 refl)
... | absam-aN<e/4 with totality 0R ((index (CauchyCompletion.elts a) m) + inverse (index (CauchyCompletion.elts a) (succ N)))
ans2 m <m | am-aN<e/4 | inl (inl 0<am-aN) = SetoidPartialOrder.transitive pOrder (<WellDefined groupIsAbelian (Equivalence.transitive eq groupIsAbelian +Associative) (orderRespectsAddition (<WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) invLeft)) identRight) (Equivalence.reflexive eq) (orderRespectsAddition am-aN<e/4 (index (CauchyCompletion.elts a) (succ N)))) e/8)) (<WellDefined (Equivalence.reflexive eq) groupIsAbelian (orderRespectsAddition (<WellDefined (Equivalence.reflexive eq) prE/4 (orderRespectsAddition (halfLess e/8 e/4 0<e/4 prE/8) e/4)) (index (CauchyCompletion.elts a) (succ N))))

View File

@@ -50,12 +50,14 @@ CauchyCompletion.converges (record { elts = a ; converges = aConv } *C record {
where
boundBoth : A
boundBoth = aBound + bBound
abstract
ab<bb : aBound < boundBoth
ab<bb = <WellDefined identLeft groupIsAbelian (orderRespectsAddition {0R} {bBound} (greaterThanAbsImpliesGreaterThan0 (prBBound (succ Nb) (le 0 refl))) aBound)
bb<bb : bBound < boundBoth
bb<bb = <WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition {0R} {aBound} (greaterThanAbsImpliesGreaterThan0 (prABound (succ Na) (le 0 refl))) bBound)
0<boundBoth : 0R < boundBoth
0<boundBoth = <WellDefined identLeft (Equivalence.reflexive eq) (ringAddInequalities (greaterThanAbsImpliesGreaterThan0 (prABound (succ Na) (le 0 refl))) (greaterThanAbsImpliesGreaterThan0 (prBBound (succ Nb) (le 0 refl))))
abstract
1/boundBoothAndPr : Sg A λ i i * (aBound + bBound) 1R
1/boundBoothAndPr = allInvertible boundBoth λ pr irreflexive (<WellDefined (Equivalence.reflexive eq) pr 0<boundBoth)
1/boundBooth : A
@@ -66,6 +68,7 @@ CauchyCompletion.converges (record { elts = a ; converges = aConv } *C record {
... | _ , pr = pr
0<1/boundBooth : 0G < 1/boundBooth
0<1/boundBooth = inversePositiveIsPositive 1/boundBoothPr 0<boundBoth
abstract
miniEAndPr : Sg A (λ i (i + i) (e * 1/boundBooth))
miniEAndPr = halve charNot2 (e * 1/boundBooth)
miniE : A
@@ -76,6 +79,7 @@ CauchyCompletion.converges (record { elts = a ; converges = aConv } *C record {
... | _ , pr = pr
0<miniE : 0R < miniE
0<miniE = halvePositive miniE (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq miniEPr) (orderRespectsMultiplication 0<e 0<1/boundBooth))
abstract
reallyNAAndPr : Sg (λ N {m n : } N <N m N <N n abs (index a m + inverse (index a n)) < miniE)
reallyNAAndPr = aConv miniE 0<miniE
reallyNa :
@@ -101,6 +105,7 @@ CauchyCompletion.converges (record { elts = a ; converges = aConv } *C record {
Na<m = inequalityShrinkLeft N<m
Nb<n : Nb <N n
Nb<n = inequalityShrinkLeft (inequalityShrinkRight {Na} N<n)
abstract
sum : {m n : } (reallyNa +N reallyNb) <N m (reallyNa +N reallyNb) <N n (boundBoth * ((abs (index b m + inverse (index b n))) + (abs (index a m + inverse (index a n))))) < e
sum {m} {n} <m <n = <WellDefined *Commutative (Equivalence.transitive eq (*WellDefined miniEPr (Equivalence.reflexive eq)) (Equivalence.transitive eq (Equivalence.symmetric eq *Associative) (Equivalence.transitive eq (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) 1/boundBoothPr) *Commutative) (identIsIdent)))) (ringCanMultiplyByPositive {c = boundBoth} 0<boundBoth (ringAddInequalities (reallyNbPr {m} {n} (inequalityShrinkRight <m) (inequalityShrinkRight <n)) (reallyNaPr {m} {n} (inequalityShrinkLeft <m) (inequalityShrinkLeft <n))))
q : ((boundBoth * (abs (index b m + inverse (index b n)))) + (boundBoth * (abs (index a m + inverse (index a n))))) < e
@@ -152,11 +157,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 order)))) (Equivalence.reflexive eq) 0<e
absBoundedImpliesBounded : {a b : A} abs a < b a < b
absBoundedImpliesBounded {a} {b} a<b with SetoidTotalOrder.totality tOrder 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
abstract
multiplicationWellDefinedLeft' : (0!=1 : 0R 1R False) (a b c : CauchyCompletion) Setoid.__ cauchyCompletionSetoid a b Setoid.__ cauchyCompletionSetoid (a *C c) (b *C c)
multiplicationWellDefinedLeft' 0!=1 a b c a=b ε 0<e = N , ans

View File

@@ -16,6 +16,8 @@ open import Numbers.Naturals.Naturals
module Fields.Lemmas {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A A A} {_*_ : A A A} {R : Ring S _+_ _*_} (F : Field R) where
abstract
open Setoid S
open Field F
open Ring R

View File

@@ -19,6 +19,8 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Fields.Orders.Lemmas {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A A A} {_*_ : A A A} {_<_ : Rel {_} {o} A} {R : Ring S _+_ _*_} {pOrder : SetoidPartialOrder S _<_} {tOrder : SetoidTotalOrder pOrder} {F : Field R} (oF : OrderedField F tOrder) where
abstract
open Ring R
open Group additiveGroup
open OrderedRing (OrderedField.oRing oF)

View File

@@ -12,6 +12,8 @@ open import Sets.EquivalenceRelations
module Rings.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} {_*_ : A A A} (R : Ring S _+_ _*_) where
abstract
open Setoid S
open Ring R
open Group additiveGroup

View File

@@ -28,6 +28,8 @@ record OrderedRing {p : _} {_<_ : Rel {_} {p} A} {pOrder : SetoidPartialOrder S
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 SetoidTotalOrder.totality order 0R a
absWellDefined a b a=b | inl (inl 0<a) with SetoidTotalOrder.totality order 0R b

View File

@@ -16,6 +16,8 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Orders.Lemmas {n m p : _} {A : Set n} {S : Setoid {n} {m} A} {_+_ : A A A} {_*_ : A A A} {_<_ : Rel {_} {p} A} {R : Ring S _+_ _*_} {pOrder : SetoidPartialOrder S _<_} {tOrder : SetoidTotalOrder pOrder} (order : OrderedRing R tOrder) where
abstract
open OrderedRing order
open Setoid S
open SetoidPartialOrder pOrder
@@ -377,3 +379,9 @@ abs1Is1 | inr 0=1 = Equivalence.reflexive eq
charNot2ImpliesNontrivial : ((1R + 1R) 0R False) (0R 1R) False
charNot2ImpliesNontrivial charNot2 0=1 = charNot2 (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq 0=1) (Equivalence.symmetric eq 0=1)) identRight)
absBoundedImpliesBounded : {a b : A} abs a < b a < b
absBoundedImpliesBounded {a} {b} a<b with SetoidTotalOrder.totality tOrder 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