open import Level module Ordinals where open import zf open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) open import Data.Empty open import Relation.Binary.PropositionalEquality open import logic open import nat open import Data.Unit using ( ⊤ ) open import Relation.Nullary open import Relation.Binary open import Relation.Binary.Core record Oprev {n : Level} (ord : Set n) (osuc : ord → ord ) (x : ord ) : Set (suc n) where field oprev : ord oprev=x : osuc oprev ≡ x record IsOrdinals {n : Level} (ord : Set n) (o∅ : ord ) (osuc : ord → ord ) (_o<_ : ord → ord → Set n) (next : ord → ord ) : Set (suc (suc n)) where field Otrans : {x y z : ord } → x o< y → y o< z → x o< z OTri : Trichotomous {n} _≡_ _o<_ ¬x<0 : { x : ord } → ¬ ( x o< o∅ ) <-osuc : { x : ord } → x o< osuc x osuc-≡< : { a x : ord } → x o< osuc a → (x ≡ a ) ∨ (x o< a) Oprev-p : ( x : ord ) → Dec ( Oprev ord osuc x ) TransFinite : { ψ : ord → Set n } → ( (x : ord) → ( (y : ord ) → y o< x → ψ y ) → ψ x ) → ∀ (x : ord) → ψ x TransFinite1 : { ψ : ord → Set (suc n) } → ( (x : ord) → ( (y : ord ) → y o< x → ψ y ) → ψ x ) → ∀ (x : ord) → ψ x record IsNext {n : Level } (ord : Set n) (o∅ : ord ) (osuc : ord → ord ) (_o<_ : ord → ord → Set n) (next : ord → ord ) : Set (suc (suc n)) where field x ¬a ¬b c = ¬b eq o<> : {x y : Ordinal } → y o< x → x o< y → ⊥ o<> {ox} {oy} lt tl with trio< ox oy o<> {ox} {oy} lt tl | tri< a ¬b ¬c = ¬c lt o<> {ox} {oy} lt tl | tri≈ ¬a b ¬c = ¬a tl o<> {ox} {oy} lt tl | tri> ¬a ¬b c = ¬a tl osuc-< : { x y : Ordinal } → y o< osuc x → x o< y → ⊥ osuc-< {x} {y} y x x < osuc x -> y = x ∨ x < y → ⊥ osucc {ox} {oy} oy ¬a ¬b c with osuc-≡< c osucc {ox} {oy} oy ¬a ¬b c | case1 eq = ⊥-elim (o<¬≡ (sym eq) oy ¬a ¬b c | case2 lt = ⊥-elim (o<> lt oy ¬a ¬b c = ⊥-elim (o<> (osucc c) oy ¬a ¬b c = ⊥-elim ( o<¬≡ refl (a→b c ) ) maxα : Ordinal → Ordinal → Ordinal maxα x y with trio< x y maxα x y | tri< a ¬b ¬c = y maxα x y | tri> ¬a ¬b c = x maxα x y | tri≈ ¬a refl ¬c = x omin : Ordinal → Ordinal → Ordinal omin x y with trio< x y omin x y | tri< a ¬b ¬c = x omin x y | tri> ¬a ¬b c = y omin x y | tri≈ ¬a refl ¬c = x min1 : {x y z : Ordinal } → z o< x → z o< y → z o< omin x y min1 {x} {y} {z} z ¬a ¬b c = z ¬a ¬b c = osuc x omax x y | tri≈ ¬a refl ¬c = osuc x omax< : ( x y : Ordinal ) → x o< y → osuc y ≡ omax x y omax< x y lt with trio< x y omax< x y lt | tri< a ¬b ¬c = refl omax< x y lt | tri≈ ¬a b ¬c = ⊥-elim (¬a lt ) omax< x y lt | tri> ¬a ¬b c = ⊥-elim (¬a lt ) omax≤ : ( x y : Ordinal ) → x o≤ y → osuc y ≡ omax x y omax≤ x y le with trio< x y omax≤ x y le | tri< a ¬b ¬c = refl omax≤ x y le | tri≈ ¬a refl ¬c = refl omax≤ x y le | tri> ¬a ¬b c with osuc-≡< le omax≤ x y le | tri> ¬a ¬b c | case1 eq = ⊥-elim (¬b eq) omax≤ x y le | tri> ¬a ¬b c | case2 x ¬a ¬b c = ⊥-elim (¬b eq ) omax-x : ( x y : Ordinal ) → x o< omax x y omax-x x y with trio< x y omax-x x y | tri< a ¬b ¬c = ordtrans a <-osuc omax-x x y | tri> ¬a ¬b c = <-osuc omax-x x y | tri≈ ¬a refl ¬c = <-osuc omax-y : ( x y : Ordinal ) → y o< omax x y omax-y x y with trio< x y omax-y x y | tri< a ¬b ¬c = <-osuc omax-y x y | tri> ¬a ¬b c = ordtrans c <-osuc omax-y x y | tri≈ ¬a refl ¬c = <-osuc omxx : ( x : Ordinal ) → omax x x ≡ osuc x omxx x with trio< x x omxx x | tri< a ¬b ¬c = ⊥-elim (¬b refl ) omxx x | tri> ¬a ¬b c = ⊥-elim (¬b refl ) omxx x | tri≈ ¬a refl ¬c = refl omxxx : ( x : Ordinal ) → omax x (omax x x ) ≡ osuc (osuc x) omxxx x = trans ( cong ( λ k → omax x k ) (omxx x )) (sym ( omax< x (osuc x) <-osuc )) open _∧_ o≤-refl : { i j : Ordinal } → i ≡ j → i o≤ j o≤-refl {i} {j} eq = subst (λ k → i o< osuc k ) eq <-osuc OrdTrans : Transitive _o≤_ OrdTrans a≤b b≤c with osuc-≡< a≤b | osuc-≡< b≤c OrdTrans a≤b b≤c | case1 refl | case1 refl = <-osuc OrdTrans a≤b b≤c | case1 refl | case2 a≤c = ordtrans a≤c <-osuc OrdTrans a≤b b≤c | case2 a≤c | case1 refl = ordtrans a≤c <-osuc OrdTrans a≤b b≤c | case2 a ¬a ¬b c = ⊥-elim ( ¬x<0 c ) next< : {x y z : Ordinal} → x o< next z → y o< next x → y o< next z next< {x} {y} {z} x ¬a ¬b c = ⊥-elim (¬nx osuc x o< next x o< next (osuc x) -> next x ≡ osuc z -> z o o< next x -> osuc z o< next x -> next x o< next x nexto≡ {x} | tri< a ¬b ¬c = ⊥-elim (¬nx osuc x o< next (osuc x) o< next x -> next (osuc x) ≡ osuc z -> z o o< next (osuc x) ... nexto≡ {x} | tri> ¬a ¬b c = ⊥-elim (¬nx ¬a ¬b c = -- x < y < next y < next x ⊥-elim ( ¬nx ¬a ¬b c with osuc-≡< x≤y ≤next {x} {y} x≤y | tri> ¬a ¬b c | case1 refl = o≤-refl refl -- x = y < next x ≤next {x} {y} x≤y | tri> ¬a ¬b c | case2 x ¬a ¬b c = o≤-refl (sym ( x ¬a ¬b c = subst (λ k → osuc x o< k ) nexto≡ (osuc ¬a ¬b c = osuc ¬a ¬b c = subst (λ k → x o< k ) nexto≡ xnx record OrdinalSubset (maxordinal : Ordinal) : Set (suc n) where field os→ : (x : Ordinal) → x o< maxordinal → Ordinal os← : Ordinal → Ordinal os←limit : (x : Ordinal) → os← x o< maxordinal os-iso← : {x : Ordinal} → os→ (os← x) (os←limit x) ≡ x os-iso→ : {x : Ordinal} → (lt : x o< maxordinal ) → os← (os→ x lt) ≡ x module o≤-Reasoning {n : Level} (O : Ordinals {n} ) where open inOrdinal O resp-o< : Ordinals._o<_ O Respects₂ _≡_ resp-o< = resp₂ _o<_ trans1 : {i j k : Ordinal} → i o< j → j o< osuc k → i o< k trans1 {i} {j} {k} i