mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-11 14:48:42 +00:00
Missed one
This commit is contained in:
39
Rings/Quotients/Definition.agda
Normal file
39
Rings/Quotients/Definition.agda
Normal file
@@ -0,0 +1,39 @@
|
||||
{-# OPTIONS --warning=error --safe --without-K #-}
|
||||
|
||||
open import Functions
|
||||
open import Sets.FinSet
|
||||
open import LogicalFormulae
|
||||
open import Groups.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.FiniteGroups.Definition
|
||||
open import Rings.Homomorphisms.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.QuotientGroup.Definition
|
||||
|
||||
module Rings.Quotients.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 → A} {_+B_ _*B_ : B → B → B} (R : Ring S _+A_ _*A_) (R2 : Ring T _+B_ _*B_) {underf : A → B} (f : RingHom R R2 underf) where
|
||||
|
||||
open import Groups.QuotientGroup.Lemmas (Ring.additiveGroup R) (Ring.additiveGroup R2) (RingHom.groupHom f)
|
||||
|
||||
quotientRing : Ring (quotientGroupSetoid (Ring.additiveGroup R) (RingHom.groupHom f)) _+A_ _*A_
|
||||
Ring.additiveGroup quotientRing = quotientGroupByHom (Ring.additiveGroup R) (RingHom.groupHom f)
|
||||
Ring.*WellDefined quotientRing fr=ft fs=fu = quotientGroupLemma (Ring.additiveGroup R) (RingHom.groupHom f) (transitive (RingHom.ringHom f) (transitive (Ring.*WellDefined R2 (quotientGroupLemma' (Ring.additiveGroup R) (RingHom.groupHom f) fr=ft) (quotientGroupLemma' (Ring.additiveGroup R) (RingHom.groupHom f) fs=fu)) (symmetric (RingHom.ringHom f))))
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
Ring.1R quotientRing = Ring.1R R
|
||||
Ring.groupIsAbelian quotientRing = quotientGroupLemma (Ring.additiveGroup R) (RingHom.groupHom f) (GroupHom.wellDefined (RingHom.groupHom f) (Ring.groupIsAbelian R))
|
||||
Ring.*Associative quotientRing = quotientGroupLemma (Ring.additiveGroup R) (RingHom.groupHom f) (GroupHom.wellDefined (RingHom.groupHom f) (Ring.*Associative R))
|
||||
Ring.*Commutative quotientRing = quotientGroupLemma (Ring.additiveGroup R) (RingHom.groupHom f) (GroupHom.wellDefined (RingHom.groupHom f) (Ring.*Commutative R))
|
||||
Ring.*DistributesOver+ quotientRing = quotientGroupLemma (Ring.additiveGroup R) (RingHom.groupHom f) (GroupHom.wellDefined (RingHom.groupHom f) (Ring.*DistributesOver+ R))
|
||||
Ring.identIsIdent quotientRing = quotientGroupLemma (Ring.additiveGroup R) (RingHom.groupHom f) (GroupHom.wellDefined (RingHom.groupHom f) (Ring.identIsIdent R))
|
||||
|
||||
projectionMapIsHom : RingHom R quotientRing id
|
||||
RingHom.preserves1 projectionMapIsHom = quotientGroupLemma (Ring.additiveGroup R) (RingHom.groupHom f) (Equivalence.reflexive (Setoid.eq T))
|
||||
RingHom.ringHom projectionMapIsHom = quotientGroupLemma (Ring.additiveGroup R) (RingHom.groupHom f) (Equivalence.reflexive (Setoid.eq T))
|
||||
RingHom.groupHom projectionMapIsHom = projectionMapIsGroupHom
|
Reference in New Issue
Block a user