changeset 966:39c680188738

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 05 Nov 2022 13:21:42 +0900
parents 1c1c6a6ed4fa
children cd0ef83189c5
files src/zorn.agda
diffstat 1 files changed, 381 insertions(+), 278 deletions(-) [+]
line wrap: on
line diff
--- a/src/zorn.agda	Sat Nov 05 10:23:57 2022 +0900
+++ b/src/zorn.agda	Sat Nov 05 13:21:42 2022 +0900
@@ -1,30 +1,30 @@
 {-# OPTIONS --allow-unsolved-metas #-}
 open import Level hiding ( suc ; zero )
 open import Ordinals
-open import Relation.Binary 
+open import Relation.Binary
 open import Relation.Binary.Core
 open import Relation.Binary.PropositionalEquality
-import OD 
+import OD
 module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) (PO : IsStrictPartialOrder _≡_ _<_ ) where
 
 --
--- Zorn-lemma : { A : HOD } 
---     → o∅ o< & A 
+-- Zorn-lemma : { A : HOD }
+--     → o∅ o< & A
 --     → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B   ) -- SUP condition
---     → Maximal A 
+--     → Maximal A
 --
 
 open import zf
 open import logic
 -- open import partfunc {n} O
 
-open import Relation.Nullary 
-open import Data.Empty 
-import BAlgbra 
+open import Relation.Nullary
+open import Data.Empty
+import BAlgbra
 
 open import Data.Nat hiding ( _<_ ; _≤_ )
-open import Data.Nat.Properties 
-open import nat 
+open import Data.Nat.Properties
+open import nat
 
 
 open inOrdinal O
@@ -52,42 +52,42 @@
 -- Partial Order on HOD ( possibly limited in A )
 --
 
-_<<_ : (x y : Ordinal ) → Set n 
+_<<_ : (x y : Ordinal ) → Set n
 x << y = * x < * y
 
 _<=_ : (x y : Ordinal ) → Set n    -- we can't use * x ≡ * y, it is Set (Level.suc n). Level (suc n) troubles Chain
 x <= y = (x ≡ y ) ∨ ( * x < * y )
 
-POO : IsStrictPartialOrder _≡_ _<<_ 
-POO = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans } 
-   ; trans = IsStrictPartialOrder.trans PO 
+POO : IsStrictPartialOrder _≡_ _<<_
+POO = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans }
+   ; trans = IsStrictPartialOrder.trans PO
    ; irrefl = λ x=y x<y → IsStrictPartialOrder.irrefl PO (cong (*) x=y) x<y
-   ; <-resp-≈ =  record { fst = λ {x} {y} {y1} y=y1 xy1 → subst (λ k → x << k ) y=y1 xy1 ; snd = λ {x} {x1} {y} x=x1 x1y → subst (λ k → k << x ) x=x1 x1y } } 
- 
+   ; <-resp-≈ =  record { fst = λ {x} {y} {y1} y=y1 xy1 → subst (λ k → x << k ) y=y1 xy1 ; snd = λ {x} {x1} {y} x=x1 x1y → subst (λ k → k << x ) x=x1 x1y } }
+
 _≤_ : (x y : HOD) → Set (Level.suc n)
 x ≤ y = ( x ≡ y ) ∨ ( x < y )
 
-≤-ftrans : {x y z : HOD} → x ≤ y → y ≤ z → x ≤ z 
+≤-ftrans : {x y z : HOD} → x ≤ y → y ≤ z → x ≤ z
 ≤-ftrans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl
 ≤-ftrans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z
 ≤-ftrans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y
 ≤-ftrans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z )
 
-<=-trans : {x y z : Ordinal } →  x <=  y →  y <=  z →  x <=  z 
+<=-trans : {x y z : Ordinal } →  x <=  y →  y <=  z →  x <=  z
 <=-trans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl
 <=-trans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z
 <=-trans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y
 <=-trans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z )
 
-ftrans<=-< : {x y z : Ordinal } →  x <=  y →  y << z →  x <<  z 
+ftrans<=-< : {x y z : Ordinal } →  x <=  y →  y << z →  x <<  z
 ftrans<=-< {x} {y} {z} (case1 eq) y<z = subst (λ k → k < * z) (sym (cong (*) eq))  y<z
-ftrans<=-< {x} {y} {z} (case2 lt) y<z = IsStrictPartialOrder.trans PO lt y<z 
+ftrans<=-< {x} {y} {z} (case2 lt) y<z = IsStrictPartialOrder.trans PO lt y<z
 
-<=to≤ : {x y : Ordinal } → x <= y → * x ≤ * y 
+<=to≤ : {x y : Ordinal } → x <= y → * x ≤ * y
 <=to≤ (case1 eq) = case1 (cong (*) eq)
 <=to≤ (case2 lt) = case2 lt
 
-≤to<= : {x y : Ordinal } → * x ≤ * y → x <= y 
+≤to<= : {x y : Ordinal } → * x ≤ * y → x <= y
 ≤to<= (case1 eq) = case1 ( subst₂ (λ j k → j ≡ k ) &iso &iso  (cong (&) eq) )
 ≤to<= (case2 lt) = case2 lt
 
@@ -101,7 +101,7 @@
 open _==_
 open _⊆_
 
--- <-TransFinite : {A x : HOD} → {P : HOD → Set n} → x ∈ A 
+-- <-TransFinite : {A x : HOD} → {P : HOD → Set n} → x ∈ A
 --     → ({x : HOD} → A ∋ x →  ({y : HOD} → A ∋  y → y < x → P y ) → P x) → P x
 -- <-TransFinite = ?
 
@@ -122,13 +122,13 @@
 
 A∋fcs : {A : HOD} (s : Ordinal) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → odef A s
 A∋fcs {A} s f mf (init as refl) = as
-A∋fcs {A} s f mf (fsuc y fcy) = A∋fcs {A} s f mf fcy 
+A∋fcs {A} s f mf (fsuc y fcy) = A∋fcs {A} s f mf fcy
 
 s≤fc : {A : HOD} (s : Ordinal ) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → * s ≤ * y
 s≤fc {A} s {.s} f mf (init x refl ) = case1 refl
 s≤fc {A} s {.(f x)} f mf (fsuc x fcy) with proj1 (mf x (A∋fc s f mf fcy ) )
 ... | case1 x=fx = subst (λ k → * s ≤ * k ) (*≡*→≡ x=fx) ( s≤fc {A} s f mf fcy )
-... | case2 x<fx with s≤fc {A} s f mf fcy 
+... | case2 x<fx with s≤fc {A} s f mf fcy
 ... | case1 s≡x = case2 ( subst₂ (λ j k → j < k ) (sym s≡x) refl x<fx )
 ... | case2 s<x = case2 ( IsStrictPartialOrder.trans PO s<x x<fx )
 
@@ -138,12 +138,12 @@
 ... | case1 eq = fcn s mf p
 ... | case2 y<fy = suc (fcn s mf p )
 
-fcn-inject : {A : HOD} (s : Ordinal) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) 
+fcn-inject : {A : HOD} (s : Ordinal) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f)
      → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx  ≡ fcn s mf cy → * x ≡ * y
 fcn-inject {A} s {x} {y} {f} mf cx cy eq = fc00 (fcn s mf cx) (fcn s mf cy) eq cx cy refl refl where
      fc06 : {y : Ordinal } { as : odef A s } (eq : s ≡ y  ) { j : ℕ } →  ¬ suc j ≡ fcn {A} s {y} {f} mf (init as eq )
      fc06 {x} {y} refl {j} not = fc08 not where
-        fc08 :  {j : ℕ} → ¬ suc j ≡ 0 
+        fc08 :  {j : ℕ} → ¬ suc j ≡ 0
         fc08 ()
      fc07 : {x : Ordinal } (cx : FClosure A f s x ) → 0 ≡ fcn s mf cx → * s ≡ * x
      fc07 {x} (init as refl) eq = refl
@@ -174,7 +174,7 @@
           fc03 : (y1 : Ordinal) → (cy1 : FClosure A f s y1 ) →  suc j ≡ fcn s mf cy1 → * (f x)  ≡ * y1
           fc03 y1 (init x₁ x₂) x = ⊥-elim (fc06 x₂ x)
           fc03 .(f y1) (fsuc y1 cy1) j=y1 with proj1 (mf y1 (A∋fc s f mf cy1 ) )
-          ... | case1 eq = trans ( fc03  y1 cy1 j=y1 ) eq 
+          ... | case1 eq = trans ( fc03  y1 cy1 j=y1 ) eq
           ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc05) where
                fc05 : * x ≡ * y1
                fc05 = fc00 i j (cong pred i=j) cx cy1 (cong pred i=x) (cong pred j=y1)
@@ -186,23 +186,23 @@
 fcn-< {A} s {x} {y} {f} mf cx cy x<y = fc01 (fcn s mf cy) cx cy refl x<y where
      fc06 : {y : Ordinal } { as : odef A s } (eq : s ≡ y  ) { j : ℕ } →  ¬ suc j ≡ fcn {A} s {y} {f} mf (init as eq )
      fc06 {x} {y} refl {j} not = fc08 not where
-        fc08 :  {j : ℕ} → ¬ suc j ≡ 0 
+        fc08 :  {j : ℕ} → ¬ suc j ≡ 0
         fc08 ()
      fc01 : (i : ℕ ) → {y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → (i ≡ fcn s mf cy ) → fcn s mf cx Data.Nat.< i → * x < * y
      fc01 (suc i) cx (init x₁ x₂) x (s≤s x₃) = ⊥-elim (fc06 x₂ x)
      fc01 (suc i) {y} cx (fsuc y1 cy) i=y (s≤s x<i) with proj1 (mf y1 (A∋fc s f mf cy ) )
-     ... | case1 y=fy = subst (λ k → * x < k ) y=fy ( fc01 (suc i) {y1} cx cy i=y (s≤s x<i)  ) 
+     ... | case1 y=fy = subst (λ k → * x < k ) y=fy ( fc01 (suc i) {y1} cx cy i=y (s≤s x<i)  )
      ... | case2 y<fy with <-cmp (fcn s mf cx ) i
      ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> x<i c )
-     ... | tri≈ ¬a b ¬c = subst (λ k → k < * (f y1) ) (fcn-inject s mf cy cx (sym (trans b (cong pred i=y) ))) y<fy 
+     ... | tri≈ ¬a b ¬c = subst (λ k → k < * (f y1) ) (fcn-inject s mf cy cx (sym (trans b (cong pred i=y) ))) y<fy
      ... | tri< a ¬b ¬c = IsStrictPartialOrder.trans PO fc02 y<fy where
           fc03 :  suc i ≡ suc (fcn s mf cy) → i ≡ fcn s mf cy
-          fc03 eq = cong pred eq 
-          fc02 :  * x < * y1 
+          fc03 eq = cong pred eq
+          fc02 :  * x < * y1
           fc02 =  fc01 i cx cy (fc03 i=y ) a
 
 
-fcn-cmp : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) 
+fcn-cmp : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f)
     → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → Tri (* x < * y) (* x ≡ * y) (* y < * x )
 fcn-cmp {A} s {x} {y} f mf cx cy with <-cmp ( fcn s mf cx ) (fcn s mf cy )
 ... | tri< a ¬b ¬c = tri< fc11 (λ eq → <-irr (case1 (sym eq)) fc11) (λ lt → <-irr (case2 fc11) lt) where
@@ -211,7 +211,7 @@
 ... | tri≈ ¬a b ¬c = tri≈ (λ lt → <-irr (case1 (sym fc10)) lt) fc10 (λ lt → <-irr (case1 fc10) lt) where
       fc10 : * x ≡ * y
       fc10 = fcn-inject {A} s {x} {y} {f} mf cx cy b
-... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12  where 
+... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12  where
       fc12 : * y < * x
       fc12 = fcn-< {A} s {y} {x} {f} mf cy cx c
 
@@ -238,7 +238,7 @@
       ax : odef A x
       y : Ordinal
       ay : odef B y
-      x=fy :  x ≡ f y 
+      x=fy :  x ≡ f y
 
 record IsSUP (A B : HOD) {x : Ordinal } (xa : odef A x)     : Set n where
    field
@@ -247,7 +247,7 @@
 record IsMinSUP (A B : HOD) ( f : Ordinal → Ordinal ) {x : Ordinal } (xa : odef A x)     : Set n where
    field
       x≤sup   : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x )
-      minsup : { sup1 : Ordinal } → odef A sup1 
+      minsup : { sup1 : Ordinal } → odef A sup1
          →  ( {z : Ordinal  } → odef B z → (z ≡ sup1 ) ∨ (z << sup1 ))  → x o≤ sup1
       not-hp : ¬ ( HasPrev A B f x )
 
@@ -261,18 +261,19 @@
 --  sup and its fclosure is in a chain HOD
 --    chain HOD is sorted by sup as Ordinal and <-ordered
 --    whole chain is a union of separated Chain
---    minimum index is sup of y not ϕ 
+--    minimum index is sup of y not ϕ
 --
 
 record ChainP (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal) (u : Ordinal) : Set n where
    field
-      fcy<sup  : {z : Ordinal } → FClosure A f y z → (z ≡ supf u) ∨ ( z << supf u ) 
+      fcy<sup  : {z : Ordinal } → FClosure A f y z → (z ≡ supf u) ∨ ( z << supf u )
       order    : {s z1 : Ordinal} → (lt : supf s o< supf u ) → FClosure A f (supf s ) z1 → (z1 ≡ supf u ) ∨ ( z1 << supf u )
       supu=u   : supf u ≡ u
 
 data UChain  ( A : HOD )    ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y )
-       (supf : Ordinal → Ordinal) (x : Ordinal) : (z : Ordinal) → Set n where 
-    ch-is-sup  : (u : Ordinal) {z : Ordinal }  (u<x : supf u o< supf x) ( is-sup : ChainP A f mf ay supf u ) 
+       (supf : Ordinal → Ordinal) (x : Ordinal) : (z : Ordinal) → Set n where
+    ch-init : {z : Ordinal } (fc : FClosure A f y z) → UChain A f mf ay supf x z
+    ch-is-sup  : (u : Ordinal) {z : Ordinal }  (u<x : supf u o< supf x) ( is-sup : ChainP A f mf ay supf u )
         ( fc : FClosure A f (supf u) z ) → UChain A f mf ay supf x z
 
 --
@@ -287,36 +288,67 @@
    {s s1 a b : Ordinal } ( ca : UChain A f mf ay supf s a ) ( cb : UChain A f mf ay supf s1 b ) → Tri (* a < * b) (* a ≡ * b) (* b < * a )
 chain-total A f mf {y} ay supf {xa} {xb} {a} {b} ca cb = ct-ind xa xb ca cb where
      ct-ind : (xa xb : Ordinal) → {a b : Ordinal} → UChain A f mf ay supf xa a → UChain A f mf ay supf xb b → Tri (* a < * b) (* a ≡ * b) (* b < * a)
