Martin Escardo 1st May 2020.
This is ported from the Midlands Graduate School 2019 lecture notes
https://www.cs.bham.ac.uk/~mhe/HoTT-UF.in-Agda-Lecture-Notes/HoTT-UF-Agda.html
https://github.com/martinescardo/HoTT-UF.Agda-Lecture-Notes
\begin{code}
{-# OPTIONS --safe --without-K #-}
module MGS.Powerset where
open import MGS.More-FunExt-Consequences public
propext : โ ๐ค โ ๐ค โบ ฬ
propext ๐ค = {P Q : ๐ค ฬ } โ is-prop P โ is-prop Q โ (P โ Q) โ (Q โ P) โ P ๏ผ Q
global-propext : ๐คฯ
global-propext = โ {๐ค} โ propext ๐ค
univalence-gives-propext : is-univalent ๐ค โ propext ๐ค
univalence-gives-propext ua {P} {Q} i j f g = EqโId ua P Q ฮณ
where
ฮณ : P โ Q
ฮณ = logically-equivalent-subsingletons-are-equivalent P Q i j (f , g)
Id-from-subsingleton : propext ๐ค โ dfunext ๐ค ๐ค
โ (P : ๐ค ฬ )
โ is-subsingleton P
โ (X : ๐ค ฬ ) โ is-subsingleton (P ๏ผ X)
Id-from-subsingleton {๐ค} pe fe P i = Hedberg P (ฮป X โ h X , k X)
where
module _ (X : ๐ค ฬ ) where
f : P ๏ผ X โ is-subsingleton X ร (P โ X)
f p = transport is-subsingleton p i , Idโfun p , (Idโfun (p โปยน))
g : is-subsingleton X ร (P โ X) โ P ๏ผ X
g (l , ฯ , ฯ) = pe i l ฯ ฯ
h : P ๏ผ X โ P ๏ผ X
h = g โ f
j : is-subsingleton (is-subsingleton X ร (P โ X))
j = ร-is-subsingleton'
((ฮป (_ : P โ X) โ being-subsingleton-is-subsingleton fe) ,
(ฮป (l : is-subsingleton X) โ ร-is-subsingleton
(ฮ -is-subsingleton fe (ฮป p โ l))
(ฮ -is-subsingleton fe (ฮป x โ i))))
k : wconstant h
k p q = ap g (j (f p) (f q))
subsingleton-univalence : propext ๐ค โ dfunext ๐ค ๐ค
โ (P : ๐ค ฬ )
โ is-subsingleton P
โ (X : ๐ค ฬ ) โ is-equiv (IdโEq P X)
subsingleton-univalence pe fe P i X = ฮณ
where
l : P โ X โ is-subsingleton X
l e = equiv-to-subsingleton (โ-sym e) i
eqtoid : P โ X โ P ๏ผ X
eqtoid e = pe i (equiv-to-subsingleton (โ-sym e) i)
โ e โ โ โ-sym e โ
m : is-subsingleton (P โ X)
m (f , k) (f' , k') = to-subtype-๏ผ
(being-equiv-is-subsingleton fe fe)
(fe (ฮป x โ j (f x) (f' x)))
where
j : is-subsingleton X
j = equiv-to-subsingleton (โ-sym (f , k)) i
ฮต : (e : P โ X) โ IdโEq P X (eqtoid e) ๏ผ e
ฮต e = m (IdโEq P X (eqtoid e)) e
ฮท : (q : P ๏ผ X) โ eqtoid (IdโEq P X q) ๏ผ q
ฮท q = Id-from-subsingleton pe fe P i X (eqtoid (IdโEq P X q)) q
ฮณ : is-equiv (IdโEq P X)
ฮณ = invertibles-are-equivs (IdโEq P X) (eqtoid , ฮท , ฮต)
subsingleton-univalence-โ : propext ๐ค โ dfunext ๐ค ๐ค
โ (X P : ๐ค ฬ ) โ is-subsingleton P โ (P ๏ผ X) โ (P โ X)
subsingleton-univalence-โ pe fe X P i = IdโEq P X ,
subsingleton-univalence pe fe P i X
ฮฉ : (๐ค : Universe) โ ๐ค โบ ฬ
ฮฉ ๐ค = ฮฃ P ๊ ๐ค ฬ , is-subsingleton P
_holds : ฮฉ ๐ค โ ๐ค ฬ
_holds (P , i) = P
holds-is-subsingleton : (p : ฮฉ ๐ค) โ is-subsingleton (p holds)
holds-is-subsingleton (P , i) = i
ฮฉ-ext : dfunext ๐ค ๐ค โ propext ๐ค โ {p q : ฮฉ ๐ค}
โ (p holds โ q holds) โ (q holds โ p holds) โ p ๏ผ q
ฮฉ-ext {๐ค} fe pe {p} {q} f g = to-subtype-๏ผ
(ฮป _ โ being-subsingleton-is-subsingleton fe)
(pe (holds-is-subsingleton p) (holds-is-subsingleton q) f g)
ฮฉ-is-set : dfunext ๐ค ๐ค โ propext ๐ค โ is-set (ฮฉ ๐ค)
ฮฉ-is-set {๐ค} fe pe = types-with-wconstant-๏ผ-endomaps-are-sets (ฮฉ ๐ค) c
where
A : (p q : ฮฉ ๐ค) โ ๐ค ฬ
A p q = (p holds โ q holds) ร (q holds โ p holds)
i : (p q : ฮฉ ๐ค) โ is-subsingleton (A p q)
i p q = ฮฃ-is-subsingleton
(ฮ -is-subsingleton fe
(ฮป _ โ holds-is-subsingleton q))
(ฮป _ โ ฮ -is-subsingleton fe (ฮป _ โ holds-is-subsingleton p))
g : (p q : ฮฉ ๐ค) โ p ๏ผ q โ A p q
g p q e = (u , v)
where
a : p holds ๏ผ q holds
a = ap _holds e
u : p holds โ q holds
u = Idโfun a
v : q holds โ p holds
v = Idโfun (a โปยน)
h : (p q : ฮฉ ๐ค) โ A p q โ p ๏ผ q
h p q (u , v) = ฮฉ-ext fe pe u v
f : (p q : ฮฉ ๐ค) โ p ๏ผ q โ p ๏ผ q
f p q e = h p q (g p q e)
k : (p q : ฮฉ ๐ค) (d e : p ๏ผ q) โ f p q d ๏ผ f p q e
k p q d e = ap (h p q) (i p q (g p q d) (g p q e))
c : (p q : ฮฉ ๐ค) โ ฮฃ f ๊ (p ๏ผ q โ p ๏ผ q), wconstant f
c p q = (f p q , k p q)
powersets-are-sets : hfunext ๐ค (๐ฅ โบ) โ dfunext ๐ฅ ๐ฅ โ propext ๐ฅ
โ {X : ๐ค ฬ } โ is-set (X โ ฮฉ ๐ฅ)
powersets-are-sets fe fe' pe = ฮ -is-set fe (ฮป x โ ฮฉ-is-set fe' pe)
๐ : ๐ค ฬ โ ๐ค โบ ฬ
๐ {๐ค} X = X โ ฮฉ ๐ค
powersets-are-sets' : Univalence
โ {X : ๐ค ฬ } โ is-set (๐ X)
powersets-are-sets' {๐ค} ua = powersets-are-sets
(univalence-gives-hfunext' (ua ๐ค) (ua (๐ค โบ)))
(univalence-gives-dfunext (ua ๐ค))
(univalence-gives-propext (ua ๐ค))
_โ_ : {X : ๐ค ฬ } โ X โ ๐ X โ ๐ค ฬ
x โ A = A x holds
_โ_ : {X : ๐ค ฬ } โ ๐ X โ ๐ X โ ๐ค ฬ
A โ B = โ x โ x โ A โ x โ B
โ-is-subsingleton : {X : ๐ค ฬ } (A : ๐ X) (x : X) โ is-subsingleton (x โ A)
โ-is-subsingleton A x = holds-is-subsingleton (A x)
โ-is-subsingleton : dfunext ๐ค ๐ค
โ {X : ๐ค ฬ } (A B : ๐ X) โ is-subsingleton (A โ B)
โ-is-subsingleton fe A B = ฮ -is-subsingleton fe
(ฮป x โ ฮ -is-subsingleton fe
(ฮป _ โ โ-is-subsingleton B x))
โ-refl : {X : ๐ค ฬ } (A : ๐ X) โ A โ A
โ-refl A x = ๐๐ (x โ A)
โ-refl-consequence : {X : ๐ค ฬ } (A B : ๐ X)
โ A ๏ผ B โ (A โ B) ร (B โ A)
โ-refl-consequence {X} A A (refl A) = โ-refl A , โ-refl A
subset-extensionality : propext ๐ค โ dfunext ๐ค ๐ค โ dfunext ๐ค (๐ค โบ)
โ {X : ๐ค ฬ } {A B : ๐ X}
โ A โ B โ B โ A โ A ๏ผ B
subset-extensionality pe fe fe' {X} {A} {B} h k = fe' ฯ
where
ฯ : (x : X) โ A x ๏ผ B x
ฯ x = to-subtype-๏ผ
(ฮป _ โ being-subsingleton-is-subsingleton fe)
(pe (holds-is-subsingleton (A x)) (holds-is-subsingleton (B x))
(h x) (k x))
subset-extensionality' : Univalence
โ {X : ๐ค ฬ } {A B : ๐ X}
โ A โ B โ B โ A โ A ๏ผ B
subset-extensionality' {๐ค} ua = subset-extensionality
(univalence-gives-propext (ua ๐ค))
(univalence-gives-dfunext (ua ๐ค))
(univalence-gives-dfunext' (ua ๐ค) (ua (๐ค โบ)))
infix 40 _โ_
\end{code}