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} _∋_ : { 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 ) 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}) ∅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 ¬a ¬b c = ¬b refl ∅8 : {n : Level} → ( x : Ordinal {n} ) → ¬ x o< o∅ {n} ∅8 {n} x (case1 ()) ∅8 {n} x (case2 ()) -- ∅10 : {n : Level} → (x : OD {n} ) → ¬ ( ( y : OD {n} ) → Lift (suc n) ( x ∋ y)) → x ≡ od∅ -- ∅10 {n} x not = ? open Ordinal ∋-subst : {n : Level} {X Y x y : OD {suc n} } → X ≡ x → Y ≡ y → X ∋ Y → x ∋ y ∋-subst refl refl x = x -- ∅77 : {n : Level} → (x : OD {suc n} ) → ¬ ( ord→od (o∅ {suc n}) ∋ x ) -- ∅77 {n} x lt = {!!} where ∅7' : {n : Level} → ord→od (o∅ {n}) ≡ od∅ {n} ∅7' {n} = cong ( λ k → record { def = k }) ( ∅-base-def ) where ∅7 : {n : Level} → ( x : OD {n} ) → od→ord x ≡ o∅ {n} → x ≡ od∅ {n} ∅7 {n} x eq = trans ( trans (sym oiso)( cong ( λ k → ord→od k ) eq ) ) ∅7' ∅9 : {n : Level} → (x : OD {n} ) → ¬ x ≡ od∅ → o∅ o< od→ord x ∅9 x not = ∅5 ( od→ord x) lemma where lemma : ¬ od→ord x ≡ o∅ lemma eq = not ( ∅7 x eq ) OD→ZF : {n : Level} → ZF {suc n} {suc n} OD→ZF {n} = record { ZFSet = OD {n} ; _∋_ = λ a x → Lift (suc n) ( a ∋ x ) ; _≈_ = _≡_ ; ∅ = 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 {n} → (OD {n} → OD {n} ) → OD {n} Replace X ψ = sup-od ψ Select : OD {n} → (OD {n} → Set (suc n) ) → OD {n} Select X ψ = record { def = λ x → select ( ord→od x ) } where select : OD {n} → Set n select x with ψ x ... | t = Lift n ⊤ _,_ : OD {n} → OD {n} → OD {n} x , y = record { def = λ z → ( (z ≡ od→ord x ) ∨ ( z ≡ od→ord y )) } Union : OD {n} → OD {n} Union x = record { def = λ y → {z : Ordinal {n}} → def x z → def (ord→od z) y } Power : OD {n} → OD {n} Power x = record { def = λ y → (z : Ordinal {n} ) → ( def x y ∧ def (ord→od z) y ) } ZFSet = OD {n} _∈_ : ( A B : ZFSet ) → Set n A ∈ B = B ∋ A _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set n _⊆_ A B {x} = A ∋ x → B ∋ x _∩_ : ( A B : ZFSet ) → ZFSet A ∩ B = Select (A , B) ( λ x → (Lift (suc n) ( A ∋ x )) ∧ (Lift (suc n) ( B ∋ x ) )) _∪_ : ( A B : ZFSet ) → ZFSet A ∪ B = Select (A , B) ( λ x → (Lift (suc n) ( A ∋ x )) ∨ (Lift (suc n) ( B ∋ x ) )) infixr 200 _∈_ infixr 230 _∩_ _∪_ infixr 220 _⊆_ isZF : IsZF (OD {n}) (λ a x → Lift (suc n) ( a ∋ x )) _≡_ od∅ _,_ Union Power Select Replace (record { def = λ x → x ≡ record { lv = Suc Zero ; ord = ℵ Zero } }) isZF = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans } ; pair = pair ; union→ = {!!} ; union← = {!!} ; empty = empty ; power→ = {!!} ; power← = {!!} ; extentionality = {!!} ; minimul = minimul ; regularity = {!!} ; infinity∅ = {!!} ; infinity = {!!} ; selection = {!!} ; replacement = {!!} } where open _∧_ open Minimumo pair : (A B : OD {n} ) → Lift (suc n) ((A , B) ∋ A) ∧ Lift (suc n) ((A , B) ∋ B) proj1 (pair A B ) = lift ( case1 refl ) proj2 (pair A B ) = lift ( case2 refl ) empty : (x : OD {n} ) → ¬ Lift (suc n) (od∅ ∋ x) empty x (lift (lift ())) union→ : (X x y : OD {n} ) → Lift (suc n) (X ∋ x) → Lift (suc n) (x ∋ y) → Lift (suc n) (Union X ∋ y) union→ X x y (lift X∋x) (lift x∋y) = lift {!!} where lemma : {z : Ordinal {n} } → def X z → z ≡ od→ord y lemma {z} X∋z = {!!} minord : (x : OD {n} ) → ¬ x ≡ od∅ → Minimumo (od→ord x) minord x not = ominimal (od→ord x) (∅9 x not) minimul : (x : OD {n} ) → ¬ x ≡ od∅ → OD {n} minimul x not = ord→od ( mino (minord x not)) minimul