mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-08 13:28:39 +00:00
Move towards base-n expansions (#112)
This commit is contained in:
35
Rings/Orders/Partial/Bounded.agda
Normal file
35
Rings/Orders/Partial/Bounded.agda
Normal file
@@ -0,0 +1,35 @@
|
||||
{-# 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.Orders.Partial.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Groups.Definition
|
||||
|
||||
module Rings.Orders.Partial.Bounded {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} (pRing : PartiallyOrderedRing R pOrder) where
|
||||
|
||||
open Group (Ring.additiveGroup R)
|
||||
open import Groups.Lemmas (Ring.additiveGroup R)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open SetoidPartialOrder pOrder
|
||||
|
||||
BoundedAbove : Sequence A → Set (m ⊔ o)
|
||||
BoundedAbove x = Sg A (λ K → (n : ℕ) → index x n < K)
|
||||
|
||||
BoundedBelow : Sequence A → Set (m ⊔ o)
|
||||
BoundedBelow x = Sg A (λ K → (n : ℕ) → K < index x n)
|
||||
|
||||
Bounded : Sequence A → Set (m ⊔ o)
|
||||
Bounded x = Sg A (λ K → (n : ℕ) → ((Group.inverse (Ring.additiveGroup R) K) < index x n) && (index x n < K))
|
||||
|
||||
boundNonzero : {s : Sequence A} → (b : Bounded s) → underlying b ∼ 0G → False
|
||||
boundNonzero {s} (a , b) isEq with b 0
|
||||
... | bad1 ,, bad2 = irreflexive (<Transitive bad1 (<WellDefined reflexive (transitive isEq (symmetric (transitive (inverseWellDefined isEq) invIdent))) bad2))
|
31
Rings/Orders/Total/Bounded.agda
Normal file
31
Rings/Orders/Total/Bounded.agda
Normal file
@@ -0,0 +1,31 @@
|
||||
{-# 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.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Groups.Definition
|
||||
|
||||
module Rings.Orders.Total.Bounded {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (tRing : TotallyOrderedRing pRing) where
|
||||
|
||||
open import Rings.Orders.Partial.Bounded pRing
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open import Groups.Lemmas (Ring.additiveGroup R)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open SetoidPartialOrder pOrder
|
||||
open import Rings.Orders.Total.Lemmas tRing
|
||||
open PartiallyOrderedRing pRing
|
||||
|
||||
boundGreaterThanZero : {s : Sequence A} → (b : Bounded s) → 0G < underlying b
|
||||
boundGreaterThanZero {s} (a , b) with b 0
|
||||
... | (l ,, r) = halvePositive a (<WellDefined invLeft reflexive (orderRespectsAddition (<Transitive l r) a))
|
30
Rings/Orders/Total/Cauchy.agda
Normal file
30
Rings/Orders/Total/Cauchy.agda
Normal file
@@ -0,0 +1,30 @@
|
||||
{-# 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.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Definition
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Sequences
|
||||
open import Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Naturals.Order
|
||||
|
||||
module Rings.Orders.Total.Cauchy {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {m} {o} A} {pOrder : SetoidPartialOrder S _<_} {R : Ring S _+_ _*_} {pRing : PartiallyOrderedRing R pOrder} (order : TotallyOrderedRing pRing) where
|
||||
|
||||
open Setoid S
|
||||
open SetoidTotalOrder (TotallyOrderedRing.total order)
|
||||
open SetoidPartialOrder pOrder
|
||||
open Equivalence eq
|
||||
open TotallyOrderedRing order
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
|
||||
open import Rings.Orders.Total.Lemmas order
|
||||
|
||||
cauchy : Sequence A → Set (m ⊔ o)
|
||||
cauchy s = ∀ (ε : A) → (0R < ε) → Sg ℕ (λ N → ∀ {m n : ℕ} → (N <N m) → (N <N n) → abs ((index s m) -R (index s n)) < ε)
|
@@ -363,3 +363,15 @@ fromNPreservesOrder : (0R ∼ 1R → False) → {a b : ℕ} → (a <N b) → (fr
|
||||
fromNPreservesOrder 0!=1 {zero} {succ zero} a<b = fromNIncreasing 0!=1 0
|
||||
fromNPreservesOrder 0!=1 {zero} {succ (succ b)} a<b = <Transitive (fromNPreservesOrder 0!=1 (succIsPositive b)) (fromNIncreasing 0!=1 (succ b))
|
||||
fromNPreservesOrder 0!=1 {succ a} {succ b} a<b = <WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (fromNPreservesOrder 0!=1 (canRemoveSuccFrom<N a<b)) 1R)
|
||||
|
||||
reciprocalPositive : (a b : A) → .(0<a : 0R < a) → (a * b ∼ 1R) → 0R < b
|
||||
reciprocalPositive a 1/a 0<a ab=1 with totality 0G 1/a
|
||||
... | inl (inl x) = x
|
||||
... | inl (inr x) = exFalso (1<0False (<WellDefined (transitive *Commutative ab=1) timesZero' (ringCanMultiplyByPositive 0<a x)))
|
||||
... | inr x = exFalso (anyComparisonImpliesNontrivial 0<a (transitive (transitive (symmetric timesZero) (*WellDefined reflexive x)) ab=1))
|
||||
|
||||
reciprocal<1 : (a b : A) → .(1<a : 1R < a) → (a * b ∼ 1R) → b < 1R
|
||||
reciprocal<1 a b 0<a ab=1 with totality b 1R
|
||||
... | inl (inl x) = x
|
||||
... | inr b=1 = exFalso (irreflexive (<WellDefined (symmetric ab=1) (transitive (symmetric identIsIdent) (transitive *Commutative ((*WellDefined reflexive (symmetric b=1))))) 0<a))
|
||||
... | inl (inr x) = exFalso (irreflexive (<WellDefined identIsIdent ab=1 (ringMultiplyPositives (0<1 (anyComparisonImpliesNontrivial 0<a)) (0<1 (anyComparisonImpliesNontrivial 0<a)) 0<a x)))
|
||||
|
Reference in New Issue
Block a user