mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-08 05:18:40 +00:00
41 lines
1.6 KiB
Agda
41 lines
1.6 KiB
Agda
{-# 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.Actions.Definition
|
||
open import Groups.Groups2
|
||
open import Sets.EquivalenceRelations
|
||
open import Groups.Actions.Definition
|
||
|
||
module Groups.Actions.Stabiliser {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 _+_} (act : GroupAction G T) where
|
||
|
||
open GroupAction act
|
||
open Setoid T
|
||
|
||
stabiliserPred : (x : B) → (g : A) → Set d
|
||
stabiliserPred x g = (action g x) ∼ x
|
||
|
||
stabiliserWellDefined : (x : B) → {g h : A} → Setoid._∼_ S g h → (stabiliserPred x g) → stabiliserPred x h
|
||
stabiliserWellDefined x {g} {h} g=h gx=x = transitive (actionWellDefined1 (Equivalence.symmetric (Setoid.eq S) g=h)) gx=x
|
||
where
|
||
open Equivalence eq
|
||
|
||
stabiliserSubgroup : (x : B) → subgroup G (stabiliserWellDefined x)
|
||
_&_&_.one (stabiliserSubgroup x) gx=x hx=x = transitive associativeAction (transitive (actionWellDefined2 hx=x) gx=x)
|
||
where
|
||
open Equivalence eq
|
||
_&_&_.two (stabiliserSubgroup x) = identityAction
|
||
_&_&_.three (stabiliserSubgroup x) {g = g} gx=x = transitive (transitive (transitive (actionWellDefined2 (symmetric gx=x)) (symmetric associativeAction)) (actionWellDefined1 (invLeft {g}))) identityAction
|
||
where
|
||
open Equivalence eq
|
||
open Group G
|