mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-13 15:48:39 +00:00
Definitions for the reals, classically (#95)
This commit is contained in:
76
Rings/InitialRing.agda
Normal file
76
Rings/InitialRing.agda
Normal file
@@ -0,0 +1,76 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Subset
|
||||
open import Setoids.Setoids
|
||||
open import Fields.Fields
|
||||
open import Rings.Definition
|
||||
open import Groups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Semirings.Definition
|
||||
|
||||
open import Numbers.Naturals.Semiring
|
||||
open import Numbers.Integers.Definition
|
||||
open import Numbers.Integers.RingStructure.Ring
|
||||
|
||||
module Rings.InitialRing {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A → A → A} (R : Ring S _+_ _*_) where
|
||||
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open import Rings.Lemmas R
|
||||
open import Groups.Lemmas additiveGroup
|
||||
|
||||
fromN : ℕ → A
|
||||
fromN zero = 0R
|
||||
fromN (succ n) = 1R + fromN n
|
||||
|
||||
fromNPreserves+ : (a b : ℕ) → fromN (a +N b) ∼ (fromN a) + (fromN b)
|
||||
fromNPreserves+ zero b = symmetric identLeft
|
||||
fromNPreserves+ (succ a) b = transitive (+WellDefined reflexive (fromNPreserves+ a b)) +Associative
|
||||
|
||||
fromNPreserves* : (a b : ℕ) → fromN (a *N b) ∼ (fromN a) * (fromN b)
|
||||
fromNPreserves* zero b = symmetric (transitive *Commutative timesZero)
|
||||
fromNPreserves* (succ a) b = transitive (transitive (fromNPreserves+ b (a *N b)) (+WellDefined (symmetric identIsIdent) (fromNPreserves* a b))) (symmetric *DistributesOver+')
|
||||
|
||||
fromNWellDefined : {a b : ℕ} → a ≡ b → fromN a ∼ fromN b
|
||||
fromNWellDefined refl = reflexive
|
||||
|
||||
fromZ : ℤ → A
|
||||
fromZ (nonneg x) = fromN x
|
||||
fromZ (negSucc x) = inverse (fromN (succ x))
|
||||
|
||||
private
|
||||
lemma : (a : ℕ) → (b : ℕ) → (1R + fromN (succ a *N b +N a)) ∼ ((1R + fromN a) * (1R + fromN b))
|
||||
lemma a b = transitive (transitive (transitive (+WellDefined reflexive (transitive (fromNPreserves+ ((succ a) *N b) a) (transitive groupIsAbelian (+WellDefined reflexive (transitive (fromNPreserves+ b (a *N b)) (+WellDefined (symmetric identIsIdent) (fromNPreserves* a b))))))) +Associative) (+WellDefined (transitive (symmetric identIsIdent) *Commutative) (symmetric *DistributesOver+'))) (symmetric *DistributesOver+)
|
||||
|
||||
fromZPreserves+ : (a b : ℤ) → fromZ (a +Z b) ∼ (fromZ a) + (fromZ b)
|
||||
fromZPreserves+ (nonneg zero) (nonneg y) = symmetric identLeft
|
||||
fromZPreserves+ (nonneg (succ x)) (nonneg y) = transitive (+WellDefined reflexive (fromNPreserves+ x y)) +Associative
|
||||
fromZPreserves+ (nonneg zero) (negSucc y) = symmetric identLeft
|
||||
fromZPreserves+ (nonneg (succ x)) (negSucc zero) = transitive (transitive (transitive (transitive (symmetric identLeft) (+WellDefined (symmetric invLeft) reflexive)) (symmetric +Associative)) (+WellDefined (inverseWellDefined (symmetric identRight)) reflexive)) (groupIsAbelian)
|
||||
fromZPreserves+ (nonneg (succ x)) (negSucc (succ y)) = transitive (fromZPreserves+ (nonneg x) (negSucc y)) (transitive (transitive (transitive (transitive (transitive (transitive (transitive (symmetric identLeft) +Associative) groupIsAbelian) (+WellDefined reflexive (+WellDefined (symmetric invLeft) reflexive))) (+WellDefined reflexive (symmetric +Associative))) +Associative) groupIsAbelian) (+WellDefined reflexive (symmetric invContravariant)))
|
||||
fromZPreserves+ (negSucc zero) (nonneg zero) = symmetric identRight
|
||||
fromZPreserves+ (negSucc zero) (nonneg (succ y)) = transitive (transitive (symmetric identLeft) (+WellDefined (symmetric (transitive (+WellDefined reflexive (symmetric identRight)) invLeft)) reflexive)) (symmetric +Associative)
|
||||
fromZPreserves+ (negSucc (succ x)) (nonneg zero) = symmetric identRight
|
||||
fromZPreserves+ (negSucc (succ x)) (nonneg (succ y)) = transitive (fromZPreserves+ (negSucc x) (nonneg y)) (transitive (transitive (+WellDefined reflexive (transitive (symmetric identLeft) (transitive (+WellDefined (symmetric invLeft) reflexive) (symmetric +Associative)))) +Associative) (+WellDefined (symmetric invContravariant) reflexive))
|
||||
fromZPreserves+ (negSucc x) (negSucc y) = transitive (transitive invContravariant (transitive (+WellDefined invContravariant reflexive) (transitive (+WellDefined (transitive (transitive (+WellDefined (transitive (inverseWellDefined (fromNPreserves+ x y)) invContravariant) reflexive) (symmetric +Associative)) groupIsAbelian) reflexive) (symmetric +Associative)))) (+WellDefined (symmetric invContravariant) (symmetric invContravariant))
|
||||
|
||||
fromZPreserves* : (a b : ℤ) → fromZ (a *Z b) ∼ (fromZ a) * (fromZ b)
|
||||
fromZPreserves* (nonneg a) (nonneg b) = fromNPreserves* a b
|
||||
fromZPreserves* (nonneg zero) (negSucc b) = symmetric (transitive *Commutative timesZero)
|
||||
fromZPreserves* (nonneg (succ a)) (negSucc b) = transitive (inverseWellDefined (lemma a b)) (symmetric ringMinusExtracts)
|
||||
fromZPreserves* (negSucc a) (nonneg zero) = symmetric timesZero
|
||||
fromZPreserves* (negSucc a) (nonneg (succ b)) = transitive (inverseWellDefined (transitive (+WellDefined reflexive (fromNWellDefined {(a +N b *N a) +N b} {(b +N a *N b) +N a} (transitivity (Semiring.commutative ℕSemiring _ b) (transitivity (applyEquality (b +N_) (transitivity (Semiring.commutative ℕSemiring a _) (applyEquality (_+N a) (multiplicationNIsCommutative b a)))) (Semiring.+Associative ℕSemiring b _ a))))) (lemma a b))) (symmetric ringMinusExtracts')
|
||||
fromZPreserves* (negSucc a) (negSucc b) = transitive (transitive (+WellDefined reflexive (fromNWellDefined {b +N a *N succ b} {(b +N a *N b) +N a} (transitivity (applyEquality (b +N_) (transitivity (multiplicationNIsCommutative a (succ b)) (transitivity (Semiring.commutative ℕSemiring a _) (applyEquality (_+N a) (multiplicationNIsCommutative b a))))) (Semiring.+Associative ℕSemiring b _ a)))) (lemma a b)) (symmetric twoNegativesTimes)
|
||||
|
||||
ringHom : RingHom ℤRing R fromZ
|
||||
RingHom.preserves1 ringHom = identRight
|
||||
RingHom.ringHom ringHom {r} {s} = fromZPreserves* r s
|
||||
GroupHom.groupHom (RingHom.groupHom ringHom) {r} {s} = fromZPreserves+ r s
|
||||
GroupHom.wellDefined (RingHom.groupHom ringHom) refl = reflexive
|
@@ -16,7 +16,7 @@ abstract
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
|
||||
ringMinusExtracts : {x y : A} → Setoid._∼_ S (x * Group.inverse (Ring.additiveGroup R) y) (Group.inverse (Ring.additiveGroup R) (x * y))
|
||||
ringMinusExtracts : {x y : A} → (x * Group.inverse (Ring.additiveGroup R) y) ∼ (Group.inverse (Ring.additiveGroup R) (x * y))
|
||||
ringMinusExtracts {x = x} {y} = transferToRight' additiveGroup (transitive (symmetric *DistributesOver+) (transitive (*WellDefined reflexive invLeft) (Ring.timesZero R)))
|
||||
where
|
||||
open Equivalence eq
|
||||
|
@@ -20,6 +20,7 @@ abstract
|
||||
open SetoidPartialOrder pOrder
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open Equivalence eq
|
||||
|
||||
open import Rings.Lemmas R
|
||||
|
||||
@@ -29,25 +30,23 @@ abstract
|
||||
ringCanMultiplyByPositive : {x y c : A} → (Ring.0R R) < c → x < y → (x * c) < (y * c)
|
||||
ringCanMultiplyByPositive {x} {y} {c} 0<c x<y = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.identRight additiveGroup) q'
|
||||
where
|
||||
open Equivalence eq
|
||||
have : 0R < (y + Group.inverse additiveGroup x)
|
||||
have = SetoidPartialOrder.<WellDefined pOrder (Group.invRight additiveGroup) reflexive (orderRespectsAddition x<y (Group.inverse additiveGroup x))
|
||||
p1 : 0R < ((y * c) + ((Group.inverse additiveGroup x) * c))
|
||||
p1 = SetoidPartialOrder.<WellDefined pOrder reflexive (Equivalence.transitive eq *Commutative (Equivalence.transitive eq *DistributesOver+ ((Group.+WellDefined additiveGroup) *Commutative *Commutative))) (orderRespectsMultiplication have 0<c)
|
||||
p1 = SetoidPartialOrder.<WellDefined pOrder reflexive (transitive *Commutative (transitive *DistributesOver+ ((Group.+WellDefined additiveGroup) *Commutative *Commutative))) (orderRespectsMultiplication have 0<c)
|
||||
p' : 0R < ((y * c) + (Group.inverse additiveGroup (x * c)))
|
||||
p' = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.+WellDefined additiveGroup reflexive (Equivalence.transitive eq (Equivalence.transitive eq *Commutative ringMinusExtracts) (inverseWellDefined additiveGroup *Commutative))) p1
|
||||
p' = SetoidPartialOrder.<WellDefined pOrder reflexive (Group.+WellDefined additiveGroup reflexive (transitive (transitive *Commutative ringMinusExtracts) (inverseWellDefined additiveGroup *Commutative))) p1
|
||||
q : (0R + (x * c)) < (((y * c) + (Group.inverse additiveGroup (x * c))) + (x * c))
|
||||
q = orderRespectsAddition p' (x * c)
|
||||
q' : (x * c) < ((y * c) + 0R)
|
||||
q' = SetoidPartialOrder.<WellDefined pOrder (Group.identLeft additiveGroup) (Equivalence.transitive eq (symmetric (Group.+Associative additiveGroup)) (Group.+WellDefined additiveGroup reflexive (Group.invLeft additiveGroup))) q
|
||||
q' = SetoidPartialOrder.<WellDefined pOrder (Group.identLeft additiveGroup) (transitive (symmetric (Group.+Associative additiveGroup)) (Group.+WellDefined additiveGroup reflexive (Group.invLeft additiveGroup))) q
|
||||
|
||||
ringMultiplyPositives : {x y a b : A} → 0R < x → 0R < a → (x < y) → (a < b) → (x * a) < (y * b)
|
||||
ringMultiplyPositives {x} {y} {a} {b} 0<x 0<a x<y a<b = SetoidPartialOrder.<Transitive pOrder (ringCanMultiplyByPositive 0<a x<y) (<WellDefined *Commutative *Commutative (ringCanMultiplyByPositive (SetoidPartialOrder.<Transitive pOrder 0<x x<y) a<b))
|
||||
|
||||
ringSwapNegatives : {x y : A} → (Group.inverse (Ring.additiveGroup R) x) < (Group.inverse (Ring.additiveGroup R) y) → y < x
|
||||
ringSwapNegatives {x} {y} -x<-y = SetoidPartialOrder.<WellDefined pOrder (Equivalence.transitive eq (symmetric (Group.+Associative additiveGroup)) (Equivalence.transitive eq (Group.+WellDefined additiveGroup reflexive (Group.invLeft additiveGroup)) (Group.identRight additiveGroup))) (Group.identLeft additiveGroup) v
|
||||
ringSwapNegatives {x} {y} -x<-y = SetoidPartialOrder.<WellDefined pOrder (transitive (symmetric (Group.+Associative additiveGroup)) (transitive (Group.+WellDefined additiveGroup reflexive (Group.invLeft additiveGroup)) (Group.identRight additiveGroup))) (Group.identLeft additiveGroup) v
|
||||
where
|
||||
open Equivalence eq
|
||||
t : ((Group.inverse additiveGroup x) + y) < ((Group.inverse additiveGroup y) + y)
|
||||
t = orderRespectsAddition -x<-y y
|
||||
u : (y + (Group.inverse additiveGroup x)) < 0R
|
||||
@@ -73,3 +72,6 @@ abstract
|
||||
|
||||
anyComparisonImpliesNontrivial : {a b : A} → a < b → (0R ∼ 1R) → False
|
||||
anyComparisonImpliesNontrivial {a} {b} a<b 0=1 = irreflexive (<WellDefined (oneZeroImpliesAllZero 0=1) (oneZeroImpliesAllZero 0=1) a<b)
|
||||
|
||||
moveInequality : {a b : A} → a < b → 0R < (b + inverse a)
|
||||
moveInequality {a} {b} a<b = <WellDefined invRight reflexive (orderRespectsAddition a<b (inverse a))
|
||||
|
@@ -344,3 +344,6 @@ abstract
|
||||
absBoundedImpliesBounded {a} {b} a<b | inl (inl x) = a<b
|
||||
absBoundedImpliesBounded {a} {b} a<b | inl (inr x) = SetoidPartialOrder.<Transitive pOrder x (SetoidPartialOrder.<Transitive pOrder (lemm2 a x) a<b)
|
||||
absBoundedImpliesBounded {a} {b} a<b | inr x = a<b
|
||||
|
||||
orderedImpliesCharNot2 : (0R ∼ 1R → False) → 1R + 1R ∼ 0R → False
|
||||
orderedImpliesCharNot2 0!=1 x = irreflexive (<WellDefined (identRight {0R}) x (ringAddInequalities (0<1 0!=1) (0<1 0!=1)))
|
||||
|
Reference in New Issue
Block a user