changeset 536:c43375ade2c5

remove unsed in zorn
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 23 Apr 2022 18:39:07 +0900
parents b83dde5dbd33
children e12add1519ec
files src/zorn.agda
diffstat 1 files changed, 0 insertions(+), 279 deletions(-) [+]
line wrap: on
line diff
--- a/src/zorn.agda	Sat Apr 23 18:35:20 2022 +0900
+++ b/src/zorn.agda	Sat Apr 23 18:39:07 2022 +0900
@@ -143,191 +143,6 @@
 -- OS<-cmp : {x : HOD} → Trichotomous {_} {IsTotalOrderSet x} _OS≈_ _OS<_ 
 -- OS<-cmp A B = {!!}
 
--- tree structure
-data IChain  (A : HOD) : Ordinal → Set n where
-    ifirst : {ox : Ordinal} → odef A ox → IChain A ox
-    inext  : {ox oy : Ordinal} → odef A oy → * ox < * oy → IChain A ox → IChain A oy
-
---   * ox < .. < * oy
-ic-connect :  {A : HOD} {oy : Ordinal} → (ox : Ordinal) →  (iy : IChain A oy) → Set n 
-ic-connect {A} ox (ifirst {oy} ay) = Lift n ⊥
-ic-connect {A} ox (inext {oy} {oz} ay y<z iz) = (ox ≡ oy) ∨ ic-connect ox iz 
-
-ic→odef : {A : HOD} {ox : Ordinal} → IChain A ox → odef A ox
-ic→odef {A} {ox} (ifirst ax) = ax
-ic→odef {A} {ox} (inext ax x<y ic) = ax
-
-ic→< :  {A : HOD} → IsPartialOrderSet A → (x : Ordinal) → odef A x → {y : Ordinal} → (iy : IChain A y) → ic-connect {A} {y} x iy → * x < * y
-ic→< {A} PO x ax {y} (ifirst ay) ()
-ic→< {A} PO x ax {y} (inext ay x<y iy) (case1 refl) = x<y
-ic→< {A} PO x ax {y} (inext {oy} ay x<y iy) (case2 ic) = IsStrictPartialOrder.trans PO 
-     {me (subst (λ k → odef A k) (sym &iso) ax )} {me (subst (λ k → odef A k) (sym &iso) (ic→odef {A} {oy} iy) ) }  {me (subst (λ k → odef A k) (sym &iso) ay) }
-    (ic→< {A} PO x ax iy ic ) x<y
-
-record IChained (A : HOD) (x y : Ordinal) : Set n where
-   field
-      iy : IChain A y
-      ic : ic-connect x iy 
-
---
--- all tree from x
---
-IChainSet : (A : HOD) {x : Ordinal} → odef A  x → HOD
-IChainSet A {x} ax = record { od = record { def = λ y → odef A y ∧ IChained A x y }
-    ; odmax = & A ; <odmax = λ {y} lt → subst (λ k → k o< & A) &iso (c<→o< (subst (λ k → odef A k) (sym &iso) (proj1 lt))) } 
-
-IChainSet⊆A :  {A : HOD} → {x : Ordinal } → (ax : odef A x ) → IChainSet A ax ⊆ A
-IChainSet⊆A {A} x = record { incl = λ {oy} y → proj1 y }
-
-¬IChained-refl : (A : HOD) {x : Ordinal} → IsPartialOrderSet A → ¬ IChained A x x
-¬IChained-refl A {x} PO record { iy = iy ; ic = ic } = IsStrictPartialOrder.irrefl PO
-        {me (subst (λ k → odef A k ) (sym &iso) ic0) } {me (subst (λ k → odef A k ) (sym &iso) ic0) } refl (ic→< {A} PO x ic0 iy ic )  where
-     ic0 : odef A x
-     ic0 = ic→odef {A} iy
-
--- there is a y, & y > & x
-
-record OSup> (A : HOD) {x : Ordinal} (ax : A ∋ * x) : Set n where
-   field
-      y : Ordinal
-      icy : odef (IChainSet A ax ) y 
-      y>x : x o< y
-
-record IChainSup> (A : HOD) {x : Ordinal} (ax : A ∋ * x) : Set n where
-   field
-      y : Ordinal
-      A∋y : odef A y
-      y>x  : * x < * y
-
--- finite IChain
---
---    tree structured
-
-ic→A∋y : (A : HOD) {x y : Ordinal}  (ax : A ∋ * x) → odef (IChainSet A ax) y → A ∋ * y
-ic→A∋y A {x} {y} ax ⟪ ay , _ ⟫ = subst (λ k → odef A k) (sym &iso) ay
-
-record InfiniteChain (A : HOD) (max : Ordinal) {x : Ordinal}  (ax : A ∋ * x) : Set n where
-   field
-      chain<x : (y : Ordinal ) → odef (IChainSet A ax) y →  y o< max
-      c-infinite : (y : Ordinal ) → (cy : odef (IChainSet A ax) y  )
-          → IChainSup> A (ic→A∋y A ax cy)
-
-open import Data.Nat hiding (_<_ ; _≤_ ) 
-import Data.Nat.Properties as NP
-open import nat
-
-data Chain (A : HOD) (s : Ordinal) (next : Ordinal  → Ordinal ) : ( x : Ordinal  ) → Set n where
-     cfirst : odef A s → Chain A s next s
-     csuc : (x : Ordinal ) → (ax : odef A x ) → Chain A s next x → odef A (next x) → Chain A s next (next x )
-
-ct∈A : (A : HOD ) (s : Ordinal)  → (next : Ordinal  → Ordinal ) → {x : Ordinal} → Chain A s next x → odef A x
-ct∈A A s next {x} (cfirst x₁) = x₁ 
-ct∈A A s next {.(next x )} (csuc x ax t anx) = anx
-
---
--- extract single chain from countable infinite chains
---
-TransitiveClosure : (A : HOD) (s : Ordinal) →  (next : Ordinal → Ordinal ) → HOD
-TransitiveClosure A s next = record { od = record { def = λ x → Chain A s next x } ; odmax = & A ; <odmax = cc01 } where
-    cc01 : {y : Ordinal} → Chain A s next y → y o< & A
-    cc01 {y} cy = subst (λ k → k o< & A ) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso) ( ct∈A A s next cy ) ) )
-
-cton0 : (A : HOD ) (s : Ordinal) → (next : Ordinal  → Ordinal )  {y : Ordinal } → Chain A s next y → ℕ
-cton0 A s next (cfirst _)  = zero
-cton0 A s next (csuc x ax z _) = suc (cton0 A s next z) 
-cton : (A : HOD ) (s : Ordinal)   → (next : Ordinal → Ordinal ) → Element (TransitiveClosure A s next) → ℕ
-cton A s next y = cton0 A s next (is-elm y)
-
-cinext :  (A : HOD) {x max : Ordinal } → (ax : A ∋ * x ) → (ifc : InfiniteChain A max ax ) → Ordinal  →  Ordinal
-cinext A ax ifc y with ODC.∋-p O (IChainSet A ax) (* y)
-... | yes ics-y = IChainSup>.y ( InfiniteChain.c-infinite ifc y (subst (λ k → odef (IChainSet A ax) k) &iso ics-y ))
-... | no _ = o∅
-
-InFCSet : (A : HOD) → {x max : Ordinal}  (ax : A ∋ * x)
-     → (ifc : InfiniteChain A max ax ) → HOD
-InFCSet A {x} ax ifc =  TransitiveClosure (IChainSet A ax) x (cinext A ax ifc ) 
-
-InFCSet⊆A : (A : HOD) → {x max : Ordinal}  (ax : A ∋ * x) →  (ifc : InfiniteChain A max ax ) → InFCSet A ax ifc ⊆ A
-InFCSet⊆A A {x} ax ifc = record { incl = λ {y} lt → incl (IChainSet⊆A ax) (
-     ct∈A (IChainSet A ax) x (cinext A ax ifc) lt ) }
-
-cinext→IChainSup :  (A : HOD) {x max : Ordinal } → (ax : A ∋ * x ) → (ifc : InfiniteChain A max ax ) → (y : Ordinal )
-    → (ay1 : IChainSet A ax ∋ * y ) → IChainSup> A (subst (odef A) (sym &iso) (proj1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1)))
-cinext→IChainSup A {x} ax ifc y ay with ODC.∋-p O (IChainSet A ax) (* y)
-... | no not = ⊥-elim ( not ay )
-... | yes ay1 = InfiniteChain.c-infinite ifc y (subst (λ k → odef (IChainSet A ax) k) &iso ay )
-
-TransitiveClosure-is-total : (A : HOD) → {x max : Ordinal}  (ax : A ∋ * x)
-     → IsPartialOrderSet A 
-     → (ifc : InfiniteChain A max ax )
-     → IsTotalOrderSet ( InFCSet A ax ifc )
-TransitiveClosure-is-total A {x} ax PO ifc = record { isEquivalence = IsStrictPartialOrder.isEquivalence IPO
-   ; trans = λ {x} {y} {z} → IsStrictPartialOrder.trans IPO {x} {y} {z}  ; compare = cmp } where
-    IPO : IsPartialOrderSet (InFCSet A ax ifc )
-    IPO = ⊆-IsPartialOrderSet record { incl = λ {y} lt → incl (InFCSet⊆A A {x} ax ifc) lt} PO
-    B = IChainSet A ax
-    cnext = cinext A ax ifc
-    ct02 : {oy : Ordinal} → (y : Chain B x cnext oy ) → A ∋ * oy 
-    ct02 y = incl (IChainSet⊆A {A} ax) (subst (λ k → odef (IChainSet A ax) k) (sym &iso) (ct∈A B x cnext y) ) 
-    ct-inject : {ox oy : Ordinal} → (x1 : Chain B x cnext ox ) → (y : Chain B x cnext oy )
-       → (cton0 B x cnext x1) ≡ (cton0 B x cnext y) → ox ≡ oy
-    ct-inject {ox} {ox} (cfirst x) (cfirst x₁) refl = refl
-    ct-inject {.(cnext x₀ )} {.(cnext x₃ )} (csuc x₀ ax x₁ x₂) (csuc x₃ ax₁ y x₄) eq = cong cnext ct05 where
-        ct06 : {x y : ℕ} → suc x ≡ suc y → x ≡ y
-        ct06 refl = refl 
-        ct05 : x₀ ≡ x₃
-        ct05 = ct-inject x₁ y (ct06 eq)
-    ct-monotonic : {ox oy : Ordinal} → (x1 : Chain B x cnext ox ) → (y : Chain B x cnext oy )
-       → (cton0 B x cnext x1) Data.Nat.< (cton0 B x cnext y) → * ox < * oy
-    ct-monotonic {ox} {oy} x1 (csuc oy1 ay y _) (s≤s lt) with NP.<-cmp ( cton0 B x cnext x1 ) ( cton0 B x cnext y )
-    ... | tri< a ¬b ¬c = ct07 where
-        ct07 :  * ox < * (cnext oy1)
-        ct07 with ODC.∋-p O (IChainSet A ax) (* oy1)
-        ... | no not  = ⊥-elim ( not (subst (λ k → odef (IChainSet A ax) k ) (sym &iso) ay ) )
-        ... | yes ay1 = IsStrictPartialOrder.trans PO {me (ct02 x1) } {me (ct02 y)} {me ct031 } (ct-monotonic x1 y a ) ct011 where
-           ct031 :  A ∋ * (IChainSup>.y (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 ) )) 
-           ct031 = subst (λ k → odef A k ) (sym &iso) (
-              IChainSup>.A∋y (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 )) )
-           ct011 :  * oy1 < * ( IChainSup>.y (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 )) )
-           ct011 = IChainSup>.y>x (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 )) 
-    ... | tri≈ ¬a b ¬c = ct11 where
-           ct11 : * ox < * (cnext oy1)
-           ct11 with ODC.∋-p O (IChainSet A ax) (* oy1)
-           ... | no not  = ⊥-elim ( not (subst (λ k → odef (IChainSet A ax) k ) (sym &iso) ay ) )
-           ... | yes ay1 = subst (λ k → * k < _) (sym (ct-inject _ _ b)) ct011  where
-              ct011 :  * oy1 < * ( IChainSup>.y (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 )) )
-              ct011 = IChainSup>.y>x (InfiniteChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet A ax) k) &iso ay1 )) 
-    ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> lt c ) 
-    ct12 : {y z : Element (TransitiveClosure B x cnext) } → elm y ≡ elm z → elm y < elm z → ⊥ 
-    ct12 {y} {z} y=z y<z = IsStrictPartialOrder.irrefl IPO {y} {z} y=z y<z
-    ct13 : {y z : Element (TransitiveClosure B x cnext) } → elm y < elm z → elm z < elm y → ⊥ 
-    ct13 {y} {z} y<z y>z = IsStrictPartialOrder.irrefl IPO {y} {y} refl ( IsStrictPartialOrder.trans IPO {y} {z} {y} y<z y>z )
-    ct17 : (x1 : Element (TransitiveClosure B x cnext)) → Chain B x cnext (& (elm x1))
-    ct17 x1 = is-elm x1
-    cmp : Trichotomous _ _ 
-    cmp x1 y with NP.<-cmp (cton B x cnext x1) (cton B x cnext y)
-    ... | tri< a ¬b ¬c = tri< ct04 ct14 ct15 where
-        ct04 : elm x1 < elm y
-        ct04 = subst₂ (λ j k → j < k  ) *iso *iso (ct-monotonic (is-elm x1) (is-elm y) a)
-        ct14 : ¬ elm x1 ≡  elm y
-        ct14 eq = ct12 {x1} {y} eq (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm x1) (is-elm y) a )  )
-        ct15 : ¬ (elm y <  elm x1)
-        ct15 lt = ct13 {y} {x1} lt (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm x1) (is-elm y) a )  )
-    ... | tri≈ ¬a b ¬c = tri≈ (ct12 {x1} {y} ct16)  ct16 (ct12 {y} {x1} (sym ct16)) where
-        ct16 :  elm x1 ≡ elm y
-        ct16 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) (ct-inject {& (elm x1)} {& (elm y)} (is-elm x1) (is-elm y) b ))
-    ... | tri> ¬a ¬b c = tri> ct15 ct14 ct04 where
-        ct04 : elm y < elm x1
-        ct04 = subst₂ (λ j k → j < k  ) *iso *iso (ct-monotonic (is-elm y) (is-elm x1) c)
-        ct14 : ¬ elm x1 ≡  elm y
-        ct14 eq = ct12 {y} {x1} (sym eq) (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm y) (is-elm x1) c )  )
-        ct15 : ¬ (elm x1 <  elm y)
-        ct15 lt = ct13 {x1} {y} lt (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm y) (is-elm x1) c )  )
-
-record IsFC (A : HOD) {x : Ordinal}  (ax : A ∋ * x) (y : Ordinal) : Set n where
-   field
-      icy : odef (IChainSet A ax) y  
-      c-finite : ¬ IChainSup> A (subst (λ k → odef A k ) (sym &iso) (proj1 icy) )
       
 record Maximal ( A : HOD )  : Set (Level.suc n) where
    field
