changeset 149:e022c0716936 release

only ordinal-definable.agda is finished. it assmues all ZF Set are Ordinals. HOD is incomplete, but we leave this for a while.
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Mon, 08 Jul 2019 19:48:47 +0900
parents 02d421f1cc06 (current diff) 6e767ad3edc2 (diff)
children a1b5b890b796
files
diffstat 6 files changed, 1315 insertions(+), 45 deletions(-) [+]
line wrap: on
line diff
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/HOD.agda	Mon Jul 08 19:48:47 2019 +0900
@@ -0,0 +1,511 @@
+open import Level
+module HOD 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
+
+-- Ordinal Definable Set
+
+record OD {n : Level}  : Set (suc n) where
+  field
+    def : (x : Ordinal {n} ) → Set n
+
+open OD
+open import Data.Unit
+
+open Ordinal
+open _∧_
+
+record _==_ {n : Level} ( a b :  OD {n} ) : Set n where
+  field
+     eq→ : ∀ { x : Ordinal {n} } → def a x → def b x 
+     eq← : ∀ { x : Ordinal {n} } → def b x → def a x 
+
+id : {n : Level} {A : Set n} → A → A
+id x = x
+
+eq-refl : {n : Level} {  x :  OD {n} } → x == x
+eq-refl {n} {x} = record { eq→ = id ; eq← = id }
+
+open  _==_ 
+
+eq-sym : {n : Level} {  x y :  OD {n} } → x == y → y == x
+eq-sym eq = record { eq→ = eq← eq ; eq← = eq→ eq }
+
+eq-trans : {n : Level} {  x y z :  OD {n} } → x == y → y == z → x == z
+eq-trans x=y y=z = record { eq→ = λ t → eq→ y=z ( eq→ x=y  t) ; eq← = λ t → eq← x=y ( eq← y=z t) }
+
+⇔→== : {n : Level} {  x y :  OD {suc n} } → ( {z : Ordinal {suc n}} → def x z ⇔  def y z) → x == y 
+eq→ ( ⇔→== {n} {x} {y}  eq ) {z} m = proj1 eq m 
+eq← ( ⇔→== {n} {x} {y}  eq ) {z} m = proj2 eq m 
+
+-- Ordinal in OD ( and ZFSet )
+Ord : { n : Level } → ( a : Ordinal {n} ) → OD {n}
+Ord {n} a = record { def = λ y → y o< a }  
+
+od∅ : {n : Level} → OD {n} 
+od∅ {n} = Ord o∅ 
+
+postulate      
+  -- OD can be iso to a subset of Ordinal ( by means of Godel Set )
+  od→ord : {n : Level} → OD {n} → Ordinal {n}
+  ord→od : {n : Level} → Ordinal {n} → OD {n} 
+  c<→o<  : {n : Level} {x y : OD {n} }      → def y ( od→ord x ) → od→ord x o< od→ord y
+  oiso   : {n : Level} {x : OD {n}}     → ord→od ( od→ord x ) ≡ x
+  diso   : {n : Level} {x : Ordinal {n}} → od→ord ( ord→od x ) ≡ x
+  ==→o≡ : {n : Level} →  { x y : OD {suc n} } → (x == y) → x ≡ y
+  -- next assumption causes ∀ x ∋ ∅ . It menas only an ordinal becomes a set
+  -- o<→c<  : {n : Level} {x y : Ordinal {n} } → x o< y             → def (ord→od y) x 
+  -- supermum as Replacement Axiom
+  sup-o  : {n : Level } → ( Ordinal {n} → Ordinal {n}) →  Ordinal {n}
+  sup-o< : {n : Level } → { ψ : Ordinal {n} →  Ordinal {n}} → ∀ {x : Ordinal {n}} →  ψ x  o<  sup-o ψ 
+  -- contra-position of mimimulity of supermum required in Power Set Axiom
+  sup-x  : {n : Level } → ( Ordinal {n} → Ordinal {n}) →  Ordinal {n}
+  sup-lb : {n : Level } → { ψ : Ordinal {n} →  Ordinal {n}} → {z : Ordinal {n}}  →  z o< sup-o ψ → z o< osuc (ψ (sup-x ψ))
+  -- sup-lb : {n : Level } → ( ψ : Ordinal {n} →  Ordinal {n}) → ( ∀ {x : Ordinal {n}} →  ψx  o<  z ) →  z o< osuc ( sup-o ψ ) 
+  minimul : {n : Level } → (x : OD {suc n} ) → ¬ (x == od∅ )→ OD {suc n} 
+  -- this should be ¬ (x == od∅ )→ ∃ ox → x ∋ Ord ox  ( minimum of x )
+  x∋minimul : {n : Level } → (x : OD {suc n} ) → ( ne : ¬ (x == od∅ ) ) → def x ( od→ord ( minimul x ne ) )
+  minimul-1 : {n : Level } → (x : OD {suc n} ) → ( ne : ¬ (x == od∅ ) ) → (y : OD {suc n}) → ¬ ( def (minimul x ne) (od→ord y)) ∧ (def x (od→ord  y) )
+  -- we should prove this in agda, but simply put here
+  ===-≡ : {n : Level} { x y : OD {suc n}} → x == y  → x ≡ y
+
+_∋_ : { n : Level } → ( a x : OD {n} ) → Set n
+_∋_ {n} a x  = def a ( od→ord x )
+
+_c<_ : { n : Level } → ( x a : 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 )
+
+cseq : {n : Level} →  OD {n} →  OD {n}
+cseq x = record { def = λ y → def x (osuc y) } where
+
+def-subst : {n : Level } {Z : OD {n}} {X : Ordinal {n} }{z : OD {n}} {x : Ordinal {n} }→ def Z X → Z ≡ z  →  X ≡ x  →  def z x
+def-subst df refl refl = df
+
+sup-od : {n : Level } → ( OD {n} → OD {n}) →  OD {n}
+sup-od ψ = Ord ( sup-o ( λ x → od→ord (ψ (ord→od x ))) )
+
+sup-c< : {n : Level } → ( ψ : OD {n} →  OD {n}) → ∀ {x : OD {n}} → def ( sup-od ψ ) (od→ord ( ψ x ))
+sup-c< {n} ψ {x} = def-subst {n} {_} {_} {Ord ( sup-o ( λ x → od→ord (ψ (ord→od x ))) )} {od→ord ( ψ x )}
+        lemma refl (cong ( λ k → od→ord (ψ k) ) oiso) where
+    lemma : od→ord (ψ (ord→od (od→ord x))) o< sup-o (λ x → od→ord (ψ (ord→od x)))
+    lemma = subst₂ (λ j k → j o< k ) refl diso (o<-subst sup-o< refl (sym diso)  )
+
+otrans : {n : Level} {a x : Ordinal {n} } → def (Ord a) x → { y : Ordinal {n} } → y o< x → def (Ord a) y
+otrans {n} {a} {x} x<a {y} y<x = ordtrans y<x x<a
+
+∅3 : {n : Level} →  { x : Ordinal {n}}  → ( ∀(y : Ordinal {n}) → ¬ (y o< x ) ) → x ≡ o∅ {n}
+∅3 {n} {x} = TransFinite {n} c2 c3 x where
+   c0 : Nat →  Ordinal {n}  → Set n
+   c0 lx x = (∀(y : Ordinal {n}) → ¬ (y o< x))  → x ≡ o∅ {n}
+   c2 : (lx : Nat) → c0 lx (record { lv = lx ; ord = Φ lx } )
+   c2 Zero not = refl
+   c2 (Suc lx) not with not ( record { lv = lx ; ord = Φ lx } )
+   ... | t with t (case1 ≤-refl )
+   c2 (Suc lx) not | t | ()
+   c3 : (lx : Nat) (x₁ : OrdinalD lx) → c0 lx  (record { lv = lx ; ord = x₁ })  → c0 lx (record { lv = lx ; ord = OSuc lx x₁ })
+   c3 lx (Φ .lx) d not with not ( record { lv = lx ; ord = Φ lx } )
+   ... | t with t (case2 Φ< )
+   c3 lx (Φ .lx) d not | t | ()
+   c3 lx (OSuc .lx x₁) d not with not (  record { lv = lx ; ord = OSuc lx x₁ } )
+   ... | t with t (case2 (s< s<refl ) )
+   c3 lx (OSuc .lx x₁) d not | t | ()
+
+∅5 : {n : Level} →  { x : Ordinal {n} }  → ¬ ( x  ≡ o∅ {n} ) → o∅ {n} o< x
+∅5 {n} {record { lv = Zero ; ord = (Φ .0) }} not = ⊥-elim (not refl) 
+∅5 {n} {record { lv = Zero ; ord = (OSuc .0 ord) }} not = case2 Φ<
+∅5 {n} {record { lv = (Suc lv) ; ord = ord }} not = case1 (s≤s z≤n)
+
+ord-iso : {n : Level} {y : Ordinal {n} } → record { lv = lv (od→ord (ord→od y)) ; ord = ord (od→ord (ord→od y)) } ≡ record { lv = lv y ; ord = ord y }
+ord-iso = cong ( λ k → record { lv = lv k ; ord = ord k } ) diso
+
+-- avoiding lv != Zero error
+orefl : {n : Level} →  { x : OD {n} } → { y : Ordinal {n} } → od→ord x ≡ y → od→ord x ≡ y
+orefl refl = refl
+
+==-iso : {n : Level} →  { x y : OD {n} } → ord→od (od→ord x) == ord→od (od→ord y)  →  x == y
+==-iso {n} {x} {y} eq = record {
+      eq→ = λ d →  lemma ( eq→  eq (def-subst d (sym oiso) refl )) ;
+      eq← = λ d →  lemma ( eq←  eq (def-subst d (sym oiso) refl ))  }
+        where
+           lemma : {x : OD {n} } {z : Ordinal {n}} → def (ord→od (od→ord x)) z → def x z
+           lemma {x} {z} d = def-subst d oiso refl
+
+=-iso : {n : Level } {x y : OD {suc n} } → (x == y) ≡ (ord→od (od→ord x) == y)
+=-iso {_} {_} {y} = cong ( λ k → k == y ) (sym oiso)
+
+ord→== : {n : Level} →  { x y : OD {n} } → od→ord x ≡  od→ord y →  x == y
+ord→== {n} {x} {y} eq = ==-iso (lemma (od→ord x) (od→ord y) (orefl eq)) where
+   lemma : ( ox oy : Ordinal {n} ) → ox ≡ oy →  (ord→od ox) == (ord→od oy)
+   lemma ox ox  refl = eq-refl
+
+o≡→== : {n : Level} →  { x y : Ordinal {n} } → x ≡  y →  ord→od x == ord→od y
+o≡→== {n} {x} {.x} refl = eq-refl
+
+>→¬< : {x y : Nat  } → (x < y ) → ¬ ( y < x )
+>→¬< (s≤s x<y) (s≤s y<x) = >→¬< x<y y<x
+
+c≤-refl : {n : Level} →  ( x : OD {n} ) → x c≤ x
+c≤-refl x = case1 refl
+
+∋→o< : {n : Level} →  { a x : OD {suc n} } → a ∋ x → od→ord x o< od→ord a
+∋→o< {n} {a} {x} lt = t where
+         t : (od→ord x) o< (od→ord a)
+         t = c<→o< {suc n} {x} {a} lt 
+
+-- o∅≡od∅ : {n : Level} → ord→od (o∅ {suc n}) ≡ od∅ {suc n}
+
+o<→¬c> : {n : Level} →  { x y : OD {n} } → (od→ord x ) o< ( od→ord y) →  ¬ (y c< x )
+o<→¬c> {n} {x} {y} olt clt = o<> olt (c<→o< clt ) where
+
+o≡→¬c< : {n : Level} →  { x y : OD {n} } →  (od→ord x ) ≡ ( od→ord y) →   ¬ x c< y
+o≡→¬c< {n} {x} {y} oeq lt  = o<¬≡  (orefl oeq ) (c<→o< lt) 
+
+∅0 : {n : Level} →  record { def = λ x →  Lift n ⊥ } == od∅ {n} 
+eq→ ∅0 {w} (lift ())
+eq← ∅0 {w} (case1 ())
+eq← ∅0 {w} (case2 ())
+
+∅< : {n : Level} →  { x y : OD {n} } → def x (od→ord y ) → ¬ (  x  == od∅ {n} )
+∅< {n} {x} {y} d eq with eq→ (eq-trans eq (eq-sym ∅0) ) d
+∅< {n} {x} {y} d eq | lift ()
+       
+∅6 : {n : Level} →  { x : OD {suc n} }  → ¬ ( x ∋ x )    --  no Russel paradox
+∅6 {n} {x} x∋x = o<¬≡ refl ( c<→o< {suc n} {x} {x} x∋x )
+
+def-iso : {n : Level} {A B : OD {n}} {x y : Ordinal {n}} → x ≡ y  → (def A y → def B y)  → def A x → def B x
+def-iso refl t = t
+
+is-o∅ : {n : Level} →  ( x : Ordinal {suc n} ) → Dec ( x ≡ o∅ {suc n} )
+is-o∅ {n} record { lv = Zero ; ord = (Φ .0) } = yes refl
+is-o∅ {n} record { lv = Zero ; ord = (OSuc .0 ord₁) } = no ( λ () )
+is-o∅ {n} record { lv = (Suc lv₁) ; ord = ord } = no (λ())
+
+
+-- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) 
+-- postulate f-extensionality : { n : Level}  → Relation.Binary.PropositionalEquality.Extensionality (suc n) (suc (suc n))
+
+in-codomain : {n : Level} → (X : OD {suc n} ) → ( ψ : OD {suc n} → OD {suc n} ) → OD {suc n}
+in-codomain {n} X ψ = record { def = λ x → ¬ ( (y : Ordinal {suc n}) → ¬ ( def X y ∧  ( x ≡ od→ord (ψ (ord→od y )))))  }
+
+-- Power Set of X ( or constructible by λ y → def X (od→ord y )
+
+ZFSubset : {n : Level} → (A x : OD {suc n} ) → OD {suc n}
+ZFSubset A x =  record { def = λ y → def A y ∧  def x y }   where
+
+Def :  {n : Level} → (A :  OD {suc n}) → OD {suc n}
+Def {n} A = Ord ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) )  
+
+OrdSubset : {n : Level} → (A x : Ordinal {suc n} ) → ZFSubset (Ord A) (Ord x) ≡ Ord ( minα A x )
+OrdSubset {n} A x = ===-≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where
+  lemma1 :  {y : Ordinal} → def (ZFSubset (Ord A) (Ord x)) y → def (Ord (minα A x)) y
+  lemma1 {y} s with trio< A x
+  lemma1 {y} s | tri< a ¬b ¬c = proj1 s
+  lemma1 {y} s | tri≈ ¬a refl ¬c = proj1 s
+  lemma1 {y} s | tri> ¬a ¬b c = proj2 s
+  lemma2 : {y : Ordinal} → def (Ord (minα A x)) y → def (ZFSubset (Ord A) (Ord x)) y
+  lemma2 {y} lt with trio< A x
+  lemma2 {y} lt | tri< a ¬b ¬c = record { proj1 = lt ; proj2 = ordtrans lt a }
+  lemma2 {y} lt | tri≈ ¬a refl ¬c = record { proj1 = lt ; proj2 = lt }
+  lemma2 {y} lt | tri> ¬a ¬b c = record { proj1 = ordtrans lt c ; proj2 = lt }
+
+-- Constructible Set on α
+-- Def X = record { def = λ y → ∀ (x : OD ) → y < X ∧ y <  od→ord x } 
+-- L (Φ 0) = Φ
+-- L (OSuc lv n) = { Def ( L n )  } 
+-- L (Φ (Suc n)) = Union (L α) (α < Φ (Suc n) )
+L : {n : Level} → (α : Ordinal {suc n}) → OD {suc n}
+L {n}  record { lv = Zero ; ord = (Φ .0) } = od∅
+L {n}  record { lv = lx ; ord = (OSuc lv ox) } = Def ( L {n} ( record { lv = lx ; ord = ox } ) ) 
+L {n}  record { lv = (Suc lx) ; ord = (Φ (Suc lx)) } = -- Union ( L α )
+    cseq ( Ord (od→ord (L {n}  (record { lv = lx ; ord = Φ lx }))))
+
+-- L0 :  {n : Level} → (α : Ordinal {suc n}) → α o< β → L (osuc α) ∋ L α
+-- L1 :  {n : Level} → (α β : Ordinal {suc n}) → α o< β → ∀ (x : OD {suc n})  → L α ∋ x → L β ∋ x 
+
+omega : { n : Level } → Ordinal {n}
+omega = record { lv = Suc Zero ; ord = Φ 1 }
+
+OD→ZF : {n : Level} → ZF {suc (suc n)} {suc n}
+OD→ZF {n}  = record { 
+    ZFSet = OD {suc n}
+    ; _∋_ = _∋_ 
+    ; _≈_ = _==_ 
+    ; ∅  = od∅
+    ; _,_ = _,_
+    ; Union = Union
+    ; Power = Power
+    ; Select = Select
+    ; Replace = Replace
+    ; infinite = Ord omega
+    ; isZF = isZF 
+ } where
+    ZFSet = OD {suc n}
+    Select : (X : OD {suc n} ) → ((x : OD {suc n} ) → Set (suc n) ) → OD {suc n}
+    Select X ψ = record { def = λ x →  ( def X  x ∧  ψ ( ord→od x )) }
+    Replace : OD {suc n} → (OD {suc n} → OD {suc n} ) → OD {suc n}
+    Replace X ψ = record { def = λ x → (x o< sup-o ( λ x → od→ord (ψ (ord→od x )))) ∧ def (in-codomain X ψ) x }
+    _,_ : OD {suc n} → OD {suc n} → OD {suc n}
+    x , y = Ord (omax (od→ord x) (od→ord y))
+    _∩_ : ( A B : ZFSet  ) → ZFSet
+    A ∩ B = record { def = λ x → def A x ∧ def B x } 
+    Union : OD {suc n} → OD {suc n}
+    Union U = record { def = λ y  → def U (osuc y) }
+    _∈_ : ( A B : ZFSet  ) → Set (suc n)
+    A ∈ B = B ∋ A
+    _⊆_ : ( A B : ZFSet  ) → ∀{ x : ZFSet } →  Set (suc n)
+    _⊆_ A B {x} = A ∋ x →  B ∋ x
+    Power : OD {suc n} → OD {suc n}
+    Power A = Replace (Def (Ord (od→ord A))) ( λ x → A ∩ x )
+    {_} : ZFSet → ZFSet
+    { x } = ( x ,  x )
+
+    infixr  200 _∈_
+    -- infixr  230 _∩_ _∪_
+    infixr  220 _⊆_
+    isZF : IsZF (OD {suc n})  _∋_  _==_ od∅ _,_ Union Power Select Replace (Ord omega)
+    isZF = record {
+           isEquivalence  = record { refl = eq-refl ; sym = eq-sym; trans = eq-trans }
+       ;   pair  = pair
+       ;   union-u = λ X z UX∋z → union-u {X} {z} UX∋z
+       ;   union→ = union→
+       ;   union← = union←
+       ;   empty = empty
+       ;   power→ = power→  
+       ;   power← = power← 
+       ;   extensionality = extensionality
+       ;   minimul = minimul
+       ;   regularity = regularity
+       ;   infinity∅ = infinity∅
+       ;   infinity = λ _ → infinity
+       ;   selection = λ {X} {ψ} {y} → selection {X} {ψ} {y}
+       ;   replacement← = replacement←
+       ;   replacement→ = replacement→
+     } where
+
+         pair : (A B : OD {suc n} ) → ((A , B) ∋ A) ∧  ((A , B) ∋ B)
+         proj1 (pair A B ) = omax-x {n} (od→ord A) (od→ord B)
+         proj2 (pair A B ) = omax-y {n} (od→ord A) (od→ord B)
+
+         empty : (x : OD {suc n} ) → ¬  (od∅ ∋ x)
+         empty x (case1 ())
+         empty x (case2 ())
+
+         union-d : (X : OD {suc n}) → OD {suc n}
+         union-d X = record { def = λ y → def X (osuc y) }
+         union-u : {X z : OD {suc n}} → (U>z : Union X ∋ z ) → OD {suc n}
+         union-u {X} {z} U>z = Ord ( osuc ( od→ord z ) )
+         union→ :  (X z u : OD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z
+         union→ X z u xx with trio< ( od→ord u ) ( osuc ( od→ord z ))
+         union→ X z u xx | tri< a ¬b ¬c with  osuc-< a (c<→o< (proj2 xx))
+         union→ X z u xx | tri< a ¬b ¬c | ()
+         union→ X z u xx | tri≈ ¬a b ¬c =  def-subst {suc n} {_} {_} {X} {osuc (od→ord z)} (proj1 xx) refl b
+         union→ X z u xx | tri> ¬a ¬b c =  def-subst lemma1 (sym lemma0) diso where
+             lemma0 : X ≡ Ord (od→ord X)
+             lemma0 = sym {!!}
+             lemma : osuc (od→ord z) o< od→ord X
+             lemma = ordtrans c ( c<→o< ( proj1 xx ) )
+             lemma1 : Ord ( od→ord X) ∋ ord→od (osuc (od→ord z) )
+             lemma1 = o<-subst lemma (sym diso) refl
+         union← :  (X z : OD) (X∋z : Union X ∋ z) → (X ∋  union-u {X} {z} X∋z ) ∧ (union-u {X} {z} X∋z ∋ z )
+         union← X z UX∋z = record { proj1 = lemma ; proj2 = <-osuc } where
+             lemma : X ∋ union-u {X} {z} UX∋z
+             lemma = def-subst {suc n} {_} {_} {X} {od→ord (Ord (osuc ( od→ord z )))} UX∋z refl {!!}
+
+         ψiso :  {ψ : OD {suc n} → Set (suc n)} {x y : OD {suc n}} → ψ x → x ≡ y   → ψ y
+         ψiso {ψ} t refl = t
+         selection : {ψ : OD → Set (suc 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))
+                 {!!} } ))
+         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 y))))
+                    → ¬ ((y : Ordinal) → ¬ def X y ∧ (ord→od (od→ord x) == ψ (Ord 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 y))) → (ord→od (od→ord x) == ψ (Ord 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 y)) )
+            lemma not y not2 = not (Ord y) (subst (λ k → k == ψ (Ord 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
+         --- Power X = ord→od ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) )       Power X is a sup of all subset of A
+         --
+         --  if Power A ∋ t, from a propertiy of minimum sup there is osuc ZFSubset A ∋ t 
+         --    then ZFSubset A ≡ t or ZFSubset A ∋ t. In the former case ZFSubset A ∋ x implies A ∋ x
+         --    In case of later, ZFSubset A ∋ t and t ∋ x implies A ∋ x by transitivity
+         --
+         POrd : {a : Ordinal } {t : OD} → Def (Ord a) ∋ t → Def (Ord a) ∋ Ord (od→ord t)
+         POrd {a} {t} P∋t = {!!}
+         ∩-≡ :  { a b : OD {suc n} } → ({x : OD {suc n} } → (a ∋ x → b ∋ x)) → a == ( b ∩ a )
+         ∩-≡ {a} {b} inc = record {
+            eq→ = λ {x} x<a → record { proj2 = x<a ;
+                 proj1 = def-subst {suc n} {_} {_} {b} {x} (inc (def-subst {suc n} {_} {_} {a} {_} x<a refl (sym diso) )) refl diso  } ;
+            eq← = λ {x} x<a∩b → proj2 x<a∩b }
+         ord-power→ : (a : Ordinal ) ( t : OD) → Def (Ord a) ∋ t → {x : OD} → t ∋ x → Ord a ∋ x
+         ord-power→ a t P∋t {x} t∋x with osuc-≡<  (sup-lb  (POrd P∋t))
+         ... | case1 eq = proj1 (def-subst (Ltx t∋x) (sym (subst₂ (λ j k → j ≡ k ) oiso oiso ( cong (λ k → ord→od k) (sym eq) ))) refl )  where
+              Ltx :   {n : Level} → {x t : OD {suc n}} → t ∋ x → Ord (od→ord t) ∋ x
+              Ltx {n} {x} {t} lt = c<→o< lt
+         ... | case2 lt = lemma3 where
+              sp =  sup-x (λ x → od→ord ( ZFSubset (Ord a) (ord→od x)))
+              minsup :  OD
+              minsup =  ZFSubset (Ord a) ( ord→od ( sup-x (λ x → od→ord ( ZFSubset (Ord a) (ord→od x))))) 
+              Ltx :   {n : Level} → {x t : OD {suc n}} → t ∋ x → Ord (od→ord t) ∋ x
+              Ltx {n} {x} {t} lt = c<→o< lt
+              -- lemma1 hold because minsup is Ord (minα a sp) 
+              lemma1 : od→ord (Ord (od→ord t)) o< od→ord minsup → minsup ∋ Ord (od→ord t)
+              lemma1 lt with OrdSubset a ( sup-x (λ x → od→ord ( ZFSubset (Ord a) (ord→od x))))
+              ... | eq with subst ( λ k →  ZFSubset (Ord a) k ≡ Ord (minα a sp)) {!!} eq
+              ... | eq1 = def-subst {suc n} {_} {_} {minsup} {od→ord (Ord (od→ord t))} {!!} lemma2 {!!} where
+                  lemma2 : Ord (od→ord (ZFSubset (Ord a) (ord→od sp))) ≡ minsup
+                  lemma2 =  let open ≡-Reasoning in begin
+                      Ord (od→ord (ZFSubset (Ord a) (ord→od sp)))
+                    ≡⟨ cong (λ k → Ord (od→ord k)) eq1 ⟩
+                      Ord (od→ord (Ord (minα a sp))) 
+                    ≡⟨ cong (λ k → Ord (od→ord k)) {!!}  ⟩
+                      Ord (od→ord (ord→od (minα a sp))) 
+                    ≡⟨ cong (λ k → Ord k) diso ⟩
+                      Ord (minα a sp)
+                    ≡⟨ sym eq1 ⟩
+                      minsup
+                    ∎
+              lemma3 : od→ord x o< a
+              lemma3 = otrans (proj1 (lemma1 lt) ) (c<→o< {suc n} {x} {Ord (od→ord t)} (Ltx t∋x) )
+         -- 
+         -- we have t ∋ x → A ∋ x means t is a subset of A, that is ZFSubset A t == t
+         -- Power A is a sup of ZFSubset A t, so Power A ∋ t
+         -- 
+         ord-power← : (a : Ordinal ) (t : OD) → ({x : OD} → (t ∋ x → (Ord a) ∋ x)) → Def (Ord a) ∋ t
+         ord-power← a t t→A  = def-subst {suc n} {_} {_} {Def (Ord a)} {od→ord t}
+                 lemma refl (lemma1 lemma-eq )where
+              lemma-eq :  ZFSubset (Ord a) t == t
+              eq→ lemma-eq {z} w = proj2 w 
+              eq← lemma-eq {z} w = record { proj2 = w  ;
+                 proj1 = def-subst {suc n} {_} {_} {(Ord a)} {z}
+                    ( t→A (def-subst {suc n} {_} {_} {t} {od→ord (ord→od z)} w refl (sym diso) )) refl diso  }
+              lemma1 : {n : Level } {a : Ordinal {suc n}} { t : OD {suc n}}
+                 → (eq : ZFSubset (Ord a) t == t)  → od→ord (ZFSubset (Ord a) (ord→od (od→ord t))) ≡ od→ord t
+              lemma1 {n} {a} {t} eq = subst (λ k → od→ord (ZFSubset (Ord a) k) ≡ od→ord t ) (sym oiso) (cong (λ k → od→ord k ) (===-≡ eq ))
+              lemma :  od→ord (ZFSubset (Ord a) (ord→od (od→ord t)) ) o< sup-o (λ x → od→ord (ZFSubset (Ord a) (ord→od x)))
+              lemma = sup-o<   
+
+         -- 
+         -- Every set in OD is a subset of Ordinals
+         -- 
+         -- Power A = Replace (Def (Ord (od→ord A))) ( λ y → A ∩ y )
+         power→ :  ( A t : OD) → Power A ∋ t → {x : OD} → t ∋ x → A ∋ x
+         power→ A t P∋t {x} t∋x = TransFiniteExists {suc n} {λ y → (t ==  (A ∩ ord→od y))}
+                 lemma4 lemma5  where
+              a = od→ord A
+              lemma2 : ¬ ( (y : OD) → ¬ (t ==  (A ∩ y)))
+              lemma2 = replacement→ (Def (Ord (od→ord A))) t P∋t
+              lemma3 : (y : OD) → t == ( A ∩ y ) → A ∋ x
+              lemma3 y eq = proj1 (eq→ eq t∋x)
+              lemma4 : ¬ ((y : Ordinal) → ¬ (t == (A ∩ ord→od y)))
+              lemma4 not = lemma2 ( λ y not1 → not (od→ord y) (subst (λ k → t == ( A ∩ k )) (sym oiso) not1 ))
+              lemma5 : {y : Ordinal} → t == (A ∩ ord→od y) → def A (od→ord x)
+              lemma5 {y} eq = lemma3 (ord→od y) eq
+         power← :  (A t : OD) → ({x : OD} → (t ∋ x → A ∋ x)) → Power A ∋ t
+         power← A t t→A = record { proj1 = lemma1 ; proj2 = lemma2 } where 
+              a = od→ord A
+              lemma0 : {x : OD} → t ∋ x → Ord a ∋ x
+              lemma0 {x} t∋x = c<→o< (t→A t∋x)
+              lemma3 : Def (Ord a) ∋ t
+              lemma3 = ord-power← a t lemma0
+              lemma4 : od→ord t ≡ od→ord (A ∩ Ord (od→ord t))
+              lemma4 = cong ( λ k → od→ord k ) ( ===-≡ (subst (λ k → t == (A ∩ k)) {!!} {!!} ))
+              lemma1 : od→ord t o< sup-o (λ x → od→ord (A ∩ ord→od x))
+              lemma1 with sup-o< {suc n} {λ x → od→ord (A ∩ ord→od x)} {od→ord t}
+              ... | lt = o<-subst {suc n} {_} {_} {_} {_} lt (sym (subst (λ k → od→ord t ≡ k) lemma5 lemma4 )) refl where
+                  lemma5 : od→ord (A ∩ Ord (od→ord t)) ≡ od→ord (A ∩ ord→od (od→ord t))
+                  lemma5 = cong (λ k → od→ord (A ∩ k )) {!!}
+              lemma2 :  def (in-codomain (Def (Ord (od→ord A))) (_∩_ A)) (od→ord t)
+              lemma2 not = ⊥-elim ( not (od→ord t) (record { proj1 = lemma3 ; proj2 = {!!} }) ) where
+
+         ∅-iso :  {x : OD} → ¬ (x == od∅) → ¬ ((ord→od (od→ord x)) == od∅) 
+         ∅-iso {x} neq = subst (λ k → ¬ k) (=-iso {n} ) neq  
+         regularity :  (x : OD) (not : ¬ (x == od∅)) →
+            (x ∋ minimul x not) ∧ (Select (minimul x not) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅)
+         proj1 (regularity x not ) = x∋minimul x not
+         proj2 (regularity x not ) = record { eq→ = lemma1 ; eq← = λ {y} d → lemma2 {y} d } where
+             lemma1 : {x₁ : Ordinal} → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) x₁ → def od∅ x₁
+             lemma1 {x₁} s = ⊥-elim  ( minimul-1 x not (ord→od x₁) lemma3 ) where
+                 lemma3 : def (minimul x not) (od→ord (ord→od x₁)) ∧ def x (od→ord (ord→od x₁))
+                 lemma3 = record { proj1 = def-subst {suc n} {_} {_} {minimul x not} {_} (proj1 s) refl (sym diso)
+                                 ; proj2 = proj2 (proj2 s) } 
+             lemma2 : {x₁ : Ordinal} → def od∅ x₁ → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) x₁
+             lemma2 {y} d = ⊥-elim (empty (ord→od y) (def-subst {suc n} {_} {_} {od∅} {od→ord (ord→od y)} d refl (sym diso) ))
+
+         extensionality : {A B : OD {suc n}} → ((z : OD) → (A ∋ z) ⇔ (B ∋ z)) → A == B
+         eq→ (extensionality {A} {B} eq ) {x} d = def-iso {suc n} {A} {B} (sym diso) (proj1 (eq (ord→od x))) d  
+         eq← (extensionality {A} {B} eq ) {x} d = def-iso {suc n} {B} {A} (sym diso) (proj2 (eq (ord→od x))) d  
+
+         open  import  Relation.Binary.PropositionalEquality
+         uxxx-ord : {x  : OD {suc n}} → {y : Ordinal {suc n}} → def (Union (x , (x , x))) y ⇔ ( y o< osuc (od→ord x) )
+         uxxx-ord {x} {y} = subst (λ k → k ⇔ ( y o< osuc (od→ord x) )) (sym lemma) ( osuc2 y (od→ord x))  where
+              lemma : {y : Ordinal {suc n}} → def (Union (x , (x , x))) y ≡ osuc y o< osuc (osuc (od→ord x)) 
+              lemma {y} = let open ≡-Reasoning in begin
+                   def (Union (x , (x , x))) y  
+                ≡⟨⟩
+                   def ( Ord ( omax (od→ord x) (od→ord (Ord (omax (od→ord x)  (od→ord x)  )) ))) ( osuc y )
+                ≡⟨⟩
+                   osuc y o<  omax (od→ord x) (od→ord (Ord (omax (od→ord x)  (od→ord x)  )) )
+                ≡⟨ cong (λ k → osuc y o<  omax (od→ord x) k ) {!!}  ⟩
+                   osuc y o<  omax (od→ord x) (omax (od→ord x)  (od→ord x)  ) 
+                ≡⟨ cong (λ k → osuc y o<  k ) (omxxx  (od→ord x) )  ⟩
+                   osuc y o< osuc (osuc (od→ord x))
+                ∎ 
+         infinite : OD {suc n}
+         infinite = Ord omega 
+         infinity∅ : Ord omega  ∋ od∅ {suc n}
+         infinity∅ = o<-subst (case1 (s≤s z≤n) ) {!!} refl
+         infinity : (x : OD) → infinite ∋ x → infinite ∋ Union (x , (x , x ))
+         infinity x lt = o<-subst ( lemma (od→ord x) lt ) eq refl where
+              eq :  osuc (od→ord x) ≡ od→ord (Union (x , (x , x)))
+              eq = let open ≡-Reasoning in begin
+                    osuc (od→ord x)
+                 ≡⟨ {!!}  ⟩
+                    od→ord (Ord (osuc (od→ord x)))
+                 ≡⟨ cong ( λ k → od→ord  k ) ( sym (==→o≡ ( ⇔→==  uxxx-ord ))) ⟩
+                    od→ord (Union (x , (x , x)))
+                 ∎
+              lemma : (ox : Ordinal {suc n} ) → ox o< omega → osuc ox o< omega 
+              lemma record { lv = Zero ; ord = (Φ .0) } (case1 (s≤s x)) = case1 (s≤s z≤n)
+              lemma record { lv = Zero ; ord = (OSuc .0 ord₁) } (case1 (s≤s x)) = case1 (s≤s z≤n)
+              lemma record { lv = (Suc lv₁) ; ord = (Φ .(Suc lv₁)) } (case1 (s≤s ()))
+              lemma record { lv = (Suc lv₁) ; ord = (OSuc .(Suc lv₁) ord₁) } (case1 (s≤s ()))
+              lemma record { lv = 1 ; ord = (Φ 1) } (case2 c2) with d<→lv c2
+              lemma record { lv = (Suc Zero) ; ord = (Φ .1) } (case2 ()) | refl
+         -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] -- this form is no good since X is a transitive set
+         -- ∀ z [ ∀ x ( x ∈ z  → ¬ ( x ≈ ∅ ) )  ∧ ∀ x ∀ y ( x , y ∈ z ∧ ¬ ( x ≈ y )  → x ∩ y ≈ ∅  ) → ∃ u ∀ x ( x ∈ z → ∃ t ( u ∩ x) ≈ { t }) ]
+         record Choice (z : OD {suc n}) : Set (suc (suc n)) where
+             field
+                 u : {x : OD {suc n}} ( x∈z  : x ∈ z ) → OD {suc n}
+                 t : {x : OD {suc n}} ( x∈z  : x ∈ z ) → (x : OD {suc n} ) → OD {suc n}
+                 choice : { x : OD {suc n} } → ( x∈z  : x ∈ z ) → ( u x∈z ∩ x) == { t x∈z x }
+         -- choice : {x :  OD {suc n}} ( x ∈ z  → ¬ ( x ≈ ∅ ) ) →
+         -- axiom-of-choice : { X : OD } → ( ¬x∅ : ¬ ( X == od∅ ) ) → { A : OD } → (A∈X : A ∈ X ) →  choice ¬x∅ A∈X ∈ A 
+         -- axiom-of-choice {X} nx {A} lt = ¬∅=→∅∈ {!!}
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/LICENSE	Mon Jul 08 19:48:47 2019 +0900
@@ -0,0 +1,674 @@
+                    GNU GENERAL PUBLIC LICENSE
+                       Version 3, 29 June 2007
+
+ Copyright (C) 2007 Free Software Foundation, Inc. <https://fsf.org/>
+ Everyone is permitted to copy and distribute verbatim copies
+ of this license document, but changing it is not allowed.
+
+                            Preamble
+
+  The GNU General Public License is a free, copyleft license for
+software and other kinds of works.
+
+  The licenses for most software and other practical works are designed
+to take away your freedom to share and change the works.  By contrast,
+the GNU General Public License is intended to guarantee your freedom to
+share and change all versions of a program--to make sure it remains free
+software for all its users.  We, the Free Software Foundation, use the
+GNU General Public License for most of our software; it applies also to
+any other work released this way by its authors.  You can apply it to
+your programs, too.
+
+  When we speak of free software, we are referring to freedom, not
+price.  Our General Public Licenses are designed to make sure that you
+have the freedom to distribute copies of free software (and charge for
+them if you wish), that you receive source code or can get it if you
+want it, that you can change the software or use pieces of it in new
+free programs, and that you know you can do these things.
+
+  To protect your rights, we need to prevent others from denying you
+these rights or asking you to surrender the rights.  Therefore, you have
+certain responsibilities if you distribute copies of the software, or if
+you modify it: responsibilities to respect the freedom of others.
+
+  For example, if you distribute copies of such a program, whether
+gratis or for a fee, you must pass on to the recipients the same
+freedoms that you received.  You must make sure that they, too, receive
+or can get the source code.  And you must show them these terms so they
+know their rights.
+
+  Developers that use the GNU GPL protect your rights with two steps:
+(1) assert copyright on the software, and (2) offer you this License
+giving you legal permission to copy, distribute and/or modify it.
+
+  For the developers' and authors' protection, the GPL clearly explains
+that there is no warranty for this free software.  For both users' and
+authors' sake, the GPL requires that modified versions be marked as
+changed, so that their problems will not be attributed erroneously to
+authors of previous versions.
+
+  Some devices are designed to deny users access to install or run
+modified versions of the software inside them, although the manufacturer
+can do so.  This is fundamentally incompatible with the aim of
+protecting users' freedom to change the software.  The systematic
+pattern of such abuse occurs in the area of products for individuals to
+use, which is precisely where it is most unacceptable.  Therefore, we
+have designed this version of the GPL to prohibit the practice for those
+products.  If such problems arise substantially in other domains, we
+stand ready to extend this provision to those domains in future versions
+of the GPL, as needed to protect the freedom of users.
+
+  Finally, every program is threatened constantly by software patents.
+States should not allow patents to restrict development and use of
+software on general-purpose computers, but in those that do, we wish to
+avoid the special danger that patents applied to a free program could
+make it effectively proprietary.  To prevent this, the GPL assures that
+patents cannot be used to render the program non-free.
+
+  The precise terms and conditions for copying, distribution and
+modification follow.
+
+                       TERMS AND CONDITIONS
+
+  0. Definitions.
+
+  "This License" refers to version 3 of the GNU General Public License.
+
+  "Copyright" also means copyright-like laws that apply to other kinds of
+works, such as semiconductor masks.
+
+  "The Program" refers to any copyrightable work licensed under this
+License.  Each licensee is addressed as "you".  "Licensees" and
+"recipients" may be individuals or organizations.
+
+  To "modify" a work means to copy from or adapt all or part of the work
+in a fashion requiring copyright permission, other than the making of an
+exact copy.  The resulting work is called a "modified version" of the
+earlier work or a work "based on" the earlier work.
+
+  A "covered work" means either the unmodified Program or a work based
+on the Program.
+
+  To "propagate" a work means to do anything with it that, without
+permission, would make you directly or secondarily liable for
+infringement under applicable copyright law, except executing it on a
+computer or modifying a private copy.  Propagation includes copying,
+distribution (with or without modification), making available to the
+public, and in some countries other activities as well.
+
+  To "convey" a work means any kind of propagation that enables other
+parties to make or receive copies.  Mere interaction with a user through
+a computer network, with no transfer of a copy, is not conveying.
+
+  An interactive user interface displays "Appropriate Legal Notices"
+to the extent that it includes a convenient and prominently visible
+feature that (1) displays an appropriate copyright notice, and (2)
+tells the user that there is no warranty for the work (except to the
+extent that warranties are provided), that licensees may convey the
+work under this License, and how to view a copy of this License.  If
+the interface presents a list of user commands or options, such as a
+menu, a prominent item in the list meets this criterion.
+
+  1. Source Code.
+
+  The "source code" for a work means the preferred form of the work
+for making modifications to it.  "Object code" means any non-source
+form of a work.
+
+  A "Standard Interface" means an interface that either is an official
+standard defined by a recognized standards body, or, in the case of
+interfaces specified for a particular programming language, one that
+is widely used among developers working in that language.
+
+  The "System Libraries" of an executable work include anything, other
+than the work as a whole, that (a) is included in the normal form of
+packaging a Major Component, but which is not part of that Major
+Component, and (b) serves only to enable use of the work with that
+Major Component, or to implement a Standard Interface for which an
+implementation is available to the public in source code form.  A
+"Major Component", in this context, means a major essential component
+(kernel, window system, and so on) of the specific operating system
+(if any) on which the executable work runs, or a compiler used to
+produce the work, or an object code interpreter used to run it.
+
+  The "Corresponding Source" for a work in object code form means all
+the source code needed to generate, install, and (for an executable
+work) run the object code and to modify the work, including scripts to
+control those activities.  However, it does not include the work's
+System Libraries, or general-purpose tools or generally available free
+programs which are used unmodified in performing those activities but
+which are not part of the work.  For example, Corresponding Source
+includes interface definition files associated with source files for
+the work, and the source code for shared libraries and dynamically
+linked subprograms that the work is specifically designed to require,
+such as by intimate data communication or control flow between those
+subprograms and other parts of the work.
+
+  The Corresponding Source need not include anything that users
+can regenerate automatically from other parts of the Corresponding
+Source.
+
+  The Corresponding Source for a work in source code form is that
+same work.
+
+  2. Basic Permissions.
+
+  All rights granted under this License are granted for the term of
+copyright on the Program, and are irrevocable provided the stated
+conditions are met.  This License explicitly affirms your unlimited
+permission to run the unmodified Program.  The output from running a
+covered work is covered by this License only if the output, given its
+content, constitutes a covered work.  This License acknowledges your
+rights of fair use or other equivalent, as provided by copyright law.
+
+  You may make, run and propagate covered works that you do not
+convey, without conditions so long as your license otherwise remains
+in force.  You may convey covered works to others for the sole purpose
+of having them make modifications exclusively for you, or provide you
+with facilities for running those works, provided that you comply with
+the terms of this License in conveying all material for which you do
+not control copyright.  Those thus making or running the covered works
+for you must do so exclusively on your behalf, under your direction
+and control, on terms that prohibit them from making any copies of
+your copyrighted material outside their relationship with you.
+
+  Conveying under any other circumstances is permitted solely under
+the conditions stated below.  Sublicensing is not allowed; section 10
+makes it unnecessary.
+
+  3. Protecting Users' Legal Rights From Anti-Circumvention Law.
+
+  No covered work shall be deemed part of an effective technological
+measure under any applicable law fulfilling obligations under article
+11 of the WIPO copyright treaty adopted on 20 December 1996, or
+similar laws prohibiting or restricting circumvention of such
+measures.
+
+  When you convey a covered work, you waive any legal power to forbid
+circumvention of technological measures to the extent such circumvention
+is effected by exercising rights under this License with respect to
+the covered work, and you disclaim any intention to limit operation or
+modification of the work as a means of enforcing, against the work's
+users, your or third parties' legal rights to forbid circumvention of
+technological measures.
+
+  4. Conveying Verbatim Copies.
+
+  You may convey verbatim copies of the Program's source code as you
+receive it, in any medium, provided that you conspicuously and
+appropriately publish on each copy an appropriate copyright notice;
+keep intact all notices stating that this License and any
+non-permissive terms added in accord with section 7 apply to the code;
+keep intact all notices of the absence of any warranty; and give all
+recipients a copy of this License along with the Program.
+
+  You may charge any price or no price for each copy that you convey,
+and you may offer support or warranty protection for a fee.
+
+  5. Conveying Modified Source Versions.
+
+  You may convey a work based on the Program, or the modifications to
+produce it from the Program, in the form of source code under the
+terms of section 4, provided that you also meet all of these conditions:
+
+    a) The work must carry prominent notices stating that you modified
+    it, and giving a relevant date.
+
+    b) The work must carry prominent notices stating that it is
+    released under this License and any conditions added under section
+    7.  This requirement modifies the requirement in section 4 to
+    "keep intact all notices".
+
+    c) You must license the entire work, as a whole, under this
+    License to anyone who comes into possession of a copy.  This
+    License will therefore apply, along with any applicable section 7
+    additional terms, to the whole of the work, and all its parts,
+    regardless of how they are packaged.  This License gives no
+    permission to license the work in any other way, but it does not
+    invalidate such permission if you have separately received it.
+
+    d) If the work has interactive user interfaces, each must display
+    Appropriate Legal Notices; however, if the Program has interactive
+    interfaces that do not display Appropriate Legal Notices, your
+    work need not make them do so.
+
+  A compilation of a covered work with other separate and independent
+works, which are not by their nature extensions of the covered work,
+and which are not combined with it such as to form a larger program,
+in or on a volume of a storage or distribution medium, is called an
+"aggregate" if the compilation and its resulting copyright are not
+used to limit the access or legal rights of the compilation's users
+beyond what the individual works permit.  Inclusion of a covered work
+in an aggregate does not cause this License to apply to the other
+parts of the aggregate.
+
+  6. Conveying Non-Source Forms.
+
+  You may convey a covered work in object code form under the terms
+of sections 4 and 5, provided that you also convey the
+machine-readable Corresponding Source under the terms of this License,
+in one of these ways:
+
+    a) Convey the object code in, or embodied in, a physical product
+    (including a physical distribution medium), accompanied by the
+    Corresponding Source fixed on a durable physical medium
+    customarily used for software interchange.
+
+    b) Convey the object code in, or embodied in, a physical product
+    (including a physical distribution medium), accompanied by a
+    written offer, valid for at least three years and valid for as
+    long as you offer spare parts or customer support for that product
+    model, to give anyone who possesses the object code either (1) a
+    copy of the Corresponding Source for all the software in the
+    product that is covered by this License, on a durable physical
+    medium customarily used for software interchange, for a price no
+    more than your reasonable cost of physically performing this
+    conveying of source, or (2) access to copy the
+    Corresponding Source from a network server at no charge.
+
+    c) Convey individual copies of the object code with a copy of the
+    written offer to provide the Corresponding Source.  This
+    alternative is allowed only occasionally and noncommercially, and
+    only if you received the object code with such an offer, in accord
+    with subsection 6b.
+
+    d) Convey the object code by offering access from a designated
+    place (gratis or for a charge), and offer equivalent access to the
+    Corresponding Source in the same way through the same place at no
+    further charge.  You need not require recipients to copy the
+    Corresponding Source along with the object code.  If the place to
+    copy the object code is a network server, the Corresponding Source
+    may be on a different server (operated by you or a third party)
+    that supports equivalent copying facilities, provided you maintain
+    clear directions next to the object code saying where to find the
+    Corresponding Source.  Regardless of what server hosts the
+    Corresponding Source, you remain obligated to ensure that it is
+    available for as long as needed to satisfy these requirements.
+
+    e) Convey the object code using peer-to-peer transmission, provided
+    you inform other peers where the object code and Corresponding
+    Source of the work are being offered to the general public at no
+    charge under subsection 6d.
+
+  A separable portion of the object code, whose source code is excluded
+from the Corresponding Source as a System Library, need not be
+included in conveying the object code work.
+
+  A "User Product" is either (1) a "consumer product", which means any
+tangible personal property which is normally used for personal, family,
+or household purposes, or (2) anything designed or sold for incorporation
+into a dwelling.  In determining whether a product is a consumer product,
+doubtful cases shall be resolved in favor of coverage.  For a particular
+product received by a particular user, "normally used" refers to a
+typical or common use of that class of product, regardless of the status
+of the particular user or of the way in which the particular user
+actually uses, or expects or is expected to use, the product.  A product
+is a consumer product regardless of whether the product has substantial
+commercial, industrial or non-consumer uses, unless such uses represent
+the only significant mode of use of the product.
+
+  "Installation Information" for a User Product means any methods,
+procedures, authorization keys, or other information required to install
+and execute modified versions of a covered work in that User Product from
+a modified version of its Corresponding Source.  The information must
+suffice to ensure that the continued functioning of the modified object
+code is in no case prevented or interfered with solely because
+modification has been made.
+
+  If you convey an object code work under this section in, or with, or
+specifically for use in, a User Product, and the conveying occurs as
+part of a transaction in which the right of possession and use of the
+User Product is transferred to the recipient in perpetuity or for a
+fixed term (regardless of how the transaction is characterized), the
+Corresponding Source conveyed under this section must be accompanied
+by the Installation Information.  But this requirement does not apply
+if neither you nor any third party retains the ability to install
+modified object code on the User Product (for example, the work has
+been installed in ROM).
+
+  The requirement to provide Installation Information does not include a
+requirement to continue to provide support service, warranty, or updates
+for a work that has been modified or installed by the recipient, or for
+the User Product in which it has been modified or installed.  Access to a
+network may be denied when the modification itself materially and
+adversely affects the operation of the network or violates the rules and
+protocols for communication across the network.
+
+  Corresponding Source conveyed, and Installation Information provided,
+in accord with this section must be in a format that is publicly
+documented (and with an implementation available to the public in
+source code form), and must require no special password or key for
+unpacking, reading or copying.
+
+  7. Additional Terms.
+
+  "Additional permissions" are terms that supplement the terms of this
+License by making exceptions from one or more of its conditions.
+Additional permissions that are applicable to the entire Program shall
+be treated as though they were included in this License, to the extent
+that they are valid under applicable law.  If additional permissions
+apply only to part of the Program, that part may be used separately
+under those permissions, but the entire Program remains governed by
+this License without regard to the additional permissions.
+
+  When you convey a copy of a covered work, you may at your option
+remove any additional permissions from that copy, or from any part of
+it.  (Additional permissions may be written to require their own
+removal in certain cases when you modify the work.)  You may place
+additional permissions on material, added by you to a covered work,
+for which you have or can give appropriate copyright permission.
+
+  Notwithstanding any other provision of this License, for material you
+add to a covered work, you may (if authorized by the copyright holders of
+that material) supplement the terms of this License with terms:
+
+    a) Disclaiming warranty or limiting liability differently from the
+    terms of sections 15 and 16 of this License; or
+
+    b) Requiring preservation of specified reasonable legal notices or
+    author attributions in that material or in the Appropriate Legal
+    Notices displayed by works containing it; or
+
+    c) Prohibiting misrepresentation of the origin of that material, or
+    requiring that modified versions of such material be marked in
+    reasonable ways as different from the original version; or
+
+    d) Limiting the use for publicity purposes of names of licensors or
+    authors of the material; or
+
+    e) Declining to grant rights under trademark law for use of some
+    trade names, trademarks, or service marks; or
+
+    f) Requiring indemnification of licensors and authors of that
+    material by anyone who conveys the material (or modified versions of
+    it) with contractual assumptions of liability to the recipient, for
+    any liability that these contractual assumptions directly impose on
+    those licensors and authors.
+
+  All other non-permissive additional terms are considered "further
+restrictions" within the meaning of section 10.  If the Program as you
+received it, or any part of it, contains a notice stating that it is
+governed by this License along with a term that is a further
+restriction, you may remove that term.  If a license document contains
+a further restriction but permits relicensing or conveying under this
+License, you may add to a covered work material governed by the terms
+of that license document, provided that the further restriction does
+not survive such relicensing or conveying.
+
+  If you add terms to a covered work in accord with this section, you
+must place, in the relevant source files, a statement of the
+additional terms that apply to those files, or a notice indicating
+where to find the applicable terms.
+
+  Additional terms, permissive or non-permissive, may be stated in the
+form of a separately written license, or stated as exceptions;
+the above requirements apply either way.
+
+  8. Termination.
+
+  You may not propagate or modify a covered work except as expressly
+provided under this License.  Any attempt otherwise to propagate or
+modify it is void, and will automatically terminate your rights under
+this License (including any patent licenses granted under the third
+paragraph of section 11).
+
+  However, if you cease all violation of this License, then your
+license from a particular copyright holder is reinstated (a)
+provisionally, unless and until the copyright holder explicitly and
+finally terminates your license, and (b) permanently, if the copyright
+holder fails to notify you of the violation by some reasonable means
+prior to 60 days after the cessation.
+
+  Moreover, your license from a particular copyright holder is
+reinstated permanently if the copyright holder notifies you of the
+violation by some reasonable means, this is the first time you have
+received notice of violation of this License (for any work) from that
+copyright holder, and you cure the violation prior to 30 days after
+your receipt of the notice.
+
+  Termination of your rights under this section does not terminate the
+licenses of parties who have received copies or rights from you under
+this License.  If your rights have been terminated and not permanently
+reinstated, you do not qualify to receive new licenses for the same
+material under section 10.
+
+  9. Acceptance Not Required for Having Copies.
+
+  You are not required to accept this License in order to receive or
+run a copy of the Program.  Ancillary propagation of a covered work
+occurring solely as a consequence of using peer-to-peer transmission
+to receive a copy likewise does not require acceptance.  However,
+nothing other than this License grants you permission to propagate or
+modify any covered work.  These actions infringe copyright if you do
+not accept this License.  Therefore, by modifying or propagating a
+covered work, you indicate your acceptance of this License to do so.
+
+  10. Automatic Licensing of Downstream Recipients.
+
+  Each time you convey a covered work, the recipient automatically
+receives a license from the original licensors, to run, modify and
+propagate that work, subject to this License.  You are not responsible
+for enforcing compliance by third parties with this License.
+
+  An "entity transaction" is a transaction transferring control of an
+organization, or substantially all assets of one, or subdividing an
+organization, or merging organizations.  If propagation of a covered
+work results from an entity transaction, each party to that
+transaction who receives a copy of the work also receives whatever
+licenses to the work the party's predecessor in interest had or could
+give under the previous paragraph, plus a right to possession of the
+Corresponding Source of the work from the predecessor in interest, if
+the predecessor has it or can get it with reasonable efforts.
+
+  You may not impose any further restrictions on the exercise of the
+rights granted or affirmed under this License.  For example, you may
+not impose a license fee, royalty, or other charge for exercise of
+rights granted under this License, and you may not initiate litigation
+(including a cross-claim or counterclaim in a lawsuit) alleging that
+any patent claim is infringed by making, using, selling, offering for
+sale, or importing the Program or any portion of it.
+
+  11. Patents.
+
+  A "contributor" is a copyright holder who authorizes use under this
+License of the Program or a work on which the Program is based.  The
+work thus licensed is called the contributor's "contributor version".
+
+  A contributor's "essential patent claims" are all patent claims
+owned or controlled by the contributor, whether already acquired or
+hereafter acquired, that would be infringed by some manner, permitted
+by this License, of making, using, or selling its contributor version,
+but do not include claims that would be infringed only as a
+consequence of further modification of the contributor version.  For
+purposes of this definition, "control" includes the right to grant
+patent sublicenses in a manner consistent with the requirements of
+this License.
+
+  Each contributor grants you a non-exclusive, worldwide, royalty-free
+patent license under the contributor's essential patent claims, to
+make, use, sell, offer for sale, import and otherwise run, modify and
+propagate the contents of its contributor version.
+
+  In the following three paragraphs, a "patent license" is any express
+agreement or commitment, however denominated, not to enforce a patent
+(such as an express permission to practice a patent or covenant not to
+sue for patent infringement).  To "grant" such a patent license to a
+party means to make such an agreement or commitment not to enforce a
+patent against the party.
+
+  If you convey a covered work, knowingly relying on a patent license,
+and the Corresponding Source of the work is not available for anyone
+to copy, free of charge and under the terms of this License, through a
+publicly available network server or other readily accessible means,
+then you must either (1) cause the Corresponding Source to be so
+available, or (2) arrange to deprive yourself of the benefit of the
+patent license for this particular work, or (3) arrange, in a manner
+consistent with the requirements of this License, to extend the patent
+license to downstream recipients.  "Knowingly relying" means you have
+actual knowledge that, but for the patent license, your conveying the
+covered work in a country, or your recipient's use of the covered work
+in a country, would infringe one or more identifiable patents in that
+country that you have reason to believe are valid.
+
+  If, pursuant to or in connection with a single transaction or
+arrangement, you convey, or propagate by procuring conveyance of, a
+covered work, and grant a patent license to some of the parties
+receiving the covered work authorizing them to use, propagate, modify
+or convey a specific copy of the covered work, then the patent license
+you grant is automatically extended to all recipients of the covered
+work and works based on it.
+
+  A patent license is "discriminatory" if it does not include within
+the scope of its coverage, prohibits the exercise of, or is
+conditioned on the non-exercise of one or more of the rights that are
+specifically granted under this License.  You may not convey a covered
+work if you are a party to an arrangement with a third party that is
+in the business of distributing software, under which you make payment
+to the third party based on the extent of your activity of conveying
+the work, and under which the third party grants, to any of the
+parties who would receive the covered work from you, a discriminatory
+patent license (a) in connection with copies of the covered work
+conveyed by you (or copies made from those copies), or (b) primarily
+for and in connection with specific products or compilations that
+contain the covered work, unless you entered into that arrangement,
+or that patent license was granted, prior to 28 March 2007.
+
+  Nothing in this License shall be construed as excluding or limiting
+any implied license or other defenses to infringement that may
+otherwise be available to you under applicable patent law.
+
+  12. No Surrender of Others' Freedom.
+
+  If conditions are imposed on you (whether by court order, agreement or
+otherwise) that contradict the conditions of this License, they do not
+excuse you from the conditions of this License.  If you cannot convey a
+covered work so as to satisfy simultaneously your obligations under this
+License and any other pertinent obligations, then as a consequence you may
+not convey it at all.  For example, if you agree to terms that obligate you
+to collect a royalty for further conveying from those to whom you convey
+the Program, the only way you could satisfy both those terms and this
+License would be to refrain entirely from conveying the Program.
+
+  13. Use with the GNU Affero General Public License.
+
+  Notwithstanding any other provision of this License, you have
+permission to link or combine any covered work with a work licensed
+under version 3 of the GNU Affero General Public License into a single
+combined work, and to convey the resulting work.  The terms of this
+License will continue to apply to the part which is the covered work,
+but the special requirements of the GNU Affero General Public License,
+section 13, concerning interaction through a network will apply to the
+combination as such.
+
+  14. Revised Versions of this License.
+
+  The Free Software Foundation may publish revised and/or new versions of
+the GNU General Public License from time to time.  Such new versions will
+be similar in spirit to the present version, but may differ in detail to
+address new problems or concerns.
+
+  Each version is given a distinguishing version number.  If the
+Program specifies that a certain numbered version of the GNU General
+Public License "or any later version" applies to it, you have the
+option of following the terms and conditions either of that numbered
+version or of any later version published by the Free Software
+Foundation.  If the Program does not specify a version number of the
+GNU General Public License, you may choose any version ever published
+by the Free Software Foundation.
+
+  If the Program specifies that a proxy can decide which future
+versions of the GNU General Public License can be used, that proxy's
+public statement of acceptance of a version permanently authorizes you
+to choose that version for the Program.
+
+  Later license versions may give you additional or different
+permissions.  However, no additional obligations are imposed on any
+author or copyright holder as a result of your choosing to follow a
+later version.
+
+  15. Disclaimer of Warranty.
+
+  THERE IS NO WARRANTY FOR THE PROGRAM, TO THE EXTENT PERMITTED BY
+APPLICABLE LAW.  EXCEPT WHEN OTHERWISE STATED IN WRITING THE COPYRIGHT
+HOLDERS AND/OR OTHER PARTIES PROVIDE THE PROGRAM "AS IS" WITHOUT WARRANTY
+OF ANY KIND, EITHER EXPRESSED OR IMPLIED, INCLUDING, BUT NOT LIMITED TO,
+THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR
+PURPOSE.  THE ENTIRE RISK AS TO THE QUALITY AND PERFORMANCE OF THE PROGRAM
+IS WITH YOU.  SHOULD THE PROGRAM PROVE DEFECTIVE, YOU ASSUME THE COST OF
+ALL NECESSARY SERVICING, REPAIR OR CORRECTION.
+
+  16. Limitation of Liability.
+
+  IN NO EVENT UNLESS REQUIRED BY APPLICABLE LAW OR AGREED TO IN WRITING
+WILL ANY COPYRIGHT HOLDER, OR ANY OTHER PARTY WHO MODIFIES AND/OR CONVEYS
+THE PROGRAM AS PERMITTED ABOVE, BE LIABLE TO YOU FOR DAMAGES, INCLUDING ANY
+GENERAL, SPECIAL, INCIDENTAL OR CONSEQUENTIAL DAMAGES ARISING OUT OF THE
+USE OR INABILITY TO USE THE PROGRAM (INCLUDING BUT NOT LIMITED TO LOSS OF
+DATA OR DATA BEING RENDERED INACCURATE OR LOSSES SUSTAINED BY YOU OR THIRD
+PARTIES OR A FAILURE OF THE PROGRAM TO OPERATE WITH ANY OTHER PROGRAMS),
+EVEN IF SUCH HOLDER OR OTHER PARTY HAS BEEN ADVISED OF THE POSSIBILITY OF
+SUCH DAMAGES.
+
+  17. Interpretation of Sections 15 and 16.
+
+  If the disclaimer of warranty and limitation of liability provided
+above cannot be given local legal effect according to their terms,
+reviewing courts shall apply local law that most closely approximates
+an absolute waiver of all civil liability in connection with the
+Program, unless a warranty or assumption of liability accompanies a
+copy of the Program in return for a fee.
+
+                     END OF TERMS AND CONDITIONS
+
+            How to Apply These Terms to Your New Programs
+
+  If you develop a new program, and you want it to be of the greatest
+possible use to the public, the best way to achieve this is to make it
+free software which everyone can redistribute and change under these terms.
+
+  To do so, attach the following notices to the program.  It is safest
+to attach them to the start of each source file to most effectively
+state the exclusion of warranty; and each file should have at least
+the "copyright" line and a pointer to where the full notice is found.
+
+    <one line to give the program's name and a brief idea of what it does.>
+    Copyright (C) <year>  <name of author>
+
+    This program is free software: you can redistribute it and/or modify
+    it under the terms of the GNU General Public License as published by
+    the Free Software Foundation, either version 3 of the License, or
+    (at your option) any later version.
+
+    This program is distributed in the hope that it will be useful,
+    but WITHOUT ANY WARRANTY; without even the implied warranty of
+    MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE.  See the
+    GNU General Public License for more details.
+
+    You should have received a copy of the GNU General Public License
+    along with this program.  If not, see <https://www.gnu.org/licenses/>.
+
+Also add information on how to contact you by electronic and paper mail.
+
+  If the program does terminal interaction, make it output a short
+notice like this when it starts in an interactive mode:
+
+    <program>  Copyright (C) <year>  <name of author>
+    This program comes with ABSOLUTELY NO WARRANTY; for details type `show w'.
+    This is free software, and you are welcome to redistribute it
+    under certain conditions; type `show c' for details.
+
+The hypothetical commands `show w' and `show c' should show the appropriate
+parts of the General Public License.  Of course, your program's commands
+might be different; for a GUI interface, you would use an "about box".
+
+  You should also get your employer (if you work as a programmer) or school,
+if any, to sign a "copyright disclaimer" for the program, if necessary.
+For more information on this, and how to apply and follow the GNU GPL, see
+<https://www.gnu.org/licenses/>.
+
+  The GNU General Public License does not permit incorporating your program
+into proprietary programs.  If your program is a subroutine library, you
+may consider it more useful to permit linking proprietary applications with
+the library.  If this is what you want to do, use the GNU Lesser General
+Public License instead of this License.  But first, please read
+<https://www.gnu.org/licenses/why-not-lgpl.html>.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Todo	Mon Jul 08 19:48:47 2019 +0900
@@ -0,0 +1,6 @@
+Mon Jul  8 19:43:37 JST 2019
+
+    ordinal-definable.agda assumes all ZF Set are ordinals, that it too restrictive
+
+    remove ord-Ord  and prove with some assuption in HOD.agda
+        union, power set, replace, inifinite
--- a/ordinal-definable.agda	Mon Jun 10 09:53:45 2019 +0900
+++ b/ordinal-definable.agda	Mon Jul 08 19:48:47 2019 +0900
@@ -1,3 +1,5 @@
+{-# OPTIONS --allow-unsolved-metas #-}
+
 open import Level
 module ordinal-definable where
 
@@ -23,6 +25,15 @@
 
 open Ordinal
 
+-- Ordinal in OD ( and ZFSet )
+Ord : { n : Level } → ( a : Ordinal {n} ) → OD {n}
+Ord {n} a = record { def = λ y → y o< a }  
+
+-- od∅ : {n : Level} → OD {n} 
+-- od∅ {n} = record { def = λ _ → Lift n ⊥ }
+od∅ : {n : Level} → OD {n} 
+od∅ {n} = Ord o∅ 
+
 record _==_ {n : Level} ( a b :  OD {n} ) : Set n where
   field
      eq→ : ∀ { x : Ordinal {n} } → def a x → def b x 
@@ -42,16 +53,16 @@
 eq-trans : {n : Level} {  x y z :  OD {n} } → x == y → y == z → x == z
 eq-trans x=y y=z = record { eq→ = λ t → eq→ y=z ( eq→ x=y  t) ; eq← = λ t → eq← x=y ( eq← y=z t) }
 
-od∅ : {n : Level} → OD {n} 
-od∅ {n} = record { def = λ _ → Lift n ⊥ }
+ord→od : {n : Level} → Ordinal {n} → OD {n} 
+ord→od a = Ord a
 
+o<→c<  : {n : Level} {x y : Ordinal {n} } → x o< y             → def (ord→od y) x 
+o<→c< {n} {x} {y} lt = lt 
 
 postulate      
   -- OD can be iso to a subset of Ordinal ( by means of Godel Set )
   od→ord : {n : Level} → OD {n} → Ordinal {n}
-  ord→od : {n : Level} → Ordinal {n} → OD {n} 
   c<→o<  : {n : Level} {x y : OD {n} }      → def y ( od→ord x ) → od→ord x o< od→ord y
-  o<→c<  : {n : Level} {x y : Ordinal {n} } → x o< y             → def (ord→od y) 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
   -- supermum as Replacement Axiom
@@ -82,7 +93,8 @@
         ( o<→c< sup-o< ) refl (cong ( λ k → od→ord (ψ k) ) oiso)
 
 ∅1 : {n : Level} →  ( x : OD {n} )  → ¬ ( x c< od∅ {n} )
-∅1 {n} x (lift ())
+∅1 {n} x (case1 ())
+∅1 {n} x (case2 ())
 
 ∅3 : {n : Level} →  { x : Ordinal {n}}  → ( ∀(y : Ordinal {n}) → ¬ (y o< x ) ) → x ≡ o∅ {n}
 ∅3 {n} {x} = TransFinite {n} c2 c3 x where
@@ -109,11 +121,6 @@
    lemma0 :  def ( ord→od ( od→ord z )) ( od→ord x) →  def z (od→ord x)
    lemma0 dz  = def-subst {suc n} { ord→od ( od→ord z )} { od→ord x} dz (oiso)  refl
 
-record Minimumo {n : Level } (x : Ordinal {n}) : Set (suc n) where
-  field
-     mino : Ordinal {n}
-     min<x :  mino o< x
-
 ∅5 : {n : Level} →  { x : Ordinal {n} }  → ¬ ( x  ≡ o∅ {n} ) → o∅ {n} o< x
 ∅5 {n} {record { lv = Zero ; ord = (Φ .0) }} not = ⊥-elim (not refl) 
 ∅5 {n} {record { lv = Zero ; ord = (OSuc .0 ord) }} not = case2 Φ<
@@ -154,11 +161,11 @@
 o<→o> : {n : Level} →  { x y : OD {n} } →  (x == y) → (od→ord x ) o< ( od→ord y) → ⊥
 o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case1 lt) with
      yx (def-subst {n} {ord→od (od→ord y)} {od→ord x} (o<→c< (case1 lt )) oiso refl )