+     ct-ind xa xb {a} {b} (ch-init fca) (ch-init fcb) = fcn-cmp y f mf fca fcb
+     ct-ind xa xb {a} {b} (ch-init fca) (ch-is-sup ub u<x supb fcb) with ChainP.fcy<sup supb fca
+     ... | case1 eq with s≤fc (supf ub) f mf fcb
+     ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00  (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
+          ct00 : * a ≡ * b
+          ct00 = trans (cong (*) eq) eq1
+     ... | case2 lt = tri< ct01  (λ eq → <-irr (case1 (sym eq)) ct01) (λ lt → <-irr (case2 ct01) lt)  where
+          ct01 : * a < * b
+          ct01 = subst (λ k → * k < * b ) (sym eq) lt
+     ct-ind xa xb {a} {b} (ch-init fca) (ch-is-sup ub u<x supb fcb) | case2 lt = tri< ct01  (λ eq → <-irr (case1 (sym eq)) ct01) (λ lt → <-irr (case2 ct01) lt)  where
+          ct00 : * a < * (supf ub)
+          ct00 = lt
+          ct01 : * a < * b
+          ct01 with s≤fc (supf ub) f mf fcb
+          ... | case1 eq =  subst (λ k → * a < k ) eq ct00
+          ... | case2 lt =  IsStrictPartialOrder.trans POO ct00 lt
+     ct-ind xa xb {a} {b} (ch-is-sup ua u<x supa fca) (ch-init fcb) with ChainP.fcy<sup supa fcb
+     ... | case1 eq with s≤fc (supf ua) f mf fca
+     ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00  (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
+          ct00 : * a ≡ * b
+          ct00 = sym (trans (cong (*) eq) eq1 )
+     ... | case2 lt = tri> (λ lt → <-irr (case2 ct01) lt) (λ eq → <-irr (case1 eq) ct01) ct01    where
+          ct01 : * b < * a
+          ct01 = subst (λ k → * k < * a ) (sym eq) lt
+     ct-ind xa xb {a} {b} (ch-is-sup ua u<x supa fca) (ch-init fcb) | case2 lt = tri> (λ lt → <-irr (case2 ct01) lt) (λ eq → <-irr (case1 eq) ct01) ct01    where
+          ct00 : * b < * (supf ua)
+          ct00 = lt
+          ct01 : * b < * a
+          ct01 with s≤fc (supf ua) f mf fca
+          ... | case1 eq =  subst (λ k → * b < k ) eq ct00
+          ... | case2 lt =  IsStrictPartialOrder.trans POO ct00 lt
      ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) with trio< ua ub
-     ... | tri< a₁ ¬b ¬c with ChainP.order supb  (subst₂ (λ j k → j o< k ) (sym (ChainP.supu=u supa )) (sym (ChainP.supu=u supb )) a₁)  fca 
-     ... | case1 eq with s≤fc (supf ub) f mf fcb 
+     ... | tri< a₁ ¬b ¬c with ChainP.order supb  (subst₂ (λ j k → j o< k ) (sym (ChainP.supu=u supa )) (sym (ChainP.supu=u supb )) a₁)  fca
+     ... | case1 eq with s≤fc (supf ub) f mf fcb
      ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00  (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
           ct00 : * a ≡ * b
           ct00 = trans (cong (*) eq) eq1
      ... | case2 lt =  tri< ct02  (λ eq → <-irr (case1 (sym eq)) ct02) (λ lt → <-irr (case2 ct02) lt)  where
-          ct02 : * a < * b 
+          ct02 : * a < * b
           ct02 = subst (λ k → * k < * b ) (sym eq) lt
      ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) | tri< a₁ ¬b ¬c | case2 lt = tri< ct02  (λ eq → <-irr (case1 (sym eq)) ct02) (λ lt → <-irr (case2 ct02) lt)  where
           ct03 : * a < * (supf ub)
           ct03 = lt
-          ct02 : * a < * b 
+          ct02 : * a < * b
           ct02 with s≤fc (supf ub) f mf fcb
           ... | case1 eq =  subst (λ k → * a < k ) eq ct03
           ... | case2 lt =  IsStrictPartialOrder.trans POO ct03 lt
-     ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x  supb fcb) | tri≈ ¬a  eq ¬c 
+     ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x  supb fcb) | tri≈ ¬a  eq ¬c
          = fcn-cmp (supf ua) f mf fca (subst (λ k → FClosure A f k b ) (cong supf (sym eq)) fcb )
-     ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) | tri> ¬a ¬b c with ChainP.order supa (subst₂ (λ j k → j o< k ) (sym (ChainP.supu=u supb )) (sym (ChainP.supu=u supa )) c) fcb 
-     ... | case1 eq with s≤fc (supf ua) f mf fca 
+     ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) | tri> ¬a ¬b c with ChainP.order supa (subst₂ (λ j k → j o< k ) (sym (ChainP.supu=u supb )) (sym (ChainP.supu=u supa )) c) fcb
+     ... | case1 eq with s≤fc (supf ua) f mf fca
      ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00  (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where
           ct00 : * a ≡ * b
           ct00 = sym (trans (cong (*) eq) eq1)
      ... | case2 lt =  tri> (λ lt → <-irr (case2 ct02) lt) (λ eq → <-irr (case1 eq) ct02) ct02    where
-          ct02 : * b < * a 
+          ct02 : * b < * a
           ct02 = subst (λ k → * k < * a ) (sym eq) lt
      ct-ind xa xb {a} {b} (ch-is-sup ua ua<x supa fca) (ch-is-sup ub ub<x supb fcb) | tri> ¬a ¬b c | case2 lt = tri> (λ lt → <-irr (case2 ct04) lt) (λ eq → <-irr (case1 (eq)) ct04) ct04    where
           ct05 : * b < * (supf ua)
           ct05 = lt
-          ct04 : * b < * a 
+          ct04 : * b < * a
           ct04 with s≤fc (supf ua) f mf fca
           ... | case1 eq =  subst (λ k → * b < k ) eq ct05
           ... | case2 lt =  IsStrictPartialOrder.trans POO ct05 lt
@@ -326,13 +358,13 @@
 
 -- Union of supf z which o< x
 --
-UnionCF : ( A : HOD )    ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y ) 
+UnionCF : ( A : HOD )    ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y )
     ( supf : Ordinal → Ordinal ) ( x : Ordinal ) → HOD
 UnionCF A f mf ay supf x
    = record { od = record { def = λ z → odef A z ∧ UChain A f mf ay supf x z } ; odmax = & A ; <odmax = λ {y} sy → ∈∧P→o< sy }
 
-supf-inject0 : {x y : Ordinal } {supf : Ordinal → Ordinal } → (supf-mono : {x y : Ordinal } →  x o≤  y  → supf x o≤ supf y )   
-   → supf x o< supf y → x o<  y 
+supf-inject0 : {x y : Ordinal } {supf : Ordinal → Ordinal } → (supf-mono : {x y : Ordinal } →  x o≤  y  → supf x o≤ supf y )
+   → supf x o< supf y → x o<  y
 supf-inject0 {x} {y} {supf} supf-mono sx<sy with trio< x y
 ... | tri< a ¬b ¬c = a
 ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy )
@@ -344,8 +376,8 @@
    field
       sup : Ordinal
       asm : odef A sup
-      x≤sup : {x : Ordinal } → odef B x → (x ≡ sup ) ∨ (x << sup )   
-      minsup : { sup1 : Ordinal } → odef A sup1 
+      x≤sup : {x : Ordinal } → odef B x → (x ≡ sup ) ∨ (x << sup )
+      minsup : { sup1 : Ordinal } → odef A sup1
          →  ( {x : Ordinal  } → odef B x → (x ≡ sup1 ) ∨ (x << sup1 ))  → sup o≤ sup1
 
 z09 : {b : Ordinal } { A : HOD } → odef A b → b o< & A
@@ -353,37 +385,38 @@
 
 M→S  : { A : HOD } { f : Ordinal → Ordinal } {mf : ≤-monotonic-f A f} {y : Ordinal} {ay : odef A y}  { x : Ordinal }
       →  (supf : Ordinal → Ordinal )
-      →  MinSUP A (UnionCF A f mf ay supf x)  
-      → SUP A (UnionCF A f mf ay supf x) 
-M→S {A} {f} {mf} {y} {ay} {x} supf ms = record { sup = * (MinSUP.sup ms) 
+      →  MinSUP A (UnionCF A f mf ay supf x)
+      → SUP A (UnionCF A f mf ay supf x)
+M→S {A} {f} {mf} {y} {ay} {x} supf ms = record { sup = * (MinSUP.sup ms)
         ; as = subst (λ k → odef A k) (sym &iso) (MinSUP.asm ms) ; x≤sup = ms00 } where
    msup = MinSUP.sup ms
    ms00 : {z : HOD} → UnionCF A f mf ay supf x ∋ z → (z ≡ * msup) ∨ (z < * msup)
-   ms00 {z} uz with MinSUP.x≤sup ms uz 
+   ms00 {z} uz with MinSUP.x≤sup ms uz
    ... | case1 eq = case1 (subst (λ k → k ≡ _) *iso ( cong (*) eq))
    ... | case2 lt = case2 (subst₂ (λ j k →  j < k ) *iso refl lt )
 
 
-chain-mono : {A : HOD}  ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f )  {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal ) 
+chain-mono : {A : HOD}  ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f )  {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal )
    (supf-mono : {x y : Ordinal } →  x o≤  y  → supf x o≤ supf y ) {a b c : Ordinal} → a o≤ b
         → odef (UnionCF A f mf ay supf a) c → odef (UnionCF A f mf ay supf b) c
+chain-mono f mf ay supf supf-mono {a} {b} {c} a≤b ⟪ ua , ch-init fc  ⟫ =
+        ⟪ ua , ch-init fc  ⟫
 chain-mono f mf ay supf supf-mono {a} {b} {c} a≤b ⟪ uaa ,  ch-is-sup ua ua<x is-sup fc  ⟫ =
-        ⟪ uaa , ch-is-sup ua (ordtrans<-≤ ua<x (supf-mono a≤b ) ) is-sup fc  ⟫ 
+        ⟪ uaa , ch-is-sup ua (ordtrans<-≤ ua<x (supf-mono a≤b ) ) is-sup fc  ⟫
 
-record ZChain ( A : HOD )    ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f) 
+record ZChain ( A : HOD )    ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f)
         {y : Ordinal} (ay : odef A y)  ( z : Ordinal ) : Set (Level.suc n) where
    field
-      supf :  Ordinal → Ordinal 
-      sup=u : {b : Ordinal} → (ab : odef A b) → b o≤ z  
-           → IsSUP A (UnionCF A f mf ay supf b) ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf b) f b ) → supf b ≡ b 
+      supf :  Ordinal → Ordinal
+      sup=u : {b : Ordinal} → (ab : odef A b) → b o≤ z
+           → IsSUP A (UnionCF A f mf ay supf b) ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf b) f b ) → supf b ≡ b
 
       asupf :  {x : Ordinal } → odef A (supf x)
       supf-mono : {x y : Ordinal } → x o≤  y → supf x o≤ supf y
       supf-< : {x y : Ordinal } → supf x o< supf  y → supf x << supf y
       supfmax : {x : Ordinal } → z o< x → supf x ≡ supf z
-      chain∋init : odef (UnionCF A f mf ay supf z) y
 
-      minsup : {x : Ordinal } → x o≤ z  → MinSUP A (UnionCF A f mf ay supf x) 
+      minsup : {x : Ordinal } → x o≤ z  → MinSUP A (UnionCF A f mf ay supf x)
       supf-is-minsup : {x : Ordinal } → (x≤z : x o≤ z) → supf x ≡ MinSUP.sup ( minsup x≤z )
       csupf : {b : Ordinal } → supf b o< z → odef (UnionCF A f mf ay supf z) (supf b) -- supf z is not an element of this chain
 
@@ -392,23 +425,27 @@
    chain⊆A : chain ⊆' A
    chain⊆A = λ lt → proj1 lt
 
-   sup : {x : Ordinal } → x o≤ z  → SUP A (UnionCF A f mf ay supf x) 
-   sup {x} x≤z = M→S supf (minsup x≤z) 
+   sup : {x : Ordinal } → x o≤ z  → SUP A (UnionCF A f mf ay supf x)
+   sup {x} x≤z = M→S supf (minsup x≤z)
 
    s=ms : {x : Ordinal } → (x≤z : x o≤ z ) → & (SUP.sup (sup x≤z)) ≡ MinSUP.sup (minsup x≤z)
    s=ms {x} x≤z = &iso
 
+   chain∋init : odef chain y
+   chain∋init = ⟪ ay , ch-init (init ay refl)    ⟫
    f-next : {a z : Ordinal} → odef (UnionCF A f mf ay supf z) a → odef (UnionCF A f mf ay supf z) (f a)
+   f-next {a} ⟪ aa , ch-init fc ⟫ = ⟪ proj2 (mf a aa) , ch-init (fsuc _ fc)  ⟫
    f-next {a} ⟪ aa , ch-is-sup u u<x is-sup fc ⟫ = ⟪ proj2 (mf a aa) , ch-is-sup u u<x is-sup (fsuc _ fc ) ⟫
    initial : {z : Ordinal } → odef chain z → * y ≤ * z
    initial {a} ⟪ aa , ua ⟫  with  ua
+   ... | ch-init fc = s≤fc y f mf fc
    ... | ch-is-sup u u<x is-sup fc = ≤-ftrans (<=to≤ zc7) (s≤fc _ f mf fc)  where
-        zc7 : y <= supf u 
+        zc7 : y <= supf u
         zc7 = ChainP.fcy<sup is-sup (init ay refl)
    f-total : IsTotalOrderSet chain
-   f-total {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where 
+   f-total {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
                uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
-               uz01 = chain-total A f mf ay supf ( (proj2 ca)) ( (proj2 cb)) 
+               uz01 = chain-total A f mf ay supf ( (proj2 ca)) ( (proj2 cb))
 
    supf-<= : {x y : Ordinal } → supf x <= supf  y → supf x o≤ supf y
    supf-<= {x} {y} (case1 sx=sy) = o≤-refl0 sx=sy
@@ -417,7 +454,7 @@
    ... | tri≈ ¬a b ¬c = o≤-refl0 b
    ... | tri> ¬a ¬b c = ⊥-elim (<-irr (case2 sx<sy ) (supf-< c) )
 
-   supf-inject : {x y : Ordinal } → supf x o< supf y → x o<  y 
+   supf-inject : {x y : Ordinal } → supf x o< supf y → x o<  y
    supf-inject {x} {y} sx<sy with trio< x y
    ... | tri< a ¬b ¬c = a
    ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy )
@@ -425,36 +462,68 @@
    ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy )
    ... | case2 lt = ⊥-elim ( o<> sx<sy lt )
 
-   fcy<sup  : {u w : Ordinal } → u o≤ z  → FClosure A f y w → (w ≡ supf u ) ∨ ( w << supf u ) -- different from order because y o< supf 
-   fcy<sup  {u} {w} u≤z fc with chain∋init
-   ... | ⟪ ay1 , ch-is-sup uy uy<x is-sup fcy ⟫ = <=-trans (ChainP.fcy<sup is-sup fc) ?
+   fcy<sup  : {u w : Ordinal } → u o≤ z  → FClosure A f y w → (w ≡ supf u ) ∨ ( w << supf u ) -- different from order because y o< supf
+   fcy<sup  {u} {w} u≤z fc with MinSUP.x≤sup (minsup u≤z) ⟪ subst (λ k → odef A k ) (sym &iso) (A∋fc {A} y f mf fc)
+       , ch-init (subst (λ k → FClosure A f y k) (sym &iso) fc ) ⟫
+   ... | case1 eq = case1 (subst (λ k → k ≡ supf u ) &iso  (trans eq (sym (supf-is-minsup u≤z ) ) ))
+   ... | case2 lt = case2 (subst₂ (λ j k → j << k ) &iso (sym (supf-is-minsup u≤z )) lt )
 
-   -- with MinSUP.x≤sup (minsup u≤z) ⟪ subst (λ k → odef A k ) (sym &iso) (A∋fc {A} y f mf fc)  
-   --     , ? ⟫ --ch-init (subst (λ k → FClosure A f y k) (sym &iso) fc ) ⟫ 
-   -- ... | case1 eq = case1 (subst (λ k → k ≡ supf u ) &iso  (trans eq (sym (supf-is-minsup u≤z ) ) ))
-   -- ... | case2 lt = case2 (subst₂ (λ j k → j << k ) &iso (sym (supf-is-minsup u≤z )) lt )
+   -- ordering is not proved here but in ZChain1
 
