{-

Proof of the standard formulation of the univalence theorem and
various consequences of univalence

- The ua constant and its computation rule (up to a path)
- Proof of univalence using that unglue is an equivalence ([EquivContr])
- Equivalence induction ([EquivJ], [elimEquiv])
- Univalence theorem ([univalence])
- The computation rule for ua ([uaβ])
- Isomorphism induction ([elimIso])

-}
{-# OPTIONS --cubical --safe #-}
module Cubical.Foundations.Univalence where

open import Cubical.Core.Glue

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Function
open import Cubical.Foundations.Isomorphism
open import Cubical.Foundations.Equiv
open import Cubical.Foundations.GroupoidLaws

private
  variable
     ℓ' : Level

-- The ua constant
ua :  {A B : Type }  A  B  A  B
ua {A = A} {B = B} e i = Glue B  { (i = i0)  (A , e)
                                   ; (i = i1)  (B , idEquiv B) })

uaIdEquiv : {A : Type }  ua (idEquiv A)  refl
uaIdEquiv {A = A} i j = Glue A {φ = i  ~ j  j}  _  A , idEquiv A)

-- Give detailed type to unglue, mainly for documentation purposes
unglueua :  {A B : Type }  (e : A  B)  (i : I) (x : ua e i)
            B [ _   { (i = i0)  e .fst x ; (i = i1)  x }) ]
unglueua e i x = inS (unglue (i  ~ i) x)

-- Proof of univalence using that unglue is an equivalence:

-- unglue is an equivalence
unglueIsEquiv :  (A : Type ) (φ : I)
                (f : PartialP φ  o  Σ[ T  Type  ] T  A)) 
                isEquiv {A = Glue A f} (unglue φ)
equiv-proof (unglueIsEquiv A φ f) = λ (b : A) 
  let u : I  Partial φ A
      u i = λ{ (φ = i1)  equivCtr (f 1=1 .snd) b .snd (~ i) }
      ctr : fiber (unglue φ) b
      ctr = ( glue  { (φ = i1)  equivCtr (f 1=1 .snd) b .fst }) (hcomp u b)
            , λ j  hfill u (inS b) (~ j))
  in ( ctr
     , λ (v : fiber (unglue φ) b) i 
         let u' : I  Partial (φ  ~ i  i) A
             u' j = λ { (φ = i1)  equivCtrPath (f 1=1 .snd) b v i .snd (~ j)
                      ; (i = i0)  hfill u (inS b) j
                      ; (i = i1)  v .snd (~ j) }
         in ( glue  { (φ = i1)  equivCtrPath (f 1=1 .snd) b v i .fst }) (hcomp u' b)
            , λ j  hfill u' (inS b) (~ j)))

-- Any partial family of equivalences can be extended to a total one
-- from Glue [ φ ↦ (T,f) ] A to A
unglueEquiv :  (A : Type ) (φ : I)
              (f : PartialP φ  o  Σ[ T  Type  ] T  A)) 
              (Glue A f)  A
unglueEquiv A φ f = ( unglue φ , unglueIsEquiv A φ f )


-- The following is a formulation of univalence proposed by Martín Escardó:
-- https://groups.google.com/forum/#!msg/homotopytypetheory/HfCB_b-PNEU/Ibb48LvUMeUJ
-- See also Theorem 5.8.4 of the HoTT Book.
--
-- The reason we have this formulation in the core library and not the
-- standard one is that this one is more direct to prove using that
-- unglue is an equivalence. The standard formulation can be found in
-- Cubical/Basics/Univalence.
--
EquivContr :  (A : Type )  isContr (Σ[ T  Type  ] T  A)
EquivContr { = } A =
  ( (A , idEquiv A)
  , idEquiv≡ )
 where
  idEquiv≡ : (y : Σ (Type )  T  T  A))  (A , idEquiv A)  y
  idEquiv≡ w = \ { i .fst                    Glue A (f i)
                 ; i .snd .fst               unglueEquiv _ _ (f i) .fst
                 ; i .snd .snd .equiv-proof  unglueEquiv _ _ (f i) .snd .equiv-proof
                 }
    where
      f :  i  PartialP (~ i  i)  x  Σ[ T  Type  ] T  A)
      f i = λ { (i = i0)  A , idEquiv A ; (i = i1)  w }

contrSinglEquiv : {A B : Type } (e : A  B)  (B , idEquiv B)  (A , e)
contrSinglEquiv {A = A} {B = B} e =
  isContr→isProp (EquivContr B) (B , idEquiv B) (A , e)

-- Equivalence induction
EquivJ : (P : (A B : Type )  (e : B  A)  Type ℓ')
        (r : (A : Type )  P A A (idEquiv A))
        (A B : Type )  (e : B  A)  P A B e
EquivJ P r A B e = subst  x  P A (x .fst) (x .snd)) (contrSinglEquiv e) (r A)

