changeset 938:93a49ffa9183

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Fri, 28 Oct 2022 18:37:05 +0900
parents 3a511519bd10
children 187594116449
files src/zorn.agda src/zorn1.agda
diffstat 2 files changed, 45 insertions(+), 818 deletions(-) [+]
line wrap: on
line diff
--- a/src/zorn.agda	Mon Oct 24 19:11:19 2022 +0900
+++ b/src/zorn.agda	Fri Oct 28 18:37:05 2022 +0900
@@ -278,7 +278,7 @@
 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
+     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
@@ -286,14 +286,14 @@
      ... | 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
+     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
+     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
@@ -301,7 +301,7 @@
      ... | 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
+     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 
@@ -424,11 +424,11 @@
    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 ) ⟫
+   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
+   ... | 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
@@ -590,8 +590,8 @@
             → * 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))  ⟫ 
+        ... | ⟪ 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
 
@@ -606,7 +606,7 @@
                 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
+                ... | ⟪ 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
@@ -617,7 +617,7 @@
                 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) ⟫ 
+                ... | ⟪ 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) ) )
@@ -889,7 +889,7 @@
                         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-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₁) ⟫ 
+                        ... | ⟪ 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 
                              ?  
@@ -928,7 +928,7 @@
                         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-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₁) ⟫ 
+                        ... | ⟪ 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  osuc-≡< ( subst (λ k → supf0 px o< k ) (sym (Oprev.oprev=x op)) sfpx<x )
                         ... | case1 sfpx=px =  ⟪ asp , ch-is-sup px ? -- (pxo<x op)
                                      record {fcy<sup = zc13 ; order = zc17 ; supu=u = zc15 } zc14 ⟫ where
@@ -956,7 +956,7 @@
 
                         ... | case2 sfp<px  with ZChain.csupf zc sfp<px --  odef (UnionCF A f mf ay supf0 px) (supf0 px)
                         ... | ⟪ ua1 , ch-init fc₁ ⟫ = ⟪ ua1 , ch-init fc₁ ⟫ 
-                        ... | ⟪ ua1 , ch-is-sup u u≤x is-sup fc₁ ⟫ = ⟪ ua1 , ch-is-sup u ?  
+                        ... | ⟪ ua1 , ch-is-sup u u<x is-sup fc₁ ⟫ = ⟪ ua1 , ch-is-sup u ?  
                               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} fc = subst (λ k → (z ≡ k) ∨ ( z << k )) (sym (sf1=sf0 ? )) (ChainP.fcy<sup is-sup fc)
@@ -967,7 +967,7 @@
                                        s<u : s o< u
                                        s<u = supf-inject0 supf1-mono lt
                                        s≤px : s o≤ px
-                                       s≤px = ordtrans s<u ? -- (o<→≤ u≤x)
+                                       s≤px = ordtrans s<u ? -- (o<→≤ u<x)
                                        lt0 : supf0 s o< supf0 u
                                        lt0 = subst₂ (λ j k → j o< k ) (sf1=sf0 s≤px) (sf1=sf0 ? ) lt
                                  z12 : supf1 u ≡ u
@@ -1111,7 +1111,7 @@
                         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) ? ))
+                        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
@@ -1235,9 +1235,9 @@
                 * 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 , 
+          ... | ⟪ 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 x f → ¬ xSUP pchain x 
           zc70 pr xsup = ?
@@ -1251,23 +1251,23 @@
                  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
+                     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₁) ⟫ 
+                          ... | ⟪ 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
+                     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₁) ⟫ 
+                          ... | ⟪ 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 = ?
-                          ... | tri> ¬a ¬b c = ? -- ⊥-elim ( o≤> u≤x 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) {!!}) {!!} )
@@ -1469,20 +1469,38 @@
              → supf mc << MinSUP.sup spd
           sc<<d {mc} {asc} spd = z25 where
                 d1 :  Ordinal
-                d1 = MinSUP.sup spd
+                d1 = MinSUP.sup spd -- supf d1 ≡ d
                 z24 : (supf mc ≡ d1) ∨ ( supf mc << d1 )
                 z24 = MinSUP.x<sup spd (init asc refl)
-                z26 :  odef (ZChain.chain zc) (supf mc)
-                z26 = ?
                 z28 : supf mc o< & A
                 z28 = z09 (ZChain.asupf zc)
                 z25 : supf mc << d1
                 z25 with z24
                 ... | case2 lt = lt
-                ... | case1 eq = ? where
+                ... | case1 eq = ⊥-elim ( <-irr z29 (proj1 (cf-is-<-monotonic nmx d1 (MinSUP.asm spd)) ) )  where
+                    --  supf mc ≡ d1
                     z27 :  odef (ZChain.chain zc) (cf nmx d1)
                     z27 = ZChain.f-next zc (subst (λ k → odef (ZChain.chain zc) k ) eq (ZChain.csupf zc z28))