-   -- ordering is not proved here but in ZChain1 
-
-   IsMinSUP→NotHasPrev : {x sp : Ordinal } → odef A sp 
+   IsMinSUP→NotHasPrev : {x sp : Ordinal } → odef A sp
        → ({y : Ordinal} → odef (UnionCF A f mf ay supf x) y → (y ≡ sp ) ∨ (y << sp ))
        → ( {a : Ordinal } → a << f a )
        → ¬ ( HasPrev A (UnionCF A f mf ay supf x) f sp )
    IsMinSUP→NotHasPrev {x} {sp} asp is-sup <-mono-f hp = ⊥-elim (<-irr ( <=to≤  fsp≤sp) sp<fsp ) where
        sp<fsp : sp << f sp
-       sp<fsp = <-mono-f 
-       pr = HasPrev.y hp 
+       sp<fsp = <-mono-f
+       pr = HasPrev.y hp
        im00 : f (f pr) <= sp
        im00 = is-sup ( f-next (f-next (HasPrev.ay hp)))
        fsp≤sp : f sp <=  sp
        fsp≤sp = subst (λ k → f k <= sp ) (sym (HasPrev.x=fy hp)) im00
 
-record ZChain1 ( A : HOD )    ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f) 
+   UChain⊆ : { supf1 : Ordinal → Ordinal }
+        → ( { x : Ordinal } → x o< z → supf x ≡ supf1 x)
+        → ( { x : Ordinal } → z o≤ x → supf z o≤ supf1 x)
+        → UnionCF A f mf ay supf z ⊆' UnionCF A f mf ay supf1 z
+   UChain⊆ {supf1} eq<x lex ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫ 
+   UChain⊆ {supf1} eq<x lex ⟪ az , ch-is-sup u {x} u<x is-sup fc ⟫ = ⟪ az , ch-is-sup u u<x1 cp1 fc1 ⟫  where
+          u<x0 : u o< z
+          u<x0 = supf-inject u<x
+          u<x1 : supf1 u o< supf1 z
+          u<x1 = subst (λ k → k o< supf1 z ) (eq<x u<x0) (ordtrans<-≤ u<x (lex o≤-refl ) )
+          fc1 : FClosure A f (supf1 u) x
+          fc1 = subst (λ k → FClosure A f k x ) (eq<x u<x0) fc
+          uc01 : {s : Ordinal } →  supf1 s o< supf1 u → s o< z
+          uc01 {s} s<u with trio< s z
+          ... | tri< a ¬b ¬c = a
+          ... | tri≈ ¬a b ¬c = ⊥-elim ( o≤> uc02  s<u ) where -- (supf-mono (o<→≤ u<x0))
+               uc02 :  supf1 u o≤ supf1 s
+               uc02 = begin
+                 supf1 u  <⟨ u<x1 ⟩
+                 supf1 z  ≡⟨ cong supf1 (sym b) ⟩
+                 supf1 s ∎  where open o≤-Reasoning O
+          ... | tri> ¬a ¬b c = ⊥-elim ( o≤> uc03 s<u ) where
+               uc03 :  supf1 u o≤ supf1 s
+               uc03 = begin
+                 supf1 u  ≡⟨ sym (eq<x u<x0) ⟩
+                 supf u  <⟨ u<x ⟩
+                 supf z  ≤⟨ lex (o<→≤ c) ⟩
+                 supf1 s ∎  where open o≤-Reasoning O
+          cp1 : ChainP A f mf ay supf1 u
+          cp1 = record { fcy<sup = λ {z} fc → subst (λ k → (z ≡ k) ∨ ( z << k ) ) (eq<x u<x0) (ChainP.fcy<sup is-sup fc )  
+               ; order =  λ {s} {z} s<u fc  → subst (λ k → (z ≡ k) ∨ ( z << k ) ) (eq<x u<x0) 
+                  (ChainP.order is-sup (subst₂ (λ j k → j o< k ) (sym (eq<x (uc01 s<u) )) (sym (eq<x u<x0)) s<u)  
+                    (subst (λ k → FClosure A f k z ) (sym (eq<x (uc01 s<u) )) fc ))
+               ; supu=u = trans (sym (eq<x u<x0)) (ChainP.supu=u is-sup)  }
+
+record ZChain1 ( A : HOD )    ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f)
         {y : Ordinal} (ay : odef A y)  (zc : ZChain A f mf ay (& A)) ( z : Ordinal ) : Set (Level.suc n) where
    supf = ZChain.supf zc
    field
-      is-max :  {a b : Ordinal } → (ca : odef (UnionCF A f mf ay supf z) a ) → supf b o< supf z  → (ab : odef A b) 
-          → HasPrev A (UnionCF A f mf ay supf z) f b ∨  IsSUP A (UnionCF A f mf ay supf b) ab  
+      is-max :  {a b : Ordinal } → (ca : odef (UnionCF A f mf ay supf z) a ) → supf b o< supf z  → (ab : odef A b)
+          → HasPrev A (UnionCF A f mf ay supf z) f b ∨  IsSUP A (UnionCF A f mf ay supf b) ab
           → * a < * b  → odef ((UnionCF A f mf ay supf z)) b
       order : {b s z1 : Ordinal} → b o< & A → supf s o< supf b → FClosure A f (supf s) z1 → (z1 ≡ supf b) ∨ (z1 << supf b)
 
@@ -464,17 +533,21 @@
       as : A ∋ maximal
       ¬maximal<x : {x : HOD} → A ∋ x  → ¬ maximal < x       -- A is Partial, use negative
 
-Zorn-lemma : { A : HOD } 
-    → o∅ o< & A 
+init-uchain : (A : HOD)  ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f) {y : Ordinal } → (ay : odef A y )
+    { supf : Ordinal → Ordinal } { x : Ordinal } → odef (UnionCF A f mf ay supf x) y
+init-uchain A f mf ay = ⟪ ay , ch-init (init ay refl)   ⟫
+
+Zorn-lemma : { A : HOD }
+    → o∅ o< & A
     → ( ( B : HOD) → (B⊆A : B ⊆' A) → IsTotalOrderSet B → SUP A B   ) -- SUP condition
-    → Maximal A 
+    → Maximal A
 Zorn-lemma {A}  0<A supP = zorn00 where
      <-irr0 : {a b : HOD} → A ∋ a → A ∋ b  → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
      <-irr0 {a} {b} A∋a A∋b = <-irr
      z07 :   {y : Ordinal} {A : HOD } → {P : Set n} → odef A y ∧ P → y o< & A
      z07 {y} {A} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))
      s : HOD
-     s = ODC.minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A )) 
+     s = ODC.minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A ))
      as : A ∋ * ( & s  )
      as =  subst (λ k → odef A (& k) ) (sym *iso) ( ODC.x∋minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A ))  )
      as0 : odef A  (& s  )
@@ -482,20 +555,20 @@
      s<A : & s o< & A
      s<A = c<→o< (subst (λ k → odef A (& k) ) *iso as )
      HasMaximal : HOD
-     HasMaximal = record { od = record { def = λ x → odef A x ∧ ( (m : Ordinal) →  odef A m → ¬ (* x < * m)) }  ; odmax = & A ; <odmax = z07 } 
+     HasMaximal = record { od = record { def = λ x → odef A x ∧ ( (m : Ordinal) →  odef A m → ¬ (* x < * m)) }  ; odmax = & A ; <odmax = z07 }
      no-maximum : HasMaximal =h= od∅ → (x : Ordinal) → odef A x ∧ ((m : Ordinal) →  odef A m →  odef A x ∧ (¬ (* x < * m) )) →  ⊥
-     no-maximum nomx x P = ¬x<0 (eq→ nomx {x} ⟪ proj1 P , (λ m ma p → proj2 ( proj2 P m ma ) p ) ⟫ )  
+     no-maximum nomx x P = ¬x<0 (eq→ nomx {x} ⟪ proj1 P , (λ m ma p → proj2 ( proj2 P m ma ) p ) ⟫ )
      Gtx : { x : HOD} → A ∋ x → HOD
-     Gtx {x} ax = record { od = record { def = λ y → odef A y ∧ (x < (* y)) } ; odmax = & A ; <odmax = z07 } 
+     Gtx {x} ax = record { od = record { def = λ y → odef A y ∧ (x < (* y)) } ; odmax = & A ; <odmax = z07 }
      z08  : ¬ Maximal A →  HasMaximal =h= od∅
      z08 nmx  = record { eq→  = λ {x} lt → ⊥-elim ( nmx record {maximal = * x ; as = subst (λ k → odef A k) (sym &iso) (proj1 lt)
          ; ¬maximal<x = λ {y} ay → subst (λ k → ¬ (* x < k)) *iso (proj2 lt (& y) ay) } ) ; eq← =  λ {y} lt → ⊥-elim ( ¬x<0 lt )}
      x-is-maximal : ¬ Maximal A → {x : Ordinal} → (ax : odef A x) → & (Gtx (subst (λ k → odef A k ) (sym &iso) ax)) ≡ o∅ → (m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m))
      x-is-maximal nmx {x} ax nogt m am  = ⟪ subst (λ k → odef A k) &iso (subst (λ k → odef A k ) (sym &iso) ax) ,  ¬x<m  ⟫ where
         ¬x<m :  ¬ (* x < * m)
-        ¬x<m x<m = ∅< {Gtx (subst (λ k → odef A k ) (sym &iso) ax)} {* m} ⟪ subst (λ k → odef A k) (sym &iso) am , subst (λ k → * x < k ) (cong (*) (sym &iso)) x<m ⟫  (≡o∅→=od∅ nogt) 
+        ¬x<m x<m = ∅< {Gtx (subst (λ k → odef A k ) (sym &iso) ax)} {* m} ⟪ subst (λ k → odef A k) (sym &iso) am , subst (λ k → * x < k ) (cong (*) (sym &iso)) x<m ⟫  (≡o∅→=od∅ nogt)
 
-     minsupP :  ( B : HOD) → (B⊆A : B ⊆' A) → IsTotalOrderSet B → MinSUP A B  
+     minsupP :  ( B : HOD) → (B⊆A : B ⊆' A) → IsTotalOrderSet B → MinSUP A B
      minsupP B B⊆A total = m02 where
          xsup : (sup : Ordinal ) → Set n
          xsup sup = {w : Ordinal } → odef B w → (w ≡ sup ) ∨ (w << sup )
@@ -536,7 +609,7 @@
              m05 {w} bw with SUP.x≤sup S {* w} (subst (λ k → odef B k) (sym &iso) bw )
              ... | case1 eq = case1 ( subst₂ (λ j k → j ≡ k ) &iso refl (cong (&) eq) )
              ... | case2 lt = case2 ( subst (λ k → _ < k ) (sym *iso) lt )
-         m02 : MinSUP A B 
+         m02 : MinSUP A B
          m02 = dont-or (m00 (& A)) m03
 
      -- Uncountable ascending chain by axiom of choice
@@ -568,18 +641,19 @@
      --     supf is fixed for z ≡ & A , we can prove order and is-max
      --
 
-     SZ1 : ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f) 
+     SZ1 : ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f)
         {init : Ordinal} (ay : odef A init) (zc : ZChain A f mf ay (& A)) (x : Ordinal)   → ZChain1 A f mf ay zc x
      SZ1 f mf {y} ay zc x = zc1 x where
         chain-mono1 :  {a b c : Ordinal} → a o≤ b
             → odef (UnionCF A f mf ay (ZChain.supf zc) a) c → odef (UnionCF A f mf ay (ZChain.supf zc) b) c
         chain-mono1  {a} {b} {c} a≤b = chain-mono f mf ay (ZChain.supf zc) (ZChain.supf-mono zc) a≤b
-        is-max-hp : (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a → (ab : odef A b) 
-            → HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) f b  
+        is-max-hp : (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a → (ab : odef A b)
+            → HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) f b
             → * a < * b → odef (UnionCF A f mf ay (ZChain.supf zc) x) b
         is-max-hp x {a} {b} ua ab has-prev a<b with HasPrev.ay has-prev
+        ... | ⟪ ab0 , ch-init fc ⟫ = ⟪ ab , ch-init ( subst (λ k → FClosure A f y k) (sym (HasPrev.x=fy has-prev)) (fsuc _ fc )) ⟫
         ... | ⟪ ab0 , ch-is-sup u u<x is-sup fc ⟫ = ⟪ ab , subst (λ k → UChain A f mf ay (ZChain.supf zc) x k )
-                      (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x is-sup (fsuc _ fc))  ⟫ 
+                      (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x is-sup (fsuc _ fc))  ⟫
 
         supf = ZChain.supf zc
 
@@ -593,17 +667,19 @@
                 zc04 = ZChain.csupf zc (z09 (ZChain.asupf zc))
                 zc05 : odef (UnionCF A f mf ay supf b)  (supf s)
                 zc05 with zc04
+                ... | ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc  ⟫
                 ... | ⟪ as , ch-is-sup u u<x is-sup fc ⟫ = ⟪ as , ch-is-sup u zc08 is-sup fc ⟫ where
                     zc07 : FClosure A f (supf u) (supf s)   -- supf u ≤ supf s  → supf u o≤ supf s
                     zc07 = fc
                     zc06 : supf u ≡ u
                     zc06 = ChainP.supu=u is-sup
-                    zc08 : supf u o< supf b  
+                    zc08 : supf u o< supf b
                     zc08 = ordtrans≤-< (ZChain.supf-<= zc (≤to<= ( s≤fc _ f mf fc ))) ss<sb
         csupf-fc {b} {s} {z1} b<z ss≤sb (fsuc x fc) = subst (λ k → odef (UnionCF A f mf ay supf b) k ) (sym &iso) zc04  where
                 zc04 : odef (UnionCF A f mf ay supf b) (f x)
                 zc04 with subst (λ k → odef (UnionCF A f mf ay supf b) k ) &iso (csupf-fc b<z ss≤sb fc  )
-                ... | ⟪ as , ch-is-sup u u<x is-sup fc ⟫ = ⟪ proj2 (mf _ as) , ch-is-sup u u<x is-sup (fsuc _ fc) ⟫ 
+                ... | ⟪ as , ch-init fc ⟫ = ⟪ proj2 (mf _ as) , ch-init (fsuc _ fc) ⟫
+                ... | ⟪ as , ch-is-sup u u<x is-sup fc ⟫ = ⟪ proj2 (mf _ as) , ch-is-sup u u<x is-sup (fsuc _ fc) ⟫
         order : {b s z1 : Ordinal} → b o< & A → supf s o< supf b → FClosure A f (supf s) z1 → (z1 ≡ supf b) ∨ (z1 << supf b)
         order {b} {s} {z1} b<z ss<sb fc = zc04 where
           zc00 : ( z1  ≡  MinSUP.sup (ZChain.minsup zc (o<→≤ b<z) )) ∨ ( z1  << MinSUP.sup ( ZChain.minsup zc (o<→≤ b<z) ) )
@@ -625,24 +701,24 @@
               HasPrev A (UnionCF A f mf ay supf x) f b  ∨ IsSUP A (UnionCF A f mf ay supf b) ab →
               * a < * b → odef (UnionCF A f mf ay supf x) b
            is-max {a} {b} ua b<x ab P a<b with ODC.or-exclude O P
