mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-06 12:28:39 +00:00
Comparison on the reals (#55)
This commit is contained in:
@@ -19,6 +19,8 @@ open import Groups.Groups2
|
||||
open import Groups.SymmetryGroups
|
||||
|
||||
open import Fields.Fields
|
||||
open import Fields.Orders.Definition
|
||||
open import Fields.Orders.Lemmas
|
||||
open import Fields.FieldOfFractions
|
||||
open import Fields.FieldOfFractionsOrder
|
||||
|
||||
@@ -53,7 +55,7 @@ open import Monoids.Definition
|
||||
open import Semirings.Definition
|
||||
open import Semirings.Solver
|
||||
|
||||
open import Fields.CauchyCompletion.Definition
|
||||
open import Fields.CauchyCompletion.Addition
|
||||
open import Fields.CauchyCompletion.Group
|
||||
open import Fields.CauchyCompletion.Comparison
|
||||
|
||||
module Everything.Safe where
|
||||
|
@@ -26,30 +26,17 @@ open Ring R
|
||||
open Group additiveGroup
|
||||
open Field F
|
||||
|
||||
open import Fields.Lemmas F
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Rings.Orders.Lemmas(order)
|
||||
|
||||
halve : (a : A) → Sg A (λ i → i + i ∼ a)
|
||||
-- TODO: we need to know the characteristic already
|
||||
halve a with allInvertible (1R + 1R) charNot2
|
||||
... | 1/2 , pr1/2 = (a * 1/2) , Equivalence.transitive eq (+WellDefined *Commutative *Commutative) (Equivalence.transitive eq (Equivalence.symmetric eq (*DistributesOver+ {1/2} {a} {a})) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) r) (Equivalence.transitive eq (*Associative) (Equivalence.transitive eq (*WellDefined pr1/2 (Equivalence.reflexive eq)) identIsIdent))))
|
||||
where
|
||||
r : a + a ∼ (1R + 1R) * a
|
||||
r = Equivalence.symmetric eq (Equivalence.transitive eq *Commutative (Equivalence.transitive eq *DistributesOver+ (+WellDefined (Equivalence.transitive eq *Commutative identIsIdent) (Equivalence.transitive eq *Commutative identIsIdent))))
|
||||
|
||||
halvePositive : (a : A) → 0R < (a + a) → 0R < a
|
||||
halvePositive a 0<2a with totality 0R a
|
||||
halvePositive a 0<2a | inl (inl x) = x
|
||||
halvePositive a 0<2a | inl (inr a<0) = exFalso (irreflexive {a + a} (SetoidPartialOrder.transitive pOrder (<WellDefined (Equivalence.reflexive eq) identRight (ringAddInequalities a<0 a<0)) 0<2a))
|
||||
halvePositive a 0<2a | inr x = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq x) (Equivalence.symmetric eq x)) identRight) 0<2a))
|
||||
|
||||
lemm : (m : ℕ) (a b : Sequence A) → index (apply _+_ a b) m ≡ (index a m) + (index b m)
|
||||
lemm zero a b = refl
|
||||
lemm (succ m) a b = lemm m (Sequence.tail a) (Sequence.tail b)
|
||||
|
||||
_+C_ : CauchyCompletion → CauchyCompletion → CauchyCompletion
|
||||
CauchyCompletion.elts (record { elts = a ; converges = convA } +C record { elts = b ; converges = convB }) = apply _+_ a b
|
||||
CauchyCompletion.converges (record { elts = a ; converges = convA } +C record { elts = b ; converges = convB }) ε 0<e with halve ε
|
||||
CauchyCompletion.converges (record { elts = a ; converges = convA } +C record { elts = b ; converges = convB }) ε 0<e with halve charNot2 ε
|
||||
... | e/2 , e/2Pr with convA e/2 (halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq e/2Pr) 0<e))
|
||||
CauchyCompletion.converges (record { elts = a ; converges = convA } +C record { elts = b ; converges = convB }) ε 0<e | e/2 , e/2Pr | Na , prA with convB e/2 (halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq e/2Pr) 0<e))
|
||||
CauchyCompletion.converges (record { elts = a ; converges = convA } +C record { elts = b ; converges = convB }) ε 0<e | e/2 , e/2Pr | Na , prA | Nb , prB = (Na +N Nb) , t
|
||||
@@ -60,3 +47,14 @@ CauchyCompletion.converges (record { elts = a ; converges = convA } +C record {
|
||||
... | bm-bn<e/2 with triangleInequality (index a m + inverse (index a n)) (index b m + inverse (index b n))
|
||||
... | inl tri rewrite lemm m a b | lemm n a b = SetoidPartialOrder.<WellDefined pOrder (Equivalence.reflexive eq) e/2Pr (SetoidPartialOrder.transitive pOrder {_} {(abs ((index a m) + (inverse (index a n)))) + (abs ((index b m) + (inverse (index b n))))} (<WellDefined (absWellDefined _ _ (Equivalence.transitive eq (Equivalence.symmetric eq (+Associative {index a m})) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq {index a m}) (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq (+Associative {index b m})) (+WellDefined (Equivalence.reflexive eq {index b m}) (Equivalence.symmetric eq (invContravariant additiveGroup)))))) (+Associative {index a m})))) (Equivalence.reflexive eq) tri) (ringAddInequalities am-an<e/2 bm-bn<e/2))
|
||||
... | inr tri rewrite lemm m a b | lemm n a b = SetoidPartialOrder.<WellDefined pOrder (Equivalence.reflexive eq) e/2Pr (<WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq tri) (absWellDefined _ _ (Equivalence.transitive eq (Equivalence.symmetric eq (+Associative {index a m})) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq {index a m}) (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq (+Associative {index b m})) (+WellDefined (Equivalence.reflexive eq {index b m}) (Equivalence.symmetric eq (invContravariant additiveGroup)))))) (+Associative {index a m}))))) (Equivalence.reflexive eq) (ringAddInequalities am-an<e/2 bm-bn<e/2))
|
||||
|
||||
inverseDistributes : {r s : A} → inverse (r + s) ∼ inverse r + inverse s
|
||||
inverseDistributes = Equivalence.transitive eq (invContravariant additiveGroup) groupIsAbelian
|
||||
|
||||
-C_ : CauchyCompletion → CauchyCompletion
|
||||
CauchyCompletion.elts (-C a) = map inverse (CauchyCompletion.elts a)
|
||||
CauchyCompletion.converges (-C record { elts = elts ; converges = converges }) ε 0<e with converges ε 0<e
|
||||
CauchyCompletion.converges (-C record { elts = elts ; converges = converges }) ε 0<e | N , prN = N , ans
|
||||
where
|
||||
ans : {m n : ℕ} → (N <N m) → (N <N n) → abs ((index (map inverse elts) m) + inverse (index (map inverse elts) n)) < ε
|
||||
ans {m} {n} N<m N<n = <WellDefined (Equivalence.transitive eq (absWellDefined _ _ (Equivalence.reflexive eq)) (Equivalence.transitive eq (absNegation (index elts m + inverse (index elts n))) (absWellDefined _ _ (Equivalence.transitive eq inverseDistributes (+WellDefined (identityOfIndiscernablesLeft _∼_ (Equivalence.reflexive eq) (equalityCommutative (mapAndIndex elts inverse m))) (inverseWellDefined additiveGroup (identityOfIndiscernablesLeft _∼_ (Equivalence.reflexive eq) (equalityCommutative (mapAndIndex elts inverse n))))))))) (Equivalence.reflexive eq) (prN N<m N<n)
|
||||
|
35
Fields/CauchyCompletion/Approximation.agda
Normal file
35
Fields/CauchyCompletion/Approximation.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.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.Approximation {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 Ring R
|
||||
open Group additiveGroup
|
||||
open Field F
|
||||
|
||||
open import Rings.Orders.Lemmas(order)
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Addition order F charNot2
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
|
||||
approximate : (a : CauchyCompletion) → (ε : A) → Sg A (λ b → a +C (-C (injection b)) < ε)
|
||||
approximate a ε = ?
|
131
Fields/CauchyCompletion/Comparison.agda
Normal file
131
Fields/CauchyCompletion/Comparison.agda
Normal file
@@ -0,0 +1,131 @@
|
||||
{-# 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
|
||||
open import Semirings.Definition
|
||||
|
||||
module Fields.CauchyCompletion.Comparison {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 Ring R
|
||||
open Group additiveGroup
|
||||
open Field F
|
||||
open import Fields.Orders.Lemmas {F = F} record { oRing = order }
|
||||
|
||||
open import Rings.Orders.Lemmas order
|
||||
open import Fields.Lemmas F
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
|
||||
shrinkInequality : {a b c : ℕ} → a +N b <N c → a <N c
|
||||
shrinkInequality {a} {b} {c} (le x pr) = le (x +N b) (transitivity (applyEquality succ (transitivity (equalityCommutative (Semiring.+Associative ℕSemiring x _ _)) (applyEquality (x +N_) (Semiring.commutative ℕSemiring b a)))) pr)
|
||||
|
||||
shrinkInequality' : {a b c : ℕ} → a +N b <N c → b <N c
|
||||
shrinkInequality' {a} {b} {c} (le x pr) = le (x +N a) (transitivity (applyEquality succ (equalityCommutative (Semiring.+Associative ℕSemiring x _ _))) pr)
|
||||
|
||||
_<C'_ : CauchyCompletion → A → Set (m ⊔ o)
|
||||
a <C' b = Sg A (λ ε → (0G < ε) && Sg ℕ (λ N → ((m : ℕ) → (N<m : N <N m) → (ε + index (CauchyCompletion.elts a) m) < b)))
|
||||
|
||||
<C'WellDefinedRight : (a : CauchyCompletion) (b c : A) → b ∼ c → a <C' b → a <C' c
|
||||
<C'WellDefinedRight a b c b=c (ε , (0<e ,, (N , pr))) = ε , (0<e ,, (N , λ m N<m → <WellDefined (Equivalence.reflexive eq) b=c (pr m N<m)))
|
||||
|
||||
halfLess : (e/2 e : A) → (0<e : 0G < e) → (pr : e/2 + e/2 ∼ e) → e/2 < e
|
||||
halfLess e/2 e 0<e pr with halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq pr) 0<e)
|
||||
... | 0<e/2 = <WellDefined identLeft pr (orderRespectsAddition 0<e/2 e/2)
|
||||
|
||||
<C'WellDefinedLemma : (a b e/2 e : A) (0<e : 0G < e) (pr : e/2 + e/2 ∼ e) → abs (a + inverse b) < e/2 → (e/2 + b) < (e + a)
|
||||
<C'WellDefinedLemma a b e/2 e 0<e pr a-b<e with SetoidTotalOrder.totality tOrder 0G (a + inverse b)
|
||||
<C'WellDefinedLemma a b e/2 e 0<e pr a-b<e | inl (inl 0<a-b) = ringAddInequalities (halfLess e/2 e 0<e pr) (<WellDefined identLeft (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) invLeft) identRight)) (orderRespectsAddition 0<a-b b))
|
||||
<C'WellDefinedLemma a b e/2 e 0<e pr a-b<e | inl (inr a-b<0) = <WellDefined (Equivalence.transitive eq (+WellDefined (invContravariant (Ring.additiveGroup R)) groupIsAbelian) (Equivalence.transitive eq (Equivalence.transitive eq +Associative (+WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (invTwice additiveGroup b) invLeft) identRight)) (Equivalence.reflexive eq))) groupIsAbelian)) (Equivalence.transitive eq +Associative (+WellDefined pr (Equivalence.reflexive eq))) (orderRespectsAddition a-b<e (e/2 + a))
|
||||
<C'WellDefinedLemma a b e/2 e 0<e pr a-b<e | inr 0=a-b = <WellDefined (Equivalence.reflexive eq) (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (transferToRight (Ring.additiveGroup R) (Equivalence.symmetric eq 0=a-b)))) (orderRespectsAddition (halfLess e/2 e 0<e pr) b)
|
||||
|
||||
<C'WellDefinedLeft : (a b : CauchyCompletion) (c : A) → Setoid._∼_ cauchyCompletionSetoid a b → a <C' c → b <C' c
|
||||
<C'WellDefinedLeft a b c a=b (ε , (0<e ,, (N , pr))) with halve charNot2 ε
|
||||
... | e/2 , prE/2 with halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prE/2) 0<e)
|
||||
... | 0<e/2 with a=b e/2 0<e/2
|
||||
... | N2 , prN2 = e/2 , (0<e/2 ,, ((N +N N2) , ans))
|
||||
where
|
||||
ans : (m : ℕ) → (N<m : (N +N N2) <N m) → (e/2 + index (CauchyCompletion.elts b) m) < c
|
||||
ans m N<m with prN2 {m} (shrinkInequality' N<m)
|
||||
... | bl rewrite indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts b)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts b) inverse m) = SetoidPartialOrder.transitive pOrder (<C'WellDefinedLemma (index (CauchyCompletion.elts a) m) (index (CauchyCompletion.elts b) m) e/2 ε 0<e prE/2 bl) (pr m (shrinkInequality N<m))
|
||||
|
||||
<C'WellDefined : (a b : CauchyCompletion) (c d : A) → Setoid._∼_ cauchyCompletionSetoid a b → c ∼ d → a <C' c → b <C' d
|
||||
<C'WellDefined a b c d a=b c=d a<c = <C'WellDefinedLeft a b d a=b (<C'WellDefinedRight a c d c=d a<c)
|
||||
|
||||
_<C''_ : A → CauchyCompletion → Set (m ⊔ o)
|
||||
a <C'' b = Sg A (λ ε → (0G < ε) && Sg ℕ (λ N → ((m : ℕ) → (N<m : N <N m) → (a + ε) < index (CauchyCompletion.elts b) m)))
|
||||
|
||||
<C''WellDefinedLeft : (a b : A) (c : CauchyCompletion) → (a ∼ b) → (a <C'' c) → b <C'' c
|
||||
<C''WellDefinedLeft a b c a=b (e , (0<e ,, (N , pr))) = e , (0<e ,, (N , λ m N<m → <WellDefined (+WellDefined a=b (Equivalence.reflexive eq)) (Equivalence.reflexive eq) (pr m N<m)))
|
||||
|
||||
<C''WellDefinedLemma : (a b c e e/2 : A) (_ : e/2 + e/2 ∼ e) (0<e : 0G < e) (_ : abs (a + inverse b) < e/2) (_ : (c + e) < a) → (c + e/2) < b
|
||||
<C''WellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a with SetoidTotalOrder.totality tOrder 0G (a + inverse b)
|
||||
<C''WellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a | inl (inl 0<a-b) = SetoidPartialOrder.transitive pOrder (<WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq prE/2) (Equivalence.reflexive eq)) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) invRight)) identRight)))) (Equivalence.reflexive eq) (orderRespectsAddition c+e<a (inverse e/2))) (<WellDefined (Equivalence.transitive eq +Associative (+WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) invLeft) identRight)) (Equivalence.reflexive eq))) (Equivalence.transitive eq groupIsAbelian (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (invLeft)) identRight))) (orderRespectsAddition pr (b + inverse e/2)))
|
||||
<C''WellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a | inl (inr a-b<0) = SetoidPartialOrder.transitive pOrder (SetoidPartialOrder.transitive pOrder (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition (halfLess e/2 e 0<e prE/2) c)) c+e<a) (<WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) invLeft)) identRight) identLeft (orderRespectsAddition a-b<0 b))
|
||||
<C''WellDefinedLemma a b c e e/2 prE/2 0<e pr c+e<a | inr 0=a-b = SetoidPartialOrder.transitive pOrder (<WellDefined groupIsAbelian (Equivalence.reflexive eq) (orderRespectsAddition (halfLess e/2 e 0<e prE/2) c)) (<WellDefined (groupIsAbelian {c} {e}) (transferToRight additiveGroup (Equivalence.symmetric eq 0=a-b)) c+e<a)
|
||||
|
||||
<C''WellDefinedRight : (a b : CauchyCompletion) (c : A) → Setoid._∼_ cauchyCompletionSetoid a b → c <C'' a → c <C'' b
|
||||
<C''WellDefinedRight a b c a=b (e , (0<e ,, (N , pr))) with halve charNot2 e
|
||||
... | e/2 , prE/2 with halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prE/2) 0<e)
|
||||
... | 0<e/2 with a=b e/2 0<e/2
|
||||
... | N2 , prN2 = e/2 , (0<e/2 ,, ((N +N N2) , ans))
|
||||
where
|
||||
ans : (m : ℕ) → (N<m : (N +N N2) <N m) → (c + e/2) < index (CauchyCompletion.elts b) m
|
||||
ans m N<m with prN2 {m} (shrinkInequality' N<m)
|
||||
... | bl rewrite indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts b)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts b) inverse m) = <C''WellDefinedLemma (index (CauchyCompletion.elts a) m) (index (CauchyCompletion.elts b) m) c e e/2 prE/2 0<e bl (pr m (shrinkInequality N<m))
|
||||
|
||||
_<C_ : CauchyCompletion → CauchyCompletion → Set (m ⊔ o)
|
||||
a <C b = Sg A (λ i → (a <C' i) && (i <C'' b))
|
||||
|
||||
<CWellDefined : {a b c d : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid c d → a <C c → b <C d
|
||||
<CWellDefined {a} {b} {c} {d} a=b c=d (bound , (a<bound ,, bound<b)) = bound , (<C'WellDefinedLeft a b bound a=b a<bound ,, <C''WellDefinedRight c d bound c=d bound<b)
|
||||
|
||||
<C'Preserves : {a b : A} → a < b → injection a <C' b
|
||||
<C'Preserves {a} {b} a<b with dense charNot2 a<b
|
||||
... | between , (a<between ,, between<b) = (between + inverse a) , (<WellDefined invRight (Equivalence.reflexive eq) (orderRespectsAddition a<between (inverse a)) ,, (0 , λ m _ → <WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq identRight) (Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (ans a m)) +Associative)) (Equivalence.reflexive eq) between<b))
|
||||
where
|
||||
ans : (x : A) (m : ℕ) → 0G ∼ inverse x + index (constSequence x) m
|
||||
ans x m rewrite indexAndConst x m = Equivalence.symmetric eq invLeft
|
||||
|
||||
<C''Preserves : {a b : A} → a < b → a <C'' injection b
|
||||
<C''Preserves {a} {b} a<b with dense charNot2 a<b
|
||||
... | between , (a<between ,, between<b) = (between + inverse a) , (<WellDefined invRight (Equivalence.reflexive eq) (orderRespectsAddition a<between (inverse a)) ,, (0 , λ m _ → <WellDefined (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq identRight) (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq invLeft))) +Associative) groupIsAbelian) (identityOfIndiscernablesLeft _∼_ (Equivalence.reflexive eq) (indexAndConst b m)) between<b))
|
||||
|
||||
<CInherited : {a b : A} → a < b → (injection a) <C (injection b)
|
||||
<CInherited {a} {b} a<b with dense charNot2 a<b
|
||||
... | between , (a<between ,, between<b) = between , (<C'Preserves a<between ,, <C''Preserves between<b)
|
||||
|
||||
<CInherited' : {a b : A} → (injection a) <C (injection b) → a < b
|
||||
<CInherited' {a} {b} (bound , ((Na , (0<Na ,, (Na2 , prA))) ,, ((Nb , (0<Nb ,, (Nb2 , prB)))))) with prA (succ Na2) (le 0 refl)
|
||||
... | bl with prB (succ Nb2) (le 0 refl)
|
||||
... | bl2 rewrite indexAndConst a Na2 | indexAndConst b Nb2 = SetoidPartialOrder.transitive pOrder (SetoidPartialOrder.transitive pOrder (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<Na a)) bl) (SetoidPartialOrder.transitive pOrder (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<Nb bound)) bl2)
|
||||
|
||||
<CIrreflexive : {a : CauchyCompletion} → a <C a → False
|
||||
<CIrreflexive {a} (bound , ((bA , (0<bA ,, (Na , prA))) ,, ((bB , (0<bB ,, (Nb , prB)))))) with prA (succ Na +N Nb) (le Nb (applyEquality succ (Semiring.commutative ℕSemiring Nb Na)))
|
||||
... | bl with prB (succ Na +N Nb) (le Na refl)
|
||||
... | bl2 with <WellDefined (Equivalence.transitive eq +Associative (+WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) groupIsAbelian) (Equivalence.reflexive eq))) groupIsAbelian (ringAddInequalities bl bl2)
|
||||
... | bad = irreflexive (SetoidPartialOrder.transitive pOrder (<WellDefined identLeft +Associative (<WellDefined (Equivalence.reflexive eq) groupIsAbelian (orderRespectsAddition (<WellDefined identLeft (Equivalence.reflexive eq) (ringAddInequalities 0<bA 0<bB)) (index (CauchyCompletion.elts a) (succ Na +N Nb) + bound)))) bad)
|
||||
|
||||
<CTransitive : {a b c : CauchyCompletion} → a <C b → b <C c → a <C c
|
||||
<CTransitive {a} {b} {c} (boundA , ((eL , (0<eL ,, prL)) ,, (eR , (0<eR ,, (Nr , prR))))) (boundB , ((fL , (0<fL ,, (Nfl , prFl))) ,, (fR , (0<fR ,, (N , prFR))))) = boundA , ((eL , (0<eL ,, prL)) ,, (eR , (0<eR ,, (((Nr +N Nfl) +N N) , λ m N<m → SetoidPartialOrder.transitive pOrder (SetoidPartialOrder.transitive pOrder (SetoidPartialOrder.transitive pOrder (SetoidPartialOrder.transitive pOrder (prR m (shrinkInequality {Nr} (shrinkInequality N<m))) (<WellDefined identLeft (Equivalence.reflexive eq) (orderRespectsAddition 0<fL (index (CauchyCompletion.elts b) m)))) (prFl m (shrinkInequality' {Nr} (shrinkInequality N<m)))) (<WellDefined identLeft groupIsAbelian (orderRespectsAddition 0<fR boundB))) (prFR m (shrinkInequality' N<m))))))
|
||||
|
||||
<COrder : SetoidPartialOrder cauchyCompletionSetoid _<C_
|
||||
SetoidPartialOrder.<WellDefined <COrder {a} {b} {c} {d} a=b c=d a<c = <CWellDefined {a} {b} {c} {d} a=b c=d a<c
|
||||
SetoidPartialOrder.irreflexive <COrder {a} a<a = <CIrreflexive {a} a<a
|
||||
SetoidPartialOrder.transitive <COrder {a} {b} {c} a<b b<c = <CTransitive {a} {b} {c} a<b b<c
|
66
Fields/CauchyCompletion/Group.agda
Normal file
66
Fields/CauchyCompletion/Group.agda
Normal file
@@ -0,0 +1,66 @@
|
||||
{-# 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.Group {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.Addition order F charNot2
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
|
||||
Cassoc : {a b c : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a +C (b +C c)) ((a +C b) +C c)
|
||||
Cassoc {a} {b} {c} ε 0<e = 0 , ans
|
||||
where
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (CauchyCompletion.elts ((a +C (b +C c)) +C (-C ((a +C b) +C c)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (a +C (b +C c))) (map inverse (CauchyCompletion.elts ((a +C b) +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)) (CauchyCompletion.elts c)) inverse m) | indexAndApply (CauchyCompletion.elts b) (CauchyCompletion.elts c) _+_ {m} | indexAndApply (apply _+_ (CauchyCompletion.elts a) (CauchyCompletion.elts b)) (CauchyCompletion.elts c) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts b) _+_ {m} = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (transferToRight'' (Ring.additiveGroup R) +Associative)) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) (absZero order)))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
CidentRight : {a : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a +C injection 0G) a
|
||||
CidentRight {a} ε 0<e = 0 , ans
|
||||
where
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (a +C injection 0G)) (map inverse (CauchyCompletion.elts a))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (a +C injection 0G)) (map inverse (CauchyCompletion.elts a)) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (constSequence 0G) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts a) inverse m) | indexAndConst 0G m = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (Equivalence.transitive eq (+WellDefined (identRight) (Equivalence.reflexive eq)) (invRight))) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) (absZero order)))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
CidentLeft : {a : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (injection 0G +C a) a
|
||||
CidentLeft {a} = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection 0G +C a} {a +C injection 0G} {a} (+CCommutative {injection 0G} {a}) (CidentRight {a})
|
||||
|
||||
CinvRight : {a : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a +C (-C a)) (injection 0G)
|
||||
CinvRight {a} ε 0<e = 0 , ans
|
||||
where
|
||||
ans : {m : ℕ} → (0 <N m) → abs (index (apply _+_ (CauchyCompletion.elts (a +C (-C a))) (map inverse (CauchyCompletion.elts (injection 0G)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (a +C (-C a))) (map inverse (CauchyCompletion.elts (injection 0G))) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (map inverse (CauchyCompletion.elts a)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts a) inverse m) | equalityCommutative (mapAndIndex (constSequence 0G) inverse m) | indexAndConst 0G m = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (Equivalence.transitive eq (+WellDefined invRight (invIdentity (Ring.additiveGroup R))) identRight)) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) (absZero order)))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
CGroup : Group cauchyCompletionSetoid _+C_
|
||||
Group.+WellDefined CGroup {a} {b} {c} {d} x y = additionWellDefined {a} {c} {b} {d} x y
|
||||
Group.0G CGroup = injection 0G
|
||||
Group.inverse CGroup = -C_
|
||||
Group.+Associative CGroup {a} {b} {c} = Cassoc {a} {b} {c}
|
||||
Group.identRight CGroup {a} = CidentRight {a}
|
||||
Group.identLeft CGroup {a} = CidentLeft {a}
|
||||
Group.invLeft CGroup {a} = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {(-C a) +C a} {a +C (-C a)} {injection 0G} (+CCommutative { -C a} {a}) (CinvRight {a})
|
||||
Group.invRight CGroup {a} = CinvRight {a}
|
||||
|
||||
CInjectionGroupHom : GroupHom (Ring.additiveGroup R) CGroup injection
|
||||
GroupHom.groupHom CInjectionGroupHom {x} {y} = additionHom x y
|
||||
GroupHom.wellDefined CInjectionGroupHom {x} {y} x=y = SetoidInjection.wellDefined CInjection {x} {y} x=y
|
64
Fields/CauchyCompletion/Multiplication.agda
Normal file
64
Fields/CauchyCompletion/Multiplication.agda
Normal file
@@ -0,0 +1,64 @@
|
||||
{-# 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.Multiplication {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 Ring R
|
||||
open Group additiveGroup
|
||||
open Field F
|
||||
|
||||
open import Rings.Orders.Lemmas(order)
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Setoid order F charNot2
|
||||
|
||||
_*C_ : CauchyCompletion → CauchyCompletion → CauchyCompletion
|
||||
CauchyCompletion.elts (record { elts = a ; converges = aConv } *C record { elts = b ; converges = bConv }) = apply _*_ a b
|
||||
CauchyCompletion.converges (record { elts = a ; converges = aConv } *C record { elts = b ; converges = bConv }) ε 0<e = {!!}
|
||||
|
||||
*CCommutative : {a b : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a *C b) (b *C a)
|
||||
*CCommutative {a} {b} ε 0<e = 0 , ans
|
||||
where
|
||||
foo : {x y : A} → (x * y) + inverse (y * x) ∼ 0G
|
||||
foo = Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (inverseWellDefined additiveGroup *Commutative)) invRight
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (a *C b)) (map inverse (CauchyCompletion.elts (b *C a)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts b)) (map inverse (apply _*_ (CauchyCompletion.elts b) (CauchyCompletion.elts a))) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts b) _*_ {m} | equalityCommutative (mapAndIndex (apply _*_ (CauchyCompletion.elts b) (CauchyCompletion.elts a)) inverse m) | indexAndApply (CauchyCompletion.elts b) (CauchyCompletion.elts a) _*_ {m} = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ foo) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) (absZero order)))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
multiplicationWellDefinedLeft : (a b c : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid (a *C c) (b *C c)
|
||||
multiplicationWellDefinedLeft a b c a=b ε 0<e = 0 , ans
|
||||
where
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts c)) (map inverse (apply _*_ (CauchyCompletion.elts b) (CauchyCompletion.elts c)))) m) < ε
|
||||
ans {m} N<m rewrite indexAndApply (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts c)) (map inverse (apply _*_ (CauchyCompletion.elts b) (CauchyCompletion.elts c))) _+_ {m} | equalityCommutative (mapAndIndex (apply _*_ (CauchyCompletion.elts b) (CauchyCompletion.elts c)) inverse m) | indexAndApply (CauchyCompletion.elts b) (CauchyCompletion.elts c) _*_ {m} | indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts c) _*_ {m} = {!!}
|
||||
|
||||
multiplicationPreservedLeft : {a b : A} {c : CauchyCompletion} → (a ∼ b) → Setoid._∼_ cauchyCompletionSetoid (injection a *C c) (injection b *C c)
|
||||
multiplicationPreservedLeft {a} {b} {c} a=b = multiplicationWellDefinedLeft (injection a) (injection b) c (injectionPreservesSetoid a b a=b)
|
||||
|
||||
multiplicationPreservedRight : {a b : A} {c : CauchyCompletion} → (a ∼ b) → Setoid._∼_ cauchyCompletionSetoid (c *C injection a) (c *C injection b)
|
||||
multiplicationPreservedRight {a} {b} {c} a=b = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {c *C injection a} {injection a *C c} {c *C injection b} (*CCommutative {c} {injection a}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection a *C c} {injection b *C c} {c *C injection b} (multiplicationPreservedLeft {a} {b} {c} a=b) (*CCommutative {injection b} {c}))
|
||||
|
||||
multiplicationPreserved : {a b c d : A} → (a ∼ b) → (c ∼ d) → Setoid._∼_ cauchyCompletionSetoid (injection a *C injection c) (injection b *C injection d)
|
||||
multiplicationPreserved {a} {b} {c} {d} a=b c=d = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection a *C injection c} {injection a *C injection d} {injection b *C injection d} (multiplicationPreservedRight {c} {d} {injection a} c=d) (multiplicationPreservedLeft {a} {b} {injection d} a=b)
|
||||
|
||||
multiplicationWellDefinedRight : (a b c : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid b c → Setoid._∼_ cauchyCompletionSetoid (a *C b) (a *C c)
|
||||
multiplicationWellDefinedRight a b c b=c = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {a *C b} {b *C a} {a *C c} (*CCommutative {a} {b}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {b *C a} {c *C a} {a *C c} (multiplicationWellDefinedLeft b c a b=c) (*CCommutative {c} {a}))
|
||||
|
||||
multiplicationWellDefined : {a b c d : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid c d → Setoid._∼_ cauchyCompletionSetoid (a *C c) (b *C d)
|
||||
multiplicationWellDefined {a} {b} {c} {d} a=b c=d = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {a *C c} {a *C d} {b *C d} (multiplicationWellDefinedRight a c d c=d) (multiplicationWellDefinedLeft a b d a=b)
|
55
Fields/CauchyCompletion/Ring.agda
Normal file
55
Fields/CauchyCompletion/Ring.agda
Normal file
@@ -0,0 +1,55 @@
|
||||
{-# OPTIONS --allow-unsolved-metas --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.Ring {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
|
||||
|
||||
c*Assoc : {a b c : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a *C (b *C c)) ((a *C b) *C c)
|
||||
c*Assoc {a} {b} {c} ε 0<e = 0 , ans
|
||||
where
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (CauchyCompletion.elts ((a *C (b *C c)) +C (-C ((a *C b) *C c)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (a *C (b *C c))) (CauchyCompletion.elts (-C ((a *C b) *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)) (CauchyCompletion.elts c)) inverse m) | indexAndApply (CauchyCompletion.elts b) (CauchyCompletion.elts c) _*_ {m} | indexAndApply (apply _*_ (CauchyCompletion.elts a) (CauchyCompletion.elts b)) (CauchyCompletion.elts c) _*_ {m} | indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts b) _*_ {m} = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (transferToRight'' (Ring.additiveGroup R) (Ring.*Associative R))) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) (absZero order)))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
c*Ident : {a : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (injection (Ring.1R R) *C a) a
|
||||
c*Ident {a} ε 0<e = 0 , ans
|
||||
where
|
||||
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
|
||||
|
||||
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
|
||||
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.identIsIdent CRing {a} = c*Ident {a}
|
133
Fields/CauchyCompletion/Setoid.agda
Normal file
133
Fields/CauchyCompletion/Setoid.agda
Normal file
@@ -0,0 +1,133 @@
|
||||
{-# 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
|
||||
open import Semirings.Definition
|
||||
|
||||
module Fields.CauchyCompletion.Setoid {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 Ring R
|
||||
open Group additiveGroup
|
||||
open Field F
|
||||
|
||||
open import Fields.Lemmas F
|
||||
open import Fields.CauchyCompletion.Definition order F
|
||||
open import Fields.CauchyCompletion.Addition order F charNot2
|
||||
open import Rings.Orders.Lemmas(order)
|
||||
|
||||
isZero : CauchyCompletion → Set (m ⊔ o)
|
||||
isZero record { elts = elts ; converges = converges } = ∀ ε → 0R < ε → Sg ℕ (λ N → ∀ {m : ℕ} → (N <N m) → (abs (index elts m)) < ε)
|
||||
|
||||
transitiveLemma : {a b c e/2 : A} → abs (a + inverse b) < e/2 → abs (b + inverse c) < e/2 → (abs (a + inverse c)) < (e/2 + e/2)
|
||||
transitiveLemma {a} {b} {c} {e/2} a-b<e/2 b-c<e/2 with triangleInequality (a + inverse b) (b + inverse c)
|
||||
transitiveLemma {a} {b} {c} {e/2} a-b<e/2 b-c<e/2 | inl x = SetoidPartialOrder.transitive pOrder (<WellDefined (absWellDefined _ _ (Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq +Associative (Equivalence.transitive eq (+WellDefined invLeft (Equivalence.reflexive eq)) identLeft))))) (Equivalence.reflexive eq) x) (ringAddInequalities a-b<e/2 b-c<e/2)
|
||||
transitiveLemma {a} {b} {c} {e/2} a-b<e/2 b-c<e/2 | inr x = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ (Equivalence.symmetric eq ((Equivalence.transitive eq (Equivalence.symmetric eq +Associative) (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq +Associative (Equivalence.transitive eq (+WellDefined invLeft (Equivalence.reflexive eq)) identLeft))))))) x)) (Equivalence.reflexive eq) (ringAddInequalities a-b<e/2 b-c<e/2)
|
||||
|
||||
cauchyCompletionSetoid : Setoid CauchyCompletion
|
||||
(cauchyCompletionSetoid Setoid.∼ a) b = isZero (a +C (-C b))
|
||||
Equivalence.reflexive (Setoid.eq cauchyCompletionSetoid) {x} ε 0<e = 0 , t
|
||||
where
|
||||
t : {m : ℕ} → (0 <N m) → abs (index (apply _+_ (CauchyCompletion.elts x) (map inverse (CauchyCompletion.elts x))) m) < ε
|
||||
t {m} 0<m rewrite indexAndApply (CauchyCompletion.elts x) (map inverse (CauchyCompletion.elts x)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts x) inverse m) = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ invRight) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) (absZero order)))) (Equivalence.reflexive eq) 0<e
|
||||
Equivalence.symmetric (Setoid.eq cauchyCompletionSetoid) {x} {y} x=y ε 0<e with x=y ε 0<e
|
||||
Equivalence.symmetric (Setoid.eq cauchyCompletionSetoid) {x} {y} x=y ε 0<e | N , pr = N , t
|
||||
where
|
||||
t : {m : ℕ} → N <N m → abs (index (apply _+_ (CauchyCompletion.elts y) (map inverse (CauchyCompletion.elts x))) m) < ε
|
||||
t {m} N<m = <WellDefined (Equivalence.transitive eq (absNegation _) (absWellDefined _ _ (identityOfIndiscernablesRight _∼_ (identityOfIndiscernablesLeft _∼_ (identityOfIndiscernablesLeft _∼_ (Equivalence.transitive eq (inverseDistributes) (Equivalence.transitive eq groupIsAbelian (+WellDefined (Equivalence.transitive eq (inverseWellDefined additiveGroup (identityOfIndiscernablesLeft _∼_ (Equivalence.reflexive eq) (mapAndIndex _ inverse m))) (invTwice additiveGroup (index (CauchyCompletion.elts y) m))) (identityOfIndiscernablesLeft _∼_ (Equivalence.reflexive eq) (equalityCommutative (mapAndIndex (CauchyCompletion.elts x) inverse m)))))) (equalityCommutative (mapAndApply (CauchyCompletion.elts x) (map inverse (CauchyCompletion.elts y)) _+_ inverse m))) (equalityCommutative (mapAndIndex _ inverse m))) (equalityCommutative (indexAndApply (CauchyCompletion.elts y) (map inverse (CauchyCompletion.elts x)) _+_ {m}))))) (Equivalence.reflexive eq) (pr N<m)
|
||||
Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {x} {y} {z} x=y y=z ε 0<e with halve charNot2 ε
|
||||
... | e/2 , prHalf with x=y e/2 (halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prHalf) 0<e))
|
||||
... | Nx , prX with y=z e/2 (halvePositive e/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prHalf) 0<e))
|
||||
... | Ny , prY = (Nx +N Ny) , t
|
||||
where
|
||||
t : {m : ℕ} → (Nx +N Ny <N m) → abs (index (apply _+_ (CauchyCompletion.elts x) (map inverse (CauchyCompletion.elts z))) m) < ε
|
||||
t {m} nsPr with prX {m} (le (_<N_.x nsPr +N Ny) (transitivity (applyEquality succ (transitivity (equalityCommutative (Semiring.+Associative ℕSemiring _ Ny Nx)) (applyEquality ( λ i → (_<N_.x nsPr) +N i) (Semiring.commutative ℕSemiring Ny Nx)))) (_<N_.proof nsPr)))
|
||||
... | x-y<e/2 with prY {m} (le (_<N_.x nsPr +N Nx) (transitivity (applyEquality succ (equalityCommutative (Semiring.+Associative ℕSemiring _ Nx Ny))) (_<N_.proof nsPr)))
|
||||
... | y-z<e/2 rewrite indexAndApply (CauchyCompletion.elts x) (map inverse (CauchyCompletion.elts z)) _+_ {m} | indexAndApply (CauchyCompletion.elts x) (map inverse (CauchyCompletion.elts y)) _+_ {m} | indexAndApply (CauchyCompletion.elts y) (map inverse (CauchyCompletion.elts z)) _+_ {m} | equalityCommutative (mapAndIndex (CauchyCompletion.elts z) inverse m) | equalityCommutative (mapAndIndex (CauchyCompletion.elts y) inverse m) = <WellDefined (Equivalence.reflexive eq) prHalf (transitiveLemma x-y<e/2 y-z<e/2)
|
||||
|
||||
injectionPreservesSetoid : (a b : A) → (a ∼ b) → Setoid._∼_ cauchyCompletionSetoid (injection a) (injection b)
|
||||
injectionPreservesSetoid a b a=b ε 0<e = 0 , t
|
||||
where
|
||||
t : {m : ℕ} → 0 <N m → abs (index (apply _+_ (constSequence a) (map inverse (constSequence b))) m) < ε
|
||||
t {m} 0<m = <WellDefined (identityOfIndiscernablesLeft _∼_ (absWellDefined 0G _ (identityOfIndiscernablesRight _∼_ (Equivalence.transitive eq (Equivalence.symmetric eq (transferToRight'' additiveGroup a=b)) (+WellDefined (identityOfIndiscernablesLeft _∼_ (Equivalence.reflexive eq) (indexAndConst a m)) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) (transitivity (applyEquality inverse (equalityCommutative (indexAndConst _ m))) (mapAndIndex _ inverse m))))) (equalityCommutative (indexAndApply (constSequence a) _ _+_ {m})))) (absZero order)) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
infinitesimalImplies0 : (a : A) → ({ε : A} → (0R < ε) → a < ε) → (a ∼ 0R) || (a < 0R)
|
||||
infinitesimalImplies0 a pr with SetoidTotalOrder.totality tOrder 0R a
|
||||
infinitesimalImplies0 a pr | inl (inl 0<a) with halve charNot2 a
|
||||
infinitesimalImplies0 a pr | inl (inl 0<a) | a/2 , prA/2 with pr {a/2} (halvePositive a/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prA/2) 0<a))
|
||||
... | bl with halvePositive a/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq prA/2) 0<a)
|
||||
... | 0<a/2 = exFalso (irreflexive {a} (<WellDefined identLeft prA/2 (ringAddInequalities 0<a/2 bl)))
|
||||
infinitesimalImplies0 a pr | inl (inr a<0) = inr a<0
|
||||
infinitesimalImplies0 a pr | inr 0=a = inl (Equivalence.symmetric eq 0=a)
|
||||
|
||||
injectionPreservesSetoid' : (a b : A) → Setoid._∼_ cauchyCompletionSetoid (injection a) (injection b) → a ∼ b
|
||||
injectionPreservesSetoid' a b a=b = transferToRight additiveGroup (absZeroImpliesZero ans)
|
||||
where
|
||||
infinitesimal : {ε : A} → 0G < ε → abs (a + inverse b) < ε
|
||||
infinitesimal {ε} 0<e with a=b ε 0<e
|
||||
... | N , pr with pr {succ N} (le 0 refl)
|
||||
... | bl rewrite indexAndApply (constSequence a) (map inverse (constSequence b)) _+_ {N} | indexAndConst a N | equalityCommutative (mapAndIndex (constSequence b) inverse N) | indexAndConst b N = bl
|
||||
ans : (abs (a + inverse b)) ∼ 0G
|
||||
ans with infinitesimalImplies0 _ infinitesimal
|
||||
... | inl x = x
|
||||
... | inr x = exFalso (absNonnegative x)
|
||||
|
||||
+CCommutative : {a b : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid (a +C b) (b +C a)
|
||||
+CCommutative {a} {b} ε 0<e = 0 , ans
|
||||
where
|
||||
foo : {x y : A} → (x + y) + inverse (y + x) ∼ 0G
|
||||
foo = Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (inverseWellDefined additiveGroup groupIsAbelian)) invRight
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (a +C b)) (map inverse (CauchyCompletion.elts (b +C a)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (a +C b)) (map inverse (CauchyCompletion.elts (b +C a))) _+_ {m} | indexAndApply (CauchyCompletion.elts a) (CauchyCompletion.elts b) _+_ {m} | equalityCommutative (mapAndIndex (apply _+_ (CauchyCompletion.elts b) (CauchyCompletion.elts a)) inverse m) | indexAndApply (CauchyCompletion.elts b) (CauchyCompletion.elts a) _+_ {m} = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ foo) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) (absZero order)) )) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
additionWellDefinedLeft : (a b c : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid (a +C c) (b +C c)
|
||||
additionWellDefinedLeft record { elts = a ; converges = aConv } record { elts = b ; converges = bConv } record { elts = c ; converges = cConv } a=b ε 0<e with a=b ε 0<e
|
||||
... | Na-b , prA-b = Na-b , ans
|
||||
where
|
||||
ans : {m : ℕ} → Na-b <N m → abs (index (apply _+_ (apply _+_ a c) (map inverse (apply _+_ b c))) m) < ε
|
||||
ans {m} mBig with prA-b {m} mBig
|
||||
... | bl rewrite indexAndApply (apply _+_ a c) (map inverse (apply _+_ b c)) _+_ {m} | indexAndApply a c _+_ {m} | equalityCommutative (mapAndIndex (apply _+_ b c) inverse m) | indexAndApply b c _+_ {m} = <WellDefined (absWellDefined _ _ t) (Equivalence.reflexive eq) bl
|
||||
where
|
||||
t : index (apply _+_ a (map inverse b)) m ∼ ((index a m + index c m) + inverse (index b m + index c m))
|
||||
t rewrite indexAndApply a (map inverse b) _+_ {m} | equalityCommutative (mapAndIndex b inverse m) = Equivalence.transitive eq (+WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq identLeft) (+WellDefined (Equivalence.symmetric eq invRight) (Equivalence.reflexive eq))) (Equivalence.symmetric eq +Associative)) (+WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq (invContravariant additiveGroup))))) (+Associative {index a m})
|
||||
|
||||
additionPreservedLeft : {a b : A} {c : CauchyCompletion} → (a ∼ b) → Setoid._∼_ cauchyCompletionSetoid (injection a +C c) (injection b +C c)
|
||||
additionPreservedLeft {a} {b} {c} a=b = additionWellDefinedLeft (injection a) (injection b) c (injectionPreservesSetoid a b a=b)
|
||||
|
||||
additionPreservedRight : {a b : A} {c : CauchyCompletion} → (a ∼ b) → Setoid._∼_ cauchyCompletionSetoid (c +C injection a) (c +C injection b)
|
||||
additionPreservedRight {a} {b} {c} a=b = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {c +C injection a} {injection a +C c} {c +C injection b} (+CCommutative {c} {injection a}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection a +C c} {injection b +C c} {c +C injection b} (additionPreservedLeft {a} {b} {c} a=b) (+CCommutative {injection b} {c}))
|
||||
|
||||
additionPreserved : {a b c d : A} → (a ∼ b) → (c ∼ d) → Setoid._∼_ cauchyCompletionSetoid (injection a +C injection c) (injection b +C injection d)
|
||||
additionPreserved {a} {b} {c} {d} a=b c=d = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {injection a +C injection c} {injection a +C injection d} {injection b +C injection d} (additionPreservedRight {c} {d} {injection a} c=d) (additionPreservedLeft {a} {b} {injection d} a=b)
|
||||
|
||||
additionWellDefinedRight : (a b c : CauchyCompletion) → Setoid._∼_ cauchyCompletionSetoid b c → Setoid._∼_ cauchyCompletionSetoid (a +C b) (a +C c)
|
||||
additionWellDefinedRight a b c b=c = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {a +C b} {b +C a} {a +C c} (+CCommutative {a} {b}) (Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {b +C a} {c +C a} {a +C c} (additionWellDefinedLeft b c a b=c) (+CCommutative {c} {a}))
|
||||
|
||||
additionWellDefined : {a b c d : CauchyCompletion} → Setoid._∼_ cauchyCompletionSetoid a b → Setoid._∼_ cauchyCompletionSetoid c d → Setoid._∼_ cauchyCompletionSetoid (a +C c) (b +C d)
|
||||
additionWellDefined {a} {b} {c} {d} a=b c=d = Equivalence.transitive (Setoid.eq cauchyCompletionSetoid) {a +C c} {a +C d} {b +C d} (additionWellDefinedRight a c d c=d) (additionWellDefinedLeft a b d a=b)
|
||||
|
||||
additionHom : (x y : A) → Setoid._∼_ cauchyCompletionSetoid (injection (x + y)) (injection x +C injection y)
|
||||
additionHom x y ε 0<e = 0 , ans
|
||||
where
|
||||
ans : {m : ℕ} → 0 <N m → abs (index (apply _+_ (CauchyCompletion.elts (injection (x + y))) (map inverse (CauchyCompletion.elts (injection x +C injection y)))) m) < ε
|
||||
ans {m} 0<m rewrite indexAndApply (CauchyCompletion.elts (injection (x + y))) (map inverse (CauchyCompletion.elts (injection x +C injection y))) _+_ {m} | equalityCommutative (mapAndIndex (apply _+_ (constSequence x) (constSequence y)) inverse m) | indexAndConst (x + y) m | indexAndApply (constSequence x) (constSequence y) _+_ {m} | indexAndConst x m | indexAndConst y m = <WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq (absWellDefined _ _ invRight) (identityOfIndiscernablesRight _∼_ (Equivalence.reflexive eq) (absZero order)))) (Equivalence.reflexive eq) 0<e
|
||||
|
||||
CInjection : SetoidInjection S cauchyCompletionSetoid injection
|
||||
SetoidInjection.wellDefined CInjection {x} {y} x=y = injectionPreservesSetoid x y x=y
|
||||
SetoidInjection.injective CInjection {x} {y} x=y = injectionPreservesSetoid' x y x=y
|
@@ -15,179 +15,180 @@ open import Sets.EquivalenceRelations
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.FieldOfFractions where
|
||||
fieldOfFractionsSet : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} → {R : Ring S _+_ _*_} → IntegralDomain R → Set (a ⊔ b)
|
||||
fieldOfFractionsSet {A = A} {S = S} {R = R} I = (A && (Sg A (λ a → (Setoid._∼_ S a (Ring.0R R) → False))))
|
||||
|
||||
fieldOfFractionsSetoid : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} → {R : Ring S _+_ _*_} → (I : IntegralDomain R) → Setoid (fieldOfFractionsSet I)
|
||||
Setoid._∼_ (fieldOfFractionsSetoid {S = S} {_*_ = _*_} I) (a ,, (b , b!=0)) (c ,, (d , d!=0)) = Setoid._∼_ S (a * d) (b * c)
|
||||
Equivalence.reflexive (Setoid.eq (fieldOfFractionsSetoid {R = R} I)) {a ,, (b , b!=0)} = Ring.*Commutative R
|
||||
Equivalence.symmetric (Setoid.eq (fieldOfFractionsSetoid {S = S} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} ad=bc = transitive (Ring.*Commutative R) (transitive (symmetric ad=bc) (Ring.*Commutative R))
|
||||
where
|
||||
open Equivalence (Setoid.eq S)
|
||||
Equivalence.transitive (Setoid.eq (fieldOfFractionsSetoid {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} ad=bc cf=de = p5
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence eq
|
||||
p : (a * d) * f ∼ (b * c) * f
|
||||
p = Ring.*WellDefined R ad=bc reflexive
|
||||
p2 : (a * f) * d ∼ b * (d * e)
|
||||
p2 = transitive (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) (transitive p (transitive (symmetric *Associative) (*WellDefined reflexive cf=de)))
|
||||
p3 : (a * f) * d ∼ (b * e) * d
|
||||
p3 = transitive p2 (transitive (*WellDefined reflexive *Commutative) *Associative)
|
||||
p4 : (d ∼ 0R) || ((a * f) ∼ (b * e))
|
||||
p4 = cancelIntDom I (transitive *Commutative (transitive p3 *Commutative))
|
||||
p5 : (a * f) ∼ (b * e)
|
||||
p5 with p4
|
||||
p5 | inl d=0 = exFalso (d!=0 d=0)
|
||||
p5 | inr x = x
|
||||
fieldOfFractionsSet : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} → {R : Ring S _+_ _*_} → IntegralDomain R → Set (a ⊔ b)
|
||||
fieldOfFractionsSet {A = A} {S = S} {R = R} I = (A && (Sg A (λ a → (Setoid._∼_ S a (Ring.0R R) → False))))
|
||||
|
||||
fieldOfFractionsPlus : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} → (I : IntegralDomain R) → fieldOfFractionsSet I → fieldOfFractionsSet I → fieldOfFractionsSet I
|
||||
fieldOfFractionsPlus {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I (a ,, (b , b!=0)) (c ,, (d , d!=0)) = (((a * d) + (b * c)) ,, ((b * d) , ans))
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
ans : ((b * d) ∼ Ring.0R R) → False
|
||||
ans pr with IntegralDomain.intDom I pr
|
||||
ans pr | inl x = b!=0 x
|
||||
ans pr | inr x = d!=0 x
|
||||
fieldOfFractionsSetoid : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} → {R : Ring S _+_ _*_} → (I : IntegralDomain R) → Setoid (fieldOfFractionsSet I)
|
||||
Setoid._∼_ (fieldOfFractionsSetoid {S = S} {_*_ = _*_} I) (a ,, (b , b!=0)) (c ,, (d , d!=0)) = Setoid._∼_ S (a * d) (b * c)
|
||||
Equivalence.reflexive (Setoid.eq (fieldOfFractionsSetoid {R = R} I)) {a ,, (b , b!=0)} = Ring.*Commutative R
|
||||
Equivalence.symmetric (Setoid.eq (fieldOfFractionsSetoid {S = S} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} ad=bc = transitive (Ring.*Commutative R) (transitive (symmetric ad=bc) (Ring.*Commutative R))
|
||||
where
|
||||
open Equivalence (Setoid.eq S)
|
||||
Equivalence.transitive (Setoid.eq (fieldOfFractionsSetoid {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} ad=bc cf=de = p5
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence eq
|
||||
p : (a * d) * f ∼ (b * c) * f
|
||||
p = Ring.*WellDefined R ad=bc reflexive
|
||||
p2 : (a * f) * d ∼ b * (d * e)
|
||||
p2 = transitive (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) (transitive p (transitive (symmetric *Associative) (*WellDefined reflexive cf=de)))
|
||||
p3 : (a * f) * d ∼ (b * e) * d
|
||||
p3 = transitive p2 (transitive (*WellDefined reflexive *Commutative) *Associative)
|
||||
p4 : (d ∼ 0R) || ((a * f) ∼ (b * e))
|
||||
p4 = cancelIntDom I (transitive *Commutative (transitive p3 *Commutative))
|
||||
p5 : (a * f) ∼ (b * e)
|
||||
p5 with p4
|
||||
p5 | inl d=0 = exFalso (d!=0 d=0)
|
||||
p5 | inr x = x
|
||||
|
||||
fieldOfFractionsTimes : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} → (I : IntegralDomain R) → fieldOfFractionsSet I → fieldOfFractionsSet I → fieldOfFractionsSet I
|
||||
fieldOfFractionsTimes {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I (a ,, (b , b!=0)) (c ,, (d , d!=0)) = (a * c) ,, ((b * d) , ans)
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
ans : ((b * d) ∼ Ring.0R R) → False
|
||||
ans pr with IntegralDomain.intDom I pr
|
||||
ans pr | inl x = b!=0 x
|
||||
ans pr | inr x = d!=0 x
|
||||
fieldOfFractionsPlus : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} → (I : IntegralDomain R) → fieldOfFractionsSet I → fieldOfFractionsSet I → fieldOfFractionsSet I
|
||||
fieldOfFractionsPlus {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I (a ,, (b , b!=0)) (c ,, (d , d!=0)) = (((a * d) + (b * c)) ,, ((b * d) , ans))
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
ans : ((b * d) ∼ Ring.0R R) → False
|
||||
ans pr with IntegralDomain.intDom I pr
|
||||
ans pr | inl x = b!=0 x
|
||||
ans pr | inr x = d!=0 x
|
||||
|
||||
fieldOfFractionsRing : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} → (I : IntegralDomain R) → Ring (fieldOfFractionsSetoid I) (fieldOfFractionsPlus I) (fieldOfFractionsTimes I)
|
||||
Group.+WellDefined (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} {g ,, (h , h!=0)} af=be ch=dg = need
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence eq
|
||||
have1 : (c * h) ∼ (d * g)
|
||||
have1 = ch=dg
|
||||
have2 : (a * f) ∼ (b * e)
|
||||
have2 = af=be
|
||||
need : (((a * d) + (b * c)) * (f * h)) ∼ ((b * d) * (((e * h) + (f * g))))
|
||||
need = transitive (transitive (Ring.*Commutative R) (transitive (Ring.*DistributesOver+ R) (Group.+WellDefined (Ring.additiveGroup R) (transitive *Associative (transitive (*WellDefined (*Commutative) reflexive) (transitive (*WellDefined *Associative reflexive) (transitive (*WellDefined (*WellDefined have2 reflexive) reflexive) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) (transitive *Associative (transitive (*WellDefined (transitive (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative)) *Associative) reflexive) (symmetric *Associative))))))))) (transitive *Commutative (transitive (transitive (symmetric *Associative) (*WellDefined reflexive (transitive (*WellDefined reflexive *Commutative) (transitive *Associative (transitive (*WellDefined have1 reflexive) (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative))))))) *Associative))))) (symmetric (Ring.*DistributesOver+ R))
|
||||
Group.0G (Ring.additiveGroup (fieldOfFractionsRing {R = R} I)) = Ring.0R R ,, (Ring.1R R , IntegralDomain.nontrivial I)
|
||||
Group.inverse (Ring.additiveGroup (fieldOfFractionsRing {R = R} I)) (a ,, b) = Group.inverse (Ring.additiveGroup R) a ,, b
|
||||
Group.+Associative (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((a * (d * f)) + (b * ((c * f) + (d * e)))) * ((b * d) * f)) ∼ ((b * (d * f)) * ((((a * d) + (b * c)) * f) + ((b * d) * e)))
|
||||
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (symmetric (Ring.*Associative R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Ring.*DistributesOver+ R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Associative R) (Ring.*Associative R))) (transitive (Group.+Associative (Ring.additiveGroup R)) (Group.+WellDefined (Ring.additiveGroup R) (transitive (transitive (Group.+WellDefined (Ring.additiveGroup R) (transitive (Ring.*Associative R) (Ring.*Commutative R)) (Ring.*Commutative R)) (symmetric (Ring.*DistributesOver+ R))) (Ring.*Commutative R)) reflexive)))))
|
||||
Group.identRight (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((a * Ring.1R R) + (b * Group.0G (Ring.additiveGroup R))) * b) ∼ ((b * Ring.1R R) * a)
|
||||
need = transitive (transitive (Ring.*WellDefined R (transitive (Group.+WellDefined (Ring.additiveGroup R) (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) reflexive) (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Ring.timesZero R)) (Group.identRight (Ring.additiveGroup R)))) reflexive) (Ring.*Commutative R)) (symmetric (Ring.*WellDefined R (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) reflexive))
|
||||
Group.identLeft (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((Group.0G (Ring.additiveGroup R) * b) + (Ring.1R R * a)) * b) ∼ ((Ring.1R R * b) * a)
|
||||
need = transitive (transitive (Ring.*WellDefined R (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Ring.identIsIdent R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) (transitive (Ring.*Commutative R) (Ring.timesZero R)) reflexive) (Group.identLeft (Ring.additiveGroup R)))) reflexive) (Ring.*Commutative R)) (Ring.*WellDefined R (symmetric (Ring.identIsIdent R)) reflexive)
|
||||
Group.invLeft (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((Group.inverse (Ring.additiveGroup R) a * b) + (b * a)) * Ring.1R R) ∼ ((b * b) * Group.0G (Ring.additiveGroup R))
|
||||
need = transitive (transitive (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) reflexive) (transitive (symmetric (Ring.*DistributesOver+ R)) (transitive (Ring.*WellDefined R reflexive (Group.invLeft (Ring.additiveGroup R))) (Ring.timesZero R))))) (symmetric (Ring.timesZero R))
|
||||
Group.invRight (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((a * b) + (b * Group.inverse (Ring.additiveGroup R) a)) * Ring.1R R) ∼ ((b * b) * Group.0G (Ring.additiveGroup R))
|
||||
need = transitive (transitive (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) reflexive) (transitive (symmetric (Ring.*DistributesOver+ R)) (transitive (Ring.*WellDefined R reflexive (Group.invRight (Ring.additiveGroup R))) (Ring.timesZero R))))) (symmetric (Ring.timesZero R))
|
||||
Ring.*WellDefined (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} {g ,, (h , _)} af=be ch=dg = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : ((a * c) * (f * h)) ∼ ((b * d) * (e * g))
|
||||
need = transitive (Ring.*WellDefined R reflexive (Ring.*Commutative R)) (transitive (Ring.*Associative R) (transitive (Ring.*WellDefined R (symmetric (Ring.*Associative R)) reflexive) (transitive (Ring.*WellDefined R (Ring.*WellDefined R reflexive ch=dg) reflexive) (transitive (Ring.*Commutative R) (transitive (Ring.*Associative R) (transitive (Ring.*WellDefined R (Ring.*Commutative R) reflexive) (transitive (Ring.*WellDefined R af=be reflexive) (transitive (Ring.*Associative R) (transitive (Ring.*WellDefined R (transitive (symmetric (Ring.*Associative R)) (transitive (Ring.*WellDefined R reflexive (Ring.*Commutative R)) (Ring.*Associative R))) reflexive) (symmetric (Ring.*Associative R)))))))))))
|
||||
Ring.1R (fieldOfFractionsRing {R = R} I) = Ring.1R R ,, (Ring.1R R , IntegralDomain.nontrivial I)
|
||||
Ring.groupIsAbelian (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((a * d) + (b * c)) * (d * b)) ∼ ((b * d) * ((c * b) + (d * a)))
|
||||
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (Ring.*Commutative R) (transitive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) (Ring.*Commutative R)) (Ring.groupIsAbelian R)))
|
||||
Ring.*Associative (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : ((a * (c * e)) * ((b * d) * f)) ∼ ((b * (d * f)) * ((a * c) * e))
|
||||
need = transitive (Ring.*WellDefined R (Ring.*Associative R) (symmetric (Ring.*Associative R))) (Ring.*Commutative R)
|
||||
Ring.*Commutative (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : ((a * c) * (d * b)) ∼ ((b * d) * (c * a))
|
||||
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (Ring.*Commutative R) (Ring.*Commutative R))
|
||||
Ring.*DistributesOver+ (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence eq
|
||||
inter : b * (a * ((c * f) + (d * e))) ∼ (((a * c) * (b * f)) + ((b * d) * (a * e)))
|
||||
inter = transitive *Associative (transitive *DistributesOver+ (Group.+WellDefined additiveGroup (transitive *Associative (transitive (*WellDefined (transitive (*WellDefined (*Commutative) reflexive) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative))) reflexive) (symmetric *Associative))) (transitive *Associative (transitive (*WellDefined (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) reflexive) (symmetric *Associative)))))
|
||||
need : ((a * ((c * f) + (d * e))) * ((b * d) * (b * f))) ∼ ((b * (d * f)) * (((a * c) * (b * f)) + ((b * d) * (a * e))))
|
||||
need = transitive (Ring.*WellDefined R reflexive (Ring.*WellDefined R reflexive (Ring.*Commutative R))) (transitive (Ring.*WellDefined R reflexive (Ring.*Associative R)) (transitive (Ring.*Commutative R) (transitive (Ring.*WellDefined R (Ring.*WellDefined R (symmetric (Ring.*Associative R)) reflexive) reflexive) (transitive (symmetric (Ring.*Associative R)) (Ring.*WellDefined R reflexive inter)))))
|
||||
Ring.identIsIdent (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((Ring.1R R) * a) * b) ∼ (((Ring.1R R * b)) * a)
|
||||
need = transitive (Ring.*WellDefined R (Ring.identIsIdent R) reflexive) (transitive (Ring.*Commutative R) (Ring.*WellDefined R (symmetric (Ring.identIsIdent R)) reflexive))
|
||||
fieldOfFractionsTimes : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} → (I : IntegralDomain R) → fieldOfFractionsSet I → fieldOfFractionsSet I → fieldOfFractionsSet I
|
||||
fieldOfFractionsTimes {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I (a ,, (b , b!=0)) (c ,, (d , d!=0)) = (a * c) ,, ((b * d) , ans)
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
ans : ((b * d) ∼ Ring.0R R) → False
|
||||
ans pr with IntegralDomain.intDom I pr
|
||||
ans pr | inl x = b!=0 x
|
||||
ans pr | inr x = d!=0 x
|
||||
|
||||
fieldOfFractions : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} → (I : IntegralDomain R) → Field (fieldOfFractionsRing I)
|
||||
Field.allInvertible (fieldOfFractions {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) (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)))
|
||||
Field.nontrivial (fieldOfFractions {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) pr = IntegralDomain.nontrivial I (symmetric (transitive (symmetric (Ring.timesZero R)) (transitive (Ring.*Commutative R) (transitive pr (Ring.identIsIdent R)))))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
pr' : (Ring.0R R) * (Ring.1R R) ∼ (Ring.1R R) * (Ring.1R R)
|
||||
pr' = pr
|
||||
fieldOfFractionsRing : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} → (I : IntegralDomain R) → Ring (fieldOfFractionsSetoid I) (fieldOfFractionsPlus I) (fieldOfFractionsTimes I)
|
||||
Group.+WellDefined (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} {g ,, (h , h!=0)} af=be ch=dg = need
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence eq
|
||||
have1 : (c * h) ∼ (d * g)
|
||||
have1 = ch=dg
|
||||
have2 : (a * f) ∼ (b * e)
|
||||
have2 = af=be
|
||||
need : (((a * d) + (b * c)) * (f * h)) ∼ ((b * d) * (((e * h) + (f * g))))
|
||||
need = transitive (transitive (Ring.*Commutative R) (transitive (Ring.*DistributesOver+ R) (Group.+WellDefined (Ring.additiveGroup R) (transitive *Associative (transitive (*WellDefined (*Commutative) reflexive) (transitive (*WellDefined *Associative reflexive) (transitive (*WellDefined (*WellDefined have2 reflexive) reflexive) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) (transitive *Associative (transitive (*WellDefined (transitive (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative)) *Associative) reflexive) (symmetric *Associative))))))))) (transitive *Commutative (transitive (transitive (symmetric *Associative) (*WellDefined reflexive (transitive (*WellDefined reflexive *Commutative) (transitive *Associative (transitive (*WellDefined have1 reflexive) (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative))))))) *Associative))))) (symmetric (Ring.*DistributesOver+ R))
|
||||
Group.0G (Ring.additiveGroup (fieldOfFractionsRing {R = R} I)) = Ring.0R R ,, (Ring.1R R , IntegralDomain.nontrivial I)
|
||||
Group.inverse (Ring.additiveGroup (fieldOfFractionsRing {R = R} I)) (a ,, b) = Group.inverse (Ring.additiveGroup R) a ,, b
|
||||
Group.+Associative (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((a * (d * f)) + (b * ((c * f) + (d * e)))) * ((b * d) * f)) ∼ ((b * (d * f)) * ((((a * d) + (b * c)) * f) + ((b * d) * e)))
|
||||
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (symmetric (Ring.*Associative R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Ring.*DistributesOver+ R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Associative R) (Ring.*Associative R))) (transitive (Group.+Associative (Ring.additiveGroup R)) (Group.+WellDefined (Ring.additiveGroup R) (transitive (transitive (Group.+WellDefined (Ring.additiveGroup R) (transitive (Ring.*Associative R) (Ring.*Commutative R)) (Ring.*Commutative R)) (symmetric (Ring.*DistributesOver+ R))) (Ring.*Commutative R)) reflexive)))))
|
||||
Group.identRight (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , b!=0)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((a * Ring.1R R) + (b * Group.0G (Ring.additiveGroup R))) * b) ∼ ((b * Ring.1R R) * a)
|
||||
need = transitive (transitive (Ring.*WellDefined R (transitive (Group.+WellDefined (Ring.additiveGroup R) (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) reflexive) (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Ring.timesZero R)) (Group.identRight (Ring.additiveGroup R)))) reflexive) (Ring.*Commutative R)) (symmetric (Ring.*WellDefined R (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) reflexive))
|
||||
Group.identLeft (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((Group.0G (Ring.additiveGroup R) * b) + (Ring.1R R * a)) * b) ∼ ((Ring.1R R * b) * a)
|
||||
need = transitive (transitive (Ring.*WellDefined R (transitive (Group.+WellDefined (Ring.additiveGroup R) reflexive (Ring.identIsIdent R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) (transitive (Ring.*Commutative R) (Ring.timesZero R)) reflexive) (Group.identLeft (Ring.additiveGroup R)))) reflexive) (Ring.*Commutative R)) (Ring.*WellDefined R (symmetric (Ring.identIsIdent R)) reflexive)
|
||||
Group.invLeft (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((Group.inverse (Ring.additiveGroup R) a * b) + (b * a)) * Ring.1R R) ∼ ((b * b) * Group.0G (Ring.additiveGroup R))
|
||||
need = transitive (transitive (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) reflexive) (transitive (symmetric (Ring.*DistributesOver+ R)) (transitive (Ring.*WellDefined R reflexive (Group.invLeft (Ring.additiveGroup R))) (Ring.timesZero R))))) (symmetric (Ring.timesZero R))
|
||||
Group.invRight (Ring.additiveGroup (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {a ,, (b , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((a * b) + (b * Group.inverse (Ring.additiveGroup R) a)) * Ring.1R R) ∼ ((b * b) * Group.0G (Ring.additiveGroup R))
|
||||
need = transitive (transitive (transitive (Ring.*Commutative R) (Ring.identIsIdent R)) (transitive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) reflexive) (transitive (symmetric (Ring.*DistributesOver+ R)) (transitive (Ring.*WellDefined R reflexive (Group.invRight (Ring.additiveGroup R))) (Ring.timesZero R))))) (symmetric (Ring.timesZero R))
|
||||
Ring.*WellDefined (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} {g ,, (h , _)} af=be ch=dg = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : ((a * c) * (f * h)) ∼ ((b * d) * (e * g))
|
||||
need = transitive (Ring.*WellDefined R reflexive (Ring.*Commutative R)) (transitive (Ring.*Associative R) (transitive (Ring.*WellDefined R (symmetric (Ring.*Associative R)) reflexive) (transitive (Ring.*WellDefined R (Ring.*WellDefined R reflexive ch=dg) reflexive) (transitive (Ring.*Commutative R) (transitive (Ring.*Associative R) (transitive (Ring.*WellDefined R (Ring.*Commutative R) reflexive) (transitive (Ring.*WellDefined R af=be reflexive) (transitive (Ring.*Associative R) (transitive (Ring.*WellDefined R (transitive (symmetric (Ring.*Associative R)) (transitive (Ring.*WellDefined R reflexive (Ring.*Commutative R)) (Ring.*Associative R))) reflexive) (symmetric (Ring.*Associative R)))))))))))
|
||||
Ring.1R (fieldOfFractionsRing {R = R} I) = Ring.1R R ,, (Ring.1R R , IntegralDomain.nontrivial I)
|
||||
Ring.groupIsAbelian (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((a * d) + (b * c)) * (d * b)) ∼ ((b * d) * ((c * b) + (d * a)))
|
||||
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (Ring.*Commutative R) (transitive (Group.+WellDefined (Ring.additiveGroup R) (Ring.*Commutative R) (Ring.*Commutative R)) (Ring.groupIsAbelian R)))
|
||||
Ring.*Associative (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : ((a * (c * e)) * ((b * d) * f)) ∼ ((b * (d * f)) * ((a * c) * e))
|
||||
need = transitive (Ring.*WellDefined R (Ring.*Associative R) (symmetric (Ring.*Associative R))) (Ring.*Commutative R)
|
||||
Ring.*Commutative (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : ((a * c) * (d * b)) ∼ ((b * d) * (c * a))
|
||||
need = transitive (Ring.*Commutative R) (Ring.*WellDefined R (Ring.*Commutative R) (Ring.*Commutative R))
|
||||
Ring.*DistributesOver+ (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Equivalence eq
|
||||
inter : b * (a * ((c * f) + (d * e))) ∼ (((a * c) * (b * f)) + ((b * d) * (a * e)))
|
||||
inter = transitive *Associative (transitive *DistributesOver+ (Group.+WellDefined additiveGroup (transitive *Associative (transitive (*WellDefined (transitive (*WellDefined (*Commutative) reflexive) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative))) reflexive) (symmetric *Associative))) (transitive *Associative (transitive (*WellDefined (transitive (symmetric *Associative) (transitive (*WellDefined reflexive *Commutative) *Associative)) reflexive) (symmetric *Associative)))))
|
||||
need : ((a * ((c * f) + (d * e))) * ((b * d) * (b * f))) ∼ ((b * (d * f)) * (((a * c) * (b * f)) + ((b * d) * (a * e))))
|
||||
need = transitive (Ring.*WellDefined R reflexive (Ring.*WellDefined R reflexive (Ring.*Commutative R))) (transitive (Ring.*WellDefined R reflexive (Ring.*Associative R)) (transitive (Ring.*Commutative R) (transitive (Ring.*WellDefined R (Ring.*WellDefined R (symmetric (Ring.*Associative R)) reflexive) reflexive) (transitive (symmetric (Ring.*Associative R)) (Ring.*WellDefined R reflexive inter)))))
|
||||
Ring.identIsIdent (fieldOfFractionsRing {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) {a ,, (b , _)} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : (((Ring.1R R) * a) * b) ∼ (((Ring.1R R * b)) * a)
|
||||
need = transitive (Ring.*WellDefined R (Ring.identIsIdent R) reflexive) (transitive (Ring.*Commutative R) (Ring.*WellDefined R (symmetric (Ring.identIsIdent R)) reflexive))
|
||||
|
||||
embedIntoFieldOfFractions : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) → A → fieldOfFractionsSet I
|
||||
embedIntoFieldOfFractions {R = R} I a = a ,, (Ring.1R R , IntegralDomain.nontrivial I)
|
||||
fieldOfFractions : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} → (I : IntegralDomain R) → Field (fieldOfFractionsRing I)
|
||||
Field.allInvertible (fieldOfFractions {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) (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)))
|
||||
Field.nontrivial (fieldOfFractions {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I) pr = IntegralDomain.nontrivial I (symmetric (transitive (symmetric (Ring.timesZero R)) (transitive (Ring.*Commutative R) (transitive pr (Ring.identIsIdent R)))))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
pr' : (Ring.0R R) * (Ring.1R R) ∼ (Ring.1R R) * (Ring.1R R)
|
||||
pr' = pr
|
||||
|
||||
homIntoFieldOfFractions : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) → RingHom R (fieldOfFractionsRing I) (embedIntoFieldOfFractions I)
|
||||
RingHom.preserves1 (homIntoFieldOfFractions {S = S} I) = Equivalence.reflexive (Setoid.eq S)
|
||||
RingHom.ringHom (homIntoFieldOfFractions {S = S} {R = R} I) {a} {b} = Equivalence.transitive (Setoid.eq S) (Ring.*WellDefined R (Equivalence.reflexive (Setoid.eq S)) (Ring.identIsIdent R)) (Ring.*Commutative R)
|
||||
GroupHom.groupHom (RingHom.groupHom (homIntoFieldOfFractions {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {x} {y} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : ((x + y) * (Ring.1R R * Ring.1R R)) ∼ (Ring.1R R * ((x * Ring.1R R) + (Ring.1R R * y)))
|
||||
need = transitive (transitive (Ring.*WellDefined R reflexive (Ring.identIsIdent R)) (transitive (Ring.*Commutative R) (transitive (Ring.identIsIdent R) (Group.+WellDefined (Ring.additiveGroup R) (symmetric (transitive (Ring.*Commutative R) (Ring.identIsIdent R))) (symmetric (Ring.identIsIdent R)))))) (symmetric (Ring.identIsIdent R))
|
||||
GroupHom.wellDefined (RingHom.groupHom (homIntoFieldOfFractions {S = S} {R = R} I)) x=y = transitive (Ring.*Commutative R) (Ring.*WellDefined R reflexive x=y)
|
||||
where
|
||||
open Equivalence (Setoid.eq S)
|
||||
embedIntoFieldOfFractions : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) → A → fieldOfFractionsSet I
|
||||
embedIntoFieldOfFractions {R = R} I a = a ,, (Ring.1R R , IntegralDomain.nontrivial I)
|
||||
|
||||
homIntoFieldOfFractionsIsInj : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) → SetoidInjection S (fieldOfFractionsSetoid I) (embedIntoFieldOfFractions I)
|
||||
SetoidInjection.wellDefined (homIntoFieldOfFractionsIsInj {S = S} {R = R} I) x=y = transitive (Ring.*Commutative R) (Ring.*WellDefined R reflexive x=y)
|
||||
where
|
||||
open Equivalence (Setoid.eq S)
|
||||
SetoidInjection.injective (homIntoFieldOfFractionsIsInj {S = S} {R = R} I) x~y = transitive (symmetric identIsIdent) (transitive *Commutative (transitive x~y identIsIdent))
|
||||
where
|
||||
open Ring R
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
homIntoFieldOfFractions : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) → RingHom R (fieldOfFractionsRing I) (embedIntoFieldOfFractions I)
|
||||
RingHom.preserves1 (homIntoFieldOfFractions {S = S} I) = Equivalence.reflexive (Setoid.eq S)
|
||||
RingHom.ringHom (homIntoFieldOfFractions {S = S} {R = R} I) {a} {b} = Equivalence.transitive (Setoid.eq S) (Ring.*WellDefined R (Equivalence.reflexive (Setoid.eq S)) (Ring.identIsIdent R)) (Ring.*Commutative R)
|
||||
GroupHom.groupHom (RingHom.groupHom (homIntoFieldOfFractions {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I)) {x} {y} = need
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
need : ((x + y) * (Ring.1R R * Ring.1R R)) ∼ (Ring.1R R * ((x * Ring.1R R) + (Ring.1R R * y)))
|
||||
need = transitive (transitive (Ring.*WellDefined R reflexive (Ring.identIsIdent R)) (transitive (Ring.*Commutative R) (transitive (Ring.identIsIdent R) (Group.+WellDefined (Ring.additiveGroup R) (symmetric (transitive (Ring.*Commutative R) (Ring.identIsIdent R))) (symmetric (Ring.identIsIdent R)))))) (symmetric (Ring.identIsIdent R))
|
||||
GroupHom.wellDefined (RingHom.groupHom (homIntoFieldOfFractions {S = S} {R = R} I)) x=y = transitive (Ring.*Commutative R) (Ring.*WellDefined R reflexive x=y)
|
||||
where
|
||||
open Equivalence (Setoid.eq S)
|
||||
|
||||
homIntoFieldOfFractionsIsInj : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) → SetoidInjection S (fieldOfFractionsSetoid I) (embedIntoFieldOfFractions I)
|
||||
SetoidInjection.wellDefined (homIntoFieldOfFractionsIsInj {S = S} {R = R} I) x=y = transitive (Ring.*Commutative R) (Ring.*WellDefined R reflexive x=y)
|
||||
where
|
||||
open Equivalence (Setoid.eq S)
|
||||
SetoidInjection.injective (homIntoFieldOfFractionsIsInj {S = S} {R = R} I) x~y = transitive (symmetric identIsIdent) (transitive *Commutative (transitive x~y identIsIdent))
|
||||
where
|
||||
open Ring R
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
@@ -16,73 +16,74 @@ open import Sets.EquivalenceRelations
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.Fields where
|
||||
record Field {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} (R : Ring S _+_ _*_) : Set (lsuc m ⊔ n) where
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
|
||||
record Field {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} (R : Ring S _+_ _*_) : Set (lsuc m ⊔ n) where
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open Setoid S
|
||||
field
|
||||
allInvertible : (a : A) → ((a ∼ Group.0G (Ring.additiveGroup R)) → False) → Sg A (λ t → t * a ∼ 1R)
|
||||
nontrivial : (0R ∼ 1R) → False
|
||||
|
||||
orderedFieldIsIntDom : {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {_} {c} A} {R : Ring S _+_ _*_} {pOrder : SetoidPartialOrder S _<_} {tOrder : SetoidTotalOrder pOrder} (O : OrderedRing R tOrder) (F : Field R) → IntegralDomain R
|
||||
IntegralDomain.intDom (orderedFieldIsIntDom {S = S} {_*_ = _*_} {R = R} {tOrder = tOrder} O F) {a} {b} ab=0 with SetoidTotalOrder.totality tOrder (Ring.0R R) a
|
||||
IntegralDomain.intDom (orderedFieldIsIntDom {A = A} {S = S} {_*_ = _*_} {R = R} {pOrder = pOrder} {tOrder = tOrder} O F) {a} {b} ab=0 | inl (inl x) = inr (transitive (transitive (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Setoid S
|
||||
field
|
||||
allInvertible : (a : A) → ((a ∼ Group.0G (Ring.additiveGroup R)) → False) → Sg A (λ t → t * a ∼ 1R)
|
||||
nontrivial : (0R ∼ 1R) → False
|
||||
open Equivalence eq
|
||||
open Ring R
|
||||
a!=0 : (a ∼ Group.0G additiveGroup) → False
|
||||
a!=0 pr = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<WellDefined pOrder (symmetric pr) reflexive x)
|
||||
invA : A
|
||||
invA = underlying (Field.allInvertible F a a!=0)
|
||||
q : 1R ∼ (invA * a)
|
||||
q with Field.allInvertible F a a!=0
|
||||
... | invA , pr = symmetric pr
|
||||
p : invA * (a * b) ∼ invA * Group.0G additiveGroup
|
||||
p = *WellDefined reflexive ab=0
|
||||
p' : (invA * a) * b ∼ Group.0G additiveGroup
|
||||
p' = transitive (symmetric *Associative) (transitive p (Ring.timesZero R))
|
||||
IntegralDomain.intDom (orderedFieldIsIntDom {A = A} {S = S} {_*_ = _*_} {R = R} {pOrder = pOrder} {tOrder = tOrder} O F) {a} {b} ab=0 | inl (inr x) = inr (transitive (transitive (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Ring R
|
||||
a!=0 : (a ∼ Group.0G additiveGroup) → False
|
||||
a!=0 pr = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<WellDefined pOrder reflexive (symmetric pr) x)
|
||||
invA : A
|
||||
invA = underlying (Field.allInvertible F a a!=0)
|
||||
q : 1R ∼ (invA * a)
|
||||
q with Field.allInvertible F a a!=0
|
||||
... | invA , pr = symmetric pr
|
||||
p : invA * (a * b) ∼ invA * Group.0G additiveGroup
|
||||
p = *WellDefined reflexive ab=0
|
||||
p' : (invA * a) * b ∼ Group.0G additiveGroup
|
||||
p' = transitive (symmetric *Associative) (transitive p (Ring.timesZero R))
|
||||
IntegralDomain.intDom (orderedFieldIsIntDom {S = S} {_*_ = _*_} {R = R} {tOrder = tOrder} O F) {a} {b} ab=0 | inr x = inl (Equivalence.symmetric (Setoid.eq S) x)
|
||||
IntegralDomain.nontrivial (orderedFieldIsIntDom {S = S} O F) pr = Field.nontrivial F (Equivalence.symmetric (Setoid.eq S) pr)
|
||||
|
||||
orderedFieldIsIntDom : {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {_} {c} A} {R : Ring S _+_ _*_} {pOrder : SetoidPartialOrder S _<_} {tOrder : SetoidTotalOrder pOrder} (O : OrderedRing R tOrder) (F : Field R) → IntegralDomain R
|
||||
IntegralDomain.intDom (orderedFieldIsIntDom {S = S} {_*_ = _*_} {R = R} {tOrder = tOrder} O F) {a} {b} ab=0 with SetoidTotalOrder.totality tOrder (Ring.0R R) a
|
||||
IntegralDomain.intDom (orderedFieldIsIntDom {A = A} {S = S} {_*_ = _*_} {R = R} {pOrder = pOrder} {tOrder = tOrder} O F) {a} {b} ab=0 | inl (inl x) = inr (transitive (transitive (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Ring R
|
||||
a!=0 : (a ∼ Group.0G additiveGroup) → False
|
||||
a!=0 pr = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<WellDefined pOrder (symmetric pr) reflexive x)
|
||||
invA : A
|
||||
invA = underlying (Field.allInvertible F a a!=0)
|
||||
q : 1R ∼ (invA * a)
|
||||
q with Field.allInvertible F a a!=0
|
||||
... | invA , pr = symmetric pr
|
||||
p : invA * (a * b) ∼ invA * Group.0G additiveGroup
|
||||
p = *WellDefined reflexive ab=0
|
||||
p' : (invA * a) * b ∼ Group.0G additiveGroup
|
||||
p' = transitive (symmetric *Associative) (transitive p (Ring.timesZero R))
|
||||
IntegralDomain.intDom (orderedFieldIsIntDom {A = A} {S = S} {_*_ = _*_} {R = R} {pOrder = pOrder} {tOrder = tOrder} O F) {a} {b} ab=0 | inl (inr x) = inr (transitive (transitive (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Ring R
|
||||
a!=0 : (a ∼ Group.0G additiveGroup) → False
|
||||
a!=0 pr = SetoidPartialOrder.irreflexive pOrder (SetoidPartialOrder.<WellDefined pOrder reflexive (symmetric pr) x)
|
||||
invA : A
|
||||
invA = underlying (Field.allInvertible F a a!=0)
|
||||
q : 1R ∼ (invA * a)
|
||||
q with Field.allInvertible F a a!=0
|
||||
... | invA , pr = symmetric pr
|
||||
p : invA * (a * b) ∼ invA * Group.0G additiveGroup
|
||||
p = *WellDefined reflexive ab=0
|
||||
p' : (invA * a) * b ∼ Group.0G additiveGroup
|
||||
p' = transitive (symmetric *Associative) (transitive p (Ring.timesZero R))
|
||||
IntegralDomain.intDom (orderedFieldIsIntDom {S = S} {_*_ = _*_} {R = R} {tOrder = tOrder} O F) {a} {b} ab=0 | inr x = inl (Equivalence.symmetric (Setoid.eq S) x)
|
||||
IntegralDomain.nontrivial (orderedFieldIsIntDom {S = S} O F) pr = Field.nontrivial F (Equivalence.symmetric (Setoid.eq S) pr)
|
||||
record Field' {m n : _} : Set (lsuc m ⊔ lsuc n) where
|
||||
field
|
||||
A : Set m
|
||||
S : Setoid {m} {n} A
|
||||
_+_ : A → A → A
|
||||
_*_ : A → A → A
|
||||
R : Ring S _+_ _*_
|
||||
isField : Field R
|
||||
|
||||
record Field' {m n : _} : Set (lsuc m ⊔ lsuc n) where
|
||||
field
|
||||
A : Set m
|
||||
S : Setoid {m} {n} A
|
||||
_+_ : A → A → A
|
||||
_*_ : A → A → A
|
||||
R : Ring S _+_ _*_
|
||||
isField : Field R
|
||||
|
||||
encapsulateField : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (F : Field R) → Field'
|
||||
encapsulateField {A = A} {S = S} {_+_} {_*_} {R} F = record { A = A ; S = S ; _+_ = _+_ ; _*_ = _*_ ; R = R ; isField = F }
|
||||
encapsulateField : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (F : Field R) → Field'
|
||||
encapsulateField {A = A} {S = S} {_+_} {_*_} {R} F = record { A = A ; S = S ; _+_ = _+_ ; _*_ = _*_ ; R = R ; isField = F }
|
||||
|
||||
|
||||
{-
|
||||
record OrderedField {n} {A : Set n} {R : Ring A} (F : Field R) : Set (lsuc n) where
|
||||
open Field F
|
||||
field
|
||||
ord : TotalOrder A
|
||||
open TotalOrder ord
|
||||
open Ring R
|
||||
field
|
||||
productPos : {a b : A} → (0R < a) → (0R < b) → (0R < (a * b))
|
||||
orderRespectsAddition : {a b c : A} → (a < b) → (a + c) < (b + c)
|
||||
record OrderedField {n} {A : Set n} {R : Ring A} (F : Field R) : Set (lsuc n) where
|
||||
open Field F
|
||||
field
|
||||
ord : TotalOrder A
|
||||
open TotalOrder ord
|
||||
open Ring R
|
||||
field
|
||||
productPos : {a b : A} → (0R < a) → (0R < b) → (0R < (a * b))
|
||||
orderRespectsAddition : {a b c : A} → (a < b) → (a + c) < (b + c)
|
||||
|
||||
-}
|
||||
|
32
Fields/Lemmas.agda
Normal file
32
Fields/Lemmas.agda
Normal file
@@ -0,0 +1,32 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
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 Setoids.Orders
|
||||
open import Functions
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Naturals
|
||||
|
||||
module Fields.Lemmas {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (F : Field R) where
|
||||
|
||||
open Setoid S
|
||||
open Field F
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
|
||||
halve : (charNot2 : ((1R + 1R) ∼ 0R) → False) → (a : A) → Sg A (λ i → i + i ∼ a)
|
||||
halve charNot2 a with allInvertible (1R + 1R) charNot2
|
||||
... | 1/2 , pr1/2 = (a * 1/2) , Equivalence.transitive eq (+WellDefined *Commutative *Commutative) (Equivalence.transitive eq (Equivalence.symmetric eq (*DistributesOver+ {1/2} {a} {a})) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) r) (Equivalence.transitive eq (*Associative) (Equivalence.transitive eq (*WellDefined pr1/2 (Equivalence.reflexive eq)) identIsIdent))))
|
||||
where
|
||||
r : a + a ∼ (1R + 1R) * a
|
||||
r = Equivalence.symmetric eq (Equivalence.transitive eq *Commutative (Equivalence.transitive eq *DistributesOver+ (+WellDefined (Equivalence.transitive eq *Commutative identIsIdent) (Equivalence.transitive eq *Commutative identIsIdent))))
|
||||
|
||||
halfHalves : {x : A} (1/2 : A) (pr : 1/2 + 1/2 ∼ 1R) → (x + x) * 1/2 ∼ x
|
||||
halfHalves {x} 1/2 pr = Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.transitive eq *Commutative (Equivalence.transitive eq (Equivalence.transitive eq *DistributesOver+ (Equivalence.transitive eq (+WellDefined *Commutative *Commutative) (Equivalence.symmetric eq *DistributesOver+))) *Commutative)) (*WellDefined pr (Equivalence.reflexive eq))) identIsIdent
|
26
Fields/Orders/Definition.agda
Normal file
26
Fields/Orders/Definition.agda
Normal file
@@ -0,0 +1,26 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Rings.Order
|
||||
open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Orders
|
||||
open import Rings.IntegralDomains
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Fields.Fields
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.Orders.Definition {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (F : Field R) where
|
||||
|
||||
open Ring R
|
||||
open import Fields.Lemmas F
|
||||
|
||||
record OrderedField {p} {_<_ : Rel {_} {p} A} {pOrder : SetoidPartialOrder S _<_} (order : SetoidTotalOrder pOrder) : Set (lsuc (m ⊔ n ⊔ p)) where
|
||||
field
|
||||
oRing : OrderedRing R order
|
48
Fields/Orders/Lemmas.agda
Normal file
48
Fields/Orders/Lemmas.agda
Normal file
@@ -0,0 +1,48 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Rings.Order
|
||||
open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Orders
|
||||
open import Rings.IntegralDomains
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Fields.Fields
|
||||
open import Fields.Orders.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.Orders.Lemmas {m n o : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {_<_ : Rel {_} {o} A} {R : Ring S _+_ _*_} {pOrder : SetoidPartialOrder S _<_} {tOrder : SetoidTotalOrder pOrder} {F : Field R} (oF : OrderedField F tOrder) where
|
||||
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open OrderedRing (OrderedField.oRing oF)
|
||||
open import Rings.Orders.Lemmas (OrderedField.oRing oF)
|
||||
open import Fields.Lemmas F
|
||||
open Setoid S
|
||||
open SetoidPartialOrder pOrder
|
||||
|
||||
clearDenominatorHalf : (x y 1/2 : A) → (1/2 + 1/2 ∼ 1R) → x < (y * 1/2) → (x + x) < y
|
||||
clearDenominatorHalf x y 1/2 pr1/2 x<1/2y = <WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (Equivalence.transitive eq (Equivalence.symmetric eq *DistributesOver+) (Equivalence.transitive eq *Commutative (*WellDefined pr1/2 (Equivalence.reflexive eq)))) identIsIdent) (ringAddInequalities x<1/2y x<1/2y)
|
||||
|
||||
clearDenominatorHalf' : (x y 1/2 : A) → (1/2 + 1/2 ∼ 1R) → (x * 1/2) < y → x < (y + y)
|
||||
clearDenominatorHalf' x y 1/2 pr1/2 1/2x<y = <WellDefined (Equivalence.transitive eq (Equivalence.symmetric eq *DistributesOver+) (Equivalence.transitive eq (Equivalence.transitive eq *Commutative (*WellDefined pr1/2 (Equivalence.reflexive eq))) identIsIdent)) (Equivalence.reflexive eq) (ringAddInequalities 1/2x<y 1/2x<y)
|
||||
|
||||
halveInequality : (x y 1/2 : A) → (1/2 + 1/2 ∼ 1R) → (x + x) < y → x < (y * 1/2)
|
||||
halveInequality x y 1/2 pr1/2 2x<y with SetoidTotalOrder.totality tOrder 0R 1R
|
||||
... | inl (inl 0<1') = <WellDefined (halfHalves 1/2 pr1/2) (Equivalence.reflexive eq) (ringCanMultiplyByPositive {_} {_} {1/2} (halvePositive 1/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq pr1/2) (0<1 λ bad → irreflexive {0R} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq bad) 0<1')))) 2x<y)
|
||||
... | inl (inr 1<0) = <WellDefined (halfHalves 1/2 pr1/2) (Equivalence.reflexive eq) (ringCanMultiplyByPositive {_} {_} {1/2} (halvePositive 1/2 (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq pr1/2) (0<1 λ bad → irreflexive {0R} (<WellDefined (Equivalence.symmetric eq bad) (Equivalence.reflexive eq) 1<0)))) 2x<y)
|
||||
... | inr 0=1 = exFalso (irreflexive {0R} (<WellDefined (oneZeroImpliesAllZero R 0=1) (oneZeroImpliesAllZero R 0=1) 2x<y))
|
||||
|
||||
halveInequality' : (x y 1/2 : A) → (1/2 + 1/2 ∼ 1R) → x < (y + y) → (x * 1/2) < y
|
||||
halveInequality' x y 1/2 pr1/2 x<2y with halveInequality (inverse y) (inverse x) 1/2 pr1/2 (<WellDefined (invContravariant additiveGroup) (Equivalence.reflexive eq) (ringSwapNegatives' x<2y))
|
||||
... | bl = ringSwapNegatives (<WellDefined (Equivalence.reflexive eq) (ringMinusExtracts' R) bl)
|
||||
|
||||
dense : (charNot2 : ((1R + 1R) ∼ 0R) → False) {x y : A} → (x < y) → Sg A (λ i → (x < i) && (i < y))
|
||||
dense charNot2 {x} {y} x<y with halve charNot2 1R
|
||||
dense charNot2 {x} {y} x<y | 1/2 , pr1/2 = ((x + y) * 1/2) , (halveInequality x (x + y) 1/2 pr1/2 (<WellDefined (Equivalence.reflexive eq) groupIsAbelian (orderRespectsAddition x<y x)) ,, halveInequality' (x + y) y 1/2 pr1/2 (orderRespectsAddition x<y y))
|
@@ -11,422 +11,422 @@ open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.Groups where
|
||||
|
||||
reflGroupWellDefined : {lvl : _} {A : Set lvl} {m n x y : A} {op : A → A → A} → m ≡ x → n ≡ y → (op m n) ≡ (op x y)
|
||||
reflGroupWellDefined {lvl} {A} {m} {n} {.m} {.n} {op} refl refl = refl
|
||||
reflGroupWellDefined : {lvl : _} {A : Set lvl} {m n x y : A} {op : A → A → A} → m ≡ x → n ≡ y → (op m n) ≡ (op x y)
|
||||
reflGroupWellDefined {lvl} {A} {m} {n} {.m} {.n} {op} refl refl = refl
|
||||
|
||||
record AbelianGroup {a} {b} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A → A → A} (G : Group S _·_) : Set (lsuc a ⊔ b) where
|
||||
open Setoid S
|
||||
field
|
||||
commutative : {a b : A} → (a · b) ∼ (b · a)
|
||||
record AbelianGroup {a} {b} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A → A → A} (G : Group S _·_) : Set (lsuc a ⊔ b) where
|
||||
open Setoid S
|
||||
field
|
||||
commutative : {a b : A} → (a · b) ∼ (b · a)
|
||||
|
||||
groupsHaveLeftCancellation : {a b : _} → {A : Set a} → {_·_ : A → A → A} → {S : Setoid {a} {b} A} → (G : Group S _·_) → (x y z : A) → (Setoid._∼_ S (x · y) (x · z)) → (Setoid._∼_ S y z)
|
||||
groupsHaveLeftCancellation {_·_ = _·_} {S} g x y z pr = o
|
||||
groupsHaveLeftCancellation : {a b : _} → {A : Set a} → {_·_ : A → A → A} → {S : Setoid {a} {b} A} → (G : Group S _·_) → (x y z : A) → (Setoid._∼_ S (x · y) (x · z)) → (Setoid._∼_ S y z)
|
||||
groupsHaveLeftCancellation {_·_ = _·_} {S} g x y z pr = o
|
||||
where
|
||||
open Group g renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
transitive = Equivalence.transitive (Setoid.eq S)
|
||||
symmetric = Equivalence.symmetric (Setoid.eq S)
|
||||
j : ((x ^-1) · x) · y ∼ (x ^-1) · (x · z)
|
||||
j = Equivalence.transitive eq (symmetric (+Associative {x ^-1} {x} {y})) (+WellDefined ~refl pr)
|
||||
k : ((x ^-1) · x) · y ∼ ((x ^-1) · x) · z
|
||||
k = transitive j +Associative
|
||||
l : 0G · y ∼ ((x ^-1) · x) · z
|
||||
l = transitive (+WellDefined (symmetric invLeft) ~refl) k
|
||||
m : 0G · y ∼ 0G · z
|
||||
m = transitive l (+WellDefined invLeft ~refl)
|
||||
n : y ∼ 0G · z
|
||||
n = transitive (symmetric identLeft) m
|
||||
o : y ∼ z
|
||||
o = transitive n identLeft
|
||||
|
||||
groupsHaveRightCancellation : {a b : _} → {A : Set a} → {_·_ : A → A → A} → {S : Setoid {a} {b} A} → (G : Group S _·_) → (x y z : A) → (Setoid._∼_ S (y · x) (z · x)) → (Setoid._∼_ S y z)
|
||||
groupsHaveRightCancellation {_·_ = _·_} {S} g x y z pr = transitive m identRight
|
||||
where
|
||||
open Group g renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
transitive = Equivalence.transitive (Setoid.eq S)
|
||||
symmetric = Equivalence.symmetric (Setoid.eq S)
|
||||
i : (y · x) · (x ^-1) ∼ (z · x) · (x ^-1)
|
||||
i = +WellDefined pr ~refl
|
||||
j : y · (x · (x ^-1)) ∼ (z · x) · (x ^-1)
|
||||
j = transitive +Associative i
|
||||
j' : y · 0G ∼ (z · x) · (x ^-1)
|
||||
j' = transitive (+WellDefined ~refl (symmetric invRight)) j
|
||||
k : y ∼ (z · x) · (x ^-1)
|
||||
k = transitive (symmetric identRight) j'
|
||||
l : y ∼ z · (x · (x ^-1))
|
||||
l = transitive k (symmetric +Associative)
|
||||
m : y ∼ z · 0G
|
||||
m = transitive l (+WellDefined ~refl invRight)
|
||||
|
||||
replaceGroupOp : {l m : _} {A : Set l} {S : Setoid {l} {m} A} {_·_ : A → A → A} → (G : Group S _·_) → {a b c d w x y z : A} → (Setoid._∼_ S a c) → (Setoid._∼_ S b d) → (Setoid._∼_ S w y) → (Setoid._∼_ S x z) → Setoid._∼_ S (a · b) (w · x) → Setoid._∼_ S (c · d) (y · z)
|
||||
replaceGroupOp {S = S} {_·_} G a~c b~d w~y x~z pr = transitive (symmetric (+WellDefined a~c b~d)) (transitive pr (+WellDefined w~y x~z))
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
replaceGroupOpRight : {l m : _} {A : Set l} {S : Setoid {l} {m} A} {_·_ : A → A → A} → (G : Group S _·_) → {a b c x : A} → (Setoid._∼_ S a (b · c)) → (Setoid._∼_ S c x) → (Setoid._∼_ S a (b · x))
|
||||
replaceGroupOpRight {S = S} {_·_} G a~bc c~x = transitive a~bc (+WellDefined reflexive c~x)
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
inverseWellDefined : {l m : _} {A : Set l} {S : Setoid {l} {m} A} {_·_ : A → A → A} (G : Group S _·_) → {x y : A} → (Setoid._∼_ S x y) → Setoid._∼_ S (Group.inverse G x) (Group.inverse G y)
|
||||
inverseWellDefined {S = S} {_·_} G {x} {y} x~y = groupsHaveRightCancellation G x (inverse x) (inverse y) q
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
p : inverse x · x ∼ inverse y · y
|
||||
p = transitive invLeft (symmetric invLeft)
|
||||
q : inverse x · x ∼ inverse y · x
|
||||
q = replaceGroupOpRight G {_·_ (inverse x) x} {inverse y} {y} {x} p (symmetric x~y)
|
||||
|
||||
rightInversesAreUnique : {a b : _} → {A : Set a} → {S : Setoid {a} {b} A} → {_·_ : A → A → A} → (G : Group S _·_) → (x : A) → (y : A) → Setoid._∼_ S (y · x) (Group.0G G) → Setoid._∼_ S y (Group.inverse G x)
|
||||
rightInversesAreUnique {S = S} {_·_} G x y f = transitive i (transitive j (transitive k (transitive l m)))
|
||||
where
|
||||
open Group G renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
i : y ∼ y · 0G
|
||||
j : y · 0G ∼ y · (x · (x ^-1))
|
||||
k : y · (x · (x ^-1)) ∼ (y · x) · (x ^-1)
|
||||
l : (y · x) · (x ^-1) ∼ 0G · (x ^-1)
|
||||
m : 0G · (x ^-1) ∼ x ^-1
|
||||
i = symmetric identRight
|
||||
j = +WellDefined ~refl (symmetric invRight)
|
||||
k = +Associative
|
||||
l = +WellDefined f ~refl
|
||||
m = identLeft
|
||||
|
||||
leftInversesAreUnique : {a b : _} → {A : Set a} → {S : Setoid {a} {b} A} → {_·_ : A → A → A} → (G : Group S _·_) → {x : A} → {y : A} → Setoid._∼_ S (x · y) (Group.0G G) → Setoid._∼_ S y (Group.inverse G x)
|
||||
leftInversesAreUnique {S = S} {_·_} G {x} {y} f = rightInversesAreUnique G x y l
|
||||
where
|
||||
open Group G renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
i : y · (x · y) ∼ y · 0G
|
||||
i' : y · (x · y) ∼ y
|
||||
j : (y · x) · y ∼ y
|
||||
k : (y · x) · y ∼ 0G · y
|
||||
l : y · x ∼ 0G
|
||||
i = +WellDefined ~refl f
|
||||
i' = transitive i identRight
|
||||
j = transitive (symmetric +Associative) i'
|
||||
k = transitive j (symmetric identLeft)
|
||||
l = groupsHaveRightCancellation G y (y · x) 0G k
|
||||
|
||||
invTwice : {a b : _} → {A : Set a} → {S : Setoid {a} {b} A} → {_·_ : A → A → A} → (G : Group S _·_) → (x : A) → Setoid._∼_ S (Group.inverse G (Group.inverse G x)) x
|
||||
invTwice {S = S} {_·_} G x = symmetric (rightInversesAreUnique G (x ^-1) x invRight)
|
||||
where
|
||||
open Group G renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
fourWay+Associative : {a b : _} → {A : Set a} → {S : Setoid {a} {b} A} → {_·_ : A → A → A} → (G : Group S _·_) → {r s t u : A} → (Setoid._∼_ S) (r · ((s · t) · u)) ((r · s) · (t · u))
|
||||
fourWay+Associative {S = S} {_·_} G {r} {s} {t} {u} = transitive p1 (transitive p2 p3)
|
||||
where
|
||||
open Group G renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
p1 : r · ((s · t) · u) ∼ (r · (s · t)) · u
|
||||
p2 : (r · (s · t)) · u ∼ ((r · s) · t) · u
|
||||
p3 : ((r · s) · t) · u ∼ (r · s) · (t · u)
|
||||
p1 = Group.+Associative G
|
||||
p2 = Group.+WellDefined G (Group.+Associative G) reflexive
|
||||
p3 = symmetric (Group.+Associative G)
|
||||
|
||||
fourWay+Associative' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} (G : Group S _·_) {a b c d : A} → (Setoid._∼_ S (((a · b) · c) · d) (a · ((b · c) · d)))
|
||||
fourWay+Associative' {S = S} G = transitive (symmetric +Associative) (symmetric (fourWay+Associative G))
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
fourWay+Associative'' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} (G : Group S _·_) {a b c d : A} → (Setoid._∼_ S (a · (b · (c · d))) (a · ((b · c) · d)))
|
||||
fourWay+Associative'' {S = S} {_·_ = _·_} G = transitive +Associative (symmetric (fourWay+Associative G))
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
invContravariant : {a b : _} → {A : Set a} → {S : Setoid {a} {b} A} → {_·_ : A → A → A} → (G : Group S _·_) → {x y : A} → (Setoid._∼_ S (Group.inverse G (x · y)) ((Group.inverse G y) · (Group.inverse G x)))
|
||||
invContravariant {S = S} {_·_} G {x} {y} = ans
|
||||
where
|
||||
open Group G renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
otherInv = (y ^-1) · (x ^-1)
|
||||
many+Associatives : x · ((y · (y ^-1)) · (x ^-1)) ∼ (x · y) · ((y ^-1) · (x ^-1))
|
||||
oneMult : (x · y) · otherInv ∼ x · (x ^-1)
|
||||
many+Associatives = fourWay+Associative G
|
||||
oneMult = symmetric (transitive (Group.+WellDefined G reflexive (transitive (symmetric (Group.identLeft G)) (Group.+WellDefined G (symmetric (Group.invRight G)) reflexive))) many+Associatives)
|
||||
otherInvIsInverse : (x · y) · otherInv ∼ 0G
|
||||
otherInvIsInverse = transitive oneMult (Group.invRight G)
|
||||
ans : (x · y) ^-1 ∼ (y ^-1) · (x ^-1)
|
||||
ans = symmetric (leftInversesAreUnique G otherInvIsInverse)
|
||||
|
||||
invIdentity : {l m : _} → {A : Set l} → {S : Setoid {l} {m} A} → {_·_ : A → A → A} → (G : Group S _·_) → Setoid._∼_ S ((Group.inverse G) (Group.0G G)) (Group.0G G)
|
||||
invIdentity {S = S} G = transitive (symmetric identLeft) invRight
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
transferToRight : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} → (G : Group S _·_) → {a b : A} → Setoid._∼_ S (a · (Group.inverse G b)) (Group.0G G) → Setoid._∼_ S a b
|
||||
transferToRight {S = S} {_·_ = _·_} G {a} {b} ab-1 = transitive (symmetric (invTwice G a)) (transitive u (invTwice G b))
|
||||
where
|
||||
open Setoid S
|
||||
open Group G
|
||||
open Equivalence eq
|
||||
t : inverse a ∼ inverse b
|
||||
t = symmetric (leftInversesAreUnique G ab-1)
|
||||
u : inverse (inverse a) ∼ inverse (inverse b)
|
||||
u = inverseWellDefined G t
|
||||
|
||||
transferToRight' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} → (G : Group S _·_) → {a b : A} → Setoid._∼_ S (a · b) (Group.0G G) → Setoid._∼_ S a (Group.inverse G b)
|
||||
transferToRight' {S = S} {_·_ = _·_} G {a} {b} ab-1 = transferToRight G lemma
|
||||
where
|
||||
open Setoid S
|
||||
open Group G
|
||||
open Equivalence eq
|
||||
lemma : a · (inverse (inverse b)) ∼ 0G
|
||||
lemma = transitive (+WellDefined reflexive (invTwice G b)) ab-1
|
||||
|
||||
transferToRight'' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} → (G : Group S _·_) → {a b : A} → Setoid._∼_ S a b → Setoid._∼_ S (a · (Group.inverse G b)) (Group.0G G)
|
||||
transferToRight'' {S = S} {_·_ = _·_} G {a} {b} a~b = transitive (Group.+WellDefined G a~b reflexive) invRight
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
directSumGroup : {m n o p : _} → {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} (G : Group S _·A_) (h : Group T _·B_) → Group (directSumSetoid S T) (λ x1 y1 → (((_&&_.fst x1) ·A (_&&_.fst y1)) ,, ((_&&_.snd x1) ·B (_&&_.snd y1))))
|
||||
Group.+WellDefined (directSumGroup {A = A} {B} g h) (s ,, t) (u ,, v) = Group.+WellDefined g s u ,, Group.+WellDefined h t v
|
||||
Group.0G (directSumGroup {A = A} {B} g h) = (Group.0G g ,, Group.0G h)
|
||||
Group.inverse (directSumGroup {A = A} {B} g h) (g1 ,, h1) = (Group.inverse g g1) ,, (Group.inverse h h1)
|
||||
Group.+Associative (directSumGroup {A = A} {B} g h) = Group.+Associative g ,, Group.+Associative h
|
||||
Group.identRight (directSumGroup {A = A} {B} g h) = Group.identRight g ,, Group.identRight h
|
||||
Group.identLeft (directSumGroup {A = A} {B} g h) = Group.identLeft g ,, Group.identLeft h
|
||||
Group.invLeft (directSumGroup {A = A} {B} g h) = Group.invLeft g ,, Group.invLeft h
|
||||
Group.invRight (directSumGroup {A = A} {B} g h) = Group.invRight g ,, Group.invRight h
|
||||
|
||||
directSumAbelianGroup : {m n o p : _} → {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} {underG : Group S _·A_} {underH : Group T _·B_} (G : AbelianGroup underG) (h : AbelianGroup underH) → (AbelianGroup (directSumGroup underG underH))
|
||||
AbelianGroup.commutative (directSumAbelianGroup {A = A} {B} G H) = AbelianGroup.commutative G ,, AbelianGroup.commutative H
|
||||
|
||||
record GroupHom {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} (G : Group S _·A_) (H : Group T _·B_) (f : A → B) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
open Group H
|
||||
open Setoid T
|
||||
field
|
||||
groupHom : {x y : A} → f (x ·A y) ∼ (f x) ·B (f y)
|
||||
wellDefined : {x y : A} → Setoid._∼_ S x y → f x ∼ f y
|
||||
|
||||
record InjectiveGroupHom {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_} {underf : A → B} (f : GroupHom G H underf) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
open Setoid S renaming (_∼_ to _∼A_)
|
||||
open Setoid T renaming (_∼_ to _∼B_)
|
||||
field
|
||||
injective : SetoidInjection S T underf
|
||||
|
||||
imageOfIdentityIsIdentity : {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {B : Set n} {T : Setoid {n} {p} B} {_·A_ : A → A → A} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_} {f : A → B} → (hom : GroupHom G H f) → Setoid._∼_ T (f (Group.0G G)) (Group.0G H)
|
||||
imageOfIdentityIsIdentity {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {f = f} hom = Equivalence.symmetric (Setoid.eq T) t
|
||||
where
|
||||
open Group H
|
||||
open Setoid T
|
||||
id2 : Setoid._∼_ S (Group.0G G) ((Group.0G G) ·A (Group.0G G))
|
||||
id2 = Equivalence.symmetric (Setoid.eq S) (Group.identRight G)
|
||||
r : f (Group.0G G) ∼ f (Group.0G G) ·B f (Group.0G G)
|
||||
s : 0G ·B f (Group.0G G) ∼ f (Group.0G G) ·B f (Group.0G G)
|
||||
t : 0G ∼ f (Group.0G G)
|
||||
t = groupsHaveRightCancellation H (f (Group.0G G)) 0G (f (Group.0G G)) s
|
||||
s = Equivalence.transitive (Setoid.eq T) identLeft r
|
||||
r = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined hom id2) (GroupHom.groupHom hom)
|
||||
|
||||
groupHomsCompose : {m n o r s t : _} {A : Set m} {S : Setoid {m} {r} A} {_+A_ : A → A → A} {B : Set n} {T : Setoid {n} {s} B} {_+B_ : B → B → B} {C : Set o} {U : Setoid {o} {t} C} {_+C_ : C → C → C} {G : Group S _+A_} {H : Group T _+B_} {I : Group U _+C_} {f : A → B} {g : B → C} (fHom : GroupHom G H f) (gHom : GroupHom H I g) → GroupHom G I (g ∘ f)
|
||||
GroupHom.wellDefined (groupHomsCompose {G = G} {H} {I} {f} {g} fHom gHom) {x} {y} pr = GroupHom.wellDefined gHom (GroupHom.wellDefined fHom pr)
|
||||
GroupHom.groupHom (groupHomsCompose {S = S} {_+A_ = _·A_} {T = T} {U = U} {_+C_ = _·C_} {G = G} {H} {I} {f} {g} fHom gHom) {x} {y} = answer
|
||||
where
|
||||
open Group I
|
||||
answer : (Setoid._∼_ U) ((g ∘ f) (x ·A y)) ((g ∘ f) x ·C (g ∘ f) y)
|
||||
answer = (Equivalence.transitive (Setoid.eq U)) (GroupHom.wellDefined gHom (GroupHom.groupHom fHom {x} {y}) ) (GroupHom.groupHom gHom {f x} {f y})
|
||||
|
||||
record GroupIso {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} (G : Group S _·A_) (H : Group T _·B_) (f : A → B) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
open Setoid S renaming (_∼_ to _∼G_)
|
||||
open Setoid T renaming (_∼_ to _∼H_)
|
||||
field
|
||||
groupHom : GroupHom G H f
|
||||
bij : SetoidBijection S T f
|
||||
|
||||
record GroupsIsomorphic {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} (G : Group S _·A_) (H : Group T _·B_) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
field
|
||||
isomorphism : A → B
|
||||
proof : GroupIso G H isomorphism
|
||||
|
||||
groupIsosCompose : {m n o r s t : _} {A : Set m} {S : Setoid {m} {r} A} {_+A_ : A → A → A} {B : Set n} {T : Setoid {n} {s} B} {_+B_ : B → B → B} {C : Set o} {U : Setoid {o} {t} C} {_+C_ : C → C → C} {G : Group S _+A_} {H : Group T _+B_} {I : Group U _+C_} {f : A → B} {g : B → C} (fHom : GroupIso G H f) (gHom : GroupIso H I g) → GroupIso G I (g ∘ f)
|
||||
GroupIso.groupHom (groupIsosCompose fHom gHom) = groupHomsCompose (GroupIso.groupHom fHom) (GroupIso.groupHom gHom)
|
||||
GroupIso.bij (groupIsosCompose {A = A} {S = S} {T = T} {C = C} {U = U} {f = f} {g = g} fHom gHom) = record { inj = record { injective = λ pr → (SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij fHom))) (SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij gHom)) pr) ; wellDefined = +WellDefined } ; surj = record { surjective = surj ; wellDefined = +WellDefined } }
|
||||
where
|
||||
open Setoid S renaming (_∼_ to _∼A_)
|
||||
open Setoid U renaming (_∼_ to _∼C_)
|
||||
+WellDefined : {x y : A} → (x ∼A y) → (g (f x) ∼C g (f y))
|
||||
+WellDefined = GroupHom.wellDefined (groupHomsCompose (GroupIso.groupHom fHom) (GroupIso.groupHom gHom))
|
||||
surj : {x : C} → Sg A (λ a → (g (f a) ∼C x))
|
||||
surj {x} with SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij gHom)) {x}
|
||||
surj {x} | a , prA with SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij fHom)) {a}
|
||||
... | b , prB = b , transitive (GroupHom.wellDefined (GroupIso.groupHom gHom) prB) prA
|
||||
where
|
||||
open Group g renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
transitive = Equivalence.transitive (Setoid.eq S)
|
||||
symmetric = Equivalence.symmetric (Setoid.eq S)
|
||||
j : ((x ^-1) · x) · y ∼ (x ^-1) · (x · z)
|
||||
j = Equivalence.transitive eq (symmetric (+Associative {x ^-1} {x} {y})) (+WellDefined ~refl pr)
|
||||
k : ((x ^-1) · x) · y ∼ ((x ^-1) · x) · z
|
||||
k = transitive j +Associative
|
||||
l : 0G · y ∼ ((x ^-1) · x) · z
|
||||
l = transitive (+WellDefined (symmetric invLeft) ~refl) k
|
||||
m : 0G · y ∼ 0G · z
|
||||
m = transitive l (+WellDefined invLeft ~refl)
|
||||
n : y ∼ 0G · z
|
||||
n = transitive (symmetric identLeft) m
|
||||
o : y ∼ z
|
||||
o = transitive n identLeft
|
||||
open Equivalence (Setoid.eq U)
|
||||
|
||||
groupsHaveRightCancellation : {a b : _} → {A : Set a} → {_·_ : A → A → A} → {S : Setoid {a} {b} A} → (G : Group S _·_) → (x y z : A) → (Setoid._∼_ S (y · x) (z · x)) → (Setoid._∼_ S y z)
|
||||
groupsHaveRightCancellation {_·_ = _·_} {S} g x y z pr = transitive m identRight
|
||||
record FiniteGroup {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A → A → A} (G : Group S _·_) {c : _} (quotientSet : Set c) : Set (a ⊔ b ⊔ c) where
|
||||
field
|
||||
toSet : SetoidToSet S quotientSet
|
||||
finite : FiniteSet quotientSet
|
||||
|
||||
groupOrder : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A → A → A} {underG : Group S _·_} {c : _} {quotientSet : Set c} (G : FiniteGroup underG quotientSet) → ℕ
|
||||
groupOrder record { toSet = toSet ; finite = record { size = size ; mapping = mapping ; bij = bij } } = size
|
||||
|
||||
--groupIsoInvertible : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d}} {_+A_ : A → A → A} {_+B_ : B → B → B} {G : Group S _+A_} {H : Group T _+B_} {f : A → B} → (iso : GroupIso G H f) → GroupIso H G (Invertible.inverse (bijectionImpliesInvertible (GroupIso.bijective iso)))
|
||||
--GroupHom.groupHom (GroupIso.groupHom (groupIsoInvertible {G = G} {H} {f} iso)) {x} {y} = {!!}
|
||||
-- where
|
||||
-- open Group G renaming (_·_ to _+G_)
|
||||
-- open Group H renaming (_·_ to _+H_)
|
||||
--GroupHom.wellDefined (GroupIso.groupHom (groupIsoInvertible {G = G} {H} {f} iso)) {x} {y} x∼y = {!GroupHom.wellDefined x∼y!}
|
||||
--GroupIso.bijective (groupIsoInvertible {G = G} {H} {f} iso) = invertibleImpliesBijection (inverseIsInvertible (bijectionImpliesInvertible (GroupIso.bijective iso)))
|
||||
|
||||
homRespectsInverse : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_} {underF : A → B} → (f : GroupHom G H underF) → {x : A} → Setoid._∼_ T (underF (Group.inverse G x)) (Group.inverse H (underF x))
|
||||
homRespectsInverse {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {underF = f} fHom {x} = rightInversesAreUnique H (f x) (f (Group.inverse G x)) (transitive (symmetric (GroupHom.groupHom fHom)) (transitive (GroupHom.wellDefined fHom (Group.invLeft G)) (imageOfIdentityIsIdentity fHom)))
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
quotientGroupSetoid : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_} → {underf : A → B} → (f : GroupHom G H underf) → (Setoid {a} {d} A)
|
||||
quotientGroupSetoid {A = A} {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H} {f} fHom = ansSetoid
|
||||
where
|
||||
open Setoid T
|
||||
open Group H
|
||||
open Equivalence eq
|
||||
ansSetoid : Setoid A
|
||||
Setoid._∼_ ansSetoid r s = (f (r ·A (Group.inverse G s))) ∼ 0G
|
||||
Equivalence.reflexive (Setoid.eq ansSetoid) {b} = transitive (GroupHom.wellDefined fHom (Group.invRight G)) (imageOfIdentityIsIdentity fHom)
|
||||
Equivalence.symmetric (Setoid.eq ansSetoid) {m} {n} pr = i
|
||||
where
|
||||
open Group g renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
transitive = Equivalence.transitive (Setoid.eq S)
|
||||
symmetric = Equivalence.symmetric (Setoid.eq S)
|
||||
i : (y · x) · (x ^-1) ∼ (z · x) · (x ^-1)
|
||||
i = +WellDefined pr ~refl
|
||||
j : y · (x · (x ^-1)) ∼ (z · x) · (x ^-1)
|
||||
j = transitive +Associative i
|
||||
j' : y · 0G ∼ (z · x) · (x ^-1)
|
||||
j' = transitive (+WellDefined ~refl (symmetric invRight)) j
|
||||
k : y ∼ (z · x) · (x ^-1)
|
||||
k = transitive (symmetric identRight) j'
|
||||
l : y ∼ z · (x · (x ^-1))
|
||||
l = transitive k (symmetric +Associative)
|
||||
m : y ∼ z · 0G
|
||||
m = transitive l (+WellDefined ~refl invRight)
|
||||
|
||||
replaceGroupOp : {l m : _} {A : Set l} {S : Setoid {l} {m} A} {_·_ : A → A → A} → (G : Group S _·_) → {a b c d w x y z : A} → (Setoid._∼_ S a c) → (Setoid._∼_ S b d) → (Setoid._∼_ S w y) → (Setoid._∼_ S x z) → Setoid._∼_ S (a · b) (w · x) → Setoid._∼_ S (c · d) (y · z)
|
||||
replaceGroupOp {S = S} {_·_} G a~c b~d w~y x~z pr = transitive (symmetric (+WellDefined a~c b~d)) (transitive pr (+WellDefined w~y x~z))
|
||||
g : f (Group.inverse G (m ·A Group.inverse G n)) ∼ 0G
|
||||
g = transitive (homRespectsInverse fHom {m ·A Group.inverse G n}) (transitive (inverseWellDefined H pr) (invIdentity H))
|
||||
h : f (Group.inverse G (Group.inverse G n) ·A Group.inverse G m) ∼ 0G
|
||||
h = transitive (GroupHom.wellDefined fHom (Equivalence.symmetric (Setoid.eq S) (invContravariant G))) g
|
||||
i : f (n ·A Group.inverse G m) ∼ 0G
|
||||
i = transitive (GroupHom.wellDefined fHom (Group.+WellDefined G (Equivalence.symmetric (Setoid.eq S) (invTwice G n)) (Equivalence.reflexive (Setoid.eq S)))) h
|
||||
Equivalence.transitive (Setoid.eq ansSetoid) {m} {n} {o} prmn prno = transitive (GroupHom.wellDefined fHom (Group.+WellDefined G (Equivalence.reflexive (Setoid.eq S)) (Equivalence.symmetric (Setoid.eq S) (Group.identLeft G)))) k
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
g : f (m ·A Group.inverse G n) ·B f (n ·A Group.inverse G o) ∼ 0G ·B 0G
|
||||
g = replaceGroupOp H reflexive reflexive prmn prno reflexive
|
||||
h : f (m ·A Group.inverse G n) ·B f (n ·A Group.inverse G o) ∼ 0G
|
||||
h = transitive g identLeft
|
||||
i : f ((m ·A Group.inverse G n) ·A (n ·A Group.inverse G o)) ∼ 0G
|
||||
i = transitive (GroupHom.groupHom fHom) h
|
||||
j : f (m ·A (((Group.inverse G n) ·A n) ·A Group.inverse G o)) ∼ 0G
|
||||
j = transitive (GroupHom.wellDefined fHom (fourWay+Associative G)) i
|
||||
k : f (m ·A ((Group.0G G) ·A Group.inverse G o)) ∼ 0G
|
||||
k = transitive (GroupHom.wellDefined fHom (Group.+WellDefined G (Equivalence.reflexive (Setoid.eq S)) (Group.+WellDefined G (Equivalence.symmetric (Setoid.eq S) (Group.invLeft G)) (Equivalence.reflexive (Setoid.eq S))))) j
|
||||
|
||||
replaceGroupOpRight : {l m : _} {A : Set l} {S : Setoid {l} {m} A} {_·_ : A → A → A} → (G : Group S _·_) → {a b c x : A} → (Setoid._∼_ S a (b · c)) → (Setoid._∼_ S c x) → (Setoid._∼_ S a (b · x))
|
||||
replaceGroupOpRight {S = S} {_·_} G a~bc c~x = transitive a~bc (+WellDefined reflexive c~x)
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
quotientGroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_} → {underf : A → B} → (f : GroupHom G H underf) → Group (quotientGroupSetoid G f) _·A_
|
||||
Group.+WellDefined (quotientGroup {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H = H} {underf = f} fHom) {x} {y} {m} {n} x~m y~n = ans
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence (Setoid.eq T)
|
||||
p1 : f ((x ·A y) ·A (Group.inverse G (m ·A n))) ∼ f ((x ·A y) ·A ((Group.inverse G n) ·A (Group.inverse G m)))
|
||||
p2 : f ((x ·A y) ·A ((Group.inverse G n) ·A (Group.inverse G m))) ∼ f (x ·A ((y ·A (Group.inverse G n)) ·A (Group.inverse G m)))
|
||||
p3 : f (x ·A ((y ·A (Group.inverse G n)) ·A (Group.inverse G m))) ∼ (f x) ·B f ((y ·A (Group.inverse G n)) ·A (Group.inverse G m))
|
||||
p4 : (f x) ·B f ((y ·A (Group.inverse G n)) ·A (Group.inverse G m)) ∼ (f x) ·B (f (y ·A (Group.inverse G n)) ·B f (Group.inverse G m))
|
||||
p5 : (f x) ·B (f (y ·A (Group.inverse G n)) ·B f (Group.inverse G m)) ∼ (f x) ·B (Group.0G H ·B f (Group.inverse G m))
|
||||
p6 : (f x) ·B (Group.0G H ·B f (Group.inverse G m)) ∼ (f x) ·B f (Group.inverse G m)
|
||||
p7 : (f x) ·B f (Group.inverse G m) ∼ f (x ·A (Group.inverse G m))
|
||||
p8 : f (x ·A (Group.inverse G m)) ∼ Group.0G H
|
||||
p1 = GroupHom.wellDefined fHom (Group.+WellDefined G (Equivalence.reflexive (Setoid.eq S)) (invContravariant G))
|
||||
p2 = GroupHom.wellDefined fHom (Equivalence.symmetric (Setoid.eq S) (fourWay+Associative G))
|
||||
p3 = GroupHom.groupHom fHom
|
||||
p4 = Group.+WellDefined H reflexive (GroupHom.groupHom fHom)
|
||||
p5 = Group.+WellDefined H reflexive (replaceGroupOp H (symmetric y~n) reflexive reflexive reflexive reflexive)
|
||||
p6 = Group.+WellDefined H reflexive (Group.identLeft H)
|
||||
p7 = symmetric (GroupHom.groupHom fHom)
|
||||
p8 = x~m
|
||||
ans : f ((x ·A y) ·A (Group.inverse G (m ·A n))) ∼ Group.0G H
|
||||
ans = transitive p1 (transitive p2 (transitive p3 (transitive p4 (transitive p5 (transitive p6 (transitive p7 p8))))))
|
||||
Group.0G (quotientGroup G fHom) = Group.0G G
|
||||
Group.inverse (quotientGroup G fHom) = Group.inverse G
|
||||
Group.+Associative (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} {b} {c} = ans
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence (Setoid.eq T)
|
||||
ans : f ((a ·A (b ·A c)) ·A (Group.inverse G ((a ·A b) ·A c))) ∼ Group.0G H
|
||||
ans = transitive (GroupHom.wellDefined fHom (transferToRight'' G (Group.+Associative G))) (imageOfIdentityIsIdentity fHom)
|
||||
Group.identRight (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
transitiveG = Equivalence.transitive (Setoid.eq S)
|
||||
ans : f ((a ·A 0G) ·A inverse a) ∼ Group.0G H
|
||||
ans = transitive (GroupHom.wellDefined fHom (transitiveG (Group.+WellDefined G (Group.identRight G) (Equivalence.reflexive (Setoid.eq S))) (Group.invRight G))) (imageOfIdentityIsIdentity fHom)
|
||||
Group.identLeft (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
transitiveG = Equivalence.transitive (Setoid.eq S)
|
||||
ans : f ((0G ·A a) ·A (inverse a)) ∼ Group.0G H
|
||||
ans = transitive (GroupHom.wellDefined fHom (transitiveG (Group.+WellDefined G (Group.identLeft G) (Equivalence.reflexive (Setoid.eq S))) (Group.invRight G))) (imageOfIdentityIsIdentity fHom)
|
||||
Group.invLeft (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
ans : f ((inverse x ·A x) ·A (inverse 0G)) ∼ (Group.0G H)
|
||||
ans = transitive (GroupHom.wellDefined fHom (Equivalence.transitive (Setoid.eq S) (replaceGroupOp G (Equivalence.symmetric (Setoid.eq S) (Group.invLeft G)) (Equivalence.symmetric (Setoid.eq S) (invIdentity G)) (Equivalence.reflexive (Setoid.eq S)) ((Equivalence.reflexive (Setoid.eq S))) ((Equivalence.reflexive (Setoid.eq S)))) (identRight {0G}))) (imageOfIdentityIsIdentity fHom)
|
||||
Group.invRight (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
ans : f ((x ·A inverse x) ·A (inverse 0G)) ∼ (Group.0G H)
|
||||
ans = transitive (GroupHom.wellDefined fHom (Equivalence.transitive (Setoid.eq S) (replaceGroupOp G (Equivalence.symmetric (Setoid.eq S) (Group.invRight G)) (Equivalence.symmetric (Setoid.eq S) (invIdentity G)) (Equivalence.reflexive (Setoid.eq S)) (Equivalence.reflexive (Setoid.eq S)) (Equivalence.reflexive (Setoid.eq S))) (identRight {0G}))) (imageOfIdentityIsIdentity fHom)
|
||||
|
||||
inverseWellDefined : {l m : _} {A : Set l} {S : Setoid {l} {m} A} {_·_ : A → A → A} (G : Group S _·_) → {x y : A} → (Setoid._∼_ S x y) → Setoid._∼_ S (Group.inverse G x) (Group.inverse G y)
|
||||
inverseWellDefined {S = S} {_·_} G {x} {y} x~y = groupsHaveRightCancellation G x (inverse x) (inverse y) q
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
p : inverse x · x ∼ inverse y · y
|
||||
p = transitive invLeft (symmetric invLeft)
|
||||
q : inverse x · x ∼ inverse y · x
|
||||
q = replaceGroupOpRight G {_·_ (inverse x) x} {inverse y} {y} {x} p (symmetric x~y)
|
||||
|
||||
rightInversesAreUnique : {a b : _} → {A : Set a} → {S : Setoid {a} {b} A} → {_·_ : A → A → A} → (G : Group S _·_) → (x : A) → (y : A) → Setoid._∼_ S (y · x) (Group.0G G) → Setoid._∼_ S y (Group.inverse G x)
|
||||
rightInversesAreUnique {S = S} {_·_} G x y f = transitive i (transitive j (transitive k (transitive l m)))
|
||||
where
|
||||
open Group G renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
i : y ∼ y · 0G
|
||||
j : y · 0G ∼ y · (x · (x ^-1))
|
||||
k : y · (x · (x ^-1)) ∼ (y · x) · (x ^-1)
|
||||
l : (y · x) · (x ^-1) ∼ 0G · (x ^-1)
|
||||
m : 0G · (x ^-1) ∼ x ^-1
|
||||
i = symmetric identRight
|
||||
j = +WellDefined ~refl (symmetric invRight)
|
||||
k = +Associative
|
||||
l = +WellDefined f ~refl
|
||||
m = identLeft
|
||||
|
||||
leftInversesAreUnique : {a b : _} → {A : Set a} → {S : Setoid {a} {b} A} → {_·_ : A → A → A} → (G : Group S _·_) → {x : A} → {y : A} → Setoid._∼_ S (x · y) (Group.0G G) → Setoid._∼_ S y (Group.inverse G x)
|
||||
leftInversesAreUnique {S = S} {_·_} G {x} {y} f = rightInversesAreUnique G x y l
|
||||
where
|
||||
open Group G renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
i : y · (x · y) ∼ y · 0G
|
||||
i' : y · (x · y) ∼ y
|
||||
j : (y · x) · y ∼ y
|
||||
k : (y · x) · y ∼ 0G · y
|
||||
l : y · x ∼ 0G
|
||||
i = +WellDefined ~refl f
|
||||
i' = transitive i identRight
|
||||
j = transitive (symmetric +Associative) i'
|
||||
k = transitive j (symmetric identLeft)
|
||||
l = groupsHaveRightCancellation G y (y · x) 0G k
|
||||
|
||||
invTwice : {a b : _} → {A : Set a} → {S : Setoid {a} {b} A} → {_·_ : A → A → A} → (G : Group S _·_) → (x : A) → Setoid._∼_ S (Group.inverse G (Group.inverse G x)) x
|
||||
invTwice {S = S} {_·_} G x = symmetric (rightInversesAreUnique G (x ^-1) x invRight)
|
||||
where
|
||||
open Group G renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
fourWay+Associative : {a b : _} → {A : Set a} → {S : Setoid {a} {b} A} → {_·_ : A → A → A} → (G : Group S _·_) → {r s t u : A} → (Setoid._∼_ S) (r · ((s · t) · u)) ((r · s) · (t · u))
|
||||
fourWay+Associative {S = S} {_·_} G {r} {s} {t} {u} = transitive p1 (transitive p2 p3)
|
||||
where
|
||||
open Group G renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
p1 : r · ((s · t) · u) ∼ (r · (s · t)) · u
|
||||
p2 : (r · (s · t)) · u ∼ ((r · s) · t) · u
|
||||
p3 : ((r · s) · t) · u ∼ (r · s) · (t · u)
|
||||
p1 = Group.+Associative G
|
||||
p2 = Group.+WellDefined G (Group.+Associative G) reflexive
|
||||
p3 = symmetric (Group.+Associative G)
|
||||
|
||||
fourWay+Associative' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} (G : Group S _·_) {a b c d : A} → (Setoid._∼_ S (((a · b) · c) · d) (a · ((b · c) · d)))
|
||||
fourWay+Associative' {S = S} G = transitive (symmetric +Associative) (symmetric (fourWay+Associative G))
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
fourWay+Associative'' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} (G : Group S _·_) {a b c d : A} → (Setoid._∼_ S (a · (b · (c · d))) (a · ((b · c) · d)))
|
||||
fourWay+Associative'' {S = S} {_·_ = _·_} G = transitive +Associative (symmetric (fourWay+Associative G))
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
invContravariant : {a b : _} → {A : Set a} → {S : Setoid {a} {b} A} → {_·_ : A → A → A} → (G : Group S _·_) → {x y : A} → (Setoid._∼_ S (Group.inverse G (x · y)) ((Group.inverse G y) · (Group.inverse G x)))
|
||||
invContravariant {S = S} {_·_} G {x} {y} = ans
|
||||
where
|
||||
open Group G renaming (inverse to _^-1)
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
otherInv = (y ^-1) · (x ^-1)
|
||||
many+Associatives : x · ((y · (y ^-1)) · (x ^-1)) ∼ (x · y) · ((y ^-1) · (x ^-1))
|
||||
oneMult : (x · y) · otherInv ∼ x · (x ^-1)
|
||||
many+Associatives = fourWay+Associative G
|
||||
oneMult = symmetric (transitive (Group.+WellDefined G reflexive (transitive (symmetric (Group.identLeft G)) (Group.+WellDefined G (symmetric (Group.invRight G)) reflexive))) many+Associatives)
|
||||
otherInvIsInverse : (x · y) · otherInv ∼ 0G
|
||||
otherInvIsInverse = transitive oneMult (Group.invRight G)
|
||||
ans : (x · y) ^-1 ∼ (y ^-1) · (x ^-1)
|
||||
ans = symmetric (leftInversesAreUnique G otherInvIsInverse)
|
||||
|
||||
invIdentity : {l m : _} → {A : Set l} → {S : Setoid {l} {m} A} → {_·_ : A → A → A} → (G : Group S _·_) → Setoid._∼_ S ((Group.inverse G) (Group.0G G)) (Group.0G G)
|
||||
invIdentity {S = S} G = transitive (symmetric identLeft) invRight
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
transferToRight : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} → (G : Group S _·_) → {a b : A} → Setoid._∼_ S (a · (Group.inverse G b)) (Group.0G G) → Setoid._∼_ S a b
|
||||
transferToRight {S = S} {_·_ = _·_} G {a} {b} ab-1 = transitive (symmetric (invTwice G a)) (transitive u (invTwice G b))
|
||||
where
|
||||
open Setoid S
|
||||
open Group G
|
||||
open Equivalence eq
|
||||
t : inverse a ∼ inverse b
|
||||
t = symmetric (leftInversesAreUnique G ab-1)
|
||||
u : inverse (inverse a) ∼ inverse (inverse b)
|
||||
u = inverseWellDefined G t
|
||||
|
||||
transferToRight' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} → (G : Group S _·_) → {a b : A} → Setoid._∼_ S (a · b) (Group.0G G) → Setoid._∼_ S a (Group.inverse G b)
|
||||
transferToRight' {S = S} {_·_ = _·_} G {a} {b} ab-1 = transferToRight G lemma
|
||||
where
|
||||
open Setoid S
|
||||
open Group G
|
||||
open Equivalence eq
|
||||
lemma : a · (inverse (inverse b)) ∼ 0G
|
||||
lemma = transitive (+WellDefined reflexive (invTwice G b)) ab-1
|
||||
|
||||
transferToRight'' : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} → (G : Group S _·_) → {a b : A} → Setoid._∼_ S a b → Setoid._∼_ S (a · (Group.inverse G b)) (Group.0G G)
|
||||
transferToRight'' {S = S} {_·_ = _·_} G {a} {b} a~b = transitive (Group.+WellDefined G a~b reflexive) invRight
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
directSumGroup : {m n o p : _} → {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} (G : Group S _·A_) (h : Group T _·B_) → Group (directSumSetoid S T) (λ x1 y1 → (((_&&_.fst x1) ·A (_&&_.fst y1)) ,, ((_&&_.snd x1) ·B (_&&_.snd y1))))
|
||||
Group.+WellDefined (directSumGroup {A = A} {B} g h) (s ,, t) (u ,, v) = Group.+WellDefined g s u ,, Group.+WellDefined h t v
|
||||
Group.0G (directSumGroup {A = A} {B} g h) = (Group.0G g ,, Group.0G h)
|
||||
Group.inverse (directSumGroup {A = A} {B} g h) (g1 ,, h1) = (Group.inverse g g1) ,, (Group.inverse h h1)
|
||||
Group.+Associative (directSumGroup {A = A} {B} g h) = Group.+Associative g ,, Group.+Associative h
|
||||
Group.identRight (directSumGroup {A = A} {B} g h) = Group.identRight g ,, Group.identRight h
|
||||
Group.identLeft (directSumGroup {A = A} {B} g h) = Group.identLeft g ,, Group.identLeft h
|
||||
Group.invLeft (directSumGroup {A = A} {B} g h) = Group.invLeft g ,, Group.invLeft h
|
||||
Group.invRight (directSumGroup {A = A} {B} g h) = Group.invRight g ,, Group.invRight h
|
||||
|
||||
directSumAbelianGroup : {m n o p : _} → {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} {underG : Group S _·A_} {underH : Group T _·B_} (G : AbelianGroup underG) (h : AbelianGroup underH) → (AbelianGroup (directSumGroup underG underH))
|
||||
AbelianGroup.commutative (directSumAbelianGroup {A = A} {B} G H) = AbelianGroup.commutative G ,, AbelianGroup.commutative H
|
||||
|
||||
record GroupHom {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} (G : Group S _·A_) (H : Group T _·B_) (f : A → B) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
open Group H
|
||||
open Setoid T
|
||||
field
|
||||
groupHom : {x y : A} → f (x ·A y) ∼ (f x) ·B (f y)
|
||||
wellDefined : {x y : A} → Setoid._∼_ S x y → f x ∼ f y
|
||||
|
||||
record InjectiveGroupHom {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_} {underf : A → B} (f : GroupHom G H underf) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
open Setoid S renaming (_∼_ to _∼A_)
|
||||
open Setoid T renaming (_∼_ to _∼B_)
|
||||
field
|
||||
injective : SetoidInjection S T underf
|
||||
|
||||
imageOfIdentityIsIdentity : {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {B : Set n} {T : Setoid {n} {p} B} {_·A_ : A → A → A} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_} {f : A → B} → (hom : GroupHom G H f) → Setoid._∼_ T (f (Group.0G G)) (Group.0G H)
|
||||
imageOfIdentityIsIdentity {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {f = f} hom = Equivalence.symmetric (Setoid.eq T) t
|
||||
where
|
||||
open Group H
|
||||
open Setoid T
|
||||
id2 : Setoid._∼_ S (Group.0G G) ((Group.0G G) ·A (Group.0G G))
|
||||
id2 = Equivalence.symmetric (Setoid.eq S) (Group.identRight G)
|
||||
r : f (Group.0G G) ∼ f (Group.0G G) ·B f (Group.0G G)
|
||||
s : 0G ·B f (Group.0G G) ∼ f (Group.0G G) ·B f (Group.0G G)
|
||||
t : 0G ∼ f (Group.0G G)
|
||||
t = groupsHaveRightCancellation H (f (Group.0G G)) 0G (f (Group.0G G)) s
|
||||
s = Equivalence.transitive (Setoid.eq T) identLeft r
|
||||
r = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined hom id2) (GroupHom.groupHom hom)
|
||||
|
||||
groupHomsCompose : {m n o r s t : _} {A : Set m} {S : Setoid {m} {r} A} {_+A_ : A → A → A} {B : Set n} {T : Setoid {n} {s} B} {_+B_ : B → B → B} {C : Set o} {U : Setoid {o} {t} C} {_+C_ : C → C → C} {G : Group S _+A_} {H : Group T _+B_} {I : Group U _+C_} {f : A → B} {g : B → C} (fHom : GroupHom G H f) (gHom : GroupHom H I g) → GroupHom G I (g ∘ f)
|
||||
GroupHom.wellDefined (groupHomsCompose {G = G} {H} {I} {f} {g} fHom gHom) {x} {y} pr = GroupHom.wellDefined gHom (GroupHom.wellDefined fHom pr)
|
||||
GroupHom.groupHom (groupHomsCompose {S = S} {_+A_ = _·A_} {T = T} {U = U} {_+C_ = _·C_} {G = G} {H} {I} {f} {g} fHom gHom) {x} {y} = answer
|
||||
where
|
||||
open Group I
|
||||
answer : (Setoid._∼_ U) ((g ∘ f) (x ·A y)) ((g ∘ f) x ·C (g ∘ f) y)
|
||||
answer = (Equivalence.transitive (Setoid.eq U)) (GroupHom.wellDefined gHom (GroupHom.groupHom fHom {x} {y}) ) (GroupHom.groupHom gHom {f x} {f y})
|
||||
|
||||
record GroupIso {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} (G : Group S _·A_) (H : Group T _·B_) (f : A → B) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
open Setoid S renaming (_∼_ to _∼G_)
|
||||
open Setoid T renaming (_∼_ to _∼H_)
|
||||
field
|
||||
groupHom : GroupHom G H f
|
||||
bij : SetoidBijection S T f
|
||||
|
||||
record GroupsIsomorphic {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·A_ : A → A → A} {B : Set n} {T : Setoid {n} {p} B} {_·B_ : B → B → B} (G : Group S _·A_) (H : Group T _·B_) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
field
|
||||
isomorphism : A → B
|
||||
proof : GroupIso G H isomorphism
|
||||
|
||||
groupIsosCompose : {m n o r s t : _} {A : Set m} {S : Setoid {m} {r} A} {_+A_ : A → A → A} {B : Set n} {T : Setoid {n} {s} B} {_+B_ : B → B → B} {C : Set o} {U : Setoid {o} {t} C} {_+C_ : C → C → C} {G : Group S _+A_} {H : Group T _+B_} {I : Group U _+C_} {f : A → B} {g : B → C} (fHom : GroupIso G H f) (gHom : GroupIso H I g) → GroupIso G I (g ∘ f)
|
||||
GroupIso.groupHom (groupIsosCompose fHom gHom) = groupHomsCompose (GroupIso.groupHom fHom) (GroupIso.groupHom gHom)
|
||||
GroupIso.bij (groupIsosCompose {A = A} {S = S} {T = T} {C = C} {U = U} {f = f} {g = g} fHom gHom) = record { inj = record { injective = λ pr → (SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij fHom))) (SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij gHom)) pr) ; wellDefined = +WellDefined } ; surj = record { surjective = surj ; wellDefined = +WellDefined } }
|
||||
where
|
||||
open Setoid S renaming (_∼_ to _∼A_)
|
||||
open Setoid U renaming (_∼_ to _∼C_)
|
||||
+WellDefined : {x y : A} → (x ∼A y) → (g (f x) ∼C g (f y))
|
||||
+WellDefined = GroupHom.wellDefined (groupHomsCompose (GroupIso.groupHom fHom) (GroupIso.groupHom gHom))
|
||||
surj : {x : C} → Sg A (λ a → (g (f a) ∼C x))
|
||||
surj {x} with SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij gHom)) {x}
|
||||
surj {x} | a , prA with SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij fHom)) {a}
|
||||
... | b , prB = b , transitive (GroupHom.wellDefined (GroupIso.groupHom gHom) prB) prA
|
||||
where
|
||||
open Equivalence (Setoid.eq U)
|
||||
|
||||
record FiniteGroup {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A → A → A} (G : Group S _·_) {c : _} (quotientSet : Set c) : Set (a ⊔ b ⊔ c) where
|
||||
field
|
||||
toSet : SetoidToSet S quotientSet
|
||||
finite : FiniteSet quotientSet
|
||||
|
||||
groupOrder : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_·_ : A → A → A} {underG : Group S _·_} {c : _} {quotientSet : Set c} (G : FiniteGroup underG quotientSet) → ℕ
|
||||
groupOrder record { toSet = toSet ; finite = record { size = size ; mapping = mapping ; bij = bij } } = size
|
||||
|
||||
--groupIsoInvertible : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d}} {_+A_ : A → A → A} {_+B_ : B → B → B} {G : Group S _+A_} {H : Group T _+B_} {f : A → B} → (iso : GroupIso G H f) → GroupIso H G (Invertible.inverse (bijectionImpliesInvertible (GroupIso.bijective iso)))
|
||||
--GroupHom.groupHom (GroupIso.groupHom (groupIsoInvertible {G = G} {H} {f} iso)) {x} {y} = {!!}
|
||||
-- where
|
||||
-- open Group G renaming (_·_ to _+G_)
|
||||
-- open Group H renaming (_·_ to _+H_)
|
||||
--GroupHom.wellDefined (GroupIso.groupHom (groupIsoInvertible {G = G} {H} {f} iso)) {x} {y} x∼y = {!GroupHom.wellDefined x∼y!}
|
||||
--GroupIso.bijective (groupIsoInvertible {G = G} {H} {f} iso) = invertibleImpliesBijection (inverseIsInvertible (bijectionImpliesInvertible (GroupIso.bijective iso)))
|
||||
|
||||
homRespectsInverse : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_} {underF : A → B} → (f : GroupHom G H underF) → {x : A} → Setoid._∼_ T (underF (Group.inverse G x)) (Group.inverse H (underF x))
|
||||
homRespectsInverse {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {underF = f} fHom {x} = rightInversesAreUnique H (f x) (f (Group.inverse G x)) (transitive (symmetric (GroupHom.groupHom fHom)) (transitive (GroupHom.wellDefined fHom (Group.invLeft G)) (imageOfIdentityIsIdentity fHom)))
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
quotientGroupSetoid : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_} → {underf : A → B} → (f : GroupHom G H underf) → (Setoid {a} {d} A)
|
||||
quotientGroupSetoid {A = A} {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H} {f} fHom = ansSetoid
|
||||
where
|
||||
open Setoid T
|
||||
open Group H
|
||||
open Equivalence eq
|
||||
ansSetoid : Setoid A
|
||||
Setoid._∼_ ansSetoid r s = (f (r ·A (Group.inverse G s))) ∼ 0G
|
||||
Equivalence.reflexive (Setoid.eq ansSetoid) {b} = transitive (GroupHom.wellDefined fHom (Group.invRight G)) (imageOfIdentityIsIdentity fHom)
|
||||
Equivalence.symmetric (Setoid.eq ansSetoid) {m} {n} pr = i
|
||||
where
|
||||
g : f (Group.inverse G (m ·A Group.inverse G n)) ∼ 0G
|
||||
g = transitive (homRespectsInverse fHom {m ·A Group.inverse G n}) (transitive (inverseWellDefined H pr) (invIdentity H))
|
||||
h : f (Group.inverse G (Group.inverse G n) ·A Group.inverse G m) ∼ 0G
|
||||
h = transitive (GroupHom.wellDefined fHom (Equivalence.symmetric (Setoid.eq S) (invContravariant G))) g
|
||||
i : f (n ·A Group.inverse G m) ∼ 0G
|
||||
i = transitive (GroupHom.wellDefined fHom (Group.+WellDefined G (Equivalence.symmetric (Setoid.eq S) (invTwice G n)) (Equivalence.reflexive (Setoid.eq S)))) h
|
||||
Equivalence.transitive (Setoid.eq ansSetoid) {m} {n} {o} prmn prno = transitive (GroupHom.wellDefined fHom (Group.+WellDefined G (Equivalence.reflexive (Setoid.eq S)) (Equivalence.symmetric (Setoid.eq S) (Group.identLeft G)))) k
|
||||
where
|
||||
g : f (m ·A Group.inverse G n) ·B f (n ·A Group.inverse G o) ∼ 0G ·B 0G
|
||||
g = replaceGroupOp H reflexive reflexive prmn prno reflexive
|
||||
h : f (m ·A Group.inverse G n) ·B f (n ·A Group.inverse G o) ∼ 0G
|
||||
h = transitive g identLeft
|
||||
i : f ((m ·A Group.inverse G n) ·A (n ·A Group.inverse G o)) ∼ 0G
|
||||
i = transitive (GroupHom.groupHom fHom) h
|
||||
j : f (m ·A (((Group.inverse G n) ·A n) ·A Group.inverse G o)) ∼ 0G
|
||||
j = transitive (GroupHom.wellDefined fHom (fourWay+Associative G)) i
|
||||
k : f (m ·A ((Group.0G G) ·A Group.inverse G o)) ∼ 0G
|
||||
k = transitive (GroupHom.wellDefined fHom (Group.+WellDefined G (Equivalence.reflexive (Setoid.eq S)) (Group.+WellDefined G (Equivalence.symmetric (Setoid.eq S) (Group.invLeft G)) (Equivalence.reflexive (Setoid.eq S))))) j
|
||||
|
||||
quotientGroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_} → {underf : A → B} → (f : GroupHom G H underf) → Group (quotientGroupSetoid G f) _·A_
|
||||
Group.+WellDefined (quotientGroup {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H = H} {underf = f} fHom) {x} {y} {m} {n} x~m y~n = ans
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence (Setoid.eq T)
|
||||
p1 : f ((x ·A y) ·A (Group.inverse G (m ·A n))) ∼ f ((x ·A y) ·A ((Group.inverse G n) ·A (Group.inverse G m)))
|
||||
p2 : f ((x ·A y) ·A ((Group.inverse G n) ·A (Group.inverse G m))) ∼ f (x ·A ((y ·A (Group.inverse G n)) ·A (Group.inverse G m)))
|
||||
p3 : f (x ·A ((y ·A (Group.inverse G n)) ·A (Group.inverse G m))) ∼ (f x) ·B f ((y ·A (Group.inverse G n)) ·A (Group.inverse G m))
|
||||
p4 : (f x) ·B f ((y ·A (Group.inverse G n)) ·A (Group.inverse G m)) ∼ (f x) ·B (f (y ·A (Group.inverse G n)) ·B f (Group.inverse G m))
|
||||
p5 : (f x) ·B (f (y ·A (Group.inverse G n)) ·B f (Group.inverse G m)) ∼ (f x) ·B (Group.0G H ·B f (Group.inverse G m))
|
||||
p6 : (f x) ·B (Group.0G H ·B f (Group.inverse G m)) ∼ (f x) ·B f (Group.inverse G m)
|
||||
p7 : (f x) ·B f (Group.inverse G m) ∼ f (x ·A (Group.inverse G m))
|
||||
p8 : f (x ·A (Group.inverse G m)) ∼ Group.0G H
|
||||
p1 = GroupHom.wellDefined fHom (Group.+WellDefined G (Equivalence.reflexive (Setoid.eq S)) (invContravariant G))
|
||||
p2 = GroupHom.wellDefined fHom (Equivalence.symmetric (Setoid.eq S) (fourWay+Associative G))
|
||||
p3 = GroupHom.groupHom fHom
|
||||
p4 = Group.+WellDefined H reflexive (GroupHom.groupHom fHom)
|
||||
p5 = Group.+WellDefined H reflexive (replaceGroupOp H (symmetric y~n) reflexive reflexive reflexive reflexive)
|
||||
p6 = Group.+WellDefined H reflexive (Group.identLeft H)
|
||||
p7 = symmetric (GroupHom.groupHom fHom)
|
||||
p8 = x~m
|
||||
ans : f ((x ·A y) ·A (Group.inverse G (m ·A n))) ∼ Group.0G H
|
||||
ans = transitive p1 (transitive p2 (transitive p3 (transitive p4 (transitive p5 (transitive p6 (transitive p7 p8))))))
|
||||
Group.0G (quotientGroup G fHom) = Group.0G G
|
||||
Group.inverse (quotientGroup G fHom) = Group.inverse G
|
||||
Group.+Associative (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} {b} {c} = ans
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence (Setoid.eq T)
|
||||
ans : f ((a ·A (b ·A c)) ·A (Group.inverse G ((a ·A b) ·A c))) ∼ Group.0G H
|
||||
ans = transitive (GroupHom.wellDefined fHom (transferToRight'' G (Group.+Associative G))) (imageOfIdentityIsIdentity fHom)
|
||||
Group.identRight (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
transitiveG = Equivalence.transitive (Setoid.eq S)
|
||||
ans : f ((a ·A 0G) ·A inverse a) ∼ Group.0G H
|
||||
ans = transitive (GroupHom.wellDefined fHom (transitiveG (Group.+WellDefined G (Group.identRight G) (Equivalence.reflexive (Setoid.eq S))) (Group.invRight G))) (imageOfIdentityIsIdentity fHom)
|
||||
Group.identLeft (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
transitiveG = Equivalence.transitive (Setoid.eq S)
|
||||
ans : f ((0G ·A a) ·A (inverse a)) ∼ Group.0G H
|
||||
ans = transitive (GroupHom.wellDefined fHom (transitiveG (Group.+WellDefined G (Group.identLeft G) (Equivalence.reflexive (Setoid.eq S))) (Group.invRight G))) (imageOfIdentityIsIdentity fHom)
|
||||
Group.invLeft (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
ans : f ((inverse x ·A x) ·A (inverse 0G)) ∼ (Group.0G H)
|
||||
ans = transitive (GroupHom.wellDefined fHom (Equivalence.transitive (Setoid.eq S) (replaceGroupOp G (Equivalence.symmetric (Setoid.eq S) (Group.invLeft G)) (Equivalence.symmetric (Setoid.eq S) (invIdentity G)) (Equivalence.reflexive (Setoid.eq S)) ((Equivalence.reflexive (Setoid.eq S))) ((Equivalence.reflexive (Setoid.eq S)))) (identRight {0G}))) (imageOfIdentityIsIdentity fHom)
|
||||
Group.invRight (quotientGroup {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
ans : f ((x ·A inverse x) ·A (inverse 0G)) ∼ (Group.0G H)
|
||||
ans = transitive (GroupHom.wellDefined fHom (Equivalence.transitive (Setoid.eq S) (replaceGroupOp G (Equivalence.symmetric (Setoid.eq S) (Group.invRight G)) (Equivalence.symmetric (Setoid.eq S) (invIdentity G)) (Equivalence.reflexive (Setoid.eq S)) (Equivalence.reflexive (Setoid.eq S)) (Equivalence.reflexive (Setoid.eq S))) (identRight {0G}))) (imageOfIdentityIsIdentity fHom)
|
||||
|
||||
record Subgroup {a} {b} {c} {d} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) (H : Group T _·B_) {f : B → A} (hom : GroupHom H G f) : Set (a ⊔ b ⊔ c ⊔ d) where
|
||||
open Setoid T renaming (_∼_ to _∼G_)
|
||||
open Setoid S renaming (_∼_ to _∼H_)
|
||||
field
|
||||
fInj : SetoidInjection T S f
|
||||
record Subgroup {a} {b} {c} {d} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) (H : Group T _·B_) {f : B → A} (hom : GroupHom H G f) : Set (a ⊔ b ⊔ c ⊔ d) where
|
||||
open Setoid T renaming (_∼_ to _∼G_)
|
||||
open Setoid S renaming (_∼_ to _∼H_)
|
||||
field
|
||||
fInj : SetoidInjection T S f
|
||||
{-
|
||||
quotientHom : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_} → {f : A → B} → (fHom : GroupHom G H f) → A → A
|
||||
quotientHom {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {f = f} fHom a = {!!}
|
||||
quotientHom : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_} → {f : A → B} → (fHom : GroupHom G H f) → A → A
|
||||
quotientHom {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {f = f} fHom a = {!!}
|
||||
|
||||
quotientInjection : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_} → {f : A → B} → (fHom : GroupHom G H f) → GroupHom (quotientGroup G fHom) G (quotientHom G fHom)
|
||||
GroupHom.groupHom (quotientInjection {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {f = f} fHom) {x} {y} = {!!}
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Reflexive reflexiveEq
|
||||
GroupHom.wellDefined (quotientInjection {S = S} {T = T} {_·A_ = _·A_} G {H = H} {f = f} fHom) {x} {y} x~y = {!!}
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Setoid T renaming (_∼_ to _∼T_)
|
||||
open Equivalence (Setoid.eq S)
|
||||
open Reflexive reflexiveEq
|
||||
have : f (x ·A inverse y) ∼T Group.0G H
|
||||
have = x~y
|
||||
need : x ∼ y
|
||||
need = {!!}
|
||||
quotientInjection : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} (G : Group S _·A_) {H : Group T _·B_} → {f : A → B} → (fHom : GroupHom G H f) → GroupHom (quotientGroup G fHom) G (quotientHom G fHom)
|
||||
GroupHom.groupHom (quotientInjection {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {f = f} fHom) {x} {y} = {!!}
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Reflexive reflexiveEq
|
||||
GroupHom.wellDefined (quotientInjection {S = S} {T = T} {_·A_ = _·A_} G {H = H} {f = f} fHom) {x} {y} x~y = {!!}
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Setoid T renaming (_∼_ to _∼T_)
|
||||
open Equivalence (Setoid.eq S)
|
||||
open Reflexive reflexiveEq
|
||||
have : f (x ·A inverse y) ∼T Group.0G H
|
||||
have = x~y
|
||||
need : x ∼ y
|
||||
need = {!!}
|
||||
|
||||
quotientIsSubgroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_} → {f : A → B} → {fHom : GroupHom G H f} → Subgroup G (quotientGroup G fHom) (quotientInjection G fHom)
|
||||
quotientIsSubgroup = {!!}
|
||||
quotientIsSubgroup : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_·A_ : A → A → A} {_·B_ : B → B → B} {G : Group S _·A_} {H : Group T _·B_} → {f : A → B} → {fHom : GroupHom G H f} → Subgroup G (quotientGroup G fHom) (quotientInjection G fHom)
|
||||
quotientIsSubgroup = {!!}
|
||||
|
||||
-}
|
||||
subgroupOfAbelianIsAbelian : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A → A → A} {_+B_ : B → B → B} {G : Group S _+A_} {H : Group T _+B_} {f : B → A} {fHom : GroupHom H G f} → Subgroup G H fHom → AbelianGroup G → AbelianGroup H
|
||||
AbelianGroup.commutative (subgroupOfAbelianIsAbelian {S = S} {_+B_ = _+B_} {fHom = fHom} record { fInj = fInj } record { commutative = commutative }) {x} {y} = SetoidInjection.injective fInj (transitive (GroupHom.groupHom fHom) (transitive commutative (symmetric (GroupHom.groupHom fHom))))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
subgroupOfAbelianIsAbelian : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A → A → A} {_+B_ : B → B → B} {G : Group S _+A_} {H : Group T _+B_} {f : B → A} {fHom : GroupHom H G f} → Subgroup G H fHom → AbelianGroup G → AbelianGroup H
|
||||
AbelianGroup.commutative (subgroupOfAbelianIsAbelian {S = S} {_+B_ = _+B_} {fHom = fHom} record { fInj = fInj } record { commutative = commutative }) {x} {y} = SetoidInjection.injective fInj (transitive (GroupHom.groupHom fHom) (transitive commutative (symmetric (GroupHom.groupHom fHom))))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
abelianIsGroupProperty : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A → A → A} {_+B_ : B → B → B} {G : Group S _+A_} {H : Group T _+B_} → GroupsIsomorphic G H → AbelianGroup H → AbelianGroup G
|
||||
abelianIsGroupProperty iso abH = subgroupOfAbelianIsAbelian {fHom = GroupIso.groupHom (GroupsIsomorphic.proof iso)} (record { fInj = SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof iso)) }) abH
|
||||
abelianIsGroupProperty : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+A_ : A → A → A} {_+B_ : B → B → B} {G : Group S _+A_} {H : Group T _+B_} → GroupsIsomorphic G H → AbelianGroup H → AbelianGroup G
|
||||
abelianIsGroupProperty iso abH = subgroupOfAbelianIsAbelian {fHom = GroupIso.groupHom (GroupsIsomorphic.proof iso)} (record { fInj = SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof iso)) }) abH
|
||||
|
@@ -21,7 +21,7 @@ ringMinusExtracts {x = x} {y} = transferToRight' additiveGroup (transitive (symm
|
||||
where
|
||||
open Equivalence eq
|
||||
|
||||
ringMinusExtracts' : {x y : A} → Setoid._∼_ S ((Group.inverse (Ring.additiveGroup R) x) * y) (Group.inverse (Ring.additiveGroup R) (x * y))
|
||||
ringMinusExtracts' : {x y : A} → ((inverse x) * y) ∼ inverse (x * y)
|
||||
ringMinusExtracts' {x = x} {y} = transitive *Commutative (transitive ringMinusExtracts (inverseWellDefined additiveGroup *Commutative))
|
||||
where
|
||||
open Equivalence eq
|
||||
@@ -40,3 +40,6 @@ groupLemmaMove0G {S = S} G {x} pr = transitive (symmetric (invInv G)) (transitiv
|
||||
|
||||
groupLemmaMove0G' : {a b : _} → {A : Set a} → {_·_ : A → A → A} → {S : Setoid {a} {b} A} → (G : Group S _·_) → {x : A} → Setoid._∼_ S x (Group.0G G) → (Setoid._∼_ S (Group.0G G) (Group.inverse G x))
|
||||
groupLemmaMove0G' {S = S} G {x} pr = transferToRight' G (Equivalence.transitive (Setoid.eq S) (Group.identLeft G) pr)
|
||||
|
||||
oneZeroImpliesAllZero : 0R ∼ 1R → {x : A} → x ∼ 0R
|
||||
oneZeroImpliesAllZero 0=1 = Equivalence.transitive eq (Equivalence.symmetric eq identIsIdent) (Equivalence.transitive eq (*WellDefined (Equivalence.symmetric eq 0=1) (Equivalence.reflexive eq)) (Equivalence.transitive eq *Commutative timesZero))
|
||||
|
@@ -25,7 +25,6 @@ open Group additiveGroup
|
||||
|
||||
open import Rings.Lemmas R
|
||||
|
||||
|
||||
ringAddInequalities : {w x y z : A} → w < x → y < z → (w + y) < (x + z)
|
||||
ringAddInequalities {w = w} {x} {y} {z} w<x y<z = transitive (orderRespectsAddition w<x y) (<WellDefined groupIsAbelian groupIsAbelian (orderRespectsAddition y<z x))
|
||||
|
||||
@@ -188,6 +187,9 @@ ringSwapNegatives {x} {y} -x<-y = SetoidPartialOrder.<WellDefined pOrder (Equiva
|
||||
v : ((y + (Group.inverse additiveGroup x)) + x) < (0R + x)
|
||||
v = OrderedRing.orderRespectsAddition order u x
|
||||
|
||||
ringSwapNegatives' : {x y : A} → x < y → (Group.inverse (Ring.additiveGroup R) y) < (Group.inverse (Ring.additiveGroup R) x)
|
||||
ringSwapNegatives' {x} {y} x<y = ringSwapNegatives (<WellDefined (Equivalence.symmetric eq (invTwice additiveGroup _)) (Equivalence.symmetric eq (invTwice additiveGroup _)) x<y)
|
||||
|
||||
ringCanMultiplyByNegative : {x y c : A} → c < (Ring.0R R) → x < y → (y * c) < (x * c)
|
||||
ringCanMultiplyByNegative {x} {y} {c} c<0 x<y = ringSwapNegatives u
|
||||
where
|
||||
@@ -246,12 +248,12 @@ absNegation a | inr 0=a | inl (inl 0<-a) = exFalso (irreflexive {0G} (<WellDefin
|
||||
absNegation a | inr 0=a | inl (inr -a<0) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.transitive eq (inverseWellDefined additiveGroup (Equivalence.symmetric eq 0=a)) (invIdent additiveGroup)) (Equivalence.reflexive eq) -a<0))
|
||||
absNegation a | inr 0=a | inr 0=-a = Equivalence.transitive eq (Equivalence.symmetric eq 0=a) 0=-a
|
||||
|
||||
lemm4 : (a b : A) → (0G < a) → (b < 0G) → (a * b) < 0G
|
||||
lemm4 a b 0<a b<0 with orderRespectsMultiplication 0<a (lemm2 _ b<0)
|
||||
posTimesNeg : (a b : A) → (0G < a) → (b < 0G) → (a * b) < 0G
|
||||
posTimesNeg a b 0<a b<0 with orderRespectsMultiplication 0<a (lemm2 _ b<0)
|
||||
... | bl = <WellDefined (invTwice additiveGroup _) (Equivalence.reflexive eq) (lemm2' _ (<WellDefined (Equivalence.reflexive eq) ringMinusExtracts bl))
|
||||
|
||||
lemm5 : (a b : A) → (a < 0G) → (b < 0G) → 0G < (a * b)
|
||||
lemm5 a b a<0 b<0 with orderRespectsMultiplication (lemm2 _ a<0) (lemm2 _ b<0)
|
||||
negTimesPos : (a b : A) → (a < 0G) → (b < 0G) → 0G < (a * b)
|
||||
negTimesPos a b a<0 b<0 with orderRespectsMultiplication (lemm2 _ a<0) (lemm2 _ b<0)
|
||||
... | bl = <WellDefined (Equivalence.reflexive eq) twoNegativesTimes bl
|
||||
|
||||
absRespectsTimes : (a b : A) → abs (a * b) ∼ (abs a) * (abs b)
|
||||
@@ -265,20 +267,20 @@ absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inl (inl 0<ab) with <WellDefined (Equivalence.reflexive eq) ringMinusExtracts (orderRespectsMultiplication 0<a (lemm2 b b<0))
|
||||
... | bl = exFalso (irreflexive {0G} (SetoidPartialOrder.transitive pOrder 0<ab (<WellDefined (invTwice additiveGroup _) (Equivalence.reflexive eq) (lemm2' _ bl))))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq 0=ab) (Equivalence.reflexive eq) (lemm4 a b 0<a b<0)))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inl (inr b<0) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq 0=ab) (Equivalence.reflexive eq) (posTimesNeg a b 0<a b<0)))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) 0<ab))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inl (inr ab<0) = exFalso ((irreflexive {0G} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) (Equivalence.reflexive eq) ab<0)))
|
||||
absRespectsTimes a b | inl (inl 0<a) | inr 0=b | inr 0=ab = Equivalence.reflexive eq
|
||||
absRespectsTimes a b | inl (inr a<0) with totality 0R b
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inl (inl 0<ab) = exFalso (irreflexive {0G} (SetoidPartialOrder.transitive pOrder 0<ab (<WellDefined *Commutative (Equivalence.reflexive eq) (lemm4 b a 0<b a<0))))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inl (inl 0<ab) = exFalso (irreflexive {0G} (SetoidPartialOrder.transitive pOrder 0<ab (<WellDefined *Commutative (Equivalence.reflexive eq) (posTimesNeg b a 0<b a<0))))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inl (inr ab<0) = Equivalence.symmetric eq ringMinusExtracts'
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq 0=ab *Commutative)) (Equivalence.reflexive eq) (lemm4 b a 0<b a<0)))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inl 0<b) | inr 0=ab = exFalso (irreflexive {0G} (<WellDefined (Equivalence.symmetric eq (Equivalence.transitive eq 0=ab *Commutative)) (Equivalence.reflexive eq) (posTimesNeg b a 0<b a<0)))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inl (inl 0<ab) = Equivalence.symmetric eq twoNegativesTimes
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inl (inr ab<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.transitive pOrder (lemm5 a b a<0 b<0) ab<0))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inr 0=ab = exFalso (exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=ab) (lemm5 a b a<0 b<0))))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inl (inr ab<0) = exFalso (irreflexive {0G} (SetoidPartialOrder.transitive pOrder (negTimesPos a b a<0 b<0) ab<0))
|
||||
absRespectsTimes a b | inl (inr a<0) | inl (inr b<0) | inr 0=ab = exFalso (exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=ab) (negTimesPos a b a<0 b<0))))
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b with totality 0R (a * b)
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) (timesZero {a})) 0<ab))
|
||||
absRespectsTimes a b | inl (inr a<0) | inr 0=b | inl (inr ab<0) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) (Equivalence.reflexive eq) ab<0))
|
||||
@@ -296,3 +298,33 @@ absRespectsTimes a b | inr 0=a | inr 0=b with totality 0R (a * b)
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b | inl (inl 0<ab) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) 0<ab))
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b | inl (inr ab<0) = exFalso (irreflexive {0R} (<WellDefined (Equivalence.transitive eq (*WellDefined (Equivalence.reflexive eq) (Equivalence.symmetric eq 0=b)) timesZero) (Equivalence.reflexive eq) ab<0))
|
||||
absRespectsTimes a b | inr 0=a | inr 0=b | inr 0=ab = Equivalence.reflexive eq
|
||||
|
||||
absNonnegative : {a : A} → (abs a < 0R) → False
|
||||
absNonnegative {a} pr with SetoidTotalOrder.totality tOrder 0R a
|
||||
absNonnegative {a} pr | inl (inl x) = irreflexive {0G} (SetoidPartialOrder.transitive pOrder x pr)
|
||||
absNonnegative {a} pr | inl (inr x) = irreflexive {0G} (SetoidPartialOrder.transitive pOrder (<WellDefined (Equivalence.reflexive eq) (invTwice additiveGroup a) (lemm2 (inverse a) pr)) x)
|
||||
absNonnegative {a} pr | inr x = irreflexive {0G} (<WellDefined (Equivalence.symmetric eq x) (Equivalence.reflexive eq) pr)
|
||||
|
||||
a-bPos : {a b : A} → ((a ∼ b) → False) → 0R < abs (a + inverse b)
|
||||
a-bPos {a} {b} a!=b with totality 0R (a + inverse b)
|
||||
a-bPos {a} {b} a!=b | inl (inl x) = x
|
||||
a-bPos {a} {b} a!=b | inl (inr x) = lemm2 _ x
|
||||
a-bPos {a} {b} a!=b | inr x = exFalso (a!=b (transferToRight additiveGroup (Equivalence.symmetric eq x)))
|
||||
|
||||
absZeroImpliesZero : {a : A} → abs a ∼ 0R → a ∼ 0R
|
||||
absZeroImpliesZero {a} a=0 with totality 0R a
|
||||
absZeroImpliesZero {a} a=0 | inl (inl 0<a) = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) a=0 0<a))
|
||||
absZeroImpliesZero {a} a=0 | inl (inr a<0) = Equivalence.symmetric eq (lemm3 (inverse a) a (Equivalence.symmetric eq invLeft) (Equivalence.symmetric eq a=0))
|
||||
absZeroImpliesZero {a} a=0 | inr 0=a = a=0
|
||||
|
||||
halvePositive : (a : A) → 0R < (a + a) → 0R < a
|
||||
halvePositive a 0<2a with totality 0R a
|
||||
halvePositive a 0<2a | inl (inl x) = x
|
||||
halvePositive a 0<2a | inl (inr a<0) = exFalso (irreflexive {a + a} (SetoidPartialOrder.transitive pOrder (<WellDefined (Equivalence.reflexive eq) identRight (ringAddInequalities a<0 a<0)) 0<2a))
|
||||
halvePositive a 0<2a | inr x = exFalso (irreflexive {0G} (<WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq (+WellDefined (Equivalence.symmetric eq x) (Equivalence.symmetric eq x)) identRight) 0<2a))
|
||||
|
||||
0<1 : (0R ∼ 1R → False) → 0R < 1R
|
||||
0<1 0!=1 with SetoidTotalOrder.totality tOrder 0R 1R
|
||||
0<1 0!=1 | inl (inl x) = x
|
||||
0<1 0!=1 | inl (inr x) = <WellDefined (Equivalence.reflexive eq) (Equivalence.transitive eq twoNegativesTimes identIsIdent) (orderRespectsMultiplication (lemm2 1R x) (lemm2 1R x))
|
||||
0<1 0!=1 | inr x = exFalso (0!=1 x)
|
||||
|
@@ -2,6 +2,7 @@
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Setoids.Setoids
|
||||
|
||||
module Sequences where
|
||||
|
||||
@@ -27,6 +28,10 @@ funcToSequenceReversible : {a : _} {A : Set a} (f : ℕ → A) → (n : ℕ) →
|
||||
funcToSequenceReversible f zero = refl
|
||||
funcToSequenceReversible f (succ n) = funcToSequenceReversible (λ i → f (succ i)) n
|
||||
|
||||
map : {a b : _} {A : Set a} {B : Set b} (f : A → B) (s : Sequence A) → Sequence B
|
||||
Sequence.head (map f s) = f (Sequence.head s)
|
||||
Sequence.tail (map f s) = map f (Sequence.tail s)
|
||||
|
||||
apply : {a b c : _} {A : Set a} {B : Set b} {C : Set c} (f : A → B → C) (s1 : Sequence A) (s2 : Sequence B) → Sequence C
|
||||
Sequence.head (apply f s1 s2) = f (Sequence.head s1) (Sequence.head s2)
|
||||
Sequence.tail (apply f s1 s2) = apply f (Sequence.tail s1) (Sequence.tail s2)
|
||||
@@ -34,3 +39,23 @@ Sequence.tail (apply f s1 s2) = apply f (Sequence.tail s1) (Sequence.tail s2)
|
||||
indexAndConst : {a : _} {A : Set a} (a : A) (n : ℕ) → index (constSequence a) n ≡ a
|
||||
indexAndConst a zero = refl
|
||||
indexAndConst a (succ n) = indexAndConst a n
|
||||
|
||||
mapTwice : {a b c : _} {A : Set a} {B : Set b} {C : Set c} (f : A → B) (g : B → C) (s : Sequence A) → {n : ℕ} → index (map g (map f s)) n ≡ index (map (λ i → g (f i)) s) n
|
||||
mapTwice f g s {zero} = refl
|
||||
mapTwice f g s {succ n} = mapTwice f g (Sequence.tail s) {n}
|
||||
|
||||
mapAndIndex : {a b : _} {A : Set a} {B : Set b} (s : Sequence A) (f : A → B) (n : ℕ) → f (index s n) ≡ index (map f s) n
|
||||
mapAndIndex s f zero = refl
|
||||
mapAndIndex s f (succ n) = mapAndIndex (Sequence.tail s) f n
|
||||
|
||||
indexExtensional : {a b c : _} {A : Set a} {B : Set b} (T : Setoid {_} {c} B) (s : Sequence A) (f g : A → B) → (extension : ∀ {x} → (Setoid._∼_ T (f x) (g x))) → {n : ℕ} → Setoid._∼_ T (index (map f s) n) (index (map g s) n)
|
||||
indexExtensional T s f g extension {zero} = extension
|
||||
indexExtensional T s f g extension {succ n} = indexExtensional T (Sequence.tail s) f g extension {n}
|
||||
|
||||
indexAndApply : {a b c : _} {A : Set a} {B : Set b} {C : Set c} (s1 : Sequence A) (s2 : Sequence B) (f : A → B → C) → {m : ℕ} → index (apply f s1 s2) m ≡ f (index s1 m) (index s2 m)
|
||||
indexAndApply s1 s2 f {zero} = refl
|
||||
indexAndApply s1 s2 f {succ m} = indexAndApply (Sequence.tail s1) (Sequence.tail s2) f {m}
|
||||
|
||||
mapAndApply : {a b c d : _} {A : Set a} {B : Set b} {C : Set c} {D : Set d} (s1 : Sequence A) (s2 : Sequence B) (f : A → B → C) (g : C → D) → (m : ℕ) → index (map g (apply f s1 s2)) m ≡ g (f (index s1 m) (index s2 m))
|
||||
mapAndApply s1 s2 f g zero = refl
|
||||
mapAndApply s1 s2 f g (succ m) = mapAndApply (Sequence.tail s1) (Sequence.tail s2) f g m
|
||||
|
@@ -9,34 +9,34 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Setoids.Orders where
|
||||
|
||||
record SetoidPartialOrder {a b c : _} {A : Set a} (S : Setoid {a} {b} A) (_<_ : Rel {a} {c} A) : Set (a ⊔ b ⊔ c) where
|
||||
open Setoid S
|
||||
field
|
||||
<WellDefined : {a b c d : A} → (a ∼ b) → (c ∼ d) → a < c → b < d
|
||||
irreflexive : {x : A} → (x < x) → False
|
||||
transitive : {a b c : A} → (a < b) → (b < c) → (a < c)
|
||||
record SetoidPartialOrder {a b c : _} {A : Set a} (S : Setoid {a} {b} A) (_<_ : Rel {a} {c} A) : Set (a ⊔ b ⊔ c) where
|
||||
open Setoid S
|
||||
field
|
||||
<WellDefined : {a b c d : A} → (a ∼ b) → (c ∼ d) → a < c → b < d
|
||||
irreflexive : {x : A} → (x < x) → False
|
||||
transitive : {a b c : A} → (a < b) → (b < c) → (a < c)
|
||||
|
||||
record SetoidTotalOrder {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_<_ : Rel {a} {c} A} (P : SetoidPartialOrder S _<_) : Set (a ⊔ b ⊔ c) where
|
||||
open Setoid S
|
||||
field
|
||||
totality : (a b : A) → ((a < b) || (b < a)) || (a ∼ b)
|
||||
partial : SetoidPartialOrder S _<_
|
||||
partial = P
|
||||
min : A → A → A
|
||||
min a b with totality a b
|
||||
min a b | inl (inl a<b) = a
|
||||
min a b | inl (inr b<a) = b
|
||||
min a b | inr a=b = a
|
||||
max : A → A → A
|
||||
max a b with totality a b
|
||||
max a b | inl (inl a<b) = b
|
||||
max a b | inl (inr b<a) = a
|
||||
max a b | inr a=b = b
|
||||
record SetoidTotalOrder {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_<_ : Rel {a} {c} A} (P : SetoidPartialOrder S _<_) : Set (a ⊔ b ⊔ c) where
|
||||
open Setoid S
|
||||
field
|
||||
totality : (a b : A) → ((a < b) || (b < a)) || (a ∼ b)
|
||||
partial : SetoidPartialOrder S _<_
|
||||
partial = P
|
||||
min : A → A → A
|
||||
min a b with totality a b
|
||||
min a b | inl (inl a<b) = a
|
||||
min a b | inl (inr b<a) = b
|
||||
min a b | inr a=b = a
|
||||
max : A → A → A
|
||||
max a b with totality a b
|
||||
max a b | inl (inl a<b) = b
|
||||
max a b | inl (inr b<a) = a
|
||||
max a b | inr a=b = b
|
||||
|
||||
partialOrderToSetoidPartialOrder : {a b : _} {A : Set a} (P : PartialOrder {a} {b} A) → SetoidPartialOrder (reflSetoid A) (PartialOrder._<_ P)
|
||||
SetoidPartialOrder.<WellDefined (partialOrderToSetoidPartialOrder P) a=b c=d a<c rewrite a=b | c=d = a<c
|
||||
SetoidPartialOrder.irreflexive (partialOrderToSetoidPartialOrder P) = PartialOrder.irreflexive P
|
||||
SetoidPartialOrder.transitive (partialOrderToSetoidPartialOrder P) = PartialOrder.transitive P
|
||||
partialOrderToSetoidPartialOrder : {a b : _} {A : Set a} (P : PartialOrder {a} {b} A) → SetoidPartialOrder (reflSetoid A) (PartialOrder._<_ P)
|
||||
SetoidPartialOrder.<WellDefined (partialOrderToSetoidPartialOrder P) a=b c=d a<c rewrite a=b | c=d = a<c
|
||||
SetoidPartialOrder.irreflexive (partialOrderToSetoidPartialOrder P) = PartialOrder.irreflexive P
|
||||
SetoidPartialOrder.transitive (partialOrderToSetoidPartialOrder P) = PartialOrder.transitive P
|
||||
|
||||
totalOrderToSetoidTotalOrder : {a b : _} {A : Set a} (T : TotalOrder {a} {b} A) → SetoidTotalOrder (partialOrderToSetoidPartialOrder (TotalOrder.order T))
|
||||
SetoidTotalOrder.totality (totalOrderToSetoidTotalOrder T) = TotalOrder.totality T
|
||||
totalOrderToSetoidTotalOrder : {a b : _} {A : Set a} (T : TotalOrder {a} {b} A) → SetoidTotalOrder (partialOrderToSetoidPartialOrder (TotalOrder.order T))
|
||||
SetoidTotalOrder.totality (totalOrderToSetoidTotalOrder T) = TotalOrder.totality T
|
||||
|
Reference in New Issue
Block a user