view src/zorn.agda @ 879:6222efcf3b04

MinSUP
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Fri, 30 Sep 2022 10:55:23 +0900
parents 1ec8c0cfc892
children d4839adf694d
line wrap: on
line source

{-# OPTIONS --allow-unsolved-metas #-}
open import Level hiding ( suc ; zero )
open import Ordinals
open import Relation.Binary 
open import Relation.Binary.Core
open import Relation.Binary.PropositionalEquality
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 
--     → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B   ) -- SUP condition
--     → 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 Data.Nat hiding ( _<_ ; _≤_ )
open import Data.Nat.Properties 
open import nat 


open inOrdinal O
open OD O
open OD.OD
open ODAxiom odAxiom
import OrdUtil
import ODUtil
open Ordinals.Ordinals  O
open Ordinals.IsOrdinals isOrdinal
open Ordinals.IsNext isNext
open OrdUtil O
open ODUtil O


import ODC

open _∧_
open _∨_
open Bool

open HOD

--
-- Partial Order on HOD ( possibly limited in A )
--

_<<_ : (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 
   ; 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 } } 
 
_≤_ : (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} (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 )

<-ftrans : {x y z : Ordinal } →  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 )

<=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<= (case1 eq) = case1 ( subst₂ (λ j k → j ≡ k ) &iso &iso  (cong (&) eq) )
≤to<= (case2 lt) = case2 lt

<-irr : {a b : HOD} → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
<-irr {a} {b} (case1 a=b) b<a = IsStrictPartialOrder.irrefl PO   (sym a=b) b<a
<-irr {a} {b} (case2 a<b) b<a = IsStrictPartialOrder.irrefl PO   refl
          (IsStrictPartialOrder.trans PO     b<a a<b)

ptrans =  IsStrictPartialOrder.trans PO

open _==_
open _⊆_

-- <-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 = ?

--
-- Closure of ≤-monotonic function f has total order
--

≤-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set (Level.suc n)
≤-monotonic-f A f = (x : Ordinal ) → odef A x → ( * x ≤ * (f x) ) ∧  odef A (f x )

data FClosure (A : HOD) (f : Ordinal → Ordinal ) (s : Ordinal) : Ordinal → Set n where
   init : {s1 : Ordinal } → odef A s → s ≡ s1  → FClosure A f s s1
   fsuc : (x : Ordinal) ( p :  FClosure A f s x ) → FClosure A f s (f x)

A∋fc : {A : HOD} (s : Ordinal) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → odef A y
A∋fc {A} s f mf (init as refl ) = as
A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {A} s  f mf fcy ) )

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 

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 
... | 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 )

fcn : {A : HOD} (s : Ordinal) { x : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) → FClosure A f s x → ℕ
fcn s mf (init as refl) = zero
fcn {A} s {x} {f} mf (fsuc y p) with proj1 (mf y (A∋fc s f mf p))
... | 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) 
     → (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 ()
     fc07 : {x : Ordinal } (cx : FClosure A f s x ) → 0 ≡ fcn s mf cx → * s ≡ * x
     fc07 {x} (init as refl) eq = refl
     fc07 {.(f x)} (fsuc x cx) eq with proj1 (mf x (A∋fc s f mf cx ) )
     ... | case1 x=fx = subst (λ k → * s ≡  k ) x=fx ( fc07 cx eq )
     -- ... | case2 x<fx = ?
     fc00 :  (i j : ℕ ) → i ≡ j  →  {x y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → i ≡ fcn s mf cx  → j ≡ fcn s mf cy → * x ≡ * y
     fc00 (suc i) (suc j) x cx (init x₃ x₄) x₁ x₂ = ⊥-elim ( fc06 x₄ x₂ )
     fc00 (suc i) (suc j) x (init x₃ x₄) (fsuc x₅ cy) x₁ x₂ = ⊥-elim ( fc06 x₄ x₁ )
     fc00 zero zero refl (init _ refl) (init x₁ refl) i=x i=y = refl
     fc00 zero zero refl (init as refl) (fsuc y cy) i=x i=y with proj1 (mf y (A∋fc s f mf cy ) )
     ... | case1 y=fy = subst (λ k → * s ≡ k ) y=fy (fc07 cy i=y) -- ( fc00 zero zero refl (init as refl) cy i=x i=y )
     fc00 zero zero refl (fsuc x cx) (init as refl) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) )
     ... | case1 x=fx = subst (λ k → k ≡ * s ) x=fx (sym (fc07 cx i=x)) -- ( fc00 zero zero refl cx (init as refl) i=x i=y )
     fc00 zero zero refl (fsuc x cx) (fsuc y cy) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) ) | proj1 (mf y (A∋fc s f mf cy ) )
     ... | case1 x=fx  | case1 y=fy = subst₂ (λ j k → j ≡ k ) x=fx y=fy ( fc00 zero zero refl cx cy  i=x i=y )
     fc00 (suc i) (suc j) i=j {.(f x)} {.(f y)} (fsuc x cx) (fsuc y cy) i=x j=y with proj1 (mf x (A∋fc s f mf cx ) ) | proj1 (mf y (A∋fc s f mf cy ) )
     ... | case1 x=fx | case1 y=fy = subst₂ (λ j k → j ≡ k ) x=fx y=fy ( fc00 (suc i) (suc j) i=j cx cy  i=x j=y )
     ... | case1 x=fx | case2 y<fy = subst (λ k → k ≡ * (f y)) x=fx (fc02 x cx i=x) where
          fc02 : (x1 : Ordinal) → (cx1 : FClosure A f s x1 ) →  suc i ≡ fcn s mf cx1 → * x1 ≡ * (f y)
          fc02 x1 (init x₁ x₂) x = ⊥-elim (fc06 x₂ x)
          fc02 .(f x1) (fsuc x1 cx1) i=x1 with proj1 (mf x1 (A∋fc s f mf cx1 ) )
          ... | case1 eq = trans (sym eq) ( fc02  x1 cx1 i=x1 )  -- derefence while f x ≡ x
          ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc04) where
               fc04 : * x1 ≡ * y
               fc04 = fc00 i j (cong pred i=j) cx1 cy (cong pred i=x1) (cong pred j=y)
     ... | case2 x<fx | case1 y=fy = subst (λ k → * (f x) ≡ k ) y=fy (fc03 y cy j=y) where
          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 
          ... | 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)
     ... | case2 x₁ | case2 x₂ = subst₂ (λ j k → * (f j) ≡ * (f k) ) &iso &iso (cong (λ k → * (f (& k))) (fc00 i j (cong pred i=j) cx cy (cong pred i=x) (cong pred j=y)))


