changeset 982:6d29911a9d00

... order supf o< supf is bad?
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Fri, 11 Nov 2022 19:27:15 +0900
parents 6b67e43733ca
children 6101b9bfbe57
files src/zorn.agda
diffstat 1 files changed, 46 insertions(+), 34 deletions(-) [+]
line wrap: on
line diff
--- a/src/zorn.agda	Thu Nov 10 09:07:07 2022 +0900
+++ b/src/zorn.agda	Fri Nov 11 19:27:15 2022 +0900
@@ -273,7 +273,7 @@
 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 )
+    ch-is-sup  : (u : Ordinal) {z : Ordinal }  (u<x : u o< x) ( is-sup : ChainP A f mf ay supf u )
         ( fc : FClosure A f (supf u) z ) → UChain A f mf ay supf x z
 
 --
@@ -289,7 +289,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
@@ -297,14 +297,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
@@ -312,7 +312,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
@@ -401,8 +401,8 @@
         → 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  ⟫
+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 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
@@ -435,11 +435,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
@@ -617,8 +617,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
 
@@ -633,7 +633,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 (o<→≤ zc08) is-sup fc ⟫ where
+                ... | ⟪ as , ch-is-sup u u<x is-sup fc ⟫ = ⟪ as , ch-is-sup u (ZChain.supf-inject zc 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
@@ -644,7 +644,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) ) )
@@ -668,7 +668,7 @@
            is-max {a} {b} ua b<x ab P a<b with ODC.or-exclude O P
            is-max {a} {b} ua b<x ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b
            is-max {a} {b} ua sb<sx ab P a<b | case2 is-sup
-             = ⟪ ab , ch-is-sup b (ZChain.supf-mono zc (o<→≤ b<x)) m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl)) ⟫   where
+             = ⟪ ab , ch-is-sup b b<x m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl)) ⟫   where
               b<A : b o< & A
               b<A = z09 ab
               b<x : b o< x
@@ -694,7 +694,7 @@
            is-max {a} {b} ua sb<sx ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b
            is-max {a} {b} ua sb<sx ab P a<b | case2 is-sup with IsSUP.x≤sup (proj2 is-sup) (init-uchain A f mf ay )
            ... | case1 b=y = ⟪ subst (λ k → odef A k ) b=y ay , ch-init (subst (λ k → FClosure A f y k ) b=y (init ay refl ))  ⟫
-           ... | case2 y<b = ⟪ ab , ch-is-sup b (ZChain.supf-mono zc (o<→≤ b<x)) m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl))  ⟫ where
+           ... | case2 y<b = ⟪ ab , ch-is-sup b b<x m06 (subst (λ k → FClosure A f k b) (sym m05) (init ab refl))  ⟫ where
               m09 : b o< & A
               m09 = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab))
               b<x : b o< x
@@ -1051,7 +1051,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
@@ -1175,9 +1175,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 f x  → ¬ xSUP pchain f x
           zc70 pr xsup = ?
@@ -1191,23 +1191,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) {!!}) {!!} )
@@ -1396,6 +1396,18 @@
           sz<<c : {z : Ordinal } → z o< & A → supf z <= mc
           sz<<c z<A = MinSUP.x≤sup msp1 (ZChain.csupf zc (z09 (ZChain.asupf zc)  ))
 
+          z511 : {u y mc : Ordinal} (u<x : u o< mc ) (is-sup1 : ChainP A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) u)
+                (fc  : FClosure A (cf nmx) (ZChain.supf zc u) y) (x=fy : mc ≡ cf nmx y) → 
+                   ( * (cf nmx (cf nmx y)) ≡ * (supf mc)) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) )
+          z511 {u} {y} {mc} u<x is-sup fc x=fy with osuc-≡< (ZChain.supf-mono zc (o<→≤ u<x))
+          ... | case2 lt = <=to≤ (ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) ? lt (fsuc _ ( fsuc  _ fc )))
+          ... | case1 eq = ?
+               --    u ≡ supf u ≡ supf mc ≡ supf (cf nmx y) 
+               --    supf u << cf nmx (cf nmx ( ... (supf u )) <= spuf mc ≡ u
+               --    y ≡ supf u
+               --    y ≡ cf nmx (supf u)
+               --    y ≡ cf nmx (cf nmx (supf u))
+
           sc=c : supf mc ≡ mc
           sc=c = ZChain.sup=u zc (MinSUP.asm msp1) (o<→≤ (∈∧P→o< ⟪ MinSUP.asm msp1 , lift true ⟫ )) ⟪ is-sup , not-hasprev ⟫ where
               not-hasprev :  ¬ HasPrev A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) mc) (cf nmx) mc
