open import Level module constructible-set (n : Level) where open import zf open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ) open import Relation.Binary.PropositionalEquality data OridinalD : (lv : Nat) → Set n where Φ : {lv : Nat} → OridinalD lv OSuc : {lv : Nat} → OridinalD lv → OridinalD lv ℵ_ : (lv : Nat) → OridinalD (Suc lv) record Ordinal : Set n where field lv : Nat ord : OridinalD lv data _o<_ : {lx ly : Nat} → OridinalD lx → OridinalD ly → Set n where l< : {lx ly : Nat } → {x : OridinalD lx } → {y : OridinalD ly } → lx < ly → x o< y Φ< : {lx : Nat} → {x : OridinalD lx} → Φ {lx} o< OSuc {lx} x s< : {lx : Nat} → {x y : OridinalD lx} → x o< y → OSuc {lx} x o< OSuc {lx} y ℵΦ< : {lx : Nat} → {x : OridinalD (Suc lx) } → Φ {Suc lx} o< (ℵ lx) ℵ< : {lx : Nat} → {x : OridinalD (Suc lx) } → OSuc {Suc lx} x o< (ℵ lx) open import Data.Nat.Properties open import Data.Empty open import Relation.Nullary open import Relation.Binary open import Relation.Binary.Core ≡→¬< : { x y : Nat } → x ≡ y → x < y → ⊥ ≡→¬< {Zero} {Zero} refl () ≡→¬< {Suc x} {.(Suc x)} refl (s≤s t) = ≡→¬< {x} {x} refl t x≤x : { x : Nat } → x ≤ x x≤x {Zero} = z≤n x≤x {Suc x} = s≤s ( x≤x ) x<>y : { x y : Nat } → x > y → x < y → ⊥ x<>y {.(Suc _)} {.(Suc _)} (s≤s lt) (s≤s lt1) = x<>y lt lt1 triO> : {lx ly : Nat} {x : OridinalD lx } { y : OridinalD ly } → ly < lx → x o< y → ⊥ triO> {lx} {ly} {x} {y} y {lx} {ly} {x} {y} y {lx} {ly} {x} {y} y {lx} {ly} {x} {y} y ¬a ¬b c = ¬a x₁ triO> {lx} {ly} {Φ} {OSuc _} y ¬a ¬b c = ¬b refl triO> {lx} {ly} {OSuc px} {OSuc py} y ¬a ¬b c = triO> y {lx} {ly} {Φ {u}} {ℵ w} y ¬a ¬b c = ¬b refl triO> {lx} {ly} {(OSuc _)} {ℵ u} y ¬a ¬b c = ¬b refl ≡→¬o< : {lv : Nat} → {x : OridinalD lv } → x o< x → ⊥ ≡→¬o< {lx} {x} (l< y) = ≡→¬< refl y ≡→¬o< {lx} {OSuc y} (s< t) = ≡→¬o< t trio<> : {lx : Nat} {x : OridinalD lx } { y : OridinalD lx } → y o< x → x o< y → ⊥ trio<> {lx} {x} {y} (l< lt) _ = ≡→¬< refl lt trio<> {lx} {x} {y} _ (l< lt) = ≡→¬< refl lt trio<> {lx} {.(OSuc _)} {.(OSuc _)} (s< s) (s< t) = trio<> s t trio<≡ : {lx : Nat} {x : OridinalD lx } { y : OridinalD lx } → x ≡ y → x o< y → ⊥ trio<≡ refl = ≡→¬o< trio>≡ : {lx : Nat} {x : OridinalD lx } { y : OridinalD lx } → x ≡ y → y o< x → ⊥ trio>≡ refl = ≡→¬o< triO : {lx ly : Nat} → OridinalD lx → OridinalD ly → Tri (lx < ly) ( lx ≡ ly ) ( lx > ly ) triO {lx} {ly} x y = <-cmp lx ly triOonSameLevel : {lx : Nat} → Trichotomous _≡_ ( _o<_ {lx} {lx} ) triOonSameLevel {lv} Φ Φ = tri≈ ≡→¬o< refl ≡→¬o< triOonSameLevel {Suc lv} (ℵ lv) (ℵ lv) = tri≈ ≡→¬o< refl ≡→¬o< triOonSameLevel {lv} Φ (OSuc y) = tri< Φ< (λ ()) ( λ lt → trio<> lt Φ< ) triOonSameLevel {.(Suc lv)} Φ (ℵ lv) = tri< (ℵΦ< {lv} {Φ} ) (λ ()) ( λ lt → trio<> lt ((ℵΦ< {lv} {Φ} )) ) triOonSameLevel {Suc lv} (ℵ lv) Φ = tri> ( λ lt → trio<> lt (ℵΦ< {lv} {Φ} ) ) (λ ()) (ℵΦ< {lv} {Φ} ) triOonSameLevel {Suc lv} (ℵ lv) (OSuc y) = tri> ( λ lt → trio<> lt (ℵ< {lv} {y} ) ) (λ ()) (ℵ< {lv} {y} ) triOonSameLevel {lv} (OSuc x) Φ = tri> (λ lt → trio<> lt Φ<) (λ ()) Φ< triOonSameLevel {.(Suc lv)} (OSuc x) (ℵ lv) = tri< ℵ< (λ ()) (λ lt → trio<> lt ℵ< ) triOonSameLevel {lv} (OSuc x) (OSuc y) with triOonSameLevel x y triOonSameLevel {lv} (OSuc x) (OSuc y) | tri< a ¬b ¬c = tri< (s< a) (λ tx=ty → trio<≡ tx=ty (s< a) ) ( λ lt → trio<> lt (s< a) ) triOonSameLevel {lv} (OSuc x) (OSuc x) | tri≈ ¬a refl ¬c = tri≈ ≡→¬o< refl ≡→¬o< triOonSameLevel {lv} (OSuc x) (OSuc y) | tri> ¬a ¬b c = tri> ( λ lt → trio<> lt (s< c) ) (λ tx=ty → trio>≡ tx=ty (s< c) ) (s< c) <→≤ : {lx ly : Nat} → lx < ly → (Suc lx ≤ ly) <→≤ {Zero} {Suc ly} (s≤s lt) = s≤s z≤n <→≤ {Suc lx} {Zero} () <→≤ {Suc lx} {Suc ly} (s≤s lt) = s≤s (<→≤ lt) orddtrans : {lx ly lz : Nat} {x : OridinalD lx } { y : OridinalD ly } { z : OridinalD lz } → x o< y → y o< z → x o< z orddtrans {lx} {ly} {lz} x ¬a ¬b₁ c = l< {!!} -- ⊥-elim ( ¬a c ) orddtrans {lx} {ly} {lz} x ¬a ¬b c | tri< a ¬b₁ ¬c = l< {!!} orddtrans {lx} {ly} {lz} x ¬a ¬b c | tri≈ ¬a₁ refl ¬c = l< {!!} orddtrans {lx} {ly} {lz} x ¬a ¬b c | tri> ¬a₁ ¬b₁ c₁ = {!!} orddtrans {lx} {ly} {lz} x ¬a₁ ¬b c = l< {!!} orddtrans {lx} {lx} {lx} x : { lx ly : Nat } → OridinalD lx → OridinalD ly → lx > ly → OridinalD lx maxα> x y _ = x maxα= : { lx : Nat } → OridinalD lx → OridinalD lx → OridinalD lx maxα= x y with triOonSameLevel x y maxα= x y | tri< a ¬b ¬c = y maxα= x y | tri≈ ¬a b ¬c = x maxα= x y | tri> ¬a ¬b c = x OrdTrans : Transitive (λ ( a b : Ordinal ) → (a ≡ b) ∨ (Ordinal.lv a < Ordinal.lv b) ∨ (Ordinal.ord a o< Ordinal.ord b) ) OrdTrans (case1 refl) (case1 refl) = case1 refl OrdTrans (case1 refl) (case2 lt2) = case2 lt2 OrdTrans (case2 lt1) (case1 refl) = case2 lt1 OrdTrans (case2 (case1 x)) (case2 (case1 y)) = case2 ( case1 ( <-trans x y ) ) OrdTrans (case2 (case1 x)) (case2 (case2 y)) = case2 {!!} OrdTrans (case2 (case2 x)) (case2 (case1 y)) = case2 {!!} OrdTrans (case2 (case2 x)) (case2 (case2 y)) = case2 {!!} OrdPreorder : Preorder n n n OrdPreorder = record { Carrier = Ordinal ; _≈_ = _≡_ ; _∼_ = λ a b → (a ≡ b) ∨ (Ordinal.lv a < Ordinal.lv b) ∨ (Ordinal.ord a o< Ordinal.ord b ) ; isPreorder = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans } ; reflexive = case1 ; trans = OrdTrans } } -- X' = { x ∈ X | ψ x } ∪ X , Mα = ( ∪ (β < α) Mβ ) ' data Constructible {lv : Nat} ( α : OridinalD lv ) : Set (suc n) where fsub : ( ψ : OridinalD lv → Set n ) → Constructible α xself : OridinalD lv → Constructible α record ConstructibleSet : Set (suc n) where field level : Nat α : OridinalD level constructible : Constructible α open ConstructibleSet data _c∋_ : {lv lv' : Nat} {α : OridinalD lv } {α' : OridinalD lv' } → Constructible {lv} α → Constructible {lv'} α' → Set n where c> : {lv lv' : Nat} {α : OridinalD lv } {α' : OridinalD lv' } (ta : Constructible {lv} α ) ( tx : Constructible {lv'} α' ) → α' o< α → ta c∋ tx xself-fsub : {lv : Nat} {α : OridinalD lv } (ta : OridinalD lv ) ( ψ : OridinalD lv → Set n ) → _c∋_ {_} {_} {α} {α} (xself ta ) ( fsub ψ) fsub-fsub : {lv lv' : Nat} {α : OridinalD lv } ( ψ : OridinalD lv → Set n ) ( ψ₁ : OridinalD lv → Set n ) → ( ∀ ( x : OridinalD lv ) → ψ x → ψ₁ x ) → _c∋_ {_} {_} {α} {α} ( fsub ψ ) ( fsub ψ₁) _∋_ : (ConstructibleSet ) → (ConstructibleSet ) → Set n a ∋ x = constructible a c∋ constructible x transitiveness : (a b c : ConstructibleSet ) → a ∋ b → b ∋ c → a ∋ c transitiveness a b c a∋b b∋c with constructible a c∋ constructible b | constructible b c∋ constructible c ... | t1 | t2 = {!!} data _c≈_ : {lv lv' : Nat} {α : OridinalD lv } {α' : OridinalD lv' } → Constructible {lv} α → Constructible {lv'} α' → Set n where crefl : {lv : Nat} {α : OridinalD lv } → _c≈_ {_} {_} {α} {α} (xself α ) (xself α ) feq : {lv : Nat} {α : OridinalD lv } → ( ψ : OridinalD lv → Set n ) ( ψ₁ : OridinalD lv → Set n ) → (∀ ( x : OridinalD lv ) → ψ x ⇔ ψ₁ x ) → _c≈_ {_} {_} {α} {α} ( fsub ψ ) ( fsub ψ₁) _≈_ : (ConstructibleSet ) → (ConstructibleSet ) → Set n a ≈ x = constructible a c≈ constructible x ConstructibleSet→ZF : ZF {suc n} ConstructibleSet→ZF = record { ZFSet = ConstructibleSet ; _∋_ = _∋_ ; _≈_ = _≈_ ; ∅ = record { level = Zero ; α = Φ ; constructible = xself Φ } ; _×_ = {!!} ; Union = {!!} ; Power = {!!} ; Select = {!!} ; Replace = {!!} ; infinite = {!!} ; isZF = {!!} }