-           is-max {a} {b} ua b<x ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b 
-           is-max {a} {b} ua sb<sx ab P a<b | case2 is-sup 
+           is-max {a} {b} ua b<x ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b
+           is-max {a} {b} ua sb<sx ab P a<b | case2 is-sup
              = ⟪ ab , ch-is-sup b sb<sx m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl)) ⟫   where
               b<A : b o< & A
               b<A = z09 ab
               b<x : b o< x
               b<x  = ZChain.supf-inject zc sb<sx
-              m04 : ¬ HasPrev A (UnionCF A f mf ay supf b) f b 
-              m04 nhp = proj1 is-sup ( record { ax = HasPrev.ax nhp ; y = HasPrev.y nhp ; ay = 
+              m04 : ¬ HasPrev A (UnionCF A f mf ay supf b) f b
+              m04 nhp = proj1 is-sup ( record { ax = HasPrev.ax nhp ; y = HasPrev.y nhp ; ay =
                     chain-mono1 (o<→≤ b<x) (HasPrev.ay  nhp) ; x=fy = HasPrev.x=fy nhp } )
               m05 : ZChain.supf zc b ≡ b
-              m05 =  ZChain.sup=u zc ab (o<→≤ (z09 ab) ) ⟪ record { x≤sup = λ {z} uz → IsSUP.x≤sup (proj2 is-sup) uz  }  , m04 ⟫ 
+              m05 =  ZChain.sup=u zc ab (o<→≤ (z09 ab) ) ⟪ record { x≤sup = λ {z} uz → IsSUP.x≤sup (proj2 is-sup) uz  }  , m04 ⟫
               m08 : {z : Ordinal} → (fcz : FClosure A f y z ) → z <= ZChain.supf zc b
               m08 {z} fcz = ZChain.fcy<sup zc (o<→≤ b<A) fcz
-              m09 : {s z1 : Ordinal} → ZChain.supf zc s o< ZChain.supf zc b 
+              m09 : {s z1 : Ordinal} → ZChain.supf zc s o< ZChain.supf zc b
                            → FClosure A f (ZChain.supf zc s) z1 → z1 <= ZChain.supf zc b
-              m09 {s} {z} s<b fcz = order b<A s<b fcz    
-              m06 : ChainP A f mf ay supf b 
+              m09 {s} {z} s<b fcz = order b<A s<b fcz
+              m06 : ChainP A f mf ay supf b
               m06 = record {  fcy<sup = m08  ; order = m09 ; supu=u = m05 }
         ... | no lim = record { is-max = is-max ; order = order }  where
            is-max :  {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay supf x) a →
@@ -650,46 +726,42 @@
               HasPrev A (UnionCF A f mf ay supf x) f b  ∨ IsSUP A (UnionCF A f mf ay supf b) ab →
               * a < * b → odef (UnionCF A f mf ay supf x) b
            is-max {a} {b} ua sb<sx ab P a<b with ODC.or-exclude O P
-           is-max {a} {b} ua sb<sx ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b 
-           is-max {a} {b} ua sb<sx ab P a<b | case2 is-sup 
-             = ⟪ ab , ch-is-sup b sb<sx m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl))  ⟫ where
+           is-max {a} {b} ua sb<sx ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b
+           is-max {a} {b} ua sb<sx ab P a<b | case2 is-sup with IsSUP.x≤sup (proj2 is-sup) (init-uchain A f mf ay )
+           ... | case1 b=y = ⟪ subst (λ k → odef A k ) b=y ay , ch-init (subst (λ k → FClosure A f y k ) b=y (init ay refl ))  ⟫
+           ... | case2 y<b = ⟪ ab , ch-is-sup b sb<sx m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl))  ⟫ where
               m09 : b o< & A
               m09 = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab))
-              b<x : b o< x 
+              b<x : b o< x
               b<x  = ZChain.supf-inject zc sb<sx
               m07 : {z : Ordinal} → FClosure A f y z → z <= ZChain.supf zc b
               m07 {z} fc = ZChain.fcy<sup zc (o<→≤ m09) fc
-              m08 : {s z1 : Ordinal} → ZChain.supf zc s o< ZChain.supf zc b 
+              m08 : {s z1 : Ordinal} → ZChain.supf zc s o< ZChain.supf zc b
                        → FClosure A f (ZChain.supf zc s) z1 → z1 <= ZChain.supf zc b
               m08 {s} {z1} s<b fc = order m09 s<b fc
               m04 : ¬ HasPrev A (UnionCF A f mf ay supf b) f b
-              m04 nhp = proj1 is-sup ( record { ax = HasPrev.ax nhp ; y = HasPrev.y nhp ; ay =  
-                      chain-mono1 (o<→≤ b<x) (HasPrev.ay  nhp) 
+              m04 nhp = proj1 is-sup ( record { ax = HasPrev.ax nhp ; y = HasPrev.y nhp ; ay =
+                      chain-mono1 (o<→≤ b<x) (HasPrev.ay  nhp)
                  ; x=fy = HasPrev.x=fy nhp } )
               m05 : ZChain.supf zc b ≡ b
               m05 = ZChain.sup=u zc ab (o<→≤  m09) ⟪ record { x≤sup = λ lt → IsSUP.x≤sup (proj2 is-sup) lt } , m04 ⟫    -- ZChain on x
-              m06 : ChainP A f mf ay supf b 
+              m06 : ChainP A f mf ay supf b
               m06 = record { fcy<sup = m07 ;  order = m08 ; supu=u = m05 }
 
      uchain : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → HOD
-     uchain f mf {y} ay = record { od = record { def = λ x → FClosure A f y x } ; odmax = & A ; <odmax = 
+     uchain f mf {y} ay = record { od = record { def = λ x → FClosure A f y x } ; odmax = & A ; <odmax =
              λ {z} cz → subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (A∋fc y f mf cz ))) }
 
-     utotal : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) 
+     utotal : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y)
         → IsTotalOrderSet (uchain f mf ay)
-     utotal f mf {y} ay {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where 
+     utotal f mf {y} ay {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
                uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
                uz01 = fcn-cmp y f mf ca cb
 
-     ysup : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) 
+     ysup : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y)
        →  MinSUP A (uchain f mf ay)
-     ysup f mf {y} ay = minsupP (uchain f mf ay) (λ lt → A∋fc y f mf lt)  (utotal f mf ay) 
+     ysup f mf {y} ay = minsupP (uchain f mf ay) (λ lt → A∋fc y f mf lt)  (utotal f mf ay)
 
-     UChainPreserve : {x z : Ordinal } { f : Ordinal → Ordinal } → {mf : ≤-monotonic-f A f } {y : Ordinal} {ay : odef A y}
-        → { supf0 supf1 : Ordinal → Ordinal }
-        → ( ( z : Ordinal ) → z o< x → supf0 z ≡ supf1 z)
-        → UnionCF A f mf ay supf0 x ≡ UnionCF A f mf ay supf1 x 
-     UChainPreserve {x} {f} {mf} {y} {ay} {supf0} {supf1} <x→eq = ?
 
      SUP⊆ : { B C : HOD } → B ⊆' C → SUP A C → SUP A B
      SUP⊆ {B} {C} B⊆C sup = record { sup = SUP.sup sup ; as = SUP.as sup ; x≤sup = λ lt → SUP.x≤sup sup (B⊆C lt)    }
@@ -703,7 +775,7 @@
      -- create all ZChains under o< x
      --
 
-     ind : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → (x : Ordinal) 
+     ind : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → (x : Ordinal)
          → ((z : Ordinal) → z o< x → ZChain A f mf ay z) → ZChain A f mf ay x
      ind f mf {y} ay x prev with Oprev-p x
      ... | yes op = zc41 where
@@ -712,39 +784,39 @@
           --
           px = Oprev.oprev op
           zc : ZChain A f mf ay (Oprev.oprev op)
-          zc = prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc ) 
+          zc = prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc )
           px<x : px o< x
           px<x = subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc
           opx=x : osuc px ≡ x
           opx=x = Oprev.oprev=x op
 
           zc-b<x : (b : Ordinal ) → b o< x → b o< osuc px
-          zc-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt 
+          zc-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt
 
           supf0 = ZChain.supf zc
           pchain  : HOD
           pchain   = UnionCF A f mf ay supf0 px
 
-          supf-mono : {a b : Ordinal } → a o≤ b → supf0 a o≤ supf0 b 
+          supf-mono : {a b : Ordinal } → a o≤ b → supf0 a o≤ supf0 b
           supf-mono = ZChain.supf-mono zc
 
           zc04 : {b : Ordinal} → b o≤ x → (b o≤ px ) ∨ (b ≡ x )
-          zc04 {b} b≤x with trio< b px 
+          zc04 {b} b≤x with trio< b px
           ... | tri< a ¬b ¬c = case1 (o<→≤ a)
           ... | tri≈ ¬a b ¬c = case1 (o≤-refl0 b)
           ... | tri> ¬a ¬b px<b with osuc-≡< b≤x
           ... | case1 eq = case2 eq
-          ... | case2 b<x = ⊥-elim ( ¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x  ⟫ ) 
+          ... | case2 b<x = ⊥-elim ( ¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x  ⟫ )
 
           --
           -- find the next value of supf
           --
 
           pchainpx : HOD
-          pchainpx = record { od = record { def = λ z →  (odef A z ∧ UChain A f mf ay supf0 px z ) 
+          pchainpx = record { od = record { def = λ z →  (odef A z ∧ UChain A f mf ay supf0 px z )
                 ∨ FClosure A f (supf0 px) z  } ; odmax = & A ; <odmax = zc00 } where
                zc00 : {z : Ordinal } → (odef A z ∧ UChain A f mf ay supf0 px z ) ∨ FClosure A f (supf0 px) z → z o< & A
-               zc00 {z} (case1 lt) = z07 lt 
+               zc00 {z} (case1 lt) = z07 lt
                zc00 {z} (case2 fc) = z09 ( A∋fc (supf0 px) f mf fc )
 
           zc02 : { a b : Ordinal } → odef A a ∧ UChain A f mf ay supf0 px a → FClosure A f (supf0 px) b → a <= b
@@ -754,15 +826,15 @@
              zc05 : {b : Ordinal } → FClosure A f (supf0 px) b → a <= b
              zc05 (fsuc b1 fb ) with proj1 ( mf b1 (A∋fc (supf0 px) f mf fb ))
              ... | case1 eq = subst (λ k → a <= k ) (subst₂ (λ j k → j ≡ k) &iso &iso (cong (&) eq)) (zc05 fb)
-             ... | case2 lt = <=-trans (zc05 fb) (case2 lt) 
-             zc05 (init b1 refl) with MinSUP.x≤sup (ZChain.minsup zc o≤-refl) 
+             ... | case2 lt = <=-trans (zc05 fb) (case2 lt)
+             zc05 (init b1 refl) with MinSUP.x≤sup (ZChain.minsup zc o≤-refl)
                 (subst (λ k → odef A k ∧ UChain A f mf ay supf0 px k) (sym &iso) ca )
              ... | case1 eq = case1 (subst₂ (λ j k → j ≡ k ) &iso zc06 eq )
-             ... | case2 lt = case2 (subst₂ (λ j k → j < k ) *iso (cong (*) zc06) lt ) 
- 
+             ... | case2 lt = case2 (subst₂ (λ j k → j < k ) *iso (cong (*) zc06) lt )
+
           ptotal : IsTotalOrderSet pchainpx
-          ptotal (case1 a) (case1 b) =  subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso 
-                   (chain-total A f mf ay supf0 (proj2 a) (proj2 b)) 
+          ptotal (case1 a) (case1 b) =  subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso
+                   (chain-total A f mf ay supf0 (proj2 a) (proj2 b))
           ptotal {a0} {b0} (case1 a) (case2 b) with zc02 a b
           ... | case1 eq = tri≈ (<-irr (case1 (sym eq1))) eq1 (<-irr (case1 eq1)) where
                eq1 : a0 ≡ b0
@@ -778,31 +850,31 @@
                lt1 : a0 < b0
                lt1 = subst₂ (λ j k → j < k ) *iso *iso lt
           ptotal (case2 a) (case2 b) = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (fcn-cmp (supf0 px) f mf a b)
- 
+
           pcha : pchainpx ⊆' A
           pcha (case1 lt) = proj1 lt
           pcha (case2 fc) = A∋fc _ f mf fc
- 
-          sup1 : MinSUP A pchainpx 
+
+          sup1 : MinSUP A pchainpx
           sup1 = minsupP pchainpx pcha ptotal
           sp1 = MinSUP.sup sup1
 
           --
           --     supf0 px o≤ sp1
-          --  
- 
-          zc41 : ZChain A f mf ay x 
+          --
+
+          zc41 : ZChain A f mf ay x
           zc41 with MinSUP.x≤sup sup1 (case2 (init (ZChain.asupf zc {px}) refl ))
           zc41 | (case2 sfpx<sp1) =  record { supf = supf1 ; sup=u = ? ; asupf = ? ; supf-mono = supf1-mono ; supf-< = ?
               ; supfmax = ? ; minsup = ? ; supf-is-minsup = ? ; csupf = csupf1    }  where
                  --  supf0 px is included by the chain of sp1
-                 --  ( UnionCF A f mf ay supf0 px ∪ FClosure (supf0 px) ) ≡ UnionCF supf1 x 
+                 --  ( UnionCF A f mf ay supf0 px ∪ FClosure (supf0 px) ) ≡ UnionCF supf1 x
                  --  supf1 x ≡ sp1, which is not included now
 
                  supf1 : Ordinal → Ordinal
-                 supf1 z with trio< z px 
+                 supf1 z with trio< z px
                  ... | tri< a ¬b ¬c = supf0 z
-                 ... | tri≈ ¬a b ¬c = supf0 z 
+                 ... | tri≈ ¬a b ¬c = supf0 z
                  ... | tri> ¬a ¬b c = sp1
 
                  sf1=sf0 : {z : Ordinal } → z o≤ px → supf1 z ≡ supf0 z
@@ -819,12 +891,12 @@
 
                  asupf1 : {z : Ordinal } → odef A (supf1 z)
                  asupf1 {z} with trio< z px
-                 ... | tri< a ¬b ¬c = ZChain.asupf zc 
-                 ... | tri≈ ¬a b ¬c = ZChain.asupf zc 
+                 ... | tri< a ¬b ¬c = ZChain.asupf zc
+                 ... | tri≈ ¬a b ¬c = ZChain.asupf zc
                  ... | tri> ¬a ¬b c = MinSUP.asm sup1
 
-                 supf1-mono : {a b : Ordinal } → a o≤ b → supf1 a o≤ supf1 b 
-                 supf1-mono {a} {b} a≤b with trio< b px 
+                 supf1-mono : {a b : Ordinal } → a o≤ b → supf1 a o≤ supf1 b
+                 supf1-mono {a} {b} a≤b with trio< b px
                  ... | tri< a ¬b ¬c =  subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (o<→≤ (ordtrans≤-< a≤b a)))) refl ( supf-mono a≤b )
                  ... | tri≈ ¬a b ¬c =  subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (subst (λ k → a o≤ k) b a≤b))) refl ( supf-mono a≤b )
                  supf1-mono {a} {b} a≤b | tri> ¬a ¬b c with trio< a px
@@ -833,25 +905,26 @@
                        zc21 = ZChain.minsup zc (o<→≤ a<px)
                        zc24 : {x₁ : Ordinal} → odef (UnionCF A f mf ay supf0 a) x₁ → (x₁ ≡ sp1) ∨ (x₁ << sp1)
                        zc24 {x₁} ux = MinSUP.x≤sup sup1 (case1 (chain-mono f mf ay supf0 (ZChain.supf-mono zc) (o<→≤ a<px) ux ) )