-
+                    z31 : {z w : Ordinal } → odef (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc) z 
+                        → (* w ≡ * z) ∨ (* w < * z)
+                        → (* w ≡ * d1) ∨ (* w < * d1)
+                    z31 {z} uz (case1 w=z) with MinSUP.x<sup spd uz
+                    ... | case1 eq = case1 (trans w=z (cong (*) eq) )
+                    ... | case2 lt = case2 (subst (λ k → k < * d1 ) (sym w=z) lt )
+                    z31 {z} {w} uz (case2 w<z) with MinSUP.x<sup spd uz
+                    ... | case1 eq = case2 (subst (λ k → * w < k ) (cong (*) eq) w<z )
+                    ... | case2 lt = case2 (  IsStrictPartialOrder.trans PO w<z lt)
+                    z29 :  (* (cf nmx d1) ≡ * d1) ∨ (* (cf nmx d1) < * d1)
+                    z29 with z27
+                    ... | ⟪ aa , ch-init fc ⟫ = ? where
+                          z30 : FClosure A (cf nmx) (& s) (cf nmx d1)
+                          z30 = fc
+                    ... | ⟪ aa , ch-is-sup u u<x is-sup fc ⟫ with trio< u (supf mc) -- u<x : supf u o< supf (& A)
+                    ... | tri< a ¬b ¬c = ?  -- u o< supf mc 
+                    ... | tri> ¬a ¬b c = ?  -- supf mc o< u
+                    ... | tri≈ ¬a b ¬c with MinSUP.x<sup spd ( subst₂ (λ j k → FClosure A (cf nmx) j k ) (trans (ChainP.supu=u is-sup) b) refl fc )
+                    ... | case1 eq = case1 (cong (*) eq)
+                    ... | case2 lt = case2 lt 
           sc<sd : {mc d : Ordinal } → supf mc << supf d → supf mc o< supf d
           sc<sd {mc} {d} sc<<sd with osuc-≡< ( ZChain.supf-<= zc (case2 sc<<sd ) )
           ... | case1 eq = ⊥-elim ( <-irr (case1 (cong (*) (sym eq) )) sc<<sd )
