mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-18 17:48:40 +00:00
Fix up the structure of Actions and SymmetricGroups (#67)
This commit is contained in:
69
Groups/SymmetricGroups/Definition.agda
Normal file
69
Groups/SymmetricGroups/Definition.agda
Normal file
@@ -0,0 +1,69 @@
|
||||
{-# 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.Groups
|
||||
open import Groups.Definition
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.SymmetricGroups.Definition where
|
||||
data SymmetryGroupElements {a b : _} {A : Set a} (S : Setoid {a} {b} A) : Set (a ⊔ b) where
|
||||
sym : {f : A → A} → SetoidBijection S S f → SymmetryGroupElements S
|
||||
|
||||
WellDefined : {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f : A → B) → Set (a ⊔ c ⊔ d)
|
||||
WellDefined {A = A} S T f = {x y : A} → Setoid._∼_ S x y → Setoid._∼_ T (f x) (f y)
|
||||
|
||||
data ExtensionallyEqual {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f g : A → B) (fWd : WellDefined S T f) (gWd : WellDefined S T g) : Set (a ⊔ b ⊔ c ⊔ d) where
|
||||
eq : ({x : A} → Setoid._∼_ T (f x) (g x)) → ExtensionallyEqual S T f g fWd gWd
|
||||
|
||||
extensionallyEqualReflexive : {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f : A → B) (fWD1 fWD2 : WellDefined S T f) → ExtensionallyEqual S T f f fWD1 fWD2
|
||||
extensionallyEqualReflexive S T f fWD1 _ = eq (Equivalence.reflexive (Setoid.eq T))
|
||||
|
||||
extensionallyEqualSymmetric : {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f g : A → B) (fWD : WellDefined S T f) (gWD : WellDefined S T g) → ExtensionallyEqual S T f g fWD gWD → ExtensionallyEqual S T g f gWD fWD
|
||||
extensionallyEqualSymmetric S T f g fWD gWD (eq pr) = eq (Equivalence.symmetric (Setoid.eq T) pr)
|
||||
|
||||
extensionallyEqualTransitive : {a b c d : _} {A : Set a} {B : Set b} (S : Setoid {a} {c} A) (T : Setoid {b} {d} B) (f g h : A → B) (fWD : WellDefined S T f) (gWD : WellDefined S T g) (hWD : WellDefined S T h) → ExtensionallyEqual S T f g fWD gWD → ExtensionallyEqual S T g h gWD hWD → ExtensionallyEqual S T f h fWD hWD
|
||||
extensionallyEqualTransitive S T f g h fWD gWD hWD (eq pr1) (eq pr2) = eq (Equivalence.transitive (Setoid.eq T) pr1 pr2)
|
||||
|
||||
symmetricSetoid : {a b : _} {A : Set a} (S : Setoid {a} {b} A) → Setoid (SymmetryGroupElements S)
|
||||
Setoid._∼_ (symmetricSetoid S) (sym {f} bijF) (sym {g} bijG) = ExtensionallyEqual S S f g (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG)
|
||||
Equivalence.reflexive (Setoid.eq (symmetricSetoid S)) {sym {f} bijF} = extensionallyEqualReflexive S S f (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijF)
|
||||
Equivalence.symmetric (Setoid.eq (symmetricSetoid S)) {sym {f} bijF} {sym {g} bijG} f~g = extensionallyEqualSymmetric S S f g (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG) f~g
|
||||
Equivalence.transitive (Setoid.eq (symmetricSetoid S)) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} f~g g~h = extensionallyEqualTransitive S S f g h (SetoidBijection.wellDefined bijF) (SetoidBijection.wellDefined bijG) (SetoidBijection.wellDefined bijH) f~g g~h
|
||||
|
||||
symmetricGroupOp : {a b : _} {A : Set a} {S : Setoid {a} {b} A} (f g : SymmetryGroupElements S) → SymmetryGroupElements S
|
||||
symmetricGroupOp (sym {f} bijF) (sym {g} bijG) = sym (setoidBijComp bijG bijF)
|
||||
|
||||
symmetricGroupInv : {a b : _} {A : Set a} (S : Setoid {a} {b} A) → SymmetryGroupElements S → SymmetryGroupElements S
|
||||
symmetricGroupInv S (sym {f} bijF) with setoidBijectiveImpliesInvertible bijF
|
||||
... | record { inverse = inverse ; inverseWellDefined = iwd ; isLeft = isLeft ; isRight = isRight } = sym (setoidInvertibleImpliesBijective (record { fWellDefined = iwd ; inverse = f ; inverseWellDefined = SetoidInjection.wellDefined (SetoidBijection.inj bijF) ; isLeft = λ b → isRight b ; isRight = λ b → isLeft b }))
|
||||
|
||||
symmetricGroupInvIsLeft : {a b : _} {A : Set a} (S : Setoid {a} {b} A) → {x : SymmetryGroupElements S} → Setoid._∼_ (symmetricSetoid S) (symmetricGroupOp (symmetricGroupInv S x) x) (sym setoidIdIsBijective)
|
||||
symmetricGroupInvIsLeft {A = A} S {sym {f = f} fBij} = ExtensionallyEqual.eq ans
|
||||
where
|
||||
ans : {x : A} → Setoid._∼_ S (SetoidInvertible.inverse (setoidBijectiveImpliesInvertible fBij) (f x)) x
|
||||
ans {x} with SetoidSurjection.surjective (SetoidBijection.surj fBij) {f x}
|
||||
ans {x} | a , b = SetoidInjection.injective (SetoidBijection.inj fBij) b
|
||||
|
||||
symmetricGroupInvIsRight : {a b : _} {A : Set a} (S : Setoid {a} {b} A) → {x : SymmetryGroupElements S} → Setoid._∼_ (symmetricSetoid S) (symmetricGroupOp x (symmetricGroupInv S x)) (sym setoidIdIsBijective)
|
||||
symmetricGroupInvIsRight {A = A} S {sym {f = f} fBij} = ExtensionallyEqual.eq ans
|
||||
where
|
||||
ans : {x : A} → Setoid._∼_ S (f (SetoidInvertible.inverse (setoidBijectiveImpliesInvertible fBij) x)) x
|
||||
ans {x} with SetoidSurjection.surjective (SetoidBijection.surj fBij) {x}
|
||||
ans {x} | a , b = b
|
||||
|
||||
symmetricGroup : {a b : _} {A : Set a} (S : Setoid {a} {b} A) → Group (symmetricSetoid S) (symmetricGroupOp {A = A})
|
||||
Group.+WellDefined (symmetricGroup A) {sym {m} bijM} {sym {n} bijN} {sym {x} bijX} {sym {y} bijY} (ExtensionallyEqual.eq m~x) (ExtensionallyEqual.eq n~y) = ExtensionallyEqual.eq (transitive m~x (SetoidBijection.wellDefined bijX n~y))
|
||||
where
|
||||
open Equivalence (Setoid.eq A)
|
||||
Group.0G (symmetricGroup A) = sym setoidIdIsBijective
|
||||
Group.inverse (symmetricGroup S) = symmetricGroupInv S
|
||||
Group.+Associative (symmetricGroup A) {sym {f} bijF} {sym {g} bijG} {sym {h} bijH} = ExtensionallyEqual.eq λ {x} → Equivalence.reflexive (Setoid.eq A)
|
||||
Group.identRight (symmetricGroup A) {sym {f} bijF} = ExtensionallyEqual.eq λ {x} → Equivalence.reflexive (Setoid.eq A)
|
||||
Group.identLeft (symmetricGroup A) {sym {f} bijF} = ExtensionallyEqual.eq λ {x} → Equivalence.reflexive (Setoid.eq A)
|
||||
Group.invLeft (symmetricGroup S) {x} = symmetricGroupInvIsLeft S {x}
|
||||
Group.invRight (symmetricGroup S) {x} = symmetricGroupInvIsRight S {x}
|
126
Groups/SymmetricGroups/Lemmas.agda
Normal file
126
Groups/SymmetricGroups/Lemmas.agda
Normal file
@@ -0,0 +1,126 @@
|
||||
{-# 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 Groups.Lemmas
|
||||
open import Groups.Groups
|
||||
open import Groups.Subgroups.Definition
|
||||
open import Groups.Homomorphisms.Definition
|
||||
open import Groups.Homomorphisms.Lemmas
|
||||
open import Groups.Actions.Definition
|
||||
open import Groups.Groups2
|
||||
open import Sets.EquivalenceRelations
|
||||
|
||||
module Groups.SymmetricGroups.Lemmas where
|
||||
|
||||
trivialAction : {m n o p : _} {A : Set m} {S : Setoid {m} {o} A} {_·_ : A → A → A} {B : Set n} (G : Group S _·_) (X : Setoid {n} {p} B) → GroupAction G X
|
||||
trivialAction G X = record { action = λ _ x → x ; actionWellDefined1 = λ _ → reflexive ; actionWellDefined2 = λ wd1 → wd1 ; identityAction = reflexive ; associativeAction = reflexive }
|
||||
where
|
||||
open Setoid X renaming (eq to setoidEq)
|
||||
open Equivalence (Setoid.eq X)
|
||||
|
||||
leftRegularAction : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} (G : Group S _·_) → GroupAction G S
|
||||
GroupAction.action (leftRegularAction {_·_ = _·_} G) g h = g · h
|
||||
where
|
||||
open Group G
|
||||
GroupAction.actionWellDefined1 (leftRegularAction {S = S} G) eq1 = +WellDefined eq1 reflexive
|
||||
where
|
||||
open Group G
|
||||
open Setoid S renaming (eq to setoidEq)
|
||||
open Equivalence setoidEq
|
||||
GroupAction.actionWellDefined2 (leftRegularAction {S = S} G) {g} {x} {y} eq1 = +WellDefined reflexive eq1
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
GroupAction.identityAction (leftRegularAction G) = identLeft
|
||||
where
|
||||
open Group G
|
||||
GroupAction.associativeAction (leftRegularAction {S = S} G) = symmetric +Associative
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
conjugationAction : {m n : _} {A : Set m} {S : Setoid {m} {n} A} {_·_ : A → A → A} → (G : Group S _·_) → GroupAction G S
|
||||
conjugationAction {S = S} {_·_ = _·_} G = record { action = λ g h → (g · h) · (inverse g) ; actionWellDefined1 = λ gh → +WellDefined (+WellDefined gh reflexive) (inverseWellDefined G gh) ; actionWellDefined2 = λ x~y → +WellDefined (+WellDefined reflexive x~y) reflexive ; identityAction = transitive (+WellDefined reflexive (invIdent G)) (transitive identRight identLeft) ; associativeAction = λ {x} {g} {h} → transitive (+WellDefined reflexive (invContravariant G)) (transitive +Associative (+WellDefined (transitive (symmetric +Associative) (transitive (symmetric (+Associative)) (+WellDefined reflexive +Associative))) reflexive)) }
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Equivalence eq
|
||||
|
||||
conjugationNormalSubgroupAction : {m n o p : _} {A : Set m} {B : Set o} {S : Setoid {m} {n} A} {T : Setoid {o} {p} 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) → GroupAction G (quotientGroupSetoid G f)
|
||||
GroupAction.action (conjugationNormalSubgroupAction {_·A_ = _·A_} G H {f} fHom) a b = a ·A (b ·A (Group.inverse G a))
|
||||
GroupAction.actionWellDefined1 (conjugationNormalSubgroupAction {S = S} {T = T} {_·A_ = _·A_} G H {f} fHom) {g} {h} {x} g~h = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence eq
|
||||
ans : f ((g ·A (x ·A (inverse g))) ·A inverse (h ·A (x ·A inverse h))) ∼ Group.0G H
|
||||
ans = transitive (GroupHom.wellDefined fHom (transferToRight'' G (Group.+WellDefined G g~h (Group.+WellDefined G (Equivalence.reflexive (Setoid.eq S)) (inverseWellDefined G g~h))))) (imageOfIdentityIsIdentity fHom)
|
||||
GroupAction.actionWellDefined2 (conjugationNormalSubgroupAction {S = S} {T = T} {_·A_ = _·A_} {_·B_ = _·B_} G H {f} fHom) {g} {x} {y} x~y = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T
|
||||
open Equivalence (Setoid.eq S)
|
||||
open Equivalence (Setoid.eq T) renaming (transitive to transitiveH ; symmetric to symmetricH ; reflexive to reflexiveH)
|
||||
input : f (x ·A inverse y) ∼ Group.0G H
|
||||
input = x~y
|
||||
p1 : Setoid._∼_ S ((g ·A (x ·A inverse g)) ·A inverse (g ·A (y ·A inverse g))) ((g ·A (x ·A inverse g)) ·A (inverse (y ·A (inverse g)) ·A inverse g))
|
||||
p1 = Group.+WellDefined G reflexive (invContravariant G)
|
||||
p2 : Setoid._∼_ S ((g ·A (x ·A inverse g)) ·A (inverse (y ·A (inverse g)) ·A inverse g)) ((g ·A (x ·A inverse g)) ·A ((inverse (inverse g) ·A inverse y) ·A inverse g))
|
||||
p2 = Group.+WellDefined G reflexive (Group.+WellDefined G (invContravariant G) reflexive)
|
||||
p3 : Setoid._∼_ S ((g ·A (x ·A inverse g)) ·A ((inverse (inverse g) ·A inverse y) ·A inverse g)) (g ·A (((x ·A inverse g) ·A (inverse (inverse g) ·A inverse y)) ·A inverse g))
|
||||
p3 = symmetric (transitive (+WellDefined reflexive (symmetric +Associative)) +Associative)
|
||||
p4 : Setoid._∼_ S (g ·A (((x ·A inverse g) ·A (inverse (inverse g) ·A inverse y)) ·A inverse g)) (g ·A ((x ·A ((inverse g ·A inverse (inverse g)) ·A inverse y)) ·A inverse g))
|
||||
p4 = Group.+WellDefined G reflexive (Group.+WellDefined G (symmetric (transitive (+WellDefined reflexive (symmetric +Associative)) +Associative)) reflexive)
|
||||
p5 : Setoid._∼_ S (g ·A ((x ·A ((inverse g ·A inverse (inverse g)) ·A inverse y)) ·A inverse g)) (g ·A ((x ·A (0G ·A inverse y)) ·A inverse g))
|
||||
p5 = Group.+WellDefined G reflexive (Group.+WellDefined G (Group.+WellDefined G reflexive (Group.+WellDefined G invRight reflexive)) reflexive)
|
||||
p6 : Setoid._∼_ S (g ·A ((x ·A (0G ·A inverse y)) ·A inverse g)) (g ·A ((x ·A inverse y) ·A inverse g))
|
||||
p6 = Group.+WellDefined G reflexive (Group.+WellDefined G (Group.+WellDefined G reflexive identLeft) reflexive)
|
||||
intermediate : Setoid._∼_ S ((g ·A (x ·A inverse g)) ·A inverse (g ·A (y ·A inverse g))) (g ·A ((x ·A inverse y) ·A inverse g))
|
||||
intermediate = transitive p1 (transitive p2 (transitive p3 (transitive p4 (transitive p5 p6))))
|
||||
p7 : f ((g ·A (x ·A inverse g)) ·A inverse (g ·A (y ·A inverse g))) ∼ f (g ·A ((x ·A inverse y) ·A inverse g))
|
||||
p7 = GroupHom.wellDefined fHom intermediate
|
||||
p8 : f (g ·A ((x ·A inverse y) ·A inverse g)) ∼ (f g) ·B (f ((x ·A inverse y) ·A inverse g))
|
||||
p8 = GroupHom.groupHom fHom
|
||||
p9 : (f g) ·B (f ((x ·A inverse y) ·A inverse g)) ∼ (f g) ·B (f (x ·A inverse y) ·B f (inverse g))
|
||||
p9 = Group.+WellDefined H reflexiveH (GroupHom.groupHom fHom)
|
||||
p10 : (f g) ·B (f (x ·A inverse y) ·B f (inverse g)) ∼ (f g) ·B (Group.0G H ·B f (inverse g))
|
||||
p10 = Group.+WellDefined H reflexiveH (Group.+WellDefined H input reflexiveH)
|
||||
p11 : (f g) ·B (Group.0G H ·B f (inverse g)) ∼ (f g) ·B (f (inverse g))
|
||||
p11 = Group.+WellDefined H reflexiveH (Group.identLeft H)
|
||||
p12 : (f g) ·B (f (inverse g)) ∼ f (g ·A (inverse g))
|
||||
p12 = symmetricH (GroupHom.groupHom fHom)
|
||||
intermediate2 : f ((g ·A (x ·A inverse g)) ·A inverse (g ·A (y ·A inverse g))) ∼ (f (g ·A (inverse g)))
|
||||
intermediate2 = transitiveH p7 (transitiveH p8 (transitiveH p9 (transitiveH p10 (transitiveH p11 p12))))
|
||||
ans : f ((g ·A (x ·A inverse g)) ·A inverse (g ·A (y ·A inverse g))) ∼ Group.0G H
|
||||
ans = transitiveH intermediate2 (transitiveH (GroupHom.wellDefined fHom invRight) (imageOfIdentityIsIdentity fHom))
|
||||
GroupAction.identityAction (conjugationNormalSubgroupAction {S = S} {T = T} {_·A_ = _·A_} G H {f} fHom) {x} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid S
|
||||
open Setoid T renaming (_∼_ to _∼T_)
|
||||
open Equivalence (Setoid.eq T)
|
||||
i : Setoid._∼_ S (x ·A inverse 0G) x
|
||||
i = Equivalence.transitive (Setoid.eq S) (+WellDefined (Equivalence.reflexive (Setoid.eq S)) (invIdent G)) identRight
|
||||
h : 0G ·A (x ·A inverse 0G) ∼ x
|
||||
h = Equivalence.transitive (Setoid.eq S) identLeft i
|
||||
g : ((0G ·A (x ·A inverse 0G)) ·A inverse x) ∼ 0G
|
||||
g = transferToRight'' G h
|
||||
ans : f ((0G ·A (x ·A inverse 0G)) ·A Group.inverse G x) ∼T Group.0G H
|
||||
ans = transitive (GroupHom.wellDefined fHom g) (imageOfIdentityIsIdentity fHom)
|
||||
GroupAction.associativeAction (conjugationNormalSubgroupAction {S = S} {T = T} {_·A_ = _·A_} G H {f} fHom) {x} {g} {h} = ans
|
||||
where
|
||||
open Group G
|
||||
open Setoid T renaming (_∼_ to _∼T_)
|
||||
open Setoid S renaming (_∼_ to _∼S_)
|
||||
open Equivalence (Setoid.eq T) renaming (transitive to transitiveH)
|
||||
open Equivalence (Setoid.eq S) renaming (transitive to transitiveG ; symmetric to symmetricG ; reflexive to reflexiveG)
|
||||
ans : f (((g ·A h) ·A (x ·A inverse (g ·A h))) ·A inverse ((g ·A ((h ·A (x ·A inverse h)) ·A inverse g)))) ∼T Group.0G H
|
||||
ans = transitiveH (GroupHom.wellDefined fHom (transferToRight'' G (transitiveG (symmetricG +Associative) (Group.+WellDefined G reflexiveG (transitiveG (+WellDefined reflexiveG (transitiveG (+WellDefined reflexiveG (invContravariant G)) +Associative)) +Associative))))) (imageOfIdentityIsIdentity fHom)
|
Reference in New Issue
Block a user