{-# OPTIONS --allow-unsolved-metas #-} open import Level open import Ordinals module OD {n : Level } (O : Ordinals {n} ) where open import zf open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) open import Relation.Binary.PropositionalEquality hiding ( [_] ) open import Data.Nat.Properties open import Data.Empty open import Relation.Nullary open import Relation.Binary hiding (_⇔_) open import Relation.Binary.Core hiding (_⇔_) open import logic import OrdUtil open import nat open Ordinals.Ordinals O open Ordinals.IsOrdinals isOrdinal open Ordinals.IsNext isNext open OrdUtil O -- Ordinal Definable Set record OD : Set (suc n ) where field def : (x : Ordinal ) → Set n open OD open _∧_ open _∨_ open Bool record _==_ ( a b : OD ) : Set n where field eq→ : ∀ { x : Ordinal } → def a x → def b x eq← : ∀ { x : Ordinal } → def b x → def a x ==-refl : { x : OD } → x == x ==-refl {x} = record { eq→ = λ x → x ; eq← = λ x → x } open _==_ ==-trans : { x y z : OD } → x == y → y == z → x == z ==-trans x=y y=z = record { eq→ = λ {m} t → eq→ y=z (eq→ x=y t) ; eq← = λ {m} t → eq← x=y (eq← y=z t) } ==-sym : { x y : OD } → x == y → y == x ==-sym x=y = record { eq→ = λ {m} t → eq← x=y t ; eq← = λ {m} t → eq→ x=y t } ⇔→== : { x y : OD } → ( {z : Ordinal } → (def x z ⇔ def y z)) → x == y eq→ ( ⇔→== {x} {y} eq ) {z} m = proj1 eq m eq← ( ⇔→== {x} {y} eq ) {z} m = proj2 eq m -- next assumptions are our axiom -- -- OD is an equation on Ordinals, so it contains Ordinals. If these Ordinals have one-to-one -- correspondence to the OD then the OD looks like a ZF Set. -- -- If all ZF Set have supreme upper bound, the solutions of OD have to be bounded, i.e. -- bbounded ODs are ZF Set. Unbounded ODs are classes. -- -- In classical Set Theory, HOD is used, as a subset of OD, -- HOD = { x | TC x ⊆ OD } -- where TC x is a transitive clusure of x, i.e. Union of all elemnts of all subset of x. -- This is not possible because we don't have V yet. So we assumes HODs are bounded OD. -- -- We also assumes HODs are isomorphic to Ordinals, which is ususally proved by Goedel number tricks. -- There two contraints on the HOD order, one is ∋, the other one is ⊂. -- ODs have an ovbious maximum, but Ordinals are not, but HOD has no maximum because there is an aribtrary -- bound on each HOD. -- -- In classical Set Theory, sup is defined by Uion, since we are working on constructive logic, -- we need explict assumption on sup. -- -- ==→o≡ is necessary to prove axiom of extensionality. -- Ordinals in OD , the maximum Ords : OD Ords = record { def = λ x → One } record HOD : Set (suc n) where field od : OD odmax : Ordinal <-osuc (c<→o< {Ords} OneObj ) -- Ordinal in OD ( and ZFSet ) Transitive Set Ord : ( a : Ordinal ) → HOD Ord a = record { od = record { def = λ y → y o< a } ; odmax = a ; ¬a ¬b c = ⊥-elim ( ¬x<0 c ) ∅6 : { x : HOD } → ¬ ( x ∋ x ) -- no Russel paradox ∅6 {x} x∋x = o<¬≡ refl ( c<→o< {x} {x} x∋x ) odef-iso : {A B : HOD } {x y : Ordinal } → x ≡ y → (odef A y → odef B y) → odef A x → odef B x odef-iso refl t = t is-o∅ : ( x : Ordinal ) → Dec ( x ≡ o∅ ) is-o∅ x with trio< x o∅ is-o∅ x | tri< a ¬b ¬c = no ¬b is-o∅ x | tri≈ ¬a b ¬c = yes b is-o∅ x | tri> ¬a ¬b c = no ¬b odef< : {b : Ordinal } { A : HOD } → odef A b → b o< & A odef< {b} {A} ab = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab)) odef∧< : {A : HOD } {y : Ordinal} {n : Level } → {P : Set n} → odef A y ∧ P → y o< & A odef∧< {A } {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p ))) -- the pair _,_ : HOD → HOD → HOD x , y = record { od = record { def = λ t → (t ≡ & x ) ∨ ( t ≡ & y ) } ; odmax = omax (& x) (& y) ; ¬a ¬b c = ⊥-elim ( o<> (⊆→o≤ {x , x} {y} y⊆x,x ) lemma1 ) where lemma : {z : Ordinal} → (z ≡ & x) ∨ (z ≡ & x) → & x ≡ z lemma (case1 refl) = refl lemma (case2 refl) = refl y⊆x,x : {z : Ordinal} → def (od (x , x)) z → def (od y) z y⊆x,x {z} lt = subst (λ k → def (od y) k ) (lemma lt) y∋x lemma1 : osuc (& y) o< & (x , x) lemma1 = subst (λ k → osuc (& y) o< k ) (sym (peq {x})) (osucc c ) ε-induction : { ψ : HOD → Set (suc n)} → ( {x : HOD } → ({ y : HOD } → x ∋ y → ψ y ) → ψ x ) → (x : HOD ) → ψ x ε-induction {ψ} ind x = subst (λ k → ψ k ) *iso (ε-induction-ord (osuc (& x)) <-osuc ) where induction : (ox : Ordinal) → ((oy : Ordinal) → oy o< ox → ψ (* oy)) → ψ (* ox) induction ox prev = ind ( λ {y} lt → subst (λ k → ψ k ) *iso (prev (& y) (o<-subst (c<→o< lt) refl &iso ))) ε-induction-ord : (ox : Ordinal) { oy : Ordinal } → oy o< ox → ψ (* oy) ε-induction-ord ox {oy} lt = TransFinite {λ oy → ψ (* oy)} induction oy -- Open supreme upper bound leads a contradition, so we use domain restriction on sup ¬open-sup : ( sup-o : (Ordinal → Ordinal ) → Ordinal) → ((ψ : Ordinal → Ordinal ) → (x : Ordinal) → ψ x o< sup-o ψ ) → ⊥ ¬open-sup sup-o sup-o< = o<> <-osuc (sup-o< next-ord (sup-o next-ord)) where next-ord : Ordinal → Ordinal next-ord x = osuc x Select : (X : HOD ) → ((x : HOD ) → Set n ) → HOD Select X ψ = record { od = record { def = λ x → ( odef X x ∧ ψ ( * x )) } ; odmax = odmax X ; ¬a ¬b c with lt (o<-subst c (sym &iso) refl ) ... | ttt = ⊥-elim ( o<¬≡ refl (o<-subst ttt &iso refl )) ψiso : {ψ : HOD → Set n} {x y : HOD } → ψ x → x ≡ y → ψ y ψiso {ψ} t refl = t selection : {ψ : HOD → Set n} {X y : HOD} → ((X ∋ y) ∧ ψ y) ⇔ (Select X ψ ∋ y) selection {ψ} {X} {y} = ⟪ ( λ cond → ⟪ proj1 cond , ψiso {ψ} (proj2 cond) (sym *iso) ⟫ ) , ( λ select → ⟪ proj1 select , ψiso {ψ} (proj2 select) *iso ⟫ ) ⟫ selection-in-domain : {ψ : HOD → Set n} {X y : HOD} → Select X ψ ∋ y → X ∋ y selection-in-domain {ψ} {X} {y} lt = proj1 ((proj2 (selection {ψ} {X} )) lt) sup-c≤ : (ψ : HOD → HOD) → {X x : HOD} → X ∋ x → & (ψ x) o≤ (sup-o X (λ y X∋y → & (ψ (* y)))) sup-c≤ ψ {X} {x} lt = subst (λ k → & (ψ k) o< _ ) *iso (sup-o≤ X lt ) --- --- Power Set --- --- First consider ordinals in HOD --- --- A ∩ x = record { def = λ y → odef A y ∧ odef x y } subset of A -- -- ∩-≡ : { a b : HOD } → ({x : HOD } → (a ∋ x → b ∋ x)) → a =h= ( b ∩ a ) ∩-≡ {a} {b} inc = record { eq→ = λ {x} x