-                       zc19 : supf0 a o≤ sp1 
+                       zc19 : supf0 a o≤ sp1
                        zc19 = subst (λ k → k o≤ sp1) (sym (ZChain.supf-is-minsup zc  (o<→≤ a<px))) ( MinSUP.minsup zc21 (MinSUP.asm sup1) zc24 )
                  ... | tri≈ ¬a b ¬c = zc18 where
                        zc21 : MinSUP A (UnionCF A f mf ay supf0 a)
                        zc21 = ZChain.minsup zc (o≤-refl0 b)
                        zc20 : MinSUP.sup zc21 ≡ supf0 a
-                       zc20 = sym (ZChain.supf-is-minsup zc (o≤-refl0 b)) 
+                       zc20 = sym (ZChain.supf-is-minsup zc (o≤-refl0 b))
                        zc24 : {x₁ : Ordinal} → odef (UnionCF A f mf ay supf0 a) x₁ → (x₁ ≡ sp1) ∨ (x₁ << sp1)
                        zc24 {x₁} ux = MinSUP.x≤sup sup1 (case1 (chain-mono f mf ay supf0 (ZChain.supf-mono zc) (o≤-refl0 b) ux ) )
-                       zc18 : supf0 a o≤ sp1 
+                       zc18 : supf0 a o≤ sp1
                        zc18 = subst (λ k → k o≤ sp1) zc20( MinSUP.minsup zc21 (MinSUP.asm sup1) zc24 )
                  ... | tri> ¬a ¬b c = o≤-refl
 
 
-                 fcup : {u z : Ordinal } → FClosure A f (supf1 u) z → u o≤ px → FClosure A f (supf0 u) z 
+                 fcup : {u z : Ordinal } → FClosure A f (supf1 u) z → u o≤ px → FClosure A f (supf0 u) z
                  fcup {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sf1=sf0 u≤px) fc
-                 fcpu : {u z : Ordinal } → FClosure A f (supf0 u) z → u o≤ px →  FClosure A f (supf1 u) z 
+                 fcpu : {u z : Ordinal } → FClosure A f (supf0 u) z → u o≤ px →  FClosure A f (supf1 u) z
                  fcpu {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sym (sf1=sf0 u≤px)) fc
                  zc11 : {z : Ordinal} → odef (UnionCF A f mf ay supf1 x) z → odef pchainpx z
+                 zc11 {z} ⟪ az , ch-init fc ⟫ = case1 ⟪ az , ch-init fc ⟫
                  zc11 {z} ⟪ az , ch-is-sup u su<sx is-sup fc ⟫ = zc21 fc where
                     u<x :  u o< x
                     u<x =  supf-inject0 supf1-mono su<sx
@@ -859,18 +932,19 @@
                     u≤px = zc-b<x _ u<x
                     zc21 : {z1 : Ordinal } → FClosure A f (supf1 u) z1 → odef pchainpx z1
                     zc21 {z1} (fsuc z2 fc ) with zc21 fc
-                    ... | case1 ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = case1 ⟪ proj2 ( mf _ ua1)  , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫ 
-                    ... | case2 fc = case2 (fsuc _ fc) 
+                    ... | case1 ⟪ ua1 , ch-init fc₁ ⟫ = case1 ⟪ proj2 ( mf _ ua1)  , ch-init (fsuc _ fc₁)  ⟫
+                    ... | case1 ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = case1 ⟪ proj2 ( mf _ ua1)  , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫
+                    ... | case2 fc = case2 (fsuc _ fc)
                     zc21 (init asp refl ) with trio< (supf0 u) (supf0 px) | inspect supf1 u
                     ... | tri< a ¬b ¬c | _ = case1 ⟪ asp , ch-is-sup u a record {fcy<sup = zc13 ; order = zc17
                          ; supu=u = trans (sym (sf1=sf0 (o<→≤ u<px))) (ChainP.supu=u is-sup) } (init asp0 (sym (sf1=sf0 (o<→≤ u<px))) ) ⟫ where
                         u<px :  u o< px
                         u<px =  ZChain.supf-inject zc a
                         asp0 : odef A (supf0 u)
-                        asp0 = ZChain.asupf zc 
+                        asp0 = ZChain.asupf zc
                         zc17 :  {s : Ordinal} {z1 : Ordinal} → supf0 s o< supf0 u →
                             FClosure A f (supf0 s) z1 → (z1 ≡ supf0 u) ∨ (z1 << supf0 u)
-                        zc17 {s} {z1} ss<spx fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k)) ((sf1=sf0 u≤px)) ( ChainP.order is-sup 
+                        zc17 {s} {z1} ss<spx fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k)) ((sf1=sf0 u≤px)) ( ChainP.order is-sup
                          (subst₂ (λ j k → j o< k ) (sym (sf1=sf0 zc18)) (sym (sf1=sf0 u≤px)) ss<spx) (fcpu fc zc18) ) where
                             zc18 : s o≤ px
                             zc18 = ordtrans (ZChain.supf-inject zc ss<spx) u≤px
@@ -879,6 +953,7 @@
                     ... | tri≈ ¬a b ¬c | _ = case2 (init (subst (λ k → odef A k) b (ZChain.asupf zc) ) (sym (trans (sf1=sf0 u≤px) b )))
                     ... | tri> ¬a ¬b c | _ = ⊥-elim ( ¬p<x<op ⟪ ZChain.supf-inject zc c , subst (λ k → u o< k ) (sym (Oprev.oprev=x op)) u<x  ⟫ )
                  zc12 : {z : Ordinal} → odef pchainpx z → odef (UnionCF A f mf ay supf1 x) z
+                 zc12 {z} (case1 ⟪ az , ch-init fc ⟫ ) = ⟪ az , ch-init fc ⟫
                  zc12 {z} (case1 ⟪ az , ch-is-sup u su<sx is-sup fc ⟫ ) = zc21 fc where
                         u<px :  u o< px
                         u<px = ZChain.supf-inject zc su<sx
@@ -887,17 +962,18 @@
                         u≤px : u o≤ px
                         u≤px = o<→≤ u<px
                         s1u<s1x : supf1 u o< supf1 x
-                        s1u<s1x = ordtrans<-≤ (subst₂ (λ j k → j o< k ) (sym (sf1=sf0 u≤px )) (sym (sf1=sf0 o≤-refl)) su<sx) (supf1-mono (o<→≤ px<x) ) 
+                        s1u<s1x = ordtrans<-≤ (subst₂ (λ j k → j o< k ) (sym (sf1=sf0 u≤px )) (sym (sf1=sf0 o≤-refl)) su<sx) (supf1-mono (o<→≤ px<x) )
                         zc21 : {z1 : Ordinal } → FClosure A f (supf0 u) z1 → odef (UnionCF A f mf ay supf1 x) z1
                         zc21 {z1} (fsuc z2 fc ) with zc21 fc
-                        ... | ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫ 
+                        ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-init (fsuc _ fc₁)  ⟫
+                        ... | ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫
                         zc21 (init asp refl ) with trio< u px | inspect supf1 u
-                        ... | tri< a ¬b ¬c | _ = ⟪ asp , ch-is-sup u 
-                             s1u<s1x 
+                        ... | tri< a ¬b ¬c | _ = ⟪ asp , ch-is-sup u
+                             s1u<s1x
                                 record {fcy<sup = zc13 ; order = zc17 ; supu=u = trans (sf1=sf0 u≤px ) (ChainP.supu=u is-sup) } zc14 ⟫  where
                             zc17 :  {s : Ordinal} {z1 : Ordinal} → supf1 s o< supf1 u →
                                   FClosure A f (supf1 s) z1 → (z1 ≡ supf1 u) ∨ (z1 << supf1 u)
-                            zc17 {s} {z1} ss<su fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k)) (sym (sf1=sf0 (o<→≤ u<px))) ( ChainP.order is-sup 
+                            zc17 {s} {z1} ss<su fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k)) (sym (sf1=sf0 (o<→≤ u<px))) ( ChainP.order is-sup
                                (subst₂ (λ j k → j o< k ) (sf1=sf0 s≤px) (sf1=sf0 (o<→≤ u<px)) ss<su) (fcup fc s≤px) ) where
                                   s≤px : s o≤ px -- ss<su : supf1 s o< supf1 u
                                   s≤px = ordtrans ( supf-inject0 supf1-mono ss<su ) (o<→≤ u<px)
@@ -905,7 +981,7 @@
                             zc14 = init (subst (λ k → odef A k ) (sym (sf1=sf0 u≤px)) asp) (sf1=sf0 u≤px)
                             zc13 : {z : Ordinal } → FClosure A f y z → (z ≡ supf1 u) ∨ ( z << supf1 u )
                             zc13 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 u≤px )) ( ChainP.fcy<sup is-sup fc )
-                        ... | tri≈ ¬a b ¬c | _ = ⟪ asp , ch-is-sup px (subst (λ k → supf1 k o< supf1 x) b s1u<s1x)  record { fcy<sup = zc13 
+                        ... | tri≈ ¬a b ¬c | _ = ⟪ asp , ch-is-sup px (subst (λ k → supf1 k o< supf1 x) b s1u<s1x)  record { fcy<sup = zc13
                               ; order = zc17 ; supu=u = zc18   } (init asupf1 (trans (sf1=sf0 o≤-refl ) (cong supf0 (sym b)))  ) ⟫ where
                             zc13 : {z : Ordinal } → FClosure A f y z → (z ≡ supf1 px) ∨ ( z << supf1 px )
                             zc13 {z} fc = subst (λ k → (z ≡ k) ∨ (z << k)) (trans (cong supf0 b) (sym (sf1=sf0 o≤-refl))) (ChainP.fcy<sup is-sup fc )
@@ -931,9 +1007,10 @@
                         zc20 = ?
                         zc21 : {z1 : Ordinal } → FClosure A f (supf0 px) z1 → odef (UnionCF A f mf ay supf1 x) z1
                         zc21 {z1} (fsuc z2 fc ) with zc21 fc
-                        ... | ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫ 
+                        ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-init (fsuc _ fc₁)  ⟫
+                        ... | ⟪ ua1 , ch-is-sup u u<x u1-is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-is-sup u u<x u1-is-sup (fsuc _ fc₁) ⟫
                         zc21 (init asp refl ) with  zc20
-                        ... | case1 sfpx=px =  ⟪ asp , ch-is-sup px zc18 
+                        ... | case1 sfpx=px =  ⟪ asp , ch-is-sup px zc18
                                      record {fcy<sup = zc13 ; order = zc17 ; supu=u = zc15 } zc14 ⟫ where
                               zc15 : supf1 px ≡ px
                               zc15 = trans (sf1=sf0 o≤-refl ) (sfpx=px)
@@ -945,28 +1022,29 @@
                               zc13 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 o≤-refl)) ( ZChain.fcy<sup zc o≤-refl fc )
                               zc17 :  {s : Ordinal} {z1 : Ordinal} → supf1 s o< supf1 px →
                                     FClosure A f (supf1 s) z1 → (z1 ≡ supf1 px) ∨ (z1 << supf1 px)
-                              zc17 {s} {z1} ss<spx fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k )) mins-is-spx 
+                              zc17 {s} {z1} ss<spx fc = subst (λ k → (z1 ≡ k) ∨ (z1 << k )) mins-is-spx
                                      (MinSUP.x≤sup mins (csupf17 (fcup fc (o<→≤ s<px) )) ) where
                                   mins : MinSUP A (UnionCF A f mf ay supf0 px)
                                   mins = ZChain.minsup zc o≤-refl
-                                  mins-is-spx : MinSUP.sup mins ≡ supf1 px 
+                                  mins-is-spx : MinSUP.sup mins ≡ supf1 px
                                   mins-is-spx = trans (sym ( ZChain.supf-is-minsup zc o≤-refl ) ) (sym (sf1=sf0 o≤-refl ))
                                   s<px : s o< px
                                   s<px = supf-inject0 supf1-mono ss<spx
                                   sf<px : supf0 s o< px
                                   sf<px = subst₂ (λ j k → j o< k ) (sf1=sf0 (o<→≤ s<px)) (trans (sf1=sf0 o≤-refl) (sfpx=px)) ss<spx
                                   csupf17 : {z1 : Ordinal } → FClosure A f (supf0 s) z1 → odef (UnionCF A f mf ay supf0 px) z1
-                                  csupf17 (init as refl ) = ZChain.csupf zc sf<px 
+                                  csupf17 (init as refl ) = ZChain.csupf zc sf<px
                                   csupf17 (fsuc x fc) = ZChain.f-next zc (csupf17 fc)
 
                         ... | case2 sfp<px  with ZChain.csupf zc sfp<px --  odef (UnionCF A f mf ay supf0 px) (supf0 px)
-                        ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ ua1 , ch-is-sup u zc18  
+                        ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ ua1 , ch-init fc₁ ⟫
+                        ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ ua1 , ch-is-sup u zc18
                               record { fcy<sup = z10 ; order = z11 ; supu=u = z12 } (fcpu fc₁ ? ) ⟫  where
-                                 z10 :  {z : Ordinal } → FClosure A f y z → (z ≡ supf1 u) ∨ ( z << supf1 u ) 
+                                 z10 :  {z : Ordinal } → FClosure A f y z → (z ≡ supf1 u) ∨ ( z << supf1 u )
                                  z10 {z} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 ? )) (ChainP.fcy<sup is-sup fc)
-                                 z11  : {s z1 : Ordinal} → (lt : supf1 s o< supf1 u ) → FClosure A f (supf1 s ) z1 
+                                 z11  : {s z1 : Ordinal} → (lt : supf1 s o< supf1 u ) → FClosure A f (supf1 s ) z1
                                       → (z1 ≡ supf1 u ) ∨ ( z1 << supf1 u )
-                                 z11 {s} {z} lt fc  = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 ? )) 
+                                 z11 {s} {z} lt fc  = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 ? ))
                                      (ChainP.order is-sup lt0 (fcup fc s≤px )) where
                                        s<u : s o< u
                                        s<u = supf-inject0 supf1-mono lt
@@ -984,11 +1062,11 @@
                  record STMP {z : Ordinal} (z≤x : z o≤ x ) : Set (Level.suc n) where
                      field
                          tsup : MinSUP A (UnionCF A f mf ay supf1 z)
-                         tsup=sup : supf1 z ≡ MinSUP.sup tsup 
+                         tsup=sup : supf1 z ≡ MinSUP.sup tsup
 
                  sup : {z : Ordinal} → (z≤x : z o≤ x ) → STMP z≤x
