{-# 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 open import Relation.Binary.Core open import logic open import nat open inOrdinal 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 (sup-o< next-ord (sup-o next-ord)) where next-ord : Ordinal → Ordinal next-ord x = osuc x -- 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 = no ¬b -- 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 : Ordinals.ord O} → 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 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 -- level trick (what'a shame) for LEM / minimal ε-induction1 : { ψ : HOD → Set (suc n)} → ( {x : HOD } → ({ y : HOD } → x ∋ y → ψ y ) → ψ x ) → (x : HOD ) → ψ x ε-induction1 {ψ} 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 = TransFinite1 {λ oy → ψ (* oy)} induction oy Select : (X : HOD ) → ((x : HOD ) → Set n ) → HOD Select X ψ = record { od = record { def = λ x → ( odef X x ∧ ψ ( * x )) } ; odmax = odmax X ; u ¬a ¬b c = ⊥-elim (not c) _∈_ : ( A B : HOD ) → Set n A ∈ B = B ∋ A OPwr : (A : HOD ) → HOD OPwr A = Ord ( sup-o (Ord (osuc (& A))) ( λ x A∋x → & ( A ∩ (* x)) ) ) Power : HOD → HOD Power A = Replace (OPwr (Ord (& A))) ( λ x → A ∩ x ) -- {_} : ZFSet → ZFSet -- { x } = ( x , x ) -- better to use (x , x) directly union→ : (X z u : HOD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z union→ X z u xx not = ⊥-elim ( not (& u) ( ⟪ proj1 xx , subst ( λ k → odef k (& z)) (sym *iso) (proj2 xx) ⟫ )) union← : (X z : HOD) (X∋z : Union X ∋ z) → ¬ ( (u : HOD ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) union← X z UX∋z = FExists _ lemma UX∋z where lemma : {y : Ordinal} → odef X y ∧ odef (* y) (& z) → ¬ ((u : HOD) → ¬ (X ∋ u) ∧ (u ∋ z)) lemma {y} xx not = not (* y) ⟪ d→∋ X (proj1 xx) , proj2 xx ⟫ data infinite-d : ( x : Ordinal ) → Set n where iφ : infinite-d o∅ isuc : {x : Ordinal } → infinite-d x → infinite-d (& ( Union (* x , (* x , * x ) ) )) -- ω can be diverged in our case, since we have no restriction on the corresponding ordinal of a pair. -- We simply assumes infinite-d y has a maximum. -- -- This means that many of OD may not be HODs because of the & mapping divergence. -- We should have some axioms to prevent this such as & x o< next (odmax x). -- -- postulate -- ωmax : Ordinal -- <ωmax : {y : Ordinal} → infinite-d y → y o< ωmax -- -- infinite : HOD -- infinite = record { od = record { def = λ x → infinite-d x } ; odmax = ωmax ; ¬a ¬b c with (incl 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 ) replacement← : {ψ : HOD → HOD} (X x : HOD) → X ∋ x → Replace X ψ ∋ ψ x replacement← {ψ} X x lt = ⟪ sup-c< ψ {X} {x} lt , lemma ⟫ where lemma : def (in-codomain X ψ) (& (ψ x)) lemma not = ⊥-elim ( not ( & x ) ⟪ lt , cong (λ k → & (ψ k)) (sym *iso)⟫ ) replacement→ : {ψ : HOD → HOD} (X x : HOD) → (lt : Replace X ψ ∋ x) → ¬ ( (y : HOD) → ¬ (x =h= ψ y)) replacement→ {ψ} X x lt = contra-position lemma (lemma2 (proj2 lt)) where lemma2 : ¬ ((y : Ordinal) → ¬ odef X y ∧ ((& x) ≡ & (ψ (* y)))) → ¬ ((y : Ordinal) → ¬ odef X y ∧ (* (& x) =h= ψ (* y))) lemma2 not not2 = not ( λ y d → not2 y ⟪ proj1 d , lemma3 (proj2 d)⟫) where lemma3 : {y : Ordinal } → (& x ≡ & (ψ (* y))) → (* (& x) =h= ψ (* y)) lemma3 {y} eq = subst (λ k → * (& x) =h= k ) *iso (o≡→== eq ) lemma : ( (y : HOD) → ¬ (x =h= ψ y)) → ( (y : Ordinal) → ¬ odef X y ∧ (* (& x) =h= ψ (* y)) ) lemma not y not2 = not (* y) (subst (λ k → k =h= ψ (* y)) *iso ( proj2 not2 )) --- --- 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