diff ordinal-definable.agda @ 29:fce60b99dc55

posturate OD is isomorphic to Ordinal
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Mon, 20 May 2019 18:18:43 +0900
parents constructible-set.agda@f36e40d5d2c3
children 3b0fdb95618e
line wrap: on
line diff
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/ordinal-definable.agda	Mon May 20 18:18:43 2019 +0900
@@ -0,0 +1,151 @@
+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
+
+
+-- X' = { x ∈ X |  ψ  x } ∪ X , Mα = ( ∪ (β < α) Mβ ) '
+
+-- Ordinal Definable Set
+
+-- o∋  : {n : Level} → {A : Ordinal {n}} → (OrdinalDefinable {n} A ) → (x : Ordinal {n} ) → (x o< A) → Set n
+-- o∋ a x x<A  = def a x x<A
+
+-- TC u : Transitive Closure of OD u
+--
+--    all elements of u or elements of elements of u, etc...
+--
+-- TC Zero = u
+-- TC (suc n) = ∪ (TC n)
+--
+-- TC u = TC ω u = ∪ ( TC n ) n ∈ ω
+-- 
+--     u ∪ ( ∪ u ) ∪ ( ∪ (∪ u ) ) ....
+--
+-- Heritic Ordinal Definable
+--
+--    ( HOD = {x | TC x ⊆ OD } ) ⊆ OD     x ∈ OD here
+-- 
+
+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} 
+ord→od x = record { def = λ y → x ≡ y }
+
+_∋_ : { 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 )
+
+postulate      
+  c<→o< : {n : Level} {x y : OD {n} } → x c< y → od→ord x o< od→ord x
+  o<→c< : {n : Level} {x y : Ordinal {n} } → x o< y → ord→od x c< ord→od x
+  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 ψ
+
+HOD = OD
+
+od∅ : {n : Level} → HOD {n} 
+od∅ {n} = record { def = λ _ → Lift n ⊥ }
+
+∅1 : {n : Level} →  ( x : OD {n} )  → ¬ ( x c< od∅ {n} )
+∅1 {n} x (lift ())
+
+HOD→ZF : {n : Level} → ZF {suc n} {suc n}
+HOD→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 = {!!}
+       ;   regularity = {!!}
+       ;   infinity∅ = {!!}
+       ;   infinity = {!!}
+       ;   selection = {!!}
+       ;   replacement = {!!}
+     } where
+         open _∧_ 
+         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 lemma  where
+            lemma : {z : Ordinal {n} } → def X z → z ≡ od→ord y
+            lemma {z} X∋z = {!!}
+
+
+
+
+
+