mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-11 22:58:40 +00:00
Quotient ring (#81)
This commit is contained in:
@@ -31,10 +31,10 @@ stabiliserWellDefined x {g} {h} g=h gx=x = transitive (actionWellDefined1 (Equiv
|
||||
open Setoid T
|
||||
open Equivalence (Setoid.eq T)
|
||||
|
||||
stabiliserSubgroup : (x : B) → subgroup G (stabiliserPred x)
|
||||
_&&_.fst (stabiliserSubgroup x) = stabiliserWellDefined x
|
||||
_&_&_.one (_&&_.snd (stabiliserSubgroup x)) gx=x hx=x = transitive associativeAction (transitive (actionWellDefined2 hx=x) gx=x)
|
||||
_&_&_.two (_&&_.snd (stabiliserSubgroup x)) = identityAction
|
||||
_&_&_.three (_&&_.snd (stabiliserSubgroup x)) {g} gx=x = transitive (transitive (transitive (actionWellDefined2 (symmetric gx=x)) (symmetric associativeAction)) (actionWellDefined1 (invLeft {g}))) identityAction
|
||||
stabiliserSubgroup : (x : B) → Subgroup G (stabiliserPred x)
|
||||
Subgroup.isSubset (stabiliserSubgroup x) = stabiliserWellDefined x
|
||||
Subgroup.closedUnderPlus (stabiliserSubgroup x) gx=x hx=x = transitive associativeAction (transitive (actionWellDefined2 hx=x) gx=x)
|
||||
Subgroup.containsIdentity (stabiliserSubgroup x) = identityAction
|
||||
Subgroup.closedUnderInverse (stabiliserSubgroup x) {g} gx=x = transitive (transitive (transitive (actionWellDefined2 (symmetric gx=x)) (symmetric associativeAction)) (actionWellDefined1 (invLeft {g}))) identityAction
|
||||
where
|
||||
open Group G
|
||||
|
@@ -10,21 +10,18 @@ open import Groups.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Subgroups.Normal.Definition
|
||||
|
||||
module Groups.Cosets {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} (G : Group S _+_) {c : _} {pred : A → Set c} (subgrp : subgroup G pred) where
|
||||
module Groups.Cosets {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} (G : Group S _+_) {c : _} {pred : A → Set c} (subgrp : Subgroup G pred) where
|
||||
|
||||
open Equivalence (Setoid.eq S)
|
||||
open import Groups.Lemmas G
|
||||
open _&_&_ (_&&_.snd subgrp)
|
||||
open Group G
|
||||
|
||||
private
|
||||
subs = _&&_.fst subgrp
|
||||
open Subgroup subgrp
|
||||
|
||||
cosetSetoid : Setoid A
|
||||
Setoid._∼_ cosetSetoid g h = pred ((Group.inverse G h) + g)
|
||||
Equivalence.reflexive (Setoid.eq cosetSetoid) = subs (symmetric (Group.invLeft G)) two
|
||||
Equivalence.symmetric (Setoid.eq cosetSetoid) yx = subs (transitive invContravariant (+WellDefined reflexive invInv)) (three yx)
|
||||
Equivalence.transitive (Setoid.eq cosetSetoid) yx zy = subs (transitive +Associative (+WellDefined (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight)) reflexive)) (one zy yx)
|
||||
Equivalence.reflexive (Setoid.eq cosetSetoid) = isSubset (symmetric (Group.invLeft G)) containsIdentity
|
||||
Equivalence.symmetric (Setoid.eq cosetSetoid) yx = isSubset (transitive invContravariant (+WellDefined reflexive invInv)) (closedUnderInverse yx)
|
||||
Equivalence.transitive (Setoid.eq cosetSetoid) yx zy = isSubset (transitive +Associative (+WellDefined (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight)) reflexive)) (closedUnderPlus zy yx)
|
||||
|
||||
cosetGroup : normalSubgroup G subgrp → Group cosetSetoid _+_
|
||||
Group.+WellDefined (cosetGroup norm) {m} {n} {x} {y} m=x n=y = ans
|
||||
@@ -34,17 +31,17 @@ Group.+WellDefined (cosetGroup norm) {m} {n} {x} {y} m=x n=y = ans
|
||||
u : pred (inverse x + m)
|
||||
u = m=x
|
||||
v : pred (m + inverse x)
|
||||
v = subs (+WellDefined reflexive (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight))) (norm u)
|
||||
v = isSubset (+WellDefined reflexive (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight))) (norm u)
|
||||
ans' : pred ((inverse y) + ((inverse x + m) + inverse (inverse y)))
|
||||
ans' = norm u
|
||||
ans'' : pred ((inverse y) + ((inverse x + m) + y))
|
||||
ans'' = subs (+WellDefined reflexive (+WellDefined reflexive (invTwice y))) ans'
|
||||
ans'' = isSubset (+WellDefined reflexive (+WellDefined reflexive (invTwice y))) ans'
|
||||
ans : pred (inverse (x + y) + (m + n))
|
||||
ans = subs (transitive (transitive +Associative (transitive (+WellDefined (transitive (symmetric +Associative) (transitive (+WellDefined reflexive (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight))) +Associative)) reflexive) (symmetric +Associative))) (symmetric (+WellDefined invContravariant reflexive))) (one ans'' t)
|
||||
ans = isSubset (transitive (transitive +Associative (transitive (+WellDefined (transitive (symmetric +Associative) (transitive (+WellDefined reflexive (transitive (symmetric +Associative) (transitive (+WellDefined reflexive invRight) identRight))) +Associative)) reflexive) (symmetric +Associative))) (symmetric (+WellDefined invContravariant reflexive))) (closedUnderPlus ans'' t)
|
||||
Group.0G (cosetGroup norm) = 0G
|
||||
Group.inverse (cosetGroup norm) = inverse
|
||||
Group.+Associative (cosetGroup norm) {a} {b} {c} = subs (symmetric (transitive (+WellDefined (inverseWellDefined (symmetric +Associative)) reflexive) (invLeft {a + (b + c)}))) two
|
||||
Group.identRight (cosetGroup norm) = subs (symmetric (transitive +Associative (transitive (+WellDefined invLeft reflexive) identRight))) two
|
||||
Group.identLeft (cosetGroup norm) = subs (symmetric (transitive (+WellDefined reflexive identLeft) invLeft)) two
|
||||
Group.invLeft (cosetGroup norm) = subs (symmetric (transitive (+WellDefined reflexive invLeft) invLeft)) two
|
||||
Group.invRight (cosetGroup norm) = subs (symmetric (transitive (+WellDefined reflexive invRight) invLeft)) two
|
||||
Group.+Associative (cosetGroup norm) {a} {b} {c} = isSubset (symmetric (transitive (+WellDefined (inverseWellDefined (symmetric +Associative)) reflexive) (invLeft {a + (b + c)}))) containsIdentity
|
||||
Group.identRight (cosetGroup norm) = isSubset (symmetric (transitive +Associative (transitive (+WellDefined invLeft reflexive) identRight))) containsIdentity
|
||||
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
|
||||
|
@@ -32,14 +32,14 @@ imageGroupSubset {x} {y} x=y (a , fa=x) = a , transitive fa=x x=y
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
imageGroupSubgroup : subgroup H imageGroupPred
|
||||
_&&_.fst imageGroupSubgroup = imageGroupSubset
|
||||
_&_&_.one (_&&_.snd imageGroupSubgroup) {x} {y} (a , fa=x) (b , fb=y) = (a +A b) , transitive (GroupHom.groupHom fHom) (Group.+WellDefined H fa=x fb=y)
|
||||
imageGroupSubgroup : Subgroup H imageGroupPred
|
||||
Subgroup.isSubset imageGroupSubgroup = imageGroupSubset
|
||||
Subgroup.closedUnderPlus imageGroupSubgroup {x} {y} (a , fa=x) (b , fb=y) = (a +A b) , transitive (GroupHom.groupHom fHom) (Group.+WellDefined H fa=x fb=y)
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
_&_&_.two (_&&_.snd imageGroupSubgroup) = Group.0G G , imageOfIdentityIsIdentity fHom
|
||||
_&_&_.three (_&&_.snd imageGroupSubgroup) {x} (a , fa=x) = Group.inverse G a , transitive (homRespectsInverse fHom) (inverseWellDefined H fa=x)
|
||||
Subgroup.containsIdentity imageGroupSubgroup = Group.0G G , imageOfIdentityIsIdentity fHom
|
||||
Subgroup.closedUnderInverse imageGroupSubgroup {x} (a , fa=x) = Group.inverse G a , transitive (homRespectsInverse fHom) (inverseWellDefined H fa=x)
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
@@ -34,11 +34,11 @@ groupKernelPred a = Setoid._∼_ T (f a) (Group.0G H)
|
||||
groupKernelPredWd : {x y : A} → (Setoid._∼_ S x y) → groupKernelPred x → groupKernelPred y
|
||||
groupKernelPredWd x=y fx=0 = transitive (GroupHom.wellDefined fHom (Equivalence.symmetric (Setoid.eq S) x=y)) fx=0
|
||||
|
||||
groupKernelIsSubgroup : subgroup G groupKernelPred
|
||||
_&_&_.one (_&&_.snd groupKernelIsSubgroup) fg=0 fh=0 = transitive (transitive (GroupHom.groupHom fHom) (Group.+WellDefined H fg=0 fh=0)) (Group.identLeft H)
|
||||
_&_&_.two (_&&_.snd groupKernelIsSubgroup) = imageOfIdentityIsIdentity fHom
|
||||
_&_&_.three (_&&_.snd groupKernelIsSubgroup) fg=0 = transitive (homRespectsInverse fHom) (transitive (inverseWellDefined H fg=0) (invIdent H))
|
||||
_&&_.fst groupKernelIsSubgroup = groupKernelPredWd
|
||||
groupKernelIsSubgroup : Subgroup G groupKernelPred
|
||||
Subgroup.closedUnderPlus groupKernelIsSubgroup fg=0 fh=0 = transitive (transitive (GroupHom.groupHom fHom) (Group.+WellDefined H fg=0 fh=0)) (Group.identLeft H)
|
||||
Subgroup.containsIdentity groupKernelIsSubgroup = imageOfIdentityIsIdentity fHom
|
||||
Subgroup.closedUnderInverse groupKernelIsSubgroup fg=0 = transitive (homRespectsInverse fHom) (transitive (inverseWellDefined H fg=0) (invIdent H))
|
||||
Subgroup.isSubset groupKernelIsSubgroup = groupKernelPredWd
|
||||
|
||||
groupKernelIsNormalSubgroup : normalSubgroup G groupKernelIsSubgroup
|
||||
groupKernelIsNormalSubgroup {g} fk=0 = transitive (transitive (transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H reflexive (transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H fk=0 reflexive) (Group.identLeft H)))) (symmetric (GroupHom.groupHom fHom)))) (GroupHom.wellDefined fHom (Group.invRight G {g}))) (imageOfIdentityIsIdentity fHom)
|
||||
|
@@ -14,16 +14,20 @@ module Groups.Subgroups.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A}
|
||||
open import Setoids.Subset S
|
||||
open Group G
|
||||
|
||||
subgroup : {c : _} (pred : A → Set c) → Set (a ⊔ b ⊔ c)
|
||||
subgroup pred = subset pred && (({g h : A} → (pred g) → (pred h) → pred (g + h)) & pred 0G & ({g : A} → (pred g) → (pred (inverse g))))
|
||||
record Subgroup {c : _} (pred : A → Set c) : Set (a ⊔ b ⊔ c) where
|
||||
field
|
||||
isSubset : subset pred
|
||||
closedUnderPlus : {g h : A} → (pred g) → (pred h) → pred (g + h)
|
||||
containsIdentity : pred 0G
|
||||
closedUnderInverse : ({g : A} → (pred g) → (pred (inverse g)))
|
||||
|
||||
subgroupOp : {c : _} {pred : A → Set c} → (s : subgroup pred) → Sg A pred → Sg A pred → Sg A pred
|
||||
subgroupOp {pred = pred} (_ ,, record { one = one ; two = two ; three = three }) (a , prA) (b , prB) = (a + b) , one prA prB
|
||||
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)
|
||||
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 { two = two })) = 0G , two
|
||||
Group.inverse (subgroupIsGroup _ (_ ,, record { three = three })) (a , prA) = (inverse a) , three prA
|
||||
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
|
||||
|
@@ -21,5 +21,5 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Subgroups.Normal.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} (G : Group S _+_) where
|
||||
|
||||
normalSubgroup : {c : _} {pred : A → Set c} (sub : subgroup G pred) → Set (a ⊔ c)
|
||||
normalSubgroup : {c : _} {pred : A → Set c} (sub : Subgroup G pred) → Set (a ⊔ c)
|
||||
normalSubgroup {pred = pred} sub = {g k : A} → pred k → pred (g + (k + Group.inverse G g))
|
||||
|
@@ -64,11 +64,11 @@ groupKernelGroupPred {T = T} G {H = H} {f = f} hom a = Setoid._∼_ T (f a) (Gro
|
||||
groupKernelGroupPredWd : {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 : A → B} (hom : GroupHom G H f) → {x y : A} → (Setoid._∼_ S x y) → (groupKernelGroupPred G hom x → groupKernelGroupPred G hom y)
|
||||
groupKernelGroupPredWd {S = S} {T = T} G hom {x} {y} x=y fx=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined hom (Equivalence.symmetric (Setoid.eq S) x=y)) fx=0
|
||||
|
||||
groupKernelGroupIsSubgroup : {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 : A → B} (hom : GroupHom G H f) → subgroup G (groupKernelGroupPred G hom)
|
||||
_&_&_.one (_&&_.snd (groupKernelGroupIsSubgroup {S = S} {T = T} G {H = H} hom)) g=0 h=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.groupHom hom) (Equivalence.transitive (Setoid.eq T) (Group.+WellDefined H g=0 h=0) (Group.identLeft H))
|
||||
_&_&_.two (_&&_.snd (groupKernelGroupIsSubgroup G hom)) = imageOfIdentityIsIdentity hom
|
||||
_&_&_.three (_&&_.snd (groupKernelGroupIsSubgroup {S = S} {T = T} G {H = H} hom)) g=0 = Equivalence.transitive (Setoid.eq T) (homRespectsInverse hom) (Equivalence.transitive (Setoid.eq T) (inverseWellDefined H g=0) (invIdent H))
|
||||
_&&_.fst (groupKernelGroupIsSubgroup G hom) = groupKernelGroupPredWd G hom
|
||||
groupKernelGroupIsSubgroup : {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 : A → B} (hom : GroupHom G H f) → Subgroup G (groupKernelGroupPred G hom)
|
||||
Subgroup.closedUnderPlus (groupKernelGroupIsSubgroup {S = S} {T = T} G {H = H} hom) g=0 h=0 = Equivalence.transitive (Setoid.eq T) (GroupHom.groupHom hom) (Equivalence.transitive (Setoid.eq T) (Group.+WellDefined H g=0 h=0) (Group.identLeft H))
|
||||
Subgroup.containsIdentity (groupKernelGroupIsSubgroup G hom) = imageOfIdentityIsIdentity hom
|
||||
Subgroup.closedUnderInverse (groupKernelGroupIsSubgroup {S = S} {T = T} G {H = H} hom) g=0 = Equivalence.transitive (Setoid.eq T) (homRespectsInverse hom) (Equivalence.transitive (Setoid.eq T) (inverseWellDefined H g=0) (invIdent H))
|
||||
Subgroup.isSubset (groupKernelGroupIsSubgroup G hom) = groupKernelGroupPredWd G hom
|
||||
|
||||
groupKernelGroupIsNormalSubgroup : {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 : A → B} (hom : GroupHom G H f) → normalSubgroup G (groupKernelGroupIsSubgroup G hom)
|
||||
groupKernelGroupIsNormalSubgroup {T = T} G {H = H} hom k=0 = transitive groupHom (transitive (+WellDefined reflexive groupHom) (transitive (+WellDefined reflexive (transitive (+WellDefined k=0 reflexive) identLeft)) (transitive (symmetric groupHom) (transitive (wellDefined (Group.invRight G)) (imageOfIdentityIsIdentity hom)))))
|
||||
@@ -78,8 +78,8 @@ groupKernelGroupIsNormalSubgroup {T = T} G {H = H} hom k=0 = transitive groupHom
|
||||
open Equivalence eq
|
||||
open GroupHom hom
|
||||
|
||||
abelianGroupSubgroupIsNormal : {a b c : _} {A : Set a} {B : Set b} {S : Setoid {a} {b} A} {_+_ : A → A → A} {G : Group S _+_} {pred : A → Set c} → (s : subgroup G pred) → AbelianGroup G → normalSubgroup G s
|
||||
abelianGroupSubgroupIsNormal {S = S} {_+_ = _+_} {G = G} (predWd ,, record { one = respectsPlus ; two = respectsId ; three = respectsInv }) abelian {k} {l} prK = predWd (transitive (transitive (transitive (symmetric identLeft) (+WellDefined (symmetric invRight) reflexive)) (symmetric +Associative)) (+WellDefined reflexive commutative)) prK
|
||||
abelianGroupSubgroupIsNormal : {a b c : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A → A → A} {G : Group S _+_} {pred : A → Set c} → (s : Subgroup G pred) → AbelianGroup G → normalSubgroup G s
|
||||
abelianGroupSubgroupIsNormal {S = S} {_+_ = _+_} {G = G} record { isSubset = predWd ; closedUnderPlus = respectsPlus ; containsIdentity = respectsId ; closedUnderInverse = respectsInv } abelian {k} {l} prK = predWd (transitive (transitive (transitive (symmetric identLeft) (+WellDefined (symmetric invRight) reflexive)) (symmetric +Associative)) (+WellDefined reflexive commutative)) prK
|
||||
where
|
||||
open Group G
|
||||
open AbelianGroup abelian
|
||||
|
Reference in New Issue
Block a user