Elevate the real numbers to actually existing (#65)

This commit is contained in:
Patrick Stevens
2019-11-04 07:34:59 +00:00
committed by GitHub
parent d95f510cdd
commit 82b81b4bab
5 changed files with 55 additions and 49 deletions

View File

@@ -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

View File

@@ -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

View File

@@ -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 ()

View File

@@ -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 = {!!}

View 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