--- a/src/zorn1.agda	Mon Oct 24 19:11:19 2022 +0900
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,791 +0,0 @@
-{-# 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 zorn1 {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 : 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
-
---
---         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 : 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) 
-        ; 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 
-   ... | 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 ) 
-   (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  ⟫ 
-
-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) b f) → 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
-
-      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
-
-   chain : HOD
-   chain = UnionCF A f mf ay supf z
-   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) 
-
-   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 = 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 )
-
-   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 )
-
-   -- 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 ) → supf b o< supf 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  ) ⟫
-
-     --
-     -- 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 = ?
-
-     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) 
-       →  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) 
-
-     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 = ?
-         
-     ---
-     --- the maximum chain  has fix point of any ≤-monotonic function
-     ---
-
-     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
-
-     data ZChainP ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y) 
-            ( supf : Ordinal → Ordinal )  (z : Ordinal) : Set n where
-          zchain : (uz : Ordinal ) → odef (UnionCF A f mf ay supf uz) z → ZChainP f mf ay supf z
-     
-     auzc :  ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y)  
-            (supf : Ordinal → Ordinal )  → {x : Ordinal } → ZChainP f mf ay supf x → odef A x
-     auzc f mf {y} ay supf {x} (zchain uz ucf) = proj1 ucf
-
-     zp-uz : ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y)  
-            (supf : Ordinal → Ordinal )  → {x : Ordinal } → ZChainP f mf ay supf x → Ordinal
-     zp-uz f mf ay supf (zchain uz _) = uz
-
-     uzc⊆zc :  ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y)  
-            (supf : Ordinal → Ordinal )  → {x : Ordinal } → (zp : ZChainP f mf ay supf x ) → UChain A f mf ay supf (zp-uz f mf ay supf zp) x 
-     uzc⊆zc f mf {y} ay supf {x} (zchain uz ⟪ ua , ch-init fc ⟫) = ch-init fc 
-     uzc⊆zc f mf {y} ay supf {x} (zchain uz ⟪ ua , ch-is-sup u u<x is-sup fc ⟫) with ChainP.supu=u is-sup
-     ... | eq = ch-is-sup u u<x is-sup fc 
-
-     UnionZF : ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y)  
-            (supf : Ordinal → Ordinal )  → HOD
-     UnionZF f mf {y} ay supf = record { od = record { def = λ x → ZChainP f mf ay supf x } 
-         ; odmax = & A ; <odmax = λ lt →  ∈∧P→o< ⟪ auzc f mf ay supf lt , lift true ⟫ }
-     
-     uzctotal : ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y)  
-         → ( supf : Ordinal → Ordinal )
-         → IsTotalOrderSet (UnionZF f mf ay supf )
-     uzctotal f mf ay supf {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (uz01 ca cb) where 
-          uz01 : {ua ub : Ordinal } → ZChainP f mf ay supf ua → ZChainP f mf ay supf ub 
-             → Tri (* ua < * ub) (* ua ≡ * ub) (* ub < * ua )
-          uz01 {ua} {ub} (zchain uza uca) (zchain uzb ucb) = chain-total A f mf ay supf (proj2 uca) (proj2 ucb)
-
-     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) 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)) 
-
-     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 (UnionCF A f mf ay (ZChain.supf zc) x)
-     sp0 f mf ay zc = M→S (ZChain.supf zc) (msp0 f mf ay zc )
-
-     fixpoint :  ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f )  (zc : ZChain A f mf as0 (& A) )
-            → ZChain.supf zc (& (SUP.sup (sp0 f mf as0 zc))) o< ZChain.supf zc (& A)
-            → f (& (SUP.sup (sp0 f mf as0 zc ))) ≡ & (SUP.sup (sp0 f mf as0 zc  ))
-     fixpoint f mf zc ss<sa = ?
-
-
-     -- 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 ss<sa ))) -- x ≡ f x ̄
-           (proj1 (cf-is-<-monotonic nmx c (SUP.as sp1 ))) where          -- x < f x
-          supf = ZChain.supf zc
-          msp1 : MinSUP A (ZChain.chain zc)
-          msp1 = msp0 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc 
-          sp1 : SUP A (ZChain.chain zc)
-          sp1 = sp0 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc 
-          c : Ordinal 
-          c = & ( SUP.sup sp1 )
-          mc = MinSUP.sup msp1
-          c=mc : c ≡ mc
-          c=mc = &iso
-          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 = 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) 
-          sd=ms : supf d ≡ MinSUP.sup ( ZChain.minsup zc (o<→≤ d<A) )
-          sd=ms = ZChain.supf-is-minsup zc (o<→≤ d<A)
-          -- z26 : {x : Ordinal } → odef (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) d) x
-          --     → odef (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) c) x ∨ odef (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc ) x 
-          -- z26 = ?
-          is-sup : IsSup A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) d) (MinSUP.asm spd)
-          is-sup = record { x<sup = z22 } where
-               z23 : {z : Ordinal } → odef (uchain (cf nmx) (cf-is-≤-monotonic nmx) asc) z → (z ≡ MinSUP.sup spd) ∨ (z << MinSUP.sup spd)
-               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 ⟫ = ?
-               z22 {a} ⟪ aa , ch-is-sup u u<x is-sup fc ⟫ = ?
-                    -- u<x    : ZChain.supf zc u o< ZChain.supf zc d
-                    --     supf u o< spuf c → order
-          not-hasprev : ¬ HasPrev A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) d) d (cf nmx)
-          not-hasprev hp = ? where
-               y : Ordinal
-               y = HasPrev.y hp
-               z24 : y << d
-               z24 = subst (λ k → y <<  k) (sym (HasPrev.x=fy hp)) ( proj1 (cf-is-<-monotonic nmx y (proj1 (HasPrev.ay hp) ) ))
-               -- z26 : {x : Ordinal } → odef (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc ) x → (x ≡ d ) ∨ (x << d )  
-               -- z26 lt with MinSUP.x<sup spd (subst (λ k → odef _ k ) ? lt)
-               -- ... | case1 eq = ?
-               -- ... | case2 lt = ?
-               -- z25 : {x : Ordinal } → odef (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc ) x → (x ≡ y ) ∨ (x << y )  
-               -- z25 {x} (init au eq ) = ?   -- sup c = x, cf y ≡ d, sup c =< d
-               -- z25  (fsuc x lt) = ?        -- cf (sup c) 
-
-          sd=d : supf d ≡ d
-          sd=d = ZChain.sup=u zc (MinSUP.asm spd) (o<→≤ d<A) ⟪ is-sup , not-hasprev ⟫
-
-          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
-                d1 = MinSUP.sup spd
-                z24 : (supf mc ≡ d1) ∨ ( supf mc << d1 )
-                z24 = MinSUP.x<sup spd (init asc refl)
-                z25 : supf mc << d1
-                z25 with z24
-                ... | case2 lt = lt
-                ... | case1 eq = ?
-
-          sc<sd : {mc d : Ordinal } → supf mc << supf d → supf mc o< supf d
-          sc<sd {mc} {d} sc<<sd with osuc-≡< ( ZChain.supf-<= zc (case2 sc<<sd ) )
-          ... | case1 eq = ⊥-elim ( <-irr (case1 (cong (*) (sym eq) )) sc<<sd )
-          ... | case2 lt = lt
-
-          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)) ) 
-              ( 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 
-     ... | 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 (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 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) ) ⟫
-
--- 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