view src/zorn.agda @ 704:01a88eeb9d00

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Tue, 12 Jul 2022 23:05:31 +0900
parents 0278f0d151f2
children 799963f352d6
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    -- Set n order
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 )

<-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 _⊆_

--
-- 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 : odef A s → FClosure A f s s
   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) = as
A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {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) = 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) = 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
     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 zero zero refl (init _) (init x₁) i=x i=y = refl
     fc00 zero zero refl (init as) (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 ( fc00 zero zero refl (init as) cy i=x i=y )
     fc00 zero zero refl (fsuc x cx) (init as) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) )
     ... | case1 x=fx = subst (λ k → k ≡ * s ) x=fx ( fc00 zero zero refl cx (init as) 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 .(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 .(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
     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) {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


fcn-imm : {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 ) → ¬ ( ( * x < * y ) ∧ ( * y < * (f x )) ) 
fcn-imm {A} s {x} {y} f mf cx cy ⟪ x<y , y<fx ⟫ = fc21 where
      fc20 : fcn s mf cy Data.Nat.< suc (fcn s mf cx) → (fcn s mf cy ≡ fcn s mf cx) ∨ ( fcn s mf cy Data.Nat.< fcn s mf cx )
      fc20 y<sx with <-cmp ( fcn s mf cy ) (fcn s mf cx )
      ... | tri< a ¬b ¬c = case2 a
      ... | tri≈ ¬a b ¬c = case1 b
      ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> y<sx (s≤s c))
      fc17 : {x y : Ordinal } → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → suc (fcn s mf cx) ≡ fcn s mf cy → * (f x ) ≡ * y
      fc17 {x} {y} cx cy sx=y = fc18 (fcn s mf cy) cx cy refl sx=y  where
          fc18 : (i : ℕ ) → {y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → (i ≡ fcn s mf cy ) → suc (fcn s mf cx) ≡  i → * (f x) ≡ * y
          fc18 (suc i) {y} cx (fsuc y1 cy)  i=y sx=i with proj1 (mf y1 (A∋fc s f mf cy ) )
          ... | case1 y=fy = subst (λ k → * (f x) ≡  k ) y=fy ( fc18 (suc i) {y1} cx cy i=y sx=i)  -- dereference
          ... | case2 y<fy = subst₂ (λ j k → * (f j) ≡ * (f k) ) &iso &iso (cong (λ k → * (f (& k) ) ) fc19) where
               fc19 : * x ≡ * y1
               fc19 = fcn-inject s mf cx cy (cong pred ( trans sx=i i=y ))
      fc21 :  ⊥
      fc21 with <-cmp (suc ( fcn s mf cx )) (fcn s mf cy )
      ... | tri< a ¬b ¬c = <-irr (case2 y<fx) (fc22 a) where -- suc ncx < ncy
           cxx :  FClosure A f s (f x)
           cxx = fsuc x cx 
           fc16 : (x : Ordinal ) → (cx : FClosure A f s x) → (fcn s mf cx  ≡ fcn s mf (fsuc x cx)) ∨ ( suc (fcn s mf cx ) ≡ fcn s mf (fsuc x cx))
           fc16 x (init as) with proj1 (mf s as )
           ... | case1 _ = case1 refl
           ... | case2 _ = case2 refl
           fc16 .(f x) (fsuc x cx ) with proj1 (mf (f x) (A∋fc s f mf (fsuc x cx)) )
           ... | case1 _ = case1 refl 
           ... | case2 _ = case2 refl
           fc22 : (suc ( fcn s mf cx ))  Data.Nat.< (fcn s mf cy ) → * (f x) < * y
           fc22 a with fc16 x cx
           ... | case1 eq = fcn-< s mf cxx cy (subst (λ k → k Data.Nat.< fcn s mf cy ) eq (<-trans a<sa a))
           ... | case2 eq = fcn-< s mf cxx cy (subst (λ k → k Data.Nat.< fcn s mf cy ) eq a )
      ... | tri≈ ¬a b ¬c = <-irr (case1 (fc17 cx cy b)) y<fx
      ... | tri> ¬a ¬b c with fc20 c -- ncy < suc ncx
      ... | case1 y=x = <-irr (case1 ( fcn-inject s mf cy cx  y=x ))  x<y
      ... | case2 y<x = <-irr (case2 x<y) (fcn-< s mf cy cx y<x ) 

-- 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 } (xa : odef A x)  ( f : Ordinal → Ordinal )  : Set n where
   field
      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
      A∋maximal : 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 y not ϕ 
--

record ChainP (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal) (sup z : Ordinal) : Set n where
   field
      y-init   : supf o∅  ≡ y
      asup     : (x : Ordinal) → odef A (supf x)
      fcy<sup  : {z : Ordinal } → FClosure A f y z → z << supf sup
      csupz    : FClosure A f (supf sup) z 
      order    : {sup1 z1 : Ordinal} → (lt : sup1 o< sup ) → FClosure A f (supf sup1 ) z1 → z1 << supf sup

data Chain (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal) : Ordinal →  Ordinal → Set n where
    ch-init    : (z : Ordinal) → FClosure A f y z  → Chain A f mf  ay supf y z 
    ch-is-sup : {sup z : Ordinal } 
         → ( is-sup : ChainP A f mf ay supf sup z)
         → ( fc : FClosure A f (supf sup) z ) → Chain A f mf ay supf sup z

-- Union of supf z which o< x
--

record 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 
   field
      u : Ordinal
      u<x : u o< x
      chain∋z : Chain A f mf ay supf u z

∈∧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 )))

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 }

