From c3671d181753ae7b1d65118716a50e1ee480fe15 Mon Sep 17 00:00:00 2001 From: finegeometer Date: Fri, 15 Aug 2025 18:10:04 -0700 Subject: [PATCH 1/2] Ordinals. --- src/1Lab/Path/IdentitySystem.lagda.md | 12 + src/Data/Ordinal/Base.lagda.md | 668 ++++++++++++++++++++++++++ src/HoTT.lagda.md | 36 ++ 3 files changed, 716 insertions(+) create mode 100644 src/Data/Ordinal/Base.lagda.md diff --git a/src/1Lab/Path/IdentitySystem.lagda.md b/src/1Lab/Path/IdentitySystem.lagda.md index 26cbf5ce0..208c17e7a 100644 --- a/src/1Lab/Path/IdentitySystem.lagda.md +++ b/src/1Lab/Path/IdentitySystem.lagda.md @@ -176,6 +176,18 @@ identity-system-gives-path {R = R} {r = r} ids = ∙ transport-refl _ ) ``` +In particular, swapping the arguments of the relation in an identity system +yields another identity system. + +```agda +flip-identity-system : + ∀ {ℓ ℓ'} {A : Type ℓ} {R : A → A → Type ℓ'} {r : ∀ a → R a a} + → is-identity-system R r + → is-identity-system (flip R) r +flip-identity-system ids = + equiv-path→identity-system $ identity-system-gives-path ids ∙e sym-equiv +``` + ## Based identity systems {defines="based-identity-system unary-identity-system"} It is sometimes useful to characterise the *based* identity type at diff --git a/src/Data/Ordinal/Base.lagda.md b/src/Data/Ordinal/Base.lagda.md new file mode 100644 index 000000000..326a6cc6a --- /dev/null +++ b/src/Data/Ordinal/Base.lagda.md @@ -0,0 +1,668 @@ + + +```agda +module Data.Ordinal.Base where +``` + +# Ordinals {defines="well-order ordinal"} + +An **ordinal** captures the concept of a [[well-founded]] order. +Such an order is necessarily a *strict* order (like `_<_`, not `_≤_`), +since a well-founded relation cannot be reflexive. + +The conditions for a well-order an be expressed fairly concisely. +A relation `_<_ : X → X → Type` is a well order +if all of the following hold: + +* It is propositional. + +* It is well-founded. + +* It is transitive. + +* Whenever `{z | z < x} ≡ {z | z < y}`, we have `x ≡ y`. + +Every ordinal is in fact a poset, +where `x ≤ y` iff `∀ {z} → z < x → z < y`. +We choose to use a somewhat redundant definition of ordinal, +where this poset is part of the *structure*, +since it often lines up with a preexisting order. + +```agda +record Ordinal (ℓ : Level) : Type (lsuc ℓ) where + no-eta-equality + field + poset : Poset ℓ ℓ + open Poset poset public + field + _<_ : Ob → Ob → Type ℓ + <-thin : ∀ {x y} → is-prop (x < y) + <-wf : Wf _<_ + <→≤ : ∀ {x y} → x < y → x ≤ y + ≤-transforms-< : ∀ {x y} → (x ≤ y) ≃ ∀ {z} → z < x → z < y + <-trans : ∀ {x y z} → x < y → y < z → x < z + <-trans x + +For example, the ordinal `ω` is given by the natural numbers, +under their standard ordering. + +```agda +ω : Ordinal lzero +ω = record where + poset = Nat-poset + _<_ = Nat._<_ + <-thin = Nat.≤-is-prop + <→≤ = Nat.<-weaken + ≤-transforms-< {x} {y} = prop-ext + Nat.≤-is-prop (hlevel 1) + (λ x≤y z≤x → Nat.≤-trans z≤x x≤y) + λ h → + let + h' : {z : Nat} → z Nat.≤ x → z Nat.≤ y + h' = λ { {zero} _ → 0≤x; {suc z} → h {z} } + in h' Nat.≤-refl + <-wf = Data.Wellfounded.Properties.<-wf +``` + +As mentioned above, this representation is rather redundant. +We confirm that ordinals are determined +by their underlying type and strict order, +and can be constructed from the same. + +```agda +record _≃ₒ_ (α : Ordinal ℓ₁) (β : Ordinal ℓ₂) : Type (ℓ₁ ⊔ ℓ₂) where + field + equiv : ⌞ α ⌟ ≃ ⌞ β ⌟ + pres-< : ∀ {x y} → Ordinal._<_ α x y ≃ Ordinal._<_ β (equiv · x) (equiv · y) +``` + +```agda +≃ₒ→≡ : {α β : Ordinal ℓ} → α ≃ₒ β → α ≡ β +≃ₒ→≡ {ℓ} {α} {β} e = p where +``` +
+Long proof. +```agda + module α = Ordinal α + module β = Ordinal β + open _≃ₒ_ e + p₀ : ⌞ α ⌟ ≡ ⌞ β ⌟ + p₀ = ua equiv + p₁ : PathP (λ i → p₀ i → p₀ i → Type ℓ) α._<_ β._<_ + p₁ = funext-dep λ {a} {x} p → funext-dep λ {b} {y} q → + let + fa≡x = ua-pathp→path equiv p + fb≡y = ua-pathp→path equiv q + a +```agda +module _ + (X : Type ℓ) (_<_ : X → X → Type ℓ) + (<-thin : ∀ {x y} → is-prop (x < y)) + (<-wf : Wf _<_) + (<-trans : ∀ {x y z} → x < y → y < z → x < z) + (ext : ∀ {x y} → (∀ {z} → (z < x) ≃ (z < y)) → x ≡ y) + where + + ordinal : Ordinal ℓ + ordinal = record where + poset = record where + Ob = X + _≤_ x y = ∀ {z} → z < x → z < y + ≤-thin = Π-is-hlevel' 1 λ _ → Π-is-hlevel 1 λ _ → <-thin + ≤-refl = id + ≤-trans h₁ h₂ = h₂ ∘ h₁ + ≤-antisym h₁ h₂ = ext (prop-ext <-thin <-thin h₁ h₂) + _<_ = _<_ + <-thin = <-thin + <→≤ x + +If a simulation is also an equivalence, it forces the ordinals to be equal. + +```agda +is-simulation→is-equiv→≃ₒ : {α : Ordinal ℓ₁} {β : Ordinal ℓ₂} {f : ⌞ α ⌟ → ⌞ β ⌟} + (sim : is-simulation α β f) (e : is-equiv f) → α ≃ₒ β +is-simulation→is-equiv→≃ₒ _ e ._≃ₒ_.equiv = _ , e +is-simulation→is-equiv→≃ₒ {_} {_} {α} {β} sim _ ._≃ₒ_.pres-< = prop-ext + (Ordinal.<-thin α) (Ordinal.<-thin β) + (is-simulation.pres-< sim) (is-simulation.reflects-< sim) +``` + +## Ordering + +Intriguingly, simulations are *unique*. + +```agda +simulation-unique : {α : Ordinal ℓ₁} {β : Ordinal ℓ₂} {f g : ⌞ α ⌟ → ⌞ β ⌟} → + is-simulation α β f → is-simulation α β g → f ≡ g +simulation-unique {_} {_} {α} {β} {f} {g} f-sim g-sim = + let + module α = Ordinal α + module β = Ordinal β + open is-simulation + in ext $ Wf-induction α._<_ α.<-wf (λ x → f x ≡ g x) λ x ih → + β.≤-antisym + (equiv→inverse (snd β.≤-transforms-<) λ z + + +We conclude with the ways in which subordinals interact with the order. + +```agda +module _ {α : Ordinal ℓ} {x : ⌞ α ⌟} {y : ⌞ Subordinal α x ⌟} where + open Ordinal α + Subsubordinal : Subordinal (Subordinal α x) y ≃ₒ Subordinal α (fst y) + Subsubordinal ._≃ₒ_.equiv = Iso→Equiv + $ (λ ((z , z + +* Definition 10.3.1: `Acc`{.Agda} +* Lemma 10.3.2: `Acc-is-prop`{.Agda} +* Definition 10.3.3: `Wf`{.Agda} +* Lemma 10.3.4: `Wf-is-prop`{.Agda} +* Example 10.3.5: `<-wf`{.Agda} +* Example 10.3.6: `W-well-founded`{.Agda} +* Several theorems about extensional well-founded sets are stated only for ordinals. + * Definition 10.3.11: `is-simulation`{.Agda} + * Lemma 10.3.12: `is-simulation.has-injective`{.Agda} + * Corollary 10.3.15: `Ordinal-poset`{.Agda} + * Lemma 10.3.16: `simulation-unique`{.Agda} +* Definition 10.3.17: `Ordinal`{.Agda} +* Theorem 10.3.20: `Ord`{.Agda} + ### 10.5: The cumulative hierarchy + ## Instances The usual induction principle for the natural numbers is equivalent to diff --git a/src/HoTT.lagda.md b/src/HoTT.lagda.md index ca999c100..50b6470d0 100644 --- a/src/HoTT.lagda.md +++ b/src/HoTT.lagda.md @@ -48,7 +48,6 @@ open import Data.Set.Material.HIT open import Data.Wellfounded.Base open import Data.Set.Surjection open import Data.Wellfounded.W -open import Data.Ordinal.Base open import Data.Fin.Finite using (Listing→choice) open import Data.Dec open import Data.Nat using (ℕ-well-ordered ; Discrete-Nat) @@ -68,6 +67,7 @@ open import Homotopy.Pushout open import Homotopy.Wedge open import Homotopy.Base +open import Order.Ordinal.Base open import Order.Base import Algebra.Monoid.Category as Monoid diff --git a/src/Data/Ordinal/Base.lagda.md b/src/Order/Ordinal/Base.lagda.md similarity index 75% rename from src/Data/Ordinal/Base.lagda.md rename to src/Order/Ordinal/Base.lagda.md index 326a6cc6a..88bfbef8b 100644 --- a/src/Data/Ordinal/Base.lagda.md +++ b/src/Order/Ordinal/Base.lagda.md @@ -7,15 +7,19 @@ open import Data.Wellfounded.Properties hiding (<-wf) open import Data.Wellfounded.Base open import Order.Instances.Nat +open import Order.Univalent using (Poset-path) +open import Order.Subposet open import Order.Base using (Poset) +import 1Lab.Reflection as Reflection + import Data.Nat.Order as Nat import Data.Nat.Base as Nat ``` --> ```agda -module Data.Ordinal.Base where +module Order.Ordinal.Base where ``` # Ordinals {defines="well-order ordinal"} @@ -53,9 +57,10 @@ record Ordinal (ℓ : Level) : Type (lsuc ℓ) where <-thin : ∀ {x y} → is-prop (x < y) <-wf : Wf _<_ <→≤ : ∀ {x y} → x < y → x ≤ y - ≤-transforms-< : ∀ {x y} → (x ≤ y) ≃ ∀ {z} → z < x → z < y + <→≤→< : ∀ {x y z} → x < y → y ≤ z → x < z + [<→<]→≤ : ∀ {x y} → (∀ {z} → z < x → z < y) → (x ≤ y) <-trans : ∀ {x y z} → x < y → y < z → x < z - <-trans x @@ -85,14 +103,12 @@ under their standard ordering. _<_ = Nat._<_ <-thin = Nat.≤-is-prop <→≤ = Nat.<-weaken - ≤-transforms-< {x} {y} = prop-ext - Nat.≤-is-prop (hlevel 1) - (λ x≤y z≤x → Nat.≤-trans z≤x x≤y) - λ h → - let - h' : {z : Nat} → z Nat.≤ x → z Nat.≤ y - h' = λ { {zero} _ → 0≤x; {suc z} → h {z} } - in h' Nat.≤-refl + <→≤→< = Nat.≤-trans + [<→<]→≤ {x} {y} h = + let + h' : {z : Nat} → z Nat.≤ x → z Nat.≤ y + h' = λ { {zero} _ → 0≤x; {suc z} → h {z} } + in h' Nat.≤-refl <-wf = Data.Wellfounded.Properties.<-wf ``` @@ -126,55 +142,49 @@ instance open _≃ₒ_ e p₀ : ⌞ α ⌟ ≡ ⌞ β ⌟ p₀ = ua equiv + p₁ : PathP (λ i → p₀ i → p₀ i → Type ℓ) α._<_ β._<_ - p₁ = funext-dep λ {a} {x} p → funext-dep λ {b} {y} q → - let - fa≡x = ua-pathp→path equiv p - fb≡y = ua-pathp→path equiv q - a ```agda -module _ - (X : Type ℓ) (_<_ : X → X → Type ℓ) - (<-thin : ∀ {x y} → is-prop (x < y)) - (<-wf : Wf _<_) - (<-trans : ∀ {x y z} → x < y → y < z → x < z) - (ext : ∀ {x y} → (∀ {z} → (z < x) ≃ (z < y)) → x ≡ y) - where - - ordinal : Ordinal ℓ - ordinal = record where +record make-ordinal (ℓ : Level) : Type (lsuc ℓ) where + no-eta-equality + field + Ob : Type ℓ + _<_ : Ob → Ob → Type ℓ + <-thin : ∀ {x y} → is-prop (x < y) + <-wf : Wf _<_ + <-trans : ∀ {x y z} → x < y → y < z → x < z + <-extensional : ∀ {x y} → (∀ {z} → (z < x → z < y) × (z < y → z < x)) → x ≡ y + + to-ordinal : Ordinal ℓ + to-ordinal = record where poset = record where - Ob = X + Ob = Ob _≤_ x y = ∀ {z} → z < x → z < y ≤-thin = Π-is-hlevel' 1 λ _ → Π-is-hlevel 1 λ _ → <-thin ≤-refl = id ≤-trans h₁ h₂ = h₂ ∘ h₁ - ≤-antisym h₁ h₂ = ext (prop-ext <-thin <-thin h₁ h₂) + ≤-antisym h₁ h₂ = <-extensional (h₁ , h₂) _<_ = _<_ <-thin = <-thin <→≤ x