-... | oyx with o<¬≡ (od→ord x) (od→ord x) refl (c<→o< oyx )
+... | oyx with o<¬≡ refl (c<→o< {n} {x} oyx )
 ... | ()
 o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case2 lt) with
      yx (def-subst {n} {ord→od (od→ord y)} {od→ord x} (o<→c< (case2 lt )) oiso refl )
-... | oyx with o<¬≡ (od→ord x) (od→ord x) refl (c<→o< oyx )
+... | oyx with o<¬≡ refl (c<→o< {n} {x} oyx )
 ... | ()
 
 ==→o≡ : {n : Level} →  { x y : Ordinal {suc n} } → ord→od x == ord→od y →  x ≡ y 
@@ -202,7 +209,8 @@
 o∅≡od∅ {n} | tri< a ¬b ¬c = ⊥-elim (lemma a) where
     lemma :  o∅ {suc n } o< (od→ord (od∅ {suc n} )) → ⊥
     lemma lt with def-subst (o<→c< lt) oiso refl
-    lemma lt | lift ()
+    lemma lt | case1 ()
+    lemma lt | case2 ()
 o∅≡od∅ {n} | tri≈ ¬a b ¬c = trans (cong (λ k → ord→od k ) b ) oiso
 o∅≡od∅ {n} | tri> ¬a ¬b c = ⊥-elim (¬x<0 c)
 
