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

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

View File

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

View File

@@ -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 = {!!}

View File

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

View File

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

View File

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

View File

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