-- Eliminate equivalences by just looking at the underlying function
elimEquivFun : (B : Type ) (P : (A : Type )  (A  B)  Type ℓ')
              (r : P B  x  x))
              (A : Type )  (e : A  B)  P A (e .fst)
elimEquivFun B P r a e = subst  x  P (x .fst) (x .snd .fst)) (contrSinglEquiv e) r

-- Assuming that we have an inverse to ua we can easily prove univalence
module Univalence (au :  {} {A B : Type }  A  B  A  B)
                  (auid :  {} {A B : Type }  au refl  idEquiv A) where
  thm :  {} {A B : Type }  isEquiv au
  thm {A = A} {B = B} =
    isoToIsEquiv {B = A  B}
      (iso au ua
        (EquivJ  _ _ e  au (ua e)  e)  X  (cong au uaIdEquiv)  (auid {B = B})) _ _)
        (J  X p  ua (au p)  p) ((cong ua (auid {B = B}))  uaIdEquiv)))

pathToEquiv : {A B : Type }  A  B  A  B
pathToEquiv p = lineToEquiv  i  p i)

pathToEquivRefl : {A : Type }  pathToEquiv refl  idEquiv A
pathToEquivRefl {A = A} = equivEq _ _  i x  transp  _  A) i x)

-- Univalence
univalence : {A B : Type }  (A  B)  (A  B)
univalence = ( pathToEquiv , Univalence.thm pathToEquiv pathToEquivRefl  )

-- The original map from UniMath/Foundations
eqweqmap : {A B : Type }  A  B  A  B
eqweqmap {A = A} e = J  X _  A  X) (idEquiv A) e

eqweqmapid : {A : Type }  eqweqmap refl  idEquiv A
eqweqmapid {A = A} = JRefl  X _  A  X) (idEquiv A)

univalenceStatement : {A B : Type }  isEquiv (eqweqmap {} {A} {B})
univalenceStatement = Univalence.thm eqweqmap eqweqmapid

univalenceUAH : {A B : Type }  (A  B)  (A  B)
univalenceUAH = ( _ , univalenceStatement )

univalencePath : {A B : Type }  (A  B)  Lift (A  B)
univalencePath = ua (compEquiv univalence LiftEquiv)

-- The computation rule for ua. Because of "ghcomp" it is now very
-- simple compared to cubicaltt:
-- https://github.com/mortberg/cubicaltt/blob/master/examples/univalence.ctt#L202
uaβ : {A B : Type } (e : A  B) (x : A)  transport (ua e) x  e .fst x
uaβ e x = transportRefl (e .fst x)

-- Alternative version of EquivJ that only requires a predicate on
-- functions
elimEquiv : {B : Type } (P : {A : Type }  (A  B)  Type ℓ') 
            (d : P (idfun B))  {A : Type }  (e : A  B)  P (e .fst)
elimEquiv P d e = subst  x  P (x .snd .fst)) (contrSinglEquiv e) d

-- Isomorphism induction
elimIso : {B : Type }  (Q : {A : Type }  (A  B)  (B  A)  Type ℓ') 
          (h : Q (idfun B) (idfun B))  {A : Type } 
          (f : A  B)  (g : B  A)  section f g  retract f g  Q f g
elimIso {} {ℓ'} {B} Q h {A} f g sfg rfg = rem1 f g sfg rfg
  where
  P : {A : Type }  (f : A  B)  Type (ℓ-max ℓ' )
  P {A} f = (g : B  A)  section f g  retract f g  Q f g

  rem : P (idfun B)
  rem g sfg rfg = subst (Q (idfun B))  i b  (sfg b) (~ i)) h

  rem1 : {A : Type }  (f : A  B)  P f
  rem1 f g sfg rfg = elimEquiv P rem (f , isoToIsEquiv (iso f g sfg rfg)) g sfg rfg


uaInvEquiv :  {A B : Type }  (e : A  B)  ua (invEquiv e)  sym (ua e)
uaInvEquiv e = EquivJ  _ _ e  ua (invEquiv e)  sym (ua e)) rem _ _ e
  where
  rem : (A : Type )  ua (invEquiv (idEquiv A))  sym (ua (idEquiv A))
  rem A = cong ua (invEquivIdEquiv A)

uaCompEquiv :  {A B C : Type }  (e : A  B) (f : B  C)  ua (compEquiv e f)  ua e  ua f
uaCompEquiv {C = C} = EquivJ  A B e  (f : A  C)  ua (compEquiv e f)  ua e  ua f) rem _ _
  where
  rem : (A : Type _) (f : A  C)  ua (compEquiv (idEquiv A) f)  ua (idEquiv A)  ua f
  rem _ f = cong ua (compEquivIdEquiv f)  sym (cong  x  x  ua f) uaIdEquiv  sym (lUnit (ua f)))