Martin Escardo, 18 January 2021.
Several additions by Tom de Jong in May and June 2024 and July 2025.
\begin{code}
{-# OPTIONS --safe --without-K --lossy-unification #-}
open import UF.Univalence
module Ordinals.AdditionProperties
(ua : Univalence)
where
open import UF.Base
open import UF.ClassicalLogic
open import UF.Embeddings hiding (β_β)
open import UF.Equiv
open import UF.EquivalenceExamples
open import UF.FunExt
open import UF.Subsingletons
open import UF.Subsingletons-FunExt
open import UF.UA-FunExt
private
fe : FunExt
fe = Univalence-gives-FunExt ua
fe' : Fun-Ext
fe' {π€} {π₯} = fe π€ π₯
pe : PropExt
pe = Univalence-gives-PropExt ua
open import MLTT.Plus-Properties
open import MLTT.Spartan
open import Notation.CanonicalMap
open import Ordinals.Arithmetic fe
open import Ordinals.Equivalence
open import Ordinals.Maps
open import Ordinals.Notions
open import Ordinals.OrdinalOfOrdinals ua
open import Ordinals.Propositions ua
open import Ordinals.Type
open import Ordinals.Underlying
πβ-left-neutral : (Ξ± : Ordinal π€) β πβ +β Ξ± οΌ Ξ±
πβ-left-neutral {π€} Ξ± = eqtoidβ (ua π€) fe' (πβ +β Ξ±) Ξ± h
where
f : π + β¨ Ξ± β© β β¨ Ξ± β©
f = β π-lneutral β
f-preserves-order : (x y : π + β¨ Ξ± β©) β x βΊβ¨ πβ +β Ξ± β© y β f x βΊβ¨ Ξ± β© f y
f-preserves-order (inr x) (inr y) l = l
f-reflects-order : (x y : π + β¨ Ξ± β©) β f x βΊβ¨ Ξ± β© f y β x βΊβ¨ πβ +β Ξ± β© y
f-reflects-order (inr x) (inr y) l = l
h : (πβ +β Ξ±) ββ Ξ±
h = f , order-preserving-reflecting-equivs-are-order-equivs (πβ +β Ξ±) Ξ± f
(ββ-is-equiv π-lneutral) f-preserves-order f-reflects-order
πβ-right-neutral : (Ξ± : Ordinal π€) β Ξ± +β πβ οΌ Ξ±
πβ-right-neutral Ξ± = eqtoidβ (ua _) fe' (Ξ± +β πβ) Ξ± h
where
f : β¨ Ξ± β© + π β β¨ Ξ± β©
f = β π-rneutral' β
f-preserves-order : is-order-preserving (Ξ± +β πβ) Ξ± f
f-preserves-order (inl x) (inl y) l = l
f-reflects-order : is-order-reflecting (Ξ± +β πβ) Ξ± f
f-reflects-order (inl x) (inl y) l = l
h : (Ξ± +β πβ) ββ Ξ±
h = f , order-preserving-reflecting-equivs-are-order-equivs (Ξ± +β πβ) Ξ± f
(ββ-is-equiv π-rneutral') f-preserves-order f-reflects-order
+β-assoc : (Ξ± Ξ² Ξ³ : Ordinal π€) β (Ξ± +β Ξ²) +β Ξ³ οΌ Ξ± +β (Ξ² +β Ξ³)
+β-assoc Ξ± Ξ² Ξ³ = eqtoidβ (ua _) fe' ((Ξ± +β Ξ²) +β Ξ³) (Ξ± +β (Ξ² +β Ξ³)) h
where
f : β¨ (Ξ± +β Ξ²) +β Ξ³ β© β β¨ Ξ± +β (Ξ² +β Ξ³) β©
f = β +assoc β
f-preserves-order : is-order-preserving ((Ξ± +β Ξ²) +β Ξ³) (Ξ± +β (Ξ² +β Ξ³)) f
f-preserves-order (inl (inl x)) (inl (inl y)) l = l
f-preserves-order (inl (inl x)) (inl (inr y)) l = l
f-preserves-order (inl (inr x)) (inl (inr y)) l = l
f-preserves-order (inl (inl x)) (inr y) l = l
f-preserves-order (inl (inr x)) (inr y) l = l
f-preserves-order (inr x) (inr y) l = l
f-reflects-order : is-order-reflecting ((Ξ± +β Ξ²) +β Ξ³) (Ξ± +β (Ξ² +β Ξ³)) f
f-reflects-order (inl (inl x)) (inl (inl y)) l = l
f-reflects-order (inl (inl x)) (inl (inr y)) l = l
f-reflects-order (inl (inr x)) (inl (inr y)) l = l
f-reflects-order (inl (inl x)) (inr y) l = l
f-reflects-order (inl (inr x)) (inr y) l = l
f-reflects-order (inr x) (inl (inl y)) l = l
f-reflects-order (inr x) (inl (inr y)) l = l
f-reflects-order (inr x) (inr y) l = l
h : ((Ξ± +β Ξ²) +β Ξ³) ββ (Ξ± +β (Ξ² +β Ξ³))
h = f , order-preserving-reflecting-equivs-are-order-equivs
((Ξ± +β Ξ²) +β Ξ³) (Ξ± +β (Ξ² +β Ξ³))
f (ββ-is-equiv +assoc) f-preserves-order f-reflects-order
+β-β-left : {Ξ± Ξ² : Ordinal π€} (a : β¨ Ξ± β©)
β (Ξ± β a) οΌ ((Ξ± +β Ξ²) β inl a)
+β-β-left {π€} {Ξ±} {Ξ²} a = h
where
Ξ³ = Ξ± β a
Ξ΄ = (Ξ± +β Ξ²) β inl a
f : β¨ Ξ³ β© β β¨ Ξ΄ β©
f (x , l) = inl x , l
g : β¨ Ξ΄ β© β β¨ Ξ³ β©
g (inl x , l) = x , l
Ξ· : g β f βΌ id
Ξ· u = refl
Ξ΅ : f β g βΌ id
Ξ΅ (inl x , l) = refl
f-is-equiv : is-equiv f
f-is-equiv = qinvs-are-equivs f (g , Ξ· , Ξ΅)
f-is-order-preserving : is-order-preserving Ξ³ Ξ΄ f
f-is-order-preserving (x , _) (x' , _) l = l
g-is-order-preserving : is-order-preserving Ξ΄ Ξ³ g
g-is-order-preserving (inl x , _) (inl x' , _) l = l
h : Ξ³ οΌ Ξ΄
h = eqtoidβ (ua π€) fe' Ξ³ Ξ΄
(f , f-is-order-preserving , f-is-equiv , g-is-order-preserving)
+β-left-β΄ : (Ξ± Ξ² : Ordinal π€)
β Ξ± β΄ Ξ± +β Ξ²
+β-left-β΄ Ξ± Ξ² = to-β΄ Ξ± (Ξ± +β Ξ²) (Ξ» a β inl a , +β-β-left a)
+β-β-right : {Ξ± Ξ² : Ordinal π€} (b : β¨ Ξ² β©)
β (Ξ± +β (Ξ² β b)) οΌ ((Ξ± +β Ξ²) β inr b)
+β-β-right {π€} {Ξ±} {Ξ²} b = h
where
Ξ³ = Ξ± +β (Ξ² β b)
Ξ΄ = (Ξ± +β Ξ²) β inr b
f : β¨ Ξ³ β© β β¨ Ξ΄ β©
f (inl a) = inl a , β
f (inr (y , l)) = inr y , l
g : β¨ Ξ΄ β© β β¨ Ξ³ β©
g (inl a , β) = inl a
g (inr y , l) = inr (y , l)
Ξ· : g β f βΌ id
Ξ· (inl a) = refl
Ξ· (inr (y , l)) = refl
Ξ΅ : f β g βΌ id
Ξ΅ (inl a , β) = refl
Ξ΅ (inr y , l) = refl
f-is-equiv : is-equiv f
f-is-equiv = qinvs-are-equivs f (g , Ξ· , Ξ΅)
f-is-order-preserving : is-order-preserving Ξ³ Ξ΄ f
f-is-order-preserving (inl _) (inl _) l = l
f-is-order-preserving (inl _) (inr _) l = l
f-is-order-preserving (inr _) (inr _) l = l
g-is-order-preserving : is-order-preserving Ξ΄ Ξ³ g
g-is-order-preserving (inl _ , _) (inl _ , _) l = l
g-is-order-preserving (inl _ , _) (inr _ , _) l = l
g-is-order-preserving (inr _ , _) (inr _ , _) l = l
h : Ξ³ οΌ Ξ΄
h = eqtoidβ (ua π€) fe' Ξ³ Ξ΄
(f , f-is-order-preserving , f-is-equiv , g-is-order-preserving)
+β-β²-left : {Ξ± Ξ² : Ordinal π€} (a : β¨ Ξ± β©)
β (Ξ± β a) β² (Ξ± +β Ξ²)
+β-β²-left {π€} {Ξ±} {Ξ²} a = inl a , +β-β-left a
+β-β²-right : {Ξ± Ξ² : Ordinal π€} (b : β¨ Ξ² β©)
β (Ξ± +β (Ξ² β b)) β² (Ξ± +β Ξ²)
+β-β²-right {π€} {Ξ±} {Ξ²} b = inr b , +β-β-right {π€} {Ξ±} {Ξ²} b
+β-increasing-on-right : {Ξ± Ξ² Ξ³ : Ordinal π€}
β Ξ² β² Ξ³
β (Ξ± +β Ξ²) β² (Ξ± +β Ξ³)
+β-increasing-on-right {π€} {Ξ±} {Ξ²} {Ξ³} (c , p) = inr c , q
where
q = (Ξ± +β Ξ²) οΌβ¨ ap (Ξ± +β_) p β©
(Ξ± +β (Ξ³ β c)) οΌβ¨ +β-β-right c β©
((Ξ± +β Ξ³) β inr c) β
+β-right-monotone : (Ξ± Ξ² Ξ³ : Ordinal π€)
β Ξ² βΌ Ξ³
β (Ξ± +β Ξ²) βΌ (Ξ± +β Ξ³)
+β-right-monotone Ξ± Ξ² Ξ³ m = to-βΌ Ο
where
l : (a : β¨ Ξ± β©) β ((Ξ± +β Ξ²) β inl a) β² (Ξ± +β Ξ³)
l a = o
where
n : (Ξ± β a) β² (Ξ± +β Ξ³)
n = +β-β²-left a
o : ((Ξ± +β Ξ²) β inl a) β² (Ξ± +β Ξ³)
o = transport (_β² (Ξ± +β Ξ³)) (+β-β-left a) n
r : (b : β¨ Ξ² β©) β ((Ξ± +β Ξ²) β inr b) β² (Ξ± +β Ξ³)
r b = s
where
o : (Ξ² β b) β² Ξ³
o = from-βΌ m b
p : (Ξ± +β (Ξ² β b)) β² (Ξ± +β Ξ³)
p = +β-increasing-on-right o
q : Ξ± +β (Ξ² β b) οΌ (Ξ± +β Ξ²) β inr b
q = +β-β-right b
s : ((Ξ± +β Ξ²) β inr b) β² (Ξ± +β Ξ³)
s = transport (_β² (Ξ± +β Ξ³)) q p
Ο : (x : β¨ Ξ± +β Ξ² β©) β ((Ξ± +β Ξ²) β x) β² (Ξ± +β Ξ³)
Ο = dep-cases l r
+β-right-monotone-β΄ : (Ξ± : Ordinal π€) β is-β΄-preserving (Ξ± +β_)
+β-right-monotone-β΄ Ξ± Ξ² Ξ³ l =
βΌ-gives-β΄ (Ξ± +β Ξ²) (Ξ± +β Ξ³) (+β-right-monotone Ξ± Ξ² Ξ³ (β΄-gives-βΌ Ξ² Ξ³ l))
\end{code}
Added by Tom de Jong in July/October 2025. The following proof has
better computational properties (and is arguably simpler) than the one
above.
\begin{code}
+β-right-monotone-β΄' : (Ξ± : Ordinal π€) β is-β΄-preserving (Ξ± +β_)
+β-right-monotone-β΄' Ξ± Ξ² Ξ³ π@(f , f-sim) = g , g-init-seg , g-order-pres
where
g : β¨ Ξ± +β Ξ² β© β β¨ Ξ± +β Ξ³ β©
g (inl a) = inl a
g (inr b) = inr (f b)
g-order-pres : is-order-preserving (Ξ± +β Ξ²) (Ξ± +β Ξ³) g
g-order-pres (inl a) (inl a') l = l
g-order-pres (inl a) (inr b) l = l
g-order-pres (inr b) (inr b') l =
simulations-are-order-preserving Ξ² Ξ³ f f-sim b b' l
g-init-seg : is-initial-segment (Ξ± +β Ξ²) (Ξ± +β Ξ³) g
g-init-seg (inl a) (inl a') l = inl a' , l , refl
g-init-seg (inr b) (inl a) l = inl a , β , refl
g-init-seg (inr b) (inr b') l =
inr (prβ I) , prβ (prβ I) , ap inr (prβ (prβ I))
where
I : Ξ£ b'' κ β¨ Ξ² β© , (b'' βΊβ¨ Ξ² β© b) Γ (f b'' οΌ b')
I = simulations-are-initial-segments Ξ² Ξ³ f f-sim b b' l
\end{code}
TODO. Find better names for the following lemmas.
\begin{code}
AP-lemmaβ : {Ξ± Ξ² : Ordinal π€}
β Ξ± βΌ (Ξ± +β Ξ²)
AP-lemmaβ {π€} {Ξ±} {Ξ²} = to-βΌ Ο
where
Ο : (a : β¨ Ξ± β©) β Ξ£ z κ β¨ Ξ± +β Ξ² β© , (Ξ± β a) οΌ ((Ξ± +β Ξ²) β z)
Ο a = inl a , (+β-β-left a)
AP-lemmaβ : {Ξ± Ξ² : Ordinal π€}
(a : β¨ Ξ± β©)
β (Ξ± +β Ξ²) β (Ξ± β a)
AP-lemmaβ {π€} {Ξ±} {Ξ²} a p = irrefl (OO π€) (Ξ± +β Ξ²) m
where
l : (Ξ± +β Ξ²) β² Ξ±
l = (a , p)
m : (Ξ± +β Ξ²) β² (Ξ± +β Ξ²)
m = AP-lemmaβ (Ξ± +β Ξ²) l
AP-lemmaβ : {Ξ± Ξ² : Ordinal π€} (a : β¨ Ξ± β©)
β Ξ± οΌ Ξ²
β Ξ£ b κ β¨ Ξ² β© , (Ξ± β a) οΌ (Ξ² β b)
AP-lemmaβ a refl = a , refl
AP-lemmaβ : {Ξ± Ξ² Ξ³ : Ordinal π€} (b : β¨ Ξ² β©) (z : β¨ Ξ± +β Ξ³ β©)
β ((Ξ± +β Ξ²) β inr b) οΌ ((Ξ± +β Ξ³) β z)
β Ξ£ c κ β¨ Ξ³ β© , z οΌ inr c
AP-lemmaβ {π€} {Ξ±} {Ξ²} {Ξ³} b (inl a) p = π-elim (AP-lemmaβ a q)
where
q : (Ξ± +β (Ξ² β b)) οΌ (Ξ± β a)
q = +β-β-right b β p β (+β-β-left a)β»ΒΉ
AP-lemmaβ b (inr c) p = c , refl
+β-left-cancellable : (Ξ± Ξ² Ξ³ : Ordinal π€)
β (Ξ± +β Ξ²) οΌ (Ξ± +β Ξ³)
β Ξ² οΌ Ξ³
+β-left-cancellable {π€} Ξ± = g
where
P : Ordinal π€ β π€ βΊ Μ
P Ξ² = β Ξ³ β (Ξ± +β Ξ²) οΌ (Ξ± +β Ξ³) β Ξ² οΌ Ξ³
Ο : β Ξ²
β (β b β P (Ξ² β b))
β P Ξ²
Ο Ξ² f Ξ³ p = Extensionality (OO π€) Ξ² Ξ³ (to-βΌ u) (to-βΌ v)
where
u : (b : β¨ Ξ² β©) β (Ξ² β b) β² Ξ³
u b = c , t
where
z : β¨ Ξ± +β Ξ³ β©
z = prβ (AP-lemmaβ (inr b) p)
r : ((Ξ± +β Ξ²) β inr b) οΌ ((Ξ± +β Ξ³) β z)
r = prβ (AP-lemmaβ (inr b) p)
c : β¨ Ξ³ β©
c = prβ (AP-lemmaβ b z r)
s : z οΌ inr c
s = prβ (AP-lemmaβ b z r)
q = (Ξ± +β (Ξ² β b)) οΌβ¨ +β-β-right b β©
((Ξ± +β Ξ²) β inr b) οΌβ¨ r β©
((Ξ± +β Ξ³) β z) οΌβ¨ ap ((Ξ± +β Ξ³) β_) s β©
((Ξ± +β Ξ³) β inr c) οΌβ¨ (+β-β-right c)β»ΒΉ β©
(Ξ± +β (Ξ³ β c)) β
t : (Ξ² β b) οΌ (Ξ³ β c)
t = f b (Ξ³ β c) q
v : (c : β¨ Ξ³ β©) β (Ξ³ β c) β² Ξ²
v c = b , (t β»ΒΉ)
where
z : β¨ Ξ± +β Ξ² β©
z = prβ (AP-lemmaβ (inr c) (p β»ΒΉ))
r : ((Ξ± +β Ξ³) β inr c) οΌ ((Ξ± +β Ξ²) β z)
r = prβ (AP-lemmaβ (inr c) (p β»ΒΉ))
b : β¨ Ξ² β©
b = prβ (AP-lemmaβ c z r)
s : z οΌ inr b
s = prβ (AP-lemmaβ c z r)
q = (Ξ± +β (Ξ³ β c)) οΌβ¨ +β-β-right c β©
((Ξ± +β Ξ³) β inr c) οΌβ¨ r β©
((Ξ± +β Ξ²) β z) οΌβ¨ ap ((Ξ± +β Ξ²) β_) s β©
((Ξ± +β Ξ²) β inr b) οΌβ¨ (+β-β-right b)β»ΒΉ β©
(Ξ± +β (Ξ² β b)) β
t : (Ξ² β b) οΌ (Ξ³ β c)
t = f b (Ξ³ β c) (q β»ΒΉ)
g : (Ξ² : Ordinal π€) β P Ξ²
g = transfinite-induction-on-OO P Ο
left-+β-is-embedding : (Ξ± : Ordinal π€) β is-embedding (Ξ± +β_)
left-+β-is-embedding Ξ± = lc-maps-into-sets-are-embeddings (Ξ± +β_)
(Ξ» {Ξ²} {Ξ³} β +β-left-cancellable Ξ± Ξ² Ξ³)
(the-type-of-ordinals-is-a-set (ua _) fe')
\end{code}
This implies that the function Ξ± +β_ reflects the _β²_ ordering:
\begin{code}
+β-left-reflects-β² : (Ξ± Ξ² Ξ³ : Ordinal π€)
β (Ξ± +β Ξ²) β² (Ξ± +β Ξ³)
β Ξ² β² Ξ³
+β-left-reflects-β² Ξ± Ξ² Ξ³ (inl a , p) = π-elim (AP-lemmaβ a q)
where
q : (Ξ± +β Ξ²) οΌ (Ξ± β a)
q = p β (+β-β-left a)β»ΒΉ
+β-left-reflects-β² Ξ± Ξ² Ξ³ (inr c , p) = l
where
q : (Ξ± +β Ξ²) οΌ (Ξ± +β (Ξ³ β c))
q = p β (+β-β-right c)β»ΒΉ
r : Ξ² οΌ (Ξ³ β c)
r = +β-left-cancellable Ξ± Ξ² (Ξ³ β c) q
l : Ξ² β² Ξ³
l = c , r
\end{code}
And in turn this implies that the function Ξ± +β_ reflects the _βΌ_
partial ordering:
\begin{code}
+β-left-reflects-βΌ : (Ξ± Ξ² Ξ³ : Ordinal π€)
β (Ξ± +β Ξ²) βΌ (Ξ± +β Ξ³)
β Ξ² βΌ Ξ³
+β-left-reflects-βΌ {π€} Ξ± Ξ² Ξ³ l = to-βΌ (Ο Ξ² l)
where
Ο : (Ξ² : Ordinal π€)
β (Ξ± +β Ξ²) βΌ (Ξ± +β Ξ³)
β (b : β¨ Ξ² β©) β (Ξ² β b) β² Ξ³
Ο Ξ² l b = o
where
m : (Ξ± +β (Ξ² β b)) β² (Ξ± +β Ξ²)
m = +β-β²-right b
n : (Ξ± +β (Ξ² β b)) β² (Ξ± +β Ξ³)
n = l (Ξ± +β (Ξ² β b)) m
o : (Ξ² β b) β² Ξ³
o = +β-left-reflects-β² Ξ± (Ξ² β b) Ξ³ n
+β-left-reflects-β΄ : (Ξ± Ξ² Ξ³ : Ordinal π€)
β (Ξ± +β Ξ²) β΄ (Ξ± +β Ξ³)
β Ξ² β΄ Ξ³
+β-left-reflects-β΄ Ξ± Ξ² Ξ³ l =
βΌ-gives-β΄ Ξ² Ξ³ (+β-left-reflects-βΌ Ξ± Ξ² Ξ³ (β΄-gives-βΌ (Ξ± +β Ξ²) (Ξ± +β Ξ³) l))
\end{code}
Added 4th April 2022.
\begin{code}
πβ-least-β΄ : (Ξ± : Ordinal π€) β πβ {π€} β΄ Ξ±
πβ-least-β΄ Ξ± = unique-from-π , (Ξ» x y l β π-elim x) , (Ξ» x y l β π-elim x)
πβ-least : (Ξ± : Ordinal π€) β πβ {π€} βΌ Ξ±
πβ-least Ξ± = β΄-gives-βΌ πβ Ξ± (πβ-least-β΄ Ξ±)
\end{code}
Added 11 April 2025 by Tom de Jong.
\begin{code}
initial-segment-of-least-element-is-πβ : (Ξ± : Ordinal π€) (a : β¨ Ξ± β©)
β is-least Ξ± a
β Ξ± β a οΌ πβ
initial-segment-of-least-element-is-πβ Ξ± a a-is-least =
β΄-antisym (Ξ± β a) πβ
(to-β΄ (Ξ± β a) πβ (Ξ» (x , l) β π-elim (least-is-minimal Ξ± a a-is-least x l)))
(πβ-least-β΄ (Ξ± β a))
\end{code}
Originally added 21st April 2022 by MartΓn EscardΓ³.
Moved up here on 27th May 2024 by Tom de Jong.
\begin{code}
successor-lemma-left : (Ξ± : Ordinal π€) (x : β¨ Ξ± β©) β ((Ξ± +β πβ) β inl x) β΄ Ξ±
successor-lemma-left Ξ± x = III
where
I : (Ξ± β x) β΄ Ξ±
I = segment-β΄ Ξ± x
II : (Ξ± β x) οΌ ((Ξ± +β πβ) β inl x)
II = +β-β-left x
III : ((Ξ± +β πβ) β inl x) β΄ Ξ±
III = transport (_β΄ Ξ±) II I
successor-lemma-right : (Ξ± : Ordinal π€) β (Ξ± +β πβ) β inr β οΌ Ξ±
successor-lemma-right Ξ± = III
where
I : (πβ β β) β΄ πβ
I = (Ξ» x β π-elim (prβ x)) , (Ξ» x β π-elim (prβ x)) , (Ξ» x β π-elim (prβ x))
II : (πβ β β) οΌ πβ
II = β΄-antisym _ _ I (πβ-least-β΄ (πβ β β))
III : (Ξ± +β πβ) β inr β οΌ Ξ±
III = (Ξ± +β πβ) β inr β οΌβ¨ (+β-β-right β)β»ΒΉ β©
Ξ± +β (πβ β β) οΌβ¨ ap (Ξ± +β_) II β©
Ξ± +β πβ οΌβ¨ πβ-right-neutral Ξ± β©
Ξ± β
successor-increasing : (Ξ± : Ordinal π€) β Ξ± β² (Ξ± +β πβ)
successor-increasing Ξ± = inr β , ((successor-lemma-right Ξ±)β»ΒΉ)
\end{code}
Added on 24th May 2024 by Tom de Jong.
\begin{code}
upper-bound-of-successors-of-initial-segments :
(Ξ± : Ordinal π€) (a : β¨ Ξ± β©) β ((Ξ± β a) +β πβ) β΄ Ξ±
upper-bound-of-successors-of-initial-segments Ξ± a = to-β΄ ((Ξ± β a) +β πβ) Ξ± I
where
I : (x : β¨ (Ξ± β a) +β πβ β©) β (((Ξ± β a) +β πβ) β x) β² Ξ±
I (inl (b , l)) = b , (((Ξ± β a) +β πβ) β inl (b , l) οΌβ¨ eβ β©
(Ξ± β a) β (b , l) οΌβ¨ eβ β©
Ξ± β b β)
where
eβ = (+β-β-left (b , l)) β»ΒΉ
eβ = iterated-β Ξ± a b l
I (inr β) = a , successor-lemma-right (Ξ± β a)
\end{code}
End of addition.
Classically, if Ξ± βΌ Ξ² then there is (a necessarily unique) Ξ³ with
Ξ± +β Ξ³ οΌ Ξ². But this not necessarily the case constructively. For
that purpose, we first characterize the order of subsingleton
ordinals.
\begin{code}
module _ {π€ : Universe}
(P Q : π€ Μ )
(P-is-prop : is-prop P)
(Q-is-prop : is-prop Q)
where
private
p q : Ordinal π€
p = prop-ordinal P P-is-prop
q = prop-ordinal Q Q-is-prop
factβ : p β² q β Β¬ P Γ Q
factβ (y , r) = u , y
where
s : P οΌ (Q Γ π)
s = ap β¨_β© r
u : Β¬ P
u p = π-elim (prβ (β idtoeq P (Q Γ π) s β p))
factβ-converse : Β¬ P Γ Q β p β² q
factβ-converse (u , y) = (y , g)
where
r : P οΌ Q Γ π
r = univalence-gives-propext (ua π€)
P-is-prop
Γ-π-is-prop
(Ξ» p β π-elim (u p))
(Ξ» (q , z) β π-elim z)
g : p οΌ (q β y)
g = to-Ξ£-οΌ (r ,
to-Ξ£-οΌ (dfunext fe' (Ξ» (y , z) β π-elim z) ,
being-well-order-is-prop (underlying-order (q β y)) fe _ _))
factβ : p βΌ q β (P β Q)
factβ l x = prβ (from-βΌ {π€} {p} {q} l x)
factβ-converse : (P β Q) β p βΌ q
factβ-converse f = to-βΌ {π€} {p} {q} Ο
where
r : P Γ π οΌ Q Γ π
r = univalence-gives-propext (ua π€)
Γ-π-is-prop
Γ-π-is-prop
(Ξ» (p , z) β π-elim z)
(Ξ» (q , z) β π-elim z)
Ο : (x : β¨ p β©) β (p β x) β² q
Ο x = f x , s
where
s : ((P Γ π) , (Ξ» x x' β π) , _) οΌ ((Q Γ π) , (Ξ» y y' β π) , _)
s = to-Ξ£-οΌ (r ,
to-Ξ£-οΌ (dfunext fe' (Ξ» z β π-elim (prβ z)) ,
being-well-order-is-prop (underlying-order (q β f x)) fe _ _))
\end{code}
The existence of ordinal subtraction implies excluded middle.
\begin{code}
existence-of-subtraction : (π€ : Universe) β π€ βΊ Μ
existence-of-subtraction π€ = (Ξ± Ξ² : Ordinal π€)
β Ξ± βΌ Ξ²
β Ξ£ Ξ³ κ Ordinal π€ , Ξ± +β Ξ³ οΌ Ξ²
existence-of-subtraction-is-prop : is-prop (existence-of-subtraction π€)
existence-of-subtraction-is-prop = Ξ β-is-prop fe'
(Ξ» Ξ± Ξ² l β left-+β-is-embedding Ξ± Ξ²)
ordinal-subtraction-gives-excluded-middle : existence-of-subtraction π€ β EM π€
ordinal-subtraction-gives-excluded-middle {π€} Ο P P-is-prop = g
where
Ξ± = prop-ordinal P P-is-prop
Ξ² = prop-ordinal π π-is-prop
Ο = Ο Ξ± Ξ² (factβ-converse {π€} P π P-is-prop π-is-prop (Ξ» _ β β))
Ξ³ : Ordinal π€
Ξ³ = prβ Ο
r : Ξ± +β Ξ³ οΌ Ξ²
r = prβ Ο
s : P + β¨ Ξ³ β© οΌ π
s = ap β¨_β© r
t : P + β¨ Ξ³ β©
t = idtofun π (P + β¨ Ξ³ β©) (s β»ΒΉ) β
f : β¨ Ξ³ β© β Β¬ P
f c p = z
where
A : π€ Μ β π€ Μ
A X = Ξ£ x κ X , Ξ£ y κ X , x β y
u : A (P + β¨ Ξ³ β©)
u = inl p , inr c , +disjoint
v : Β¬ A π
v (x , y , d) = d (π-is-prop x y)
w : A (P + β¨ Ξ³ β©) β A π
w = transport A s
z : π
z = v (w u)
g : P + Β¬ P
g = Cases t inl (Ξ» c β inr (f c))
\end{code}
TODO. Another example where subtraction doesn't necessarily exist is
the situation (Ο +β πβ) βΌ βββ, discussed in the module
OrdinalOfOrdinals. The types Ο +β πβ and βββ are equal if and only if
LPO holds. Without assuming LPO, the image of the inclusion (Ο +β πβ)
β βββ, has empty complement, and so there is nothing that can be added
to (Ο +β πβ) to get βββ, unless LPO holds.
\begin{code}
open import UF.Retracts
open import UF.SubtypeClassifier
retract-Ξ©-of-Ordinal : retract (Ξ© π€) of (Ordinal π€)
retract-Ξ©-of-Ordinal {π€} = r , s , Ξ·
where
s : Ξ© π€ β Ordinal π€
s (P , i) = prop-ordinal P i
r : Ordinal π€ β Ξ© π€
r Ξ± = has-least Ξ± , having-least-is-prop fe' Ξ±
Ξ· : r β s βΌ id
Ξ· (P , i) = to-subtype-οΌ (Ξ» _ β being-prop-is-prop fe') t
where
f : P β has-least (prop-ordinal P i)
f p = p , (Ξ» x u β id)
g : has-least (prop-ordinal P i) β P
g (p , _) = p
t : has-least (prop-ordinal P i) οΌ P
t = pe π€ (having-least-is-prop fe' (prop-ordinal P i)) i g f
\end{code}
Added 17 September 2024 by Fredrik Nordvall Forsberg.
\begin{code}
left-preserves-least : (Ξ± Ξ² : Ordinal π€)
β (aβ : β¨ Ξ± β©) β is-least Ξ± aβ β is-least (Ξ± +β Ξ²) (inl aβ)
left-preserves-least Ξ± Ξ² aβ aβ-least (inl x) (inl u) l = aβ-least x u l
left-preserves-least Ξ± Ξ² aβ aβ-least (inr x) (inl u) l = β
\end{code}
End of addition.
Added 29 March 2022.
It is not the case in general that Ξ² βΌ Ξ± +β Ξ². We work with the
equivalent order _β΄_.
\begin{code}
module _ {π€ : Universe} where
open import Ordinals.OrdinalOfTruthValues fe π€ (pe π€)
open import UF.DiscreteAndSeparated
β΄-add-taboo : Ξ©β β΄ (πβ +β Ξ©β) β typal-WEM π€
β΄-add-taboo (f , s) = VI
where
I : is-least (πβ +β Ξ©β) (inl β)
I = left-preserves-least πβ Ξ©β β (Ξ» β β ())
II : f β₯ οΌ inl β
II = simulations-preserve-least Ξ©β (πβ +β Ξ©β) β₯ (inl β) f s β₯-is-least I
III : is-isolated (f β₯)
III = transportβ»ΒΉ is-isolated II (inl-is-isolated β (π-is-discrete β))
IV : is-isolated β₯
IV = lc-maps-reflect-isolatedness
f
(simulations-are-lc Ξ©β (πβ +β Ξ©β) f s)
β₯
III
V : β P β is-prop P β Β¬ P + ¬¬ P
V P i = Cases (IV (P , i))
(Ξ» (e : β₯ οΌ (P , i))
β inl (equal-π-is-empty (ap prβ (e β»ΒΉ))))
(Ξ» (Ξ½ : β₯ β (P , i))
β inr (contrapositive
(Ξ» (u : Β¬ P)
β to-subtype-οΌ (Ξ» _ β being-prop-is-prop fe')
(empty-types-are-οΌ-π (pe π€) u)β»ΒΉ) Ξ½))
VI : β P β Β¬ P + ¬¬ P
VI = WEM-gives-typal-WEM fe' V
\end{code}
Added 5th April 2022.
Successor reflects order:
\begin{code}
succβ-reflects-β΄ : (Ξ± : Ordinal π€) (Ξ² : Ordinal π₯)
β (Ξ± +β πβ) β΄ (Ξ² +β πβ) β Ξ± β΄ Ξ²
succβ-reflects-β΄ Ξ± Ξ² (f , i , p) = g , j , q
where
k : (x : β¨ Ξ± β©) (t : β¨ Ξ² β© + π)
β f (inl x) οΌ t β Ξ£ y κ β¨ Ξ² β© , f (inl x) οΌ inl y
k x (inl y) e = y , e
k x (inr β) e = π-elim (III (f (inr β)) II)
where
I : f (inl x) βΊβ¨ Ξ² +β πβ β© (f (inr β))
I = p (inl x) (inr β) β
II : inr β βΊβ¨ Ξ² +β πβ β© (f (inr β))
II = transport (Ξ» - β - βΊβ¨ Ξ² +β πβ β© (f (inr β))) e I
III : (t : β¨ Ξ² β© + π) β Β¬ (inr β βΊβ¨ Ξ² +β πβ β© t)
III (inl y) l = π-elim l
III (inr β) l = π-elim l
h : (x : β¨ Ξ± β©) β Ξ£ y κ β¨ Ξ² β© , f (inl x) οΌ inl y
h x = k x (f (inl x)) refl
g : β¨ Ξ± β© β β¨ Ξ² β©
g x = prβ (h x)
Ο : (x : β¨ Ξ± β©) β f (inl x) οΌ inl (g x)
Ο x = prβ (h x)
j : is-initial-segment Ξ± Ξ² g
j x y l = II I
where
m : inl y βΊβ¨ Ξ² +β πβ β© f (inl x)
m = transportβ»ΒΉ (Ξ» - β inl y βΊβ¨ Ξ² +β πβ β© -) (Ο x) l
I : Ξ£ z κ β¨ Ξ± +β πβ β© , (z βΊβ¨ Ξ± +β πβ β© inl x) Γ (f z οΌ inl y)
I = i (inl x) (inl y) m
II : type-of I β Ξ£ x' κ β¨ Ξ± β© , (x' βΊβ¨ Ξ± β© x) Γ (g x' οΌ y)
II (inl x' , n , e) = x' , n , inl-lc (inl (g x') οΌβ¨ (Ο x')β»ΒΉ β©
f (inl x') οΌβ¨ e β©
inl y β)
q : is-order-preserving Ξ± Ξ² g
q x x' l = transportβ (Ξ» y y' β y βΊβ¨ Ξ² +β πβ β© y') (Ο x) (Ο x') I
where
I : f (inl x) βΊβ¨ Ξ² +β πβ β© f (inl x')
I = p (inl x) (inl x') l
succβ-reflects-βΌ : (Ξ± Ξ² : Ordinal π€) β (Ξ± +β πβ) βΌ (Ξ² +β πβ) β Ξ± βΌ Ξ²
succβ-reflects-βΌ Ξ± Ξ² l = β΄-gives-βΌ Ξ± Ξ²
(succβ-reflects-β΄ Ξ± Ξ²
(βΌ-gives-β΄ (Ξ± +β πβ) (Ξ² +β πβ) l))
succβ-preserves-βΎ : (Ξ± Ξ² : Ordinal π€) β Ξ± βΎ Ξ² β (Ξ± +β πβ) βΎ (Ξ² +β πβ)
succβ-preserves-βΎ Ξ± Ξ² = contrapositive (succβ-reflects-βΌ Ξ² Ξ±)
\end{code}
TODO. Actually (Ξ± +β πβ) β΄ (Ξ² +β πβ) is equivalent to
Ξ± ββ Ξ² or Ξ² ββ Ξ± +β πβ + Ξ³ for some (necessarily unique) Ξ³.
This condition in turn implies Ξ± β΄ Ξ² (and is equivalent to Ξ± β΄ Ξ² under
excluded middle).
However, the successor function does not preserve _β΄_ in general:
\begin{code}
succ-not-necessarily-monotone : ((Ξ± Ξ² : Ordinal π€)
β Ξ± β΄ Ξ²
β (Ξ± +β πβ) β΄ (Ξ² +β πβ))
β typal-WEM π€
succ-not-necessarily-monotone {π€} Ο = XII
where
module _ (P : π€ Μ ) (isp : is-prop P) where
Ξ± : Ordinal π€
Ξ± = prop-ordinal P isp
I : (Ξ± +β πβ) β΄ πβ
I = Ο Ξ± πβ l
where
l : Ξ± β΄ πβ
l = unique-to-π ,
(Ξ» x y (l : y βΊβ¨ πβ β© β) β π-elim l) ,
(Ξ» x y l β l)
II : type-of I β Β¬ P + ¬¬ P
II (f , f-is-initial , f-is-order-preserving) = III (f (inr β)) refl
where
III : (y : β¨ πβ β©) β f (inr β) οΌ y β Β¬ P + ¬¬ P
III (inl β) e = inl VII
where
IV : (p : P) β f (inl p) βΊβ¨ πβ β© f (inr β)
IV p = f-is-order-preserving (inl p) (inr β) β
V : (p : P) β f (inl p) βΊβ¨ πβ β© inl β
V p = transport (Ξ» - β f (inl p) βΊβ¨ πβ β© -) e (IV p)
VI : (z : β¨ πβ β©) β Β¬ (z βΊβ¨ πβ β© inl β)
VI (inl β) l = π-elim l
VI (inr β) l = π-elim l
VII : Β¬ P
VII p = VI (f (inl p)) (V p)
III (inr β) e = inr IX
where
VIII : Ξ£ x' κ β¨ Ξ± +β πβ β© , (x' βΊβ¨ Ξ± +β πβ β© inr β) Γ (f x' οΌ inl β)
VIII = f-is-initial
(inr β)
(inl β)
(transportβ»ΒΉ (Ξ» - β inl β βΊβ¨ πβ β© -) e β)
IX : ¬¬ P
IX u = XI
where
X : β x' β Β¬ (x' βΊβ¨ Ξ± +β πβ β© inr β)
X (inl p) l = u p
X (inr β) l = π-elim l
XI : π
XI = X (prβ VIII) (prβ (prβ VIII))
XII : typal-WEM π€
XII = WEM-gives-typal-WEM fe' (Ξ» P isp β II P isp (I P isp))
\end{code}
TODO. Also the implication Ξ± β² Ξ² β (Ξ± +β πβ) β² (Ξ² +β πβ) fails in general.
\begin{code}
succ-monotone : EM (π€ βΊ) β (Ξ± Ξ² : Ordinal π€) β Ξ± β΄ Ξ² β (Ξ± +β πβ) β΄ (Ξ² +β πβ)
succ-monotone em Ξ± Ξ² l = II I
where
I : ((Ξ± +β πβ) β² (Ξ² +β πβ)) + ((Ξ± +β πβ) οΌ (Ξ² +β πβ)) + ((Ξ² +β πβ) β² (Ξ± +β πβ))
I = trichotomy _β²_ fe' em β²-is-well-order (Ξ± +β πβ) (Ξ² +β πβ)
II : type-of I β (Ξ± +β πβ) β΄ (Ξ² +β πβ)
II (inl m) = β²-gives-β΄ _ _ m
II (inr (inl e)) = transport ((Ξ± +β πβ) β΄_) e (β΄-refl (Ξ± +β πβ))
II (inr (inr m)) = transport ((Ξ± +β πβ) β΄_) VI (β΄-refl (Ξ± +β πβ))
where
III : (Ξ² +β πβ) β΄ (Ξ± +β πβ)
III = β²-gives-β΄ _ _ m
IV : Ξ² β΄ Ξ±
IV = succβ-reflects-β΄ Ξ² Ξ± III
V : Ξ± οΌ Ξ²
V = β΄-antisym _ _ l IV
VI : (Ξ± +β πβ) οΌ (Ξ² +β πβ)
VI = ap (_+β πβ) V
\end{code}
TODO. EM π€ is sufficient, because we can work with the resized order _β²β»_.
Added 4th May 2022.
\begin{code}
open import Ordinals.ToppedType fe
open import Ordinals.ToppedArithmetic fe
alternative-plusβ : (Οβ Οβ : Ordinalα΅ π€)
β [ Οβ +α΅ Οβ ] ββ ([ Οβ ] +β [ Οβ ])
alternative-plusβ Οβ Οβ = e
where
Ο
= cases (Ξ» β β Οβ) (Ξ» β β Οβ)
f : β¨ β πα΅ Ο
β© β β¨ [ Οβ ] +β [ Οβ ] β©
f (inl β , x) = inl x
f (inr β , y) = inr y
g : β¨ [ Οβ ] +β [ Οβ ] β© β β¨ β πα΅ Ο
β©
g (inl x) = (inl β , x)
g (inr y) = (inr β , y)
Ξ· : g β f βΌ id
Ξ· (inl β , x) = refl
Ξ· (inr β , y) = refl
Ξ΅ : f β g βΌ id
Ξ΅ (inl x) = refl
Ξ΅ (inr y) = refl
f-is-equiv : is-equiv f
f-is-equiv = qinvs-are-equivs f (g , Ξ· , Ξ΅)
f-is-op : is-order-preserving [ β πα΅ Ο
] ([ Οβ ] +β [ Οβ ]) f
f-is-op (inl β , _) (inl β , _) (inr (refl , l)) = l
f-is-op (inl β , _) (inr β , _) (inl β) = β
f-is-op (inr β , _) (inl β , _) (inl l) = l
f-is-op (inr β , _) (inr β , _) (inr (refl , l)) = l
g-is-op : is-order-preserving ([ Οβ ] +β [ Οβ ]) [ β πα΅ Ο
] g
g-is-op (inl _) (inl _) l = inr (refl , l)
g-is-op (inl _) (inr _) β = inl β
g-is-op (inr _) (inl _) ()
g-is-op (inr _) (inr _) l = inr (refl , l)
e : [ β πα΅ Ο
] ββ ([ Οβ ] +β [ Οβ ])
e = f , f-is-op , f-is-equiv , g-is-op
alternative-plus : (Οβ Οβ : Ordinalα΅ π€)
β [ Οβ +α΅ Οβ ] οΌ ([ Οβ ] +β [ Οβ ])
alternative-plus Οβ Οβ = eqtoidβ (ua _) fe' _ _ (alternative-plusβ Οβ Οβ)
\end{code}
Added 13 November 2023 by Fredrik Nordvall Forsberg.
Addition satisfies the expected recursive equations (which classically define
addition): zero is the neutral element (this is πβ-right-neutral above), addition
commutes with successors and addition preserves inhabited suprema.
Note that (the index of) the supremum indeed has to be inhabited, because
preserving the empty supremum would give the false equation
Ξ± +β π οΌ π
for any ordinal Ξ±.
\begin{code}
+β-commutes-with-successor : (Ξ± Ξ² : Ordinal π€) β Ξ± +β (Ξ² +β πβ) οΌ (Ξ± +β Ξ²) +β πβ
+β-commutes-with-successor Ξ± Ξ² = (+β-assoc Ξ± Ξ² πβ) β»ΒΉ
open import UF.PropTrunc
open import UF.Size
module _ (pt : propositional-truncations-exist)
(sr : Set-Replacement pt)
where
open import Ordinals.OrdinalOfOrdinalsSuprema ua
open suprema pt sr
open PropositionalTruncation pt
+β-preserves-inhabited-suprema : (Ξ± : Ordinal π€) {I : π€ Μ } (Ξ² : I β Ordinal π€)
β β₯ I β₯
β Ξ± +β sup Ξ² οΌ sup (Ξ» i β Ξ± +β Ξ² i)
+β-preserves-inhabited-suprema Ξ± {I} Ξ² =
β₯β₯-rec (the-type-of-ordinals-is-a-set (ua _) fe')
(Ξ» iβ β β΄-antisym _ _ (βΌ-gives-β΄ _ _ (β¦
1β¦ iβ)) β¦
2β¦)
where
β¦
2β¦ : sup (Ξ» i β Ξ± +β Ξ² i) β΄ (Ξ± +β sup Ξ²)
β¦
2β¦ = sup-is-lower-bound-of-upper-bounds (Ξ» i β Ξ± +β Ξ² i) (Ξ± +β sup Ξ²) β¦
2β¦'
where
β¦
2β¦' : (i : I) β (Ξ± +β Ξ² i) β΄ (Ξ± +β sup Ξ²)
β¦
2β¦' i = +β-right-monotone-β΄ Ξ± (Ξ² i) (sup Ξ²) (sup-is-upper-bound Ξ² i)
β¦
1β¦ : I β (Ξ± +β sup Ξ²) βΌ sup (Ξ» i β Ξ± +β Ξ² i)
β¦
1β¦ iβ _ (inl a , refl) =
transport (_β² sup (Ξ» i β Ξ± +β Ξ² i))
(+β-β-left a)
(β²-β΄-gives-β² (Ξ± β a) (Ξ± +β Ξ² iβ) (sup (Ξ» i β Ξ± +β Ξ² i))
(inl a , +β-β-left a)
(sup-is-upper-bound (Ξ» i β Ξ± +β Ξ² i) iβ))
β¦
1β¦ iβ _ (inr s , refl) =
transport (_β² sup (Ξ» i β Ξ± +β Ξ² i))
(+β-β-right s)
(β₯β₯-rec (β²-is-prop-valued _ _) β¦
1β¦'
(initial-segment-of-sup-is-initial-segment-of-some-component
Ξ² s))
where
β¦
1β¦' : Ξ£ i κ I , Ξ£ b κ β¨ Ξ² i β© , sup Ξ² β s οΌ Ξ² i β b
β (Ξ± +β (sup Ξ² β s)) β² sup (Ξ» i β Ξ± +β Ξ² i)
β¦
1β¦' (i , b , p) =
transportβ»ΒΉ (Ξ» - β (Ξ± +β -) β² sup (Ξ» j β Ξ± +β Ξ² j)) p
(β²-β΄-gives-β² (Ξ± +β (Ξ² i β b)) (Ξ± +β Ξ² i) (sup (Ξ» j β Ξ± +β Ξ² j))
(inr b , +β-β-right b)
(sup-is-upper-bound (Ξ» j β Ξ± +β Ξ² j) i))
\end{code}
Constructively, these equations do not fully characterize ordinal addition, at
least not as far as we know. If addition preserved *all* suprema, then,
expressing the ordinal Ξ² as a supremum via the result given below, we would have
the recursive equation
Ξ± +β Ξ² οΌ Ξ± +β sup (Ξ» b β (B β b) +β πβ)
οΌ sup (Ξ» b β Ξ± +β ((B β b) +β πβ))
οΌ sup (Ξ» b β (Ξ± +β (B β b)) +β πβ)
which would ensure that there is a unique operation satisfying the above
equations for successors and suprema. The problem is that constructively we
cannot, in general, make a case distinction on whether Ξ² is zero or not.
In contrast, multiplication behaves differently and is uniquely characterized by
similar equations since it does preserve all suprema, see
MultiplicationProperties.
Added 14 February 2025 by Tom de Jong.
However, we could reformulate the equations for addition to the classically
equivalent set of equations:
Ξ± +β (Ξ² +β πβ) οΌ (Ξ± +β Ξ²β) +β πβ
Ξ± +β (sup Ξ²) οΌ Ξ± β¨ sup (Ξ» i β Ξ± +β Ξ² i)
for all families Ξ² : I β Ord without any inhabitedness condition on the index
type I.
Note that the equation Ξ± +β πβ = Ξ± follows by taking the empty family in the
supremum equation.
These reformulated equations have the benefit that they uniquely characterize
addition via the recursive equation
Ξ± +β Ξ² οΌ Ξ± +β sup (Ξ» b β (B β b) +β πβ)
οΌ Ξ± β¨ sup (Ξ» b β Ξ± +β ((B β b) +β πβ))
οΌ Ξ± β¨ sup (Ξ» b β (Ξ± +β (B β b)) +β πβ)
which also gives a construction of addition via transfinite recursion.
I first realized this in the context of ordinal exponentiation, cf.
Ordinals.Exponentiation.Specification.
Added 24th May 2024 by Tom de Jong.
Every ordinal is the supremum of the successors of its initial segments.
\begin{code}
supremum-of-successors-of-initial-segments :
(Ξ± : Ordinal π€) β Ξ± οΌ sup (Ξ» x β (Ξ± β x) +β πβ)
supremum-of-successors-of-initial-segments {π€} Ξ± =
Antisymmetry (OO π€) Ξ± s (to-βΌ I) (β΄-gives-βΌ s Ξ± II)
where
s : Ordinal π€
s = sup (Ξ» x β (Ξ± β x) +β πβ)
F : β¨ Ξ± β© β Ordinal π€
F x = (Ξ± β x) +β πβ
I : (a : β¨ Ξ± β©) β (Ξ± β a) β² s
I a = f (inr β) , ((Ξ± β a) οΌβ¨ eβ β©
(F a β inr β) οΌβ¨ eβ β©
(s β f (inr β)) β)
where
f : (y : β¨ F a β©) β β¨ s β©
f = [ F a , s ]β¨ sup-is-upper-bound F a β©
eβ = (successor-lemma-right (Ξ± β a)) β»ΒΉ
eβ = (initial-segment-of-sup-at-component F a (inr β)) β»ΒΉ
II : s β΄ Ξ±
II = sup-is-lower-bound-of-upper-bounds F Ξ±
(upper-bound-of-successors-of-initial-segments Ξ±)
\end{code}
Added 14 July 2024.
We prove that Ξ± +β (sup Ξ²) οΌ Ξ± β¨ sup (Ξ» i β Ξ± +β Ξ² i), but we formulate the RHS
as a single supremum by indexing over π + I instead, sending inl β to Ξ±.
\begin{code}
+β-preserves-suprema-up-to-join
: (Ξ± : Ordinal π€) (I : π€ Μ ) (Ξ² : I β Ordinal π€)
β Ξ± +β sup Ξ² οΌ sup (cases (Ξ» β β Ξ±) (Ξ» i β Ξ± +β Ξ² i))
+β-preserves-suprema-up-to-join {π€} Ξ± I Ξ² =
β΄-antisym (Ξ± +β sup Ξ²) (sup F) β¦
1β¦ β¦
2β¦
where
F : π {π€} + I β Ordinal π€
F = cases (Ξ» _ β Ξ±) (Ξ» i β Ξ± +β Ξ² i)
β¦
1β¦ : Ξ± +β sup Ξ² β΄ sup F
β¦
1β¦ = to-β΄ (Ξ± +β sup Ξ²) (sup F) h
where
h : (z : β¨ Ξ± +β sup Ξ² β©)
β (Ξ± +β sup Ξ²) β z β² sup F
h (inl a) = (s , (Ξ± +β sup Ξ² β inl a οΌβ¨ (+β-β-left a) β»ΒΉ β©
Ξ± β a οΌβ¨ e β©
sup F β s β))
where
s : β¨ sup F β©
s = [ Ξ± , sup F ]β¨ sup-is-upper-bound F (inl β) β© a
e = (initial-segment-of-sup-at-component F (inl β) a) β»ΒΉ
h (inr y) =
β₯β₯-rec
(β²-is-prop-valued (Ξ± +β sup Ξ² β inr y) (sup F))
g
(initial-segment-of-sup-is-initial-segment-of-some-component Ξ² y)
where
g : (Ξ£ i κ I , Ξ£ x κ β¨ Ξ² i β© , sup Ξ² β y οΌ Ξ² i β x)
β Ξ± +β sup Ξ² β inr y β² sup F
g (i , x , e) = s , ((Ξ± +β sup Ξ²) β inr y οΌβ¨ (+β-β-right y) β»ΒΉ β©
Ξ± +β (sup Ξ² β y) οΌβ¨ ap (Ξ± +β_) e β©
Ξ± +β (Ξ² i β x) οΌβ¨ +β-β-right x β©
(Ξ± +β Ξ² i) β inr x οΌβ¨ e' β©
sup F β s β)
where
s : β¨ sup F β©
s = [ F (inr i) , sup F ]β¨ sup-is-upper-bound F (inr i) β© (inr x)
e' = (initial-segment-of-sup-at-component F (inr i) (inr x)) β»ΒΉ
β¦
2β¦ : sup F β΄ Ξ± +β sup Ξ²
β¦
2β¦ = sup-is-lower-bound-of-upper-bounds F (Ξ± +β sup Ξ²) h
where
h : (x : π + I) β F x β΄ Ξ± +β sup Ξ²
h (inl β) = +β-left-β΄ Ξ± (sup Ξ²)
h (inr i) = +β-right-monotone-β΄ Ξ± (Ξ² i) (sup Ξ²) (sup-is-upper-bound Ξ² i)
\end{code}
Added 22 July 2025 by Tom de Jong.
The above, together with +β-commutes-with-successor, uniquely determines
addition of ordinals, see also the comment dated 14 February 2025.
\begin{code}
private
successor-equation : (Ordinal π€ β Ordinal π€ β Ordinal π€) β π€ βΊ Μ
successor-equation {π€} _β_ =
(Ξ± Ξ² : Ordinal π€) β Ξ± β (Ξ² +β πβ) οΌ (Ξ± β Ξ²) +β πβ
suprema-equation : (Ordinal π€ β Ordinal π€ β Ordinal π€) β π€ βΊ Μ
suprema-equation {π€} _β_ =
(Ξ± : Ordinal π€) (I : π€ Μ ) (Ξ² : I β Ordinal π€)
β Ξ± β (sup Ξ²) οΌ sup (cases (Ξ» (_ : π{π€}) β Ξ±) (Ξ» i β Ξ± β Ξ² i))
recursive-equation : (Ordinal π€ β Ordinal π€ β Ordinal π€) β π€ βΊ Μ
recursive-equation {π€} _β_ =
(Ξ± Ξ² : Ordinal π€)
β Ξ± β Ξ² οΌ sup (cases (Ξ» (_ : π{π€}) β Ξ±) (Ξ» b β (Ξ± β (Ξ² β b)) +β πβ))
successor-and-suprema-equations-give-recursive-equation
: (_β_ : Ordinal π€ β Ordinal π€ β Ordinal π€)
β successor-equation _β_
β suprema-equation _β_
β recursive-equation _β_
successor-and-suprema-equations-give-recursive-equation
{π€} _β_ β-succ β-sup Ξ± Ξ² =
Ξ± β Ξ² οΌβ¨ I β©
(Ξ± β sup (Ξ» b β (Ξ² β b) +β πβ)) οΌβ¨ II β©
sup (cases (Ξ» β β Ξ±) (Ξ» b β Ξ± β ((Ξ² β b) +β πβ))) οΌβ¨ III β©
sup (cases (Ξ» β β Ξ±) (Ξ» b β (Ξ± β (Ξ² β b)) +β πβ)) β
where
I = ap (Ξ± β_) (supremum-of-successors-of-initial-segments Ξ²)
II = β-sup Ξ± β¨ Ξ² β© (Ξ» b β (Ξ² β b) +β πβ)
III = ap (Ξ» - β sup (cases (Ξ» (_ : π{π€}) β Ξ±) -))
(dfunext fe' (Ξ» b β β-succ Ξ± (Ξ² β b)))
+β-recursive-equation : recursive-equation {π€} _+β_
+β-recursive-equation =
successor-and-suprema-equations-give-recursive-equation
_+β_ +β-commutes-with-successor +β-preserves-suprema-up-to-join
+β-is-uniquely-specified'
: (_β_ : Ordinal π€ β Ordinal π€ β Ordinal π€)
β recursive-equation _β_
β (Ξ± Ξ² : Ordinal π€) β Ξ± β Ξ² οΌ Ξ± +β Ξ²
+β-is-uniquely-specified' {π€} _β_ β-rec Ξ± =
transfinite-induction-on-OO (Ξ» - β (Ξ± β -) οΌ (Ξ± +β -)) I
where
I : (Ξ² : Ordinal π€)
β ((b : β¨ Ξ² β©) β (Ξ± β (Ξ² β b)) οΌ (Ξ± +β (Ξ² β b)))
β (Ξ± β Ξ²) οΌ (Ξ± +β Ξ²)
I Ξ² IH = Ξ± β Ξ² οΌβ¨ II β©
sup (cases (Ξ» β β Ξ±) (Ξ» b β (Ξ± β (Ξ² β b)) +β πβ)) οΌβ¨ III β©
sup (cases (Ξ» β β Ξ±) (Ξ» b β (Ξ± +β (Ξ² β b)) +β πβ)) οΌβ¨ IV β©
Ξ± +β Ξ² β
where
II = β-rec Ξ± Ξ²
III = ap (Ξ» - β sup (cases (Ξ» (_ : π{π€}) β Ξ±) -))
(dfunext fe' (Ξ» b β ap (_+β πβ) (IH b)))
IV = +β-recursive-equation Ξ± Ξ² β»ΒΉ
+β-is-uniquely-specified
: β! _β_ κ (Ordinal π€ β Ordinal π€ β Ordinal π€) ,
(successor-equation _β_) Γ (suprema-equation _β_)
+β-is-uniquely-specified {π€} =
(_+β_ , (+β-commutes-with-successor , +β-preserves-suprema-up-to-join)) ,
(Ξ» (_β_ , β-succ , β-sup) β
to-subtype-οΌ
(Ξ» F β Γ-is-prop (Ξ β-is-prop fe'
(Ξ» _ _ β underlying-type-is-set fe (OO π€)))
(Ξ β-is-prop fe'
(Ξ» _ _ _ β underlying-type-is-set fe (OO π€))))
(dfunext fe'
(Ξ» Ξ± β dfunext fe'
(Ξ» Ξ² β
(+β-is-uniquely-specified' _β_
(successor-and-suprema-equations-give-recursive-equation
_β_ β-succ β-sup)
Ξ± Ξ²) β»ΒΉ))))
\end{code}
Added 2 June 2024 by Tom de Jong.
\begin{code}
no-greatest-ordinal : Β¬ (Ξ£ Ξ± κ Ordinal π€ , ((Ξ² : Ordinal π€) β Ξ² β΄ Ξ±))
no-greatest-ordinal {π€} (Ξ± , Ξ±-greatest) = irrefl (OO π€) Ξ± IV
where
I : (Ξ± +β πβ) β΄ Ξ±
I = Ξ±-greatest (Ξ± +β πβ)
II : Ξ± β΄ (Ξ± +β πβ)
II = β²-gives-β΄ Ξ± (Ξ± +β πβ) (successor-increasing Ξ±)
III : Ξ± +β πβ οΌ Ξ±
III = β΄-antisym (Ξ± +β πβ) Ξ± I II
IV : Ξ± β² Ξ±
IV = transport (Ξ± β²_) III (successor-increasing Ξ±)
\end{code}
Added 15 July 2025 by Tom de Jong after discussions with Nicolai Kraus, Fredrik
Nordvall Forsberg and Chuangjie Xu a year earlier.
\begin{code}
+β-as-large-as-right-summand-implies-EM : ((Ξ± Ξ² : Ordinal π€) β Ξ² β΄ Ξ± +β Ξ²)
β EM π€
+β-as-large-as-right-summand-implies-EM hyp P P-is-prop = IV
where
Ξ± = prop-ordinal P P-is-prop
Ξ² = πβ
π : Ξ² β΄ Ξ± +β Ξ²
π = hyp Ξ± Ξ²
f = [ Ξ² , Ξ± +β Ξ² ]β¨ π β©
I : (p : P) β f β οΌ inl p β P
I p _ = p
II : (p : P) β f β οΌ inl p
II p = simulations-preserve-least Ξ² (Ξ± +β Ξ²) β (inl p) f
[ Ξ² , Ξ± +β Ξ² ]β¨ π β©-is-simulation
πβ-least
l
where
l : is-least (Ξ± +β Ξ²) (inl p)
l = minimal-is-least (Ξ± +β Ξ²) (inl p) m
where
m : is-minimal (Ξ± +β Ξ²) (inl p)
m (inl p') = π-elim
m (inr β ) = π-elim
III : f β οΌ inr β β Β¬ P
III e p = +disjoint ((II p) β»ΒΉ β e)
IV : P + Β¬ P
IV = equality-cases (f β) (Ξ» p β inl β I p) (Ξ» _ β inr β III)
EM-implies-+β-as-large-as-right-summand : EM π€
β ((Ξ± Ξ² : Ordinal π€) β Ξ² β΄ Ξ± +β Ξ²)
EM-implies-+β-as-large-as-right-summand em Ξ± Ξ² =
βΌ-gives-β΄ Ξ² (Ξ± +β Ξ²)
(EM-implies-order-preserving-gives-βΌ em Ξ² (Ξ± +β Ξ²) (f , I))
where
f : β¨ Ξ² β© β β¨ Ξ± +β Ξ² β©
f = inr
I : is-order-preserving Ξ² (Ξ± +β Ξ²) f
I y y' l = l
\end{code}
Added 15 July 2025 by Tom de Jong.
\begin{code}
+β-minimal : (Ξ± Ξ² : Ordinal π€) (aβ : β¨ Ξ± β©)
β is-minimal Ξ± aβ β is-minimal (Ξ± +β Ξ²) (inl aβ)
+β-minimal Ξ± Ξ² aβ aβ-minimal (inl a) = aβ-minimal a
+β-minimal Ξ± Ξ² aβ aβ-minimal (inr b) = π-elim
+β-least : (Ξ± Ξ² : Ordinal π€) (aβ : β¨ Ξ± β©)
β is-least Ξ± aβ β is-least (Ξ± +β Ξ²) (inl aβ)
+β-least Ξ± Ξ² aβ aβ-least =
minimal-is-least (Ξ± +β Ξ²) (inl aβ)
(+β-minimal Ξ± Ξ² aβ (least-is-minimal Ξ± aβ aβ-least))
\end{code}
Added 26 September 2025 by Fredrik Nordvall Forsberg.
\begin{code}
πβ-is-not-πβ : πβ {π€} β πβ {π€}
πβ-is-not-πβ p = πβ-is-not-πβ (+β-left-cancellable πβ πβ πβ p' β»ΒΉ)
where
p' : πβ οΌ πβ +β πβ
p' = p β πβ-right-neutral πβ β»ΒΉ
\end{code}
Added in September 2025 by Fredrik Nordvall Forsberg.
Moved here from ArithmeticReflection by Tom de Jong in October 2025.
Some special cases of addition by Ο.
\begin{code}
πβ+βΟ-is-Ο : πβ +β Ο οΌ Ο
πβ+βΟ-is-Ο = πβ-left-neutral Ο
πβ+βΟ-is-Ο : πβ +β Ο οΌ Ο
πβ+βΟ-is-Ο = eqtoidβ (ua _) fe' (πβ +β Ο) Ο h
where
f : π + β β β
f (inl β) = 0
f (inr n) = succ n
g : β β π + β
g 0 = inl β
g (succ n) = inr n
f-equiv : is-equiv f
f-equiv = qinvs-are-equivs f (g , (Ξ· , Ο΅))
where
Ξ· : (Ξ» x β g (f x)) βΌ id
Ξ· (inl β) = refl
Ξ· (inr n) = refl
Ο΅ : (Ξ» x β f (g x)) βΌ id
Ο΅ zero = refl
Ο΅ (succ x) = refl
f-preserves-order : (x y : π + β) β x βΊβ¨ πβ +β Ο β© y β f x βΊβ¨ Ο β© f y
f-preserves-order (inl β) (inr n) p = β
f-preserves-order (inr n) (inr m) p = p
f-reflects-order : (x y : π + β) β f x βΊβ¨ Ο β© f y β x βΊβ¨ πβ +β Ο β© y
f-reflects-order (inl β) (inr n) _ = β
f-reflects-order (inr n) (inr m) p = p
h : (πβ +β Ο) ββ Ο
h = f , order-preserving-reflecting-equivs-are-order-equivs (πβ +β Ο) Ο f
f-equiv f-preserves-order f-reflects-order
\end{code}