-                 sup {z} z≤x with trio< z px 
-                 ... | tri< a ¬b ¬c = record { tsup = record { sup = MinSUP.sup m  ; asm = MinSUP.asm m  
+                 sup {z} z≤x with trio< z px
+                 ... | tri< a ¬b ¬c = record { tsup = record { sup = MinSUP.sup m  ; asm = MinSUP.asm m
                          ; x≤sup = ms00 ; minsup = ms01 } ; tsup=sup = trans (sf1=sf0 (o<→≤ a) ) (ZChain.supf-is-minsup zc (o<→≤ a)) } where
                     m = ZChain.minsup zc (o<→≤ a)
                     ms00 :  {x : Ordinal} → odef (UnionCF A f mf ay supf1 z) x → (x ≡ MinSUP.sup m) ∨ (x << MinSUP.sup m)
@@ -996,7 +1074,7 @@
                     ms01 : {sup2 : Ordinal} → odef A sup2 → ({x : Ordinal} →
                         odef (UnionCF A f mf ay supf1 z) x → (x ≡ sup2) ∨ (x << sup2)) → MinSUP.sup m o≤ sup2
                     ms01 {sup2} us P = MinSUP.minsup m ? ?
-                 ... | tri≈ ¬a b ¬c = record { tsup = record { sup = MinSUP.sup m  ; asm = MinSUP.asm m  
+                 ... | tri≈ ¬a b ¬c = record { tsup = record { sup = MinSUP.sup m  ; asm = MinSUP.asm m
                          ; x≤sup = ms00 ; minsup = ms01 } ; tsup=sup = trans (sf1=sf0 (o≤-refl0 b) ) (ZChain.supf-is-minsup zc (o≤-refl0 b))  } where
                     m = ZChain.minsup zc (o≤-refl0 b)
                     ms00 :  {x : Ordinal} → odef (UnionCF A f mf ay supf1 z) x → (x ≡ MinSUP.sup m) ∨ (x << MinSUP.sup m)
@@ -1013,14 +1091,14 @@
                         odef (UnionCF A f mf ay supf1 z) x → (x ≡ sup2) ∨ (x << sup2)) → MinSUP.sup m o≤ sup2
                     ms01 {sup2} us P = MinSUP.minsup m ? ?
 
-                 csupf1 : {z1 : Ordinal } → supf1 z1 o< x → odef (UnionCF A f mf ay supf1 x) (supf1 z1) 
+                 csupf1 : {z1 : Ordinal } → supf1 z1 o< x → odef (UnionCF A f mf ay supf1 x) (supf1 z1)
                  csupf1 {z1} sz1<x = csupf2 where
                      --  z1 o< px → supf1 z1 ≡ supf0 z1
                      --  z1 ≡ px , supf0 px o< px   .. px o< z1, x o≤ z1 ...  supf1 z1 ≡ sp1
                      --  z1 ≡ px , supf0 px ≡  px
                      psz1≤px :  supf1 z1 o≤ px
-                     psz1≤px =  subst (λ k → supf1 z1 o< k ) (sym opx=x) sz1<x 
-                     csupf2 :  odef (UnionCF A f mf ay supf1 x) (supf1 z1) 
+                     psz1≤px =  subst (λ k → supf1 z1 o< k ) (sym opx=x) sz1<x
+                     csupf2 :  odef (UnionCF A f mf ay supf1 x) (supf1 z1)
                      csupf2 with  trio< z1 px | inspect supf1 z1
                      csupf2 | tri< a ¬b ¬c | record { eq = eq1 } with osuc-≡< psz1≤px
                      ... | case2 lt  = zc12 (case1 cs03)  where
@@ -1030,7 +1108,7 @@
                      ... | case1 sfz=sfpx =  zc12 (case2 (init (ZChain.asupf zc) cs04 )) where
                            supu=u : supf1 (supf1 z1) ≡ supf1 z1
                            supu=u = trans (cong supf1 sfz=px) (sym sfz=sfpx)
-                           cs04 : supf0 px ≡ supf0 z1 
+                           cs04 : supf0 px ≡ supf0 z1
                            cs04 = begin
                              supf0 px ≡⟨ sym (sf1=sf0 o≤-refl)  ⟩
                              supf1 px ≡⟨ sym sfz=sfpx ⟩
@@ -1043,7 +1121,7 @@
                            cs06 : supf0 px o< osuc px
                            cs06 = subst (λ k → supf0 px o< k ) (sym opx=x) ?
                      csupf2 | tri≈ ¬a b ¬c | record { eq = eq1 } =  zc12 (case2 (init (ZChain.asupf zc) (cong supf0 (sym b))))
-                     csupf2 | tri> ¬a ¬b px<z1 | record { eq = eq1 } = ? 
+                     csupf2 | tri> ¬a ¬b px<z1 | record { eq = eq1 } = ?
                      -- ⊥-elim ( ¬p<x<op ⟪ px<z1 , subst (λ k → z1 o< k) (sym opx=x) z1<x ⟫ )
 
 
@@ -1054,7 +1132,7 @@
                  --     supf1 x ≡ supf0 px because of supfmax
 
                  supf1 : Ordinal → Ordinal
-                 supf1 z with trio< z px 
+                 supf1 z with trio< z px
                  ... | tri< a ¬b ¬c = supf0 z
                  ... | tri≈ ¬a b ¬c = supf0 px
                  ... | tri> ¬a ¬b c = supf0 px
@@ -1069,7 +1147,7 @@
                  zc17 = ? -- px o< z, px o< supf0 px
 
                  supf-mono1 : {z w : Ordinal } → z o≤ w → supf1 z o≤ supf1 w
-                 supf-mono1 {z} {w} z≤w with trio< w px 
+                 supf-mono1 {z} {w} z≤w with trio< w px
                  ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (ordtrans≤-< z≤w a))) refl ( supf-mono z≤w )
                  ... | tri≈ ¬a refl ¬c with trio< z px
                  ... | tri< a ¬b ¬c = zc17
@@ -1084,22 +1162,25 @@
                  pchain1 = UnionCF A f mf ay supf1 x
 
                  zc10 :  {z : Ordinal} → OD.def (od pchain) z → OD.def (od pchain1) z
+                 zc10 {z} ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
                  zc10 {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ = ⟪ az , ch-is-sup u1 ? ?  ?  ⟫
 
-                 zc111 : {z : Ordinal} → z o< px → OD.def (od pchain1) z → OD.def (od pchain) z 
+                 zc111 : {z : Ordinal} → z o< px → OD.def (od pchain1) z → OD.def (od pchain) z
+                 zc111 {z} z<px ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
                  zc111 {z} z<px ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ = ⟪ az , ch-is-sup u1 ? ?  ?  ⟫
 
                  zc11 : (¬ xSUP (UnionCF A f mf ay supf0 px) f x ) ∨  (HasPrev A pchain f x )
                         → {z : Ordinal} → OD.def (od pchain1) z → OD.def (od pchain) z ∨ ( (supf0 px ≡ px) ∧ FClosure A f px z )
-                 zc11 P {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ with trio< u1 px 
-                 ... | tri< u1<px ¬b ¬c = case1 ⟪ az , ch-is-sup u1 ? ? fc  ⟫ 
+                 zc11 P {z} ⟪ az , ch-init fc ⟫ = case1 ⟪ az , ch-init fc ⟫
+                 zc11 P {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ with trio< u1 px
+                 ... | tri< u1<px ¬b ¬c = case1 ⟪ az , ch-is-sup u1 ? ? fc  ⟫
                  ... | tri≈ ¬a eq ¬c  = case2 ⟪ subst (λ k → supf0 k ≡ k) eq s1u=u , subst (λ k → FClosure A f k z) zc12 ? ⟫ where
                         s1u=u : supf0 u1 ≡ u1
                         s1u=u = ? -- ChainP.supu=u u1-is-sup
-                        zc12 : supf0 u1 ≡ px 
+                        zc12 : supf0 u1 ≡ px
                         zc12 = trans s1u=u eq
                  zc11 (case1 ¬sp=x) {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ | tri> ¬a ¬b px<u = ⊥-elim (¬sp=x zcsup) where
-                        eq : u1 ≡ x 
+                        eq : u1 ≡ x
                         eq with trio< u1 x
                         ... | tri< a ¬b ¬c = ⊥-elim ( ¬p<x<op ⟪ px<u , subst (λ k → u1 o< k ) (sym (Oprev.oprev=x op )) a ⟫ )
                         ... | tri≈ ¬a b ¬c = b
@@ -1107,25 +1188,26 @@
                         s1u=x : supf0 u1 ≡ x
                         s1u=x = trans ? eq
                         zc13 : osuc px o< osuc u1
-                        zc13 = o≤-refl0 ( trans (Oprev.oprev=x op) (sym eq ) ) 
+                        zc13 = o≤-refl0 ( trans (Oprev.oprev=x op) (sym eq ) )
                         x≤sup : {w : Ordinal} → odef (UnionCF A f mf ay supf0 px) w → (w ≡ x) ∨ (w << x)
+                        x≤sup {w} ⟪ az , ch-init {w} fc ⟫ = subst (λ k → (w ≡ k) ∨ (w << k)) s1u=x ?
                         x≤sup {w} ⟪ az , ch-is-sup u u<x is-sup fc ⟫ with osuc-≡< ( supf-mono (ordtrans (o<→≤ u<x) ? ))
                         ... | case1 eq1 = ⊥-elim ( o<¬≡ zc14 ? ) where
                                  zc14 : u ≡ osuc px
                                  zc14 = begin
-                                    u ≡⟨ sym ( ChainP.supu=u is-sup) ⟩ 
-                                    supf0 u ≡⟨ ? ⟩ 
-                                    supf0 u1 ≡⟨ s1u=x ⟩ 
-                                    x ≡⟨ sym (Oprev.oprev=x op) ⟩ 
+                                    u ≡⟨ sym ( ChainP.supu=u is-sup) ⟩
+                                    supf0 u ≡⟨ ? ⟩
+                                    supf0 u1 ≡⟨ s1u=x ⟩
+                                    x ≡⟨ sym (Oprev.oprev=x op) ⟩
                                     osuc px ∎ where open ≡-Reasoning
                         ... | case2 lt = subst (λ k → (w ≡ k) ∨ (w << k)) s1u=x ?
                         zc12 : supf0 x ≡ u1
                         zc12 = subst  (λ k → supf0 k ≡ u1) eq ?
-                        zcsup : xSUP (UnionCF A f mf ay supf0 px) f x 
+                        zcsup : xSUP (UnionCF A f mf ay supf0 px) f x
                         zcsup = record { ax =  subst (λ k → odef A k) (trans zc12 eq) (ZChain.asupf zc)
                                  ; is-sup = record { x≤sup = x≤sup ; minsup = ? } }
                  zc11 (case2 hp) {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ | tri> ¬a ¬b px<u = case1 ? where
-                        eq : u1 ≡ x 
+                        eq : u1 ≡ x
                         eq with trio< u1 x
                         ... | tri< a ¬b ¬c = ⊥-elim ( ¬p<x<op ⟪ px<u , subst (λ k → u1 o< k ) (sym (Oprev.oprev=x op )) a ⟫ )
                         ... | tri≈ ¬a b ¬c = b
@@ -1145,10 +1227,10 @@
                  record STMP {z : Ordinal} (z≤x : z o≤ x ) : Set (Level.suc n) where
                      field
                          tsup : MinSUP A (UnionCF A f mf ay supf1 z)
-                         tsup=sup : supf1 z ≡ MinSUP.sup tsup 
+                         tsup=sup : supf1 z ≡ MinSUP.sup tsup
 
                  sup : {z : Ordinal} → (z≤x : z o≤ x ) → STMP z≤x
-                 sup {z} z≤x with trio< z px 
+                 sup {z} z≤x with trio< z px
                  ... | tri< a ¬b ¬c = ? -- jrecord { tsup = ZChain.minsup zc (o<→≤ a)  ; tsup=sup = ZChain.supf-is-minsup zc (o<→≤ a) }
                  ... | tri≈ ¬a b ¬c = ? -- record { tsup = ZChain.minsup zc (o≤-refl0 b)  ; tsup=sup = ZChain.supf-is-minsup zc (o≤-refl0 b) }
                  ... | tri> ¬a ¬b px<z = zc35 where
@@ -1156,15 +1238,15 @@
                      zc30 with osuc-≡< z≤x
                      ... | case1 eq = eq
                      ... | case2 z<x = ⊥-elim (¬p<x<op ⟪ px<z , subst (λ k → z o< k ) (sym (Oprev.oprev=x op)) z<x ⟫ )
-                     zc32 = ZChain.sup zc o≤-refl 
+                     zc32 = ZChain.sup zc o≤-refl
                      zc34 : ¬ (supf0 px ≡ px) → {w : HOD} → UnionCF A f mf ay supf0 z ∋ w → (w ≡ SUP.sup zc32) ∨ (w < SUP.sup zc32)
                      zc34 ne {w} lt with zc11 ? ⟪ proj1 lt , ? ⟫
-                     ... | case1 lt = SUP.x≤sup zc32 lt 
+                     ... | case1 lt = SUP.x≤sup zc32 lt
                      ... | case2 ⟪ spx=px  , fc ⟫ = ⊥-elim ( ne spx=px )
                      zc33 : supf0 z ≡ & (SUP.sup zc32)
                      zc33 = ? -- trans (sym (supfx (o≤-refl0 (sym zc30)))) ( ZChain.supf-is-minsup zc o≤-refl  )
                      zc36 : ¬ (supf0 px ≡ px) → STMP z≤x
-                     zc36 ne = ? -- record { tsup = record { sup = SUP.sup zc32 ; as = SUP.as zc32 ; x≤sup = zc34 ne } ; tsup=sup = zc33  } 
+                     zc36 ne = ? -- record { tsup = record { sup = SUP.sup zc32 ; as = SUP.as zc32 ; x≤sup = zc34 ne } ; tsup=sup = zc33  }
                      zc35 : STMP z≤x
                      zc35 with trio< (supf0 px) px
                      ... | tri< a ¬b ¬c = zc36 ¬b
@@ -1175,23 +1257,23 @@
                  sup=u : {b : Ordinal} (ab : odef A b) →
                     b o≤ x → IsMinSUP A (UnionCF A f mf ay supf0 b) supf0 ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf0 b) f b ) → supf0 b ≡ b
                  sup=u {b} ab b≤x is-sup with trio< b px
-                 ... | tri< a ¬b ¬c = ZChain.sup=u zc ab (o<→≤ a) ⟪ record { x≤sup = λ lt → IsMinSUP.x≤sup (proj1 is-sup) lt } , proj2 is-sup ⟫ 
-                 ... | tri≈ ¬a b ¬c = ZChain.sup=u zc ab (o≤-refl0 b) ⟪ record { x≤sup = λ lt → IsMinSUP.x≤sup (proj1 is-sup) lt } , proj2 is-sup ⟫ 
+                 ... | tri< a ¬b ¬c = ZChain.sup=u zc ab (o<→≤ a) ⟪ record { x≤sup = λ lt → IsMinSUP.x≤sup (proj1 is-sup) lt } , proj2 is-sup ⟫
+                 ... | tri≈ ¬a b ¬c = ZChain.sup=u zc ab (o≤-refl0 b) ⟪ record { x≤sup = λ lt → IsMinSUP.x≤sup (proj1 is-sup) lt } , proj2 is-sup ⟫
                  ... | tri> ¬a ¬b px<b = zc31 ? where
                      zc30 : x ≡ b
                      zc30 with osuc-≡< b≤x
                      ... | case1 eq = sym (eq)
                      ... | case2 b<x = ⊥-elim (¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ )
-                     zcsup : xSUP (UnionCF A f mf ay supf0 px) supf0 x 
+                     zcsup : xSUP (UnionCF A f mf ay supf0 px) supf0 x
                      zcsup with zc30
-                     ... | refl = record { ax = ab ; is-sup = record { x≤sup = λ {w} lt → 
-                        IsMinSUP.x≤sup (proj1 is-sup) ? ; minsup = ? } } 
+                     ... | refl = record { ax = ab ; is-sup = record { x≤sup = λ {w} lt →
+                        IsMinSUP.x≤sup (proj1 is-sup) ? ; minsup = ? } }
                      zc31 : ( (¬ xSUP (UnionCF A f mf ay supf0 px) supf0 x ) ∨  HasPrev A (UnionCF A f mf ay supf0 px) f x ) → supf0 b ≡ b
                      zc31 (case1 ¬sp=x) with zc30
                      ... | refl = ⊥-elim (¬sp=x zcsup )
                      zc31 (case2 hasPrev ) with zc30
-                     ... | refl = ⊥-elim ( proj2 is-sup record { ax = HasPrev.ax hasPrev ; y = HasPrev.y hasPrev 
-                                ; ay = ? ; x=fy = HasPrev.x=fy hasPrev } ) 
+                     ... | refl = ⊥-elim ( proj2 is-sup record { ax = HasPrev.ax hasPrev ; y = HasPrev.y hasPrev
+                                ; ay = ? ; x=fy = HasPrev.x=fy hasPrev } )
 
      ... | no lim = zc5 where
 
@@ -1210,10 +1292,10 @@
           pchain = UnionCF A f mf ay supf0 x
 
           ptotal0 : IsTotalOrderSet pchain
-          ptotal0 {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where 
+          ptotal0 {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where
                uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
-               uz01 = chain-total A f mf ay supf0 ( (proj2 ca)) ( (proj2 cb)) 
-  
+               uz01 = chain-total A f mf ay supf0 ( (proj2 ca)) ( (proj2 cb))
+
           usup : MinSUP A pchain
           usup = minsupP pchain (λ lt → proj1 lt) ptotal0
           spu = MinSUP.sup usup
@@ -1229,66 +1311,71 @@
 
           is-max-hp : (supf : Ordinal → Ordinal) (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay supf x) a →
                 b o< x → (ab : odef A b) →
-                HasPrev A (UnionCF A f mf ay supf x) f b  → 
+                HasPrev A (UnionCF A f mf ay supf x) f b  →
                 * a < * b → odef (UnionCF A f mf ay supf x) b
           is-max-hp supf x {a} {b} ua b<x ab has-prev a<b with HasPrev.ay has-prev
-          ... | ⟪ ab0 , ch-is-sup u u<x is-sup fc ⟫ = ? -- ⟪ ab , 
+          ... | ⟪ ab0 , ch-init fc ⟫ = ⟪ ab , ch-init ( subst (λ k → FClosure A f y k) (sym (HasPrev.x=fy has-prev)) (fsuc _ fc )) ⟫
+          ... | ⟪ ab0 , ch-is-sup u u<x is-sup fc ⟫ = ? -- ⟪ ab ,
                      -- subst (λ k → UChain A f mf ay supf x k )
-                     --     (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x is-sup (fsuc _ fc))  ⟫ 
+                     --     (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x is-sup (fsuc _ fc))  ⟫
 
-          zc70 : HasPrev A pchain f x  → ¬ xSUP pchain f x 
+          zc70 : HasPrev A pchain f x  → ¬ xSUP pchain f x
           zc70 pr xsup = ?
 
           no-extension : ¬ ( xSUP (UnionCF A f mf ay supf0 x) supf0 x ) → ZChain A f mf ay x
-          no-extension ¬sp=x  = ? where -- record { supf = supf1  ; sup=u = sup=u  
+          no-extension ¬sp=x  = ? where -- record { supf = supf1  ; sup=u = sup=u
                -- ; sup = sup ; supf-is-sup = sis ; supf-mono = {!!} ; asupf = ? } where
                  supfu : {u : Ordinal } → ( a : u o< x ) → (z : Ordinal) → Ordinal
                  supfu {u} a z = ZChain.supf (pzc (osuc u) (ob<x lim a)) z
                  pchain0=1 : pchain ≡ pchain1
                  pchain0=1 =  ==→o≡ record { eq→ = zc10 ; eq← = zc11 } where
                      zc10 :  {z : Ordinal} → OD.def (od pchain) z → OD.def (od pchain1) z
+                     zc10 {z} ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
                      zc10 {z} ⟪ az , ch-is-sup u u<x is-sup fc ⟫ = zc12 fc where
                           zc12 : {z : Ordinal} →  FClosure A f (supf0 u) z → odef pchain1 z
                           zc12 (fsuc x fc) with zc12 fc
-                          ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-is-sup u u<x is-sup (fsuc _ fc₁) ⟫ 
-                          zc12 (init asu su=z ) = ⟪ subst (λ k → odef A k) su=z asu , ch-is-sup u ? ? (init ? ? ) ⟫ 
+                          ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-init (fsuc _ fc₁)  ⟫
+                          ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-is-sup u u<x is-sup (fsuc _ fc₁) ⟫
+                          zc12 (init asu su=z ) = ⟪ subst (λ k → odef A k) su=z asu , ch-is-sup u ? ? (init ? ? ) ⟫
                      zc11 :  {z : Ordinal} → OD.def (od pchain1) z → OD.def (od pchain) z
+                     zc11 {z} ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫
                      zc11 {z} ⟪ az , ch-is-sup u u<x is-sup fc ⟫ = zc13 fc where
                           zc13 : {z : Ordinal} →  FClosure A f (supf1 u) z → odef pchain z
                           zc13 (fsuc x fc) with zc13 fc
-                          ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-is-sup u u<x is-sup (fsuc _ fc₁) ⟫ 
+                          ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-init (fsuc _ fc₁)  ⟫
+                          ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ proj2 ( mf _ ua1)  , ch-is-sup u u<x is-sup (fsuc _ fc₁) ⟫
                           zc13 (init asu su=z ) with trio< u x
-                          ... | tri< a ¬b ¬c = ⟪ ? , ch-is-sup u ? ? (init ? ? ) ⟫ 
+                          ... | tri< a ¬b ¬c = ⟪ ? , ch-is-sup u ? ? (init ? ? ) ⟫
                           ... | tri≈ ¬a b ¬c = ?
                           ... | tri> ¬a ¬b c = ? -- ⊥-elim ( o≤> u<x c )
                  sup : {z : Ordinal} → z o≤ x → SUP A (UnionCF A f mf ay supf1 z)
                  sup {z} z≤x with trio< z x
                  ... | tri< a ¬b ¬c = SUP⊆ ? (ZChain.sup (pzc (osuc z) {!!}) {!!} )
                  ... | tri≈ ¬a b ¬c = {!!}
-                 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x )) 
+                 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x ))
                  sis : {z : Ordinal} (x≤z : z o≤ x) → supf1 z ≡ & (SUP.sup (sup {!!}))
                  sis {z} z≤x with trio< z x
                  ... | tri< a ¬b ¬c = {!!} where
                      zc8 = ZChain.supf-is-minsup (pzc z a) {!!}
                  ... | tri≈ ¬a b ¬c = {!!}
-                 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x )) 
+                 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x ))
                  sup=u : {b : Ordinal} (ab : odef A b) → b o≤ x → IsMinSUP A (UnionCF A f mf ay supf1 b) f ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf1 b) f b ) → supf1 b ≡ b
                  sup=u {z} ab z≤x is-sup with trio< z x
                  ... | tri< a ¬b ¬c = ? -- ZChain.sup=u (pzc (osuc b)  (ob<x lim a))  ab {!!} record { x≤sup = {!!} }
                  ... | tri≈ ¬a b ¬c = {!!} -- ZChain.sup=u (pzc (osuc ?)  ?)  ab {!!} record { x≤sup = {!!} }
