mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-12 07:08:40 +00:00
Reals are a ring (#58)
This commit is contained in:
62
Fields/CauchyCompletion/Field.agda
Normal file
62
Fields/CauchyCompletion/Field.agda
Normal file
@@ -0,0 +1,62 @@
|
||||
{-# 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.Lemmas
|
||||
open import Rings.Order
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
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.Naturals
|
||||
|
||||
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 _<_} {tOrder : SetoidTotalOrder {_<_ = _<_} pOrder} {R : Ring S _+_ _*_} (order : OrderedRing R tOrder) (F : Field R) (charNot2 : Setoid._∼_ S ((Ring.1R R) + (Ring.1R R)) (Ring.0R R) → False) where
|
||||
|
||||
open Setoid S
|
||||
open SetoidTotalOrder tOrder
|
||||
open SetoidPartialOrder pOrder
|
||||
open Equivalence eq
|
||||
open OrderedRing order
|
||||
open Field F
|
||||
open Group (Ring.additiveGroup R)
|
||||
|
||||
open import Rings.Orders.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
|
||||
|
||||
Cnontrivial : (pr : Setoid._∼_ cauchyCompletionSetoid (injection (Ring.0R R)) (injection (Ring.1R R))) → False
|
||||
Cnontrivial pr with pr (Ring.1R R) (0<1 (charNot2ImpliesNontrivial charNot2))
|
||||
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)
|
||||
|
||||
boundedMap : A → A
|
||||
boundedMap a with SetoidTotalOrder.totality tOrder 0G a
|
||||
boundedMap a | inl (inl x) = underlying (allInvertible a λ pr → irreflexive (<WellDefined (Equivalence.reflexive eq) pr x))
|
||||
boundedMap a | inl (inr x) = underlying (allInvertible a λ pr → irreflexive (<WellDefined pr (Equivalence.reflexive eq) x))
|
||||
boundedMap a | inr x = Ring.1R R
|
||||
|
||||
aNonzeroImpliesBounded : (a : CauchyCompletion) → (Setoid._∼_ cauchyCompletionSetoid a (injection 0G) → False) → (a <C (injection 0G)) || (injection 0G) <C a
|
||||
aNonzeroImpliesBounded a a!=0 = ?
|
||||
|
||||
1/aConverges : (a : CauchyCompletion) → (Setoid._∼_ cauchyCompletionSetoid a (injection 0G) → False) → cauchy (map boundedMap (CauchyCompletion.elts a))
|
||||
1/aConverges a a!=0 e 0<e = {!!}
|
||||
|
||||
1/a : (a : CauchyCompletion) → (Setoid._∼_ cauchyCompletionSetoid a (injection 0G) → False) → CauchyCompletion
|
||||
1/a a a!=0 = record { elts = map boundedMap (CauchyCompletion.elts a) ; converges = 1/aConverges a a!=0 }
|
||||
|
||||
1/a*a=1 : (a : CauchyCompletion) (pr : Setoid._∼_ cauchyCompletionSetoid a (injection 0G) → False) → Setoid._∼_ cauchyCompletionSetoid ((1/a a pr) *C a) (injection (Ring.1R R))
|
||||
1/a*a=1 a a!=0 e 0<e = {!!}
|
||||
|
||||
CField : Field CRing
|
||||
Field.allInvertible CField a a!=0 = (1/a a a!=0) , 1/a*a=1 a a!=0
|
||||
Field.nontrivial CField = Cnontrivial
|
@@ -1,4 +1,4 @@
|
||||
{-# OPTIONS --allow-unsolved-metas --warning=error --without-K --guardedness #-}
|
||||
{-# OPTIONS --safe --warning=error --without-K --guardedness #-}
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Setoids.Setoids
|
||||
@@ -44,6 +44,12 @@ c*Ident {a} ε 0<e = 0 , ans
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (injection (Ring.1R R) *C a)) (map inverse (CauchyCompletion.elts a))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (injection (Ring.1R R) *C a)) (map inverse (CauchyCompletion.elts a)) _+_ {m} | indexAndApply (constSequence (Ring.1R R)) (CauchyCompletion.elts a) _*_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts a) inverse m) | indexAndConst (Ring.1R R) m = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (transferToRight'' (Ring.additiveGroup R) (Ring.identIsIdent R))) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) (absZero order)))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
*CDistribute : {a b c : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a *C (b +C c)) ((a *C b) +C (a *C c))
|
||||
*CDistribute {a} {b} {c} e 0<e = 0 , ans
|
||||
where
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (a *C (b +C c))) (map inverse (CauchyCompletion.elts ((a *C b) +C (a *C c))))) m) < e
|
||||
ans {m} N<m rewrite indexAndApply (CauchyCompletion.elts (a *C (b +C c))) (map inverse (CauchyCompletion.elts ((a *C b) +C (a *C c)))) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (apply _+_ (CauchyCompletion.elts b) (CauchyCompletion.elts c)) _*_ {m} | equalityCommutative (mapAndIndex (apply _+_ (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts b)) (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts c))) inverse m) | indexAndApply (CauchyCompletion.elts b) (CauchyCompletion.elts c) _+_ {m} | indexAndApply (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts b)) (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts c)) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts b) _*_ {m} | indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts c) _*_ {m} = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (transferToRight'' (Ring.additiveGroup R) (Ring.*DistributesOver+ R))) (absZeroIsZero))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
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
|
||||
@@ -51,5 +57,5 @@ Ring.1R CRing = injection (Ring.1R R)
|
||||
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 = {!!}
|
||||
Ring.*DistributesOver+ CRing {a} {b} {c} = *CDistribute {a} {b} {c}
|
||||
Ring.identIsIdent CRing {a} = c*Ident {a}
|
||||
|
Reference in New Issue
Block a user