mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-11 14:48:42 +00:00
Restructure towards ideals
This commit is contained in:
18
Groups/Homomorphisms/Examples.agda
Normal file
18
Groups/Homomorphisms/Examples.agda
Normal file
@@ -0,0 +1,18 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import LogicalFormulae
|
||||
open import Setoids.Setoids
|
||||
open import Functions
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Lemmas
|
||||
|
||||
module Groups.Homomorphisms.Examples where
|
||||
|
||||
identityHom : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+A_ : A → A → A} (G : Group S _+A_) → GroupHom G G id
|
||||
GroupHom.groupHom (identityHom {S = S} G) = Equivalence.reflexive (Setoid.eq S)
|
||||
GroupHom.wellDefined (identityHom G) = id
|
@@ -5,6 +5,7 @@ open import Groups.Definition
|
||||
open import Orders
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import Setoids.Subset
|
||||
open import LogicalFormulae
|
||||
open import Sets.FinSet
|
||||
open import Functions
|
||||
@@ -20,45 +21,25 @@ open import Groups.QuotientGroup.Definition
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Homomorphisms.Image where
|
||||
module Groups.Homomorphisms.Image {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
|
||||
|
||||
data GroupHomImageElement {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) : Set (a ⊔ b ⊔ c ⊔ d) where
|
||||
ofElt : (x : A) → GroupHomImageElement fHom
|
||||
imageGroupPred : B → Set (a ⊔ d)
|
||||
imageGroupPred b = Sg A (λ a → Setoid._∼_ T (f a) b)
|
||||
|
||||
imageGroupSetoid : {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) → Setoid (GroupHomImageElement fHom)
|
||||
(imageGroupSetoid {T = T} {f = f} fHom Setoid.∼ ofElt b1) (ofElt b2) = Setoid._∼_ T (f b1) (f b2)
|
||||
Equivalence.reflexive (Setoid.eq (imageGroupSetoid {T = T} fHom)) {ofElt b1} = Equivalence.reflexive (Setoid.eq T)
|
||||
Equivalence.symmetric (Setoid.eq (imageGroupSetoid {T = T} fHom)) {ofElt b1} {ofElt b2} = Equivalence.symmetric (Setoid.eq T)
|
||||
Equivalence.transitive (Setoid.eq (imageGroupSetoid {T = T} fHom)) {ofElt b1} {ofElt b2} {ofElt b3} = Equivalence.transitive (Setoid.eq T)
|
||||
|
||||
imageGroupOp : {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) → GroupHomImageElement fHom → GroupHomImageElement fHom → GroupHomImageElement fHom
|
||||
imageGroupOp {_+A_ = _+A_} fHom (ofElt a) (ofElt b) = ofElt (a +A b)
|
||||
|
||||
imageGroup : {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) → Group (imageGroupSetoid fHom) (imageGroupOp fHom)
|
||||
Group.+WellDefined (imageGroup {T = T} {_+A_ = _+A_} {H = H} {f = f} fHom) {ofElt x} {ofElt y} {ofElt a} {ofElt b} x~a y~b = ans
|
||||
imageGroupSubset : subset T imageGroupPred
|
||||
imageGroupSubset {x} {y} x=y (a , fa=x) = a , transitive fa=x x=y
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
ans : f (x +A y) ∼ f (a +A b)
|
||||
ans = transitive (GroupHom.groupHom fHom) (transitive (Group.+WellDefined H x~a y~b) (symmetric (GroupHom.groupHom fHom)))
|
||||
Group.0G (imageGroup {G = G} fHom) = ofElt (Group.0G G)
|
||||
Group.inverse (imageGroup {G = G} fHom) (ofElt a) = ofElt (Group.inverse G a)
|
||||
Group.+Associative (imageGroup {G = G} fHom) {ofElt a} {ofElt b} {ofElt c} = GroupHom.wellDefined fHom (Group.+Associative G)
|
||||
Group.identRight (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.identRight G)
|
||||
Group.identLeft (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.identLeft G)
|
||||
Group.invLeft (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.invLeft G)
|
||||
Group.invRight (imageGroup {G = G} fHom) {ofElt a} = GroupHom.wellDefined fHom (Group.invRight G)
|
||||
|
||||
groupHomImageInclusion : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A → A → A} {_+H_ : B → B → B} {G : Group S _+G_} {H : Group T _+H_} {f : A → B} (fHom : GroupHom G H f) → GroupHomImageElement fHom → B
|
||||
groupHomImageInclusion {f = f} fHom (ofElt x) = f x
|
||||
|
||||
groupHomImageIncludes : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A → A → A} {_+H_ : B → B → B} {G : Group S _+G_} {H : Group T _+H_} {f : A → B} (fHom : GroupHom G H f) → GroupHom (imageGroup fHom) H (groupHomImageInclusion fHom)
|
||||
GroupHom.groupHom (groupHomImageIncludes fHom) {ofElt x} {ofElt y} = GroupHom.groupHom fHom
|
||||
GroupHom.wellDefined (groupHomImageIncludes fHom) {ofElt x} {ofElt y} x~y = x~y
|
||||
|
||||
groupHomImageInjects : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A → A → A} {_+H_ : B → B → B} {G : Group S _+G_} {H : Group T _+H_} {f : A → B} (fHom : GroupHom G H f) → SetoidInjection (imageGroupSetoid fHom) T (groupHomImageInclusion fHom)
|
||||
groupHomImageInjects {S = S} {T} {_+G_} {_+H_} {G} {H} {f} fHom = record { wellDefined = λ {x} {y} → GroupHom.wellDefined (groupHomImageIncludes fHom) {x} {y} ; injective = λ {x} {y} → inj {x} {y} }
|
||||
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)
|
||||
where
|
||||
inj : {x y : GroupHomImageElement fHom} → (Setoid._∼_ T (groupHomImageInclusion fHom x) (groupHomImageInclusion fHom y)) → Setoid._∼_ (imageGroupSetoid fHom) x y
|
||||
inj {ofElt x} {ofElt y} x~y = x~y
|
||||
|
||||
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)
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
44
Groups/Homomorphisms/Kernel.agda
Normal file
44
Groups/Homomorphisms/Kernel.agda
Normal file
@@ -0,0 +1,44 @@
|
||||
{-# OPTIONS --safe --warning=error --without-K #-}
|
||||
|
||||
open import Groups.Groups
|
||||
open import Groups.Definition
|
||||
open import Orders
|
||||
open import Numbers.Integers.Integers
|
||||
open import Setoids.Setoids
|
||||
open import LogicalFormulae
|
||||
open import Sets.FinSet
|
||||
open import Functions
|
||||
open import Sets.EquivalenceRelations
|
||||
open import Numbers.Naturals.Naturals
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Isomorphisms.Definition
|
||||
open import Groups.Subgroups.Definition
|
||||
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 Groups.Cosets
|
||||
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Groups.Homomorphisms.Kernel {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+G_ : A → A → A} {_+H_ : B → B → B} {G : Group S _+G_} {H : Group T _+H_} {f : A → B} (fHom : GroupHom G H f) where
|
||||
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
groupKernelPred : A → Set d
|
||||
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
|
||||
|
||||
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)
|
@@ -11,36 +11,38 @@ open import Sets.EquivalenceRelations
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Lemmas
|
||||
|
||||
module Groups.Homomorphisms.Lemmas where
|
||||
module Groups.Homomorphisms.Lemmas {a b c d : _} {A : Set a} {S : Setoid {a} {c} A} {B : Set b} {T : Setoid {b} {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) where
|
||||
|
||||
imageOfIdentityIsIdentity : {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {B : Set n} {T : Setoid {n} {p} 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) → Setoid._∼_ T (f (Group.0G G)) (Group.0G H)
|
||||
imageOfIdentityIsIdentity {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {f = f} hom = Equivalence.symmetric (Setoid.eq T) t
|
||||
imageOfIdentityIsIdentity : Setoid._∼_ T (f (Group.0G G)) (Group.0G H)
|
||||
imageOfIdentityIsIdentity = Equivalence.symmetric (Setoid.eq T) t
|
||||
where
|
||||
open Group H
|
||||
open Setoid T
|
||||
id2 : Setoid._∼_ S (Group.0G G) ((Group.0G G) ·A (Group.0G G))
|
||||
id2 : Setoid._∼_ S (Group.0G G) ((Group.0G G) +A (Group.0G G))
|
||||
id2 = Equivalence.symmetric (Setoid.eq S) (Group.identRight G)
|
||||
r : f (Group.0G G) ∼ f (Group.0G G) ·B f (Group.0G G)
|
||||
s : 0G ·B f (Group.0G G) ∼ f (Group.0G G) ·B f (Group.0G G)
|
||||
r : f (Group.0G G) ∼ f (Group.0G G) +B f (Group.0G G)
|
||||
s : 0G +B f (Group.0G G) ∼ f (Group.0G G) +B f (Group.0G G)
|
||||
t : 0G ∼ f (Group.0G G)
|
||||
t = groupsHaveRightCancellation H (f (Group.0G G)) 0G (f (Group.0G G)) s
|
||||
s = Equivalence.transitive (Setoid.eq T) identLeft r
|
||||
r = Equivalence.transitive (Setoid.eq T) (GroupHom.wellDefined hom id2) (GroupHom.groupHom hom)
|
||||
|
||||
groupHomsCompose : {m n o r s t : _} {A : Set m} {S : Setoid {m} {r} A} {_+A_ : A → A → A} {B : Set n} {T : Setoid {n} {s} B} {_+B_ : B → B → B} {C : Set o} {U : Setoid {o} {t} C} {_+C_ : C → C → C} {G : Group S _+A_} {H : Group T _+B_} {I : Group U _+C_} {f : A → B} {g : B → C} (fHom : GroupHom G H f) (gHom : GroupHom H I g) → GroupHom G I (g ∘ f)
|
||||
GroupHom.wellDefined (groupHomsCompose {G = G} {H} {I} {f} {g} fHom gHom) {x} {y} pr = GroupHom.wellDefined gHom (GroupHom.wellDefined fHom pr)
|
||||
GroupHom.groupHom (groupHomsCompose {S = S} {_+A_ = _·A_} {T = T} {U = U} {_+C_ = _·C_} {G = G} {H} {I} {f} {g} fHom gHom) {x} {y} = answer
|
||||
groupHomsCompose : {o t : _} {C : Set o} {U : Setoid {o} {t} C} {_+C_ : C → C → C} {I : Group U _+C_} {g : B → C} (gHom : GroupHom H I g) → GroupHom G I (g ∘ f)
|
||||
GroupHom.wellDefined (groupHomsCompose {I} {f} gHom) {x} {y} pr = GroupHom.wellDefined gHom (GroupHom.wellDefined hom pr)
|
||||
GroupHom.groupHom (groupHomsCompose {U = U} {_+C_ = _·C_} {I} {g} gHom) {x} {y} = answer
|
||||
where
|
||||
open Group I
|
||||
answer : (Setoid._∼_ U) ((g ∘ f) (x ·A y)) ((g ∘ f) x ·C (g ∘ f) y)
|
||||
answer = (Equivalence.transitive (Setoid.eq U)) (GroupHom.wellDefined gHom (GroupHom.groupHom fHom {x} {y}) ) (GroupHom.groupHom gHom {f x} {f y})
|
||||
answer : (Setoid._∼_ U) ((g ∘ f) (x +A y)) ((g ∘ f) x ·C (g ∘ f) y)
|
||||
answer = (Equivalence.transitive (Setoid.eq U)) (GroupHom.wellDefined gHom (GroupHom.groupHom hom {x} {y}) ) (GroupHom.groupHom gHom {f x} {f y})
|
||||
|
||||
homRespectsInverse : {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 : A} → Setoid._∼_ T (underF (Group.inverse G x)) (Group.inverse H (underF x))
|
||||
homRespectsInverse {T = T} {_·A_ = _·A_} {_·B_ = _·B_} {G = G} {H = H} {underF = f} fHom {x} = rightInversesAreUnique H (f x) (f (Group.inverse G x)) (transitive (symmetric (GroupHom.groupHom fHom)) (transitive (GroupHom.wellDefined fHom (Group.invLeft G)) (imageOfIdentityIsIdentity fHom)))
|
||||
homRespectsInverse : {x : A} → Setoid._∼_ T (f (Group.inverse G x)) (Group.inverse H (f x))
|
||||
homRespectsInverse {x} = rightInversesAreUnique H (f x) (f (Group.inverse G x)) (transitive (symmetric (GroupHom.groupHom hom)) (transitive (GroupHom.wellDefined hom (Group.invLeft G)) imageOfIdentityIsIdentity))
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
||||
identityHom : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+A_ : A → A → A} (G : Group S _+A_) → GroupHom G G id
|
||||
GroupHom.groupHom (identityHom {S = S} G) = Equivalence.reflexive (Setoid.eq S)
|
||||
GroupHom.wellDefined (identityHom G) = id
|
||||
zeroImpliesInverseZero : {x : A} → Setoid._∼_ T (f x) (Group.0G H) → Setoid._∼_ T (f (Group.inverse G x)) (Group.0G H)
|
||||
zeroImpliesInverseZero {x} fx=0 = transitive homRespectsInverse (transitive (inverseWellDefined H fx=0) (invIdent H))
|
||||
where
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
|
Reference in New Issue
Block a user