-- The SIP applied to groups
{-# OPTIONS --safe #-}
module Cubical.Algebra.Group.GroupPath where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.Equiv
open import Cubical.Foundations.Equiv.HalfAdjoint
open import Cubical.Foundations.HLevels
open import Cubical.Foundations.Transport
open import Cubical.Foundations.Univalence
open import Cubical.Foundations.SIP
open import Cubical.Foundations.Function using (_∘_)
open import Cubical.Foundations.GroupoidLaws hiding (assoc)
open import Cubical.Data.Sigma

open import Cubical.Displayed.Base
open import Cubical.Displayed.Auto
open import Cubical.Displayed.Record
open import Cubical.Displayed.Universe

open import Cubical.Algebra.Semigroup
open import Cubical.Algebra.Monoid

open import Cubical.Algebra.Group.Base
open import Cubical.Algebra.Group.Properties
open import Cubical.Algebra.Group.Morphisms
open import Cubical.Algebra.Group.MorphismProperties

private
  variable
     ℓ' ℓ'' : Level

open Iso
open GroupStr
open IsGroupHom

𝒮ᴰ-Group : DUARel (𝒮-Univ ) GroupStr 
𝒮ᴰ-Group =
  𝒮ᴰ-Record (𝒮-Univ _) IsGroupEquiv
    (fields:
      data[ _·_  autoDUARel _ _  pres· ]
      data[ 1g  autoDUARel _ _  pres1 ]
      data[ inv  autoDUARel _ _  presinv ]
      prop[ isGroup   _ _  isPropIsGroup _ _ _) ])
  where
  open GroupStr
  open IsGroupHom

GroupPath : (M N : Group )  GroupEquiv M N  (M  N)
GroupPath =  𝒮ᴰ-Group .UARel.ua

-- TODO: Induced structure results are temporarily inconvenient while we transition between algebra
-- representations
module _ (G : Group ) {A : Type } (m : A  A  A)
  (e :  G   A)
  ( :  x y  e .fst (G .snd ._·_ x y)  m (e .fst x) (e .fst y))
  where

  private
    module G = GroupStr (G .snd)

    FamilyΣ : Σ[ B  Type  ] (B  B  B)  Type 
    FamilyΣ (B , n) =
      Σ[ e  B ]
      Σ[ i  (B  B) ]
      IsGroup e n i

    inducedΣ : FamilyΣ (A , m)
    inducedΣ =
      subst FamilyΣ
        (UARel.≅→≡ (autoUARel (Σ[ B  Type  ] (B  B  B))) (e , ))
        (G.1g , G.inv , G.isGroup)

  InducedGroup : Group 
  InducedGroup .fst = A
  InducedGroup .snd ._·_ = m
  InducedGroup .snd .1g = inducedΣ .fst
  InducedGroup .snd .inv = inducedΣ .snd .fst
  InducedGroup .snd .isGroup = inducedΣ .snd .snd

  InducedGroupPath : G  InducedGroup
  InducedGroupPath = GroupPath _ _ .fst (e , makeIsGroupHom )

uaGroup : {G H : Group }  GroupEquiv G H  G  H
uaGroup {G = G} {H = H} = equivFun (GroupPath G H)

-- Group-ua functoriality
Group≡ : (G H : Group )  (
  Σ[ p   G    H  ]
  Σ[ q  PathP  i  p i) (1g (snd G)) (1g (snd H)) ]
  Σ[ r  PathP  i  p i  p i  p i) (_·_ (snd G)) (_·_ (snd H)) ]
  Σ[ s  PathP  i  p i  p i) (inv (snd G)) (inv (snd H)) ]
  PathP  i  IsGroup (q i) (r i) (s i)) (isGroup (snd G)) (isGroup (snd H)))
   (G  H)
Group≡ G H = isoToEquiv theIso
  where
  theIso : Iso _ _
  fun theIso (p , q , r , s , t) i = p i , groupstr (q i) (r i) (s i) (t i)
  inv theIso x = cong ⟨_⟩ x , cong (1g  snd) x , cong (_·_  snd) x , cong (inv  snd) x , cong (isGroup  snd) x
  rightInv theIso _ = refl
  leftInv theIso _ = refl

caracGroup≡ : {G H : Group } (p q : G  H)  cong ⟨_⟩ p  cong ⟨_⟩ q  p  q
caracGroup≡ {G = G} {H = H} p q P =
  sym (transportTransport⁻ (ua (Group≡ G H)) p)
                                   ∙∙ cong (transport (ua (Group≡ G H))) helper
                                   ∙∙ transportTransport⁻ (ua (Group≡ G H)) q
    where
    helper : transport (sym (ua (Group≡ G H))) p  transport (sym (ua (Group≡ G H))) q
    helper = Σ≡Prop
                _  isPropΣ
                         (isOfHLevelPathP' 1 (is-set (snd H)) _ _)
                         λ _  isPropΣ (isOfHLevelPathP' 1 (isSetΠ2 λ _ _  is-set (snd H)) _ _)
                         λ _  isPropΣ (isOfHLevelPathP' 1 (isSetΠ λ _  is-set (snd H)) _ _)
                         λ _  isOfHLevelPathP 1 (isPropIsGroup _ _ _) _ _)
               (transportRefl (cong ⟨_⟩ p)  P  sym (transportRefl (cong ⟨_⟩ q)))

uaGroupId : (G : Group )  uaGroup (idGroupEquiv {G = G})  refl
uaGroupId G = caracGroup≡ _ _ uaIdEquiv

uaCompGroupEquiv : {F G H : Group } (f : GroupEquiv F G) (g : GroupEquiv G H)
                  uaGroup (compGroupEquiv f g)  uaGroup f  uaGroup g
uaCompGroupEquiv f g = caracGroup≡ _ _ (
  cong ⟨_⟩ (uaGroup (compGroupEquiv f g))
    ≡⟨ uaCompEquiv _ _ 
  cong ⟨_⟩ (uaGroup f)  cong ⟨_⟩ (uaGroup g)
    ≡⟨ sym (cong-∙ ⟨_⟩ (uaGroup f) (uaGroup g)) 
  cong ⟨_⟩ (uaGroup f  uaGroup g) )

-- J-rule for GroupEquivs
GroupEquivJ : {G : Group } (P : (H : Group )  GroupEquiv G H  Type ℓ')
             P G idGroupEquiv
              {H} e  P H e
GroupEquivJ {G = G} P p {H} e =
  transport  i  P (GroupPath G H .fst e i)
    (transp  j  GroupEquiv G (GroupPath G H .fst e (i  ~ j))) i e))
      (subst (P G) (sym lem) p)
  where
  lem : transport  j  GroupEquiv G (GroupPath G H .fst e (~ j))) e
        idGroupEquiv
  lem = Σ≡Prop  _  isPropIsGroupHom _ _)
       (Σ≡Prop  _  isPropIsEquiv _)
         (funExt λ x   i  fst (fst (fst e .snd .equiv-proof
                          (transportRefl (fst (fst e) (transportRefl x i)) i))))
                          retEq (fst e) x))