-                 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x )) 
+                 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x ))
 
-          zc5 : ZChain A f mf ay x 
+          zc5 : ZChain A f mf ay x
           zc5 with ODC.∋-p O A (* x)
           ... | no noax = no-extension {!!} -- ¬ A ∋ p, just skip
-          ... | yes ax with ODC.p∨¬p O ( HasPrev A pchain f x  )   
+          ... | yes ax with ODC.p∨¬p O ( HasPrev A pchain f x  )
                -- we have to check adding x preserve is-max ZChain A y f mf x
-          ... | case1 pr = no-extension {!!} 
+          ... | case1 pr = no-extension {!!}
           ... | case2 ¬fy<x with ODC.p∨¬p O (IsMinSUP A pchain f ax )
-          ... | case1 is-sup = ? -- record { supf = supf1  ; sup=u = {!!} 
-               -- ; sup = {!!} ; supf-is-sup = {!!} ; supf-mono = {!!};  asupf = {!!}  } -- where -- x is a sup of (zc ?) 
+          ... | case1 is-sup = ? -- record { supf = supf1  ; sup=u = {!!}
+               -- ; sup = {!!} ; supf-is-sup = {!!} ; supf-mono = {!!};  asupf = {!!}  } -- where -- x is a sup of (zc ?)
           ... | case2 ¬x=sup =  no-extension {!!} -- x is not f y' nor sup of former ZChain from y -- no extention
-         
+
      ---
      --- the maximum chain  has fix point of any ≤-monotonic function
      ---
@@ -1296,13 +1383,13 @@
      SZ : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) → {y : Ordinal} (ay : odef A y) → (x : Ordinal) → ZChain A f mf ay x
      SZ f mf {y} ay x = TransFinite {λ z → ZChain A f mf ay z  } (λ x → ind f mf ay x   ) x
 
-     msp0 : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {x y : Ordinal} (ay : odef A y)  
-         → (zc : ZChain A f mf ay x ) 
+     msp0 : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {x y : Ordinal} (ay : odef A y)
+         → (zc : ZChain A f mf ay x )
          → MinSUP A (UnionCF A f mf ay (ZChain.supf zc) x)
      msp0 f mf {x} ay zc = minsupP (UnionCF A f mf ay (ZChain.supf zc) x) (ZChain.chain⊆A zc) (ZChain.f-total zc)
 
      fixpoint :  ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f )  (zc : ZChain A f mf as0 (& A) )
-            → (sp1 : MinSUP A (ZChain.chain zc)) 
+            → (sp1 : MinSUP A (ZChain.chain zc))
             → (ssp<as :  ZChain.supf zc (MinSUP.sup sp1) o< ZChain.supf zc (& A))
             → f (MinSUP.sup sp1)  ≡ MinSUP.sup sp1
      fixpoint f mf zc sp1 ssp<as = z14 where
@@ -1312,7 +1399,7 @@
            sp = MinSUP.sup sp1
            asp : odef A sp
            asp = MinSUP.asm sp1
-           z10 :  {a b : Ordinal } → (ca : odef chain a ) → supf b o< supf (& A) → (ab : odef A b ) 
+           z10 :  {a b : Ordinal } → (ca : odef chain a ) → supf b o< supf (& A) → (ab : odef A b )
               →  HasPrev A chain f b  ∨  IsSUP A (UnionCF A f mf as0 (ZChain.supf zc) b) ab
               → * a < * b  → odef chain b
            z10 = ZChain1.is-max (SZ1 f mf as0 zc (& A) )
@@ -1325,22 +1412,22 @@
                z13 :  * (& s) < * sp
                z13 with MinSUP.x≤sup sp1 ( ZChain.chain∋init zc )
                ... | case1 eq = ⊥-elim ( ne eq )
-               ... | case2 lt = lt 
+               ... | case2 lt = lt
                z19 : IsSUP A (UnionCF A f mf as0 (ZChain.supf zc) sp) asp
                z19 = record {   x≤sup = z20 }  where
                    z20 : {y : Ordinal} → odef (UnionCF A f mf as0 (ZChain.supf zc) sp) y → (y ≡ sp) ∨ (y << sp)
-                   z20 {y} zy with MinSUP.x≤sup sp1 
+                   z20 {y} zy with MinSUP.x≤sup sp1
                        (subst (λ k → odef chain k ) (sym &iso) (chain-mono f mf as0 supf (ZChain.supf-mono zc) (o<→≤ z22)  zy ))
-                   ... | case1 y=p = case1 (subst (λ k → k ≡ _ ) &iso y=p ) 
+                   ... | case1 y=p = case1 (subst (λ k → k ≡ _ ) &iso y=p )
                    ... | case2 y<p = case2 (subst (λ k → * k < _ ) &iso y<p )
            z14 :  f sp ≡ sp
            z14 with ZChain.f-total zc (subst (λ k → odef chain k) (sym &iso)  (ZChain.f-next zc z12 )) (subst (λ k → odef chain k) (sym &iso) z12 )
            ... | tri< a ¬b ¬c = ⊥-elim z16 where
                z16 : ⊥
                z16 with proj1 (mf (( MinSUP.sup sp1)) ( MinSUP.asm sp1 ))
-               ... | case1 eq = ⊥-elim (¬b (sym eq) ) 
-               ... | case2 lt = ⊥-elim (¬c lt ) 
-           ... | tri≈ ¬a b ¬c = subst₂ (λ j k → j ≡ k ) &iso &iso ( cong (&) b ) 
+               ... | case1 eq = ⊥-elim (¬b (sym eq) )
+               ... | case2 lt = ⊥-elim (¬c lt )
+           ... | tri≈ ¬a b ¬c = subst₂ (λ j k → j ≡ k ) &iso &iso ( cong (&) b )
            ... | tri> ¬a ¬b c = ⊥-elim z17 where
                z15 : (f sp ≡ MinSUP.sup sp1) ∨ (* (f sp) < * (MinSUP.sup sp1) )
                z15  = MinSUP.x≤sup sp1 (ZChain.f-next zc z12 )
@@ -1367,37 +1454,37 @@
      --
 
      z04 :  (nmx : ¬ Maximal A ) → (zc : ZChain A (cf nmx) (cf-is-≤-monotonic nmx) as0 (& A)) → ⊥
-     z04 nmx zc = <-irr0  {* (cf nmx c)} {* c} 
-           (subst (λ k → odef A k ) (sym &iso) (proj1 (is-cf nmx (MinSUP.asm  msp1 )))) 
+     z04 nmx zc = <-irr0  {* (cf nmx c)} {* c}
+           (subst (λ k → odef A k ) (sym &iso) (proj1 (is-cf nmx (MinSUP.asm  msp1 ))))
            (subst (λ k → odef A k) (sym &iso) (MinSUP.asm msp1) )
            (case1 ( cong (*)( fixpoint (cf nmx) (cf-is-≤-monotonic nmx ) zc msp1  ss<sa ))) -- x ≡ f x ̄
                 (proj1 (cf-is-<-monotonic nmx c (MinSUP.asm msp1 ))) where          -- x < f x
 
           supf = ZChain.supf zc
           msp1 : MinSUP A (ZChain.chain zc)
-          msp1 = msp0 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc 
-          c : Ordinal 
-          c = MinSUP.sup msp1 
+          msp1 = msp0 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc
+          c : Ordinal
+          c = MinSUP.sup msp1
           mc = c
           mc<A : mc o< & A
           mc<A =  ∈∧P→o< ⟪ MinSUP.asm msp1 , lift true ⟫
           c=mc : c ≡ mc
           c=mc = refl
-          z20 : mc << cf nmx mc 
+          z20 : mc << cf nmx mc
           z20 = proj1 (cf-is-<-monotonic nmx mc (MinSUP.asm msp1) )
           asc : odef A (supf mc)
           asc = ZChain.asupf zc
-          spd : MinSUP A (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc ) 
+          spd : MinSUP A (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc )
           spd = ysup (cf nmx)  (cf-is-≤-monotonic nmx) asc
           d = MinSUP.sup spd
           d<A : d o< & A
           d<A =  ∈∧P→o< ⟪ MinSUP.asm spd , lift true ⟫
           msup : MinSUP  A (UnionCF A (cf nmx)  (cf-is-≤-monotonic nmx) as0 supf d)
-          msup = ZChain.minsup zc (o<→≤ d<A) 
+          msup = ZChain.minsup zc (o<→≤ d<A)
           sd=ms : supf d ≡ MinSUP.sup ( ZChain.minsup zc (o<→≤ d<A) )
           sd=ms = ZChain.supf-is-minsup zc (o<→≤ d<A)
 
-          sc<<d : {mc : Ordinal } → (asc : odef A (supf mc)) → (spd : MinSUP A (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc )) 
+          sc<<d : {mc : Ordinal } → (asc : odef A (supf mc)) → (spd : MinSUP A (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc ))
              → supf mc << MinSUP.sup spd
           sc<<d {mc} asc spd = z25 where
                 d1 :  Ordinal
