agda-gset/Algebra/G-Set.agda

46 lines
1.4 KiB
Agda
Raw Normal View History

2024-05-05 14:43:35 +02:00
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