@@ -213,7 +221,7 @@
 o<→¬c> {n} {x} {y} olt clt = o<> olt (c<→o< clt ) where
 
 o≡→¬c< : {n : Level} →  { x y : OD {n} } →  (od→ord x ) ≡ ( od→ord y) →   ¬ x c< y
-o≡→¬c< {n} {x} {y} oeq lt  = o<¬≡ (od→ord x) (od→ord y) (orefl oeq ) (c<→o< lt) 
+o≡→¬c< {n} {x} {y} oeq lt  = o<¬≡   (orefl oeq ) (c<→o< lt) 
 
 tri-c< : {n : Level} →  Trichotomous _==_ (_c<_ {suc n})
 tri-c< {n} x y with trio< {n} (od→ord x) (od→ord y) 
@@ -229,7 +237,8 @@
 
 ∅< : {n : Level} →  { x y : OD {n} } → def x (od→ord y ) → ¬ (  x  == od∅ {n} )
 ∅< {n} {x} {y} d eq with eq→ eq d
-∅< {n} {x} {y} d eq | lift ()
+∅< {n} {x} {y} d eq | case1 ()
+∅< {n} {x} {y} d eq | case2 ()
        
 ∅6 : {n : Level} →  { x : OD {suc n} }  → ¬ ( x ∋ x )    --  no Russel paradox
 ∅6 {n} {x} x∋x = c<> {n} {x} {x} x∋x x∋x