@@ -1417,10 +1504,10 @@
                     z32 = MinSUP.x≤sup spd (fsuc _ (init asc refl))
                     z29 :  (* (cf nmx d1) ≡ * d1) ∨ (* (cf nmx d1) < * d1)
                     z29  with z32
-                    ... | case1 eq1 = case1 (cong (*) (trans (cong (cf nmx) (sym eq)) eq1)  ) 
-                    ... | case2 lt = case2 (subst (λ k → * k < * d1 ) (cong (cf nmx) eq) lt) 
+                    ... | case1 eq1 = case1 (cong (*) (trans (cong (cf nmx) (sym eq)) eq1)  )
+                    ... | case2 lt = case2 (subst (λ k → * k < * d1 ) (cong (cf nmx) eq) lt)
 
-          fsc<<d : {mc z : Ordinal } → (asc : odef A (supf mc)) → (spd : MinSUP A (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc )) 
+          fsc<<d : {mc z : Ordinal } → (asc : odef A (supf mc)) → (spd : MinSUP A (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc ))
              → (fc : FClosure A (cf nmx) (supf mc) z) → z << MinSUP.sup spd
           fsc<<d {mc} {z} asc spd fc = z25 where
                 d1 :  Ordinal
@@ -1440,18 +1527,24 @@
                     z32 = MinSUP.x≤sup spd (fsuc _ fc)
                     z29 :  (* (cf nmx d1) ≡ * d1) ∨ (* (cf nmx d1) < * d1)
                     z29  with z32
-                    ... | case1 eq1 = case1 (cong (*) (trans (cong (cf nmx) (sym eq)) eq1)  ) 
-                    ... | case2 lt = case2 (subst (λ k → * k < * d1 ) (cong (cf nmx) eq) lt) 
+                    ... | case1 eq1 = case1 (cong (*) (trans (cong (cf nmx) (sym eq)) eq1)  )
+                    ... | case2 lt = case2 (subst (λ k → * k < * d1 ) (cong (cf nmx) eq) lt)
 
           smc<<d : supf mc << d
           smc<<d = sc<<d asc spd
 
-          sz<<c : {z : Ordinal } → z o< & A → supf z <= mc 
+          sz<<c : {z : Ordinal } → z o< & A → supf z <= mc
           sz<<c z<A = MinSUP.x≤sup msp1 (ZChain.csupf zc (z09 (ZChain.asupf zc)  ))
 
           sc=c : supf mc ≡ mc
           sc=c = ZChain.sup=u zc (MinSUP.asm msp1) (o<→≤ (∈∧P→o< ⟪ MinSUP.asm msp1 , lift true ⟫ )) ⟪ is-sup , not-hasprev ⟫ where
               not-hasprev :  ¬ HasPrev A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) mc) (cf nmx) mc
+              not-hasprev record { ax = ax ; y = y ; ay = ⟪ ua1 , ch-init fc ⟫ ; x=fy = x=fy } = ⊥-elim ( <-irr z31 z30 ) where
+                   z30 : * mc < * (cf nmx mc)
+                   z30 = proj1 (cf-is-<-monotonic nmx mc (MinSUP.asm msp1))
+                   z31 : ( * (cf nmx mc)  ≡  * mc ) ∨ ( * (cf nmx mc) < * mc )
+                   z31 = <=to≤ ( MinSUP.x≤sup msp1 (subst (λ k →  odef (ZChain.chain zc) (cf nmx k)) (sym x=fy)
+                       ⟪ proj2  (cf-is-≤-monotonic nmx _ (proj2 (cf-is-≤-monotonic nmx _ ua1 ) )) , ch-init (fsuc _ (fsuc _ fc))  ⟫ ))
               not-hasprev record { ax = ax ; y = y ; ay =   ⟪ ua1 , ch-is-sup u u<x is-sup1 fc ⟫; x=fy = x=fy } = ⊥-elim ( <-irr z48 z32 ) where
                    z30 : * mc < * (cf nmx mc)
                    z30 = proj1 (cf-is-<-monotonic nmx mc (MinSUP.asm msp1))
@@ -1466,6 +1559,13 @@
 
 
           not-hasprev : ¬ HasPrev A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 supf d) (cf nmx) d
+          not-hasprev record { ax = ax ; y = y ; ay = ⟪ ua1 , ch-init fc ⟫ ; x=fy = x=fy } = ⊥-elim ( <-irr z31 z30 ) where
+                z30 : * d < * (cf nmx d)
+                z30 = proj1 (cf-is-<-monotonic nmx d (MinSUP.asm spd))
+                z32 : ( cf nmx (cf nmx y)  ≡  supf mc ) ∨ ( * (cf nmx (cf nmx y)) < * (supf mc) )
+                z32 = ZChain.fcy<sup zc (o<→≤ mc<A) (fsuc _ (fsuc _ fc))
+                z31 : ( * (cf nmx d)  ≡  * d ) ∨ ( * (cf nmx d) < * d )
+                z31 = case2 ( subst (λ k → * (cf nmx k) < * d ) (sym x=fy) ( ftrans<=-< z32 ( sc<<d {mc} asc spd ) ))
           not-hasprev record { ax = ax ; y = y ; ay =   ⟪ ua1 , ch-is-sup u u<x is-sup1 fc ⟫; x=fy = x=fy } = ⊥-elim ( <-irr z46 z30 ) where
                 z45 : (* (cf nmx (cf nmx y)) ≡ * d) ∨ (* (cf nmx (cf nmx y)) < * d) → (* (cf nmx d) ≡ * d) ∨ (* (cf nmx d) < * d)
                 z45 p = subst (λ k → (* (cf nmx k) ≡ * d) ∨ (* (cf nmx k) < * d)) (sym x=fy) p
@@ -1474,7 +1574,7 @@
                 z53 : supf u o< supf (& A)
                 z53 = ordtrans<-≤ u<x (ZChain.supf-mono zc (o<→≤ d<A) )
                 z52 : ( u ≡ mc ) ∨  ( u << mc )
-                z52 = MinSUP.x≤sup msp1 ⟪ subst (λ k → odef A k ) (ChainP.supu=u is-sup1) (A∋fcs _ _ (cf-is-≤-monotonic nmx) fc) 
+                z52 = MinSUP.x≤sup msp1 ⟪ subst (λ k → odef A k ) (ChainP.supu=u is-sup1) (A∋fcs _ _ (cf-is-≤-monotonic nmx) fc)
                         , ch-is-sup u z53 is-sup1 (init (ZChain.asupf zc)  (ChainP.supu=u is-sup1)) ⟫
                 z51 : supf u o≤ supf mc
                 z51 = or z52 (λ le → o≤-refl0 (z56 le) ) z57 where
@@ -1483,15 +1583,15 @@
                     z57 : u << mc → supf u o≤ supf mc
                     z57 lt = ZChain.supf-<= zc (case2 z58) where
                         z58 : supf u << supf mc -- supf u o< supf d -- supf u << supf d
-                        z58 = subst₂ ( λ j k → j << k ) (sym (ChainP.supu=u is-sup1)) (sym sc=c)  lt 
+                        z58 = subst₂ ( λ j k → j << k ) (sym (ChainP.supu=u is-sup1)) (sym sc=c)  lt
                 z49 : supf u o< supf mc → (cf nmx (cf nmx y) ≡ supf mc) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) )
                 z49 su<smc = ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) mc<A su<smc (fsuc _ ( fsuc  _ fc ))
                 z50 : (cf nmx (cf nmx y) ≡ supf d) ∨ (* (cf nmx (cf nmx y)) < * (supf d) )
                 z50 = ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) d<A u<x (fsuc _ ( fsuc  _ fc ))
-                z47 : {mc d1 : Ordinal } {asc : odef A (supf mc)} (spd : MinSUP A (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc )) 
-                    → (cf nmx (cf nmx y) ≡ supf mc) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) ) → supf mc << d1 
+                z47 : {mc d1 : Ordinal } {asc : odef A (supf mc)} (spd : MinSUP A (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc ))
+                    → (cf nmx (cf nmx y) ≡ supf mc) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) ) → supf mc << d1
                     →  * (cf nmx (cf nmx y)) < * d1
-                z47 {mc} {d1} {asc} spd (case1 eq) smc<d = subst (λ k → k < * d1 ) (sym (cong (*) eq))  smc<d 
+                z47 {mc} {d1} {asc} spd (case1 eq) smc<d = subst (λ k → k < * d1 ) (sym (cong (*) eq))  smc<d
                 z47 {mc} {d1} {asc} spd (case2 lt) smc<d = IsStrictPartialOrder.trans PO lt smc<d
                 z30 : * d < * (cf nmx d)
                 z30 = proj1 (cf-is-<-monotonic nmx d (MinSUP.asm spd))
@@ -1511,23 +1611,26 @@
                 z23 lt = MinSUP.x≤sup spd lt
                 z22 :  {y : Ordinal} → odef (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) d) y →
                    (y ≡ MinSUP.sup spd) ∨ (y << MinSUP.sup spd)
+                z22 {a} ⟪ aa , ch-init fc ⟫ = case2 ( ( ftrans<=-< z32 ( sc<<d {mc} asc spd ) )) where
+                    z32 : ( a  ≡  supf mc ) ∨ ( * a < * (supf mc) )
+                    z32 = ZChain.fcy<sup zc (o<→≤ mc<A) fc
                 z22 {a} ⟪ aa , ch-is-sup u u<x is-sup1 fc ⟫ = tri u (supf mc)
                        z60 z61 ( λ sc<u → ⊥-elim ( o≤> ( subst (λ k → k o≤ supf mc) (ChainP.supu=u is-sup1) z51) sc<u )) where
                     z53 : supf u o< supf (& A)
                     z53 = ordtrans<-≤ u<x (ZChain.supf-mono zc (o<→≤ d<A) )
                     z52 : ( u ≡ mc ) ∨  ( u << mc )
-                    z52 = MinSUP.x≤sup msp1 ⟪ subst (λ k → odef A k ) (ChainP.supu=u is-sup1) (A∋fcs _ _ (cf-is-≤-monotonic nmx) fc) 
+                    z52 = MinSUP.x≤sup msp1 ⟪ subst (λ k → odef A k ) (ChainP.supu=u is-sup1) (A∋fcs _ _ (cf-is-≤-monotonic nmx) fc)
                            , ch-is-sup u z53 is-sup1 (init (ZChain.asupf zc)  (ChainP.supu=u is-sup1)) ⟫
                     z56 : u ≡ mc → supf u ≡  supf mc
                     z56 eq = cong supf eq
                     z57 : u << mc → supf u o≤ supf mc
                     z57 lt = ZChain.supf-<= zc (case2 z58) where
                         z58 : supf u << supf mc -- supf u o< supf d -- supf u << supf d
-                        z58 = subst₂ ( λ j k → j << k ) (sym (ChainP.supu=u is-sup1)) (sym sc=c)  lt 
+                        z58 = subst₂ ( λ j k → j << k ) (sym (ChainP.supu=u is-sup1)) (sym sc=c)  lt
                     z51 : supf u o≤ supf mc
-                    z51 = or z52 (λ le → o≤-refl0 (z56 le) ) z57 
+                    z51 = or z52 (λ le → o≤-refl0 (z56 le) ) z57
                     z60 : u o< supf mc → (a ≡ d ) ∨ ( * a < * d )
-                    z60 u<smc = case2 ( ftrans<=-< (ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) mc<A 
+                    z60 u<smc = case2 ( ftrans<=-< (ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) mc<A
                         (subst (λ k → k o< supf mc) (sym (ChainP.supu=u is-sup1))  u<smc)  fc ) smc<<d )
                     z61 : u ≡ supf mc → (a ≡ d ) ∨ ( * a < * d )
                     z61 u=sc = case2 (fsc<<d {mc} asc spd (subst (λ k → FClosure A (cf nmx) k a) (trans (ChainP.supu=u is-sup1) u=sc) fc ) )
@@ -1545,27 +1648,27 @@
           sms<sa : supf mc o< supf (& A)
           sms<sa with osuc-≡< ( ZChain.supf-mono zc (o<→≤ ( ∈∧P→o< ⟪ MinSUP.asm msp1 , lift true ⟫) ))
           ... | case2 lt = lt
-          ... | case1 eq = ⊥-elim ( o<¬≡ eq ( ordtrans<-≤ (sc<sd (subst (λ k → supf mc << k ) (sym sd=d) (sc<<d {mc} asc spd)) ) 
+          ... | case1 eq = ⊥-elim ( o<¬≡ eq ( ordtrans<-≤ (sc<sd (subst (λ k → supf mc << k ) (sym sd=d) (sc<<d {mc} asc spd)) )
               ( ZChain.supf-mono zc (o<→≤ d<A ))))
 
           ss<sa : supf c o< supf (& A)
           ss<sa = subst (λ k → supf k o< supf (& A)) (sym c=mc) sms<sa
 
-     zorn00 : Maximal A 
-     zorn00 with is-o∅ ( & HasMaximal )  -- we have no Level (suc n) LEM 
+     zorn00 : Maximal A
+     zorn00 with is-o∅ ( & HasMaximal )  -- we have no Level (suc n) LEM
      ... | no not = record { maximal = ODC.minimal O HasMaximal  (λ eq → not (=od∅→≡o∅ eq)) ; as = zorn01 ; ¬maximal<x  = zorn02 } where
          -- yes we have the maximal
          zorn03 :  odef HasMaximal ( & ( ODC.minimal O HasMaximal  (λ eq → not (=od∅→≡o∅ eq)) ) )
          zorn03 =  ODC.x∋minimal  O HasMaximal  (λ eq → not (=od∅→≡o∅ eq))   -- Axiom of choice
          zorn01 :  A ∋ ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq))
-         zorn01  = proj1  zorn03  
+         zorn01  = proj1  zorn03
          zorn02 : {x : HOD} → A ∋ x → ¬ (ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) < x)
          zorn02 {x} ax m<x = proj2 zorn03 (& x) ax (subst₂ (λ j k → j < k) (sym *iso) (sym *iso) m<x )
      ... | yes ¬Maximal = ⊥-elim ( z04 nmx (SZ (cf nmx) (cf-is-≤-monotonic nmx) as0 (& A) )) where
          -- if we have no maximal, make ZChain, which contradict SUP condition
-         nmx : ¬ Maximal A 
+         nmx : ¬ Maximal A
          nmx mx =  ∅< {HasMaximal} zc5 ( ≡o∅→=od∅  ¬Maximal ) where
-              zc5 : odef A (& (Maximal.maximal mx)) ∧ (( y : Ordinal ) →  odef A y → ¬ (* (& (Maximal.maximal mx)) < * y)) 
+              zc5 : odef A (& (Maximal.maximal mx)) ∧ (( y : Ordinal ) →  odef A y → ¬ (* (& (Maximal.maximal mx)) < * y))
               zc5 = ⟪  Maximal.as mx , (λ y ay mx<y → Maximal.¬maximal<x mx (subst (λ k → odef A k ) (sym &iso) ay) (subst (λ k → k < * y) *iso mx<y) ) ⟫
 
 -- usage (see filter.agda )
@@ -1573,7 +1676,7 @@
 -- _⊆'_ : ( A B : HOD ) → Set n
 -- _⊆'_ A B = (x : Ordinal ) → odef A x → odef B x
 
--- MaximumSubset : {L P : HOD} 
+-- MaximumSubset : {L P : HOD}
 --        → o∅ o< & L →  o∅ o< & P → P ⊆ L
 --        → IsPartialOrderSet P _⊆'_
 --        → ( (B : HOD) → B ⊆ P → IsTotalOrderSet B _⊆'_ → SUP P B _⊆'_ )