fcn-< : {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 Data.Nat.< fcn s mf cy  → * x < * y
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 ()
     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)  ) 
     ... | 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 = 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 
          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) 
    → (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
      fc11 : * x < * y
      fc11 = fcn-< {A} s {x} {y} {f} mf cx cy a
... | 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 
      fc12 : * y < * x
      fc12 = fcn-< {A} s {y} {x} {f} mf cy cx c



-- open import Relation.Binary.Properties.Poset as Poset

IsTotalOrderSet : ( A : HOD ) → Set (Level.suc n)
IsTotalOrderSet A = {a b : HOD} → odef A (& a) → odef A (& b)  → Tri (a < b) (a ≡ b) (b < a )

⊆-IsTotalOrderSet : { A B : HOD } →  B ⊆ A  → IsTotalOrderSet A → IsTotalOrderSet B
⊆-IsTotalOrderSet {A} {B} B⊆A T  ax ay = T (incl B⊆A ax) (incl B⊆A ay)

_⊆'_ : ( A B : HOD ) → Set n
_⊆'_ A B = {x : Ordinal } → odef A x → odef B x

--
-- inductive maxmum tree from x
-- tree structure
--

record HasPrev (A B : HOD) (x : Ordinal ) ( f : Ordinal → Ordinal )  : Set n where
   field
      ax : odef A x
      y : Ordinal
      ay : odef B y
      x=fy :  x ≡ f y 

record IsSup (A B : HOD) {x : Ordinal } (xa : odef A x)     : Set n where
   field
      x<sup   : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x )

record SUP ( A B : HOD )  : Set (Level.suc n) where
   field
      sup : HOD
      as : A ∋ sup
      x<sup : {x : HOD} → B ∋ x → (x ≡ sup ) ∨ (x < sup )   -- B is Total, use positive

--
--  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 ϕ 
--

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 ) 
      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-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 : u o< 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

--
--         f (f ( ... (sup y))) f (f ( ... (sup z1)))
--        /          |         /             |
--       /           |        /              |
--    sup y   <       sup z1          <    sup z2
--           o<                      o<
-- data UChain is total

chain-total : (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal )
   {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 
     ... | 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 = 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 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 
         = 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 
     ... | 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 = 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 with s≤fc (supf ua) f mf fca
          ... | case1 eq =  subst (λ k → * b < k ) eq ct05
          ... | case2 lt =  IsStrictPartialOrder.trans POO ct05 lt

∈∧P→o< :  {A : HOD } {y : Ordinal} → {P : Set n} → odef A y ∧ P → y o< & A
∈∧P→o< {A } {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))

-- Union of supf z which o< x
--
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} {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 )
... | tri> ¬a ¬b y<x with osuc-≡< (supf-mono (o<→≤ y<x) )
... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy )
... | case2 lt = ⊥-elim ( o<> sx<sy lt )

record MinSUP ( A B : HOD )  : Set n where
   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 : Ordinal  } → odef B x → (x ≡ sup1 ) ∨ (x << sup1 ))  → sup o≤ sup1

z09 : {b : Ordinal } { A : HOD } → odef A b → b o< & A
z09 {b} {A} ab = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab))

-- M→S  : {A B : HOD} → MinSUP A B → SUP A B
-- M→S {A} {B} ms = record { sup = MinSUP.sup ms ; as = MinSUP.asm ms ; x<sup = MinSUP.x<sup ms } 

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 
      asupf :  {x : Ordinal } → odef A (supf x)
   chain : HOD
   chain = UnionCF A f mf ay supf z
   chain⊆A : chain ⊆' A
   chain⊆A = λ lt → proj1 lt
   field
      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
      minsup : {x : Ordinal } → x o≤ z  → MinSUP A (UnionCF A f mf ay supf x) 
   sup : {x : Ordinal } → x o≤ z  → SUP A (UnionCF A f mf ay supf x) 
   sup {x} x≤z = ? -- M→S  (minsup x≤z )
   field
      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) b f) → supf b ≡ b 
      supf-is-sup : {x : Ordinal } → (x≤z : x o≤ z) → supf x ≡ & (SUP.sup (sup 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

      -- supf-max : {x : Ordinal } → z o< x → supf x ≡ supf z -- supf z may different from spuf pz
      x≤supfx : z o≤ supf z

   chain∋init : odef chain y
   chain∋init = ⟪ ay , ch-init (init ay refl)    ⟫
   f-next : {a : Ordinal} → odef chain a → odef chain (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 = 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 
               uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
               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
   supf-<= {x} {y} (case2 sx<sy) with trio< (supf x) (supf y)
   ... | tri< a ¬b ¬c = o<→≤ a
   ... | 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} sx<sy with trio< x y
   ... | tri< a ¬b ¬c = a
   ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy )
   ... | tri> ¬a ¬b y<x with osuc-≡< (supf-mono (o<→≤ y<x) )
   ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy )
   ... | case2 lt = ⊥-elim ( o<> sx<sy lt )

   supf∈A : {b : Ordinal} → b o≤ z → odef A (supf b)
   supf∈A {b} b≤z = subst (λ k → odef A k ) (sym (supf-is-sup b≤z)) ( SUP.as ( sup b≤z ) )

   -- supf-idem : {x : Ordinal } → supf x o≤ z  → supf (supf x) ≡ supf x
   -- supf-idem {x} sx≤z = sup=u (supf∈A ?) sx≤z ?

   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 SUP.x<sup (sup 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 (cong (&) eq) (sym (supf-is-sup u≤z ) ) ))
   ... | case2 lt = case2 (subst (λ k → * w < k ) (subst (λ k → k ≡ _ ) *iso (cong (*) (sym (supf-is-sup u≤z ))) ) lt )

   -- ordering is not proved here but in ZChain1 

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 ) →  b o< z  → (ab : odef A b) 
          → HasPrev A (UnionCF A f mf ay supf z) b f ∨  IsSup A (UnionCF A f mf ay supf z) ab  
          → * a < * b  → odef ((UnionCF A f mf ay supf z)) b

