open import Level module ordinal-definable where open import zf open import ordinal open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) open import Relation.Binary.PropositionalEquality open import Data.Nat.Properties open import Data.Empty open import Relation.Nullary open import Relation.Binary open import Relation.Binary.Core -- Ordinal Definable Set record OD {n : Level} : Set (suc n) where field def : (x : Ordinal {n} ) → Set n open OD open import Data.Unit postulate od→ord : {n : Level} → OD {n} → Ordinal {n} ord→od : {n : Level} → Ordinal {n} → OD {n} ord→od x = record { def = λ y → x ≡ y } _∋_ : { n : Level } → ( a x : OD {n} ) → Set n _∋_ {n} a x = def a ( od→ord x ) _c<_ : { n : Level } → ( a x : OD {n} ) → Set n x c< a = a ∋ x _c≤_ : {n : Level} → OD {n} → OD {n} → Set (suc n) a c≤ b = (a ≡ b) ∨ ( b ∋ a ) postulate c<→o< : {n : Level} {x y : OD {n} } → x c< y → od→ord x o< od→ord x o<→c< : {n : Level} {x y : Ordinal {n} } → x o< y → ord→od x c< ord→od x oiso : {n : Level} {x : OD {n}} → ord→od ( od→ord x ) ≡ x diso : {n : Level} {x : Ordinal {n}} → od→ord ( ord→od x ) ≡ x sup-od : {n : Level } → ( OD {n} → OD {n}) → OD {n} sup-c< : {n : Level } → ( ψ : OD {n} → OD {n}) → ∀ {x : OD {n}} → ψ x c< sup-od ψ HOD = OD od∅ : {n : Level} → HOD {n} od∅ {n} = record { def = λ _ → Lift n ⊥ } ∅1 : {n : Level} → ( x : OD {n} ) → ¬ ( x c< od∅ {n} ) ∅1 {n} x (lift ()) ∅3 : {n : Level} → ( x : Ordinal {n}) → ( ∀(y : Ordinal {n}) → ¬ (y o< x ) ) → x ≡ o∅ {n} ∅3 {n} x = TransFinite {n} c1 c2 c3 x where c0 : Nat → Ordinal {n} → Set n c0 lx x = (∀(y : Ordinal {n}) → ¬ (y o< x)) → x ≡ o∅ {n} c1 : ∀ (lx : Nat ) → c0 lx (record { lv = Suc lx ; ord = ℵ lx } ) c1 lx not with not ( record { lv = lx ; ord = Φ lx } ) ... | t with t (case1 ≤-refl ) c1 lx not | t | () c2 : (lx : Nat) → c0 lx (record { lv = lx ; ord = Φ lx } ) c2 Zero not = refl c2 (Suc lx) not with not ( record { lv = lx ; ord = Φ lx } ) ... | t with t (case1 ≤-refl ) c2 (Suc lx) not | t | () c3 : (lx : Nat) (x₁ : OrdinalD lx) → c0 lx (record { lv = lx ; ord = x₁ }) → c0 lx (record { lv = lx ; ord = OSuc lx x₁ }) c3 lx (Φ .lx) d not with not ( record { lv = lx ; ord = Φ lx } ) ... | t with t (case2 Φ< ) c3 lx (Φ .lx) d not | t | () c3 lx (OSuc .lx x₁) d not with not ( record { lv = lx ; ord = OSuc lx x₁ } ) ... | t with t (case2 (s< s