Martin Escardo
Properties of equivalences depending on function extensionality.
(Not included in UF.Equiv because the module UF.funext defines function
extensionality as the equivalence of happly for suitable parameters.)
\begin{code}
{-# OPTIONS --safe --without-K #-}
module UF.Equiv-FunExt where
open import MLTT.Spartan
open import UF.Base
open import UF.Subsingletons
open import UF.Subsingletons-FunExt
open import UF.Retracts
open import UF.FunExt
open import UF.Equiv
open import UF.EquivalenceExamples
being-vv-equiv-is-prop' : funext ๐ฅ (๐ค โ ๐ฅ)
                        โ funext (๐ค โ ๐ฅ) (๐ค โ ๐ฅ)
                        โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } (f : X โ Y)
                        โ is-prop (is-vv-equiv f)
being-vv-equiv-is-prop' {๐ค} {๐ฅ} fe fe' f = ฮ -is-prop
                                             fe
                                             (ฮป x โ being-singleton-is-prop fe')
being-vv-equiv-is-prop : FunExt
                       โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                         (f : X โ Y)
                       โ is-prop (is-vv-equiv f)
being-vv-equiv-is-prop {๐ค} {๐ฅ} fe =
 being-vv-equiv-is-prop' (fe ๐ฅ (๐ค โ ๐ฅ)) (fe (๐ค โ ๐ฅ) (๐ค โ ๐ฅ))
qinv-post' : {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {A : ๐ฆ ฬ }
          โ naive-funext ๐ฆ ๐ค
          โ naive-funext ๐ฆ ๐ฅ
          โ (f : X โ Y)
          โ qinv f
          โ qinv (ฮป (h : A โ X) โ f โ h)
qinv-post' {๐ค} {๐ฅ} {๐ฆ} {X} {Y} {A} nfe nfe' f (g , ฮท , ฮต) = (g' , ฮท' , ฮต')
 where
  f' : (A โ X) โ (A โ Y)
  f' h = f โ h
  g' : (A โ Y) โ (A โ X)
  g' k = g โ k
  ฮท' : (h : A โ X) โ g' (f' h) ๏ผ h
  ฮท' h = nfe (ฮท โ h)
  ฮต' : (k : A โ Y) โ f' (g' k) ๏ผ k
  ฮต' k = nfe' (ฮต โ k)
qinv-post : (โ ๐ค ๐ฅ โ naive-funext ๐ค ๐ฅ)
          โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {A : ๐ฆ ฬ }
            (f : X โ Y)
          โ qinv f
          โ qinv (ฮป (h : A โ X) โ f โ h)
qinv-post {๐ค} {๐ฅ} {๐ฆ} nfe = qinv-post' (nfe ๐ฆ ๐ค) (nfe ๐ฆ ๐ฅ)
equiv-post : {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {A : ๐ฆ ฬ }
           โ naive-funext ๐ฆ ๐ค
           โ naive-funext ๐ฆ ๐ฅ
           โ (f : X โ Y)
           โ is-equiv f
           โ is-equiv (ฮป (h : A โ X) โ f โ h)
equiv-post nfe nfe' f e = qinvs-are-equivs
                           (f โ_) (qinv-post' nfe nfe' f (equivs-are-qinvs f e))
qinv-pre' : {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {A : ๐ฆ ฬ }
          โ naive-funext ๐ฅ ๐ฆ
          โ naive-funext ๐ค ๐ฆ
          โ (f : X โ Y)
          โ qinv f
          โ qinv (ฮป (h : Y โ A) โ h โ f)
qinv-pre' {๐ค} {๐ฅ} {๐ฆ} {X} {Y} {A} nfe nfe' f (g , ฮท , ฮต) = (g' , ฮท' , ฮต')
 where
  f' : (Y โ A) โ (X โ A)
  f' h = h โ f
  g' : (X โ A) โ (Y โ A)
  g' k = k โ g
  ฮท' : (h : Y โ A) โ g' (f' h) ๏ผ h
  ฮท' h = nfe (ฮป y โ ap h (ฮต y))
  ฮต' : (k : X โ A) โ f' (g' k) ๏ผ k
  ฮต' k = nfe' (ฮป x โ ap k (ฮท x))
qinv-pre : (โ ๐ค ๐ฅ โ naive-funext ๐ค ๐ฅ)
         โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {A : ๐ฆ ฬ } (f : X โ Y)
         โ qinv f
         โ qinv (ฮป (h : Y โ A) โ h โ f)
qinv-pre {๐ค} {๐ฅ} {๐ฆ} nfe = qinv-pre' (nfe ๐ฅ ๐ฆ) (nfe ๐ค ๐ฆ)
retractions-have-at-most-one-section' : {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                                      โ funext ๐ฅ ๐ค
                                      โ funext ๐ฅ ๐ฅ
                                      โ (f : X โ Y)
                                      โ is-section f
                                      โ is-prop (has-section f)
retractions-have-at-most-one-section'
 {๐ค} {๐ฅ} {X} {Y} fe fe' f (g , gf) (h , fh) = singletons-are-props c (h , fh)
 where
  a : qinv f
  a = equivs-are-qinvs f ((h , fh) , g , gf)
  b : is-singleton(fiber (f โ_) id)
  b = qinvs-are-vv-equivs (f โ_) (qinv-post' (dfunext fe) (dfunext fe') f a) id
  r : fiber (f โ_) id โ has-section f
  r (h , p) = (h , happly' (f โ h) id p)
  s : has-section f โ fiber (f โ_) id
  s (h , ฮท) = (h , dfunext fe' ฮท)
  rs : (ฯ : has-section f) โ r (s ฯ) ๏ผ ฯ
  rs (h , ฮท) = ap (ฮป - โ (h , -)) q
   where
    q : happly' (f โ h) id (dfunext fe' ฮท) ๏ผ ฮท
    q = happly-funext fe' (f โ h) id ฮท
  c : is-singleton (has-section f)
  c = retract-of-singleton (r , s , rs) b
sections-have-at-most-one-retraction' : {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                                      โ funext ๐ค ๐ค
                                      โ funext ๐ฅ ๐ค
                                      โ (f : X โ Y)
                                      โ has-section f
                                      โ is-prop (is-section f)
sections-have-at-most-one-retraction'
 {๐ค} {๐ฅ} {X} {Y} fe fe' f (g , fg) (h , hf) =
 singletons-are-props c (h , hf)
 where
  a : qinv f
  a = equivs-are-qinvs f ((g , fg) , (h , hf))
  b : is-singleton(fiber (_โ f) id)
  b = qinvs-are-vv-equivs (_โ f) (qinv-pre' (dfunext fe') (dfunext fe) f a) id
  r : fiber (_โ f) id โ is-section f
  r (h , p) = (h , happly' (h โ f) id p)
  s : is-section f โ fiber (_โ f) id
  s (h , ฮท) = (h , dfunext fe ฮท)
  rs : (ฯ : is-section f) โ r (s ฯ) ๏ผ ฯ
  rs (h , ฮท) = ap (ฮป - โ (h , -)) q
   where
    q : happly' (h โ f) id (dfunext fe ฮท) ๏ผ ฮท
    q = happly-funext fe (h โ f) id ฮท
  c : is-singleton (is-section f)
  c = retract-of-singleton (r , s , rs) b
retractions-have-at-most-one-section : FunExt
                                     โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                                       (f : X โ Y)
                                     โ is-section f
                                     โ is-prop (has-section f)
retractions-have-at-most-one-section {๐ค} {๐ฅ} fe =
 retractions-have-at-most-one-section' (fe ๐ฅ ๐ค) (fe ๐ฅ ๐ฅ)
sections-have-at-most-one-retraction : FunExt
                                     โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                                       (f : X โ Y)
                                     โ has-section f
                                     โ is-prop (is-section f)
sections-have-at-most-one-retraction {๐ค} {๐ฅ} fe =
 sections-have-at-most-one-retraction' (fe ๐ค ๐ค) (fe ๐ฅ ๐ค)
being-equiv-is-prop' : {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                     โ funext ๐ฅ ๐ค
                     โ funext ๐ฅ ๐ฅ
                     โ funext ๐ค ๐ค
                     โ funext ๐ฅ ๐ค
                     โ (f : X โ Y)
                     โ is-prop (is-equiv f)
being-equiv-is-prop' fe fe' fe'' fe''' f =
 ร-prop-criterion
  (retractions-have-at-most-one-section' fe fe' f ,
   sections-have-at-most-one-retraction' fe'' fe''' f)
being-equiv-is-prop : FunExt
                    โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                      (f : X โ Y)
                    โ is-prop (is-equiv f)
being-equiv-is-prop {๐ค} {๐ฅ} fe f =
 being-equiv-is-prop' (fe ๐ฅ ๐ค) (fe ๐ฅ ๐ฅ) (fe ๐ค ๐ค) (fe ๐ฅ ๐ค) f
being-equiv-is-prop'' : {X Y : ๐ค ฬ }
                      โ funext ๐ค ๐ค
                      โ (f : X โ Y)
                      โ is-prop (is-equiv f)
being-equiv-is-prop'' fe = being-equiv-is-prop' fe fe fe fe
โ-assoc' : funext ๐ฃ ๐ค
         โ funext ๐ฃ ๐ฃ
         โ funext ๐ค ๐ค
         โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {Z : ๐ฆ ฬ } {T : ๐ฃ ฬ }
           (ฮฑ : X โ Y) (ฮฒ : Y โ Z) (ฮณ : Z โ T)
         โ ฮฑ โ (ฮฒ โ ฮณ) ๏ผ (ฮฑ โ ฮฒ) โ ฮณ
โ-assoc' feโ fโ fโ (f , a) (g , b) (h , c) = to-ฮฃ-๏ผ (p , q)
 where
  p : (h โ g) โ f ๏ผ h โ (g โ f)
  p = refl
  d e : is-equiv (h โ g โ f)
  d = โ-is-equiv a (โ-is-equiv b c)
  e = โ-is-equiv (โ-is-equiv a b) c
  q : transport is-equiv p d ๏ผ e
  q = being-equiv-is-prop' feโ fโ fโ feโ (h โ g โ f) _ _
โ-assoc : FunExt
        โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {Z : ๐ฆ ฬ } {T : ๐ฃ ฬ }
          (ฮฑ : X โ Y) (ฮฒ : Y โ Z) (ฮณ : Z โ T)
        โ ฮฑ โ (ฮฒ โ ฮณ) ๏ผ (ฮฑ โ ฮฒ) โ ฮณ
โ-assoc fe = โ-assoc' (fe _ _) (fe _ _) (fe _ _)
to-โ-๏ผ : {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
        โ Fun-Ext
        โ {f g : X โ Y} {i : is-equiv f} {j : is-equiv g}
        โ f โผ g
        โ (f , i) ๏ผ[ X โ Y ] (g , j)
to-โ-๏ผ fe h = to-subtype-๏ผ
               (being-equiv-is-prop' fe fe fe fe)
                        (dfunext fe h)
\end{code}
The above proof can be condensed to one line in the style of the
following two proofs, which exploit the fact that the identity map is
a neutral element for ordinary function composition, definitionally:
\begin{code}
โ-refl-left' : funext ๐ฅ ๐ค
             โ funext ๐ฅ ๐ฅ
             โ funext ๐ค ๐ค
             โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } (ฮฑ : X โ Y)
             โ โ-refl X โ ฮฑ ๏ผ ฮฑ
โ-refl-left' feโ feโ feโ ฮฑ =
 to-ฮฃ-๏ผ' (being-equiv-is-prop' feโ feโ feโ feโ _ _ _)
โ-refl-right' : funext ๐ฅ ๐ค
              โ funext ๐ฅ ๐ฅ
              โ funext ๐ค ๐ค
              โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                (ฮฑ : X โ Y)
              โ ฮฑ โ โ-refl Y ๏ผ ฮฑ
โ-refl-right' feโ feโ feโ  ฮฑ =
 to-ฮฃ-๏ผ' (being-equiv-is-prop' feโ feโ feโ feโ _ _ _)
โ-sym-involutive' : funext ๐ฅ ๐ค
                  โ funext ๐ฅ ๐ฅ
                  โ funext ๐ค ๐ค
                  โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                    (ฮฑ : X โ Y)
                  โ โ-sym (โ-sym ฮฑ) ๏ผ ฮฑ
โ-sym-involutive' feโ feโ feโ (f , a) =
 to-ฮฃ-๏ผ
  (inversion-involutive f a ,
   being-equiv-is-prop' feโ feโ feโ feโ f _ a)
โ-sym-involutive : FunExt
                 โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                   (ฮฑ : X โ Y)
                 โ โ-sym (โ-sym ฮฑ) ๏ผ ฮฑ
โ-sym-involutive {๐ค} {๐ฅ} fe = โ-sym-involutive' (fe ๐ฅ ๐ค) (fe ๐ฅ ๐ฅ) (fe ๐ค ๐ค)
โ-Sym' : funext ๐ฅ ๐ค
       โ funext ๐ฅ ๐ฅ
       โ funext ๐ค ๐ค
       โ funext ๐ค ๐ฅ
       โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
       โ (X โ Y) โ (Y โ X)
โ-Sym' feโ feโ feโ feโ = qinveq โ-sym
                          (โ-sym ,
                           โ-sym-involutive' feโ feโ feโ ,
                           โ-sym-involutive' feโ feโ feโ)
โ-Sym'' : funext ๐ค ๐ค
        โ {X Y : ๐ค ฬ }
        โ (X โ Y) โ (Y โ X)
โ-Sym'' fe = โ-Sym' fe fe fe fe
โ-Sym : FunExt
      โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
      โ (X โ Y) โ (Y โ X)
โ-Sym {๐ค} {๐ฅ} fe = โ-Sym' (fe ๐ฅ ๐ค) (fe ๐ฅ ๐ฅ) (fe ๐ค ๐ค) (fe ๐ค ๐ฅ)
โ-refl-left : FunExt
            โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
              (ฮฑ : X โ Y)
            โ โ-refl X โ ฮฑ ๏ผ ฮฑ
โ-refl-left fe = โ-refl-left' (fe _ _) (fe _ _) (fe _ _)
โ-refl-right : FunExt
             โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
               (ฮฑ : X โ Y)
             โ ฮฑ โ โ-refl Y ๏ผ ฮฑ
โ-refl-right fe = โ-refl-right' (fe _ _) (fe _ _) (fe _ _)
โ-sym-left-inverse' : funext ๐ฅ ๐ฅ
                    โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                      (ฮฑ : X โ Y)
                    โ โ-sym ฮฑ โ ฮฑ ๏ผ โ-refl Y
โ-sym-left-inverse' fe {X} {Y} (f , e) = ฮณ
 where
  ฮฑ = (f , e)
  p : f โ inverse f e ๏ผ id
  p = dfunext fe (inverses-are-sections f e)
  ฮณ : โ-sym ฮฑ โ ฮฑ ๏ผ โ-refl Y
  ฮณ = to-ฮฃ-๏ผ (p , being-equiv-is-prop' fe fe fe fe _ _ _)
โ-sym-left-inverse : FunExt
                   โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
                     (ฮฑ : X โ Y)
                   โ โ-sym ฮฑ โ ฮฑ ๏ผ โ-refl Y
โ-sym-left-inverse fe = โ-sym-left-inverse' (fe _ _)
โ-sym-right-inverse' : funext ๐ค ๐ค
                     โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } (ฮฑ : X โ Y)
                     โ ฮฑ โ โ-sym ฮฑ ๏ผ โ-refl X
โ-sym-right-inverse' fe {X} (f , e) = ฮณ
 where
  ฮฑ = (f , e)
  p : inverse f e โ f ๏ผ id
  p = dfunext fe (inverses-are-retractions f e)
  ฮณ : ฮฑ โ โ-sym ฮฑ ๏ผ โ-refl X
  ฮณ = to-ฮฃ-๏ผ (p , being-equiv-is-prop' fe fe fe fe _ _ _)
โ-sym-right-inverse : FunExt
                    โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } (ฮฑ : X โ Y)
                    โ ฮฑ โ โ-sym ฮฑ ๏ผ โ-refl X
โ-sym-right-inverse fe = โ-sym-right-inverse' (fe _ _)
โ-cong-left' : {๐ค ๐ฅ ๐ฆ : Universe}
             โ funext ๐ฆ ๐ค
             โ funext ๐ฆ ๐ฆ
             โ funext ๐ค ๐ค
             โ funext ๐ฆ ๐ฅ
             โ funext ๐ฅ ๐ฅ
             โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {Z : ๐ฆ ฬ }
             โ X โ Y
             โ (X โ Z) โ (Y โ Z)
โ-cong-left' feโ feโ feโ feโ feโ {X} {Y} {Z} ฮฑ = ฮณ
 where
  p = ฮป ฮณ โ ฮฑ โ (โ-sym ฮฑ โ ฮณ) ๏ผโจ โ-assoc' feโ feโ feโ ฮฑ (โ-sym ฮฑ) ฮณ โฉ
            (ฮฑ โ โ-sym ฮฑ) โ ฮณ ๏ผโจ ap (_โ ฮณ) (โ-sym-right-inverse' feโ ฮฑ) โฉ
            โ-refl _ โ ฮณ      ๏ผโจ โ-refl-left' feโ feโ feโ _ โฉ
            ฮณ                 โ
  q = ฮป ฮฒ โ โ-sym ฮฑ โ (ฮฑ โ ฮฒ) ๏ผโจ โ-assoc' feโ feโ feโ (โ-sym ฮฑ) ฮฑ ฮฒ โฉ
            (โ-sym ฮฑ โ ฮฑ) โ ฮฒ ๏ผโจ ap (_โ ฮฒ) (โ-sym-left-inverse' feโ ฮฑ) โฉ
            โ-refl _ โ ฮฒ      ๏ผโจ โ-refl-left' feโ feโ feโ _ โฉ
            ฮฒ                 โ
  ฮณ : (X โ Z) โ (Y โ Z)
  ฮณ = qinveq ((โ-sym ฮฑ) โ_) ((ฮฑ โ_), p , q)
โ-cong-left'' : funext ๐ค ๐ค
              โ {X Y Z : ๐ค ฬ }
              โ X โ Y
              โ (X โ Z) โ (Y โ Z)
โ-cong-left'' fe = โ-cong-left' fe fe fe fe fe
โ-cong-left : FunExt
            โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {Z : ๐ฆ ฬ }
            โ X โ Y
            โ (X โ Z) โ (Y โ Z)
โ-cong-left fe = โ-cong-left' (fe _ _) (fe _ _) (fe _ _) (fe _ _) (fe _ _)
โ-cong-right : FunExt
             โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {A : ๐ฆ ฬ }
             โ X โ Y
             โ (A โ X) โ (A โ Y)
โ-cong-right fe {X} {Y} {A} ฮฑ =
 (A โ X) โโจ โ-Sym fe โฉ
 (X โ A) โโจ โ-cong-left fe ฮฑ โฉ
 (Y โ A) โโจ โ-Sym fe โฉ
 (A โ Y) โ 
โ-cong : FunExt
       โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ } {A : ๐ฆ ฬ } {B : ๐ฃ ฬ }
       โ X โ A
       โ Y โ B
       โ (X โ Y) โ (A โ B)
โ-cong fe {X} {Y} {A} {B} ฮฑ ฮฒ =
 (X โ Y) โโจ โ-cong-left  fe ฮฑ โฉ
 (A โ Y) โโจ โ-cong-right fe ฮฒ โฉ
 (A โ B) โ 
โ-is-prop : FunExt
          โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
          โ is-prop Y
          โ is-prop (X โ Y)
โ-is-prop {๐ค} {๐ฅ} fe i (f , e) (f' , e') =
 to-subtype-๏ผ
  (being-equiv-is-prop fe)
  (dfunext (fe ๐ค ๐ฅ) (ฮป x โ i (f x) (f' x)))
โ-is-prop' : FunExt
           โ {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
           โ is-prop X
           โ is-prop (X โ Y)
โ-is-prop' fe i = equiv-to-prop (โ-Sym fe) (โ-is-prop fe i)
\end{code}
Propositional and functional extesionality give univalence for
propositions. Notice that P is assumed to be a proposition, but X
ranges over arbitrary types:
\begin{code}
propext-funext-give-prop-ua : propext ๐ค
                            โ funext ๐ค ๐ค
                            โ (X : ๐ค ฬ ) (P : ๐ค ฬ )
                            โ is-prop P
                            โ is-equiv (idtoeq X P)
propext-funext-give-prop-ua {๐ค} pe fe X P i = (eqtoid , ฮท) , (eqtoid , ฮต)
 where
  l : X โ P โ is-prop X
  l e = equiv-to-prop e i
  eqtoid : X โ P โ X ๏ผ P
  eqtoid (f , (r , rf) , h) = pe (l (f , (r , rf) , h)) i f r
  m : is-prop (X โ P)
  m (f , e) (f' , e') = to-ฮฃ-๏ผ (dfunext fe (ฮป x โ i (f x) (f' x)) ,
                                being-equiv-is-prop'' fe f' _ e')
  ฮท : (e : X โ P) โ idtoeq X P (eqtoid e) ๏ผ e
  ฮท e = m (idtoeq X P (eqtoid e)) e
  ฮต : (q : X ๏ผ P) โ eqtoid (idtoeq X P q) ๏ผ q
  ฮต q = identifications-with-props-are-props pe fe
         P
         i
         X
         (eqtoid (idtoeq X P q))
         q
prop-univalent-โ : propext ๐ค
                 โ funext ๐ค ๐ค
                 โ (X P : ๐ค ฬ )
                 โ is-prop P
                 โ (X ๏ผ P) โ (X โ P)
prop-univalent-โ pe fe X P i = idtoeq X P ,
                               propext-funext-give-prop-ua pe fe X P i
prop-univalent-โ' : propext ๐ค
                  โ funext ๐ค ๐ค
                  โ (X P : ๐ค ฬ )
                  โ is-prop P
                  โ (P ๏ผ X) โ (P โ X)
prop-univalent-โ' pe fe X P i = (P ๏ผ X) โโจ ๏ผ-flip โฉ
                                (X ๏ผ P) โโจ prop-univalent-โ pe fe X P i โฉ
                                (X โ P)  โโจ โ-Sym'' fe โฉ
                                (P โ X)  โ 
\end{code}
Added 24th Feb 2023.
\begin{code}
prop-โ-โ-โ : Fun-Ext
           โ {P : ๐ค ฬ } {Q : ๐ฅ ฬ }
           โ is-prop P
           โ is-prop Q
           โ (P โ Q) โ (P โ Q)
prop-โ-โ-โ fe i j = qinveq (ฮป f โ โ f โ ,  โ f โโปยน)
                     ((ฮป (g , h) โ qinveq g
                                    (h ,
                                    (ฮป p โ i (h (g p)) p) ,
                                    (ฮป q โ j (g (h q)) q))) ,
                      (ฮป f โ to-subtype-๏ผ
                              (being-equiv-is-prop (ฮป _ _ โ fe))
                               refl) ,
                      (ฮป _ โ refl))
prop-๏ผ-โ-โ : Prop-Ext
            โ Fun-Ext
            โ {P Q : ๐ค ฬ }
            โ is-prop P
            โ is-prop Q
            โ (P ๏ผ Q) โ (P โ Q)
prop-๏ผ-โ-โ pe fe i j = prop-univalent-โ pe fe _ _ j
                      โ prop-โ-โ-โ fe i j
\end{code}
Added 3rd November 2023.
\begin{code}
open import UF.Subsingletons-Properties
open import UF.Sets
open import UF.Sets-Properties
โ-is-set : {X : ๐ค ฬ } {Y : ๐ฅ ฬ }
         โ Fun-Ext
         โ is-set X
         โ is-set (X โ Y)
โ-is-set {๐ค} {๐ฅ} {X} {Y} fe X-is-set {๐} {๐} =
 ฮฃ-is-set
  (ฮ -is-set fe (ฮป _ โ equiv-to-set (โ-sym ๐) X-is-set))
  (ฮป _ โ props-are-sets (being-equiv-is-prop (ฮป _ _ โ fe) _))
\end{code}
Added 25 March 2025 by Tom de Jong.
If the domain or codomain of f is a set, then being invertible is a property of
the map f.
In particular in such cases the type expressing that f is an equivalence is
equivalent to the type expressing that f is invertible.
\begin{code}
being-qinv-is-prop : Fun-Ext
                   โ {X : ๐ค ฬ  } {Y : ๐ฅ ฬ  }
                   โ (f : X โ Y)
                   โ is-set X
                   โ is-prop (qinv f)
being-qinv-is-prop fe {X} {Y} f X-is-set = prop-criterion II
 where
  module _ (q : qinv f)
   where
    I : Y โ X
    I = โ-sym (f , qinvs-are-equivs f q)
    II : is-prop (qinv f)
    II (g , g-ret , g-sec) (h , h-ret , h-sec) =
     to-subtype-๏ผ (ฮป k โ ร-is-prop
                           (ฮ -is-prop fe (ฮป x โ X-is-set))
                           (ฮ -is-prop fe (ฮป y โ equiv-to-set I X-is-set)))
                   (dfunext fe (ฮป y โ g y         ๏ผโจ ap g ((h-sec y) โปยน) โฉ
                                      g (f (h y)) ๏ผโจ g-ret (h y) โฉ
                                      h y         โ))
being-qinv-is-prop' : Fun-Ext
                    โ {X Y : ๐ค ฬ  }
                    โ (f : X โ Y)
                    โ is-set Y
                    โ is-prop (qinv f)
being-qinv-is-prop' fe {X} {Y} f Y-is-set = prop-criterion II
 where
  module _ (q : qinv f)
   where
    I : X โ Y
    I = f , qinvs-are-equivs f q
    II : is-prop (qinv f)
    II = being-qinv-is-prop fe f (equiv-to-set I Y-is-set)
is-equiv-โ-qinv : Fun-Ext
                โ {X Y : ๐ค ฬ  }
                โ (f : X โ Y)
                โ is-set X
                โ is-equiv f โ qinv f
is-equiv-โ-qinv fe f X-is-set =
 logically-equivalent-props-are-equivalent
  (being-equiv-is-prop (ฮป _ _ โ fe) f)
  (being-qinv-is-prop fe f X-is-set)
  (equivs-are-qinvs f)
  (qinvs-are-equivs f)
is-equiv-โ-qinv' : Fun-Ext
                 โ {X Y : ๐ค ฬ  }
                 โ (f : X โ Y)
                 โ is-set Y
                 โ is-equiv f โ qinv f
is-equiv-โ-qinv' fe f Y-is-set =
 logically-equivalent-props-are-equivalent
  (being-equiv-is-prop (ฮป _ _ โ fe) f)
  (being-qinv-is-prop' fe f Y-is-set)
  (equivs-are-qinvs f)
  (qinvs-are-equivs f)
\end{code}
Added by Martin Escardo 22nd June 2025.
The function โ-sym : X โ Y โ Y โ X is an equivalence.
\begin{code}
module _
         {X : ๐ค ฬ }
         {Y : ๐ฅ ฬ }
       where
 module _
          (feuu : funext ๐ค ๐ค)
          (feuv : funext ๐ค ๐ฅ)
          (fevv : funext ๐ฅ ๐ฅ)
          (fevu : funext ๐ฅ ๐ค)
      where
  โ-sym-is-equiv' : is-equiv (โ-sym {๐ค} {๐ฅ} {X} {Y})
  โ-sym-is-equiv' =
   qinvs-are-equivs (โ-sym {๐ค} {๐ฅ} {X} {Y})
    (โ-sym {๐ฅ} {๐ค} {Y} {X} ,
    (ฮป _ โ to-subtype-๏ผ (being-equiv-is-prop' fevu fevv feuu fevu) refl) ,
    (ฮป _ โ to-subtype-๏ผ (being-equiv-is-prop' feuv feuu fevv feuv) refl))
  โ-flip' : (X โ Y) โ (Y โ X)
  โ-flip' = โ-sym , โ-sym-is-equiv'
 module _ (fe : Fun-Ext) where
  โ-sym-is-equiv : is-equiv (โ-sym {๐ค} {๐ฅ} {X} {Y})
  โ-sym-is-equiv = โ-sym-is-equiv' fe fe fe fe
  โ-flip : (X โ Y) โ (Y โ X)
  โ-flip = โ-sym , โ-sym-is-equiv
\end{code}
Added 8th September 2025 by Martin Escardo.
\begin{code}
equivalences-with-props-are-props' : funext ๐ค ๐ฅ
                                   โ funext ๐ค ๐ค
                                   โ funext ๐ฅ ๐ฅ
                                   โ funext ๐ค ๐ฅ
                                   โ funext ๐ฅ ๐ค
                                   โ (P : ๐ค ฬ )
                                  โ is-prop P
                                  โ (X : ๐ฅ ฬ ) โ is-prop (X โ P)
equivalences-with-props-are-props' {๐ค} {๐ฅ} feโ feโ feโ feโ feโ P i X (f , e) (f' , e') =
 to-subtype-๏ผ
  (ฮป ฯ โ being-equiv-is-prop' feโ feโ feโ feโ ฯ)
  (dfunext feโ (ฮป x โ i (f x) (f' x)))
equivalences-with-props-are-props : Fun-Ext
                                  โ (P : ๐ค ฬ )
                                  โ is-prop P
                                  โ (X : ๐ฅ ฬ ) โ is-prop (X โ P)
equivalences-with-props-are-props fe = equivalences-with-props-are-props' fe fe fe fe fe
\end{code}
End of addition.