@@ -250,6 +259,9 @@
 
 open _∧_
 
+--
+-- This menas OD is Ordinal here
+--
 ¬∅=→∅∈ :  {n : Level} →  { x : OD {suc n} } → ¬ (  x  == od∅ {suc n} ) → x ∋ od∅ {suc n} 
 ¬∅=→∅∈  {n} {x} ne = def-subst (lemma (od→ord x) (subst (λ k → ¬ (k == od∅ {suc n} )) (sym oiso) ne )) oiso refl where
      lemma : (ox : Ordinal {suc n}) →  ¬ (ord→od  ox  == od∅ {suc n} ) → ord→od ox ∋ od∅ {suc n}
@@ -264,7 +276,10 @@
 -- postulate f-extensionality : { n : Level}  → Relation.Binary.PropositionalEquality.Extensionality (suc n) (suc (suc n))
 
 csuc :  {n : Level} →  OD {suc n} → OD {suc n}
-csuc x = ord→od ( osuc ( od→ord x ))
+csuc x = Ord ( osuc ( od→ord x ))
+
+in-codomain : {n : Level} → (X : OD {suc n} ) → ( ψ : OD {suc n} → OD {suc n} ) → OD {suc n}
+in-codomain {n} X ψ = record { def = λ x → ¬ ( (y : Ordinal {suc n}) → ¬ ( def X y ∧  ( x ≡ od→ord (ψ (Ord y )))))  }
 
 -- Power Set of X ( or constructible by λ y → def X (od→ord y )
 
@@ -272,7 +287,7 @@
 ZFSubset A x =  record { def = λ y → def A y ∧  def x y }  
 
 Def :  {n : Level} → (A :  OD {suc n}) → OD {suc n}
-Def {n} A = ord→od ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) )  
+Def {n} A = Ord ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) )  
 
 -- Constructible Set on α
 L : {n : Level} → (α : Ordinal {suc n}) → OD {suc n}