@@ -335,100 +150,6 @@
       A∋maximal : A ∋ maximal
       ¬maximal<x : {x : HOD} → A ∋ x  → ¬ maximal < x       -- A is Partial, use negative
 
---
--- possible three cases in a limit ordinal step
--- 
--- case 1) < goes x o<
--- case 2) no > x in some chain ( maximal )
--- case 3) countably infinite chain below x 
---
-Zorn-lemma-3case : { A : HOD } 
-    → o∅ o< & A 
-    → IsPartialOrderSet A 
-    → (x : Ordinal ) → (ax : odef A x)  → OSup> A (d→∋ A ax) ∨ Maximal A ∨ InfiniteChain A x  (d→∋ A ax)
-Zorn-lemma-3case {A}  0<A PO x ax = zc2 where
-    Gtx : HOD
-    Gtx = record { od = record { def = λ y → odef ( IChainSet A ax ) y ∧  ( x o< y ) } ; odmax = & A
-        ; <odmax = λ lt → subst (λ k → k o< & A) &iso (c<→o< (d→∋  A (proj1 (proj1 lt))))  }
-    HG : HOD
-    HG = record { od = record { def = λ y → odef A y ∧ IsFC A (d→∋ A ax ) y } ; odmax = & A
-        ; <odmax = λ lt → subst (λ k → k o< & A) &iso (c<→o< (d→∋  A  (proj1 lt) ))  }
-    zc2 :  OSup> A (d→∋ A ax)  ∨ Maximal A ∨ InfiniteChain A x  (d→∋ A ax )
-    zc2 with  is-o∅ (& Gtx)
-    ... | no not = case1 record { y = & y ; icy = zc4 ; y>x = proj2 zc3 } where
-        y : HOD
-        y =  ODC.minimal O Gtx  (λ eq → not (=od∅→≡o∅ eq))
-        zc3 :  odef ( IChainSet A ax ) (& y) ∧  ( x o< (& y ))
-        zc3  = ODC.x∋minimal O Gtx  (λ eq → not (=od∅→≡o∅ eq))
-        zc4 : odef (IChainSet A (subst (OD.def (od A)) (sym &iso) ax)) (& y)
-        zc4 = ⟪ proj1 (proj1 zc3) , (subst (λ k → IChained A k (& y)) (sym &iso) (proj2 (proj1 zc3))) ⟫ 
-    ... | yes nogt with is-o∅ (& HG)
-    ... | no  finite-chain = case2 (case1 record { maximal = y ; A∋maximal = proj1 zc3 ; ¬maximal<x = zc4 } ) where
-        y : HOD
-        y =  ODC.minimal O HG (λ eq → finite-chain (=od∅→≡o∅ eq))
-        zc3 :  odef A (& y) ∧ IsFC A (d→∋ A ax ) (& y)
-        zc3  = ODC.x∋minimal O HG  (λ eq → finite-chain (=od∅→≡o∅ eq))
-        zc4 : {z : HOD} → A ∋ z → ¬ (y < z)
-        zc4 {z} az y<z = IsFC.c-finite (proj2 zc3) record { y = & z ; A∋y =  az ; y>x = subst₂ (λ j k → j < k ) (sym *iso) (sym *iso) y<z } 
-    ... | yes inifite = case2 (case2 record {    c-infinite = zc91 ; chain<x = zc10 } ) where
-        B : HOD
-        B = IChainSet A ax -- (me (subst (OD.def (od A)) (sym &iso) (is-elm x)))
-        B1 : HOD
-        B1 = IChainSet A (subst (OD.def (od A)) (sym &iso) ax)
-        Nx : (y : Ordinal) → odef A y → HOD
-        Nx y ay = record { od = record { def = λ x → odef A x ∧ ( * y < * x ) } ; odmax = & A
-              ; <odmax = λ lt → subst (λ k → k o< & A) &iso (c<→o< (d→∋  A (proj1 lt)))  }
-        zc10 : (y : Ordinal) → odef (IChainSet A (subst (OD.def (od A)) (sym &iso) ax)) y → y o< x
-        zc10 oy icsy = zc21 where
-             zc20 : (y : HOD) → (IChainSet A ax) ∋ y  → x o< & y → ⊥
-             zc20 y icsy lt = ¬A∋x→A≡od∅ Gtx ⟪ icsy , lt ⟫ nogt
-             zc22 : IChainSet A ax ∋ * oy
-             zc22 = ⟪ subst (λ k → odef A k) (sym &iso) (proj1 icsy) , subst₂ (λ j k → IChained A j k ) &iso (sym &iso) (proj2 icsy) ⟫
-             zc21 : oy o< x
-             zc21 with trio< oy  x
-             ... | tri< a ¬b ¬c = a
-             ... | tri≈ ¬a b ¬c = ⊥-elim (¬IChained-refl A PO (subst₂ (λ j k → IChained A j k ) &iso b (proj2 icsy)) ) 
-             ... | tri> ¬a ¬b c = ⊥-elim ( zc20 (* oy) zc22 (subst (λ k → x o< k) (sym &iso) c ))
-        zc91 : (y : Ordinal) (cy : odef B1 y) → IChainSup> A (ic→A∋y A (subst (OD.def (od A)) (sym &iso) ax) cy)
-        zc91 y cy with is-o∅ (& (Nx y (proj1 cy) ))
-        ... | yes no-next = ⊥-elim zc16 where
-             zc18 :   ¬ IChainSup> A (subst (odef A) (sym &iso) (proj1 (subst (λ k → odef (IChainSet A (d→∋ A ax)) k) (sym &iso) cy)))
-             zc18 ics = ¬A∋x→A≡od∅ (Nx y (proj1 cy) ) ⟪ subst (λ k → odef A k ) (sym &iso) (IChainSup>.A∋y ics)
-                  ,  subst₂ (λ j k → j < k ) *iso (cong (*) (sym &iso))( IChainSup>.y>x ics) ⟫ no-next  
-             zc17 : IsFC A {x} (d→∋ A ax) (& (* y))
-             zc17 = record { icy = subst (λ k → odef (IChainSet A (d→∋ A ax)) k) (sym &iso) cy ; c-finite = zc18 }
-             zc16 : ⊥
-             zc16 = ¬A∋x→A≡od∅ HG ⟪ subst (λ k → odef A k ) (sym &iso) (proj1 cy ) , zc17 ⟫ inifite 
-        ... | no not = record { y = & zc13 ; A∋y = proj1 zc12  ; y>x = proj2 zc12 }  where
-             zc13 = ODC.minimal O (Nx y (proj1 cy)) (λ eq → not (=od∅→≡o∅ eq ))
-             zc12 : odef A (& zc13 ) ∧ ( * y < * ( & zc13 ))
-             zc12 = ODC.x∋minimal O (Nx y (proj1 cy)) (λ eq → not (=od∅→≡o∅ eq ))
-
-all-climb-case : { A : HOD } → (0<A : o∅ o< & A) → IsPartialOrderSet A
-     → (( x : Ordinal ) → (ax : odef A (& (* x))) → OSup> A ax )
-     → (x : HOD) ( ax : A ∋ x )
-     → InfiniteChain A (& A) (d→∋ A ax)
-all-climb-case {A} 0<A PO climb x ax = record {    c-infinite = ac00 ; chain<x = ac01 }  where
-        B = IChainSet A ax
-        ac01 : (y : Ordinal) → odef (IChainSet A (d→∋ A ax)) y → y o< & A 
-        ac01 y ⟪ ay , _ ⟫ = subst (λ k → k o< & A ) &iso (c<→o< (subst (λ k → odef A k ) (sym &iso) ay) )
-        ac00 : (y : Ordinal) (cy : odef (IChainSet A (d→∋ A ax)) y) → IChainSup> A (ic→A∋y A (d→∋ A ax) cy)
-        ac00 y cy = record { y = z ; A∋y = az ; y>x = y<z} where
-             ay : odef A (& (* y))
-             ay = subst (λ k → odef A k) (sym &iso) (proj1 cy)
-             z : Ordinal
-             z = OSup>.y ( climb y  ay)
-             az : odef A z
-             az = subst (λ k → odef A k) &iso ( incl (IChainSet⊆A {A} ay ) (subst (λ k → odef (IChainSet A ay) k ) (sym &iso) (OSup>.icy ( climb y ay))))
-             icy :  odef (IChainSet A ay ) z
-             icy  = OSup>.icy ( climb y ay )
-             y<z  : * y < * z
-             y<z  = ic→< {A} PO y (subst (λ k → odef A k) &iso ay) (IChained.iy (proj2 icy))
-               (subst (λ k → ic-connect k (IChained.iy (proj2 icy))) &iso (IChained.ic (proj2 icy)))
-
--- <-TransFinite : ( A : HOD ) → IsTotalOrderSet A 
---          → ( (x : Ordinal) → ((y : Ordinal) → y o< x → ZChain A y) → ZChain A x ) → (x : Ordinal ) → ZChain A x
--- <-TransFinite A TA ind x = TransFinite {ZChain A} ind x 
 
 --
 -- inductive maxmum tree from x