mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-08 13:28:39 +00:00
Elevate the real numbers to actually existing (#65)
This commit is contained in:
@@ -1,4 +1,4 @@
|
||||
{-# OPTIONS --warning=error --safe #-}
|
||||
{-# OPTIONS --warning=error --safe --guardedness #-}
|
||||
|
||||
-- This file contains everything that is --safe, but uses K.
|
||||
|
||||
@@ -7,6 +7,8 @@ open import Numbers.Primes.IntegerFactorisation
|
||||
open import Numbers.Rationals.Definition
|
||||
open import Numbers.Rationals.Lemmas
|
||||
|
||||
open import Numbers.Reals.Definition
|
||||
|
||||
open import Logic.PropositionalLogic
|
||||
open import Logic.PropositionalLogicExamples
|
||||
open import Logic.PropositionalAxiomsTautology
|
||||
|
@@ -24,15 +24,16 @@ open import Fields.FieldOfFractions.Ring I
|
||||
fieldOfFractions : Field fieldOfFractionsRing
|
||||
Field.allInvertible fieldOfFractions (fst ,, (b , _)) prA = (b ,, (fst , ans)) , need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : ((b * fst) * Ring.1R R) ∼ ((fst * b) * Ring.1R R)
|
||||
need = Ring.*WellDefined R (Ring.*Commutative R) reflexive
|
||||
ans : fst ∼ Ring.0R R → False
|
||||
ans pr = prA need'
|
||||
where
|
||||
need' : (fst * Ring.1R R) ∼ (b * Ring.0R R)
|
||||
need' = transitive (Ring.*WellDefined R pr reflexive) (transitive (transitive (Ring.*Commutative R) (Ring.timesZero R)) (symmetric (Ring.timesZero R)))
|
||||
abstract
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : ((b * fst) * Ring.1R R) ∼ ((fst * b) * Ring.1R R)
|
||||
need = Ring.*WellDefined R (Ring.*Commutative R) reflexive
|
||||
ans : fst ∼ Ring.0R R → False
|
||||
ans pr = prA need'
|
||||
where
|
||||
need' : (fst * Ring.1R R) ∼ (b * Ring.0R R)
|
||||
need' = transitive (Ring.*WellDefined R pr reflexive) (transitive (transitive (Ring.*Commutative R) (Ring.timesZero R)) (symmetric (Ring.timesZero R)))
|
||||
Field.nontrivial fieldOfFractions pr = IntegralDomain.nontrivial I (symmetric (transitive (symmetric (Ring.timesZero R)) (transitive (Ring.*Commutative R) (transitive pr (Ring.identIsIdent R)))))
|
||||
where
|
||||
open Setoid S
|
||||
|
@@ -85,3 +85,6 @@ negateQWellDefined a b a=b = inverseWellDefined (Ring.additiveGroup ℚRing) {a}
|
||||
|
||||
ℚOrdered : TotallyOrderedRing ℚPOrdered
|
||||
ℚOrdered = fieldOfFractionsOrderedRing
|
||||
|
||||
ℚcharNot2 : ((Ring.1R ℚRing) +Q (Ring.1R ℚRing)) =Q (Ring.0R ℚRing) → False
|
||||
ℚcharNot2 ()
|
||||
|
@@ -1,39 +0,0 @@
|
||||
{-# OPTIONS --safe --warning=error #-}
|
||||
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Orders
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Rationals
|
||||
open import Numbers.RationalsLemmas
|
||||
open import Numbers.Naturals
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Numbers.Reals where
|
||||
record ℝ : Set where
|
||||
field
|
||||
f : ℕ → ℚ
|
||||
converges : {ε : ℚ} → Sg ℕ (λ x → {y : ℕ} → x <N y → (absQ ((f x) -Q (f y))) <Q ε)
|
||||
|
||||
_+R_ : ℝ → ℝ → ℝ
|
||||
ℝ.f (record { f = a ; converges = conA } +R record { f = b ; converges = conB }) n = (a n) +Q (b n)
|
||||
ℝ.converges (record { f = a ; converges = conA } +R record { f = b ; converges = conB }) {e} = {!absQ (a x +Q b x)!}
|
||||
|
||||
negateR : ℝ → ℝ
|
||||
ℝ.f (negateR record { f = f ; converges = converges }) n = negateQ (f n)
|
||||
ℝ.converges (negateR record { f = f ; converges = converges }) {e} with converges {e}
|
||||
... | n , pr = n , λ {y} x<y → SetoidPartialOrder.wellDefined ℚPartialOrder {absQ ((f n) -Q (f y))} {absQ ((negateQ (f n)) -Q (negateQ (f y)))} {e} {e} {!!} {!reflQ e!} (pr {y} x<y)
|
||||
|
||||
_-R_ : ℝ → ℝ → ℝ
|
||||
a -R b = a +R (negateR b)
|
||||
|
||||
_*R_ : ℝ → ℝ → ℝ
|
||||
ℝ.f (record { f = a ; converges = conA } *R record { f = b ; converges = conB }) n = (a n) *Q (b n)
|
||||
ℝ.converges (record { f = a ; converges = conA } *R record { f = b ; converges = conB}) {e} = {!!}
|
||||
|
||||
realsSetoid : Setoid ℝ
|
||||
(realsSetoid Setoid.∼ record { f = a ; converges = convA }) record { f = b ; converges = convB } = {!!}
|
||||
Setoid.eq realsSetoid = {!!}
|
39
Numbers/Reals/Definition.agda
Normal file
39
Numbers/Reals/Definition.agda
Normal file
@@ -0,0 +1,39 @@
|
||||
{-# OPTIONS --warning=error --safe --guardedness #-}
|
||||
|
||||
open import Setoids.Orders
|
||||
open import LogicalFormulae
|
||||
open import Rings.Definition
|
||||
open import Numbers.Rationals.Definition
|
||||
|
||||
module Numbers.Reals.Definition where
|
||||
|
||||
open import Fields.CauchyCompletion.Definition ℚOrdered ℚField
|
||||
open import Fields.CauchyCompletion.Setoid ℚOrdered ℚField ℚcharNot2
|
||||
open import Fields.CauchyCompletion.Addition ℚOrdered ℚField ℚcharNot2
|
||||
open import Fields.CauchyCompletion.Multiplication ℚOrdered ℚField ℚcharNot2
|
||||
open import Fields.CauchyCompletion.Ring ℚOrdered ℚField ℚcharNot2
|
||||
open import Fields.CauchyCompletion.Comparison ℚOrdered ℚField ℚcharNot2
|
||||
|
||||
ℝ : Set
|
||||
ℝ = CauchyCompletion
|
||||
|
||||
_+R_ : ℝ → ℝ → ℝ
|
||||
_+R_ = _+C_
|
||||
|
||||
_*R_ : ℝ → ℝ → ℝ
|
||||
_*R_ = _*C_
|
||||
|
||||
ℝRing : Ring cauchyCompletionSetoid _+R_ _*R_
|
||||
ℝRing = CRing
|
||||
|
||||
injectionR : ℚ → ℝ
|
||||
injectionR = injection
|
||||
|
||||
0R : ℝ
|
||||
0R = injection 0Q
|
||||
|
||||
_<R_ : ℝ → ℝ → Set
|
||||
_<R_ = _<C_
|
||||
|
||||
ℝPartialOrder : SetoidPartialOrder cauchyCompletionSetoid _<C_
|
||||
ℝPartialOrder = <COrder
|
Reference in New Issue
Block a user