@@ -281,8 +296,8 @@
 L {n}  record { lv = (Suc lx) ; ord = (Φ (Suc lx)) } = -- Union ( L α )
     record { def = λ y → osuc y o< (od→ord (L {n}  (record { lv = lx ; ord = Φ lx }) )) }
 
-OD→ZF : {n : Level} → ZF {suc (suc n)} {suc n}
-OD→ZF {n}  = record { 
+Ord→ZF : {n : Level} → ZF {suc (suc n)} {suc n}
+Ord→ZF {n}  = record { 
     ZFSet = OD {suc n}
     ; _∋_ = _∋_ 
     ; _≈_ = _==_ 
@@ -295,12 +310,14 @@
     ; infinite = ord→od ( record { lv = Suc Zero ; ord = Φ 1 } )
     ; isZF = isZF 
  } where
-    Replace : OD {suc n} → (OD {suc n} → OD {suc n} ) → OD {suc n}
-    Replace X ψ = sup-od ψ
     Select : OD {suc n} → (OD {suc n} → Set (suc n) ) → OD {suc n}
     Select X ψ = record { def = λ x →  ( def X  x ∧  ψ ( ord→od x )) } 
     _,_ : OD {suc n} → OD {suc n} → OD {suc n}
     x , y = record { def = λ z → z o< (omax (od→ord x) (od→ord y)) }
+    _∩_ : ( A B : OD {suc n} ) → OD
+    A ∩ B = record { def = λ x → def A x  ∧ def B x }
+    Replace : OD {suc n} → (OD {suc n} → OD {suc n} ) → OD {suc n}
+    Replace X ψ = record { def = λ x → (x o< sup-o ( λ x → od→ord (ψ (ord→od x )))) ∧ def (in-codomain X ψ) x }
     Union : OD {suc n} → OD {suc n}
     Union U = record { def = λ y → osuc y o< (od→ord U) }
     -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ ( ∀ {x} → t ∋ x →  X ∋ x )
@@ -311,8 +328,6 @@
     A ∈ B = B ∋ A
     _⊆_ : ( A B : ZFSet  ) → ∀{ x : ZFSet } →  Set (suc n)
     _⊆_ A B {x} = A ∋ x →  B ∋ x
-    -- _∩_ : ( A B : ZFSet  ) → ZFSet
-    -- A ∩ B = Select (A , B) (  λ x → ( A ∋ x ) ∧ (B ∋ x) )
     -- _∪_ : ( A B : ZFSet  ) → ZFSet
     -- A ∪ B = Select (A , B) (  λ x → (A ∋ x)  ∨ ( B ∋ x ) )
     infixr  200 _∈_
@@ -334,15 +349,15 @@
        ;   infinity∅ = infinity∅
        ;   infinity = λ _ → infinity
        ;   selection = λ {ψ} {X} {y} → selection {ψ} {X} {y}
-       ;   replacement = replacement
+       ;   replacement← = replacement←
+       ;   replacement→ = replacement→
      } where
