Lots of rings (#82)

This commit is contained in:
Patrick Stevens
2019-11-22 19:52:57 +00:00
committed by GitHub
parent b33baa5fb7
commit 660d7aa27c
40 changed files with 1246 additions and 881 deletions

View File

@@ -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

View 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
View 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)

View 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

View 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))