Martin Escardo & Tom de Jong, June 2023. Iterative ordinals. We define the type of iterative ordinals as a subtype of that of iterative sets, which in turn, is defined a subtype of that of iterative multisets. Iterative ordinals are defined in the same way as in the constructive and non-constructive set theories CZF and ZFC, following Powell [7], as transitive sets whose members are also transitive. The main theorem in this module is that the iterative ordinals coincide with the HoTT-book ordinals. This builds on [5] Tom de Jong, Nicolai Kraus, Fredrik Nordvall Forsberg and Chuangjie Xu. "Set-Theoretic and Type-Theoretic Ordinals Coincide". To appear at LICS 2023, June 2023. https://arxiv.org/abs/2301.10696 The difference is that instead of using the HoTT-book higher-inductive definition of the cumulative hierarchy, we use Gylterud's construction. [4] H. R. Gylterud, "From multisets to sets in homotopy type theory". The Journal of Symbolic Logic, vol. 83, no. 3, pp. 1132β1146, 2018. https://doi.org/10.1017/jsl.2017.84 See the module Iterative.index for more bibliographic references, which uses the same numbering as above. \begin{code} {-# OPTIONS --safe --without-K --lossy-unification #-} open import MLTT.Spartan open import UF.Univalence module Iterative.Ordinals (ua : Univalence) (π€ : Universe) where open import UF.FunExt open import UF.UA-FunExt private π€βΊ : Universe π€βΊ = π€ βΊ fe : Fun-Ext fe = Univalence-gives-Fun-Ext ua fe' : FunExt fe' π€ π₯ = fe {π€} {π₯} open import Iterative.Multisets π€ open import Iterative.Sets ua π€ open import Ordinals.Equivalence open import Ordinals.Notions open import Ordinals.OrdinalOfOrdinals ua open import Ordinals.Type hiding (Ord) open import Ordinals.Underlying open import Ordinals.WellOrderTransport open import UF.Base open import UF.Embeddings open import UF.Equiv open import UF.EquivalenceExamples open import UF.PairFun open import UF.Sets open import UF.Size open import UF.Subsingletons open import UF.Subsingletons-FunExt open import W.Type \end{code} An iterative ordinal is a transitive iterative set whose members are also transitive. This definition is due to Powell [7] in the context of ZF. Notice that it is not by induction. (von Neumann's inductive definition is that a set is an ordinal if it is a well-ordered set of all smaller ordinals.) \begin{code} is-transitive-iset : π β π€βΊ Μ is-transitive-iset A = (B C : π) β B β A β C β B β C β A has-transitive-members : π β π€βΊ Μ has-transitive-members A = (B : π) β B β A β is-transitive-iset B being-transitive-iset-is-prop : (A : π) β is-prop (is-transitive-iset A) being-transitive-iset-is-prop A = Ξ β-is-prop fe (Ξ» B C l m β β-is-prop-valued C A) having-transitive-members-is-prop : (A : π) β is-prop (has-transitive-members A) having-transitive-members-is-prop A = Ξ β-is-prop fe (Ξ» B l β being-transitive-iset-is-prop B) is-iterative-ordinal : π β π€βΊ Μ is-iterative-ordinal A = is-transitive-iset A Γ has-transitive-members A being-iordinal-is-prop : (A : π) β is-prop (is-iterative-ordinal A) being-iordinal-is-prop A = Γ-is-prop (being-transitive-iset-is-prop A) (having-transitive-members-is-prop A) \end{code} We name the projections for the sake of clarity: \begin{code} iordinals-are-transitive : (A : π) β is-iterative-ordinal A β is-transitive-iset A iordinals-are-transitive A = prβ members-of-iordinals-are-transitive : (A : π) β is-iterative-ordinal A β has-transitive-members A members-of-iordinals-are-transitive A = prβ \end{code} Although the definition of iterative ordinal is not by induction, it does satisfy the following inductive characterization, which we use a number of times. \begin{code} ordinal-is-hereditary : (A B : π) β B β A β is-iterative-ordinal A β is-iterative-ordinal B ordinal-is-hereditary A B B-in-A (A-trans , A-members-trans) = III where I : is-transitive-iset B I = A-members-trans B B-in-A II : (C : π) β C β B β is-transitive-iset C II C C-in-B = A-members-trans C (A-trans B C B-in-A C-in-B) III : is-iterative-ordinal B III = I , II \end{code} The type of iterative ordinals is defined as a subtype of that of iterative sets. \begin{code} π : π€βΊ Μ π = Ξ£ A κ π , is-iterative-ordinal A π-is-locally-small : is-locally-small π π-is-locally-small = subtype-is-locally-small being-iordinal-is-prop π-is-locally-small \end{code} We again name the projections for the sake of clarity: \begin{code} underlying-iset : π β π underlying-iset = prβ underlying-iset-is-iordinal : (Ξ± : π) β is-iterative-ordinal (underlying-iset Ξ±) underlying-iset-is-iordinal = prβ \end{code} Because the notion of iterative ordinal is property, we get that π is a subtype of π. \begin{code} underlying-iset-is-embedding : is-embedding underlying-iset underlying-iset-is-embedding = prβ-is-embedding being-iordinal-is-prop π-is-set : is-set π π-is-set = subtypes-of-sets-are-sets underlying-iset underlying-iset-is-embedding π-is-set \end{code} We define the less-than relation on ordinals to be the membership relation, as in material set theory under von Newmann's encoding: \begin{code} _<_ : π β π β π€βΊ Μ Ξ± < Ξ² = underlying-iset Ξ± β underlying-iset Ξ² \end{code} As is the case for iterative sets, there is a resized down, equivalent definition of the less-than relation on ordinals: \begin{code} _<β»_ : π β π β π€ Μ Ξ± <β» Ξ² = underlying-iset Ξ± ββ» underlying-iset Ξ² <β»β-< : (Ξ± Ξ² : π) β (Ξ± < Ξ²) β (Ξ± <β» Ξ²) <β»β-< Ξ± Ξ² = ββ»ββ (underlying-iset Ξ±) (underlying-iset Ξ²) <-is-prop-valued : (Ξ± Ξ² : π) β is-prop (Ξ± < Ξ²) <-is-prop-valued Ξ± Ξ² = β-is-prop-valued (underlying-iset Ξ±) (underlying-iset Ξ²) \end{code} We now show that this order is transitive, extensional and accessible, so that π is an ordinal in the sense of the HoTT book. This result is in de Jong, Kraus, Nordvall Forsberg and Xu [5]. As discussed above, the difference is that [5] uses the higher-inductive definition of π. \begin{code} <-is-transitive : (Ξ± Ξ² Ξ³ : π) β Ξ± < Ξ² β Ξ² < Ξ³ β Ξ± < Ξ³ <-is-transitive (A , _) (B , _) (C , C-trans , _) u v = I where I : A β C I = C-trans B A v u _β€_ : π β π β π€βΊ Μ Ξ± β€ Ξ² = β Ξ³ β Ξ³ < Ξ± β Ξ³ < Ξ² β-gives-β€ : (Ξ± Ξ² : π) β underlying-iset Ξ± β underlying-iset Ξ² β Ξ± β€ Ξ² β-gives-β€ Ξ± Ξ² u Ξ³ = u (underlying-iset Ξ³) β€-gives-β : (Ξ± Ξ² : π) β Ξ± β€ Ξ² β underlying-iset Ξ± β underlying-iset Ξ² β€-gives-β Ξ±@(A , A-is-iord) Ξ²@(B , _) u = I where I : A β B I C C-in-A = Iβ where C-is-iord : is-iterative-ordinal C C-is-iord = ordinal-is-hereditary A C C-in-A A-is-iord Iβ : is-transitive-iset C Iβ = iordinals-are-transitive C C-is-iord Iβ : (B : π) β B β C β is-transitive-iset B Iβ = members-of-iordinals-are-transitive C C-is-iord Iβ : C β B Iβ = u (C , Iβ , Iβ) C-in-A \end{code} We now define root and forest "destructors" for the type π. This will give rise to an inductive characterization of π which doesn't mention π or π, which seems to be new. \begin{code} π-root : π β π€ Μ π-root Ξ± = π-root (underlying-iset Ξ±) π-forest : (Ξ± : π) β π-root Ξ± β π π-forest Ξ± x = π-forest A x , ordinal-is-hereditary A (π-forest A x) (π-forest-β A x) (underlying-iset-is-iordinal Ξ±) where A = underlying-iset Ξ± \end{code} By definition, any (immediate) subtree of Ξ± is less than Ξ±: \begin{code} π-forest-< : (Ξ± : π) (x : π-root Ξ±) β π-forest Ξ± x < Ξ± π-forest-< Ξ± = π-forest-β (underlying-iset Ξ±) \end{code} The π-forest map is an embedding, which follows from the fact that the π-forest map is an embedding: \begin{code} π-forest-is-embedding : (Ξ± : π) β is-embedding (π-forest Ξ±) π-forest-is-embedding Ξ±@(A@(ssup _ _ , _) , _) = pair-fun-is-embedding-special (underlying-iset β π-forest Ξ±) (underlying-iset-is-iordinal β π-forest Ξ±) (π-forest-is-embedding A) being-iordinal-is-prop \end{code} The antisymmetry of the β€ relation, which amounts to the extensionality condition: \begin{code} β€-is-antisymmetric : (Ξ± Ξ² : π) β Ξ± β€ Ξ² β Ξ² β€ Ξ± β Ξ± οΌ Ξ² β€-is-antisymmetric Ξ±@(A , _) Ξ²@(B , _) u v = II where I : A οΌ B I = β-is-extensional A B (β€-gives-β Ξ± Ξ² u) (β€-gives-β Ξ² Ξ± v) II : Ξ± οΌ Ξ² II = to-subtype-οΌ (being-iordinal-is-prop) I <-is-extensional : is-extensional _<_ <-is-extensional = β€-is-antisymmetric \end{code} A characterization of the < relation: \begin{code} <-behaviour : (Ξ± Ξ² : π) β (Ξ± < Ξ²) β (Ξ£ y κ π-root Ξ² , π-forest Ξ² y οΌ Ξ±) <-behaviour Ξ±@(A@(M , _) , _) Ξ²@(B@(N@(ssup Y Ξ³) , _) , _) = II where I : (y : Y) β (Ξ³ y οΌ M) β (π-forest Ξ² y οΌ Ξ±) I y = (Ξ³ y οΌ M) ββ¨ Iβ β© (π-forest B y οΌ A) ββ¨ Iβ β© (π-forest Ξ² y οΌ Ξ±) β where Iβ = embedding-criterion-converse underlying-mset underlying-mset-is-embedding (π-forest B y) A Iβ = embedding-criterion-converse underlying-iset underlying-iset-is-embedding (π-forest Ξ² y) Ξ± II : (Ξ£ y κ Y , Ξ³ y οΌ M) β (Ξ£ y κ Y , π-forest Ξ² y οΌ Ξ±) II = Ξ£-cong I \end{code} The π-forest map is lower closed: \begin{code} is-lower-closed : {X : π€ Μ } β (X β π) β π€βΊ Μ is-lower-closed {X} Ο = (Ξ² : π) (x : X) β Ξ² < Ο x β Ξ£ y κ X , Ο y οΌ Ξ² being-lower-closed-is-prop : {X : π€ Μ } (Ο : X β π) β is-embedding Ο β is-prop (is-lower-closed Ο) being-lower-closed-is-prop Ο e = Ξ β-is-prop fe (Ξ» Ξ² _ _ β e Ξ²) π-forest-is-lower-closed : (Ξ± : π) β is-lower-closed (π-forest Ξ±) π-forest-is-lower-closed Ξ± Ξ² x l = VII where have-l : Ξ² < π-forest Ξ± x have-l = l I : π-forest Ξ± x < Ξ± I = π-forest-< Ξ± x II : Ξ² < Ξ± II = <-is-transitive Ξ² (π-forest Ξ± x) Ξ± l I VII : Ξ£ y κ π-root Ξ± , π-forest Ξ± y οΌ Ξ² VII = β <-behaviour Ξ² Ξ± β II \end{code} The "constructor" of elements of π: \begin{code} π-ssup : (X : π€ Μ ) (Ο : X β π) β is-embedding Ο β is-lower-closed Ο β π π-ssup X Ο Ο-emb Ο-lower = A , A-is-iord where Ο : X β π Ο = underlying-iset β Ο Ο-iter : (x : X) β is-iterative-ordinal (Ο x) Ο-iter = underlying-iset-is-iordinal β Ο Ο-emb : is-embedding Ο Ο-emb = β-is-embedding Ο-emb underlying-iset-is-embedding A : π A = π-ssup X Ο Ο-emb A-behaviour : (B : π) β B β A β (Ξ£ x κ X , Ο x οΌ B) A-behaviour B = β-behaviour B X Ο Ο-emb I : (B : π) β B β A β is-iterative-ordinal B I B B-in-A = Iβ where Iβ : Ξ£ x κ X , Ο x οΌ B Iβ = β A-behaviour B β B-in-A Iβ : is-iterative-ordinal B Iβ = transport is-iterative-ordinal (prβ Iβ) (Ο-iter (prβ Iβ)) II : (B C : π) β B β A β C β B β C β A II B C B-in-A C-in-B = IIβ where x : X x = prβ (β A-behaviour B β B-in-A) p : Ο x οΌ B p = prβ (β A-behaviour B β B-in-A) Ξ² Ξ³ : π Ξ² = (B , I B B-in-A) Ξ³ = (C , ordinal-is-hereditary B C C-in-B (I B B-in-A)) IIβ : Ξ³ < Ξ² IIβ = C-in-B q : Ο x οΌ Ξ² q = embeddings-are-lc underlying-iset underlying-iset-is-embedding p IIβ : Ξ³ < Ο x IIβ = transport (Ξ³ <_) (q β»ΒΉ) IIβ IIβ : Ξ£ y κ X , Ο y οΌ Ξ³ IIβ = Ο-lower Ξ³ x IIβ IIβ : type-of IIβ β Ξ£ y κ X , Ο y οΌ C IIβ (y , p) = y , ap underlying-iset p IIβ : Ξ£ x κ X , underlying-mset (Ο x) οΌ underlying-mset C IIβ = β A-behaviour C ββ»ΒΉ (IIβ IIβ) IIβ : C β A IIβ = IIβ III : (B : π) β B β A β is-transitive-iset B III B m = iordinals-are-transitive B (I B m) A-is-iord : is-iterative-ordinal A A-is-iord = II , III \end{code} It satisfies the expected properties with respect to the destructors: \begin{code} π-ssup-root : (X : π€ Μ ) (Ο : X β π) (e : is-embedding Ο) (l : is-lower-closed Ο) β π-root (π-ssup X Ο e l) οΌ X π-ssup-root X Ο e l = refl π-ssup-forest : (X : π€ Μ ) (Ο : X β π) (e : is-embedding Ο) (l : is-lower-closed Ο) β π-forest (π-ssup X Ο e l) βΌ Ο π-ssup-forest X Ο e l x = to-subtype-οΌ being-iordinal-is-prop refl π-Ξ·' : (Ξ± : π) (e : is-embedding (π-forest Ξ±)) (l : is-lower-closed (π-forest Ξ±)) β π-ssup (π-root Ξ±) (π-forest Ξ±) e l οΌ Ξ± π-Ξ·' (A@(ssup _ _ , _) , _) _ _ = to-subtype-οΌ being-iordinal-is-prop (π-Ξ·' A _) π-Ξ· : (Ξ± : π) β π-ssup (π-root Ξ±) (π-forest Ξ±) (π-forest-is-embedding Ξ±) (π-forest-is-lower-closed Ξ±) οΌ Ξ± π-Ξ· A = π-Ξ·' A (π-forest-is-embedding A) (π-forest-is-lower-closed A) \end{code} A criterion for equality in π, which is much less general than we can have, but enough for our purposes for the moment: \begin{code} to-π-οΌ-special : (X : π€ Μ ) (Ο Ο' : X β π) (e : is-embedding Ο ) (l : is-lower-closed Ο ) (e' : is-embedding Ο') (l' : is-lower-closed Ο') β Ο οΌ Ο' β π-ssup X Ο e l οΌ π-ssup X Ο' e' l' to-π-οΌ-special X Ο Ο' e l e' l' refl = to-subtype-οΌ being-iordinal-is-prop (to-subtype-οΌ being-iset-is-prop refl) \end{code} We now justify our notation "ssup" in comparison with the more traditional notation "sup" for the constructors. π-ssup X Ο e l is the unique iterative ordinal whose predecessors are precisely the members of the family Ο, which is known as the strict supremum (or successor supremum, or strong supremum) of Ο, and is also its rank in the sense of set theory. \begin{code} π-ssup-behaviour : (X : π€ Μ ) (Ο : X β π) (e : is-embedding Ο) (l : is-lower-closed Ο) (Ξ± : π) β (Ξ± < π-ssup X Ο e l) β (Ξ£ x κ X , Ο x οΌ Ξ±) π-ssup-behaviour X Ο e l Ξ± = (Ξ± < π-ssup X Ο e l) ββ¨ I β© (Ξ£ x κ X , π-forest (π-ssup X Ο e l) x οΌ Ξ±) ββ¨ II β© (Ξ£ x κ X , Ο x οΌ Ξ±) β where I = <-behaviour Ξ± (π-ssup X Ο e l) II = Ξ£-cong (Ξ» x β οΌ-cong-l _ _ (π-ssup-forest X Ο e l x)) \end{code} We now discuss various equivalent induction principles on π. \begin{code} π-induction' : (P : π β π₯ Μ ) β ((Ξ± : π) β ((x : π-root Ξ±) β P (π-forest Ξ± x)) β P Ξ±) β (Ξ± : π) β P Ξ± π-induction' P f ((M , M-is-iset) , M-is-iord) = h M M-is-iset M-is-iord where h : (M : π) (M-is-iset : is-iterative-set M) (M-is-iord : is-iterative-ordinal (M , M-is-iset)) β P ((M , M-is-iset) , M-is-iord) h M@(ssup X Ο) M-is-iset@(Ο-emb , Ο-iter) M-is-iord = I where Ξ± : π Ξ± = (M , M-is-iset) , M-is-iord IH : (x : X) β P (π-forest Ξ± x) IH x = h (Ο x) (Ο-iter x) (ordinal-is-hereditary (M , M-is-iset) (Ο x , Ο-iter x) (π-forest-β M x) M-is-iord) I : P Ξ± I = f Ξ± IH π-recursion' : (P : π₯ Μ ) β ((Ξ± : π) β (π-root Ξ± β P) β P) β π β P π-recursion' P = π-induction' (Ξ» _ β P) \end{code} TODO. Do things get nicer if use use induction on π rather than π in the above proof? It would be nice if we could define π inductively as follows: data π : π€βΊ Μ where ssup : (X : π€ Μ ) (Ο : X β π) β is-embedding Ο β is-lower-closed Ο β π However, this is not a strictly positive definition, for the criterion of strict positivity adopted by Agda, and so it is not accepted. Nevertheless, all iterative ordinals *are* generated by the "constructor" π-ssup, in the following sense, so that we can view π as really inductively defined by the above data declaration, which seems to be a new result. This is close to von Neumann's definition, with the difference that we don't require Ο to be well-ordered - this follows automatically. \begin{code} π-induction : (P : π β π₯ Μ ) β ((X : π€ Μ ) (Ο : X β π) (e : is-embedding Ο) (l : is-lower-closed Ο) β ((x : X) β P (Ο x)) β P (π-ssup X Ο e l)) β (Ξ± : π) β P Ξ± π-induction P f = π-induction' P f' where f' : (Ξ± : π) β ((x : π-root Ξ±) β P (π-forest Ξ± x)) β P Ξ± f' Ξ± IH = transport P (π-Ξ· Ξ±) I where I : P (π-ssup (π-root Ξ±) (π-forest Ξ±) (π-forest-is-embedding Ξ±) (π-forest-is-lower-closed Ξ±)) I = f (π-root Ξ±) (π-forest Ξ±) (π-forest-is-embedding Ξ±) (π-forest-is-lower-closed Ξ±) IH π-recursion : (P : π₯ Μ ) β ((X : π€ Μ ) (Ο : X β π) β is-embedding Ο β is-lower-closed Ο β (X β P) β P) β π β P π-recursion P f = π-induction (Ξ» _ β P) (Ξ» X Ο e l s β f X Ο e l s) π-iteration : (P : π₯ Μ ) β ((X : π€ Μ ) β (X β P) β P) β π β P π-iteration P f = π-induction (Ξ» _ β P) (Ξ» X Ο e l β f X) \end{code} The usual induction principle for ordinals follows directly from the above form of induction. \begin{code} <-induction : (P : π β π₯ Μ ) β ((Ξ± : π) β ((Ξ² : π) β Ξ² < Ξ± β P Ξ²) β P Ξ±) β (Ξ± : π) β P Ξ± <-induction P IH = π-induction P f where f : (X : π€ Μ ) (Ο : X β π) (e : is-embedding Ο) (l : is-lower-closed Ο) β ((x : X) β P (Ο x)) β P (π-ssup X Ο e l) f X Ο e l u = IH Ξ± s where Ξ± : π Ξ± = π-ssup X Ο e l s : (Ξ² : π) β Ξ² < Ξ± β P Ξ² s Ξ²@((.(underlying-mset (underlying-iset (Ο x))) , is) , io) (x , refl) = III where I : P (Ο x) I = u x II : Ο x οΌ Ξ² II = to-subtype-οΌ being-iordinal-is-prop (to-subtype-οΌ being-iset-is-prop refl) III : P Ξ² III = transport P II I \end{code} Which in turn gives the accessibility of the order directly: \begin{code} <-is-accessible : (Ξ± : π) β is-accessible _<_ Ξ± <-is-accessible = <-induction (is-accessible _<_) (Ξ» _ β acc) \end{code} Putting the above together we conclude that the type of iterative ordinals has the structure of an ordinal in the sense of the HoTT book, which, as mentioned above, had already been proved in [5]. \begin{code} π : Ordinal π€βΊ π = π , _<_ , <-is-prop-valued , <-is-accessible , <-is-extensional , <-is-transitive \end{code} We now want to show that π is equivalent to the "ordinal of ordinals" in the sense of the HoTT book. Every iterative ordinal can be mapped to a HoTT-book ordinal, by taking the root of the iterative ordinal to be the underlying set of the HoTT-book ordinal. This is similar to [5, Proposition 43], with the difference that we don't need to quotient because the forest is an embedding in our conception of iterative ordinals. There is also a minor improvement in avoiding propositional truncations in the definition of the order on the root of Ξ±, which, together with the avoidance of quotients, makes the proof more direct. \begin{code} Ord : π€βΊ Μ Ord = Ordinal π€ π-to-Ord : π β Ord π-to-Ord Ξ± = Ξ±' where X : π€ Μ X = π-root Ξ± _βΊ_ : X β X β π€βΊ Μ x βΊ y = π-forest Ξ± x < π-forest Ξ± y _β_ : X β X β π€βΊ Μ x β y = β z β z βΊ x β z βΊ y β-gives-β€ : (x y : X) β x β y β π-forest Ξ± x β€ π-forest Ξ± y β-gives-β€ x y l Ξ² m = IV where I : Ξ£ z κ X , π-forest Ξ± z οΌ Ξ² I = π-forest-is-lower-closed Ξ± Ξ² x m II : prβ I βΊ x II = transportβ»ΒΉ (_< π-forest Ξ± x) (prβ I) m III : prβ I βΊ y III = l (prβ I) II IV : Ξ² < (π-forest Ξ± y) IV = transport (_< (π-forest Ξ± y)) (prβ I) III β€-gives-β : (x y : X) β π-forest Ξ± x β€ π-forest Ξ± y β x β y β€-gives-β x y l z = l (π-forest Ξ± z) βΊ-is-prop-valued : (x y : X) β is-prop (x βΊ y) βΊ-is-prop-valued x y = <-is-prop-valued (π-forest Ξ± x) (π-forest Ξ± y) βΊ-is-accessible : (x : X) β is-accessible _βΊ_ x βΊ-is-accessible x = f x (<-is-accessible (π-forest Ξ± x)) where f : (x : X) β is-accessible _<_ (π-forest Ξ± x) β is-accessible _βΊ_ x f x (acc u) = acc (Ξ» y l β f y (u (π-forest Ξ± y) l)) βΊ-is-extensional : is-extensional _βΊ_ βΊ-is-extensional x y u v = II where I : π-forest Ξ± x οΌ π-forest Ξ± y I = <-is-extensional _ _ (β-gives-β€ x y u) (β-gives-β€ y x v) II : x οΌ y II = embeddings-are-lc (π-forest Ξ±) (π-forest-is-embedding Ξ±) I βΊ-is-transitive : is-transitive _βΊ_ βΊ-is-transitive x y z = <-is-transitive (π-forest Ξ± x) (π-forest Ξ± y) (π-forest Ξ± z) βΊ-is-well-order : is-well-order _βΊ_ βΊ-is-well-order = βΊ-is-prop-valued , βΊ-is-accessible , βΊ-is-extensional , βΊ-is-transitive _βΊβ»_ : X β X β π€ Μ x βΊβ» y = (π-forest Ξ± x) <β» (π-forest Ξ± y) βΊβ»β-βΊ : (x y : X) β (x βΊ y) β (x βΊβ» y) βΊβ»β-βΊ x y = <β»β-< (π-forest Ξ± x) (π-forest Ξ± y) βΊβ»-is-well-order : is-well-order _βΊβ»_ βΊβ»-is-well-order = order-transfer-lemmaβ.well-orderβ fe' X _βΊβ»_ _βΊ_ (Ξ» x y β β-sym (βΊβ»β-βΊ x y)) βΊ-is-well-order Ξ±' : Ord Ξ±' = X , _βΊβ»_ , βΊβ»-is-well-order \end{code} The order of π-to-Ord Ξ± is characterized as follows in terms of the order of π, by definition: \begin{code} π-to-Ord-order : (Ξ± : π) (x y : β¨ π-to-Ord Ξ± β©) β (π-forest Ξ± x < π-forest Ξ± y) β (x βΊβ¨ π-to-Ord Ξ± β© y) π-to-Ord-order Ξ± x y = <β»β-< (π-forest Ξ± x) (π-forest Ξ± y) \end{code} We now define the map in the other direction, essentially in the same way as in [5], in several steps, but we start mapping into π rather than into π directly. In any case, π doesn't occur in the set-up of [5]. \begin{code} Ord-to-π : Ord β π Ord-to-π = transfinite-recursion-on-OO π (Ξ» Ξ± β ssup β¨ Ξ± β©) \end{code} This is characterized by the following recursive definition, where Ξ± β x denotes the sub-ordinal of Ξ± consisting of the elements below x. \begin{code} Ord-to-π-behaviour : (Ξ± : Ord) β Ord-to-π Ξ± οΌ ssup β¨ Ξ± β© (Ξ» (x : β¨ Ξ± β©) β Ord-to-π (Ξ± β x)) Ord-to-π-behaviour = transfinite-recursion-on-OO-behaviour π (Ξ» Ξ± β ssup β¨ Ξ± β©) \end{code} This map is left cancellable and we will later conclude from this fact that it is actually an embedding. \begin{code} Ord-to-π-is-lc : left-cancellable Ord-to-π Ord-to-π-is-lc {Ξ±} {Ξ²} = transfinite-induction-on-OO _ f Ξ± Ξ² where f : (Ξ± : Ord) β ((a : β¨ Ξ± β©) (Ξ² : Ord) β Ord-to-π (Ξ± β a) οΌ Ord-to-π Ξ² β (Ξ± β a) οΌ Ξ²) β (Ξ² : Ord) β Ord-to-π Ξ± οΌ Ord-to-π Ξ² β Ξ± οΌ Ξ² f Ξ± IH Ξ² p = VII where I : ssup β¨ Ξ± β© (Ξ» (a : β¨ Ξ± β©) β Ord-to-π (Ξ± β a)) οΌ ssup β¨ Ξ² β© (Ξ» (b : β¨ Ξ² β©) β Ord-to-π (Ξ² β b)) I = transportβ (_οΌ_) (Ord-to-π-behaviour Ξ±) (Ord-to-π-behaviour Ξ²) p II : β¨ Ξ± β© οΌ β¨ Ξ² β© II = prβ (from-π-οΌ I) III : (a : β¨ Ξ± β©) β Ord-to-π (Ξ± β a) οΌ Ord-to-π (Ξ² β Idtofun II a) III = happly (prβ (from-π-οΌ I)) IV : (a : β¨ Ξ± β©) β (Ξ± β a) οΌ (Ξ² β Idtofun II a) IV a = IH a (Ξ² β Idtofun II a) (III a) V : (a : β¨ Ξ± β©) β (Ξ± β a) β² Ξ² V a = Idtofun II a , IV a VI : Ξ± βΌ Ξ² VI = to-βΌ V II' : β¨ Ξ² β© οΌ β¨ Ξ± β© II' = prβ (from-π-οΌ (I β»ΒΉ)) III' : (b : β¨ Ξ² β©) β Ord-to-π (Ξ² β b) οΌ Ord-to-π (Ξ± β Idtofun II' b) III' = happly (prβ (from-π-οΌ (I β»ΒΉ))) IV' : (b : β¨ Ξ² β©) β (Ξ² β b) οΌ (Ξ± β Idtofun II' b) IV' b = (IH (Idtofun II' b) (Ξ² β b) ((III' b)β»ΒΉ))β»ΒΉ V' : (b : β¨ Ξ² β©) β (Ξ² β b) β² Ξ± V' b = Idtofun II' b , IV' b VI' : Ξ² βΌ Ξ± VI' = to-βΌ V' VII : Ξ± οΌ Ξ² VII = Extensionality (OO π€) Ξ± Ξ² VI VI' \end{code} Using this we can show that the elements in the image of Ord-to-π are iterative sets, by induction on the ordinal of ordinals in the sense of the HoTT book. \begin{code} Ord-to-π-is-iset : (Ξ± : Ord) β is-iterative-set (Ord-to-π Ξ±) Ord-to-π-is-iset = transfinite-induction-on-OO _ f where f : (Ξ± : Ord) β ((x : β¨ Ξ± β©) β is-iterative-set (Ord-to-π (Ξ± β x))) β is-iterative-set (Ord-to-π Ξ±) f Ξ± IH = transportβ»ΒΉ is-iterative-set (Ord-to-π-behaviour Ξ±) I where I : is-iterative-set (ssup β¨ Ξ± β© (Ξ» (x : β¨ Ξ± β©) β Ord-to-π (Ξ± β x))) I = II , IH where II : is-embedding (Ξ» x β Ord-to-π (Ξ± β x)) II M = III where III : is-prop (Ξ£ a κ β¨ Ξ± β© , Ord-to-π (Ξ± β a) οΌ M) III (a , p) (b , q) = VI where IV : Ξ± β a οΌ Ξ± β b IV = Ord-to-π-is-lc (Ord-to-π (Ξ± β a) οΌβ¨ p β© M οΌβ¨ q β»ΒΉ β© Ord-to-π (Ξ± β b) β) V : a οΌ b V = β-lc Ξ± a b IV VI : (a , p) οΌ (b , q) VI = to-subtype-οΌ (Ξ» x β isets-are-h-isolated (Ord-to-π (Ξ± β x)) (IH x)) V \end{code} So we get a map Ord β π from the above map Ord β π. \begin{code} Ord-to-π : Ord β π Ord-to-π Ξ± = Ord-to-π Ξ± , Ord-to-π-is-iset Ξ± \end{code} The next step is to get a map Ord β π from this map Ord-to-π. We have the definitionally commutative triangle Ord-to-π Ord ------> π \ / \ / Ord-to-π \ / underlying-mset \ / v π We previously showed that Ord-to-π is left cancellable. Hence Ord-to-π is left cancellable as well. But π is a 0-type, so Ord-to-π is actually an embedding. Finally, the map underlying-mset is an embedding, as π is a subtype of π, so Ord-to-π is a composition of embeddings, and therefore an embedding itself. \begin{code} private commutative-triangle : Ord-to-π οΌ underlying-mset β Ord-to-π commutative-triangle = refl Ord-to-π-is-embedding : is-embedding Ord-to-π Ord-to-π-is-embedding = lc-maps-into-sets-are-embeddings Ord-to-π (factor-is-lc Ord-to-π underlying-mset Ord-to-π-is-lc) π-is-set Ord-to-π-is-embedding : is-embedding Ord-to-π Ord-to-π-is-embedding = β-is-embedding Ord-to-π-is-embedding underlying-mset-is-embedding Ord-to-πβ-is-embedding : (Ξ± : Ord) β is-embedding (Ξ» x β Ord-to-π (Ξ± β x)) Ord-to-πβ-is-embedding Ξ± = β-is-embedding (β-is-embedding Ξ±) Ord-to-π-is-embedding \end{code} The fact that Ord-to-π is an embedding seems to be new, and it is interesting because π is not a 0-type. The following gives a recursive characterization of Ord-to-π: \begin{code} Ord-to-π-behaviour : (Ξ± : Ord) β Ord-to-π Ξ± οΌ π-ssup β¨ Ξ± β© (Ξ» (x : β¨ Ξ± β©) β Ord-to-π (Ξ± β x)) (Ord-to-πβ-is-embedding Ξ±) Ord-to-π-behaviour Ξ± = to-subtype-οΌ being-iset-is-prop (Ord-to-π-behaviour Ξ±) \end{code} It is convenient to name the "body" of the definition for the sake of brevity. \begin{code} Ord-to-π-body : Ord β π Ord-to-π-body Ξ± = π-ssup β¨ Ξ± β© (Ξ» (x : β¨ Ξ± β©) β Ord-to-π (Ξ± β x)) (Ord-to-πβ-is-embedding Ξ±) \end{code} We now show that Ord-to-π Ξ± is an iterative ordinal. We begin with a useful observation. \begin{code} Ord-to-π-membership : (A : π) (Ξ± : Ord) β A β Ord-to-π-body Ξ± β (Ξ£ x κ β¨ Ξ± β© , Ord-to-π (Ξ± β x) οΌ A) Ord-to-π-membership A Ξ± = β-behaviour A β¨ Ξ± β© (Ξ» x β Ord-to-π (Ξ± β x)) (Ord-to-πβ-is-embedding Ξ±) \end{code} The map Ord-to-π (Ξ± β -) is lower closed in the following sense: \begin{code} Ord-to-πβ-is-lower : (Ξ± : Ord) (A : π) (x : β¨ Ξ± β©) β A β Ord-to-π (Ξ± β x) β Ξ£ y κ β¨ Ξ± β© , (y βΊβ¨ Ξ± β© x) Γ (A οΌ Ord-to-π (Ξ± β y)) Ord-to-πβ-is-lower Ξ± A x m = IV III where I : A β Ord-to-π-body (Ξ± β x) I = transport (A β_) (Ord-to-π-behaviour (Ξ± β x)) m II : A β Ord-to-π-body (Ξ± β x) β (Ξ£ u κ β¨ Ξ± β x β© , Ord-to-π ((Ξ± β x) β u) οΌ A) II = Ord-to-π-membership A (Ξ± β x) III : Ξ£ u κ β¨ Ξ± β x β© , Ord-to-π ((Ξ± β x) β u) οΌ A III = β II β I IV : type-of III β Ξ£ y κ β¨ Ξ± β© , (y βΊβ¨ Ξ± β© x) Γ (A οΌ Ord-to-π (Ξ± β y)) IV ((y , l) , p) = y , (l , q) where q = A οΌβ¨ p β»ΒΉ β© Ord-to-π ((Ξ± β x) β (y , l)) οΌβ¨ ap Ord-to-π (iterated-β Ξ± x y l) β© Ord-to-π (Ξ± β y) β \end{code} After the above preparation we are ready to show the desired result. Notice that it doesn't require induction. \begin{code} Ord-to-π-is-transitive-iset : (Ξ± : Ord) β is-transitive-iset (Ord-to-π Ξ±) Ord-to-π-is-transitive-iset Ξ± = transportβ»ΒΉ is-transitive-iset (Ord-to-π-behaviour Ξ±) I where g : (B : π) β B β Ord-to-π-body Ξ± β (Ξ£ x κ β¨ Ξ± β© , Ord-to-π (Ξ± β x) οΌ B) g B = Ord-to-π-membership B Ξ± I : is-transitive-iset (Ord-to-π-body Ξ±) I B C B-in-Ξ± C-in-B = Iβ Iβ where Iβ : Ξ£ x κ β¨ Ξ± β© , Ord-to-π (Ξ± β x) οΌ B Iβ = β g B β B-in-Ξ± Iβ : type-of Iβ β C β Ord-to-π-body Ξ± Iβ (x , p) = Iβ Iβ where Iβ : C β Ord-to-π (Ξ± β x) Iβ = transport (C β_) (p β»ΒΉ) C-in-B Iβ : Ξ£ y κ β¨ Ξ± β© , (y βΊβ¨ Ξ± β© x) Γ (C οΌ Ord-to-π (Ξ± β y)) Iβ = Ord-to-πβ-is-lower Ξ± C x Iβ Iβ : type-of Iβ β C β Ord-to-π-body Ξ± Iβ (y , _ , q) = β g C ββ»ΒΉ (y , (q β»ΒΉ)) Ord-to-π-has-transitive-members : (Ξ± : Ord) β has-transitive-members (Ord-to-π Ξ±) Ord-to-π-has-transitive-members Ξ± = transportβ»ΒΉ has-transitive-members (Ord-to-π-behaviour Ξ±) I where I : has-transitive-members (Ord-to-π-body Ξ±) I B B-in-Ξ± = Iβ Iβ where Iβ : Ξ£ x κ β¨ Ξ± β© , Ord-to-π (Ξ± β x) οΌ B Iβ = β Ord-to-π-membership B Ξ± β B-in-Ξ± Iβ : type-of Iβ β is-transitive-iset B Iβ (x , p) = transport is-transitive-iset p (Ord-to-π-is-transitive-iset (Ξ± β x)) Ord-to-π-is-iordinal : (Ξ± : Ord) β is-iterative-ordinal (Ord-to-π Ξ±) Ord-to-π-is-iordinal Ξ± = Ord-to-π-is-transitive-iset Ξ± , Ord-to-π-has-transitive-members Ξ± \end{code} From this we get the desired map Ord β π, which is easily seen to be an embedding from the above development: \begin{code} Ord-to-π : Ord β π Ord-to-π Ξ± = Ord-to-π Ξ± , Ord-to-π-is-iordinal Ξ± Ord-to-π-is-embedding : is-embedding Ord-to-π Ord-to-π-is-embedding = pair-fun-is-embedding-special Ord-to-π Ord-to-π-is-iordinal Ord-to-π-is-embedding being-iordinal-is-prop \end{code} In order to show that this map is an equivalence, with two sided inverse π-to-Ord, we need some preparation: \begin{code} Ord-to-πβ-is-embedding : (Ξ± : Ord) β is-embedding (Ξ» x β Ord-to-π (Ξ± β x)) Ord-to-πβ-is-embedding Ξ± = β-is-embedding (β-is-embedding Ξ±) Ord-to-π-is-embedding Ord-to-πβ-is-lower-closed : (Ξ± : Ord) β is-lower-closed (Ξ» x β Ord-to-π (Ξ± β x)) Ord-to-πβ-is-lower-closed Ξ± Ξ² x l = II I where B : π B = underlying-iset Ξ² I : Ξ£ y κ β¨ Ξ± β© , (y βΊβ¨ Ξ± β© x) Γ (B οΌ Ord-to-π (Ξ± β y)) I = Ord-to-πβ-is-lower Ξ± B x l II : type-of I β Ξ£ y κ β¨ Ξ± β© , Ord-to-π (Ξ± β y) οΌ Ξ² II (y , _ , p) = y , to-subtype-οΌ being-iordinal-is-prop (p β»ΒΉ) \end{code} TODO. Instead show that the composition of two lower closed maps is lower closed, and use this to give a direct proof of the above. We use this to obtain the following recursive characterization of the map Ord-to-π. \begin{code} Ord-to-π-behaviour : (Ξ± : Ord) β Ord-to-π Ξ± οΌ π-ssup β¨ Ξ± β© ((Ξ» (x : β¨ Ξ± β©) β Ord-to-π (Ξ± β x))) (Ord-to-πβ-is-embedding Ξ±) (Ord-to-πβ-is-lower-closed Ξ±) Ord-to-π-behaviour Ξ± = to-subtype-οΌ being-iordinal-is-prop (to-subtype-οΌ being-iset-is-prop (Ord-to-π-behaviour Ξ±)) \end{code} We now establish the following equation by nested induction, first on π and then on the ordinal π-to-Ord Ξ±, which seems to be a new observation. \begin{code} π-to-Ord-equation : (Ξ± : π) (x : π-root Ξ±) β (π-to-Ord Ξ±) β x οΌ π-to-Ord (π-forest Ξ± x) π-to-Ord-equation = π-induction' _ f where f : (Ξ± : π) β ((x : π-root Ξ±) (y : π-root (π-forest Ξ± x)) β π-to-Ord (π-forest Ξ± x) β y οΌ π-to-Ord (π-forest (π-forest Ξ± x) y)) β (x : π-root Ξ±) β (π-to-Ord Ξ± β x) οΌ π-to-Ord (π-forest Ξ± x) f Ξ± IH-f = Transfinite-induction (π-to-Ord Ξ±) _ g where g : (x : π-root Ξ±) β ((y : π-root Ξ±) β y βΊβ¨ π-to-Ord Ξ± β© x β (π-to-Ord Ξ± β y) οΌ π-to-Ord (π-forest Ξ± y)) β (π-to-Ord Ξ± β x) οΌ π-to-Ord (π-forest Ξ± x) g x IH-g = β²-is-extensional _ _ (to-βΌ I) (to-βΌ II) where I : (y : β¨ π-to-Ord Ξ± β x β©) β ((π-to-Ord Ξ± β x) β y) β² π-to-Ord (π-forest Ξ± x) I πͺ@(y , l) = (y' , eq) where Iβ : Ξ£ y' κ π-root (π-forest Ξ± x) , π-forest (π-forest Ξ± x) y' οΌ π-forest Ξ± y Iβ = β <-behaviour (π-forest Ξ± y) (π-forest Ξ± x) β (β π-to-Ord-order Ξ± y x ββ»ΒΉ l) y' = prβ Iβ eq' = prβ Iβ eq = (π-to-Ord Ξ± β x) β πͺ οΌβ¨ β¦ 1β¦ β© π-to-Ord Ξ± β y οΌβ¨ β¦ 2β¦ β© π-to-Ord (π-forest Ξ± y) οΌβ¨ β¦ 3β¦ β© π-to-Ord (π-forest (π-forest Ξ± x) y') οΌβ¨ β¦ 4β¦ β© π-to-Ord (π-forest Ξ± x) β y' β where β¦ 1β¦ = iterated-β (π-to-Ord Ξ±) x y l β¦ 2β¦ = IH-g y l β¦ 3β¦ = ap π-to-Ord (eq' β»ΒΉ) β¦ 4β¦ = (IH-f x y')β»ΒΉ II : (y : β¨ π-to-Ord (π-forest Ξ± x) β©) β (π-to-Ord (π-forest Ξ± x) β y) β² (π-to-Ord Ξ± β x) II y = (πͺ , (eq β»ΒΉ)) where note : π-root (π-forest Ξ± x) οΌ β¨ π-to-Ord (π-forest Ξ± x) β© note = refl Iβ : Ξ£ y' κ π-root Ξ± , π-forest Ξ± y' οΌ π-forest (π-forest Ξ± x) y Iβ = π-forest-is-lower-closed Ξ± (π-forest (π-forest Ξ± x) y) x (π-forest-< (π-forest Ξ± x) y) y' = prβ Iβ eq' = prβ Iβ l : π-forest Ξ± y' < π-forest Ξ± x l = β <-behaviour (π-forest Ξ± y') (π-forest Ξ± x) ββ»ΒΉ (y , (eq' β»ΒΉ)) l' : y' βΊβ¨ π-to-Ord Ξ± β© x l' = β π-to-Ord-order Ξ± y' x β l πͺ = (y' , l') eq = (π-to-Ord Ξ± β x) β πͺ οΌβ¨ β¦ 1β¦ β© π-to-Ord Ξ± β y' οΌβ¨ β¦ 2β¦ β© π-to-Ord (π-forest Ξ± y') οΌβ¨ β¦ 3β¦ β© π-to-Ord (π-forest (π-forest Ξ± x) y) οΌβ¨ β¦ 4β¦ β© π-to-Ord (π-forest Ξ± x) β y β where β¦ 1β¦ = iterated-β (π-to-Ord Ξ±) x y' l' β¦ 2β¦ = IH-g y' l' β¦ 3β¦ = ap π-to-Ord eq' β¦ 4β¦ = (IH-f x y)β»ΒΉ \end{code} From this equation and the previous results, we easily deduce that the map Ord-to-π is an equivalence, by induction on π. \begin{code} Ord-to-π-is-equiv : is-equiv Ord-to-π Ord-to-π-is-equiv = embeddings-with-sections-are-equivs Ord-to-π Ord-to-π-is-embedding (π-to-Ord , Ξ·) where f : (Ξ± : π) β ((x : π-root Ξ±) β Ord-to-π (π-to-Ord (π-forest Ξ± x)) οΌ π-forest Ξ± x) β Ord-to-π (π-to-Ord Ξ±) οΌ Ξ± f Ξ± IH = Ord-to-π (π-to-Ord Ξ±) οΌβ¨ I β© π-ssup (π-root Ξ±) (Ξ» x β Ord-to-π (π-to-Ord Ξ± β x)) e l οΌβ¨ II β© π-ssup (π-root Ξ±) (π-forest Ξ±) e' l' οΌβ¨ π-Ξ· Ξ± β© Ξ± β where e = Ord-to-πβ-is-embedding (π-to-Ord Ξ±) l = Ord-to-πβ-is-lower-closed (π-to-Ord Ξ±) I = Ord-to-π-behaviour (π-to-Ord Ξ±) e' = π-forest-is-embedding Ξ± l' = π-forest-is-lower-closed Ξ± II' = Ξ» x β Ord-to-π (π-to-Ord Ξ± β x) οΌβ¨ ap Ord-to-π (π-to-Ord-equation Ξ± x) β© Ord-to-π (π-to-Ord (π-forest Ξ± x)) οΌβ¨ IH x β© π-forest Ξ± x β II = to-π-οΌ-special (π-root Ξ±) (Ξ» x β Ord-to-π (π-to-Ord Ξ± β x)) (π-forest Ξ±) e l e' l' (dfunext fe II') Ξ· : Ord-to-π β π-to-Ord βΌ id Ξ· = π-induction' _ f \end{code} Hence Ord and π are equivalent types: \begin{code} Ordinals-β : Ord β π Ordinals-β = Ord-to-π , Ord-to-π-is-equiv \end{code} But more than this is true: the types π (HoTT-book ordinal of iterative ordinals) and OO π€ (HoTT-book ordinal of HoTT-book ordinals) are isomorphic as HoTT-book ordinals. It is easy to see that π-to-Ord reflects order: \begin{code} π-to-Ord-reflects-order : (Ξ± Ξ² : π) β π-to-Ord Ξ± β² π-to-Ord Ξ² β Ξ± < Ξ² π-to-Ord-reflects-order Ξ± Ξ² (y , p) = III where I = π-to-Ord (π-forest Ξ² y) οΌβ¨ (π-to-Ord-equation Ξ² y)β»ΒΉ β© (π-to-Ord Ξ² β y) οΌβ¨ p β»ΒΉ β© π-to-Ord Ξ± β II : π-forest Ξ² y οΌ Ξ± II = equivs-are-lc β Ordinals-β ββ»ΒΉ β Ordinals-β ββ»ΒΉ-is-equiv I III : Ξ± < Ξ² III = β <-behaviour Ξ± Ξ² ββ»ΒΉ (y , II) \end{code} And a second application of the above equation shows that π-to-Ord preserves order. \begin{code} π-to-Ord-preserves-order : (Ξ± Ξ² : π) β Ξ± < Ξ² β π-to-Ord Ξ± β² π-to-Ord Ξ² π-to-Ord-preserves-order Ξ± Ξ² l = II I where I : Ξ£ y κ π-root Ξ² , π-forest Ξ² y οΌ Ξ± I = β <-behaviour Ξ± Ξ² β l II : type-of I β π-to-Ord Ξ± β² π-to-Ord Ξ² II (y , refl) = IV where III : π-to-Ord (π-forest Ξ² y) οΌ (π-to-Ord Ξ² β y) III = (π-to-Ord-equation Ξ² y)β»ΒΉ IV : π-to-Ord (π-forest Ξ² y) β² π-to-Ord Ξ² IV = y , III \end{code} Putting this together we get our desired isomorphism of ordinals: \begin{code} Ordinals-agreementβ : π ββ OO π€ Ordinals-agreementβ = β Ordinals-β ββ»ΒΉ , order-preserving-reflecting-equivs-are-order-equivs π (OO π€) β Ordinals-β ββ»ΒΉ β Ordinals-β ββ»ΒΉ-is-equiv π-to-Ord-preserves-order π-to-Ord-reflects-order \end{code} Which then gives an identification between the two types. \begin{code} Ordinals-agreement : π οΌ OO π€ Ordinals-agreement = eqtoidβ (ua π€βΊ) fe π (OO π€) Ordinals-agreementβ \end{code} Notice that this identification lives in the identity type of the type of ordinals in the universe π€βΊ, which is a 0-type, and therefore is unique. \begin{code} Ordinals-agreement-is-unique : is-singleton (π οΌ[ Ordinal π€βΊ ] OO π€) Ordinals-agreement-is-unique = pointed-props-are-singletons Ordinals-agreement (the-type-of-ordinals-is-a-set (ua (π€ βΊ)) fe) \end{code}