Tom de Jong, 1-18 March 2021

We show that the induction principle for ๐•Šยน with propositional computation rules
follows from the universal property of ๐•Šยน.

This is claimed at the end of Section 6.2 in the HoTT Book and follows from a
general result by Sojakova in her PhD Thesis "Higher Inductive Types as
Homotopy-Initial Algebras" (CMU-CS-16-125). The proof of the general result is
quite complicated (see for instance Lemma 105 in the PhD thesis) and the
development below offers an alternative proof for ๐•Šยน.

\begin{code}

{-# OPTIONS --safe --without-K #-}

open import MLTT.Spartan
open import UF.Base
open import UF.Equiv
open import UF.FunExt
open import UF.Sets
open import UF.Sets-Properties
open import UF.Subsingletons

module Circle.Induction where

\end{code}

First some preliminaries on the free loop space.

\begin{code}

๐“› : (X : ๐“ค ฬ‡ ) โ†’ ๐“ค ฬ‡
๐“› X = ฮฃ x ๊ž‰ X , x ๏ผ x

๐“›-functor : {X : ๐“ค ฬ‡ } {Y : ๐“ฅ ฬ‡ } (f : X โ†’ Y) โ†’ ๐“› X โ†’ ๐“› Y
๐“›-functor f (x , p) = f x , ap f p

๐“›-functor-id : {X : ๐“ค ฬ‡ } โ†’ ๐“›-functor id โˆผ id {๐“ค} {๐“› X}
๐“›-functor-id {๐“ค} {X} (x , p) = to-ฮฃ-๏ผ (refl , ฮณ p)
 where
  ฮณ : {y z : X} (q : y ๏ผ z) โ†’ transport (ฮป - โ†’ y ๏ผ -) q refl ๏ผ q
  ฮณ refl = refl

๐“›-functor-comp : {X : ๐“ค ฬ‡ } {Y : ๐“ฅ ฬ‡ } {Z : ๐“ฆ ฬ‡ } (f : X โ†’ Y) (g : Y โ†’ Z)
               โ†’ ๐“›-functor g โˆ˜ ๐“›-functor f โˆผ ๐“›-functor (g โˆ˜ f)
๐“›-functor-comp f g (x , p) = to-ฮฃ-๏ผ (refl , (ap-ap f g p))

ap-๐“›-functor-lemma : {A : ๐“ค ฬ‡ } {B : ๐“ฅ ฬ‡ } (f g : A โ†’ B)
                     {a : A} (p : a ๏ผ a) {b : B} (q : b ๏ผ b)
                     (u : ๐“›-functor f (a , p) ๏ผ (b , q))
                     (v : ๐“›-functor g (a , p) ๏ผ (b , q))
                     (w : (f , u) ๏ผ (g , v))
                   โ†’ ap (ฮป - โ†’ ๐“›-functor - (a , p)) (ap prโ‚ w) ๏ผ u โˆ™ v โปยน
ap-๐“›-functor-lemma f g p q refl v refl = refl

happly-๐“›-functor-lemma : {A : ๐“ค ฬ‡ } {B : ๐“ฅ ฬ‡ } (f g : A โ†’ B)
                         {a : A} (p : a ๏ผ a) {b : B} (q : b ๏ผ b)
                         (u : ๐“›-functor f (a , p) ๏ผ (b , q))
                         (v : ๐“›-functor g (a , p) ๏ผ (b , q))
                         (w : (f , u) ๏ผ (g , v))
                       โ†’ happly (ap prโ‚ w) a ๏ผ (ap prโ‚ u) โˆ™ (ap prโ‚ v) โปยน
happly-๐“›-functor-lemma f g p q refl v refl = refl

\end{code}

Next we introduce the circle ๐•Šยน with its point base, its loop at base and its
universal property.

\begin{code}

module _
        (๐•Šยน : ๐“ค ฬ‡ )
        (base : ๐•Šยน)
        (loop : base ๏ผ base)
       where

 ๐•Šยน-universal-map : (A : ๐“ฅ ฬ‡ )
                  โ†’ (๐•Šยน โ†’ A) โ†’ ๐“› A
 ๐•Šยน-universal-map A f = (f base , ap f loop)

 module _
         (๐•Šยน-universal-property : {๐“ฅ : Universe} (A : ๐“ฅ ฬ‡ )
                                โ†’ is-equiv (๐•Šยน-universal-map A))
        where

  ๐•Šยน-uniqueness-principle : {A : ๐“ฅ ฬ‡ } (a : A) (p : a ๏ผ a)
                          โ†’ โˆƒ! r ๊ž‰ (๐•Šยน โ†’ A) , (r base , ap r loop) ๏ผ (a , p)
  ๐•Šยน-uniqueness-principle {๐“ฅ} {A} a p =
    equivs-are-vv-equivs (๐•Šยน-universal-map A)
                         (๐•Šยน-universal-property A) (a , p)

  ๐•Šยน-at-most-one-function : {A : ๐“ฅ ฬ‡ } (a : A) (p : a ๏ผ a)
                          โ†’ is-prop (ฮฃ r ๊ž‰ (๐•Šยน โ†’ A) , (r base , ap r loop) ๏ผ (a , p))
  ๐•Šยน-at-most-one-function a p = singletons-are-props (๐•Šยน-uniqueness-principle a p)

\end{code}

The recursion principle for ๐•Šยน with its computation rule follows immediately
from the universal property of ๐•Šยน.

\begin{code}

  ๐•Šยน-rec : {A : ๐“ฅ ฬ‡ } (a : A) (p : a ๏ผ a)
         โ†’ ๐•Šยน โ†’ A
  ๐•Šยน-rec {๐“ฅ} {A} a p = โˆƒ!-witness (๐•Šยน-uniqueness-principle a p)

  ๐•Šยน-rec-comp : {A : ๐“ฅ ฬ‡ } (a : A) (p : a ๏ผ a)
              โ†’ ๐“›-functor (๐•Šยน-rec a p) (base , loop) ๏ผ[ ๐“› A ] (a , p)
  ๐•Šยน-rec-comp {๐“ฅ} {A} a p = โˆƒ!-is-witness (๐•Šยน-uniqueness-principle a p)

  ๐•Šยน-rec-on-base : {A : ๐“ฅ ฬ‡ } (a : A) (p : a ๏ผ a)
                  โ†’ ๐•Šยน-rec a p base ๏ผ a
  ๐•Šยน-rec-on-base a p = ap prโ‚ (๐•Šยน-rec-comp a p)

  ๐•Šยน-rec-on-loop : {A : ๐“ฅ ฬ‡ } (a : A) (p : a ๏ผ a)
                 โ†’ transport (ฮป - โ†’ - ๏ผ -) (๐•Šยน-rec-on-base a p)
                     (ap (๐•Šยน-rec a p) loop)
                 ๏ผ p
  ๐•Šยน-rec-on-loop a p = from-ฮฃ-๏ผ' (๐•Šยน-rec-comp a p)

\end{code}

The induction principle for ๐•Šยน also follows quite directly. The idea is to turn
a type family A over ๐•Šยน to the type ฮฃ A and consider a nondependent map ๐•Šยน โ†’ ฮฃ A
as a substitute for the dependent function (x : ๐•Šยน) โ†’ A x.

What is significantly harder is showing that it obeys the computation rules.

\begin{code}

  module ๐•Šยน-induction
          (A : ๐•Šยน โ†’ ๐“ฅ ฬ‡ )
          (a : A base)
          (l : transport A loop a ๏ผ a)
         where

   lโบ : (base , a) ๏ผ[ ฮฃ A ] (base , a)
   lโบ = to-ฮฃ-๏ผ (loop , l)

   r : ๐•Šยน โ†’ ฮฃ A
   r = ๐•Šยน-rec (base , a) lโบ

\end{code}

Next we show that r is a retraction of prโ‚ : ฮฃ A โ†’ ๐•Šยน. This tells us that
r (x) = (x , prโ‚‚ (r x)), so that we can define ๐•Šยน-induction by transport.

\begin{code}

   r-retraction-lemma : ๐“›-functor (prโ‚ โˆ˜ r) (base , loop) ๏ผ[ ๐“› ๐•Šยน ] (base , loop)
   r-retraction-lemma =
    ((prโ‚ โˆ˜ r) base , ap (prโ‚ โˆ˜ r) loop) ๏ผโŸจ I   โŸฉ
    ๐“›-functor prโ‚ (r base , ap r loop)   ๏ผโŸจ II  โŸฉ
    (base , ap prโ‚ (to-ฮฃ-๏ผ (loop , l)))  ๏ผโŸจ III โŸฉ
    (base , loop)                        โˆŽ
     where
      I   = to-ฮฃ-๏ผ (refl , ((ap-ap r prโ‚ loop) โปยน))
      II  = ap (๐“›-functor prโ‚) (๐•Šยน-rec-comp (base , a) lโบ)
      III = to-ฮฃ-๏ผ (refl , (ap-prโ‚-to-ฮฃ-๏ผ (loop , l)))

   r-is-retraction-of-prโ‚ : prโ‚ โˆ˜ r ๏ผ id
   r-is-retraction-of-prโ‚ = ap prโ‚ (๐•Šยน-at-most-one-function base loop
                                     (prโ‚ โˆ˜ r , r-retraction-lemma)
                                     (id , to-ฮฃ-๏ผ (refl , ap-id-is-id loop)))

   ๐•Šยน-induction : (x : ๐•Šยน) โ†’ A x
   ๐•Šยน-induction x = transport A (happly r-is-retraction-of-prโ‚ x) (prโ‚‚ (r x))

\end{code}

Next we set out to prove the computation rules for ๐•Šยน-induction.

\begin{code}

   ฯ : ๐•Šยน โ†’ ฮฃ A
   ฯ x = (x , ๐•Šยน-induction x)

   r-comp : (r base , ap r loop) ๏ผ[ ๐“› (ฮฃ A) ] ((base , a) , lโบ)
   r-comp = ๐•Šยน-rec-comp (base , a) lโบ

   ฯ-r-homotopy : ฯ โˆผ r
   ฯ-r-homotopy x = to-ฮฃ-๏ผ ((ฮณโ‚ โปยน) , ฮณโ‚‚)
    where
     ฮณโ‚ : prโ‚ (r x) ๏ผ prโ‚ (ฯ x)
     ฮณโ‚ = happly r-is-retraction-of-prโ‚ x
     ฮณโ‚‚ = transport A (ฮณโ‚ โปยน) (prโ‚‚ (ฯ x))                  ๏ผโŸจ refl โŸฉ
          transport A (ฮณโ‚ โปยน) (transport A ฮณโ‚ (prโ‚‚ (r x))) ๏ผโŸจ I    โŸฉ
          transport A (ฮณโ‚ โˆ™ ฮณโ‚ โปยน) (prโ‚‚ (r x))             ๏ผโŸจ II   โŸฉ
          transport A refl (prโ‚‚ (r x))                     ๏ผโŸจ refl โŸฉ
          prโ‚‚ (r x)                                        โˆŽ
      where
       I  = (transport-โˆ™ A ฮณโ‚ (ฮณโ‚ โปยน)) โปยน
       II = ap (ฮป - โ†’ transport A - (prโ‚‚ (r x))) ((right-inverse ฮณโ‚) โปยน)

   ฯ-and-r-on-base-and-loop : (ฯ base , ap ฯ loop) ๏ผ[ ๐“› (ฮฃ A) ] (r base , ap r loop)
   ฯ-and-r-on-base-and-loop = to-ฮฃ-๏ผ (ฯ-r-homotopy base , ฮณ)
    where
     ฮณ = transport (ฮป - โ†’ - ๏ผ -) (ฯ-r-homotopy base) (ap ฯ loop) ๏ผโŸจ I  โŸฉ
         ฯ-r-homotopy base โปยน โˆ™ ap ฯ loop โˆ™ ฯ-r-homotopy base    ๏ผโŸจ II โŸฉ
         ap r loop                                               โˆŽ
      where
       I  = transport-along-๏ผ (ฯ-r-homotopy base) (ap ฯ loop)
       II = homotopies-are-natural'' ฯ r ฯ-r-homotopy {base} {base} {loop}

   ฯ-comp : (ฯ base , ap ฯ loop) ๏ผ[ ๐“› (ฮฃ A) ] ((base , a) , lโบ)
   ฯ-comp = ฯ-and-r-on-base-and-loop โˆ™ r-comp

\end{code}

Looking at ฯ-comp, we see that ฯ base = (base , ๐•Šยน-induction base) ๏ผ (base , a),
which looks promising, for if we can show that the equality in the first
component is refl, then ๐•Šยน-induction base ๏ผ a would follow. So that's exactly
what we do next.

\begin{code}

   ฯ-comp-lemma : ap prโ‚ (ap prโ‚ ฯ-comp) ๏ผ refl
   ฯ-comp-lemma =
    ap prโ‚ (ap prโ‚ ฯ-comp)                                          ๏ผโŸจ I   โŸฉ
    ap (prโ‚ โˆ˜ prโ‚) ฯ-comp                                           ๏ผโŸจ II  โŸฉ
    ap (prโ‚ โˆ˜ prโ‚) ฯ-and-r-on-base-and-loop โˆ™ ap (prโ‚ โˆ˜ prโ‚) r-comp ๏ผโŸจ III โŸฉ
    p โปยน โˆ™ p                                                        ๏ผโŸจ IV  โŸฉ
    refl                                                            โˆŽ
    where
     p = happly r-is-retraction-of-prโ‚ base
     I   = ap-ap prโ‚ prโ‚ ฯ-comp
     II  = ap-โˆ™ (prโ‚ โˆ˜ prโ‚) ฯ-and-r-on-base-and-loop r-comp
     IV  = left-inverse p
     III = apโ‚‚ _โˆ™_ ฮณโ‚ ฮณโ‚‚
      where
       ฮณโ‚ : ap (prโ‚ โˆ˜ prโ‚) ฯ-and-r-on-base-and-loop  ๏ผ p โปยน
       ฮณโ‚ = ap (prโ‚ โˆ˜ prโ‚) ฯ-and-r-on-base-and-loop  ๏ผโŸจ Iโ‚   โŸฉ
            ap prโ‚ (ap prโ‚ ฯ-and-r-on-base-and-loop) ๏ผโŸจ IIโ‚  โŸฉ
            ap prโ‚ (ฯ-r-homotopy base)               ๏ผโŸจ IIIโ‚ โŸฉ
            p โปยน                                     โˆŽ
        where
         Iโ‚   = (ap-ap prโ‚ prโ‚ ฯ-and-r-on-base-and-loop) โปยน
         IIโ‚  = ap (ap prโ‚) (ap-prโ‚-to-ฮฃ-๏ผ (ฯ-r-homotopy base , _))
         IIIโ‚ = ap-prโ‚-to-ฮฃ-๏ผ ((p โปยน) , _)
       ฮณโ‚‚ : ap (prโ‚ โˆ˜ prโ‚) r-comp ๏ผ p
       ฮณโ‚‚ = ฯ• โปยน
        where
         ฮบ = r-retraction-lemma
         ฯ• = p                                                     ๏ผโŸจ Iโ‚‚    โŸฉ
             ap prโ‚ ฮบ โˆ™ ap ฯ€ (to-ฮฃ-๏ผ (refl , ap-id-is-id loop)) โปยน ๏ผโŸจ IIโ‚‚   โŸฉ
             ap prโ‚ ฮบ โˆ™ refl โปยน                                    ๏ผโŸจ refl  โŸฉ
             ap prโ‚ ฮบ                                              ๏ผโŸจ IIIโ‚‚  โŸฉ
             ap prโ‚ (ap prโ‚ r-comp)                                ๏ผโŸจ IVโ‚‚   โŸฉ
             ap (prโ‚ โˆ˜ prโ‚) r-comp                                 โˆŽ
          where
           ฯ€ : ๐“› (๐•Šยน) โ†’ ๐•Šยน
           ฯ€ = prโ‚
           Iโ‚‚   = happly-๐“›-functor-lemma (prโ‚ โˆ˜ r) id loop loop
                   ฮบ (to-ฮฃ-๏ผ (refl , ap-id-is-id loop))
                   (๐•Šยน-at-most-one-function base loop
                     (prโ‚ โˆ˜ r , r-retraction-lemma)
                     (id , to-ฮฃ-๏ผ (refl , ap-id-is-id loop)))
           IIโ‚‚  = ap (ฮป - โ†’ ap prโ‚ ฮบ โˆ™ - โปยน)
                   (ap-prโ‚-to-ฮฃ-๏ผ {๐“ค} {๐“ค} {๐•Šยน} {ฮป - โ†’ (- ๏ผ -)} {_} {_}
                    (refl , ap-id-is-id loop))
           IVโ‚‚  = ap-ap prโ‚ prโ‚ r-comp
           IIIโ‚‚ = ap prโ‚ ฮบ                        ๏ผโŸจ refl โŸฉ
                  ap prโ‚ (ฮบโ‚ โˆ™ (ฮบโ‚‚ โˆ™ ฮบโ‚ƒ))         ๏ผโŸจ I'   โŸฉ
                  ap prโ‚ ฮบโ‚ โˆ™ ap prโ‚ (ฮบโ‚‚ โˆ™ ฮบโ‚ƒ)    ๏ผโŸจ II'  โŸฉ
                  refl โˆ™ ap prโ‚ (ฮบโ‚‚ โˆ™ ฮบโ‚ƒ)         ๏ผโŸจ III' โŸฉ
                  ap prโ‚ (ฮบโ‚‚ โˆ™ ฮบโ‚ƒ)                ๏ผโŸจ IV'  โŸฉ
                  ap prโ‚ ฮบโ‚‚ โˆ™ ap prโ‚ ฮบโ‚ƒ           ๏ผโŸจ V'   โŸฉ
                  ap prโ‚ ฮบโ‚‚ โˆ™ refl                ๏ผโŸจ refl โŸฉ
                  ap prโ‚ ฮบโ‚‚                       ๏ผโŸจ VI'  โŸฉ
                  ap (prโ‚ โˆ˜ ๐“›-functor prโ‚) r-comp ๏ผโŸจ refl โŸฉ
                  ap (prโ‚ โˆ˜ prโ‚) r-comp           ๏ผโŸจ VII' โŸฉ
                  ap prโ‚ (ap prโ‚ r-comp)          โˆŽ
                  where
                   ฮบโ‚ = to-ฮฃ-๏ผ (refl , ((ap-ap r prโ‚ loop) โปยน))
                   ฮบโ‚‚ = ap (๐“›-functor prโ‚) r-comp
                   ฮบโ‚ƒ = to-ฮฃ-๏ผ (refl , (ap-prโ‚-to-ฮฃ-๏ผ (loop , l)))
                   I'   = ap-โˆ™ prโ‚ ฮบโ‚ (ฮบโ‚‚ โˆ™ ฮบโ‚ƒ)
                   II'  = ap (_โˆ™ (ap prโ‚ (ฮบโ‚‚ โˆ™ ฮบโ‚ƒ)))
                           (ap-prโ‚-to-ฮฃ-๏ผ {๐“ค} {๐“ค} {๐•Šยน} {ฮป - โ†’ (- ๏ผ -)} {_} {_}
                            (refl , ((ap-ap r prโ‚ loop) โปยน)))
                   III' = refl-left-neutral
                   IV'  = ap-โˆ™ prโ‚ ฮบโ‚‚ ฮบโ‚ƒ
                   V'   = ap ((ap prโ‚ ฮบโ‚‚) โˆ™_)
                           (ap-prโ‚-to-ฮฃ-๏ผ {๐“ค} {๐“ค} {๐•Šยน} {ฮป - โ†’ (- ๏ผ -)} {_} {_}
                            (refl , ap-prโ‚-to-ฮฃ-๏ผ (loop , l)))
                   VI'  = ap-ap (๐“›-functor prโ‚) prโ‚ r-comp
                   VII' = (ap-ap prโ‚ prโ‚ r-comp) โปยน

   ๐•Šยน-induction-on-base : ๐•Šยน-induction base ๏ผ a
   ๐•Šยน-induction-on-base =
    transport (ฮป - โ†’ transport A - (๐•Šยน-induction base) ๏ผ a) ฯ-comp-lemma ฮณ
     where
      ฮณ : transport A (ap prโ‚ (ap prโ‚ ฯ-comp)) (๐•Šยน-induction base) ๏ผ a
      ฮณ = from-ฮฃ-๏ผ' (ap prโ‚ ฯ-comp)

\end{code}

This takes care of the first computation rule for ๐•Šยน-induction. We can
get a fairly direct proof of the second computation rule (the one for
loop) by assuming that base ๏ผ base is a set, because this tells us
that every element of loop ๏ผ loop must be refl.

We can satisfy this assumption for our intended application (see
CircleConstruction.lagda), because for the construction involving โ„ค-torsors it's
is quite easy to prove that base ๏ผ base is a set.

However, for completeness sake, below we also show that assuming function
extensionality and univalence, it is possible to prove that base ๏ผ base is a
set, by using both computation rules for ๐•Šยน-rec and the first computation rule
for ๐•Šยน-induction.

\begin{code}

   ๐•Šยน-induction-on-loop-lemma : (loop , transport (ฮป - โ†’ transport A loop - ๏ผ -)
                                         ๐•Šยน-induction-on-base
                                         (apd ๐•Šยน-induction loop))
                              ๏ผ (loop , l)
   ๐•Šยน-induction-on-loop-lemma =
      (fromto-ฮฃ-๏ผ (loop , transport (ฮป - โ†’ transport A loop - ๏ผ -) ฯƒ ฯ„)) โปยน
    โˆ™ (ap from-ฮฃ-๏ผ ฮณ) โˆ™ (fromto-ฮฃ-๏ผ (loop , l))
     where
      ฯƒ = ๐•Šยน-induction-on-base
      ฯ„ = apd ๐•Šยน-induction loop
      ฮณ : to-ฮฃ-๏ผ (loop , transport (ฮป - โ†’ transport A loop - ๏ผ -) ฯƒ ฯ„)
        ๏ผ to-ฮฃ-๏ผ (loop , l)
      ฮณ = to-ฮฃ-๏ผ (loop , transport (ฮป - โ†’ transport A loop - ๏ผ -) ฯƒ ฯ„)    ๏ผโŸจ I   โŸฉ
          transport (ฮป - โ†’ - ๏ผ -) (to-ฮฃ-๏ผ (refl , ฯƒ)) (to-ฮฃ-๏ผ (loop , ฯ„)) ๏ผโŸจ II  โŸฉ
          transport (ฮป - โ†’ - ๏ผ -) (ap prโ‚ ฯ-comp) (to-ฮฃ-๏ผ (loop , ฯ„))     ๏ผโŸจ III โŸฉ
          transport (ฮป - โ†’ - ๏ผ -) (ap prโ‚ ฯ-comp) (ap ฯ loop)             ๏ผโŸจ IV  โŸฉ
          to-ฮฃ-๏ผ (loop , l)                                               โˆŽ
       where
        I   = h loop ฯƒ ฯ„
         where
          h : {X : ๐“ฆ ฬ‡ } {Y : X โ†’ ๐“ฃ ฬ‡ } {x : X} (p : x ๏ผ x) {y y' : Y x}
              (q : y ๏ผ y') (q' : transport Y p y ๏ผ y)
            โ†’ to-ฮฃ-๏ผ (p , transport (ฮป - โ†’ transport Y p - ๏ผ -) q q')
            ๏ผ transport (ฮป - โ†’ - ๏ผ -) (to-ฮฃ-๏ผ (refl , q)) (to-ฮฃ-๏ผ (p , q'))
          h p refl q' = refl
        II  = ap (ฮป - โ†’ transport (ฮป - โ†’ - ๏ผ -) - (to-ฮฃ-๏ผ (loop , ฯ„))) h
         where
          h = to-ฮฃ-๏ผ (refl , ฯƒ)                 ๏ผโŸจ I'  โŸฉ
              to-ฮฃ-๏ผ (from-ฮฃ-๏ผ (ap prโ‚ ฯ-comp)) ๏ผโŸจ II' โŸฉ
              ap prโ‚ ฯ-comp                     โˆŽ
           where
            I'  = (ap to-ฮฃ-๏ผ (to-ฮฃ-๏ผ (ฯ-comp-lemma , refl))) โปยน
            II' = tofrom-ฮฃ-๏ผ (ap prโ‚ ฯ-comp)
        III = ap (transport (ฮป - โ†’ - ๏ผ -) (ap prโ‚ ฯ-comp)) (h ๐•Šยน-induction loop)
         where
          h : {X : ๐“ฆ ฬ‡ } {Y : X โ†’ ๐“ฃ ฬ‡ } (f : (x : X) โ†’ Y x)
              {x x' : X} (p : x ๏ผ x')
            โ†’ to-ฮฃ-๏ผ (p , apd f p) ๏ผ ap (ฮป x โ†’ (x , f x)) p
          h f refl = refl
        IV  = from-ฮฃ-๏ผ' ฯ-comp

   module _
           (base-sethood : is-set (base ๏ผ base))
          where

    ๐•Šยน-induction-on-loop : transport (ฮป - โ†’ transport A loop - ๏ผ -)
                            ๐•Šยน-induction-on-base (apd ๐•Šยน-induction loop)
                         ๏ผ l
    ๐•Šยน-induction-on-loop =
     ap-prโ‚-refl-lemma (ฮป - โ†’ transport A - a ๏ผ a) ๐•Šยน-induction-on-loop-lemma ฮณ
     where
      ฮณ : ap prโ‚ ๐•Šยน-induction-on-loop-lemma ๏ผ refl
      ฮณ = base-sethood (ap prโ‚ ๐•Šยน-induction-on-loop-lemma) refl

    ๐•Šยน-induction-comp : (๐•Šยน-induction base , apd ๐•Šยน-induction loop)
                      ๏ผ[ ฮฃ y ๊ž‰ A base , transport A loop y ๏ผ y ] (a , l)
    ๐•Šยน-induction-comp = to-ฮฃ-๏ผ (๐•Šยน-induction-on-base , ๐•Šยน-induction-on-loop)

\end{code}

As promised above, here follows a proof, assuming function
extensionality and univalence, that base ๏ผ base is a set, using both
computation rules for ๐•Šยน-rec and the first computation rule for
๐•Šยน-induction.

The proof uses the encode-decode method (Section 8.1.4 of the HoTT
Book) to show that base ๏ผ base is a retract of โ„ค. Since sets are
closed under retracts, the claim follows.

\begin{code}

  open import Circle.Integers
  open import Circle.Integers-Properties

  open import UF.Univalence

  module _
          (ua : is-univalent ๐“คโ‚€)
         where

   succ-โ„ค-๏ผ : โ„ค ๏ผ โ„ค
   succ-โ„ค-๏ผ = eqtoid ua โ„ค โ„ค succ-โ„ค-โ‰ƒ

   code : ๐•Šยน โ†’ ๐“คโ‚€ ฬ‡
   code = ๐•Šยน-rec โ„ค succ-โ„ค-๏ผ

\end{code}

   Using the first computation rule for ๐•Šยน-rec:

\begin{code}

   code-on-base : code base ๏ผ โ„ค
   code-on-base = ๐•Šยน-rec-on-base โ„ค succ-โ„ค-๏ผ

   โ„ค-to-code-base : โ„ค โ†’ code base
   โ„ค-to-code-base = Idtofun (code-on-base โปยน)

   code-base-to-โ„ค : code base โ†’ โ„ค
   code-base-to-โ„ค = Idtofun code-on-base

   transport-code-loop-is-succ-โ„ค : code-base-to-โ„ค
                                 โˆ˜ transport code loop
                                 โˆ˜ โ„ค-to-code-base
                                 ๏ผ succ-โ„ค
   transport-code-loop-is-succ-โ„ค =
    ฮด โˆ˜ transport code loop โˆ˜ ฮต                  ๏ผโŸจ I    โŸฉ
    ฮด โˆ˜ transport id acl โˆ˜ ฮต                     ๏ผโŸจ refl โŸฉ
    Idtofun cob โˆ˜ Idtofun acl โˆ˜ Idtofun (cob โปยน) ๏ผโŸจ II   โŸฉ
    Idtofun cob โˆ˜ Idtofun (cob โปยน โˆ™ acl)         ๏ผโŸจ III  โŸฉ
    Idtofun (cob โปยน โˆ™ acl โˆ™ cob)                 ๏ผโŸจ IV   โŸฉ
    Idtofun succ-โ„ค-๏ผ                             ๏ผโŸจ V    โŸฉ
    succ-โ„ค                                       โˆŽ
     where
      cob = code-on-base
      acl = ap code loop
      ฮต = โ„ค-to-code-base
      ฮด = code-base-to-โ„ค
      I   = ap (ฮป - โ†’ ฮด โˆ˜ - โˆ˜ ฮต) (transport-ap' id code loop)
      II  = ap (_โˆ˜_ (Idtofun cob)) ((Idtofun-โˆ™ ua (cob โปยน) acl) โปยน)
      III = (Idtofun-โˆ™ ua (cob โปยน โˆ™ acl) cob) โปยน

\end{code}

      Using the second computation rule for ๐•Šยน-rec

\begin{code}

      IV  = ap Idtofun ((transport-along-๏ผ cob acl) โปยน
                       โˆ™ (๐•Šยน-rec-on-loop โ„ค succ-โ„ค-๏ผ))
      V   = Idtofun-eqtoid ua succ-โ„ค-โ‰ƒ

   transport-code-loopโปยน-is-pred-โ„ค : code-base-to-โ„ค
                                   โˆ˜ transport code (loop โปยน)
                                   โˆ˜ โ„ค-to-code-base
                                   โˆผ pred-โ„ค
   transport-code-loopโปยน-is-pred-โ„ค x = equivs-are-lc succ-โ„ค succ-โ„ค-is-equiv ฮณ
    where
     ฮณ : (succ-โ„ค โˆ˜ code-base-to-โ„ค โˆ˜ transport code (loop โปยน) โˆ˜ โ„ค-to-code-base) x
       ๏ผ (succ-โ„ค โˆ˜ pred-โ„ค) x
     ฮณ = (succ-โ„ค โˆ˜ ฮด โˆ˜ tโปยน โˆ˜ ฮต) x    ๏ผโŸจ I   โŸฉ
         (ฮด โˆ˜ t โˆ˜ ฮต โˆ˜ ฮด โˆ˜ tโปยน โˆ˜ ฮต) x ๏ผโŸจ II  โŸฉ
         (ฮด โˆ˜ t โˆ˜ tโปยน โˆ˜ ฮต) x         ๏ผโŸจ III โŸฉ
         (ฮด โˆ˜ ฮต) x                   ๏ผโŸจ IV  โŸฉ
         x                           ๏ผโŸจ V   โŸฉ
         (succ-โ„ค โˆ˜ pred-โ„ค) x         โˆŽ
      where
       ฮต = โ„ค-to-code-base
       ฮด = code-base-to-โ„ค
       tโปยน = transport code (loop โปยน)
       t   = transport code loop
       I   = happly (transport-code-loop-is-succ-โ„ค โปยน) ((ฮด โˆ˜ tโปยน โˆ˜ ฮต) x)
       II  = ap (ฮด โˆ˜ t) (Idtofun-section code-on-base (tโปยน (ฮต x)))
       III = ap ฮด (back-and-forth-transport loop)
       IV  = Idtofun-retraction code-on-base x
       V   = (succ-โ„ค-is-retraction x) โปยน

   transport-code-loopโปยน-is-pred-โ„ค' : transport code (loop โปยน)
                                    โˆผ โ„ค-to-code-base โˆ˜ pred-โ„ค โˆ˜ code-base-to-โ„ค
   transport-code-loopโปยน-is-pred-โ„ค' x =
    transport code (loop โปยน) x                   ๏ผโŸจ I   โŸฉ
    (ฮต โˆ˜ ฮด โˆ˜ transport code (loop โปยน)) x         ๏ผโŸจ II  โŸฉ
    (ฮต โˆ˜ ฮด โˆ˜ transport code (loop โปยน) โˆ˜ ฮต โˆ˜ ฮด) x ๏ผโŸจ III โŸฉ
    (ฮต โˆ˜ pred-โ„ค โˆ˜ ฮด) x                           โˆŽ
     where
      ฮต = โ„ค-to-code-base
      ฮด = code-base-to-โ„ค
      I   = (Idtofun-section code-on-base (transport code (loop โปยน) x)) โปยน
      II  = ap (ฮต โˆ˜ ฮด โˆ˜ transport code (loop โปยน))
             ((Idtofun-section code-on-base x) โปยน)
      III = ap ฮต (transport-code-loopโปยน-is-pred-โ„ค (ฮด x))

   encode : (x : ๐•Šยน) โ†’ (base ๏ผ x) โ†’ code x
   encode x p = transport code p (โ„ค-to-code-base ๐ŸŽ)

   iterated-path : {X : ๐“ฆ ฬ‡ } {x : X} โ†’ x ๏ผ x โ†’ โ„• โ†’ x ๏ผ x
   iterated-path p zero     = refl
   iterated-path p (succ n) = p โˆ™ iterated-path p n

   iterated-path-comm : {X : ๐“ฆ ฬ‡ } {x : X} (p : x ๏ผ x) (n : โ„•)
                      โ†’ iterated-path p n โˆ™ p ๏ผ p โˆ™ iterated-path p n
   iterated-path-comm p zero = refl โˆ™ p ๏ผโŸจ refl-left-neutral โŸฉ
                               p        ๏ผโŸจ refl              โŸฉ
                               p โˆ™ refl โˆŽ
   iterated-path-comm p (succ n) = p โˆ™ iterated-path p n โˆ™ p   ๏ผโŸจ I  โŸฉ
                                   p โˆ™ (iterated-path p n โˆ™ p) ๏ผโŸจ II โŸฉ
                                   p โˆ™ (p โˆ™ iterated-path p n) โˆŽ
    where
     I  =  โˆ™assoc p (iterated-path p n) p
     II = ap (p โˆ™_) (iterated-path-comm p n)

   loops : โ„ค โ†’ base ๏ผ base
   loops ๐ŸŽ       = refl
   loops (pos n) = iterated-path loop (succ n)
   loops (neg n) = iterated-path (loop โปยน) (succ n)

   module _
           (fe : funext ๐“คโ‚€ ๐“ค)
          where

    open import UF.Lower-FunExt

    loops-lemma : (_โˆ™ loop) โˆ˜ loops โˆ˜ pred-โ„ค ๏ผ loops
    loops-lemma = dfunext fe h
     where
      h : (k : โ„ค) โ†’ loops (pred-โ„ค k) โˆ™ loop ๏ผ loops k
      h ๐ŸŽ = loop โปยน โˆ™ refl โˆ™ loop ๏ผโŸจ refl              โŸฉ
            loop โปยน โˆ™ loop        ๏ผโŸจ left-inverse loop โŸฉ
            refl                  โˆŽ
      h (pos n) = g n
       where
        g : (n : โ„•) โ†’ loops (pred-โ„ค (pos n)) โˆ™ loop ๏ผ loops (pos n)
        g zero     = iterated-path-comm loop zero
        g (succ n) = iterated-path-comm loop (succ n)
      h (neg n) =
       loop โปยน โˆ™ (loop โปยน โˆ™ iterated-path (loop โปยน) n) โˆ™ loop ๏ผโŸจ I'   โŸฉ
       loop โปยน โˆ™ (iterated-path (loop โปยน) n โˆ™ loop โปยน) โˆ™ loop ๏ผโŸจ II'  โŸฉ
       loop โปยน โˆ™ iterated-path (loop โปยน) n โˆ™ (loop โปยน โˆ™ loop) ๏ผโŸจ III' โŸฉ
       loop โปยน โˆ™ iterated-path (loop โปยน) n                    โˆŽ
        where
         I'   = ap (ฮป - โ†’ loop โปยน โˆ™ - โˆ™ loop)
                 ((iterated-path-comm (loop โปยน) n) โปยน)
         II'  = โˆ™assoc (loop โปยน) (iterated-path (loop โปยน) n โˆ™ loop โปยน) loop
              โˆ™ ap (loop โปยน โˆ™_)
                 (โˆ™assoc (iterated-path (loop โปยน) n) (loop โปยน) loop)
              โˆ™ (โˆ™assoc (loop โปยน) (iterated-path (loop โปยน) n)
                  (loop โปยน โˆ™ loop)) โปยน
         III' = ap ((loop โปยน โˆ™ iterated-path (loop โปยน) n) โˆ™_)
                 (left-inverse loop)

    transport-loops-lemma : transport (ฮป - โ†’ code - โ†’ base ๏ผ -) loop
                             (loops โˆ˜ code-base-to-โ„ค)
                          ๏ผ (loops โˆ˜ code-base-to-โ„ค)
    transport-loops-lemma =
     transport (ฮป - โ†’ code - โ†’ base ๏ผ -) loop f                     ๏ผโŸจ I   โŸฉ
     transport (ฮป - โ†’ base ๏ผ -) loop โˆ˜ f โˆ˜ transport code (loop โปยน) ๏ผโŸจ II  โŸฉ
     (_โˆ™ loop) โˆ˜ f โˆ˜ transport code (loop โปยน)                       ๏ผโŸจ III โŸฉ
     (_โˆ™ loop) โˆ˜ loops โˆ˜ ฮด โˆ˜ ฮต โˆ˜ pred-โ„ค โˆ˜ ฮด                         ๏ผโŸจ IV  โŸฉ
     (_โˆ™ loop) โˆ˜ loops โˆ˜ pred-โ„ค โˆ˜ ฮด                                 ๏ผโŸจ V   โŸฉ
     loops โˆ˜ ฮด                                                      โˆŽ
      where
       ฮต : โ„ค โ†’ code base
       ฮต = โ„ค-to-code-base
       ฮด : code base โ†’ โ„ค
       ฮด = code-base-to-โ„ค
       f : code base โ†’ base ๏ผ base
       f = loops โˆ˜ ฮด
       I   = transport-along-โ†’ code (_๏ผ_ base) loop f
       II  = refl
       III = ap ((_โˆ™ loop) โˆ˜ f โˆ˜_)
              (dfunext (lower-funext ๐“คโ‚€ ๐“ค fe) transport-code-loopโปยน-is-pred-โ„ค')
       IV  = ap (ฮป - โ†’ (_โˆ™ loop) โˆ˜ loops โˆ˜ - โˆ˜ pred-โ„ค โˆ˜ ฮด)
              (dfunext (lower-funext ๐“คโ‚€ ๐“ค fe) (Idtofun-retraction code-on-base))
       V   = ap (_โˆ˜ ฮด) loops-lemma


    open ๐•Šยน-induction (ฮป - โ†’ code - โ†’ base ๏ผ -)
                      (loops โˆ˜ code-base-to-โ„ค)
                      transport-loops-lemma

    decode : (x : ๐•Šยน) โ†’ code x โ†’ base ๏ผ x
    decode = ๐•Šยน-induction

    decode-encode : (x : ๐•Šยน) (p : base ๏ผ x) โ†’ decode x (encode x p) ๏ผ p
    decode-encode base refl =
     decode base (encode base refl)                       ๏ผโŸจ refl โŸฉ
     decode base (transport code refl (โ„ค-to-code-base ๐ŸŽ)) ๏ผโŸจ refl โŸฉ
     decode base (โ„ค-to-code-base ๐ŸŽ)                       ๏ผโŸจ I    โŸฉ
     loops (code-base-to-โ„ค (โ„ค-to-code-base ๐ŸŽ))            ๏ผโŸจ II   โŸฉ
     loops ๐ŸŽ                                              ๏ผโŸจ refl โŸฉ
     refl                                                 โˆŽ
      where

\end{code}

       Using the first computation rule for ๐•Šยน-induction

\begin{code}

       I  = happly ๐•Šยน-induction-on-base (โ„ค-to-code-base ๐ŸŽ)
       II = ap loops (Idtofun-retraction code-on-base ๐ŸŽ)

    open import UF.Retracts

    ฮฉ๐•Šยน-is-set : is-set (base ๏ผ base)
    ฮฉ๐•Šยน-is-set = subtypes-of-sets-are-sets' (encode base)
                  (sections-are-lc (encode base)
                   ((decode base) , (decode-encode base)))
                   (transport is-set (code-on-base โปยน) โ„ค-is-set)

  module ๐•Šยน-induction'
          {๐“ฅ : Universe}
          (A : ๐•Šยน โ†’ ๐“ฅ ฬ‡ )
          (a : A base)
          (l : transport A loop a ๏ผ a)
          (fe : funext ๐“คโ‚€ ๐“ค)
          (ua : is-univalent ๐“คโ‚€)
         where

   open ๐•Šยน-induction A a l

   ๐•Šยน-induction-on-loop' : transport (ฮป - โ†’ transport A loop - ๏ผ -)
                            ๐•Šยน-induction-on-base (apd ๐•Šยน-induction loop)
                         ๏ผ l
   ๐•Šยน-induction-on-loop' = ๐•Šยน-induction-on-loop (ฮฉ๐•Šยน-is-set ua fe)

   ๐•Šยน-induction-comp' : (๐•Šยน-induction base , apd ๐•Šยน-induction loop)
                      ๏ผ[ ฮฃ y ๊ž‰ A base , transport A loop y ๏ผ y ] (a , l)
   ๐•Šยน-induction-comp' = ๐•Šยน-induction-comp (ฮฉ๐•Šยน-is-set ua fe)

\end{code}