-         open _∧_ 
-         open Minimumo
          pair : (A B : OD {suc n} ) → ((A , B) ∋ A) ∧  ((A , B) ∋ B)
          proj1 (pair A B ) = omax-x {n} (od→ord A) (od→ord B)
          proj2 (pair A B ) = omax-y {n} (od→ord A) (od→ord B)
          empty : (x : OD {suc n} ) → ¬  (od∅ ∋ x)
-         empty x ()
+         empty x (case1 ())
+         empty x (case2 ())
          ---
          --- ZFSubset A x =  record { def = λ y → def A y ∧  def x y }                   subset of A
          --- Power X = ord→od ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) )       Power X is a sup of all subset of A
@@ -398,8 +413,19 @@
               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) → Replace X ψ ∋ ψ x
-         replacement {ψ} X x = sup-c< ψ {x}
+         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 (proj2 lt)) where
+            lemma2 :  ¬ ((y : Ordinal) → ¬ def X y ∧ ((od→ord x) ≡ od→ord (ψ (Ord y))))
+                    → ¬ ((y : Ordinal) → ¬ def X y ∧ (ord→od (od→ord x) == ψ (Ord 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 y))) → (ord→od (od→ord x) == ψ (Ord 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 y)) )
+            lemma not y not2 = not (Ord y) (subst (λ k → k == ψ (Ord y)) oiso  ( proj2 not2 ))
          ∅-iso :  {x : OD} → ¬ (x == od∅) → ¬ ((ord→od (od→ord x)) == od∅) 
          ∅-iso {x} neq = subst (λ k → ¬ k) (=-iso {n} ) neq  
          minimul : (x : OD {suc n} ) → ¬ (x == od∅ )→ OD {suc n} 