@@ -1405,7 +1417,7 @@
                    z31 : ( * (cf nmx mc)  ≡  * mc ) ∨ ( * (cf nmx mc) < * mc )
                    z31 = <=to≤ ( MinSUP.x≤sup msp1 (subst (λ k →  odef (ZChain.chain zc) (cf nmx k)) (sym x=fy)
                        ⟪ proj2  (cf-is-≤-monotonic nmx _ (proj2 (cf-is-≤-monotonic nmx _ ua1 ) )) , ch-init (fsuc _ (fsuc _ fc))  ⟫ ))
-              not-hasprev record { ax = ax ; y = y ; ay =   ⟪ ua1 , ch-is-sup u u≤x is-sup1 fc ⟫; x=fy = x=fy } = ⊥-elim ( <-irr z48 z32 ) where
+              not-hasprev record { ax = ax ; y = y ; ay =   ⟪ ua1 , ch-is-sup u u<x is-sup1 fc ⟫; x=fy = x=fy } = ⊥-elim ( <-irr z48 z32 ) where
                    z30 : * mc < * (cf nmx mc)
                    z30 = proj1 (cf-is-<-monotonic nmx mc (MinSUP.asm msp1))
                    z31 : ( supf mc ≡  mc ) ∨ ( * (supf mc) < * mc )
@@ -1413,7 +1425,7 @@
                    z32 : * (supf mc) < * (cf nmx (cf nmx y))
                    z32 = ftrans<=-< z31 (subst (λ k → * mc < * k ) (cong (cf nmx) x=fy) z30 )
                    z48 : ( * (cf nmx (cf nmx y)) ≡ * (supf mc)) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) )
-                   z48 = <=to≤ (ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) mc<A ? (fsuc _ ( fsuc  _ fc )))
+                   z48 = z511 u<x is-sup1 fc x=fy
               is-sup :  IsSUP A (UnionCF A (cf nmx) (cf-is-≤-monotonic nmx) as0 (ZChain.supf zc) mc)  (MinSUP.asm msp1)
               is-sup = record { x≤sup = λ zy → MinSUP.x≤sup msp1 (chain-mono (cf nmx) (cf-is-≤-monotonic nmx) as0 supf (ZChain.supf-mono zc) (o<→≤ mc<A) zy ) }
 
@@ -1426,16 +1438,17 @@
                 z32 = ZChain.fcy<sup zc (o<→≤ mc<A) (fsuc _ (fsuc _ fc))
                 z31 : ( * (cf nmx d)  ≡  * d ) ∨ ( * (cf nmx d) < * d )
                 z31 = case2 ( subst (λ k → * (cf nmx k) < * d ) (sym x=fy) ( ftrans<=-< z32 ( sc<<d {mc} asc spd ) ))
-          not-hasprev record { ax = ax ; y = y ; ay =   ⟪ ua1 , ch-is-sup u u≤x is-sup1 fc ⟫; x=fy = x=fy } = ⊥-elim ( <-irr z46 z30 ) where
+          not-hasprev record { ax = ax ; y = y ; ay =   ⟪ ua1 , ch-is-sup u u<x is-sup1 fc ⟫; x=fy = x=fy } = ⊥-elim ( <-irr z46 z30 ) where
                 z45 : (* (cf nmx (cf nmx y)) ≡ * d) ∨ (* (cf nmx (cf nmx y)) < * d) → (* (cf nmx d) ≡ * d) ∨ (* (cf nmx d) < * d)
                 z45 p = subst (λ k → (* (cf nmx k) ≡ * d) ∨ (* (cf nmx k) < * d)) (sym x=fy) p
                 z48 : supf mc << d
                 z48 = sc<<d {mc} asc spd
                 z53 : supf u o≤ supf (& A)
-                z53 = OrdTrans u≤x ?
+                z53 = ZChain.supf-mono zc (o<→≤ (z09 (subst (λ k → odef A k) (ChainP.supu=u is-sup1) (ZChain.asupf zc) ))) -- OrdTrans u<x ?
                 z52 : ( u ≡ mc ) ∨  ( u << mc )
                 z52 = MinSUP.x≤sup msp1 ⟪ subst (λ k → odef A k ) (ChainP.supu=u is-sup1) (A∋fcs _ _ (cf-is-≤-monotonic nmx) fc)
