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 open Ordinal postulate od→ord : {n : Level} → OD {n} → Ordinal {n} ord→od : {n : Level} → Ordinal {n} → OD {n} _∋_ : { 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 record _==_ {n : Level} ( a b : OD {n} ) : Set n where field eq→ : ∀ { x : Ordinal {n} } → def a x → def b x eq← : ∀ { x : Ordinal {n} } → def b x → def a x id : {n : Level} {A : Set n} → A → A id x = x eq-refl : {n : Level} { x : OD {n} } → x == x eq-refl {n} {x} = record { eq→ = id ; eq← = id } open _==_ eq-sym : {n : Level} { x y : OD {n} } → x == y → y == x eq-sym eq = record { eq→ = eq← eq ; eq← = eq→ eq } eq-trans : {n : Level} { x y z : OD {n} } → x == y → y == z → x == z eq-trans x=y y=z = record { eq→ = λ t → eq→ y=z ( eq→ x=y t) ; eq← = λ t → eq← x=y ( eq← y=z t) } _c≤_ : {n : Level} → OD {n} → OD {n} → Set (suc n) a c≤ b = (a ≡ b) ∨ ( b ∋ a ) od∅ : {n : Level} → OD {n} od∅ {n} = record { def = λ _ → Lift n ⊥ } postulate c<→o< : {n : Level} {x y : OD {n} } → x c< y → od→ord x o< od→ord y o<→c< : {n : Level} {x y : Ordinal {n} } → x o< y → ord→od x c< ord→od y 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 ψ ∅-base-def : {n : Level} → def ( ord→od (o∅ {n}) ) ≡ def (od∅ {n}) congf : {n : Level} {x y : OD {n}} → { f g : OD {n} → OD {n} } → x ≡ y → f ≡ g → f x ≡ g y congf refl refl = refl o∅→od∅ : {n : Level} → ord→od (o∅ {n}) ≡ od∅ {n} o∅→od∅ {n} = cong ( λ k → record { def = k }) ( ∅-base-def ) ∅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< sy = lift ( ∅8 y (o<-subst (c<→o< {n} {ord→od y} {x} (def-subst {n} {x} {y} x>y refl (sym diso))) ord-iso eq )) =→¬< : {x : Nat } → ¬ ( x < x ) =→¬< {Zero} () =→¬< {Suc x} (s≤s lt) = =→¬< lt >→¬< : {x y : Nat } → (x < y ) → ¬ ( y < x ) >→¬< (s≤s x→¬< x : {n : Level } ( ox oy : Ordinal {n}) → ox o< oy → oy o< ox → ⊥ o<> ox oy (case1 x→¬< x ox oy (case1 x ox oy (case2 x ox oy (case2 x x : {n : Level} → { x y : OD {n} } → (x == y) → (od→ord x ) o< ( od→ord y) → ⊥ o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case1 lt) with yx (def-subst {n} {ord→od (od→ord y)} {od→ord (ord→od (od→ord x))} (o<→c< (case1 lt )) oiso diso ) ... | oyx with o<¬≡ (od→ord x) (od→ord x) refl (c<→o< oyx ) ... | () o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case2 lt) with yx (def-subst {n} {ord→od (od→ord y)} {od→ord (ord→od (od→ord x))} (o<→c< (case2 lt )) oiso diso ) ... | oyx with o<¬≡ (od→ord x) (od→ord x) refl (c<→o< oyx ) ... | () o<→¬== : {n : Level} → { x y : OD {n} } → (od→ord x ) o< ( od→ord y) → ¬ (x == y ) o<→¬== {n} {x} {y} lt eq = o<→o> eq lt o<→¬c> : {n : Level} → { x y : OD {n} } → (od→ord x ) o< ( od→ord y) → ¬ (y c< x ) o<→¬c> {n} {x} {y} olt clt = o<> (od→ord x) (od→ord y) olt (c<→o< clt ) where o≡→¬c< : {n : Level} → { x y : OD {n} } → (od→ord x ) ≡ ( od→ord y) → ¬ x c< y o≡→¬c< {n} {x} {y} oeq lt = o<¬≡ (od→ord x) (od→ord y) (orefl oeq ) (c<→o< lt) tri-c< : {n : Level} → Trichotomous _==_ (_c<_ {suc n}) tri-c< {n} x y with trio< {n} (od→ord x) (od→ord y) tri-c< {n} x y | tri< a ¬b ¬c = tri< (def-subst (o<→c< a) oiso diso) (o<→¬== a) ( o<→¬c> a ) tri-c< {n} x y | tri≈ ¬a b ¬c = tri≈ (o≡→¬c< b) (ord→== b) (o≡→¬c< (sym b)) tri-c< {n} x y | tri> ¬a ¬b c = tri> ( o<→¬c> c) (λ eq → o<→¬== c (eq-sym eq ) ) (def-subst (o<→c< c) oiso diso) c<> : {n : Level } { x y : OD {suc n}} → x c< y → y c< x → ⊥ c<> {n} {x} {y} x {n} {x} {y} x {n} {x} {y} x b ( c<→o< x {n} {x} {y} x ¬a ¬b c = ¬a x ¬a ¬b c = no ( ∅2 c ) is-∋ : {n : Level} → ( x y : OD {suc n} ) → Dec ( x ∋ y ) is-∋ {n} x y with tri-c< x y is-∋ {n} x y | tri< a ¬b ¬c = no ¬c is-∋ {n} x y | tri≈ ¬a b ¬c = no ¬c is-∋ {n} x y | tri> ¬a ¬b c = yes c is-o∅ : {n : Level} → ( x : Ordinal {suc n} ) → Dec ( x ≡ o∅ {suc n} ) is-o∅ {n} record { lv = Zero ; ord = (Φ .0) } = yes refl is-o∅ {n} record { lv = Zero ; ord = (OSuc .0 ord₁) } = no ( λ () ) is-o∅ {n} record { lv = (Suc lv₁) ; ord = ord } = no (λ()) open _∧_ ∅9 : {n : Level} → {x : OD {n} } → ¬ x == od∅ → o∅ o< od→ord x ∅9 {_} {x} not = ∅5 lemma where lemma : ¬ od→ord x ≡ o∅ lemma eq = not ( ∅7 eq ) ∅10 : {n : Level} → {ox : Ordinal {n}} → (not : ¬ (ord→od ox == od∅)) → o∅ o< ox ∅10 {n} {ox} not = subst (λ k → o∅ o< k) diso (∅9 not) ¬∅=→∅∈ : {n : Level} → { x : OD {suc n} } → ¬ ( x == od∅ {suc n} ) → x ∋ od∅ {suc n} ¬∅=→∅∈ {n} {x} ne = def-subst (lemma (od→ord x) (subst (λ k → ¬ (k == od∅ {suc n} )) (sym oiso) ne )) oiso refl where lemma : (ox : Ordinal {suc n}) → ¬ (ord→od ox == od∅ {suc n} ) → ord→od ox ∋ od∅ {suc n} lemma ox ne with is-o∅ ox lemma ox ne | yes refl with ne ( ord→== lemma1 ) where lemma1 : od→ord (ord→od o∅) ≡ od→ord od∅ lemma1 = cong ( λ k → od→ord k ) o∅→od∅ lemma o∅ ne | yes refl | () lemma ox ne | no ¬p = subst ( λ k → def (ord→od ox) (od→ord k) ) o∅→od∅ (o<→c< (∅5 ¬p)) -- ∃-set : {n : Level} → ( x : OD {suc n} ) → ( ψ : OD {suc n} → Set (suc n) ) → Set (suc n) -- ∃-set = {!!} -- postulate f-extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality (suc (suc n)) (suc (suc n )) -- ==∅→≡ : {n : Level} → { x : OD {suc n} } → ( x == od∅ {suc n} ) → x ≡ od∅ {suc n} -- ==∅→≡ {n} {x} eq = cong ( λ k → record { def = k } ) (trans {!!} ∅-base-def ) where open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) OD→ZF : {n : Level} → ZF {suc (suc n)} {suc n} OD→ZF {n} = record { ZFSet = OD {suc n} ; _∋_ = _∋_ ; _≈_ = _==_ ; ∅ = od∅ ; _,_ = _,_ ; Union = Union ; Power = Power ; Select = Select ; Replace = Replace ; infinite = record { def = λ x → x ≡ record { lv = Suc Zero ; ord = ℵ Zero } } ; isZF = isZF } where Replace : OD {suc n} → (OD {suc n} → OD {suc n} ) → OD {suc n} Replace X ψ = sup-od ψ Select : OD {suc n} → (OD {suc n} → Set (suc n) ) → OD {suc n} Select X ψ = record { def = λ x → ( def X x ∧ ψ ( ord→od x )) } _,_ : OD {suc n} → OD {suc n} → OD {suc n} x , y = record { def = λ z → ( (z ≡ od→ord x ) ∨ ( z ≡ od→ord y )) } Union : OD {suc n} → OD {suc n} Union U = record { def = λ y → y o< (od→ord U) } Power : OD {suc n} → OD {suc n} Power x = record { def = λ y → (z : Ordinal {suc n} ) → ( def x y ∧ def (ord→od z) y ) } ZFSet = OD {suc n} _∈_ : ( A B : ZFSet ) → Set (suc n) A ∈ B = B ∋ A _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set (suc n) _⊆_ A B {x} = A ∋ x → B ∋ x _∩_ : ( A B : ZFSet ) → ZFSet A ∩ B = Select (A , B) ( λ x → ( A ∋ x ) ∧ (B ∋ x) ) _∪_ : ( A B : ZFSet ) → ZFSet A ∪ B = Select (A , B) ( λ x → (A ∋ x) ∨ ( B ∋ x ) ) infixr 200 _∈_ infixr 230 _∩_ _∪_ infixr 220 _⊆_ isZF : IsZF (OD {suc n}) _∋_ _==_ od∅ _,_ Union Power Select Replace (record { def = λ x → x ≡ record { lv = Suc Zero ; ord = ℵ Zero } }) isZF = record { isEquivalence = record { refl = eq-refl ; sym = eq-sym; trans = eq-trans } ; pair = pair ; union-u = {!!} ; union→ = {!!} ; union← = {!!} ; empty = empty ; power→ = {!!} ; power← = {!!} ; extensionality = {!!} ; minimul = minimul ; regularity = regularity ; infinity∅ = {!!} ; infinity = {!!} ; selection = λ {ψ} {X} {y} → selection {ψ} {X} {y} ; replacement = {!!} } where open _∧_ open Minimumo pair : (A B : OD {suc n} ) → ((A , B) ∋ A) ∧ ((A , B) ∋ B) proj1 (pair A B ) = case1 refl proj2 (pair A B ) = case2 refl empty : (x : OD {suc n} ) → ¬ (od∅ ∋ x) empty x () union-u : (X z : OD {suc n}) → Union X ∋ z → OD {suc n} union-u X z U>z with od→ord X | od→ord z union-u X z (case1 (s≤s x)) | record { lv = Suc lX ; ord = oX } | record { lv = lz ; ord = oz } = ord→od (mid (Suc lX) lz oX oz (s≤s x) ) where mid : (lX lz : Nat) → (oX : OrdinalD {suc n} lX ) → (oz : OrdinalD {suc n} lz ) → lX > lz → Ordinal {suc n} mid (Suc lX) 0 oX oz (s≤s z≤n) = record { lv = 0 ; ord = OSuc 0 oz } mid (Suc (Suc lX)) (Suc lz) (Φ (Suc (Suc lX))) (Φ (Suc lz)) (s≤s (s≤s lt)) = record { lv = Suc lz ; ord = OSuc (Suc lz) (Φ (Suc lz) ) } mid (Suc (Suc lX)) (Suc lz) (Φ (Suc (Suc lX))) (OSuc (Suc lz) oz) (s≤s (s≤s lt)) = record { lv = Suc lz ; ord = OSuc (Suc lz) (OSuc (Suc lz) oz) } mid (Suc (Suc lX)) (Suc lz) (Φ (Suc (Suc lX))) (ℵ lz) (s≤s (s≤s lt)) = record { lv = (Suc (Suc lX)) ; ord = OSuc (Suc (Suc lX)) (ℵ (Suc lX)) } mid (Suc (Suc lX)) (Suc lz) (OSuc (Suc (Suc lX)) oX) oz (s≤s (s≤s lt)) = record { lv = Suc (Suc lX) ; ord = oX } mid (Suc (Suc lX)) (Suc lz) (ℵ (Suc lX)) oz (s≤s (s≤s lt)) = record { lv = Suc (Suc lX) ; ord = OSuc (Suc (Suc lX)) (Φ (Suc (Suc lX))) } union-u X z (case2 Φ<) | record { lv = Zero ; ord = (OSuc 0 oX) } | record { lv = .0 ; ord = .(Φ 0) } = ord→od (record { lv = Zero ; ord = oX }) -- this case contains X = odΦ union-u X z (case2 Φ<) | record { lv = Suc lv₁ ; ord = (OSuc (Suc lv₁) oX) } | record { lv = .(Suc lv₁) ; ord = .(Φ (Suc lv₁)) } = ord→od (record { lv = Suc lv₁ ; ord = oX } ) union-u X z (case2 (s< x)) | record { lv = lv₁ ; ord = .(OSuc lv₁ _) } | record { lv = .lv₁ ; ord = .(OSuc lv₁ _) } = {!!} union-u X z (case2 ℵΦ<) | record { lv = Suc lX ; ord = .(ℵ _) } | record { lv = .(Suc _) ; ord = .(Φ (Suc _)) } = ord→od ( record { lv = Suc lX ; ord = OSuc (Suc lX ) (Φ (Suc lX))} ) union-u X z (case2 (ℵ< x)) | record { lv = Suc lX ; ord = .(ℵ _) } | record { lv = .(Suc _) ; ord = OSuc (Suc _) oz } = ord→od ( record { lv = Suc lX ; ord = OSuc (Suc lX ) (OSuc (Suc lX) (OSuc (Suc lX) oz ))} ) union-u X z (case2 ℵs<) | record { lv = Suc lX ; ord = OSuc (Suc lX) (ℵ lX) } | record { lv = Suc lX ; ord = ℵ lX } = {!!} union-u X z (case2 (ℵss< x)) | record { lv = .(Suc _) ; ord = .(OSuc (Suc _) _) } | record { lv = .(Suc _) ; ord = .(ℵ _) } = {!!} union→ : (X x y : OD {suc n}) → X ∋ x → x ∋ y → Union X ∋ y union→ X x y X∋x x∋y = c<→o< (transitive {n} {X} {x} {y} X∋x x∋y ) union← : (X z : OD {suc n}) → (X∋z : Union X ∋ z) → ((Union X ∋ union-u X z X∋z) ∧ (union-u X z X∋z ∋ z)) union← = {!!} ψiso : {ψ : OD {suc n} → Set (suc n)} {x y : OD {suc n}} → ψ x → x ≡ y → ψ y ψiso {ψ} t refl = t selection : {ψ : OD → Set (suc n)} {X y : OD} → ((X ∋ y) ∧ ψ y) ⇔ (Select X ψ ∋ y) selection {ψ} {X} {y} = record { proj1 = λ cond → record { proj1 = proj1 cond ; proj2 = ψiso {ψ} (proj2 cond) (sym oiso) } ; proj2 = λ select → record { proj1 = proj1 select ; proj2 = ψiso {ψ} (proj2 select) oiso } } ∅-iso : {x : OD} → ¬ (x == od∅) → ¬ ((ord→od (od→ord x)) == od∅) ∅-iso {x} neq = subst (λ k → ¬ k) (=-iso {n} ) neq minimul : (x : OD {suc n} ) → ¬ (x == od∅ )→ OD {suc n} minimul x not = od∅ regularity : (x : OD) (not : ¬ (x == od∅)) → (x ∋ minimul x not) ∧ (Select (minimul x not) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅) proj1 (regularity x not ) = ¬∅=→∅∈ not proj2 (regularity x not ) = record { eq→ = reg ; eq← = λ () } where reg : {y : Ordinal} → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) y → def od∅ y reg {y} t with proj1 t ... | x∈∅ = x∈∅