45 lines
1.4 KiB
Agda
45 lines
1.4 KiB
Agda
open import Algebra.Group
|
||
open import Level
|
||
open import Data.Product
|
||
open import Relation.Binary.PropositionalEquality
|
||
|
||
module Algebra.G-Set {ℓ : Level} where
|
||
open Group using () renaming (Carrier to ∣_∣)
|
||
record G-Set (G : Group ℓ) : Set (suc ℓ) where
|
||
open Group G using (ε; _∙_)
|
||
field
|
||
X : Set ℓ
|
||
_⊳_ : ∣ G ∣ → X → X
|
||
field
|
||
ε⊳ : ∀ {x : X} → ε ⊳ x ≡ x
|
||
∘⊳ : ∀ {g h : ∣ G ∣} {x : X} → (g ∙ h) ⊳ x ≡ (g ⊳ (h ⊳ x))
|
||
|
||
orb : X → Set ℓ
|
||
orb x = Σ[ y ∈ X ] Σ[ g ∈ ∣ G ∣ ] g ⊳ x ≡ y
|
||
|
||
Orb : Set ℓ
|
||
Orb = Σ[ x ∈ X ] orb x
|
||
|
||
-- record orb (x : X) : Set ℓ where
|
||
-- field
|
||
-- y : X
|
||
-- g : ∣ G ∣
|
||
-- .eq : g ⊳ x ≡ y
|
||
|
||
-- record Orb : Set ℓ where
|
||
-- field
|
||
-- x : X
|
||
-- O : orb x
|
||
|
||
open G-Set using () renaming (X to ∣_∣)
|
||
|
||
IsEquivariant : ∀ {G : Group ℓ} (X Y : G-Set G) (f : ∣ X ∣ → ∣ Y ∣) → Set ℓ
|
||
IsEquivariant {G} X Y f = ∀ {g : ∣ G ∣} {x : ∣ X ∣} → f (g ⊳ˣ x) ≡ g ⊳ʸ (f x)
|
||
where
|
||
open G-Set X using () renaming (_⊳_ to _⊳ˣ_)
|
||
open G-Set Y using () renaming (_⊳_ to _⊳ʸ_)
|
||
|
||
record G-Set-Morphism (G : Group ℓ) (X Y : G-Set G) : Set (suc ℓ) where
|
||
field
|
||
u : ∣ X ∣ → ∣ Y ∣ -- u for underlying
|
||
equivariance : IsEquivariant X Y u
|