record ZChain ( A : HOD )    ( f : Ordinal → Ordinal )  (mf : ≤-monotonic-f A f) 
        {init : Ordinal} (ay : odef A init)  ( z : Ordinal ) : Set (Level.suc n) where
   field
      supf :  Ordinal → Ordinal 
   chain : HOD
   chain = UnionCF A f mf ay supf z
   field
      chain⊆A : chain ⊆' A
      chain∋init : odef chain init
      initial : {y : Ordinal } → odef chain y → * init ≤ * y
      f-next : {a : Ordinal } → odef chain a → odef chain (f a)
      f-total : IsTotalOrderSet chain
      is-max :  {a b : Ordinal } → (ca : odef chain a ) →  b o< osuc z  → (ab : odef A b) 
          → HasPrev A chain ab f ∨  IsSup A chain ab  
          → * a < * b  → odef chain b

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

-- data Chain (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y) : Ordinal →  Ordinal → Set n where
--
-- data Chain 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 : Chain A f mf ay supf s a ) ( cb : Chain 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} → Chain A f mf ay supf xa a → Chain A f mf ay supf xb b → Tri (* a < * b) (* a ≡ * b) (* b < * a)
     ct-ind xa xb {a} {b} (ch-init a fca) (ch-init b fcb) = fcn-cmp y f mf fca fcb
     ct-ind xa xb {a} {b} (ch-init a fca) (ch-is-sup supb fcb) = tri< ct01  (λ eq → <-irr (case1 (sym eq)) ct01) (λ lt → <-irr (case2 ct01) lt)  where
          ct00 : * a < * (supf xb)
          ct00 = ChainP.fcy<sup supb fca
          ct01 : * a < * b 
          ct01 with s≤fc (supf xb) 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 supa fca) (ch-init b fcb)= tri> (λ lt → <-irr (case2 ct01) lt) (λ eq → <-irr (case1 eq) ct01) ct01    where
          ct00 : * b < * (supf xa)
          ct00 = ChainP.fcy<sup supa fcb
          ct01 : * b < * a 
          ct01 with s≤fc (supf xa) 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 supa fca) (ch-is-sup supb fcb) with trio< xa xb
     ... | tri< a₁ ¬b ¬c = tri< ct02  (λ eq → <-irr (case1 (sym eq)) ct02) (λ lt → <-irr (case2 ct02) lt)  where
          ct03 : * a < * (supf xb)
          ct03 = ChainP.order supb a₁ (ChainP.csupz supa)
          ct02 : * a < * b 
          ct02 with s≤fc (supf xb) f mf fcb
          ... | case1 eq =  subst (λ k → * a < k ) eq ct03
          ... | case2 lt =  IsStrictPartialOrder.trans POO ct03 lt
     ... | tri≈ ¬a  refl ¬c = fcn-cmp (supf xa) f mf fca fcb
     ... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 ct04) lt) (λ eq → <-irr (case1 (eq)) ct04) ct04    where
          ct05 : * b < * (supf xa)
          ct05 = ChainP.order supa c (ChainP.csupz supb)
          ct04 : * b < * a 
          ct04 with s≤fc (supf xa) f mf fca
          ... | case1 eq =  subst (λ k → * b < k ) eq ct05
          ... | case2 lt =  IsStrictPartialOrder.trans POO ct05 lt

ChainP-next : (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f)  {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal )
   → {x z : Ordinal } →  ChainP A f mf ay supf x z → ChainP A f mf ay supf x (f z )
