Martin Escardo, July 2023
A type of numbers used to measure lengths of paths in trees in W-types
(see the module W.Paths).
For an exposition of what is done here, see the post 7/6 of this thread:
https://mathstodon.xyz/@MartinEscardo/110753930251021051
\begin{code}
{-# OPTIONS --safe --without-K #-}
open import MLTT.Spartan hiding (_^_)
module W.Numbers where
open import Fin.Type hiding (suc)
open import NotionsOfDecidability.Decidable
open import UF.Base
open import UF.DiscreteAndSeparated
open import UF.Embeddings
open import UF.Equiv
open import UF.EquivalenceExamples
open import UF.ClassicalLogic
open import UF.FunExt
open import UF.PropIndexedPiSigma
open import UF.PropTrunc
open import UF.Retracts
open import UF.Sets
open import UF.SubtypeClassifier
open import UF.SubtypeClassifier-Properties
open import UF.Subsingletons
open import UF.Subsingletons-FunExt
open import W.Properties
open import W.Type
\end{code}
We work with a fixed universe π₯.
\begin{code}
module _ {π₯ : Universe} where
π : π₯ βΊ Μ
π = W (Ξ© π₯) _holds
positive : π β Ξ© π₯
positive = W-root
is-positive : π β π₯ Μ
is-positive n = positive n holds
\end{code}
The predecessor function is defined on positive numbers.
\begin{code}
Pred : (n : π) β is-positive n β π
Pred = W-forest
\end{code}
For every proposition p, there is a p-ary successor function. Notice
that we can regard an element of π ^ p as a p-indexed tuple of
elements of π, or, equivalently, as a partial element of π.
\begin{code}
_^_ : π€ Μ β Ξ© π₯ β π₯ β π€ Μ
X ^ p = p holds β X
Suc : (p : Ξ© π₯) β π ^ p β π
Suc = ssup
Suc-positivity : (p : Ξ© π₯) (ns : π ^ p) β positive (Suc p ns) οΌ p
Suc-positivity = W-ssup-root
Pred-Suc : (p : Ξ© π₯) (ns : π ^ p) β Pred (Suc p ns) οΌ ns
Pred-Suc = W-ssup-forest
Suc-Pred : (n : π) β Suc (positive n) (Pred n) οΌ n
Suc-Pred = W-Ξ·
π-induction : (P : π β π¦ Μ )
β ((p : Ξ© π₯) (ns : π ^ p)
β ((h : p holds) β P (ns h))
β P (Suc p ns))
β (n : π) β P n
π-induction = W-induction
π-induction' : (P : π β π¦ Μ )
β ((n : π)
β ((p : is-positive n) β P (Pred n p))
β P n)
β (n : π) β P n
π-induction' = W-induction'
\end{code}
A criterion for equality on π.
\begin{code}
to-π-οΌ : {p q : Ξ© π₯} {ms : π ^ p} {ns : π ^ q}
β (Ξ£ e κ p οΌ q , (ms οΌ ns β transport _holds e))
β Suc p ms οΌ Suc q ns
to-π-οΌ = to-W-οΌ (Ξ© π₯) _holds
from-π-οΌ : {p q : Ξ© π₯} {ms : π ^ p} {ns : π ^ q}
β Suc p ms οΌ Suc q ns
β (Ξ£ e κ p οΌ q , (ms οΌ ns β transport _holds e))
from-π-οΌ = from-W-οΌ (Ξ© π₯) _holds
\end{code}
The β₯-ary successor function amounts to the number zero, and the β€-ary
successor function amounts to the ordinary successor function.
\begin{code}
private
Zero' : (π β π) β π
Zero' = Suc β₯
Succ' : (π β π) β π
Succ' = Suc β€
Zero : π
Zero = Zero' unique-from-π
Succ : π β π
Succ n = Succ' (Ξ» _ β n)
being-positive-is-prop : (n : π) β is-prop (is-positive n)
being-positive-is-prop n = holds-is-prop (positive n)
Succ-is-positive : (n : π) β is-positive (Succ n)
Succ-is-positive n = β€-holds
Zero-is-not-positive : Β¬ is-positive Zero
Zero-is-not-positive = β₯-doesnt-hold
Succ-is-not-Zero : (n : π) β Succ n β Zero
Succ-is-not-Zero n e = Zero-is-not-positive
(transport
is-positive
e
(Succ-is-positive n))
Zero-is-not-Succ : (n : π) β Zero β Succ n
Zero-is-not-Succ n = β -sym (Succ-is-not-Zero n)
\end{code}
The type of positive numbers.
\begin{code}
πβΊ : π₯ βΊ Μ
πβΊ = Ξ£ n κ π , is-positive n
forget-positivity : πβΊ β π
forget-positivity = prβ
forget-positivity-is-embedding : is-embedding forget-positivity
forget-positivity-is-embedding = prβ-is-embedding being-positive-is-prop
PredβΊ : πβΊ β π
PredβΊ = uncurry Pred
SuccβΊ : π β πβΊ
SuccβΊ n = Succ n , Succ-is-positive n
PredβΊ-SuccβΊ : (n : π) β PredβΊ (SuccβΊ n) οΌ n
PredβΊ-SuccβΊ n = refl
Succ-lc : left-cancellable Succ
Succ-lc {m} {n} e = II
where
have-e : Succ m οΌ Succ n
have-e = e
I : SuccβΊ m οΌ SuccβΊ n
I = embeddings-are-lc forget-positivity forget-positivity-is-embedding e
II : m οΌ n
II = ap PredβΊ I
\end{code}
The type of natural numbers is embedded into our type of numbers.
\begin{code}
β-to-π : β β π
β-to-π zero = Zero
β-to-π (succ n) = Succ (β-to-π n)
β-to-π-lc : left-cancellable β-to-π
β-to-π-lc {zero} {zero} e = refl
β-to-π-lc {zero} {succ n} e = π-elim (Zero-is-not-Succ (β-to-π n) e)
β-to-π-lc {succ m} {zero} e = π-elim (Succ-is-not-Zero (β-to-π m) e)
β-to-π-lc {succ m} {succ n} e = ap succ (β-to-π-lc (Succ-lc e))
\end{code}
We now assume functional and propositional extensionality.
\begin{code}
module _ (fe : Fun-Ext) (pe : Prop-Ext) where
π-is-set : is-set π
π-is-set = W-is-set (Ξ© π₯) _holds fe (Ξ©-is-set fe pe)
Succ-is-embedding : is-embedding Succ
Succ-is-embedding = lc-maps-into-sets-are-embeddings Succ Succ-lc π-is-set
β-to-π-is-embedding : is-embedding β-to-π
β-to-π-is-embedding = lc-maps-into-sets-are-embeddings β-to-π β-to-π-lc π-is-set
SuccβΊ-PredβΊ : (nβΊ : πβΊ) β SuccβΊ (PredβΊ nβΊ) οΌ nβΊ
SuccβΊ-PredβΊ (n , pos) = to-subtype-οΌ being-positive-is-prop I
where
I = Succ (Pred n pos) οΌβ¨ refl β©
Suc β€ (Ξ» _ β Pred n pos) οΌβ¨ II β©
Suc (positive n) (Pred n) οΌβ¨ Suc-Pred n β©
n β
where
II = to-π-οΌ
(((true-gives-equal-β€ pe fe
(is-positive n)
(being-positive-is-prop n)
pos)β»ΒΉ) ,
dfunext fe (Ξ» h β ap (Pred n) (being-positive-is-prop n _ _)))
\end{code}
Hence πβΊ and π are equivalent.
\begin{code}
πβΊ-β-π : πβΊ β π
πβΊ-β-π = qinveq PredβΊ (SuccβΊ , SuccβΊ-PredβΊ , PredβΊ-SuccβΊ)
\end{code}
End of the anonymous submodule assuming Fun-Ext and Prop-Ext.
Our numbers "count" the number of elements of certain types.
\begin{code}
ππ²π· : π β π₯ Μ
ππ²π· (ssup p ns) = p holds + (Ξ£ h κ p holds , ππ²π· (ns h))
\end{code}
The map Fin : β β π€β factors as β-to-π : β β π followed
by ππ²π· : π β π₯.
\begin{code}
Fin-factor : (n : β) β ππ²π· (β-to-π n) β Fin n
Fin-factor zero =
π + (Ξ£ h κ π , ππ²π· (π-elim h)) ββ¨ π-lneutral β©
(Ξ£ h κ π , ππ²π· (π-elim h)) ββ¨ empty-indexed-sum-is-π id β©
π β
Fin-factor (succ n) = I
where
IH : ππ²π· (β-to-π n) β Fin n
IH = Fin-factor n
I = ππ²π· (β-to-π (succ n)) ββ¨ β-refl _ β©
π + (Ξ£ h κ π , ππ²π· (β-to-π n)) ββ¨ II β©
π + ππ²π· (β-to-π n) ββ¨ III β©
π + Fin n ββ¨ +comm β©
Fin n + π {π₯} ββ¨ IV β©
Fin n + π {π€β} ββ¨ β-refl _ β©
Fin (succ n) β
where
II = +-cong (β-refl π) π-lneutral
III = +-cong (β-refl π) IH
IV = +-cong (β-refl _) one-π-only
\end{code}
Although we can't say that β-to-π n is a surjection, its image has
empty complement.
\begin{code}
Ξ©-to-π : Ξ© π₯ β π
Ξ©-to-π p = Suc p (Ξ» _ β Zero)
Ξ©-to-π-behaviour : (p : Ξ© π₯) β is-positive (Ξ©-to-π p) οΌ (p holds)
Ξ©-to-π-behaviour p = refl
Ξ©-is-retract-of-π : retract Ξ© π₯ of π
Ξ©-is-retract-of-π = positive , Ξ©-to-π , (Ξ» x β refl)
Ξ©-to-π-is-section : is-section Ξ©-to-π
Ξ©-to-π-is-section = section-is-section Ξ©-is-retract-of-π
Ξ©-to-π-lc : left-cancellable Ξ©-to-π
Ξ©-to-π-lc = sections-are-lc Ξ©-to-π Ξ©-to-π-is-section
decidability-of-positivity-gives-LEM : ((n : π) β is-decidable (is-positive n))
β LEM π₯
decidability-of-positivity-gives-LEM f p = I
where
I : is-decidable (is-positive (Ξ©-to-π p))
I = f (Ξ©-to-π p)
\end{code}
We now assume functional and propositional extensionality
again. Sections are not necessarily embeddings
(https://doi.org/10.2168/LMCS-12(3:9)2016), but sections into sets
are:
\begin{code}
module _ (fe : Fun-Ext) (pe : Prop-Ext) where
Ξ©-to-π-is-embedding : is-embedding Ξ©-to-π
Ξ©-to-π-is-embedding = lc-maps-into-sets-are-embeddings
Ξ©-to-π
Ξ©-to-π-lc
(π-is-set fe pe)
lc-map-from-Ξ©-to-β-gives-LEM : (Ξ£ f κ (Ξ© π₯ β β) , left-cancellable f)
β LEM π₯
lc-map-from-Ξ©-to-β-gives-LEM (f , f-lc) p = I (β-is-discrete (f β€) (f p))
where
I : is-decidable (f β€ οΌ f p) β is-decidable (p holds)
I (inl e) = inl (Idtofun (ap _holds (f-lc e)) β)
I (inr Ξ½) = inr (Ξ» (h : p holds)
β Ξ½ (ap f (holds-gives-equal-β€ pe fe p h)β»ΒΉ))
\end{code}
We now further assume that propositional truncations exist.
\begin{code}
module _ (pt : propositional-truncations-exist) where
open PropositionalTruncation pt
open import UF.ImageAndSurjection pt
β-to-π-surjection-gives-LEM : is-surjection β-to-π β LEM π₯
β-to-π-surjection-gives-LEM s p = I
where
f : β β π
f = β-to-π ,
surjective-embeddings-are-equivs β-to-π
(β-to-π-is-embedding fe pe) s
g : Ξ© π₯ β β
g = β f ββ»ΒΉ β Ξ©-to-π
g-is-emb : is-embedding g
g-is-emb = β-is-embedding
Ξ©-to-π-is-embedding
(equivs-are-embeddings β f ββ»ΒΉ β f ββ»ΒΉ-is-equiv)
I : is-decidable (p holds)
I = lc-map-from-Ξ©-to-β-gives-LEM (g , embeddings-are-lc g g-is-emb) p
\end{code}
Although the above shows that β-to-π isn't necessarily a surjection,
its image has empty complement in the sense that there is no π· : π
which is different from β-to-π n for every n : β.
\begin{code}
β-to-π-has-empty-complement : Β¬ (Ξ£ π· κ π , ((n : β) β β-to-π n β π·))
β-to-π-has-empty-complement = uncurry Ο
where
Ο : (π· : π) β Β¬ ((n : β) β β-to-π n β π·)
Ο (ssup p ns) f = III IV
where
I : Zero β ssup p ns
I = f 0
II : Β¬ (p holds) β Zero οΌ ssup p ns
II Ξ½ = to-π-οΌ ((IIβ β»ΒΉ) , dfunext fe (Ξ» x β π-elim x))
where
IIβ : p οΌ β₯
IIβ = fails-gives-equal-β₯ pe fe p Ξ½
III : ¬¬ (p holds)
III h = I (II h)
IV : Β¬ (p holds)
IV h = Ο (ns h) f'
where
IVβ : p οΌ β€
IVβ = holds-gives-equal-β€ pe fe p h
f' : (n : β) β β-to-π n β ns h
f' n e = f (succ n) IVβ
where
IVβ = Succ (β-to-π n) οΌβ¨ refl β©
Suc β€ (Ξ» _ β β-to-π n) οΌβ¨ IVβ β©
Suc β€ (Ξ» _ β ns h) οΌβ¨ IVβ β©
Suc p ns β
where
IVβ = ap (Ξ» - β Suc β€ (Ξ» _ β -)) e
IVβ = to-π-οΌ
((IVβ β»ΒΉ) ,
dfunext fe (Ξ» _ β ap ns (holds-is-prop p _ _)))
\end{code}
So if excluded middle holds then β-to-π is a surjection and the types β
and π are equivalent.
\begin{code}
LEM-gives-β-to-π-surjection : LEM (π₯ βΊ) β is-surjection β-to-π
LEM-gives-β-to-π-surjection lem π· = III
where
em = LEM-gives-EM lem
I : Β¬ ((n : β) β β-to-π n β π·)
I = not-Ξ£-implies-Ξ -not β-to-π-has-empty-complement π·
II : β n κ β , ¬¬ (β-to-π n οΌ π·)
II = not-Ξ -implies-β-not pt em (Ξ» n β negations-are-props fe) I
III : β n κ β , β-to-π n οΌ π·
III = β₯β₯-functor
(Ξ» (n , Ξ½) β (n , ¬¬-elim (em (β-to-π n οΌ π·) (π-is-set fe pe)) Ξ½))
II
\end{code}
It is possible to reach the same conclusion assuming LEM π₯ rather than
LEM (π₯ βΊ).
\begin{code}
LEM-gives-β-to-π-is-equiv : LEM π₯ β is-equiv β-to-π
LEM-gives-β-to-π-is-equiv lem =
qinvs-are-equivs β-to-π
((Ξ» π· β f π· (lem (positive π·))) ,
(Ξ» n β Ξ· n (lem (positive (β-to-π n)))) ,
(Ξ» π· β Ξ΅ π· (lem (positive π·))))
where
f : (π· : π) β is-decidable (is-positive π·) β β
f (ssup p ns) (inr Ξ½) = zero
f (ssup p ns) (inl h) = succ (f (ns h) (lem (positive (ns h))))
Ξ· : (n : β) (d : is-decidable (is-positive (β-to-π n)))
β f (β-to-π n) d οΌ n
Ξ· zero (inr Ξ½) = refl
Ξ· (succ n) (inr Ξ½) = π-elim (Ξ½ (Succ-is-positive (Succ (β-to-π n))))
Ξ· (succ n) (inl h) = ap succ (Ξ· n (lem (positive (β-to-π n))))
Ξ΅ : (π· : π) (d : is-decidable (is-positive π·)) β β-to-π (f π· d) οΌ π·
Ξ΅ (ssup p ns) (inr Ξ½) = to-π-οΌ ((I β»ΒΉ) , dfunext fe (Ξ» x β π-elim x))
where
I : p οΌ β₯
I = fails-gives-equal-β₯ pe fe p Ξ½
Ξ΅ (ssup p ns) (inl h) =
to-π-οΌ ((I β»ΒΉ) ,
dfunext fe (Ξ» _ β
β-to-π (f (ns h) (lem (positive (ns h)))) οΌβ¨ II β©
ns h οΌβ¨ III β©
(ns β transport _holds (I β»ΒΉ)) _ β))
where
I : p οΌ β€
I = holds-gives-equal-β€ pe fe p h
II = Ξ΅ (ns h) (lem (positive (ns h)))
III = ap ns (holds-is-prop p h _)
\end{code}
TODO. Show that π has the structure of an ordinal. This requires more
work.