@@ -407,10 +433,12 @@
          regularity :  (x : OD) (not : ¬ (x == od∅)) →
             (x ∋ minimul x not) ∧ (Select (minimul x not) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅)
          proj1 (regularity x not ) = ¬∅=→∅∈ not 
-         proj2 (regularity x not ) = record { eq→ = reg ; eq← = λ () } where
+         proj2 (regularity x not ) = record { eq→ = reg ; eq← = lemma } where
+            lemma : {ox : Ordinal} → def od∅ ox → def (Select (minimul x not) (λ y → (minimul x not ∋ y) ∧ (x ∋ y))) ox
+            lemma (case1 ())
+            lemma (case2 ())
             reg : {y : Ordinal} → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) y → def od∅ y
-            reg {y} t with proj1 t
-            ... | x∈∅ = x∈∅
+            reg {y} t = ⊥-elim ( ¬x<0 (proj1 (proj2 t )) )
          extensionality : {A B : OD {suc n}} → ((z : OD) → (A ∋ z) ⇔ (B ∋ z)) → A == B
          eq→ (extensionality {A} {B} eq ) {x} d = def-iso {suc n} {A} {B} (sym diso) (proj1 (eq (ord→od x))) d  
          eq← (extensionality {A} {B} eq ) {x} d = def-iso {suc n} {B} {A} (sym diso) (proj2 (eq (ord→od x))) d  
