{- Basic theory about transport:

- transport is invertible
- transport is an equivalence ([transportEquiv])

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

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

-- Direct definition of transport filler, note that we have to
-- explicitly tell Agda that the type is constant (like in CHM)
transpFill :  {} {A : Type }
             (φ : I)
             (A : (i : I)  Type  [ φ   _  A) ])
             (u0 : outS (A i0))
            --------------------------------------
             PathP  i  outS (A i)) u0 (transp  i  outS (A i)) φ u0)
transpFill φ A u0 i = transp  j  outS (A (i  j))) (~ i  φ) u0

transport⁻ :  {} {A B : Type }  A  B  B  A
transport⁻ p = transport  i  p (~ i))

transport⁻Transport :  {} {A B : Type }  (p : A  B)  (a : A) 
                          transport⁻ p (transport p a)  a
transport⁻Transport p a j =
  transp  i  p (~ i  ~ j)) j (transp  i  p (i  ~ j)) j a)

transportTransport⁻ :  {} {A B : Type }  (p : A  B)  (b : B) 
                        transport p (transport⁻ p b)  b
transportTransport⁻ p b j =
  transp  i  p (i  j)) j (transp  i  p (~ i  j)) j b)

-- Transport is an equivalence
isEquivTransport :  {} {A B : Type } (p : A  B)  isEquiv (transport p)
isEquivTransport {A = A} {B = B} p =
  transport  i  isEquiv  x  transp  j  p (i  j)) (~ i) x)) (idIsEquiv A)

transportEquiv :  {} {A B : Type }  A  B  A  B
transportEquiv p = (transport p , isEquivTransport p)

pathToIso :  {} {A B : Type }  A  B  Iso A B
pathToIso x = iso (transport x) (transport⁻ x ) ( transportTransport⁻ x) (transport⁻Transport x)

isSet-subst :  { ℓ′} {A : Type } {B : A  Type ℓ′}
                 (isSet-A : isSet A)
                  {a : A}
                 (p : a  a)  (x : B a)  subst B p x  x
isSet-subst {B = B} isSet-A p x = subst  p′  subst B p′ x  x) (isSet-A _ _ refl p) (substRefl {B = B} x)

-- substituting along a composite path is equivalent to substituting twice
substComposite-□ :  { ℓ′} {A : Type }  (B : A  Type ℓ′)
                      {x y z : A} (p : x  y) (q : y  z) (u : B x)
                      subst B (p  q) u  subst B q (subst B p u)
substComposite-□ B p q Bx = sym (substRefl {B = B} _)  helper where
  compSq : I  I  _
  compSq = compPath'-filler p q
  helper : subst B refl (subst B (p  q) Bx)  subst B q (subst B p Bx)
  helper i = subst B  k  compSq (~ i  ~ k) (~ i  k)) (subst B  k  compSq (~ i  ~ k) (~ i  k)) Bx)

-- substitution commutes with morphisms in slices
substCommSlice :  { ℓ′} {A : Type }
                    (B C : A  Type ℓ′)
                    (F :  i  B i  C i)
                    {x y : A} (p : x  y) (u : B x)
                    subst C p (F x u)  F y (subst B p u)
substCommSlice B C F p Bx i = comp pathC  k  λ where
      (i = i0)  toPathP {A = pathC}  _  subst C p (F _ Bx)) k
      (i = i1)  F (p k) (toPathP {A = pathB}  _  subst B p Bx) k)
    ) (F _ Bx) where
  pathC : I  Type _
  pathC i = cong C p i
  pathB : I  Type _
  pathB i = cong B p i