open import Algebra.Group open import Algebra.G-Set open import Categories.Category.Core open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) open import Level module Category.G-Sets {ℓ : Level} where open Category open G-Set-Morphism using () renaming (u to <_>) G-Sets : Group ℓ → Category (suc ℓ) (suc ℓ) ℓ G-Sets G .Obj = G-Set G G-Sets G ._⇒_ = G-Set-Morphism G G-Sets G ._≈_ f g = ∀ {x} → f.u x ≡ g.u x where module f = G-Set-Morphism f module g = G-Set-Morphism g G-Sets G .id = record { u = λ x → x ; equivariance = ≡.refl } G-Sets G ._∘_ f g = record { u = λ x → f.u (g.u x) ; equivariance = ≡.trans (≡.cong < f > g.equivariance) f.equivariance } where module f = G-Set-Morphism f module g = G-Set-Morphism g G-Sets G .assoc = ≡.refl G-Sets G .sym-assoc = ≡.refl G-Sets G .identityˡ = ≡.refl G-Sets G .identityʳ = ≡.refl G-Sets G .identity² = ≡.refl G-Sets G .equiv = record { refl = ≡.refl ; sym = λ eq → ≡.sym eq ; trans = λ eq₁ eq₂ → ≡.trans eq₁ eq₂ } G-Sets G .∘-resp-≈ {X} {Y} {Z} {f} {h} {g} {i} f≡h g≡i = ≡.trans f≡h (≡.cong < h > g≡i)