-                        , ch-is-sup u ? is-sup1 (init (ZChain.asupf zc)  (ChainP.supu=u is-sup1)) ⟫
+                        , ch-is-sup u (z09 (subst (λ k → odef A k) (ChainP.supu=u is-sup1) (ZChain.asupf zc) )) 
+                           is-sup1 (init (ZChain.asupf zc)  (ChainP.supu=u is-sup1)) ⟫
                 z51 : supf u o≤ supf mc
                 z51 = or z52 (λ le → o≤-refl0 (z56 le) ) z57 where
                     z56 : u ≡ mc → supf u ≡  supf mc
@@ -1447,7 +1460,7 @@
                 z49 : supf u o< supf mc → (cf nmx (cf nmx y) ≡ supf mc) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) )
                 z49 su<smc = ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) mc<A su<smc (fsuc _ ( fsuc  _ fc ))
                 z50 : (cf nmx (cf nmx y) ≡ supf d) ∨ (* (cf nmx (cf nmx y)) < * (supf d) )
-                z50 = ZChain1.order (SZ1 (cf nmx) (cf-is-≤-monotonic nmx) as0 zc (& A)) d<A ? (fsuc _ ( fsuc  _ fc ))
+                z50 = ≤to<= ( z511 u<x is-sup1 fc x=fy )
                 z47 : {mc d1 : Ordinal } {asc : odef A (supf mc)} (spd : MinSUP A (uchain (cf nmx)  (cf-is-≤-monotonic nmx) asc ))
                     → (cf nmx (cf nmx y) ≡ supf mc) ∨ (* (cf nmx (cf nmx y)) < * (supf mc) ) → supf mc << d1
                     →  * (cf nmx (cf nmx y)) < * d1
@@ -1474,13 +1487,12 @@
                 z22 {a} ⟪ aa , ch-init fc ⟫ = case2 ( ( ftrans<=-< z32 ( sc<<d {mc} asc spd ) )) where
                     z32 : ( a  ≡  supf mc ) ∨ ( * a < * (supf mc) )
                     z32 = ZChain.fcy<sup zc (o<→≤ mc<A) fc
-                z22 {a} ⟪ aa , ch-is-sup u u≤x is-sup1 fc ⟫ = tri u (supf mc)
+                z22 {a} ⟪ aa , ch-is-sup u u<x is-sup1 fc ⟫ = tri u (supf mc)
                        z60 z61 ( λ sc<u → ⊥-elim ( o≤> ( subst (λ k → k o≤ supf mc) (ChainP.supu=u is-sup1) z51) sc<u )) where
-                    z53 : supf u o< supf (& A)
-                    z53 = ordtrans<-≤ u≤x ?
                     z52 : ( u ≡ mc ) ∨  ( u << mc )
                     z52 = MinSUP.x≤sup msp1 ⟪ subst (λ k → odef A k ) (ChainP.supu=u is-sup1) (A∋fcs _ _ (cf-is-≤-monotonic nmx) fc)
-                           , ch-is-sup u ? is-sup1 (init (ZChain.asupf zc)  (ChainP.supu=u is-sup1)) ⟫
+                           , ch-is-sup u (z09 (subst (λ k → odef A k) (ChainP.supu=u is-sup1) (ZChain.asupf zc) )) 
+                               is-sup1 (init (ZChain.asupf zc)  (ChainP.supu=u is-sup1)) ⟫
                     z56 : u ≡ mc → supf u ≡  supf mc
                     z56 eq = cong supf eq
                     z57 : u << mc → supf u o≤ supf mc
@@ -1494,7 +1506,7 @@
                         (subst (λ k → k o< supf mc) (sym (ChainP.supu=u is-sup1))  u<smc)  fc ) smc<<d )
                     z61 : u ≡ supf mc → (a ≡ d ) ∨ ( * a < * d )
                     z61 u=sc = case2 (fsc<<d {mc} asc spd (subst (λ k → FClosure A (cf nmx) k a) (trans (ChainP.supu=u is-sup1) u=sc) fc ) )
-                    -- u≤x    : ZChain.supf zc u o< ZChain.supf zc d
+                    -- u<x    : ZChain.supf zc u o< ZChain.supf zc d
                     --     supf u o< supf c → order
 
           sd=d : supf d ≡ d