open import Algebra.G-Set open import Algebra.Group open import Category.G-Sets open import Categories.Category.Instance.Sets open import Categories.Functor.Core open import Categories.Category.Core open import Relation.Binary.PropositionalEquality as ≡ using (_≡_) open import Relation.Binary.PropositionalEquality.WithK open import Axiom.UniquenessOfIdentityProofs.WithK open import Data.Product open import Data.Product.Properties renaming (Σ-≡,≡→≡ to peq) open import Level module Category.G-Sets.Properties.A4 {ℓ} (G : Group ℓ) where open Functor open G-Set-Morphism using () renaming (u to <_>) D : Functor (Sets ℓ) (G-Sets G) D .F₀ S = record { X = S ; _⊳_ = λ _ s → s ; ε⊳ = ≡.refl ; ∘⊳ = ≡.refl } D .F₁ {A} {B} f = record { u = f ; equivariance = ≡.refl } D .identity = ≡.refl D .homomorphism = ≡.refl D .F-resp-≈ f≡g = f≡g V : Functor (G-Sets G) (Sets ℓ) V .F₀ GS = Orb where open G-Set GS V .F₁ {A} {B} f (x , (y , (g , eq))) = (< f > x , (< f > y , (g , ≡.trans (≡.sym f.equivariance) (≡.cong < f > eq)))) where open G-Set A module f = G-Set-Morphism f V .identity {GS} {(x , (y , (g , ≡.refl)))} = ≡.refl V .homomorphism {X} {Y} {Z} {f} {g} = {!!} V .F-resp-≈ {A} {B} {f} {g} eq {(x , (y , (h , ≡.refl)))} = peq (eq , {!!}) where module f = G-Set-Morphism f module g = G-Set-Morphism g module A = G-Set A