mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-06 20:38:40 +00:00
Lots of rings (#82)
This commit is contained in:
@@ -47,7 +47,7 @@ open import Rings.Lemmas
|
||||
open import Rings.Orders.Partial.Definition
|
||||
open import Rings.Orders.Total.Lemmas
|
||||
open import Rings.Orders.Partial.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Rings.DirectSum
|
||||
open import Rings.Polynomial.Ring
|
||||
open import Rings.Polynomial.Evaluation
|
||||
@@ -55,6 +55,17 @@ open import Rings.Ideals.Definition
|
||||
open import Rings.Isomorphisms.Definition
|
||||
open import Rings.Quotients.Definition
|
||||
open import Rings.Cosets
|
||||
open import Rings.EuclideanDomains.Definition
|
||||
open import Rings.EuclideanDomains.Examples
|
||||
open import Rings.Homomorphisms.Image
|
||||
open import Rings.Homomorphisms.Kernel
|
||||
open import Rings.Ideals.FirstIsomorphismTheorem
|
||||
open import Rings.Ideals.Lemmas
|
||||
open import Rings.Ideals.Prime.Definition
|
||||
open import Rings.Ideals.Principal.Definition
|
||||
open import Rings.IntegralDomains.Examples
|
||||
open import Rings.PrincipalIdealDomain
|
||||
open import Rings.Subrings.Definition
|
||||
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.DirectSum
|
||||
|
@@ -6,7 +6,7 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
@@ -25,8 +25,7 @@ fieldOfFractionsPlus (a ,, (b , b!=0)) (c ,, (d , d!=0)) = (((a * d) + (b * c))
|
||||
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
|
||||
ans pr | f = exFalso (d!=0 (f b!=0))
|
||||
|
||||
plusWellDefined : {a b c d : fieldOfFractionsSet} → (Setoid._∼_ fieldOfFractionsSetoid a c) → (Setoid._∼_ fieldOfFractionsSetoid b d) → Setoid._∼_ fieldOfFractionsSetoid (fieldOfFractionsPlus a b) (fieldOfFractionsPlus c d)
|
||||
plusWellDefined {a ,, (b , b!=0)} {c ,, (d , d!=0)} {e ,, (f , f!=0)} {g ,, (h , h!=0)} af=be ch=dg = need
|
||||
|
@@ -6,7 +6,7 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
|
@@ -6,7 +6,7 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
|
@@ -8,7 +8,7 @@ open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
|
@@ -6,7 +6,7 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
@@ -25,8 +25,7 @@ fieldOfFractionsTimes (a ,, (b , b!=0)) (c ,, (d , d!=0)) = (a * c) ,, ((b * d)
|
||||
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
|
||||
ans pr | f = exFalso (d!=0 (f b!=0))
|
||||
|
||||
fieldOfFractionsTimesWellDefined : {a b c d : fieldOfFractionsSet} → (Setoid._∼_ fieldOfFractionsSetoid a c) → (Setoid._∼_ fieldOfFractionsSetoid b d) → (Setoid._∼_ fieldOfFractionsSetoid (fieldOfFractionsTimes a b) (fieldOfFractionsTimes c d))
|
||||
fieldOfFractionsTimesWellDefined {a ,, (b , _)} {c ,, (d , _)} {e ,, (f , _)} {g ,, (h , _)} af=be ch=dg = need
|
||||
|
File diff suppressed because it is too large
Load Diff
@@ -6,7 +6,7 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
|
@@ -6,16 +6,17 @@ open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Rings.IntegralDomains
|
||||
open import Fields.Fields
|
||||
open import Functions
|
||||
open import Setoids.Setoids
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.IntegralDomains.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Fields.FieldOfFractions.Setoid {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) where
|
||||
|
||||
|
||||
fieldOfFractionsSet : Set (a ⊔ b)
|
||||
fieldOfFractionsSet = (A && (Sg A (λ a → (Setoid._∼_ S a (Ring.0R R) → False))))
|
||||
|
||||
@@ -36,9 +37,7 @@ Equivalence.transitive (Setoid.eq fieldOfFractionsSetoid) {a ,, (b , b!=0)} {c ,
|
||||
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))
|
||||
p4 : ((d ∼ 0R) → False) → ((a * f) ∼ (b * e))
|
||||
p4 = cancelIntDom R 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
|
||||
p5 = p4 d!=0
|
||||
|
@@ -8,7 +8,7 @@ open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Orders
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
@@ -23,6 +23,8 @@ record Field {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A}
|
||||
field
|
||||
allInvertible : (a : A) → ((a ∼ Group.0G (Ring.additiveGroup R)) → False) → Sg A (λ t → t * a ∼ 1R)
|
||||
nontrivial : (0R ∼ 1R) → False
|
||||
0F : A
|
||||
0F = Ring.0R R
|
||||
|
||||
record Field' {m n : _} : Set (lsuc m ⊔ lsuc n) where
|
||||
field
|
||||
|
@@ -11,7 +11,7 @@ open import Rings.Lemmas
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Orders
|
||||
open import Orders
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Fields.Fields
|
||||
@@ -84,35 +84,39 @@ abstract
|
||||
r : (1/2 * ((1R + 1R) * x)) ∼ (1/2' * ((1R + 1R) * x))
|
||||
r = Equivalence.transitive eq *Associative (Equivalence.transitive eq q (Equivalence.symmetric eq *Associative))
|
||||
|
||||
private
|
||||
orderedFieldIntDom : {a b : A} → (a * b ∼ 0R) → (a ∼ 0R) || (b ∼ 0R)
|
||||
orderedFieldIntDom {a} {b} ab=0 with totality 0R a
|
||||
... | inl (inl x) = inr (Equivalence.transitive eq (Equivalence.transitive eq (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Equivalence eq
|
||||
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' = Equivalence.transitive eq (symmetric *Associative) (Equivalence.transitive eq p (Ring.timesZero R))
|
||||
orderedFieldIntDom {a} {b} ab=0 | inl (inr x) = inr (Equivalence.transitive eq (Equivalence.transitive eq (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Equivalence eq
|
||||
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' = Equivalence.transitive eq (symmetric *Associative) (Equivalence.transitive eq p (Ring.timesZero R))
|
||||
orderedFieldIntDom {a} {b} ab=0 | inr x = inl (Equivalence.symmetric (Setoid.eq S) x)
|
||||
|
||||
orderedFieldIsIntDom : IntegralDomain R
|
||||
IntegralDomain.intDom orderedFieldIsIntDom {a} {b} ab=0 with totality (Ring.0R R) a
|
||||
IntegralDomain.intDom orderedFieldIsIntDom {a} {b} ab=0 | inl (inl x) = inr (Equivalence.transitive eq (Equivalence.transitive eq (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Equivalence eq
|
||||
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' = Equivalence.transitive eq (symmetric *Associative) (Equivalence.transitive eq p (Ring.timesZero R))
|
||||
IntegralDomain.intDom orderedFieldIsIntDom {a} {b} ab=0 | inl (inr x) = inr (Equivalence.transitive eq (Equivalence.transitive eq (symmetric identIsIdent) (*WellDefined q reflexive)) p')
|
||||
where
|
||||
open Equivalence eq
|
||||
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' = Equivalence.transitive eq (symmetric *Associative) (Equivalence.transitive eq p (Ring.timesZero R))
|
||||
IntegralDomain.intDom orderedFieldIsIntDom {a} {b} ab=0 | inr x = inl (Equivalence.symmetric (Setoid.eq S) x)
|
||||
IntegralDomain.intDom orderedFieldIsIntDom = decidedIntDom R orderedFieldIntDom
|
||||
IntegralDomain.nontrivial orderedFieldIsIntDom pr = Field.nontrivial F (Equivalence.symmetric (Setoid.eq S) pr)
|
||||
|
@@ -9,7 +9,6 @@ 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
|
||||
|
@@ -10,7 +10,6 @@ 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
|
||||
|
@@ -7,6 +7,7 @@ open import Setoids.Subset
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Groups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Subgroups.Normal.Definition
|
||||
|
||||
@@ -45,3 +46,7 @@ Group.identRight (cosetGroup norm) = isSubset (symmetric (transitive +Associativ
|
||||
Group.identLeft (cosetGroup norm) = isSubset (symmetric (transitive (+WellDefined reflexive identLeft) invLeft)) containsIdentity
|
||||
Group.invLeft (cosetGroup norm) = isSubset (symmetric (transitive (+WellDefined reflexive invLeft) invLeft)) containsIdentity
|
||||
Group.invRight (cosetGroup norm) = isSubset (symmetric (transitive (+WellDefined reflexive invRight) invLeft)) containsIdentity
|
||||
|
||||
cosetGroupHom : (norm : normalSubgroup G subgrp) → GroupHom G (cosetGroup norm) id
|
||||
GroupHom.groupHom (cosetGroupHom norm) = isSubset (symmetric (transitive (+WellDefined invContravariant reflexive) (transitive +Associative (transitive (+WellDefined (transitive (symmetric +Associative) (+WellDefined reflexive invLeft)) reflexive) (transitive (+WellDefined identRight reflexive) invLeft))))) (Subgroup.containsIdentity subgrp)
|
||||
GroupHom.wellDefined (cosetGroupHom norm) {x} {y} x=y = isSubset (symmetric (transitive (+WellDefined reflexive x=y) invLeft)) (Subgroup.containsIdentity subgrp)
|
||||
|
@@ -18,7 +18,6 @@ open import Groups.Subgroups.Normal.Definition
|
||||
open import Groups.Subgroups.Normal.Lemmas
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Abelian.Definition
|
||||
open import Groups.QuotientGroup.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
@@ -27,6 +26,7 @@ module Groups.FirstIsomorphismTheorem {a b c d : _} {A : Set a} {B : Set b} {S :
|
||||
open import Groups.Homomorphisms.Image fHom
|
||||
open import Groups.Homomorphisms.Kernel fHom
|
||||
open import Groups.Cosets G groupKernelIsSubgroup
|
||||
open import Groups.QuotientGroup.Definition G fHom
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
open import Setoids.Subset T
|
||||
@@ -39,7 +39,7 @@ groupFirstIsomorphismTheoremWellDefined {x} {y} pr = transitive (symmetric (invT
|
||||
u : Setoid._∼_ T (Group.inverse H (Group.inverse H (f y))) (Group.inverse H (Group.inverse H (f x)))
|
||||
u = inverseWellDefined H (transferToRight' H t)
|
||||
|
||||
groupFirstIsomorphismTheorem : GroupsIsomorphic (cosetGroup groupKernelIsNormalSubgroup) (subgroupIsGroup H imageGroupSubset imageGroupSubgroup)
|
||||
groupFirstIsomorphismTheorem : GroupsIsomorphic (cosetGroup groupKernelIsNormalSubgroup) (subgroupIsGroup H imageGroupSubgroup)
|
||||
GroupsIsomorphic.isomorphism groupFirstIsomorphismTheorem x = f x , (x , reflexive)
|
||||
GroupHom.groupHom (GroupIso.groupHom (GroupsIsomorphic.proof groupFirstIsomorphismTheorem)) {x} {y} = GroupHom.groupHom fHom
|
||||
GroupHom.wellDefined (GroupIso.groupHom (GroupsIsomorphic.proof groupFirstIsomorphismTheorem)) {x} {y} = groupFirstIsomorphismTheoremWellDefined {x} {y}
|
||||
@@ -47,3 +47,15 @@ SetoidInjection.wellDefined (SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic
|
||||
SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem))) fx=fy = transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H reflexive fx=fy) (transitive (symmetric (GroupHom.groupHom fHom)) (transitive (GroupHom.wellDefined fHom (Group.invLeft G)) (imageOfIdentityIsIdentity fHom))))
|
||||
SetoidSurjection.wellDefined (SetoidBijection.surj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem))) = groupFirstIsomorphismTheoremWellDefined
|
||||
SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem))) {b , (a , fa=b)} = a , fa=b
|
||||
|
||||
groupFirstIsomorphismTheoremWellDefined' : {x y : A} → f (x +G (Group.inverse G y)) ∼ Group.0G H → Setoid._∼_ (subsetSetoid imageGroupSubset) (f x , (x , reflexive)) (f y , (y , reflexive))
|
||||
groupFirstIsomorphismTheoremWellDefined' {x} {y} pr = transferToRight H (transitive (symmetric (transitive (GroupHom.groupHom fHom) (Group.+WellDefined H reflexive (homRespectsInverse fHom)))) pr)
|
||||
|
||||
groupFirstIsomorphismTheorem' : GroupsIsomorphic (quotientGroupByHom) (subgroupIsGroup H imageGroupSubgroup)
|
||||
GroupsIsomorphic.isomorphism groupFirstIsomorphismTheorem' a = f a , (a , reflexive)
|
||||
GroupHom.groupHom (GroupIso.groupHom (GroupsIsomorphic.proof groupFirstIsomorphismTheorem')) {x} {y} = GroupHom.groupHom fHom
|
||||
GroupHom.wellDefined (GroupIso.groupHom (GroupsIsomorphic.proof groupFirstIsomorphismTheorem')) {x} {y} = groupFirstIsomorphismTheoremWellDefined'
|
||||
SetoidInjection.wellDefined (SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem'))) {x} {y} = groupFirstIsomorphismTheoremWellDefined' {x} {y}
|
||||
SetoidInjection.injective (SetoidBijection.inj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem'))) {x} {y} fx=fy = transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H reflexive (homRespectsInverse fHom)) (transferToRight'' H fx=fy))
|
||||
SetoidSurjection.wellDefined (SetoidBijection.surj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem'))) {x} {y} = groupFirstIsomorphismTheoremWellDefined'
|
||||
SetoidSurjection.surjective (SetoidBijection.surj (GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem'))) {b , (a , fa=b)} = a , fa=b
|
||||
|
@@ -43,50 +43,6 @@ fourWay+Associative'' {S = S} {_·_ = _·_} G = transitive +Associative (symmetr
|
||||
open Setoid S
|
||||
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) (invIdent 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
|
||||
|
||||
quotientGroupLemma : {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 y : A} → Setoid._∼_ T (underf x) (underf y) → Setoid._∼_ (quotientGroupSetoid G f) x y
|
||||
quotientGroupLemma {S = S} {T = T} G {H = H} fHom {x} {y} fx=fy = transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H (Equivalence.reflexive (Setoid.eq T)) (homRespectsInverse fHom)) (transferToRight'' H fx=fy))
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
quotientGroupLemma' : {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 y : A} → Setoid._∼_ (quotientGroupSetoid G f) x y → Setoid._∼_ T (underf x) (underf y)
|
||||
quotientGroupLemma' {S = S} {T = T} G {H = H} f fx=fy = transferToRight H (transitive (transitive (Group.+WellDefined H (Equivalence.reflexive (Setoid.eq T)) (symmetric (homRespectsInverse f))) (symmetric (GroupHom.groupHom f))) fx=fy)
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
{-
|
||||
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 = {!!}
|
||||
|
@@ -7,20 +7,48 @@ open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Rings.Definition
|
||||
open import Fields.FieldOfFractions.Setoid
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Subgroups.Normal.Definition
|
||||
|
||||
module Groups.QuotientGroup.Definition where
|
||||
module Groups.QuotientGroup.Definition {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) where
|
||||
|
||||
quotientGroupByHom : {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 (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G {H = H} {underf = f} fHom) {x} {y} {m} {n} x~m y~n = ans
|
||||
quotientGroupSetoid : (Setoid {a} {d} A)
|
||||
quotientGroupSetoid = 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) (invIdent 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
|
||||
|
||||
|
||||
quotientGroupByHom : Group (quotientGroupSetoid) _·A_
|
||||
Group.+WellDefined (quotientGroupByHom) {x} {y} {m} {n} x~m y~n = ans
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence (Setoid.eq T)
|
||||
@@ -42,15 +70,15 @@ Group.+WellDefined (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} {_·B_ =
|
||||
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 (quotientGroupByHom G fHom) = Group.0G G
|
||||
Group.inverse (quotientGroupByHom G fHom) = Group.inverse G
|
||||
Group.+Associative (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} {b} {c} = ans
|
||||
Group.0G (quotientGroupByHom) = Group.0G G
|
||||
Group.inverse (quotientGroupByHom) = Group.inverse G
|
||||
Group.+Associative (quotientGroupByHom) {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 (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans
|
||||
Group.identRight (quotientGroupByHom) {a} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
@@ -58,7 +86,7 @@ Group.identRight (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {
|
||||
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 (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {a} = ans
|
||||
Group.identLeft (quotientGroupByHom) {a} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
@@ -66,14 +94,14 @@ Group.identLeft (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {u
|
||||
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 (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans
|
||||
Group.invLeft (quotientGroupByHom) {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) (invIdent G)) (Equivalence.reflexive (Setoid.eq S)) ((Equivalence.reflexive (Setoid.eq S))) ((Equivalence.reflexive (Setoid.eq S)))) (identRight {0G}))) (imageOfIdentityIsIdentity fHom)
|
||||
Group.invRight (quotientGroupByHom {S = S} {T = T} {_·A_ = _·A_} G {H = H} {underf = f} fHom) {x} = ans
|
||||
Group.invRight (quotientGroupByHom) {x} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
|
@@ -9,13 +9,28 @@ open import Groups.FiniteGroups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Abelian.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Fields.FieldOfFractions.Setoid
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Lemmas
|
||||
open import Groups.QuotientGroup.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
|
||||
module Groups.QuotientGroup.Lemmas {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) where
|
||||
|
||||
quotientGroupLemma : {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 y : A} → Setoid._∼_ T (underf x) (underf y) → Setoid._∼_ (quotientGroupSetoid G f) x y
|
||||
quotientGroupLemma {S = S} {T = T} G {H = H} fHom {x} {y} fx=fy = transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H (Equivalence.reflexive (Setoid.eq T)) (homRespectsInverse fHom)) (transferToRight'' H fx=fy))
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
quotientGroupLemma' : {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 y : A} → Setoid._∼_ (quotientGroupSetoid G f) x y → Setoid._∼_ T (underf x) (underf y)
|
||||
quotientGroupLemma' {S = S} {T = T} G {H = H} f fx=fy = transferToRight H (transitive (transitive (Group.+WellDefined H (Equivalence.reflexive (Setoid.eq T)) (symmetric (homRespectsInverse f))) (symmetric (GroupHom.groupHom f))) fx=fy)
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
|
||||
projectionMapIsGroupHom : GroupHom G (quotientGroupByHom G fHom) id
|
||||
GroupHom.groupHom projectionMapIsGroupHom {x} {y} = quotientGroupLemma G fHom (Equivalence.reflexive (Setoid.eq T))
|
||||
GroupHom.wellDefined projectionMapIsGroupHom x=y = quotientGroupLemma G fHom (GroupHom.wellDefined fHom x=y)
|
||||
|
@@ -24,12 +24,12 @@ record Subgroup {c : _} (pred : A → Set c) : Set (a ⊔ b ⊔ c) where
|
||||
subgroupOp : {c : _} {pred : A → Set c} → (s : Subgroup pred) → Sg A pred → Sg A pred → Sg A pred
|
||||
subgroupOp {pred = pred} record { closedUnderPlus = one } (a , prA) (b , prB) = (a + b) , one prA prB
|
||||
|
||||
subgroupIsGroup : {c : _} {pred : A → Set c} → (subs : subset pred) → (s : Subgroup pred) → Group (subsetSetoid subs) (subgroupOp s)
|
||||
Group.+WellDefined (subgroupIsGroup _ s) {m , prM} {n , prN} {x , prX} {y , prY} m=x n=y = +WellDefined m=x n=y
|
||||
Group.0G (subgroupIsGroup _ record { containsIdentity = two }) = 0G , two
|
||||
Group.inverse (subgroupIsGroup _ record { closedUnderInverse = three }) (a , prA) = (inverse a) , three prA
|
||||
Group.+Associative (subgroupIsGroup _ s) {a , prA} {b , prB} {c , prC} = +Associative
|
||||
Group.identRight (subgroupIsGroup _ s) {a , prA} = identRight
|
||||
Group.identLeft (subgroupIsGroup _ s) {a , prA} = identLeft
|
||||
Group.invLeft (subgroupIsGroup _ s) {a , prA} = invLeft
|
||||
Group.invRight (subgroupIsGroup _ s) {a , prA} = invRight
|
||||
subgroupIsGroup : {c : _} {pred : A → Set c} → (s : Subgroup pred) → Group (subsetSetoid (Subgroup.isSubset s)) (subgroupOp s)
|
||||
Group.+WellDefined (subgroupIsGroup s) {m , prM} {n , prN} {x , prX} {y , prY} m=x n=y = +WellDefined m=x n=y
|
||||
Group.0G (subgroupIsGroup record { containsIdentity = two }) = 0G , two
|
||||
Group.inverse (subgroupIsGroup record { closedUnderInverse = three }) (a , prA) = (inverse a) , three prA
|
||||
Group.+Associative (subgroupIsGroup s) {a , prA} {b , prB} {c , prC} = +Associative
|
||||
Group.identRight (subgroupIsGroup s) {a , prA} = identRight
|
||||
Group.identLeft (subgroupIsGroup s) {a , prA} = identLeft
|
||||
Group.invLeft (subgroupIsGroup s) {a , prA} = invLeft
|
||||
Group.invRight (subgroupIsGroup s) {a , prA} = invRight
|
||||
|
@@ -11,6 +11,7 @@ open import Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.QuotientGroup.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Actions.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
@@ -9,7 +9,6 @@ open import Semirings.Definition
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Rings.IntegralDomains
|
||||
|
||||
module Numbers.Integers.Multiplication where
|
||||
|
||||
|
@@ -8,15 +8,15 @@ open import Semirings.Definition
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Rings.IntegralDomains
|
||||
open import Rings.IntegralDomains.Definition
|
||||
|
||||
module Numbers.Integers.RingStructure.IntegralDomain where
|
||||
|
||||
intDom : (a b : ℤ) → a *Z b ≡ nonneg 0 → (a ≡ nonneg 0) || (b ≡ nonneg 0)
|
||||
intDom (nonneg zero) (nonneg b) pr = inl refl
|
||||
intDom (nonneg (succ a)) (nonneg zero) pr = inr refl
|
||||
intDom (nonneg zero) (negSucc b) pr = inl refl
|
||||
intDom (negSucc a) (nonneg zero) pr = inr refl
|
||||
intDom : (a b : ℤ) → a *Z b ≡ nonneg 0 → (a ≡ nonneg 0 → False) → (b ≡ nonneg 0)
|
||||
intDom (nonneg zero) (nonneg b) x a!=0 = exFalso (a!=0 x)
|
||||
intDom (nonneg (succ a)) (nonneg zero) a!=0 _ = refl
|
||||
intDom (nonneg zero) (negSucc b) pr a!=0 = exFalso (a!=0 pr)
|
||||
intDom (negSucc a) (nonneg zero) _ _ = refl
|
||||
|
||||
ℤIntDom : IntegralDomain ℤRing
|
||||
IntegralDomain.intDom ℤIntDom {a} {b} = intDom a b
|
||||
|
@@ -7,7 +7,6 @@ open import Semirings.Definition
|
||||
open import Groups.Definition
|
||||
open import Rings.Definition
|
||||
open import Setoids.Setoids
|
||||
open import Rings.IntegralDomains
|
||||
|
||||
module Numbers.Integers.RingStructure.Ring where
|
||||
|
||||
|
@@ -47,3 +47,8 @@ Ring.*Associative cosetRing = isSubset (transitive (symmetric invLeft) (+WellDef
|
||||
Ring.*Commutative cosetRing {a} {b} = isSubset (transitive (symmetric invLeft) (+WellDefined (inverseWellDefined additiveGroup *Commutative) reflexive)) containsIdentity
|
||||
Ring.*DistributesOver+ cosetRing = isSubset (symmetric (transitive (+WellDefined (inverseWellDefined additiveGroup (symmetric *DistributesOver+)) reflexive) invLeft)) containsIdentity
|
||||
Ring.identIsIdent cosetRing = isSubset (symmetric (transitive (Group.+WellDefined additiveGroup reflexive identIsIdent) invLeft)) containsIdentity
|
||||
|
||||
cosetRingHom : RingHom R cosetRing id
|
||||
RingHom.preserves1 cosetRingHom = isSubset (symmetric invLeft) (Ideal.containsIdentity ideal)
|
||||
RingHom.ringHom cosetRingHom = isSubset (symmetric invLeft) (Ideal.containsIdentity ideal)
|
||||
RingHom.groupHom cosetRingHom = cosetGroupHom (abelianGroupSubgroupIsNormal (Ideal.isSubgroup ideal) abelianUnderlyingGroup)
|
||||
|
37
Rings/EuclideanDomains/Definition.agda
Normal file
37
Rings/EuclideanDomains/Definition.agda
Normal file
@@ -0,0 +1,37 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Rings.IntegralDomains.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.EuclideanDomains.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A → A → A} (R : Ring S _+_ _*_) where
|
||||
|
||||
open Setoid S
|
||||
open Ring R
|
||||
|
||||
record DivisionAlgorithmResult {norm : {a : A} → ((a ∼ 0R) → False) → ℕ} {x y : A} (x!=0 : (x ∼ 0R) → False) (y!=0 : (y ∼ 0R) → False) : Set (a ⊔ b) where
|
||||
field
|
||||
quotient : A
|
||||
rem : A
|
||||
remSmall : (rem ∼ 0R) || Sg ((rem ∼ 0R) → False) (λ rem!=0 → (norm rem!=0) <N (norm y!=0))
|
||||
divAlg : x ∼ ((quotient * y) + rem)
|
||||
|
||||
record EuclideanDomain : Set (a ⊔ lsuc b) where
|
||||
field
|
||||
isIntegralDomain : IntegralDomain R
|
||||
norm : {a : A} → ((a ∼ 0R) → False) → ℕ
|
||||
normSize : {a b : A} → (a!=0 : (a ∼ 0R) → False) → (b!=0 : (b ∼ 0R) → False) → (c : A) → b ∼ (a * c) → (norm a!=0) ≤N (norm b!=0)
|
||||
divisionAlg : {a b : A} → (a!=0 : (a ∼ 0R) → False) → (b!=0 : (b ∼ 0R) → False) → DivisionAlgorithmResult {norm} a!=0 b!=0
|
39
Rings/EuclideanDomains/Examples.agda
Normal file
39
Rings/EuclideanDomains/Examples.agda
Normal file
@@ -0,0 +1,39 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Definition
|
||||
open import Numbers.Naturals.Order
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Rings.IntegralDomains.Examples
|
||||
open import Rings.EuclideanDomains.Definition
|
||||
open import Fields.Fields
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.EuclideanDomains.Examples where
|
||||
|
||||
polynomialField : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A → A → A} {R : Ring S _+_ _*_} (F : Field R) → (Setoid._∼_ S (Ring.1R R) (Ring.0R R) → False) → EuclideanDomain R
|
||||
EuclideanDomain.isIntegralDomain (polynomialField F 1!=0) = fieldIsIntDom F 1!=0
|
||||
EuclideanDomain.norm (polynomialField F _) a!=0 = zero
|
||||
EuclideanDomain.normSize (polynomialField F _) a!=0 b!=0 c b=ac = inr refl
|
||||
DivisionAlgorithmResult.quotient (EuclideanDomain.divisionAlg (polynomialField {_*_ = _*_} F _) {a = a} {b} a!=0 b!=0) with Field.allInvertible F b b!=0
|
||||
... | bInv , prB = a * bInv
|
||||
DivisionAlgorithmResult.rem (EuclideanDomain.divisionAlg (polynomialField F _) a!=0 b!=0) = Field.0F F
|
||||
DivisionAlgorithmResult.remSmall (EuclideanDomain.divisionAlg (polynomialField {S = S} F _) a!=0 b!=0) = inl (Equivalence.reflexive (Setoid.eq S))
|
||||
DivisionAlgorithmResult.divAlg (EuclideanDomain.divisionAlg (polynomialField {S = S} {R = R} F _) {a = a} {b = b} a!=0 b!=0) with Field.allInvertible F b b!=0
|
||||
... | bInv , prB = transitive (transitive (transitive (symmetric identIsIdent) (transitive *Commutative (*WellDefined reflexive (symmetric prB)))) *Associative) (symmetric identRight)
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Ring R
|
||||
open Group additiveGroup
|
@@ -13,7 +13,6 @@ open import Rings.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
-- Following Part IB's course Groups, Rings, and Modules, we take rings to be commutative with one.
|
||||
module Rings.Homomorphisms.Definition where
|
||||
|
||||
record RingHom {m n o p : _} {A : Set m} {B : Set n} {SA : Setoid {m} {o} A} {SB : Setoid {n} {p} B} {_+A_ : A → A → A} {_*A_ : A → A → A} (R : Ring SA _+A_ _*A_) {_+B_ : B → B → B} {_*B_ : B → B → B} (S : Ring SB _+B_ _*B_) (f : A → B) : Set (m ⊔ n ⊔ o ⊔ p) where
|
||||
|
29
Rings/Homomorphisms/Image.agda
Normal file
29
Rings/Homomorphisms/Image.agda
Normal file
@@ -0,0 +1,29 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.Homomorphisms.Image {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_+A_ _*A_ : A → A → A} {_+B_ _*B_ : B → B → B} {R1 : Ring S _+A_ _*A_} {R2 : Ring T _+B_ _*B_} {f : A → B} (hom : RingHom R1 R2 f) where
|
||||
|
||||
open import Groups.Homomorphisms.Image (RingHom.groupHom hom)
|
||||
open import Rings.Subrings.Definition
|
||||
|
||||
imageGroupSubring : Subring R2 imageGroupPred
|
||||
Subring.isSubgroup imageGroupSubring = imageGroupSubgroup
|
||||
Subring.containsOne imageGroupSubring = Ring.1R R1 , RingHom.preserves1 hom
|
||||
Subring.closedUnderProduct imageGroupSubring {x} {y} (a , fa=x) (b , fb=y) = (a *A b) , transitive ringHom (Ring.*WellDefined R2 fa=x fb=y)
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
open RingHom hom
|
21
Rings/Homomorphisms/Kernel.agda
Normal file
21
Rings/Homomorphisms/Kernel.agda
Normal file
@@ -0,0 +1,21 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.Homomorphisms.Kernel {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_+1_ _*1_ : A → A → A} {_+2_ _*2_ : B → B → B} {R1 : Ring S _+1_ _*1_} {R2 : Ring T _+2_ _*2_} {f : A → B} (fHom : RingHom R1 R2 f) where
|
||||
|
||||
ringKernel : Set (a ⊔ d)
|
||||
ringKernel = Sg A (λ a → Setoid._∼_ T (f a) (Ring.0R R2))
|
@@ -19,9 +19,6 @@ module Rings.Ideals.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_
|
||||
|
||||
open import Groups.Subgroups.Definition (Ring.additiveGroup R)
|
||||
|
||||
ringKernel : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C → C → C} (R2 : Ring T _+2_ _*2_) {f : A → C} (fHom : RingHom R R2 f) → Set (a ⊔ d)
|
||||
ringKernel {T = T} R2 {f} fHom = Sg A (λ a → Setoid._∼_ T (f a) (Ring.0R R2))
|
||||
|
||||
record Ideal {c : _} (pred : A → Set c) : Set (a ⊔ b ⊔ c) where
|
||||
field
|
||||
isSubgroup : Subgroup pred
|
||||
@@ -30,31 +27,3 @@ record Ideal {c : _} (pred : A → Set c) : Set (a ⊔ b ⊔ c) where
|
||||
closedUnderInverse = Subgroup.closedUnderInverse isSubgroup
|
||||
containsIdentity = Subgroup.containsIdentity isSubgroup
|
||||
isSubset = Subgroup.isSubset isSubgroup
|
||||
|
||||
idealPredForKernel : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C → C → C} (R2 : Ring T _+2_ _*2_) {f : A → C} (fHom : RingHom R R2 f) → A → Set d
|
||||
idealPredForKernel {T = T} R2 {f} fHom a = Setoid._∼_ T (f a) (Ring.0R R2)
|
||||
|
||||
idealPredForKernelWellDefined : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C → C → C} (R2 : Ring T _+2_ _*2_) {f : A → C} (fHom : RingHom R R2 f) → {x y : A} → (Setoid._∼_ S x y) → (idealPredForKernel R2 fHom x → idealPredForKernel R2 fHom y)
|
||||
idealPredForKernelWellDefined {T = T} R2 {f} fHom a x=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined (RingHom.groupHom fHom) (Equivalence.symmetric (Setoid.eq S) a)) x=0
|
||||
|
||||
kernelIdealIsIdeal : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C → C → C} {R2 : Ring T _+2_ _*2_} {f : A → C} (fHom : RingHom R R2 f) → Ideal (idealPredForKernel R2 fHom)
|
||||
Subgroup.isSubset (Ideal.isSubgroup (kernelIdealIsIdeal {R2 = R2} fHom)) = idealPredForKernelWellDefined R2 fHom
|
||||
Subgroup.closedUnderPlus (Ideal.isSubgroup (kernelIdealIsIdeal {T = T} {R2 = R2} fHom)) {x} {y} fx=0 fy=0 = transitive (transitive (GroupHom.groupHom (RingHom.groupHom fHom)) (+WellDefined fx=0 fy=0)) identLeft
|
||||
where
|
||||
open Ring R2
|
||||
open Group (Ring.additiveGroup R2)
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
Subgroup.containsIdentity (Ideal.isSubgroup (kernelIdealIsIdeal fHom)) = imageOfIdentityIsIdentity (RingHom.groupHom fHom)
|
||||
Subgroup.closedUnderInverse (Ideal.isSubgroup (kernelIdealIsIdeal {T = T} {R2 = R2} fHom)) {x} fx=0 = zeroImpliesInverseZero (RingHom.groupHom fHom) fx=0
|
||||
where
|
||||
open Ring R2
|
||||
open Group (Ring.additiveGroup R2)
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
Ideal.accumulatesTimes (kernelIdealIsIdeal {T = T} {R2 = R2} {f = f} fHom) {x} {y} fx=0 = transitive (RingHom.ringHom fHom) (transitive (Ring.*WellDefined R2 fx=0 reflexive) (transitive (Ring.*Commutative R2) (Ring.timesZero R2 {f y})))
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
-- TODO : define the quotient by an ideal; note that the result is a ring
|
||||
|
34
Rings/Ideals/FirstIsomorphismTheorem.agda
Normal file
34
Rings/Ideals/FirstIsomorphismTheorem.agda
Normal file
@@ -0,0 +1,34 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Rings.Subrings.Definition
|
||||
open import Rings.Isomorphisms.Definition
|
||||
open import Groups.Isomorphisms.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.Ideals.FirstIsomorphismTheorem {a b c d : _} {A : Set a} {B : Set c} {S : Setoid {a} {b} A} {T : Setoid {c} {d} B} {_+A_ _*A_ : A → A → A} {_+B_ _*B_ : B → B → B} {R1 : Ring S _+A_ _*A_} {R2 : Ring T _+B_ _*B_} {f : A → B} (hom : RingHom R1 R2 f) where
|
||||
|
||||
open import Rings.Quotients.Definition R1 R2 hom
|
||||
open import Rings.Homomorphisms.Image hom
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
open import Groups.FirstIsomorphismTheorem (RingHom.groupHom hom)
|
||||
|
||||
ringFirstIsomorphismTheorem' : RingsIsomorphic quotientByRingHom (subringIsRing R2 imageGroupSubring)
|
||||
RingsIsomorphic.f ringFirstIsomorphismTheorem' a = f a , (a , reflexive)
|
||||
RingHom.preserves1 (RingIso.ringHom (RingsIsomorphic.iso ringFirstIsomorphismTheorem')) = RingHom.preserves1 hom
|
||||
RingHom.ringHom (RingIso.ringHom (RingsIsomorphic.iso ringFirstIsomorphismTheorem')) {r} {s} = RingHom.ringHom hom
|
||||
RingHom.groupHom (RingIso.ringHom (RingsIsomorphic.iso ringFirstIsomorphismTheorem')) = GroupIso.groupHom (GroupsIsomorphic.proof (groupFirstIsomorphismTheorem'))
|
||||
RingIso.bijective (RingsIsomorphic.iso ringFirstIsomorphismTheorem') = GroupIso.bij (GroupsIsomorphic.proof groupFirstIsomorphismTheorem')
|
58
Rings/Ideals/Lemmas.agda
Normal file
58
Rings/Ideals/Lemmas.agda
Normal file
@@ -0,0 +1,58 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Rings.Homomorphisms.Kernel
|
||||
open import Rings.Cosets
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.Ideals.Lemmas {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A → A → A} (R : Ring S _+_ _*_) where
|
||||
|
||||
open import Rings.Ideals.Definition R
|
||||
|
||||
idealPredForKernel : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C → C → C} (R2 : Ring T _+2_ _*2_) {f : A → C} (fHom : RingHom R R2 f) → A → Set d
|
||||
idealPredForKernel {T = T} R2 {f} fHom a = Setoid._∼_ T (f a) (Ring.0R R2)
|
||||
|
||||
idealPredForKernelWellDefined : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C → C → C} (R2 : Ring T _+2_ _*2_) {f : A → C} (fHom : RingHom R R2 f) → {x y : A} → (Setoid._∼_ S x y) → (idealPredForKernel R2 fHom x → idealPredForKernel R2 fHom y)
|
||||
idealPredForKernelWellDefined {T = T} R2 {f} fHom a x=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined (RingHom.groupHom fHom) (Equivalence.symmetric (Setoid.eq S) a)) x=0
|
||||
|
||||
kernelIdealIsIdeal : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C → C → C} {R2 : Ring T _+2_ _*2_} {f : A → C} (fHom : RingHom R R2 f) → Ideal (idealPredForKernel R2 fHom)
|
||||
Subgroup.isSubset (Ideal.isSubgroup (kernelIdealIsIdeal {R2 = R2} fHom)) = idealPredForKernelWellDefined R2 fHom
|
||||
Subgroup.closedUnderPlus (Ideal.isSubgroup (kernelIdealIsIdeal {T = T} {R2 = R2} fHom)) {x} {y} fx=0 fy=0 = transitive (transitive (GroupHom.groupHom (RingHom.groupHom fHom)) (+WellDefined fx=0 fy=0)) identLeft
|
||||
where
|
||||
open Ring R2
|
||||
open Group (Ring.additiveGroup R2)
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
Subgroup.containsIdentity (Ideal.isSubgroup (kernelIdealIsIdeal fHom)) = imageOfIdentityIsIdentity (RingHom.groupHom fHom)
|
||||
Subgroup.closedUnderInverse (Ideal.isSubgroup (kernelIdealIsIdeal {T = T} {R2 = R2} fHom)) {x} fx=0 = zeroImpliesInverseZero (RingHom.groupHom fHom) fx=0
|
||||
where
|
||||
open Ring R2
|
||||
open Group (Ring.additiveGroup R2)
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
Ideal.accumulatesTimes (kernelIdealIsIdeal {T = T} {R2 = R2} {f = f} fHom) {x} {y} fx=0 = transitive (RingHom.ringHom fHom) (transitive (Ring.*WellDefined R2 fx=0 reflexive) (transitive (Ring.*Commutative R2) (Ring.timesZero R2 {f y})))
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
open Setoid S
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open Equivalence eq
|
||||
open import Groups.Lemmas additiveGroup
|
||||
|
||||
idealIsKernelMap : {c : _} {pred : A → Set c} → (i : Ideal pred) → {x : A} → pred x → ringKernel {R1 = R} {R2 = cosetRing R i} (cosetRingHom R i)
|
||||
idealIsKernelMap {c} {pred} i {x} predX = x , (Ideal.isSubset i (transitive (symmetric identLeft) (+WellDefined (symmetric invIdent) reflexive)) predX)
|
22
Rings/Ideals/Prime/Definition.agda
Normal file
22
Rings/Ideals/Prime/Definition.agda
Normal file
@@ -0,0 +1,22 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Fields.Fields
|
||||
open import Rings.Ideals.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.Ideals.Prime.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A → A → A} {R : Ring S _+_ _*_} {c : _} {pred : A → Set c} (i : Ideal R pred) where
|
||||
|
||||
PrimeIdeal : Set (a ⊔ c)
|
||||
PrimeIdeal = {a b : A} → pred (a * b) → ((pred a) → False) → pred b
|
24
Rings/Ideals/Principal/Definition.agda
Normal file
24
Rings/Ideals/Principal/Definition.agda
Normal file
@@ -0,0 +1,24 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.Ideals.Principal.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A → A → A} (R : Ring S _+_ _*_) where
|
||||
|
||||
open import Rings.Ideals.Definition R
|
||||
open Setoid S
|
||||
|
||||
PrincipalIdeal : {c : _} {pred : A → Set c} (ideal : Ideal pred) → Set (a ⊔ b ⊔ c)
|
||||
PrincipalIdeal {pred = pred} ideal = Sg A (λ a → {x : A} → (pred x) → Sg A (λ c → (a * c) ∼ x))
|
@@ -1,37 +0,0 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.IntegralDomains where
|
||||
record IntegralDomain {n m : _} {A : Set n} {S : Setoid {n} {m} A} {_+_ : A → A → A} {_*_ : A → A → A} (R : Ring S _+_ _*_) : Set (lsuc m ⊔ n) where
|
||||
field
|
||||
intDom : {a b : A} → Setoid._∼_ S (a * b) (Ring.0R R) → (Setoid._∼_ S a (Ring.0R R)) || (Setoid._∼_ S b (Ring.0R R))
|
||||
nontrivial : Setoid._∼_ S (Ring.1R R) (Ring.0R R) → False
|
||||
|
||||
cancelIntDom : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_+_ : A → A → A} {_*_ : A → A → A} {R : Ring S _+_ _*_} (I : IntegralDomain R) {a b c : A} → Setoid._∼_ S (a * b) (a * c) → (Setoid._∼_ S a (Ring.0R R)) || (Setoid._∼_ S b c)
|
||||
cancelIntDom {S = S} {_+_ = _+_} {_*_ = _*_} {R = R} I {a} {b} {c} ab=ac = t4
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
t1 : (a * b) + Group.inverse (Ring.additiveGroup R) (a * c) ∼ Ring.0R R
|
||||
t1 = transferToRight'' (Ring.additiveGroup R) ab=ac
|
||||
t2 : a * (b + Group.inverse (Ring.additiveGroup R) c) ∼ Ring.0R R
|
||||
t2 = transitive (transitive (Ring.*DistributesOver+ R) (Group.+WellDefined (Ring.additiveGroup R) reflexive (transferToRight' (Ring.additiveGroup R) (transitive (symmetric (Ring.*DistributesOver+ R)) (transitive (Ring.*WellDefined R reflexive (Group.invLeft (Ring.additiveGroup R))) (Ring.timesZero R)))))) t1
|
||||
t3 : (a ∼ Ring.0R R) || ((b + Group.inverse (Ring.additiveGroup R) c) ∼ Ring.0R R)
|
||||
t3 = IntegralDomain.intDom I t2
|
||||
t4 : (a ∼ Ring.0R R) || (b ∼ c)
|
||||
t4 with t3
|
||||
... | inl t = inl t
|
||||
... | inr b = inr (transferToRight (Ring.additiveGroup R) b)
|
41
Rings/IntegralDomains/Definition.agda
Normal file
41
Rings/IntegralDomains/Definition.agda
Normal file
@@ -0,0 +1,41 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.IntegralDomains.Definition {m n : _} {A : Set n} {S : Setoid {n} {m} A} {_+_ _*_ : A → A → A} (R : Ring S _+_ _*_) where
|
||||
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Ring R
|
||||
|
||||
record IntegralDomain : Set (lsuc m ⊔ n) where
|
||||
field
|
||||
intDom : {a b : A} → (a * b) ∼ (Ring.0R R) → ((a ∼ (Ring.0R R)) → False) → b ∼ (Ring.0R R)
|
||||
nontrivial : Setoid._∼_ S (Ring.1R R) (Ring.0R R) → False
|
||||
|
||||
decidedIntDom : ({a b : A} → (a * b) ∼ (Ring.0R R) → (a ∼ 0R) || (b ∼ 0R)) → ({a b : A} → (a * b) ∼ 0R → ((a ∼ (Ring.0R R)) → False) → b ∼ (Ring.0R R))
|
||||
decidedIntDom f ab=0 a!=0 with f ab=0
|
||||
decidedIntDom f ab=0 a!=0 | inl x = exFalso (a!=0 x)
|
||||
decidedIntDom f ab=0 a!=0 | inr x = x
|
||||
|
||||
cancelIntDom : (I : IntegralDomain) {a b c : A} → (a * b) ∼ (a * c) → ((a ∼ (Ring.0R R)) → False) → (b ∼ c)
|
||||
cancelIntDom I {a} {b} {c} ab=ac a!=0 = transferToRight (Ring.additiveGroup R) t3
|
||||
where
|
||||
t1 : (a * b) + Group.inverse (Ring.additiveGroup R) (a * c) ∼ Ring.0R R
|
||||
t1 = transferToRight'' (Ring.additiveGroup R) ab=ac
|
||||
t2 : a * (b + Group.inverse (Ring.additiveGroup R) c) ∼ Ring.0R R
|
||||
t2 = transitive (transitive (Ring.*DistributesOver+ R) (Group.+WellDefined (Ring.additiveGroup R) reflexive (transferToRight' (Ring.additiveGroup R) (transitive (symmetric (Ring.*DistributesOver+ R)) (transitive (Ring.*WellDefined R reflexive (Group.invLeft (Ring.additiveGroup R))) (Ring.timesZero R)))))) t1
|
||||
t3 : (b + Group.inverse (Ring.additiveGroup R) c) ∼ Ring.0R R
|
||||
t3 = IntegralDomain.intDom I t2 a!=0
|
28
Rings/IntegralDomains/Examples.agda
Normal file
28
Rings/IntegralDomains/Examples.agda
Normal file
@@ -0,0 +1,28 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Rings.Definition
|
||||
open import Rings.Lemmas
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.IntegralDomains.Definition
|
||||
open import Fields.Fields
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.IntegralDomains.Examples where
|
||||
|
||||
fieldIsIntDom : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A → A → A} {R : Ring S _+_ _*_} (F : Field R) → (Setoid._∼_ S (Ring.1R R) (Ring.0R R) → False) → IntegralDomain R
|
||||
IntegralDomain.intDom (fieldIsIntDom F 1!=0) {a} {b} ab=0 a!=0 with Field.allInvertible F a a!=0
|
||||
IntegralDomain.intDom (fieldIsIntDom {S = S} {R = R} F _) {a} {b} ab=0 a!=0 | 1/a , prA = transitive (symmetric identIsIdent) (transitive (*WellDefined (symmetric prA) reflexive) (transitive (symmetric *Associative) (transitive (*WellDefined reflexive ab=0) timesZero)))
|
||||
where
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
open Ring R
|
||||
IntegralDomain.nontrivial (fieldIsIntDom F 1!=0) = 1!=0
|
@@ -19,4 +19,9 @@ module Rings.Isomorphisms.Definition {a b c d : _} {A : Set a} {S : Setoid {a} {
|
||||
record RingIso (f : A → B) : Set (a ⊔ b ⊔ c ⊔ d) where
|
||||
field
|
||||
ringHom : RingHom R1 R2 f
|
||||
bijective : Bijection f
|
||||
bijective : SetoidBijection S T f
|
||||
|
||||
record RingsIsomorphic : Set (a ⊔ b ⊔ c ⊔ d) where
|
||||
field
|
||||
f : A → B
|
||||
iso : RingIso f
|
||||
|
24
Rings/PrincipalIdealDomain.agda
Normal file
24
Rings/PrincipalIdealDomain.agda
Normal file
@@ -0,0 +1,24 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.PrincipalIdealDomain {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A → A → A} (R : Ring S _+_ _*_) where
|
||||
|
||||
open import Rings.Ideals.Principal.Definition R
|
||||
open import Rings.Ideals.Definition R
|
||||
|
||||
PrincipalIdealDomain : {c : _} → Set (a ⊔ b ⊔ lsuc c)
|
||||
PrincipalIdealDomain {c} = {pred : A → Set c} → (ideal : Ideal pred) → PrincipalIdeal ideal
|
44
Rings/Subrings/Definition.agda
Normal file
44
Rings/Subrings/Definition.agda
Normal file
@@ -0,0 +1,44 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Groups.Groups
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Definition
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Setoids.Orders
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Rings.Definition
|
||||
open import Rings.Homomorphisms.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Rings.Subrings.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+A_ _*A_ : A → A → A} (R : Ring S _+A_ _*A_) where
|
||||
|
||||
open import Setoids.Subset S
|
||||
open Ring R
|
||||
open Group additiveGroup
|
||||
open import Groups.Subgroups.Definition additiveGroup
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
record Subring {c : _} (pred : A → Set c) : Set (a ⊔ b ⊔ c) where
|
||||
field
|
||||
isSubgroup : Subgroup pred
|
||||
containsOne : pred 1R
|
||||
closedUnderProduct : {x y : A} → pred x → pred y → pred (x *A y)
|
||||
isSubset = Subgroup.isSubset isSubgroup
|
||||
|
||||
subringMult : {c : _} {pred : A → Set c} → (s : Subring pred) → Sg A pred → Sg A pred → Sg A pred
|
||||
subringMult s (a , prA) (b , prB) = (a *A b) , Subring.closedUnderProduct s prA prB
|
||||
|
||||
subringIsRing : {c : _} {pred : A → Set c} → (subring : Subring pred) → Ring (subsetSetoid (Subring.isSubset subring)) (subgroupOp (Subring.isSubgroup subring)) (subringMult subring)
|
||||
Ring.additiveGroup (subringIsRing sub) = subgroupIsGroup (Subring.isSubgroup sub)
|
||||
Ring.*WellDefined (subringIsRing sub) {r , prR} {s , prS} {t , prT} {u , prU} r=t s=u = *WellDefined r=t s=u
|
||||
Ring.1R (subringIsRing sub) = (1R , Subring.containsOne sub)
|
||||
Ring.groupIsAbelian (subringIsRing sub) {a , prA} {b , prB} = groupIsAbelian
|
||||
Ring.*Associative (subringIsRing sub) {a , prA} {b , prB} {c , prC} = *Associative
|
||||
Ring.*Commutative (subringIsRing sub) {a , prA} {b , prB} = *Commutative
|
||||
Ring.*DistributesOver+ (subringIsRing sub) {a , prA} {b , prB} {c , prC} = *DistributesOver+
|
||||
Ring.identIsIdent (subringIsRing sub) {a , prA} = identIsIdent
|
Reference in New Issue
Block a user