ChainP-next A f mf {y} ay supf {x} {z} cp = record { y-init = ChainP.y-init cp ; asup = ChainP.asup cp 
     ; fcy<sup = ChainP.fcy<sup cp ; csupz = fsuc _ (ChainP.csupz cp) ; order = ChainP.order cp }

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} → {P : Set n} → odef A y ∧ P → y o< & A
     z07 {y} 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 ; A∋maximal = 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) 

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

     sp0 : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f )  (zc : ZChain A f mf as0 (& A) ) 
        (total : IsTotalOrderSet (ZChain.chain zc) )  → SUP A (ZChain.chain zc)
     sp0 f mf zc total = supP (ZChain.chain zc) (ZChain.chain⊆A zc) total 
     zc< : {x y z : Ordinal} → {P : Set n} → (x o< y → P) → x o< z → z o< y → P
     zc< {x} {y} {z} {P} prev x<z z<y = prev (ordtrans x<z z<y)

     ---
     --- 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) )
            → (total : IsTotalOrderSet (ZChain.chain zc) )
            → f (& (SUP.sup (sp0 f mf zc total ))) ≡ & (SUP.sup (sp0 f mf zc  total))
     fixpoint f mf zc total = z14 where
           chain = ZChain.chain zc
           sp1 = sp0 f mf zc total
           z10 :  {a b : Ordinal } → (ca : odef chain a ) → b o< osuc (& A) → (ab : odef A b ) 
              →  HasPrev A chain ab f ∨  IsSup A chain {b} ab -- (supO  chain  (ZChain.chain⊆A zc) (ZChain.f-total zc) ≡ b )
              → * a < * b  → odef chain b
           z10 = ZChain.is-max zc
           z11 : & (SUP.sup sp1) o< & A
           z11 = c<→o< ( SUP.A∋maximal 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 ) (ordtrans z11 <-osuc ) (SUP.A∋maximal 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.A∋maximal 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 ? ) }
           z14 :  f (& (SUP.sup (sp0 f mf zc total ))) ≡ & (SUP.sup (sp0 f mf zc total ))
           z14 with total (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.A∋maximal 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)) 
           → IsTotalOrderSet (ZChain.chain zc) → ⊥
     z04 nmx zc total = <-irr0  {* (cf nmx c)} {* c} (subst (λ k → odef A k ) (sym &iso) (proj1 (is-cf nmx (SUP.A∋maximal  sp1 ))))
                                               (subst (λ k → odef A (& k)) (sym *iso) (SUP.A∋maximal sp1) )
           (case1 ( cong (*)( fixpoint (cf nmx) (cf-is-≤-monotonic nmx ) zc total ))) -- x ≡ f x ̄
           (proj1 (cf-is-<-monotonic nmx c (SUP.A∋maximal sp1 ))) where          -- x < f x
          sp1 : SUP A (ZChain.chain zc)
          sp1 = sp0 (cf nmx) (cf-is-≤-monotonic nmx) zc total
          c = & (SUP.sup sp1)

     --
     -- 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 trio< y x
     ... | tri> ¬a ¬b c = ?
     ... | tri≈ ¬a refl ¬c = ?
     ... | tri< y<x ¬b ¬c 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
          zc0-b<x : (b : Ordinal ) → b o< x → b o< osuc px
          zc0-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt 

          pchain : HOD
          pchain  = UnionCF A f mf ay (ZChain.supf zc) x
          ptotal : IsTotalOrderSet pchain
          ptotal {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) (UChain.chain∋z (proj2 ca)) (UChain.chain∋z (proj2 cb)) 
          psupf : Ordinal → Ordinal
          psupf z with trio< z x 
          ... | tri< a ¬b ¬c = ZChain.supf zc z
          ... | tri≈ ¬a b ¬c = x
          ... | tri> ¬a ¬b c = x
          pchain⊆A : {y : Ordinal} → odef pchain y →  odef A y
          pchain⊆A {y} ny = proj1 ny
          pnext : {a : Ordinal} → odef pchain a → odef pchain (f a)
          pnext {a} ⟪ aa , ua ⟫ = ⟪ afa , record { u = UChain.u ua ; u<x = UChain.u<x ua ; chain∋z = ? } ⟫ where
               afa : odef A ( f a )
               afa = proj2 ( mf a aa )
               fua : Chain A f mf ay psupf (UChain.u ua) (f a)
               fua with UChain.chain∋z ua
               ... | ch-init a fc = ch-init (f a) ( fsuc _ fc )
               ... | ch-is-sup is-sup fc = ch-is-sup (ChainP-next A f mf ay _ ? ) (fsuc _ ? )
          pinit :  {y₁ : Ordinal} → odef pchain y₁ → * y ≤ * y₁
          pinit {a} ⟪ aa , ua ⟫  with UChain.chain∋z ua
          ... | ch-init a fc = s≤fc y f mf fc
          ... | ch-is-sup is-sup fc = ≤-ftrans (case2 zc7) ? where -- (s≤fc _ f mf fc)  where
               zc7 : y << psupf (UChain.u ua)
               zc7 = ? -- ChainP.fcy<sup is-sup (init ay)
          pcy : odef pchain y
          pcy = ⟪ ay , record { u =  y ; u<x = y<x ; chain∋z = ch-init _ (init ay)  }  ⟫

          -- if previous chain satisfies maximality, we caan reuse it
          --
          no-extenion : ( {a b z : Ordinal} → (z<x : z o< x)  → odef (ZChain.chain zc) a → b o< osuc x → (ab : odef A b) →
                    HasPrev A (ZChain.chain zc ) ab f ∨  IsSup A (ZChain.chain zc ) ab →
                            * a < * b → odef (ZChain.chain zc ) b ) → ZChain A f mf ay x
          no-extenion is-max  = record { initial = ? ; chain∋init = ? }

          zc4 : ZChain A f mf ay x 
          zc4 with ODC.∋-p O A (* x)
          ... | no noax = no-extenion zc1  where -- ¬ A ∋ p, just skip
                zc1 : {a b z : Ordinal} → (z<x : z o< x) → odef (ZChain.chain zc ) a → b o< osuc x → (ab : odef A b) →
                    HasPrev A (ZChain.chain zc ) ab f ∨  IsSup A (ZChain.chain zc ) ab →
                            * a < * b → odef (ZChain.chain zc ) b
                zc1 {a} {b} z<x za b<ox ab P a<b with osuc-≡< b<ox
                ... | case1 eq = ⊥-elim ( noax (subst (λ k → odef A k) (trans eq (sym &iso)) ab ) )
                ... | case2 lt = ZChain.is-max zc za (subst (λ k → b o< k) (sym (Oprev.oprev=x op)) lt ) ab P a<b
          ... | yes ax with ODC.p∨¬p O ( HasPrev A (ZChain.chain zc ) ax f )   
               -- we have to check adding x preserve is-max ZChain A y f mf x
          ... | case1 pr = no-extenion zc7  where -- we have previous A ∋ z < x , f z ≡ x, so chain ∋ f z ≡ x because of f-next
                chain0 = ZChain.chain zc 
                zc7 :  {a b z : Ordinal} → (z<x : z o< x) → odef (ZChain.chain zc ) a → b o< osuc x → (ab : odef A b) →
                            HasPrev A (ZChain.chain zc ) ab f ∨ IsSup A (ZChain.chain zc ) ab →
                            * a < * b → odef (ZChain.chain zc ) b
                zc7 {a} {b} z<x za b<ox ab P a<b with osuc-≡< b<ox
                ... | case2 lt = ZChain.is-max zc za (subst (λ k → b o< k) (sym (Oprev.oprev=x op)) lt ) ab P a<b
                ... | case1 b=x = subst (λ k → odef chain0 k ) (trans (sym (HasPrev.x=fy pr )) (trans &iso (sym b=x)) ) ( ZChain.f-next zc (HasPrev.ay pr))
          ... | case2 ¬fy<x with ODC.p∨¬p O (IsSup A (ZChain.chain zc ) ax )
          ... | case1 is-sup = -- x is a sup of zc 
                record {  chain⊆A = pchain⊆A ; f-next = pnext ; f-total = ptotal
                     ; initial = pinit ; chain∋init  = pcy ; is-max = p-ismax } where
                p-ismax :  {a b : Ordinal} → odef pchain  a →
                    b o< osuc x → (ab : odef A b) →
                    ( HasPrev A pchain ab f ∨ IsSup A pchain ab ) →
                        * a < * b → odef pchain b
                p-ismax {a} {b} ua b<ox ab (case1 hasp) a<b = ?
                p-ismax {a} {b} ua b<ox ab (case2 sup)  a<b = ?

          ... | case2 ¬x=sup =  no-extenion z18 where -- x is not f y' nor sup of former ZChain from y -- no extention
                z18 :  {a b z : Ordinal} → (z<x : z o< x) → odef (ZChain.chain zc ) a → b o< osuc x → (ab : odef A b) →
                            HasPrev A (ZChain.chain zc ) ab f ∨ IsSup A (ZChain.chain zc )   ab →
                            * a < * b → odef (ZChain.chain zc ) b
                z18 {a} {b} z<x za b<x ab p a<b with osuc-≡< b<x
                ... | case2 lt = ZChain.is-max zc za (subst (λ k → b o< k) (sym (Oprev.oprev=x op)) lt ) ab p a<b 
                ... | case1 b=x with p
                ... | case1 pr = ⊥-elim ( ¬fy<x record {y = HasPrev.y pr ; ay = HasPrev.ay pr ; x=fy = trans (trans &iso (sym b=x) ) (HasPrev.x=fy pr ) } )
                ... | case2 b=sup = ⊥-elim ( ¬x=sup record { 
                      x<sup = λ {y} zy → subst (λ k → (y ≡ k) ∨ (y << k)) (trans b=x (sym &iso)) (IsSup.x<sup b=sup zy )  } ) 
     ... | no op = zc5 where
          pzc : (z : Ordinal) → z o< x → ZChain A f mf ay z
          pzc z z<x = prev z z<x
          psupf0 : (z : Ordinal) → Ordinal
          psupf0 z with trio< z x
          ... | tri< a ¬b ¬c = ZChain.supf (pzc z a) z
          ... | tri≈ ¬a b ¬c = y
          ... | tri> ¬a ¬b c = y
          pchain : HOD
          pchain = UnionCF A f mf ay psupf0 x
          ptotal : IsTotalOrderSet pchain
          ptotal {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 psupf0 (UChain.chain∋z (proj2 ca)) (UChain.chain∋z (proj2 cb)) 

          pchain⊆A : {y : Ordinal} → odef pchain y →  odef A y
          pchain⊆A {y} ny = proj1 ny
          pnext : {a : Ordinal} → odef pchain a → odef pchain (f a)
          pnext {a} ⟪ aa , ua ⟫ = ⟪ afa , record { u = UChain.u ua ; u<x = UChain.u<x ua ; chain∋z = ? } ⟫ where
               afa : odef A ( f a )
               afa = proj2 ( mf a aa )
               fua : Chain A f mf ay psupf0 (UChain.u ua) (f a)
               fua with UChain.chain∋z ua
               ... | ch-init a fc = ch-init (f a) ( fsuc _ fc )
               ... | ch-is-sup is-sup fc = ch-is-sup (ChainP-next A f mf ay _ ? ) (fsuc _ ? )
          pinit :  {y₁ : Ordinal} → odef pchain y₁ → * y ≤ * y₁
          pinit {a} ⟪ aa , ua ⟫  with UChain.chain∋z ua
          ... | ch-init a fc = s≤fc y f mf fc
          ... | ch-is-sup is-sup fc = ≤-ftrans (case2 zc7) ? where -- (s≤fc _ f mf fc)  where
               zc7 : y << psupf0 (UChain.u ua)
               zc7 = ? -- ChainP.fcy<sup is-sup (init ay)
          pcy : odef pchain y
          pcy = ⟪ ay , record { u =  y ; u<x = y<x ; chain∋z = ch-init _ (init ay)  }  ⟫

          usup : SUP A pchain
          usup = supP pchain (λ lt → proj1 lt) ptotal
          spu = & (SUP.sup usup)
          psupf : Ordinal → Ordinal
          psupf z with trio< z x 
          ... | tri< a ¬b ¬c = ZChain.supf (pzc z a) z
          ... | tri≈ ¬a b ¬c = spu
          ... | tri> ¬a ¬b c = spu

          zc5 : ZChain A f mf ay x 
          zc5 with ODC.∋-p O A (* x)
          ... | no noax = {!!} where -- ¬ A ∋ p, just skip
          ... | yes ax with ODC.p∨¬p O ( HasPrev A pchain ax f )   
               -- we have to check adding x preserve is-max ZChain A y f mf x
          ... | case1 pr = {!!} where -- 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 pchain ax )
          ... | case1 is-sup = {!!} -- x is a sup of (zc ?) 
          ... | case2 ¬x=sup =  {!!} -- no-extenion z18 where -- 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)) ; A∋maximal = 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 total ) 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.A∋maximal 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 ) 
         total : IsTotalOrderSet (ZChain.chain zorn04)
         total {a} {b} = zorn06  where
             zorn06 :  odef (ZChain.chain zorn04) (& a) → odef (ZChain.chain zorn04) (& b) → Tri (a < b) (a ≡ b) (b < a)
             zorn06 = ZChain.f-total (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