Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

Real numbers, based on Cauchy sequences #2487

Draft
wants to merge 16 commits into
base: master
Choose a base branch
from
Draft
39 changes: 38 additions & 1 deletion src/Codata/Guarded/Stream/Relation/Binary/Pointwise.agda
Original file line number Diff line number Diff line change
Expand Up @@ -12,7 +12,7 @@ open import Codata.Guarded.Stream as Stream using (Stream; head; tail)
open import Data.Nat.Base using (ℕ; zero; suc)
open import Function.Base using (_∘_; _on_)
open import Level using (Level; _⊔_)
open import Relation.Binary.Core using (REL; _⇒_)
open import Relation.Binary.Core using (REL; Rel; _⇒_)
open import Relation.Binary.Bundles using (Setoid)
open import Relation.Binary.Definitions
using (Reflexive; Sym; Trans; Antisym; Symmetric; Transitive)
Expand Down Expand Up @@ -104,6 +104,43 @@ drop⁺ : ∀ n → Pointwise R ⇒ (Pointwise R on Stream.drop n)
drop⁺ zero as≈bs = as≈bs
drop⁺ (suc n) as≈bs = drop⁺ n (as≈bs .tail)

------------------------------------------------------------------------
-- Algebraic properties

module _ {A : Set a} {_≈_ : Rel A ℓ} where

open import Algebra.Definitions

private
variable
_∙_ : A → A → A
_⁻¹ : A → A
ε : A

assoc : Associative _≈_ _∙_ → Associative (Pointwise _≈_) (Stream.zipWith _∙_)
head (assoc assoc₁ xs ys zs) = assoc₁ (head xs) (head ys) (head zs)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Here (and for all the newly added bits): assoc assoc₁ is super ugly. One of the two should be renamed.

tail (assoc assoc₁ xs ys zs) = assoc assoc₁ (tail xs) (tail ys) (tail zs)

comm : Commutative _≈_ _∙_ → Commutative (Pointwise _≈_) (Stream.zipWith _∙_)
head (comm comm₁ xs ys) = comm₁ (head xs) (head ys)
tail (comm comm₁ xs ys) = comm comm₁ (tail xs) (tail ys)

identityˡ : LeftIdentity _≈_ ε _∙_ → LeftIdentity (Pointwise _≈_) (Stream.repeat ε) (Stream.zipWith _∙_)
head (identityˡ identityˡ₁ xs) = identityˡ₁ (head xs)
tail (identityˡ identityˡ₁ xs) = identityˡ identityˡ₁ (tail xs)

identityʳ : RightIdentity _≈_ ε _∙_ → RightIdentity (Pointwise _≈_) (Stream.repeat ε) (Stream.zipWith _∙_)
head (identityʳ identityʳ₁ xs) = identityʳ₁ (head xs)
tail (identityʳ identityʳ₁ xs) = identityʳ identityʳ₁ (tail xs)

inverseˡ : LeftInverse _≈_ ε _⁻¹ _∙_ → LeftInverse (Pointwise _≈_) (Stream.repeat ε) (Stream.map _⁻¹) (Stream.zipWith _∙_)
head (inverseˡ inverseˡ₁ xs) = inverseˡ₁ (head xs)
tail (inverseˡ inverseˡ₁ xs) = inverseˡ inverseˡ₁ (tail xs)

inverseʳ : RightInverse _≈_ ε _⁻¹ _∙_ → RightInverse (Pointwise _≈_) (Stream.repeat ε) (Stream.map _⁻¹) (Stream.zipWith _∙_)
head (inverseʳ inverseʳ₁ xs) = inverseʳ₁ (head xs)
tail (inverseʳ inverseʳ₁ xs) = inverseʳ inverseʳ₁ (tail xs)

------------------------------------------------------------------------
-- Pointwise Equality as a Bisimilarity

Expand Down
84 changes: 84 additions & 0 deletions src/Data/Rational/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -49,6 +49,7 @@ open import Data.Sum.Base as Sum using (inj₁; inj₂; [_,_]′; _⊎_)
import Data.Sign.Base as Sign
open import Function.Base using (_∘_; _∘′_; _∘₂_; _$_; flip)
open import Function.Definitions using (Injective)
open import Function.Metric.Rational as Metric hiding (Symmetric)
open import Level using (0ℓ)
open import Relation.Binary
open import Relation.Binary.Morphism.Structures
Expand Down Expand Up @@ -1799,6 +1800,89 @@ toℚᵘ-homo-∣-∣ (mkℚ -[1+ _ ] _ _) = *≡* refl
∣∣p∣∣≡∣p∣ : ∀ p → ∣ ∣ p ∣ ∣ ≡ ∣ p ∣
∣∣p∣∣≡∣p∣ p = 0≤p⇒∣p∣≡p (0≤∣p∣ p)

