mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-13 07:38:40 +00:00
A bit of cleanup (#69)
This commit is contained in:
@@ -8,31 +8,28 @@ open import Numbers.Naturals.Naturals
|
||||
open import Sets.FinSet
|
||||
open import Groups.Definition
|
||||
open import Groups.Lemmas
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Groups
|
||||
open import Groups.SymmetricGroups.Definition
|
||||
open import Groups.Groups2
|
||||
open import Groups.Actions
|
||||
open import Groups.Actions.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.ActionIsSymmetry where
|
||||
module Groups.ActionIsSymmetry {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+_ : A → A → A} {G : Group S _+_} (gAction : GroupAction G T) where
|
||||
|
||||
actionPermutation : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+_ : A → A → A} {G : Group S _+_} → (action : GroupAction G T) → (g : A) → SymmetryGroupElements T
|
||||
actionPermutation {B = B} {T = T} {_+_ = _+_} {G = G} action g = sym {f = λ x → (GroupAction.action action g x)} (record { inj = record { injective = inj ; wellDefined = GroupAction.actionWellDefined2 action } ; surj = record { surjective = surj ; wellDefined = GroupAction.actionWellDefined2 action } })
|
||||
open Group G
|
||||
open GroupAction gAction
|
||||
|
||||
actionPermutation : (g : A) → SymmetryGroupElements T
|
||||
actionPermutation g = sym {f = λ x → action g x} (record { inj = record { injective = inj ; wellDefined = actionWellDefined2 } ; surj = record { surjective = surj ; wellDefined = actionWellDefined2 } })
|
||||
where
|
||||
open Setoid T
|
||||
open Reflexive (Equivalence.reflexiveEq (Setoid.eq T))
|
||||
open Symmetric (Equivalence.symmetricEq (Setoid.eq T))
|
||||
open Transitive (Equivalence.transitiveEq (Setoid.eq T))
|
||||
open Group G
|
||||
inj : {x y : B} → (Setoid._∼_ T (GroupAction.action action g x) (GroupAction.action action g y)) → Setoid._∼_ T x y
|
||||
inj {x} {y} gx=gy = transitive (symmetric (GroupAction.identityAction action)) (transitive (transitive (symmetric (GroupAction.actionWellDefined1 action (invLeft {g}))) (transitive (transitive (GroupAction.associativeAction action) (transitive (GroupAction.actionWellDefined2 action gx=gy) (symmetric (GroupAction.associativeAction action)))) (GroupAction.actionWellDefined1 action (invLeft {g})))) (GroupAction.identityAction action))
|
||||
surj : {x : B} → Sg B (λ a → GroupAction.action action g a ∼ x)
|
||||
surj {x} = GroupAction.action action (inverse g) x , transitive (symmetric (GroupAction.associativeAction action)) (transitive (GroupAction.actionWellDefined1 action invRight) (GroupAction.identityAction action))
|
||||
open Equivalence eq
|
||||
inj : {x y : B} → (Setoid._∼_ T (action g x) (action g y)) → Setoid._∼_ T x y
|
||||
inj {x} {y} gx=gy = transitive (symmetric identityAction) (transitive (transitive (symmetric (actionWellDefined1 (invLeft {g}))) (transitive (transitive associativeAction (transitive (actionWellDefined2 gx=gy) (symmetric associativeAction))) (actionWellDefined1 (invLeft {g})))) identityAction)
|
||||
surj : {x : B} → Sg B (λ a → action g a ∼ x)
|
||||
surj {x} = action (inverse g) x , transitive (symmetric associativeAction) (transitive (actionWellDefined1 invRight) identityAction)
|
||||
|
||||
actionPermutationMapIsHom : {a b c d : _} {A : Set a} {B : Set b} {S : Setoid {a} {c} A} {T : Setoid {b} {d} B} {_+_ : A → A → A} {G : Group S _+_} (action : GroupAction G T) → GroupHom G (symmetricGroup T) (actionPermutation action)
|
||||
GroupHom.groupHom (actionPermutationMapIsHom {T = T} action) = ExtensionallyEqual.eq λ {z} → GroupAction.associativeAction action
|
||||
where
|
||||
open Setoid T
|
||||
open Reflexive (Equivalence.reflexiveEq (Setoid.eq T))
|
||||
open Symmetric (Equivalence.symmetricEq (Setoid.eq T))
|
||||
open Transitive (Equivalence.transitiveEq (Setoid.eq T))
|
||||
GroupHom.wellDefined (actionPermutationMapIsHom action) x=y = ExtensionallyEqual.eq λ {z} → GroupAction.actionWellDefined1 action x=y
|
||||
actionPermutationMapIsHom : GroupHom G (symmetricGroup T) actionPermutation
|
||||
GroupHom.groupHom actionPermutationMapIsHom = associativeAction
|
||||
GroupHom.wellDefined actionPermutationMapIsHom x=y = actionWellDefined1 x=y
|
||||
|
Reference in New Issue
Block a user