record Maximal ( A : HOD )  : Set (Level.suc n) where
   field
      maximal : HOD
      as : A ∋ maximal
      ¬maximal<x : {x : HOD} → A ∋ x  → ¬ maximal < x       -- A is Partial, use negative

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 
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 )) 
     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  )
     as0 =  subst (λ k → odef A k ) &iso as
     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 } 
     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 ) ⟫ )  
     Gtx : { x : HOD} → A ∋ x → HOD
     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) 

     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 )
         ∀-imply-or :  {A : Ordinal  → Set n } {B : Set n }
                        → ((x : Ordinal ) → A x ∨ B) →  ((x : Ordinal ) → A x) ∨ B
         ∀-imply-or  {A} {B} ∀AB with ODC.p∨¬p O ((x : Ordinal ) → A x) -- LEM
         ∀-imply-or  {A} {B} ∀AB | case1 t = case1 t
         ∀-imply-or  {A} {B} ∀AB | case2 x  = case2 (lemma (λ not → x not )) where
               lemma : ¬ ((x : Ordinal ) → A x) →  B
               lemma not with ODC.p∨¬p O B
               lemma not | case1 b = b
               lemma not | case2 ¬b = ⊥-elim  (not (λ x → dont-orb (∀AB x) ¬b ))
         m00 : (x : Ordinal ) → ( ( z : Ordinal) → z o< x →  ¬ (odef A z ∧ xsup z) ) ∨ MinSUP A B
         m00 x = TransFinite0 ind x where
            ind : (x : Ordinal) → ((z : Ordinal) → z o< x → ( ( w : Ordinal) → w o< z →  ¬ (odef A w ∧ xsup w ))  ∨ MinSUP A B)
                  → ( ( w : Ordinal) → w o< x →  ¬ (odef A w ∧ xsup w) )  ∨ MinSUP A B
            ind x prev  =  ∀-imply-or m01 where
                m01 : (z : Ordinal) → (z o< x → ¬ (odef A z ∧ xsup z)) ∨ MinSUP A B
                m01 z with trio< z x
                ... | tri≈ ¬a b ¬c = case1 ( λ lt →  ⊥-elim ( ¬a lt )  )
                ... | tri> ¬a ¬b c = case1 ( λ lt →  ⊥-elim ( ¬a lt )  )
                ... | tri< a ¬b ¬c with prev z a
                ... | case2 mins = case2 mins
                ... | case1 not with ODC.p∨¬p O (odef A z ∧ xsup z)
                ... | case1 mins = case2 record { sup = z ; asm = proj1 mins ; x<sup = proj2 mins ; minsup = m04 } where
                  m04 : {sup1 : Ordinal} → odef A sup1 → ({w : Ordinal} → odef B w → (w ≡ sup1) ∨ (w << sup1)) → z o≤ sup1
                  m04 {s} as lt with trio< z s
                  ... | tri< a ¬b ¬c = o<→≤ a
                  ... | tri≈ ¬a b ¬c = o≤-refl0 b
                  ... | tri> ¬a ¬b s<z = ⊥-elim ( not s s<z ⟪ as , lt ⟫  )
                ... | case2 notz = case1 (λ _ → notz )
         m03 : ¬ ((z : Ordinal) → z o< & A → ¬ odef A z ∧ xsup z)
         m03 not = ⊥-elim ( not s1 (z09 (SUP.as S)) ⟪ SUP.as S , m05 ⟫ ) where
             S : SUP A B
             S = supP B  B⊆A total
             s1 = & (SUP.sup S)
             m05 : {w : Ordinal } → odef B w → (w ≡ s1 ) ∨ (w << s1 )
             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 = dont-or (m00 (& A)) m03

     -- Uncountable ascending chain by axiom of choice
     cf : ¬ Maximal A → Ordinal → Ordinal
     cf  nmx x with ODC.∋-p O A (* x)
     ... | no _ = o∅
     ... | yes ax with is-o∅ (& ( Gtx ax ))
     ... | yes nogt = -- no larger element, so it is maximal
         ⊥-elim (no-maximum (z08 nmx) x ⟪ subst (λ k → odef A k) &iso ax , x-is-maximal nmx (subst (λ k → odef A k ) &iso ax) nogt ⟫ )
     ... | no not =  & (ODC.minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq)))
     is-cf : (nmx : ¬ Maximal A ) → {x : Ordinal} → odef A x → odef A (cf nmx x) ∧ ( * x < * (cf nmx x) )
     is-cf nmx {x} ax with ODC.∋-p O A (* x)
     ... | no not = ⊥-elim ( not (subst (λ k → odef A k ) (sym &iso) ax ))
     ... | yes ax with is-o∅ (& ( Gtx ax ))
     ... | yes nogt = ⊥-elim (no-maximum (z08 nmx) x ⟪ subst (λ k → odef A k) &iso ax , x-is-maximal nmx (subst (λ k → odef A k ) &iso ax) nogt ⟫ )
     ... | no not = ODC.x∋minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq))

     ---
     --- infintie ascention sequence of f
     ---
     cf-is-<-monotonic : (nmx : ¬ Maximal A ) → (x : Ordinal) →  odef A x → ( * x < * (cf nmx x) ) ∧  odef A (cf nmx x )
     cf-is-<-monotonic nmx x ax = ⟪ proj2 (is-cf nmx ax ) , proj1 (is-cf nmx ax ) ⟫
     cf-is-≤-monotonic : (nmx : ¬ Maximal A ) →  ≤-monotonic-f A ( cf nmx )
     cf-is-≤-monotonic nmx x ax = ⟪ case2 (proj1 ( cf-is-<-monotonic nmx x ax  ))  , proj2 ( cf-is-<-monotonic nmx x ax  ) ⟫

     chain-mono :  ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f )  {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal ) 
       {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 {a} {b} {c} a≤b ⟪ ua , ch-init fc  ⟫ = 
            ⟪ ua , ch-init fc  ⟫ 
     chain-mono f mf ay supf {a} {b} {c} a≤b ⟪ uaa ,  ch-is-sup ua ua<x is-sup fc  ⟫ = 
            ⟪ uaa , ch-is-sup ua (ordtrans<-≤ ua<x a≤b ) is-sup fc  ⟫ 

     sp0 : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {x y : Ordinal} (ay : odef A y)  
         (zc : ZChain A f mf ay x ) → SUP A (ZChain.chain zc)
     sp0 f mf ay zc = supP (ZChain.chain zc) (ZChain.chain⊆A zc) ztotal where
        ztotal : IsTotalOrderSet (ZChain.chain zc) 
        ztotal {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 (ZChain.supf zc) ( (proj2 ca)) ( (proj2 cb)) 


     --
     -- Second TransFinite Pass for maximality
     --

     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 = TransFinite { λ x →  ZChain1 A f mf 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) a≤b
        is-max-hp : (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a →
            b o< x → (ab : odef A b) →
            HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) b f → 
            * a < * b → odef (UnionCF A f mf ay (ZChain.supf zc) x) b
        is-max-hp x {a} {b} ua b<x 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))  ⟫ 

        supf = ZChain.supf zc

        csupf-fc : {b s z1 : Ordinal} → b o≤ & A → supf s o< supf b → FClosure A f (supf s) z1 → UnionCF A f mf ay supf b ∋ * z1
        csupf-fc {b} {s} {z1} b≤z ss<sb (init x refl ) = subst (λ k → odef (UnionCF A f mf ay supf b) k ) (sym &iso) zc05  where
                s<b : s o< b
                s<b = ZChain.supf-inject zc ss<sb
                s≤<z : s o≤ & A
                s≤<z = ordtrans s<b b≤z
                zc04 : odef (UnionCF A f mf ay supf (& A))  (supf s)
                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 (zc09 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 : u o≤ supf s  
                    zc08 = subst (λ k → k o≤ supf s) zc06 (ZChain.supf-<= zc (≤to<= ( s≤fc _ f mf fc )))
                    zc09 : u o≤ supf s  →  u o< b 
                    zc09 u≤s with osuc-≡< u≤s
                    ... | case1 u=ss = ZChain.supf-inject zc (subst (λ k → k o< supf b) (sym (trans zc06 u=ss)) ss<sb )
                    ... | case2 u<ss = ordtrans (ZChain.supf-inject zc (subst (λ k → k o< supf s) (sym zc06) u<ss)) s<b
        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-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  ≡  SUP.sup (ZChain.sup zc (o<→≤ b<z) )) ∨ ( * z1  < SUP.sup ( ZChain.sup zc (o<→≤ b<z) ) )
          zc00 = SUP.x<sup (ZChain.sup zc (o<→≤  b<z) ) (csupf-fc (o<→≤ b<z) ss<sb fc )
          -- supf (supf b) ≡ supf b
          zc04 : (z1 ≡ supf b) ∨ (z1 << supf b)
          zc04 with zc00
          ... | case1 eq = case1 (subst₂ (λ j k → j ≡ k ) &iso (sym (ZChain.supf-is-sup zc (o<→≤ b<z))  ) (cong (&) eq) )
          ... | case2 lt = case2 (subst₂ (λ j k → j < k ) refl (subst₂ (λ j k → j ≡ k ) *iso refl (cong (*) (sym (ZChain.supf-is-sup zc (o<→≤ b<z) ) )))  lt )

        zc1 :  (x : Ordinal) →  ((y₁ : Ordinal) → y₁ o< x → ZChain1 A f mf ay zc y₁) → ZChain1 A f mf ay zc x
        zc1 x prev with Oprev-p x
        ... | yes op = record { is-max = is-max } where
           px = Oprev.oprev 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 
           is-max :  {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a →
              b o< x → (ab : odef A b) →
              HasPrev A (UnionCF A f mf ay supf x) b f ∨ IsSup A (UnionCF A f mf ay supf x) 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 b<x ab has-prev a<b 
           is-max {a} {b} ua b<x ab P a<b | case2 is-sup 
             = ⟪ ab , ch-is-sup b b<x m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl)) ⟫   where
              b<A : b o< & A
              b<A = z09 ab
              m04 : ¬ HasPrev A (UnionCF A f mf ay supf b) b f
              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) (chain-mono1 (o<→≤ b<x) 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 
                           → 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 
              m06 = record {  fcy<sup = m08  ; order = m09 ; supu=u = m05 }
        ... | no lim = record { is-max = is-max }  where
           is-max :  {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) b f ∨ IsSup A (UnionCF A f mf ay supf x) 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 b<x ab has-prev a<b 
           is-max {a} {b} ua b<x 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 = chain-mono1 (osucc b<x) 
             ⟪ ab , ch-is-sup b <-osuc 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))
              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 
                       → 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) b f
              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) (chain-mono1 (o<→≤ b<x) lt )} , m04 ⟫    -- ZChain on x
              m06 : ChainP A f mf ay supf b 
              m06 = record { fcy<sup = m07 ;  order = m08 ; supu=u = m05 }

     ---
     --- the maximum chain  has fix point of any ≤-monotonic function
     ---
     fixpoint :  ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f )  (zc : ZChain A f mf as0 (& A) )
            → f (& (SUP.sup (sp0 f mf as0 zc ))) ≡ & (SUP.sup (sp0 f mf as0 zc  ))
     fixpoint f mf zc = z14 where
           chain = ZChain.chain zc
           sp1 = sp0 f mf as0 zc 
           z10 :  {a b : Ordinal } → (ca : odef chain a ) → b o< & A → (ab : odef A b ) 
              →  HasPrev A chain b f ∨  IsSup A chain {b} ab -- (supO  chain  (ZChain.chain⊆A zc) (ZChain.f-total zc) ≡ b )
              → * a < * b  → odef chain b
           z10 = ZChain1.is-max (SZ1 f mf as0 zc (& A) )
           z11 : & (SUP.sup sp1) o< & A
           z11 = c<→o< ( SUP.as sp1)
           z12 : odef chain (& (SUP.sup sp1))
           z12 with o≡? (& s) (& (SUP.sup sp1))
           ... | yes eq = subst (λ k → odef chain k) eq ( ZChain.chain∋init zc )
           ... | no ne = z10 {& s} {& (SUP.sup sp1)} ( ZChain.chain∋init zc ) z11 (SUP.as sp1)
                (case2 z19 ) z13 where
               z13 :  * (& s) < * (& (SUP.sup sp1))
               z13 with SUP.x<sup sp1 ( ZChain.chain∋init zc )
               ... | case1 eq = ⊥-elim ( ne (cong (&) eq) )
               ... | case2 lt = subst₂ (λ j k → j < k ) (sym *iso) (sym *iso) lt
               z19 : IsSup A chain {& (SUP.sup sp1)} (SUP.as sp1)
               z19 = record {   x<sup = z20 }  where
                   z20 :  {y : Ordinal} → odef chain y → (y ≡ & (SUP.sup sp1)) ∨ (y << & (SUP.sup sp1))
                   z20 {y} zy with SUP.x<sup sp1 (subst (λ k → odef chain k ) (sym &iso) zy)
                   ... | case1 y=p = case1 (subst (λ k → k ≡ _ ) &iso ( cong (&) y=p ))
                   ... | case2 y<p = case2 (subst (λ k → * y < k ) (sym *iso) y<p )
                   -- λ {y} zy → subst (λ k → (y ≡ & k ) ∨ (y << & k)) ?  (SUP.x<sup sp1 ? ) }
           ztotal : IsTotalOrderSet (ZChain.chain zc)
           ztotal {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 as0 (ZChain.supf zc) ( (proj2 ca)) ( (proj2 cb)) 

           z14 :  f (& (SUP.sup (sp0 f mf as0 zc ))) ≡ & (SUP.sup (sp0 f mf as0 zc ))
           z14 with ztotal (subst (λ k → odef chain k) (sym &iso)  (ZChain.f-next zc z12 )) z12 
           ... | tri< a ¬b ¬c = ⊥-elim z16 where
               z16 : ⊥
               z16 with proj1 (mf (& ( SUP.sup sp1)) ( SUP.as sp1 ))
               ... | case1 eq = ⊥-elim (¬b (subst₂ (λ j k → j ≡ k ) refl *iso (sym eq) ))
               ... | case2 lt = ⊥-elim (¬c (subst₂ (λ j k → k < j ) refl *iso lt ))
           ... | tri≈ ¬a b ¬c = subst ( λ k → k ≡ & (SUP.sup sp1) ) &iso ( cong (&) b )
           ... | tri> ¬a ¬b c = ⊥-elim z17 where
               z15 : (* (f ( & ( SUP.sup sp1 ))) ≡ SUP.sup sp1) ∨ (* (f ( & ( SUP.sup sp1 ))) <  SUP.sup sp1)
               z15  = SUP.x<sup sp1 (subst (λ k → odef chain k ) (sym &iso) (ZChain.f-next zc z12 ))
               z17 : ⊥
               z17 with z15
               ... | case1 eq = ¬b eq
               ... | case2 lt = ¬a lt

     -- ZChain contradicts ¬ Maximal
     --
     -- ZChain forces fix point on any ≤-monotonic function (fixpoint)
     -- ¬ Maximal create cf which is a <-monotonic function by axiom of choice. This contradicts fix point of ZChain
     --
     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 (SUP.as  sp1 ))))
                                               (subst (λ k → odef A (& k)) (sym *iso) (SUP.as sp1) )
           (case1 ( cong (*)( fixpoint (cf nmx) (cf-is-≤-monotonic nmx ) zc ))) -- x ≡ f x ̄
           (proj1 (cf-is-<-monotonic nmx c (SUP.as sp1 ))) where          -- x < f x
          sp1 : SUP A (ZChain.chain zc)
          sp1 = sp0 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc 
          c = & (SUP.sup sp1)

     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 = 
             λ {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) 
        → 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 
               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) 
       →  SUP A (uchain f mf ay)
     ysup f mf {y} ay = supP (uchain f mf ay) (λ lt → A∋fc y f mf lt)  (utotal f mf ay) 

     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)    }

     record xSUP (B : HOD) (x : Ordinal) : Set n where
        field
           ax : odef A x
           is-sup : IsSup A B ax

     --
     -- create all ZChains under o< x
     --

     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 = zc4 where
          --
          -- we have previous ordinal to use induction
          --
          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 ) 
          px<x : px o< x
          px<x = subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc
          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 

          supf0 = ZChain.supf zc
          pchain  : HOD
          pchain   = UnionCF A f mf ay supf0 px

          -- ¬ supf0 px ≡ px → UnionCF supf0 px ≡ UnionCF supf1 x 
          --       supf1 x ≡ supf0 px
          --   supf0 px ≡ px → ( UnionCF A f mf ay supf0 px ∪ FClosure px ) ≡ UnionCF supf1 x 
          --       supf1 x ≡ sup of ( UnionCF A f mf ay supf0 px ∪ FClosure px (= UnionCF supf1 x))) ≥ supf0 px

          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 
          ... | 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  ⟫ ) 

          -- if previous chain satisfies maximality, we caan reuse it
          --
          --    UninCF supf0 px  previous chain u o< px, supf0 px is not included
          --    UninCF supf0 x   supf0 px is included
          --           supf0 px ≡ px 
          --           supf0 px ≡ supf0 x 

          no-extension : ( (¬ xSUP (UnionCF A f mf ay supf0 px) x ) ∨  HasPrev A pchain x f) → ZChain A f mf ay x
          no-extension P with osuc-≡< (ZChain.x≤supfx zc )
          ... | case1 sfpx=px = ? where
                 pchainpx : HOD
                 pchainpx = record { od = record { def = λ z →  (odef A z ∧ UChain A f mf ay supf0 px z ) ∨ FClosure A f px z  } ; odmax = & A ; <odmax = zc00 } where
                      zc00 : {z : Ordinal } → (odef A z ∧ UChain A f mf ay supf0 px z ) ∨ FClosure A f px z → z o< & A
                      zc00 {z} (case1 lt) = z07 lt 
                      zc00 {z} (case2 fc) = z09 ( A∋fc px f mf fc )
                 zc01 : {z : Ordinal } → (odef A z ∧ UChain A f mf ay supf0 px z ) ∨ FClosure A f px z → odef A z
                 zc01 {z} (case1 lt) = proj1 lt 
                 zc01 {z} (case2 fc) = ( A∋fc px f mf fc )

                 zc02 : { a b : Ordinal } → odef A a ∧ UChain A f mf ay supf0 px a → FClosure A f px b → a <= b
                 zc02 {a} {b} ca fb = zc05 fb where
                    zc06 : & (SUP.sup (ZChain.sup zc o≤-refl)) ≡ px
                    zc06 = trans (sym ( ZChain.supf-is-sup zc o≤-refl )) (sym sfpx=px)
                    zc05 : {b : Ordinal } → FClosure A f px b → a <= b
                    zc05 (fsuc b1 fb ) with proj1 ( mf b1 (A∋fc px f mf fb ))
                    ... | case1 eq = subst (λ k → a <= k ) (subst₂ (λ j k → j ≡ k) &iso &iso (cong (&) eq)) (zc05 fb)
                    ... | case2 lt = <-ftrans (zc05 fb) (case2 lt) 
                    zc05 (init b1 refl) with SUP.x<sup (ZChain.sup 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 (cong (&) eq))
                    ... | case2 lt = case2 (subst (λ k → (* a) < k ) (trans (sym *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 {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
                      eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq )
                 ... | case2 lt = tri< lt1 (λ eq → <-irr (case1 (sym eq)) lt1) (<-irr (case2 lt1)) where
                      lt1 : a0 < b0
                      lt1 = subst₂ (λ j k → j < k ) *iso *iso lt
                 ptotal {b0} {a0} (case2 b) (case1 a) with zc02 a b
                 ... | case1 eq = tri≈ (<-irr (case1 eq1)) (sym eq1) (<-irr (case1 (sym eq1))) where
                      eq1 : a0 ≡ b0
                      eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq )
                 ... | case2 lt = tri> (<-irr (case2 lt1)) (λ eq → <-irr (case1 eq) lt1) lt1  where
                      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 px f mf a b)

                 sup1 : SUP A pchainpx 
                 sup1 = supP pchainpx zc01 ptotal

                 sp1 : Ordinal
                 sp1 = & (SUP.sup sup1 )

                 supf1 : Ordinal → Ordinal
                 supf1 z with trio< z px 
                 ... | tri< a ¬b ¬c = supf0 z
                 ... | tri≈ ¬a b ¬c = px   --- supf px ≡ px
                 ... | tri> ¬a ¬b c = sp1 

                 pchainp : HOD
                 pchainp = UnionCF A f mf ay supf1  x

                 zc16 : {z : Ordinal } → z o< px → supf1 z ≡ supf0 z
                 zc16 {z} z<px with trio< z px
                 ... | tri< a ¬b ¬c = refl
                 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬a z<px )
                 ... | tri> ¬a ¬b c = ⊥-elim ( ¬a z<px )

                 supf-mono1 : {z w : Ordinal } → z o≤ w → supf1 z o≤ supf1 w
                 supf-mono1 {z} {w} z≤w with trio< w px 
                 ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (zc16 (ordtrans≤-< z≤w a))) refl ( supf-mono z≤w )
                 ... | tri≈ ¬a refl ¬c with osuc-≡< z≤w 
                 ... | case1 refl = o≤-refl0 zc17 where
                       zc17 : supf1 px ≡ px
                       zc17 with trio< px px
                       ... | tri< a ¬b ¬c = ⊥-elim ( ¬b refl )
                       ... | tri≈ ¬a b ¬c = refl
                       ... | tri> ¬a ¬b c = ⊥-elim ( ¬b refl )
                 ... | case2 lt = subst₂ (λ j k → j o≤ k ) (sym (zc16 lt)) (sym sfpx=px) ( supf-mono z≤w )
                 supf-mono1 {z} {w} z≤w | tri> ¬a ¬b c with trio< z px
                 ... | tri< a ¬b ¬c = zc19 where
                       zc19 : supf0 z o≤ sp1
                       zc19 = ?
                 ... | tri≈ ¬a b ¬c = zc18 where
                       zc18 : px o≤ sp1
                       zc18 = ?
                 ... | tri> ¬a ¬b c = o≤-refl

                 zc10 :  {z : Ordinal} → OD.def (od pchain) z → OD.def (od pchainp) 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 (ordtrans u1<x (pxo<x op)) zc15  zc14  ⟫ where
                    zc14 : FClosure A f (supf1 u1) z
                    zc14 = subst (λ k → FClosure A f k z) (sym (zc16 u1<x)) fc
                    zc15 : ChainP A f mf ay supf1 u1
                    zc15 = record { fcy<sup = λ {z} fcy → subst (λ k →  (z ≡ k) ∨ ( z << k ) ) (sym (zc16 u1<x)) (ChainP.fcy<sup u1-is-sup fcy)
                       ; order = λ {s} {z1} lt fc1 → subst (λ k → (z1 ≡ k) ∨ ( z1 << k ) ) (sym (zc16 u1<x)) (
                           ChainP.order u1-is-sup (subst₂ (λ j k → j o< k) (zc17 u1<x lt) (zc16 u1<x) lt) (subst (λ k → FClosure A f k z1) (zc17 u1<x lt) fc1) )
                       ; supu=u = trans (zc16 u1<x) (ChainP.supu=u u1-is-sup) } where
                               zc17 : {s u : Ordinal } → u o< px → supf1 s o< supf1 u → supf1 s ≡ supf0 s
                               zc17 u1<x lt = zc16 (ordtrans ( supf-inject0 supf-mono1 lt) u1<x) 

                 zc11 :  {z : Ordinal} → z o< px  → OD.def (od pchainp) z → OD.def (od pchain) z
                 zc11 {z} z<px ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫ 
                 zc11 {z} z<px ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ = ⟪ az , ch-is-sup u1 (ordtrans u1<x ?) ?  ?  ⟫

                 zc12 :  {z : Ordinal} → OD.def (od pchainp) z → OD.def (od pchainpx) z
                 zc12 {z} ⟪ az , ch-init fc ⟫ = case1 ⟪ az , ch-init fc ⟫ 
                 zc12 {z} ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ = ?

                 zc13 :  {z : Ordinal} → OD.def (od pchainpx) z → OD.def (od pchainp) z
                 zc13 {z} (case1 ⟪ az , ch-init fc ⟫ ) = ⟪ az , ch-init fc ⟫ 
                 zc13 {z} (case1 ⟪ az , ch-is-sup u1 u1<x u1-is-sup fc ⟫ ) = ⟪ az , ch-is-sup u1 (ordtrans u1<x (pxo<x op)) ?  ?  ⟫
                 zc13 {z} (case2 fc) = ⟪ ? , ch-is-sup ? ? ?  ?  ⟫

                 record STMP {z : Ordinal} (z≤x : z o≤ x ) : Set (Level.suc n) where
                     field
                         tsup : SUP A (UnionCF A f mf ay supf1 z)
                         tsup=sup : supf1 z ≡ & (SUP.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 = ? ; tsup=sup = ? }
                 ... | tri≈ ¬a b ¬c = record { tsup = ?  ; tsup=sup = ? }
                 ... | tri> ¬a ¬b px<z = record { tsup = ? ; tsup=sup = ? }

                 csupf1 : {b : Ordinal } → supf1 b o< x → odef (UnionCF A f mf ay supf1 x) (supf1 b) 
                 csupf1 {b} sb<z = ⟪ asb , ch-is-sup (supf1 b) sb<z 
                     record { fcy<sup = ?  ; order = order ; supu=u = ? }  
                        (init ? ? ) ⟫ where
                    b<x : b o< x
                    b<x = ZChain.supf-inject zc ?
                    asb : odef A (supf1 b)
                    asb = ? --  ZChain.supf∈A zc ? -- (o<→≤ b<x)
                    supb : SUP A (UnionCF A f mf ay supf1 (supf1 b))
                    supb = ? -- ZChain.sup zc (o<→≤ sb<z)
                    supb-is-b : supf1 (supf1 b) ≡ & (SUP.sup supb)
                    supb-is-b = ? -- ZChain.supf-is-sup zc ? -- (o<→≤ sb<z)
                    order : {s z1 : Ordinal} → supf1 s o< supf1 (supf1 b) →
                        FClosure A f (supf1 s) z1 → (z1 ≡ supf1 (supf1 b)) ∨ (z1 << supf1 (supf1 b))
                    order {s} {z1} ss<ssb fs with SUP.x<sup supb ?
                    ... | case1 eq = ?
                    ... | case2 lt = ?

          ... | case2 px<spfx = ? where
            -- record { supf = supf0 ;  asupf = ZChain.asupf zc ; sup = λ lt → STMP.tsup (sup lt ) ; supf-mono = supf-mono 
            --              ; supf-< = ? ; sup=u = sup=u ; supf-is-sup = λ lt → STMP.tsup=sup (sup lt) }  where
                 supf1 : Ordinal → Ordinal
                 supf1 z with trio< z px 
                 ... | tri< a ¬b ¬c = supf0 z
                 ... | tri≈ ¬a b ¬c = supf0 px
                 ... | tri> ¬a ¬b c = supf0 px

                 zc16 : {z : Ordinal } → z o< px → supf1 z ≡ supf0 z
                 zc16 {z} z<px with trio< z px
                 ... | tri< a ¬b ¬c = refl
                 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬a z<px )
                 ... | tri> ¬a ¬b c = ⊥-elim ( ¬a z<px )

                 zc17 : {z : Ordinal } → supf0 z o≤ supf0 px
                 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 
                 ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (zc16 (ordtrans≤-< z≤w a))) refl ( supf-mono z≤w )
                 ... | tri≈ ¬a refl ¬c with trio< z px
                 ... | tri< a ¬b ¬c = zc17
                 ... | tri≈ ¬a refl ¬c = o≤-refl
                 ... | tri> ¬a ¬b c = o≤-refl
                 supf-mono1 {z} {w} z≤w | tri> ¬a ¬b c with trio< z px
                 ... | tri< a ¬b ¬c = zc17
                 ... | tri≈ ¬a b ¬c = o≤-refl
                 ... | tri> ¬a ¬b c = o≤-refl

                 pchain1 : HOD
                 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 (ordtrans u1<x (pxo<x op)) ?  ?  ⟫

                 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 (ordtrans u1<x ?) ?  ?  ⟫

                 zc11 : (¬ xSUP (UnionCF A f mf ay supf0 px) x ) ∨  (HasPrev A pchain x f )
                        → {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-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 u1<px ? 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 = 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 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
                        ... | tri> ¬a ¬b c = ⊥-elim ( o<> u1<x c )
                        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 ) ) 
                        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) zc13 ))
                        ... | case1 eq1 = ⊥-elim ( o<¬≡ zc14 (o<→≤ u<x) ) where
                                 zc14 : u ≡ osuc px
                                 zc14 = begin
                                    u ≡⟨ sym ( ChainP.supu=u is-sup) ⟩ 
                                    supf0 u ≡⟨ eq1 ⟩ 
                                    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) x 
                        zcsup = record { ax =  subst (λ k → odef A k) (trans zc12 eq) (ZChain.asupf zc)
                                 ; is-sup = record { x<sup = x<sup } }
                 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 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
                        ... | tri> ¬a ¬b c = ⊥-elim ( o<> u1<x c )
                        zc20 : {z : Ordinal} → FClosure A f (supf0 u1) z → OD.def (od pchain) z
                        zc20 {z} (init asu su=z ) = zc13 where
                          zc14 : x ≡ z
                          zc14 = begin
                             x ≡⟨ sym eq ⟩
                             u1 ≡⟨  sym ? ⟩
                             supf0 u1 ≡⟨ su=z ⟩
                             z ∎ where open ≡-Reasoning
                          zc13 : odef pchain z
                          zc13 = subst (λ k → odef pchain k) (trans (sym (HasPrev.x=fy hp)) zc14) ( ZChain.f-next zc (HasPrev.ay hp) )
                        zc20 {.(f w)} (fsuc w fc) = ZChain.f-next zc (zc20 fc)
                 record STMP {z : Ordinal} (z≤x : z o≤ x ) : Set (Level.suc n) where
                     field
                         tsup : SUP A (UnionCF A f mf ay supf0 z)
                         tsup=sup : supf0 z ≡ & (SUP.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 = ZChain.sup zc (o<→≤ a)  ; tsup=sup = ZChain.supf-is-sup zc (o<→≤ a) }
                 ... | tri≈ ¬a b ¬c = record { tsup = ZChain.sup zc (o≤-refl0 b)  ; tsup=sup = ZChain.supf-is-sup zc (o≤-refl0 b) }
                 ... | tri> ¬a ¬b px<z = zc35 where
                     zc30 : z ≡ x
                     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 
                     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 P ? 
                     ... | 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-sup 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  } 
                     zc35 : STMP z≤x
                     zc35 with trio< (supf0 px) px
                     ... | tri< a ¬b ¬c = zc36 ¬b
                     ... | tri> ¬a ¬b c = zc36 ¬b
                     ... | tri≈ ¬a b ¬c = record { tsup = zc37 ; tsup=sup = ?  } where
                          zc37 : SUP A (UnionCF A f mf ay supf0 z)
                          zc37 = record { sup = ? ; as = ? ; x<sup = ? }
                 sup=u : {b : Ordinal} (ab : odef A b) →
                    b o≤ x → IsSup A (UnionCF A f mf ay supf0 b) ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf0 b) b f ) → 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 → IsSup.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 → IsSup.x<sup (proj1 is-sup) lt } , proj2 is-sup ⟫ 
                 ... | tri> ¬a ¬b px<b = zc31 P 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) x 
                     zcsup with zc30
                     ... | refl = record { ax = ab ; is-sup = record { x<sup = λ {w} lt → 
                        IsSup.x<sup (proj1 is-sup) ?} } 
                     zc31 : ( (¬ xSUP (UnionCF A f mf ay supf0 px) x ) ∨  HasPrev A (UnionCF A f mf ay supf0 px) x f) → 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 } ) 

          zc4 : ZChain A f mf ay x 
          zc4 with ODC.∋-p O A (* x)
          ... | no noax = no-extension (case1 ( λ s → noax (subst (λ k → odef A k ) (sym &iso) (xSUP.ax s) ))) -- ¬ A ∋ p, just skip
          ... | yes ax with ODC.p∨¬p O ( HasPrev A (ZChain.chain zc ) x f )   
               -- we have to check adding x preserve is-max ZChain A y f mf x
          ... | case1 pr = no-extension (case2 pr) -- we have previous A ∋ z < x , f z ≡ x, so chain ∋ f z ≡ x because of f-next
          ... | case2 ¬fy<x with ODC.p∨¬p O (IsSup A (ZChain.chain zc ) ax )
          ... | case1 is-sup = ? -- x is a sup of zc 
          ... | case2 ¬x=sup =  no-extension (case1 nsup) where -- px is not f y' nor sup of former ZChain from y -- no extention
             nsup : ¬ xSUP (UnionCF A f mf ay supf0 px) x 
             nsup s = ¬x=sup z12 where
                z12 : IsSup A (UnionCF A f mf ay supf0 px) ax
                z12 = record { x<sup = λ {z} lt → subst (λ k → (z ≡ k) ∨ (z << k )) (sym &iso) ( IsSup.x<sup ( xSUP.is-sup s ) lt ) }

     ... | no lim = zc5 where

          pzc : (z : Ordinal) → z o< x → ZChain A f mf ay z
          pzc z z<x = prev z z<x

          ysp =  & (SUP.sup (ysup f mf ay))

          supf0 : Ordinal → Ordinal
          supf0 z with trio< z x
          ... | tri< a ¬b ¬c = ZChain.supf (pzc (osuc z) (ob<x lim a)) z
          ... | tri≈ ¬a b ¬c = ysp
          ... | tri> ¬a ¬b c = ysp

          pchain : HOD
          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 
               uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) )
               uz01 = chain-total A f mf ay supf0 ( (proj2 ca)) ( (proj2 cb)) 
  
          usup : SUP A pchain
          usup = supP pchain (λ lt → proj1 lt) ptotal0
          spu = & (SUP.sup usup)

          supf1 : Ordinal → Ordinal
          supf1 z with trio< z x
          ... | tri< a ¬b ¬c = ZChain.supf (pzc (osuc z) (ob<x lim a)) z
          ... | tri≈ ¬a b ¬c = spu
          ... | tri> ¬a ¬b c = spu

          pchain1 : HOD
          pchain1 = UnionCF A f mf ay supf1 x

          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) b f → 
                * 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-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))  ⟫ 

          zc70 : HasPrev A pchain x f → ¬ xSUP pchain x 
          zc70 pr xsup = ?

          no-extension : ¬ ( xSUP (UnionCF A f mf ay supf0 x) x ) → ZChain A f mf ay x
          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-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 u≤x ? (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-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 u<x ? (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 )) 
                 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-sup (pzc z a) {!!}
                 ... | tri≈ ¬a b ¬c = {!!}
                 ... | tri> ¬a ¬b x<z = ⊥-elim (o<¬≡ refl (ordtrans<-≤ x<z z≤x )) 
                 sup=u : {b : Ordinal} (ab : odef A b) → b o≤ x → IsSup A (UnionCF A f mf ay supf1 b) ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf1 b) b f ) → 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 )) 

          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 x f )   
               -- we have to check adding x preserve is-max ZChain A y f mf x
          ... | case1 pr = no-extension {!!} 
          ... | case2 ¬fy<x with ODC.p∨¬p O (IsSup A pchain ax )
          ... | 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
         
     SZ : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) → {y : Ordinal} (ay : odef A y) → ZChain A f mf ay   (& A)
     SZ f mf {y} ay = TransFinite {λ z → ZChain A f mf ay   z  } (λ x → ind f mf ay x   ) (& A)

     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  
         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 zorn04 ) where
         -- if we have no maximal, make ZChain, which contradict SUP condition
         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 = ⟪  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) ) ⟫
         zorn04 : ZChain A (cf nmx) (cf-is-≤-monotonic nmx) as0 (& A)
         zorn04 = SZ (cf nmx) (cf-is-≤-monotonic nmx) (subst (λ k → odef A k ) &iso as ) 

-- usage (see filter.agda )
--
-- _⊆'_ : ( A B : HOD ) → Set n
-- _⊆'_ A B = (x : Ordinal ) → odef A x → odef B x

-- MaximumSubset : {L P : HOD} 
--        → o∅ o< & L →  o∅ o< & P → P ⊆ L
--        → IsPartialOrderSet P _⊆'_
--        → ( (B : HOD) → B ⊆ P → IsTotalOrderSet B _⊆'_ → SUP P B _⊆'_ )
--        → Maximal P (_⊆'_)
-- MaximumSubset {L} {P} 0<L 0<P P⊆L PO SP  = Zorn-lemma {P} {_⊆'_} 0<P PO SP