Mercurial > hg > Members > kono > Proof > ZF-in-agda
view src/zorn.agda @ 751:71ad137dd101
..
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Fri, 22 Jul 2022 19:18:05 +0900 |
parents | b96491f7c775 |
children | 2be90b90deb3 |
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 ) ) 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) = 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) = 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 ) fc-conv : (A : HOD ) (f : Ordinal → Ordinal) {b u : Ordinal } → {p0 p1 : Ordinal → Ordinal} → p0 u ≡ p1 u → FClosure A f (p0 u) b → FClosure A f (p1 u) b fc-conv A f {.(p0 u)} {u} {p0} {p1} p0u=p1u (init ap0u) = subst (λ k → FClosure A f (p1 u) k) (sym p0u=p1u) ( init (subst (λ k → odef A k) p0u=p1u ap0u )) fc-conv A f {_} {u} {p0} {p1} p0u=p1u (fsuc z fc) = fsuc z (fc-conv A f {_} {u} {p0} {p1} p0u=p1u fc) -- 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) (u z : Ordinal) : Set n where field fcy<sup : {z : Ordinal } → FClosure A f y z → z << supf u csupz : FClosure A f (supf u) z order : {sup1 z1 : Ordinal} → (lt : sup1 o< u ) → FClosure A f (supf sup1 ) z1 → z1 << supf u y<s : y << supf u -- not a initial chain supfu=u : supf u ≡ u -- Union of supf z which o< x -- data UChain ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y ) (supf : Ordinal → Ordinal) (x : Ordinal) : (z : Ordinal) → Set n where ch-init : {z : Ordinal } (fc : FClosure A f y z) → UChain A f mf ay supf x z ch-is-sup : (u : Ordinal) {z : Ordinal } (u<x : u o< x ) ( is-sup : ChainP A f mf ay supf u z) ( fc : FClosure A f (supf u) z ) → UChain A f mf ay supf x 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 sup=u : {b : Ordinal} → {ab : odef A b} → b o< z → IsSup A chain ab → supf b ≡ b order : {b sup1 z1 : Ordinal} → b o< z → sup1 o< b → FClosure A f (supf sup1) z1 → z1 << supf b record ZChain1 ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {init : Ordinal} (ay : odef A init) (zc : ZChain A f mf ay (& A)) ( z : Ordinal ) : Set (Level.suc n) where field chain-mono2 : { a b c : Ordinal } → a o≤ b → b o≤ z → odef (UnionCF A f mf ay (ZChain.supf zc) a) c → odef (UnionCF A f mf ay (ZChain.supf zc) b) c is-max : {a b : Ordinal } → (ca : odef (UnionCF A f mf ay (ZChain.supf zc) z) a ) → b o< z → (ab : odef A b) → HasPrev A (UnionCF A f mf ay (ZChain.supf zc) z) ab f ∨ IsSup A (UnionCF A f mf ay (ZChain.supf zc) z) ab → * a < * b → odef ((UnionCF A f mf ay (ZChain.supf zc) z)) b fcy<sup : {u w : Ordinal } → u o< z → FClosure A f init w → w << (ZChain.supf zc) u sup=u : {b : Ordinal} → {ab : odef A b} → b o< z → IsSup A (UnionCF A f mf ay (ZChain.supf zc) (osuc b)) ab → (ZChain.supf zc) b ≡ b order : {b sup1 z1 : Ordinal} → b o< z → sup1 o< b → FClosure A f ((ZChain.supf zc) sup1) z1 → z1 << (ZChain.supf zc) 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 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 ub<x supb fcb) = tri< ct01 (λ eq → <-irr (case1 (sym eq)) ct01) (λ lt → <-irr (case2 ct01) lt) where ct00 : * a < * (supf ub) ct00 = ChainP.fcy<sup supb fca 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 ua<x supa fca) (ch-init fcb)= tri> (λ lt → <-irr (case2 ct01) lt) (λ eq → <-irr (case1 eq) ct01) ct01 where ct00 : * b < * (supf ua) ct00 = ChainP.fcy<sup supa fcb 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 = tri< ct02 (λ eq → <-irr (case1 (sym eq)) ct02) (λ lt → <-irr (case2 ct02) lt) where ct03 : * a < * (supf ub) ct03 = ChainP.order supb a₁ (ChainP.csupz supa) 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 ... | tri≈ ¬a refl ¬c = fcn-cmp (supf ua) f mf fca fcb ... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 ct04) lt) (λ eq → <-irr (case1 (eq)) ct04) ct04 where ct05 : * b < * (supf ua) ct05 = ChainP.order supa c (ChainP.csupz supb) 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 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) ⟫ 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<s = ChainP.y<s cp ; supfu=u = ChainP.supfu=u 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) SZ1 :( A : HOD ) ( 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 A f mf {y} ay zc x = TransFinite { λ x → ZChain1 A f mf ay zc x } zc1 x where chain-mono2 : (x : Ordinal) {a b c : Ordinal} → a o≤ b → b o≤ x → odef (UnionCF A f mf ay (ZChain.supf zc) a) c → odef (UnionCF A f mf ay (ZChain.supf zc) b) c chain-mono2 x {a} {b} {c} a≤b b≤x ⟪ ua , ch-init fc ⟫ = ⟪ ua , ch-init fc ⟫ chain-mono2 x {a} {b} {c} a≤b b≤x ⟪ uaa , ch-is-sup ua u<x is-sup fc ⟫ = ⟪ uaa , ch-is-sup ua (ordtrans<-≤ u<x a≤b ) is-sup fc ⟫ chain<ZA : {x : Ordinal } → UnionCF A f mf ay (ZChain.supf zc) x ⊆' UnionCF A f mf ay (ZChain.supf zc) (& A) chain<ZA {x} ux with proj2 ux ... | ch-init fc = ⟪ proj1 ux , ch-init fc ⟫ ... | ch-is-sup u pu<x is-sup fc = ⟪ proj1 ux , ch-is-sup u u<x is-sup fc ⟫ where u<A : (& ( * ( ZChain.supf zc u))) o< & A u<A = c<→o< (subst (λ k → odef A k ) (sym &iso) (A∋fcs _ f mf fc) ) u<x : u o< & A u<x = subst (λ k → k o< & A ) (trans &iso (ChainP.supfu=u is-sup)) u<A is-max-hp : (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a → b o< x → (ab : odef A b) → HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) ab f → * a < * b → odef (UnionCF A f mf ay (ZChain.supf zc) x) b is-max-hp x {a} {b} ua b<x ab has-prev a<b with HasPrev.ay has-prev ... | ⟪ ab0 , ch-init fc ⟫ = ⟪ ab , ch-init ( subst (λ k → FClosure A f y k) (sym (HasPrev.x=fy has-prev)) (fsuc _ fc )) ⟫ ... | ⟪ ab0 , ch-is-sup u u<x is-sup fc ⟫ = ⟪ ab , subst (λ k → UChain A f mf ay (ZChain.supf zc) x k ) (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x (ChainP-next A f mf ay _ is-sup) (fsuc _ fc)) ⟫ zc1 : (x : Ordinal) → ((y₁ : Ordinal) → y₁ o< x → ZChain1 A f mf ay zc y₁) → ZChain1 A f mf ay zc x zc1 x prev with Oprev-p x ... | yes op = record { is-max = is-max ; chain-mono2 = chain-mono2 x ; fcy<sup = fcy<sup ; sup=u = ? ; order = order } where px = Oprev.oprev op px<x : px o< x px<x = subst (λ k → px o< k ) (Oprev.oprev=x op) <-osuc zc-b<x : (b : Ordinal ) → b o< x → b o< osuc px zc-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt fcy<sup : {u w : Ordinal} → u o< x → FClosure A f y w → w << ZChain.supf zc u fcy<sup {a} {b} a<x fcb with trio< a px ... | tri< a<px ¬b ¬c = ZChain1.fcy<sup (prev px px<x) a<px fcb ... | tri≈ ¬a a=px ¬c = ? ... | tri> ¬a ¬b c = ⊥-elim ( ¬p<x<op ⟪ c , subst (λ k → a o< k ) (sym (Oprev.oprev=x op)) a<x ⟫ ) order : {b sup1 z1 : Ordinal} → b o< x → sup1 o< b → FClosure A f (ZChain.supf zc sup1) z1 → z1 << ZChain.supf zc b order {b} {sup1} {z1} b<x s<b fcz with trio< b px ... | tri< b<px ¬b ¬c = ZChain1.order (prev px px<x) b<px s<b fcz ... | tri≈ ¬a b=px ¬c = ? ... | tri> ¬a ¬b c = ⊥-elim ( ¬p<x<op ⟪ c , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ ) sup=u : {b : Ordinal} {ab : odef A b} → b o< x → IsSup A (UnionCF A f mf ay (ZChain.supf zc) (osuc b)) ab → ZChain.supf zc b ≡ b sup=u {b} {ab} b<x is-sup with trio< b px ... | tri< b<px ¬b ¬c = ZChain1.sup=u (prev px px<x) b<px is-sup ... | tri≈ ¬a b=px ¬c = ZChain.sup=u zc ? is-sup ... | tri> ¬a ¬b c = ⊥-elim ( ¬p<x<op ⟪ c , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ ) is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a → b o< x → (ab : odef A b) → HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) ab f ∨ IsSup A (UnionCF A f mf ay (ZChain.supf zc) x) ab → * a < * b → odef (UnionCF A f mf ay (ZChain.supf zc) x) b is-max {a} {b} ua b<x ab (case1 has-prev) a<b = is-max-hp x {a} {b} ua b<x ab has-prev a<b is-max {a} {b} ua b<x ab (case2 is-sup) a<b with ODC.p∨¬p O ( HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) ab f ) ... | case1 has-prev = is-max-hp x {a} {b} ua b<x ab has-prev a<b ... | case2 ¬fy<x = m01 where m01 : odef (UnionCF A f mf ay (ZChain.supf zc) x) b m01 with trio< b px --- px < b < x ... | tri> ¬a ¬b c = ⊥-elim (¬p<x<op ⟪ c , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫) ... | tri< b<px ¬b ¬c = chain-mono2 x ( o<→≤ (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc )) o≤-refl m04 where m03 : odef (UnionCF A f mf ay (ZChain.supf zc) px) a m03 with proj2 ua ... | ch-init fc = ⟪ proj1 ua , ch-init fc ⟫ ... | ch-is-sup u u<x is-sup fc with trio< u px ... | tri< a ¬b ¬c = ⟪ proj1 ua , ch-is-sup u a is-sup fc ⟫ ... | tri≈ ¬a u=px ¬c = ? --- supf u < a < b , ... | tri> ¬a ¬b c = ⊥-elim ( ¬p<x<op ⟪ c , subst (λ k → u o< k ) (sym (Oprev.oprev=x op)) u<x ⟫ ) -- = ⟪ proj1 ua , ? ⟫ m04 : odef (UnionCF A f mf ay (ZChain.supf zc) px) b m04 = ZChain1.is-max (prev px px<x) m03 b<px ab (case2 record {x<sup = λ {z} lt → IsSup.x<sup is-sup (chain-mono2 x ( o<→≤ (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc )) o≤-refl lt) } ) a<b ... | tri≈ ¬a b=px ¬c = ? -- b = px case, u = px u<x ... | no lim = record { is-max = is-max ; chain-mono2 = chain-mono2 x ; fcy<sup = fcy<sup ; sup=u = sup=u ; order = order } where fcy<sup : {u w : Ordinal} → u o< x → FClosure A f y w → w << ZChain.supf zc u fcy<sup {u} {w} u<x fc = ZChain1.fcy<sup (prev (osuc u) (ob<x lim u<x)) <-osuc fc sup=u : {b : Ordinal} {ab : odef A b} → b o< x → IsSup A (UnionCF A f mf ay (ZChain.supf zc) (osuc b)) ab → ZChain.supf zc b ≡ b sup=u {b} {ab} b<x is-sup = ZChain1.sup=u (prev (osuc b) (ob<x lim b<x)) <-osuc is-sup order : {b sup1 z1 : Ordinal} → b o< x → sup1 o< b → FClosure A f (ZChain.supf zc sup1) z1 → z1 << ZChain.supf zc b order {b} b<x s<b fc = ZChain1.order (prev (osuc b) (ob<x lim b<x)) <-osuc s<b fc is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) x) a → b o< x → (ab : odef A b) → HasPrev A (UnionCF A f mf ay (ZChain.supf zc) x) ab f ∨ IsSup A (UnionCF A f mf ay (ZChain.supf zc) x) ab → * a < * b → odef (UnionCF A f mf ay (ZChain.supf zc) x) b is-max {a} {b} ua b<x ab (case1 has-prev) a<b = is-max-hp x {a} {b} ua b<x ab has-prev a<b is-max {a} {b} ua b<x ab (case2 is-sup) a<b with IsSup.x<sup is-sup (init-uchain A f mf ay ) ... | case1 b=y = ⊥-elim ( <-irr ( ZChain.initial zc (chain<ZA (chain-mono2 (osuc x) (o<→≤ <-osuc ) o≤-refl ua )) ) (subst (λ k → * a < * k ) (sym b=y) a<b ) ) ... | case2 y<b = chain-mono2 x (o<→≤ (ob<x lim b<x) ) o≤-refl m04 where y<s : y << ZChain.supf zc b y<s = y<s m07 : {z : Ordinal} → FClosure A f y z → z << ZChain.supf zc b m07 {z} fc = ZChain1.fcy<sup (prev (osuc b) (ob<x lim b<x)) <-osuc fc m08 : {sup1 z1 : Ordinal} → sup1 o< b → FClosure A f (ZChain.supf zc sup1) z1 → z1 << ZChain.supf zc b m08 {sup1} {z1} s<b fc = ZChain1.order (prev (osuc b) (ob<x lim b<x) ) <-osuc s<b fc m05 : b ≡ ZChain.supf zc b m05 = sym (ZChain1.sup=u (prev (osuc b) (ob<x lim b<x)) {_} {ab} <-osuc record { x<sup = λ lt → IsSup.x<sup is-sup (chain-mono2 x (o<→≤ (ob<x lim b<x)) o≤-refl lt )} ) -- ZChain on x m06 : ChainP A f mf ay (ZChain.supf zc) b b m06 = record { fcy<sup = m07 ; csupz = subst (λ k → FClosure A f k b ) m05 (init ab) ; order = m08 ; y<s = y<s ; supfu=u = sym m05 } m04 : odef (UnionCF A f mf ay (ZChain.supf zc) (osuc b)) b m04 = ⟪ ab , ch-is-sup b <-osuc m06 (subst (λ k → FClosure A f k b) m05 (init ab)) ⟫ --- --- 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< & 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 = ZChain1.is-max (SZ1 A f mf as0 zc (& A) ) 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 ) z11 (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) inititalChain : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → ZChain A f mf ay o∅ inititalChain f mf {y} ay = record { supf = isupf ; chain⊆A = λ lt → proj1 lt ; chain∋init = cy ; initial = isy ; f-next = inext ; f-total = itotal ; sup=u = λ b<0 → ⊥-elim (¬x<0 b<0) ; order = λ b<0 → ⊥-elim (¬x<0 b<0) } where isupf : Ordinal → Ordinal isupf z = y cy : odef (UnionCF A f mf ay isupf o∅) y cy = ⟪ ay , ch-init (init ay) ⟫ isy : {z : Ordinal } → odef (UnionCF A f mf ay isupf o∅) z → * y ≤ * z isy {z} ⟪ az , uz ⟫ with uz ... | ch-init fc = s≤fc y f mf fc ... | ch-is-sup u u<x is-sup fc = ⊥-elim ( <-irr (case1 refl) ( ChainP.fcy<sup is-sup (init ay) ) ) inext : {a : Ordinal} → odef (UnionCF A f mf ay isupf o∅) a → odef (UnionCF A f mf ay isupf o∅) (f a) inext {a} ua with (proj2 ua) ... | ch-init fc = ⟪ proj2 (mf _ (proj1 ua)) , ch-init (fsuc _ fc ) ⟫ ... | ch-is-sup u u<x is-sup fc = ⊥-elim ( <-irr (case1 refl) ( ChainP.fcy<sup is-sup (init ay) ) ) itotal : IsTotalOrderSet (UnionCF A f mf ay isupf o∅) itotal {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 isupf (proj2 ca) (proj2 cb) imax : {a b : Ordinal} → odef (UnionCF A f mf ay isupf o∅) a → b o< o∅ → (ab : odef A b) → HasPrev A (UnionCF A f mf ay isupf o∅) ab f ∨ IsSup A (UnionCF A f mf ay isupf o∅) ab → * a < * b → odef (UnionCF A f mf ay isupf o∅) b imax {a} {b} ua b<x ab (case1 hasp) a<b = subst (λ k → odef (UnionCF A f mf ay isupf o∅) k ) (sym (HasPrev.x=fy hasp)) ( inext (HasPrev.ay hasp) ) imax {a} {b} ua b<x ab (case2 sup) a<b = ⊥-elim ( ¬x<0 b<x ) -- -- create all ZChains under o< x -- ind : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → (x : Ordinal) → ((z : Ordinal) → z o< x → ZChain A f mf ay z) → ZChain A f mf ay x ind f mf {y} ay x prev with Oprev-p x ... | yes op = zc4 where -- -- we have previous ordinal to use induction -- px = Oprev.oprev op zc : ZChain A f mf ay (Oprev.oprev op) zc = prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc ) px<x : px o< x px<x = subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc zc-b<x : (b : Ordinal ) → b o< x → b o< osuc px zc-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt 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) ( (proj2 ca)) ( (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 , ch-init fc ⟫ = ⟪ proj2 (mf a aa) , ch-init (fsuc _ fc) ⟫ pnext {a} ⟪ aa , ch-is-sup u u<x is-sup fc ⟫ = ⟪ proj2 (mf a aa) , ch-is-sup u u<x (ChainP-next A f mf ay _ is-sup ) (fsuc _ fc ) ⟫ pinit : {y₁ : Ordinal} → odef pchain y₁ → * y ≤ * y₁ pinit {a} ⟪ aa , ua ⟫ with ua ... | ch-init fc = s≤fc y f mf fc ... | ch-is-sup u u<x is-sup fc = ≤-ftrans (case2 zc7) (s≤fc _ f mf fc) where zc7 : y << (ZChain.supf zc) u zc7 = ChainP.fcy<sup is-sup (init ay) pcy : odef pchain y pcy = ⟪ ay , ch-init (init ay) ⟫ supf0 = ZChain.supf (prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc )) -- if previous chain satisfies maximality, we caan reuse it -- no-extension : ZChain A f mf ay x no-extension = record { supf = supf0 ; initial = pinit ; chain∋init = pcy ; sup=u = sup=u ; chain⊆A = pchain⊆A ; f-next = pnext ; f-total = ptotal ; order = order } where sup=u : {b : Ordinal} {ab : odef A b} → b o< x → IsSup A (UnionCF A f mf ay supf0 x) ab → supf0 b ≡ b sup=u {b} {ab} b<x is-sup with trio< b px ... | tri< a ¬b ¬c = ZChain.sup=u zc {b} {ab} a record { x<sup = pc20 } where pc20 : {z : Ordinal} → odef (UnionCF A f mf ay (ZChain.supf zc) (Oprev.oprev op)) z → (z ≡ b) ∨ (z << b) pc20 {z} ⟪ az , ch-init fc ⟫ = IsSup.x<sup is-sup ⟪ az , ch-init fc ⟫ pc20 {z} ⟪ az , ch-is-sup u u<x is-sup-c fc ⟫ = IsSup.x<sup is-sup ⟪ az , ch-is-sup u (subst (λ k → u o< k) (Oprev.oprev=x op) (ordtrans u<x <-osuc)) is-sup-c fc ⟫ ... | tri≈ ¬a refl ¬c = ? ... | tri> ¬a ¬b c = ⊥-elim ( ¬p<x<op ⟪ c , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ ) order : {b sup1 z1 : Ordinal} → b o< x → sup1 o< b → FClosure A f (supf0 sup1) z1 → z1 << supf0 b order {b} {s} {z1} b<x s<b fc with trio< b px ... | tri< a ¬b ¬c = ZChain.order zc a s<b fc ... | tri≈ ¬a refl ¬c = ? ... | tri> ¬a ¬b c = ⊥-elim ( ¬p<x<op ⟪ c , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ ) zc4 : ZChain A f mf ay x zc4 with ODC.∋-p O A (* px) ... | no noapx = no-extension -- ¬ A ∋ p, just skip ... | yes apx with ODC.p∨¬p O ( HasPrev A (ZChain.chain zc ) apx f ) -- we have to check adding x preserve is-max ZChain A y f mf x ... | case1 pr = no-extension -- we have previous A ∋ z < x , f z ≡ x, so chain ∋ f z ≡ x because of f-next ... | case2 ¬fy<x with ODC.p∨¬p O (IsSup A (ZChain.chain zc ) apx ) ... | case1 is-sup = -- x is a sup of zc record { supf = psupf1 ; chain⊆A = ? ; f-next = ? ; f-total = ? ; initial = ? ; chain∋init = ? ; sup=u = ? ; order = ? } where psupf1 : Ordinal → Ordinal psupf1 z with trio< z x ... | tri< a ¬b ¬c = ZChain.supf zc z ... | tri≈ ¬a b ¬c = x ... | tri> ¬a ¬b c = x ... | case2 ¬x=sup = no-extension -- px is not f y' nor sup of former ZChain from y -- no extention ... | no lim = 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 = & A ... | tri> ¬a ¬b c = & A pchain : HOD pchain = UnionCF A f mf ay psupf0 x psupf0=pzc : {z : Ordinal} → (z<x : z o< x) → psupf0 z ≡ ZChain.supf (pzc z z<x) z psupf0=pzc {z} z<x with trio< z x ... | tri≈ ¬a b ¬c = ⊥-elim (¬a z<x) ... | tri> ¬a ¬b c = ⊥-elim (¬a z<x) ... | tri< a ¬b ¬c with o<-irr z<x a ... | refl = refl 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 ( (proj2 ca)) ( (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 , ch-init fc ⟫ = ⟪ proj2 ( mf a aa ) , ch-init (fsuc _ fc) ⟫ pnext {a} ⟪ aa , ch-is-sup u u<x is-sup fc ⟫ = ⟪ proj2 ( mf a aa ) , ch-is-sup u u<x (ChainP-next A f mf ay _ is-sup ) (fsuc _ fc) ⟫ pinit : {y₁ : Ordinal} → odef pchain y₁ → * y ≤ * y₁ pinit {a} ⟪ aa , ua ⟫ with ua ... | ch-init fc = s≤fc y f mf fc ... | ch-is-sup u u<x is-sup fc = ≤-ftrans (case2 zc7) (s≤fc _ f mf fc) where zc7 : y << psupf0 ? zc7 = ChainP.fcy<sup is-sup (init ay) pcy : odef pchain y pcy = ⟪ ay , ch-init (init ay) ⟫ no-extension : ZChain A f mf ay x no-extension = record { initial = pinit ; chain∋init = pcy ; supf = psupf0 ; chain⊆A = pchain⊆A ; f-next = pnext ; f-total = ptotal ; sup=u = ? ; order = ? } 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 = no-extension -- ¬ 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 = no-extension ... | case2 ¬fy<x with ODC.p∨¬p O (IsSup A pchain ax ) ... | case1 is-sup = record { initial = {!!} ; chain∋init = {!!} ; supf = psupf1 ; sup=u = ? ; order = ? ; chain⊆A = {!!} ; f-next = {!!} ; f-total = ? } where -- x is a sup of (zc ?) psupf1 : Ordinal → Ordinal psupf1 z with trio< z x ... | tri< a ¬b ¬c = ZChain.supf (pzc z a) z ... | tri≈ ¬a b ¬c = x ... | tri> ¬a ¬b c = x ... | case2 ¬x=sup = no-extension -- x is not f y' nor sup of former ZChain from y -- no extention SZ : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) → {y : Ordinal} (ay : odef A y) → ZChain A f mf ay (& A) SZ f mf {y} ay = TransFinite {λ z → ZChain A f mf ay z } (λ x → ind f mf ay x ) (& A) zorn00 : Maximal A zorn00 with is-o∅ ( & HasMaximal ) -- we have no Level (suc n) LEM ... | no not = record { maximal = ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ; 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