--- a/ordinal.agda	Mon Jun 10 09:53:45 2019 +0900
+++ b/ordinal.agda	Mon Jul 08 19:48:47 2019 +0900
@@ -97,6 +97,11 @@
 osuc-lveq : {n : Level} { x : Ordinal {n} } → lv x ≡ lv ( osuc x )
 osuc-lveq {n} = refl
 
+osucc : {n : Level} {ox oy : Ordinal {n}} → oy o< ox  → osuc oy o< osuc ox
+osucc {n} {ox} {oy} (case1 x) = case1 x
+osucc {n} {ox} {oy} (case2 x) with d<→lv x
+... | refl = case2 (s< x)
+
 nat-<> : { x y : Nat } → x < y → y < x → ⊥
 nat-<>  (s≤s x<y) (s≤s y<x) = nat-<> x<y y<x
 
@@ -113,9 +118,17 @@
 =→¬< {Zero} ()
 =→¬< {Suc x} (s≤s lt) = =→¬< lt
 
-o<¬≡ : {n : Level } ( ox oy : Ordinal {n}) → ox ≡ oy  → ox o< oy  → ⊥
-o<¬≡ ox ox refl (case1 lt) =  =→¬< lt
-o<¬≡ ox ox refl (case2 (s< lt)) = trio<≡ refl lt
+case12-⊥ : {n : Level} {x y : Ordinal {suc n}} → lv x < lv y → ord x d< ord y → ⊥
+case12-⊥ {x} {y} lt1 lt2 with d<→lv lt2
+... | refl = nat-≡< refl lt1
+
+case21-⊥ : {n : Level} {x y : Ordinal {suc n}} → lv x < lv y → ord y d< ord x → ⊥
+case21-⊥ {x} {y} lt1 lt2 with d<→lv lt2
+... | refl = nat-≡< refl lt1
+
+o<¬≡ : {n : Level } { ox oy : Ordinal {n}} → ox ≡ oy  → ox o< oy  → ⊥
+o<¬≡ {_} {ox} {ox} refl (case1 lt) =  =→¬< lt
+o<¬≡ {_} {ox} {ox} refl (case2 (s< lt)) = trio<≡ refl lt
 
 ¬x<0 : {n : Level} →  { x  : Ordinal {suc n} } → ¬ ( x o< o∅ {suc n} )
 ¬x<0 {n} {x} (case1 ())
@@ -165,6 +178,12 @@
 maxαd x y | tri≈ ¬a b ¬c = x
 maxαd x y | tri> ¬a ¬b c = x
 
+minαd : {n : Level} → { lx : Nat } → OrdinalD {n} lx  →  OrdinalD  lx  →  OrdinalD  lx
+minαd x y with triOrdd x y
+minαd x y | tri< a ¬b ¬c = x
+minαd x y | tri≈ ¬a b ¬c = y
+minαd x y | tri> ¬a ¬b c = x
+
 _o≤_ : {n : Level} → Ordinal → Ordinal → Set (suc n)
 a o≤ b  = (a ≡ b)  ∨ ( a o< b )
 
@@ -205,11 +224,23 @@
    lemma1 (case1 x) = ¬a x
    lemma1 (case2 x) = ≡→¬d< x
 
-maxα : {n : Level} →  Ordinal {n} →  Ordinal  → Ordinal
-maxα x y with <-cmp (lv x) (lv y)
-maxα x y | tri< a ¬b ¬c = x
-maxα x y | tri> ¬a ¬b c = y
-maxα x y | tri≈ ¬a refl ¬c = record { lv = lv x ; ord = maxαd (ord x) (ord y) }
+maxα : {n : Level} →  Ordinal {suc n} →  Ordinal  → Ordinal
+maxα x y with trio< x y
+maxα x y | tri< a ¬b ¬c = y
+maxα x y | tri> ¬a ¬b c = x
+maxα x y | tri≈ ¬a refl ¬c = x
+
+minα : {n : Level} →  Ordinal {suc n} →  Ordinal  → Ordinal
+minα {n} x y with trio< {n} x  y
+minα x y | tri< a ¬b ¬c = x
+minα x y | tri> ¬a ¬b c = y
+minα x y | tri≈ ¬a refl ¬c = x
+
+min1 : {n : Level} →  {x y z : Ordinal {suc n} } → z o< x → z o< y → z o< minα x y
+min1 {n} {x} {y} {z} z<x z<y with trio< {n} x y
+min1 {n} {x} {y} {z} z<x z<y | tri< a ¬b ¬c = z<x
+min1 {n} {x} {y} {z} z<x z<y | tri≈ ¬a refl ¬c = z<x
+min1 {n} {x} {y} {z} z<x z<y | tri> ¬a ¬b c = z<y
 
 --
 --  max ( osuc x , osuc y )
@@ -290,3 +321,15 @@
 TransFinite caseΦ caseOSuc record { lv = lx ; ord = (OSuc lx ox) } = 
       caseOSuc lx ox (TransFinite caseΦ caseOSuc  record { lv = lx ; ord = ox })
 
+--  (¬ (∀ y → ¬ ( ψ y ))) → (ψ y → p )  → p
+--      may be we can prove this...
+postulate 
+ TransFiniteExists : {n : Level} → { ψ : Ordinal {n} → Set n } 
+  → (exists : ¬ (∀ y → ¬ ( ψ y ) ))
+  → {p : Set n} ( P : { y : Ordinal {n} } →  ψ y → p )
+  → p
+
+-- TransFiniteExists {n} {ψ} exists {p} P = ⊥-elim ( exists lemma ) where
+--     lemma : (y : Ordinal {n} ) → ¬ ψ y
+--     lemma y ψy = ( TransFinite {n} {{!!}} {!!} {!!} y ) ψy
+   
--- a/zf.agda	Mon Jun 10 09:53:45 2019 +0900
+++ b/zf.agda	Mon Jul 08 19:48:47 2019 +0900
@@ -15,12 +15,16 @@
    case1 : A → A ∨ B
    case2 : B → A ∨ B
 
-_⇔_ : {n : Level } → ( A B : Set n )  → Set  n
+_⇔_ : {n m : Level } → ( A : Set n ) ( B : Set m )  → Set (n ⊔ m)
 _⇔_ A B =  ( A → B ) ∧ ( B → A )
 
+
 open import Relation.Nullary
 open import Relation.Binary
 
+contra-position : {n m : Level } {A : Set n} {B : Set m} → (A → B) → ¬ B → ¬ A 
+contra-position {n} {m} {A} {B}  f ¬b a = ¬b ( f a ) 
+
 infixr  130 _∧_
 infixr  140 _∨_
 infixr  150 _⇔_
@@ -33,7 +37,7 @@
      (_,_ : ( A B : ZFSet  ) → ZFSet)
      (Union : ( A : ZFSet  ) → ZFSet)
      (Power : ( A : ZFSet  ) → ZFSet)
-     (Select : ZFSet → ( ZFSet → Set m ) → ZFSet )
+     (Select :  (X : ZFSet  ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet ) 
      (Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet )
      (infinite : ZFSet)
        : Set (suc (n ⊔ m)) where
@@ -52,7 +56,7 @@
   _∩_ : ( A B : ZFSet  ) → ZFSet
   A ∩ B = Select A (  λ x → ( A ∋ x ) ∧ ( B ∋ x )  )
   _∪_ : ( A B : ZFSet  ) → ZFSet
-  A ∪ B = Union (A , B) 
+  A ∪ B = Union (A , B)    -- Select A (  λ x → ( A ∋ x ) ∨ ( B ∋ x )  ) is easer
   {_} : ZFSet → ZFSet
   { x } = ( x ,  x )
   infixr  200 _∈_
@@ -73,7 +77,11 @@
      infinity :  ∀( X x : ZFSet  ) → x ∈ infinite →  ( x ∪ { x }) ∈ infinite 
      selection : { ψ : ZFSet → Set m } → ∀ { X y : ZFSet  } →  ( ( y ∈ X ) ∧ ψ y ) ⇔ (y ∈  Select X ψ ) 
      -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) )
-     replacement : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet  ) →  ( ψ x ∈  Replace X ψ )  
+     replacement← : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet  ) → x ∈ X → ψ x ∈  Replace X ψ 
+     replacement→ : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet  ) →  ( lt : x ∈  Replace X ψ ) → ¬ ( ∀ (y : ZFSet)  →  ¬ ( x ≈ ψ y ) )
+   -- -- ∀ z [ ∀ x ( x ∈ z  → ¬ ( x ≈ ∅ ) )  ∧ ∀ x ∀ y ( x , y ∈ z ∧ ¬ ( x ≈ y )  → x ∩ y ≈ ∅  ) → ∃ u ∀ x ( x ∈ z → ∃ t ( u ∩ x) ≈ { t }) ]
+   -- axiom-of-choice : Set (suc n) 
+   -- axiom-of-choice = ?
 
 record ZF {n m : Level } : Set (suc (n ⊔ m)) where
   infixr  210 _,_
@@ -88,7 +96,7 @@
      _,_ : ( A B : ZFSet  ) → ZFSet
      Union : ( A : ZFSet  ) → ZFSet
      Power : ( A : ZFSet  ) → ZFSet
-     Select : ZFSet → ( ZFSet → Set m ) → ZFSet
+     Select :  (X : ZFSet  ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet 
      Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet
      infinite : ZFSet
      isZF : IsZF ZFSet _∋_ _≈_ ∅ _,_ Union Power Select Replace infinite