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 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 id : {A : Set n} → A → A id x = x ==-refl : { x : OD } → x == x ==-refl {x} = record { eq→ = id ; eq← = id } 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 -- 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. -- We simply assume V=OD here. -- -- We also assumes ODs are isomorphic to Ordinals, which is ususally proved by Goedel number tricks. -- ODs have an ovbious maximum, but Ordinals are not. This means, od→ord is not an on-to mapping. -- -- ==→o≡ is necessary to prove axiom of extensionality. -- -- In classical Set Theory, sup is defined by Uion. Since we are working on constructive logic, -- we need explict assumption on sup. record ODAxiom : Set (suc n) where -- OD can be iso to a subset of Ordinal ( by means of Godel Set ) field maxod : Ordinal od→ord : OD → Ordinal ord→od : (x : Ordinal ) → x o< maxod → OD -- ZFSet has bounded solution of OD o ¬a ¬b c = ⊥-elim (¬x<0 c) o∅≡od∅ : ord→od o∅ o∅ ¬a ¬b c = no ¬b _,_ : OD → OD → OD x , y = record { def = λ t → (t ≡ od→ord x ) ∨ ( t ≡ od→ord y ) } -- Ord (omax (od→ord x) (od→ord y)) -- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) -- postulate f-extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality n (suc n) in-codomain : (X : OD ) → ( ψ : OD → OD ) → OD in-codomain X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( (lt : def X y ) → (lt : y o< maxod) → ( x ≡ od→ord (ψ (ord→od y lt ))))) } where -- Power Set of X ( or constructible by λ y → def X (od→ord y ) ZFSubset : (A x : OD ) → OD ZFSubset A x = record { def = λ y → def A y ∧ def x y } -- roughly x = A → Set Def : (A : OD ) → OD Def A = Ord ( sup-o ( λ x → od→ord ( ZFSubset A x) ) ) -- _⊆_ : ( A B : OD ) → ∀{ x : OD } → Set n -- _⊆_ A B {x} = A ∋ x → B ∋ x record _⊆_ ( A B : OD ) : Set (suc n) where field incl : { x : OD } → A ∋ x → B ∋ x open _⊆_ infixr 220 _⊆_ subset-lemma : {A x : OD } → ( {y : OD } → x ∋ y → ZFSubset A x ∋ y ) ⇔ ( x ⊆ A ) subset-lemma {A} {x} = record { proj1 = λ lt → record { incl = λ x∋z → proj1 (lt x∋z) } ; proj2 = λ x⊆A lt → record { proj1 = incl x⊆A lt ; proj2 = lt } } open import Data.Unit ε-induction : { ψ : OD → Set (suc n)} → ( {x : OD } → ({ y : OD } → x ∋ y → ψ y ) → ψ x ) → (x : OD ) → ψ x ε-induction {ψ} ind x = subst (λ k → ψ k ) oiso (ε-induction-ord (osuc (od→ord x)) <-osuc ) where induction : (ox : Ordinal) → ((oy : Ordinal) → oy o< ox → ψ (ord→od oy {!!} )) → ψ (ord→od ox {!!} ) induction ox prev = ind ( λ {y} lt → subst (λ k → ψ k ) oiso (prev (od→ord y) (o<-subst (c<→o< lt) refl (diso {!!}) ))) ε-induction-ord : (ox : Ordinal) { oy : Ordinal } → oy o< ox → ψ (ord→od oy {!!} ) ε-induction-ord ox {oy} lt = TransFinite {λ oy → ψ (ord→od oy {!!} )} induction oy -- minimal-2 : (x : OD ) → ( ne : ¬ (x == od∅ ) ) → (y : OD ) → ¬ ( def (minimal x ne) (od→ord y)) ∧ (def x (od→ord y) ) -- minimal-2 x ne y = contra-position ( ε-induction (λ {z} ind → {!!} ) x ) ( λ p → {!!} ) OD→ZF : ZF OD→ZF = record { ZFSet = OD ; _∋_ = _∋_ ; _≈_ = _==_ ; ∅ = od∅ ; _,_ = _,_ ; Union = Union ; Power = Power ; Select = Select ; Replace = Replace ; infinite = infinite ; isZF = isZF } where ZFSet = OD -- is less than Ords because of maxod Select : (X : OD ) → ((x : OD ) → Set n ) → OD Select X ψ = record { def = λ x → ( def X x ∧ ψ ( ord→od x {!!} )) } Replace : OD → (OD → OD ) → OD Replace X ψ = record { def = λ x → (x o< sup-o ( λ x → od→ord (ψ x))) ∧ def (in-codomain X ψ) x } _∩_ : ( A B : ZFSet ) → ZFSet A ∩ B = record { def = λ x → def A x ∧ def B x } Union : OD → OD Union U = record { def = λ x → ¬ (∀ (u : Ordinal ) → ¬ ((def U u) ∧ (def (ord→od u {!!} ) x))) } _∈_ : ( A B : ZFSet ) → Set n A ∈ B = B ∋ A Power : OD → OD Power A = Replace (Def (Ord (od→ord A))) ( λ x → A ∩ x ) -- {_} : ZFSet → ZFSet -- { x } = ( x , x ) -- it works but we don't use data infinite-d : ( x : Ordinal ) → Set n where iφ : infinite-d o∅ isuc : {x : Ordinal } → infinite-d x → infinite-d (od→ord ( Union (ord→od x {!!} , (ord→od x {!!} , ord→od x {!!} ) ) )) infinite : OD infinite = record { def = λ x → infinite-d x } infixr 200 _∈_ -- infixr 230 _∩_ _∪_ isZF : IsZF (OD ) _∋_ _==_ od∅ _,_ Union Power Select Replace infinite isZF = record { isEquivalence = record { refl = ==-refl ; sym = ==-sym; trans = ==-trans } ; pair→ = pair→ ; pair← = pair← ; union→ = union→ ; union← = union← ; empty = empty ; power→ = power→ ; power← = power← ; extensionality = λ {A} {B} {w} → extensionality {A} {B} {w} ; ε-induction = ε-induction ; infinity∅ = infinity∅ ; infinity = infinity ; selection = λ {X} {ψ} {y} → selection {X} {ψ} {y} ; replacement← = replacement← ; replacement→ = replacement→ -- ; choice-func = choice-func -- ; choice = choice } where pair→ : ( x y t : ZFSet ) → (x , y) ∋ t → ( t == x ) ∨ ( t == y ) pair→ x y t (case1 t≡x ) = case1 (subst₂ (λ j k → j == k ) oiso oiso (o≡→== t≡x {!!} {!!} )) pair→ x y t (case2 t≡y ) = case2 (subst₂ (λ j k → j == k ) oiso oiso (o≡→== t≡y {!!} {!!} )) pair← : ( x y t : ZFSet ) → ( t == x ) ∨ ( t == y ) → (x , y) ∋ t pair← x y t (case1 t==x) = case1 (cong (λ k → od→ord k ) (==→o≡ t==x)) pair← x y t (case2 t==y) = case2 (cong (λ k → od→ord k ) (==→o≡ t==y)) empty : (x : OD ) → ¬ (od∅ ∋ x) empty x = ¬x<0 o<→c< : {x y : Ordinal } → x o< y → (Ord x) ⊆ (Ord y) o<→c< lt = record { incl = λ z → ordtrans z lt } ⊆→o< : {x y : Ordinal } → (Ord x) ⊆ (Ord y) → x o< osuc y ⊆→o< {x} {y} lt with trio< x y ⊆→o< {x} {y} lt | tri< a ¬b ¬c = ordtrans a <-osuc ⊆→o< {x} {y} lt | tri≈ ¬a b ¬c = subst ( λ k → k o< osuc y) (sym b) <-osuc ⊆→o< {x} {y} lt | tri> ¬a ¬b c with (incl lt) (o<-subst c (sym (diso {!!})) refl ) ... | ttt = ⊥-elim ( o<¬≡ refl (o<-subst ttt (diso {!!}) refl )) union→ : (X z u : OD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z union→ X z u xx not = ⊥-elim ( not (od→ord u) ( record { proj1 = proj1 xx ; proj2 = subst ( λ k → def k (od→ord z)) (sym oiso) (proj2 xx) } )) union← : (X z : OD) (X∋z : Union X ∋ z) → ¬ ( (u : OD ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) union← X z UX∋z = FExists _ lemma UX∋z where lemma : {y : Ordinal} → def X y ∧ def (ord→od y {!!} ) (od→ord z) → ¬ ((u : OD) → ¬ (X ∋ u) ∧ (u ∋ z)) lemma {y} xx not = not (ord→od y {!!} ) record { proj1 = subst ( λ k → def X k ) (sym (diso {!!})) (proj1 xx ) ; proj2 = proj2 xx } ψiso : {ψ : OD → Set n} {x y : OD } → ψ x → x ≡ y → ψ y ψiso {ψ} t refl = t selection : {ψ : OD → Set 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 } } replacement← : {ψ : OD → OD} (X x : OD) → X ∋ x → Replace X ψ ∋ ψ x replacement← {ψ} X x lt = record { proj1 = sup-c< ψ {x} ; proj2 = lemma } where lemma : def (in-codomain X ψ) (od→ord (ψ x)) lemma not = ⊥-elim ( not ( od→ord x ) {!!} ) -- (record { proj1 = lt ; proj2 = cong (λ k → od→ord (ψ k)) (sym oiso)} )) replacement→ : {ψ : OD → OD} (X x : OD) → (lt : Replace X ψ ∋ x) → ¬ ( (y : OD) → ¬ (x == ψ y)) replacement→ {ψ} X x lt = contra-position lemma (lemma2 {!!}) where lemma2 : ¬ ((y : Ordinal) → ¬ def X y ∧ ((od→ord x) ≡ od→ord (ψ (ord→od y {!!} )))) → ¬ ((y : Ordinal) → ¬ def X y ∧ (ord→od (od→ord x) {!!} == ψ (ord→od y {!!} ))) lemma2 not not2 = not ( λ y d → not2 y (record { proj1 = proj1 d ; proj2 = lemma3 (proj2 d)})) where lemma3 : {y : Ordinal } → (od→ord x ≡ od→ord (ψ (ord→od y {!!} ))) → (ord→od (od→ord x) {!!} == ψ (ord→od y {!!} )) lemma3 {y} eq = subst (λ k → ord→od (od→ord x) {!!} == k ) oiso (o≡→== eq {!!} {!!} ) lemma : ( (y : OD) → ¬ (x == ψ y)) → ( (y : Ordinal) → ¬ def X y ∧ (ord→od (od→ord x) {!!} == ψ (ord→od y {!!} )) ) lemma not y not2 = not (ord→od y {!!} ) (subst (λ k → k == ψ (ord→od y {!!} )) oiso ( proj2 not2 )) --- --- Power Set --- --- First consider ordinals in OD --- --- ZFSubset A x = record { def = λ y → def A y ∧ def x y } subset of A -- -- ∩-≡ : { a b : OD } → ({x : OD } → (a ∋ x → b ∋ x)) → a == ( b ∩ a ) ∩-≡ {a} {b} inc = record { eq→ = λ {x} x