------------------------------------------------------------------------
-- Metric space
------------------------------------------------------------------------

private
d : ℚ → ℚ → ℚ
d p q = ∣ p - q ∣
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

This is really very frustrating. Anyone have any nice ideas how we can make this a function in Data.Rational.Base? I guess we could use some weird unicode bars?

Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

No idea on how to name this properly in Data.Rational.Base. But I was surprised that it is ℚ-valued - do we not have non-negative rationals as a type? That way this would be nonNegative by construction instead of requiring a proof.

Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

See my global comment: we should have ℚ⁺ as a first class thing in the library before we proceed!

Copy link
Contributor

@jamesmckinna jamesmckinna Oct 3, 2024

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Re: notation. UPDATED since I learnt how to write the symbols ;-)
It is a norm, so does using the following offer any assistance?

Suggested change
d p q = ∣ p - q ∣
-- distance function
∥_─_∥ :
∥ p ─ q ∥ = ∣ p - q ∣

which reuses the 'long minus' \ - - - (used in Data.List.Relation.Unary.Any so hopefully the two uses won't be in scope at the same time...), and whose "weird Unicode bars" here are in fact simply given by \ | |... so they aren't too bad, either...


d-cong : Congruent _≡_ d
d-cong = cong₂ _

d-nonNegative : ∀ {p q} → 0ℚ ≤ d p q
d-nonNegative {p} {q} = nonNegative⁻¹ _ {{∣-∣-nonNeg (p - q)}}

d-definite : Definite _≡_ d
d-definite {p} refl = cong ∣_∣ (+-inverseʳ p)

d-indiscernable : Indiscernable _≡_ d
d-indiscernable {p} {q} ∣p-q∣≡0 = begin
p ≡⟨ +-identityʳ p ⟨
p - 0ℚ ≡⟨ cong (_-_ p) (∣p∣≡0⇒p≡0 (p - q) ∣p-q∣≡0) ⟨
p - (p - q) ≡⟨ cong (_+_ p) (neg-distrib-+ p (- q)) ⟩
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

It feels like p - (p - q) ≡ q should be a lemma in GroupProperties?

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

It actually requires commutativity! (this is hidden in neg-distrib-+). As far as I can tell it doesn't exist yet.

Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

The following proof looks commutativity-free to me?

p - (p - q) = p + - (p + - q) = p + (-p + q) = (p + -p) + q = 0# + q = q

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

The step - (p + - q) = - p + q requires commutativity

Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Oh whoops. Yes, you don't get distributivity for free.

Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

See Algebra.Properties.AbelianGroup.⁻¹-anti-homo‿- added in #2349 ?

p + (- p - - q) ≡⟨ +-assoc p (- p) (- - q) ⟨
(p - p) - - q ≡⟨ cong₂ _+_ (+-inverseʳ p) (⁻¹-involutive q) ⟩
0ℚ + q ≡⟨ +-identityˡ q ⟩
q ∎
where
open ≡-Reasoning
open GroupProperties +-0-group

d-sym : Metric.Symmetric d
d-sym p q = begin
∣ p - q ∣ ≡˘⟨ ∣-p∣≡∣p∣ (p - q) ⟩
∣ - (p - q) ∣ ≡⟨ cong ∣_∣ (⁻¹-anti-homo-// p q) ⟩
∣ q - p ∣ ∎
where
open ≡-Reasoning
open GroupProperties +-0-group

d-triangle : TriangleInequality d
d-triangle p q r = begin
∣ p - r ∣ ≡⟨ cong (λ # → ∣ # - r ∣) (+-identityʳ p) ⟨
∣ p + 0ℚ - r ∣ ≡⟨ cong (λ # → ∣ p + # - r ∣) (+-inverseˡ q) ⟨
∣ p + (- q + q) - r ∣ ≡⟨ cong (λ # → ∣ # - r ∣) (+-assoc p (- q) q) ⟨
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Isn't this 'middle4' for semigroup? (i.e. all these +-assoc moves).

∣ ((p - q) + q) - r ∣ ≡⟨ cong ∣_∣ (+-assoc (p - q) q (- r)) ⟩
∣ (p - q) + (q - r) ∣ ≤⟨ ∣p+q∣≤∣p∣+∣q∣ (p - q) (q - r) ⟩
∣ p - q ∣ + ∣ q - r ∣ ∎
where open ≤-Reasoning

d-isProtoMetric : IsProtoMetric _≡_ d
d-isProtoMetric = record
{ isPartialOrder = ≤-isPartialOrder
; ≈-isEquivalence = isEquivalence
; cong = cong₂ _
; nonNegative = λ {p q} → d-nonNegative {p} {q}
}

d-isPreMetric : IsPreMetric _≡_ d
d-isPreMetric = record
{ isProtoMetric = d-isProtoMetric
; ≈⇒0 = d-definite
}

d-isQuasiSemiMetric : IsQuasiSemiMetric _≡_ d
d-isQuasiSemiMetric = record
{ isPreMetric = d-isPreMetric
; 0⇒≈ = d-indiscernable
}

d-isSemiMetric : IsSemiMetric _≡_ d
d-isSemiMetric = record
{ isQuasiSemiMetric = d-isQuasiSemiMetric
; sym = d-sym
}

d-isMetric : IsMetric _≡_ d
d-isMetric = record
{ isSemiMetric = d-isSemiMetric
; triangle = d-triangle
}

d-metric : Metric _ _
d-metric = record { isMetric = d-isMetric }

------------------------------------------------------------------------
-- DEPRECATED NAMES
Expand Down
217 changes: 217 additions & 0 deletions src/Data/Real/Base.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,217 @@
------------------------------------------------------------------------
-- The Agda standard library
--
-- Real numbers
------------------------------------------------------------------------

{-# OPTIONS --safe --cubical-compatible --guardedness #-}

module Data.Real.Base where

open import Codata.Guarded.Stream
open import Codata.Guarded.Stream.Properties
open import Data.Integer.Base using (+<+)
open import Data.Nat.Base as ℕ using (z≤n; s≤s)
import Data.Nat.Properties as ℕ
open import Data.Product.Base hiding (map)
open import Data.Rational.Base as ℚ hiding (_+_; -_)
open import Data.Rational.Properties
open import Data.Rational.Solver
open import Data.Rational.Unnormalised using (*<*)
open import Function.Base using (_$_)
open import Relation.Binary.PropositionalEquality using (_≡_; refl; cong; cong₂)
open import Relation.Nullary

open import Function.Metric.Rational.CauchySequence d-metric public renaming (CauchySequence to ℝ)

fromℚ : ℚ → ℝ
fromℚ = embed

0ℝ : ℝ
0ℝ = fromℚ 0ℚ

_+_ : ℝ → ℝ → ℝ
sequence (x + y) = zipWith ℚ._+_ (sequence x) (sequence y)
isCauchy (x + y) ε = proj₁ [x] ℕ.+ proj₁ [y] , λ {m} {n} m≥N n≥N → begin-strict
∣ lookup (zipWith ℚ._+_ (sequence x) (sequence y)) m - lookup (zipWith ℚ._+_ (sequence x) (sequence y)) n ∣
≡⟨ cong₂ (λ a b → ∣ a - b ∣)
(lookup-zipWith m ℚ._+_ (sequence x) (sequence y))
(lookup-zipWith n ℚ._+_ (sequence x) (sequence y))
∣ (lookup (sequence x) m ℚ.+ lookup (sequence y) m) - (lookup (sequence x) n ℚ.+ lookup (sequence y) n) ∣
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Looks like some notation for lookup (sequence x) m would make all of this way more readable.

≡⟨ cong ∣_∣ (lemma (lookup (sequence x) m) (lookup (sequence y) m) (lookup (sequence x) n) (lookup (sequence y) n)) ⟩
∣ (lookup (sequence x) m - lookup (sequence x) n) ℚ.+ (lookup (sequence y) m - lookup (sequence y) n) ∣
≤⟨ ∣p+q∣≤∣p∣+∣q∣
(lookup (sequence x) m - lookup (sequence x) n)
(lookup (sequence y) m - lookup (sequence y) n)
∣ lookup (sequence x) m - lookup (sequence x) n ∣ ℚ.+ ∣ lookup (sequence y) m - lookup (sequence y) n ∣
<⟨ +-mono-<
(proj₂ [x]
(ℕ.≤-trans (ℕ.m≤m+n (proj₁ [x]) (proj₁ [y])) m≥N)
(ℕ.≤-trans (ℕ.m≤m+n (proj₁ [x]) (proj₁ [y])) n≥N)
)
(proj₂ [y]
(ℕ.≤-trans (ℕ.m≤n+m (proj₁ [y]) (proj₁ [x])) m≥N)
(ℕ.≤-trans (ℕ.m≤n+m (proj₁ [y]) (proj₁ [x])) n≥N)
)
½ * ε ℚ.+ ½ * ε
≡⟨ *-distribʳ-+ ε ½ ½ ⟨
1ℚ * ε
≡⟨ *-identityˡ ε ⟩
ε ∎
where
open ≤-Reasoning
instance _ : Positive (½ * ε)
_ = pos*pos⇒pos ½ ε
[x] = isCauchy x (½ * ε)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

you name isCauchy but only use its first component -- maybe name that instead, to save all the projections above?

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I do use the second component, on lines 50 and 54. But I use the first component a lot more

[y] = isCauchy y (½ * ε)

lemma : ∀ a b c d → (a ℚ.+ b) - (c ℚ.+ d) ≡ (a - c) ℚ.+ (b - d)
lemma = solve 4 (λ a b c d → ((a :+ b) :- (c :+ d)) , ((a :- c) :+ (b :- d))) refl
where open +-*-Solver

-_ : ℝ → ℝ
sequence (- x) = map ℚ.-_ (sequence x)
isCauchy (- x) ε = proj₁ (isCauchy x ε) , λ {m} {n} m≥N n≥N → begin-strict
∣ lookup (map ℚ.-_ (sequence x)) m - lookup (map ℚ.-_ (sequence x)) n ∣
≡⟨ cong₂ (λ a b → ∣ a - b ∣) (lookup-map m ℚ.-_ (sequence x)) (lookup-map n ℚ.-_ (sequence x)) ⟩
∣ ℚ.- lookup (sequence x) m - ℚ.- lookup (sequence x) n ∣
≡⟨ cong ∣_∣ (lemma (lookup (sequence x) m) (lookup (sequence x) n)) ⟩
∣ ℚ.- (lookup (sequence x) m - lookup (sequence x) n) ∣
≡⟨ ∣-p∣≡∣p∣ (lookup (sequence x) m - lookup (sequence x) n) ⟩
∣ lookup (sequence x) m - lookup (sequence x) n ∣
<⟨ proj₂ (isCauchy x ε) m≥N n≥N ⟩
ε ∎
where
open ≤-Reasoning
lemma : ∀ a b → ℚ.- a - ℚ.- b ≡ ℚ.- (a - b)
lemma = solve 2 (λ a b → (:- a) :- (:- b) , (:- (a :- b))) refl
where open +-*-Solver

_*ₗ_ : ℚ → ℝ → ℝ
sequence (p *ₗ x) = map (p *_) (sequence x)
isCauchy (p *ₗ x) ε with p ≟ 0ℚ
... | yes p≡0 = 0 , λ {m} {n} _ _ → begin-strict
∣ lookup (map (p *_) (sequence x)) m - lookup (map (p *_) (sequence x)) n ∣
≡⟨ cong₂ (λ a b → ∣ a - b ∣) (lookup-map m (p *_) (sequence x)) (lookup-map n (p *_) (sequence x)) ⟩
∣ p * lookup (sequence x) m - p * lookup (sequence x) n ∣
≡⟨ cong (λ # → ∣ p * lookup (sequence x) m ℚ.+ # ∣) (neg-distribʳ-* p (lookup (sequence x) n)) ⟩
∣ p * lookup (sequence x) m ℚ.+ p * ℚ.- lookup (sequence x) n ∣
≡⟨ cong ∣_∣ (*-distribˡ-+ p (lookup (sequence x) m) (ℚ.- lookup (sequence x) n)) ⟨
∣ p * (lookup (sequence x) m - lookup (sequence x) n) ∣
≡⟨ cong (λ # → ∣ # * (lookup (sequence x) m - lookup (sequence x) n) ∣) p≡0 ⟩
∣ 0ℚ * (lookup (sequence x) m - lookup (sequence x) n) ∣
≡⟨ cong ∣_∣ (*-zeroˡ (lookup (sequence x) m - lookup (sequence x) n)) ⟩
∣ 0ℚ ∣
≡⟨⟩
0ℚ
<⟨ positive⁻¹ ε ⟩
ε ∎
where open ≤-Reasoning
... | no p≢0 = proj₁ (isCauchy x (1/ ∣ p ∣ * ε)) , λ {m} {n} m≥N n≥N → begin-strict
∣ lookup (map (p *_) (sequence x)) m - lookup (map (p *_) (sequence x)) n ∣
≡⟨ cong₂ (λ a b → ∣ a - b ∣) (lookup-map m (p *_) (sequence x)) (lookup-map n (p *_) (sequence x)) ⟩
∣ p * lookup (sequence x) m - p * lookup (sequence x) n ∣
≡⟨ cong (λ # → ∣ p * lookup (sequence x) m ℚ.+ # ∣) (neg-distribʳ-* p (lookup (sequence x) n)) ⟩
∣ p * lookup (sequence x) m ℚ.+ p * ℚ.- lookup (sequence x) n ∣
≡⟨ cong ∣_∣ (*-distribˡ-+ p (lookup (sequence x) m) (ℚ.- lookup (sequence x) n)) ⟨
∣ p * (lookup (sequence x) m - lookup (sequence x) n) ∣
≡⟨ ∣p*q∣≡∣p∣*∣q∣ p (lookup (sequence x) m - lookup (sequence x) n) ⟩
∣ p ∣ * ∣ lookup (sequence x) m - lookup (sequence x) n ∣
<⟨ *-monoʳ-<-pos ∣ p ∣ (proj₂ (isCauchy x (1/ ∣ p ∣ * ε)) m≥N n≥N) ⟩
∣ p ∣ * (1/ ∣ p ∣ * ε)
≡⟨ *-assoc ∣ p ∣ (1/ ∣ p ∣) ε ⟨
(∣ p ∣ * 1/ ∣ p ∣) * ε
≡⟨ cong (_* ε) (*-inverseʳ ∣ p ∣) ⟩
1ℚ * ε
≡⟨ *-identityˡ ε ⟩
ε ∎
where
open ≤-Reasoning
instance _ : NonZero ∣ p ∣
_ = ≢-nonZero {∣ p ∣} λ ∣p∣≡0 → p≢0 (∣p∣≡0⇒p≡0 p ∣p∣≡0)
instance _ : Positive ∣ p ∣
_ = nonNeg∧nonZero⇒pos ∣ p ∣ {{∣-∣-nonNeg p}}
instance _ : Positive (1/ ∣ p ∣)
_ = 1/pos⇒pos ∣ p ∣
instance _ : Positive (1/ ∣ p ∣ * ε)
_ = pos*pos⇒pos (1/ ∣ p ∣) ε

square : ℝ → ℝ
sequence (square x) = map (λ p → p * p) (sequence x)
isCauchy (square x) ε = B ℕ.⊔ proj₁ (isCauchy x (1/ (b ℚ.+ b) * ε)) , λ {m} {n} m≥N n≥N → begin-strict
∣ lookup (map (λ p → p * p) (sequence x)) m - lookup (map (λ p → p * p) (sequence x)) n ∣
≡⟨ cong₂ (λ a b → ∣ a - b ∣) (lookup-map m _ (sequence x)) (lookup-map n _ (sequence x)) ⟩
∣ lookup (sequence x) m * lookup (sequence x) m - lookup (sequence x) n * lookup (sequence x) n ∣
≡⟨ cong ∣_∣ (lemma (lookup (sequence x) m) (lookup (sequence x) n)) ⟩
∣ (lookup (sequence x) m ℚ.+ lookup (sequence x) n) * (lookup (sequence x) m - lookup (sequence x) n) ∣
≡⟨ ∣p*q∣≡∣p∣*∣q∣ (lookup (sequence x) m ℚ.+ lookup (sequence x) n) (lookup (sequence x) m - lookup (sequence x) n) ⟩
∣ lookup (sequence x) m ℚ.+ lookup (sequence x) n ∣ * ∣ lookup (sequence x) m - lookup (sequence x) n ∣
≤⟨ *-monoʳ-≤-nonNeg ∣ lookup (sequence x) m - lookup (sequence x) n ∣
{{∣-∣-nonNeg (lookup (sequence x) m - lookup (sequence x) n)}}
(∣p+q∣≤∣p∣+∣q∣ (lookup (sequence x) m) (lookup (sequence x) n))
(∣ lookup (sequence x) m ∣ ℚ.+ ∣ lookup (sequence x) n ∣) * ∣ lookup (sequence x) m - lookup (sequence x) n ∣
≤⟨ *-monoʳ-≤-nonNeg ∣ lookup (sequence x) m - lookup (sequence x) n ∣
{{∣-∣-nonNeg (lookup (sequence x) m - lookup (sequence x) n)}}
(<⇒≤ (+-mono-<
(b-prop (ℕ.≤-trans (ℕ.m≤m⊔n B (proj₁ (isCauchy x (1/ (b ℚ.+ b) * ε)))) m≥N))
(b-prop (ℕ.≤-trans (ℕ.m≤m⊔n B (proj₁ (isCauchy x (1/ (b ℚ.+ b) * ε)))) n≥N))
))
(b ℚ.+ b) * ∣ lookup (sequence x) m - lookup (sequence x) n ∣
<⟨ *-monoʳ-<-pos (b ℚ.+ b) (proj₂ (isCauchy x (1/ (b ℚ.+ b) * ε))
(ℕ.≤-trans (ℕ.m≤n⊔m B (proj₁ (isCauchy x (1/ (b ℚ.+ b) * ε)))) m≥N)
(ℕ.≤-trans (ℕ.m≤n⊔m B (proj₁ (isCauchy x (1/ (b ℚ.+ b) * ε)))) n≥N)
) ⟩
(b ℚ.+ b) * (1/ (b ℚ.+ b) * ε)
≡⟨ *-assoc (b ℚ.+ b) (1/ (b ℚ.+ b)) ε ⟨
((b ℚ.+ b) * 1/ (b ℚ.+ b)) * ε
≡⟨ cong (_* ε) (*-inverseʳ (b ℚ.+ b)) ⟩
1ℚ * ε
≡⟨ *-identityˡ ε ⟩
ε ∎
where
open ≤-Reasoning

B : ℕ.ℕ
B = proj₁ (isCauchy x 1ℚ)

b : ℚ
b = 1ℚ ℚ.+ ∣ lookup (sequence x) B ∣

instance _ : Positive b
_ = pos+nonNeg⇒pos 1ℚ ∣ lookup (sequence x) B ∣ {{∣-∣-nonNeg (lookup (sequence x) B)}}

instance _ : NonZero b
_ = pos⇒nonZero b

instance _ : Positive (b ℚ.+ b)
_ = pos+pos⇒pos b b

instance _ : NonZero (b ℚ.+ b)
_ = pos⇒nonZero (b ℚ.+ b)

instance _ : Positive (1/ (b ℚ.+ b) * ε)
_ = pos*pos⇒pos (1/ (b ℚ.+ b)) {{1/pos⇒pos (b ℚ.+ b)}} ε

b-prop : ∀ {n} → n ℕ.≥ B → ∣ lookup (sequence x) n ∣ < b
b-prop {n} n≥B = begin-strict
∣ lookup (sequence x) n ∣
≡⟨ cong ∣_∣ (+-identityʳ (lookup (sequence x) n)) ⟨
∣ lookup (sequence x) n ℚ.+ 0ℚ ∣
≡⟨ cong (λ # → ∣ lookup (sequence x) n ℚ.+ # ∣) (+-inverseˡ (lookup (sequence x) B)) ⟨
∣ lookup (sequence x) n ℚ.+ (ℚ.- lookup (sequence x) B ℚ.+ lookup (sequence x) B) ∣
≡⟨ cong ∣_∣ (+-assoc (lookup (sequence x) n) (ℚ.- lookup (sequence x) B) (lookup (sequence x) B)) ⟨
∣ (lookup (sequence x) n - lookup (sequence x) B) ℚ.+ lookup (sequence x) B ∣
≤⟨ ∣p+q∣≤∣p∣+∣q∣ (lookup (sequence x) n - lookup (sequence x) B) (lookup (sequence x) B) ⟩
∣ lookup (sequence x) n - lookup (sequence x) B ∣ ℚ.+ ∣ lookup (sequence x) B ∣
<⟨ +-monoˡ-< ∣ lookup (sequence x) B ∣ (proj₂ (isCauchy x 1ℚ) n≥B ℕ.≤-refl) ⟩
1ℚ ℚ.+ ∣ lookup (sequence x) B ∣ ≡⟨⟩
b ∎

lemma : ∀ a b → a * a - b * b ≡ (a ℚ.+ b) * (a - b)
lemma = solve 2 (λ a b → (a :* a :- b :* b) , ((a :+ b) :* (a :- b))) refl
where open +-*-Solver
Loading
Loading