{-# OPTIONS --allow-unsolved-metas #-} open import Level hiding ( suc ; zero ) open import Ordinals open import Relation.Binary open import Relation.Binary.Core open import Relation.Binary.PropositionalEquality import OD module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) (PO : IsStrictPartialOrder _≡_ _<_ ) where -- -- Zorn-lemma : { A : HOD } -- → o∅ o< & A -- → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition -- → Maximal A -- open import zf open import logic -- open import partfunc {n} O open import Relation.Nullary open import Data.Empty import BAlgbra open import Data.Nat hiding ( _<_ ; _≤_ ) open import Data.Nat.Properties open import nat open inOrdinal O open OD O open OD.OD open ODAxiom odAxiom import OrdUtil import ODUtil open Ordinals.Ordinals O open Ordinals.IsOrdinals isOrdinal open Ordinals.IsNext isNext open OrdUtil O open ODUtil O import ODC open _∧_ open _∨_ open Bool open HOD -- -- Partial Order on HOD ( possibly limited in A ) -- _<<_ : (x y : Ordinal ) → Set n x << y = * x < * y _≤_ : (x y : Ordinal ) → Set n -- we can't use * x ≡ * y, it is Set (Level.suc n). Level (suc n) troubles Chain x ≤ y = (x ≡ y ) ∨ ( * x < * y ) POO : IsStrictPartialOrder _≡_ _<<_ POO = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans } ; trans = IsStrictPartialOrder.trans PO ; irrefl = λ x=y x ¬a ¬b c = ⊥-elim ( nat-≤> x ¬a ¬b c = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where fc12 : * y < * x fc12 = fcn-< {A} s {y} {x} {f} mf cy cx c -- open import Relation.Binary.Properties.Poset as Poset IsTotalOrderSet : ( A : HOD ) → Set (Level.suc n) IsTotalOrderSet A = {a b : HOD} → odef A (& a) → odef A (& b) → Tri (a < b) (a ≡ b) (b < a ) ⊆-IsTotalOrderSet : { A B : HOD } → B ⊆ A → IsTotalOrderSet A → IsTotalOrderSet B ⊆-IsTotalOrderSet {A} {B} B⊆A T ax ay = T (incl B⊆A ax) (incl B⊆A ay) _⊆'_ : ( A B : HOD ) → Set n _⊆'_ A B = {x : Ordinal } → odef A x → odef B x -- -- inductive masum tree from x -- tree structure -- record HasPrev (A B : HOD) ( f : Ordinal → Ordinal ) (x : Ordinal ) : Set n where field ax : odef A x y : Ordinal ay : odef B y x=fy : x ≡ f y record IsSUP (A B : HOD) (x : Ordinal ) : Set n where field ax : odef A x x≤sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x ) -- B is Total, use positive record SUP ( A B : HOD ) : Set (Level.suc n) where field sup : HOD isSUP : IsSUP A B (& sup) ax = IsSUP.ax isSUP x≤sup = IsSUP.x≤sup isSUP -- -- -- f (f ( ... (supf y))) f (f ( ... (supf z1))) -- / | / | -- / | / | -- supf y < supf z1 < supf z2 -- o< o< -- -- if sup z1 ≡ sup z2, the chain is stopped at sup z1, then f (sup z1) ≡ sup z1 -- this means sup z1 is the Maximal, so f is <-monotonic if we have no Maximal. -- fc-stop : ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) { a b : Ordinal } → (aa : odef A a ) →( {y : Ordinal} → FClosure A f a y → (y ≡ b ) ∨ (y << b )) → a ≡ b → f a ≡ a fc-stop A f mf {a} {b} aa x≤sup a=b with x≤sup (fsuc a (init aa refl )) ... | case1 eq = trans eq (sym a=b) ... | case2 lt = ⊥-elim (<-irr (case1 (cong (λ k → * (f k) ) (sym a=b))) (ftrans<-≤ lt fc00 ) ) where fc00 : b ≤ (f b) fc00 = proj1 (mf _ (subst (λ k → odef A k) a=b aa )) ∈∧P→o< : {A : HOD } {y : Ordinal} → {P : Set n} → odef A y ∧ P → y o< & A ∈∧P→o< {A } {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p ))) -- Union of supf z and FClosure A f y data UChain { A : HOD } { f : Ordinal → Ordinal } {supf : Ordinal → Ordinal} {y : Ordinal } (ay : odef A y ) (x : Ordinal) : (z : Ordinal) → Set n where ch-init : {z : Ordinal } (fc : FClosure A f y z) → UChain ay x z ch-is-sup : (u : Ordinal) {z : Ordinal } (u ¬a ¬b y sx ¬a ¬b y sx ¬a ¬b c with ≤-ftrans (order (o<→≤ sua (λ lt → <-irr (case2 b (λ lt → <-irr (case2 b ( IsMinSUP.minsup (is-minsup b≤z) asupf (λ {z} uzb → IsMinSUP.x≤sup (is-minsup a≤z) (subst (λ k → odef k z) (sym ua=ub) uzb)) ) sa ¬a ¬b sb ( IsMinSUP.minsup (is-minsup a≤z) asupf (λ {z} uza → IsMinSUP.x≤sup (is-minsup b≤z) (subst (λ k → odef k z) ua=ub uza)) ) sb z47 sb ¬a ¬b b z48 b b≤sa ( supf-inject ( osucprev ( begin osuc (supf (supf a)) ≡⟨ cong osuc (supf-idem (ordtrans (supf-inject sa ( begin supfb x ≡⟨ sbx=spb ⟩ spb ≤⟨ MinSUP.minsup mb (MinSUP.as ma) (λ {z} uzb → MinSUP.x≤sup ma (z53 uzb)) ⟩ spa ≡⟨ sym sax=spa ⟩ supfa x ∎ ) a ) where open o≤-Reasoning O z53 : {z : Ordinal } → odef (UnionCF A f ay (ZChain.supf zb) x) z → odef (UnionCF A f ay (ZChain.supf za) x) z z53 ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc ⟫ z53 {z} ⟪ as , ch-is-sup u u ¬a ¬b c = ⊥-elim ( o≤> ( begin supfa x ≡⟨ sax=spa ⟩ spa ≤⟨ MinSUP.minsup ma (MinSUP.as mb) (λ uza → MinSUP.x≤sup mb (z53 uza)) ⟩ spb ≡⟨ sym sbx=spb ⟩ supfb x ∎ ) c ) where open o≤-Reasoning O z53 : {z : Ordinal } → odef (UnionCF A f ay (ZChain.supf za) x) z → odef (UnionCF A f ay (ZChain.supf zb) x) z z53 ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc ⟫ z53 {z} ⟪ as , ch-is-sup u u ¬a ¬b c = case1 ( λ lt → ⊥-elim ( ¬a lt ) ) ... | tri< a ¬b ¬c with prev z a ... | case2 mins = case2 mins ... | case1 not with ODC.p∨¬p O (odef A z ∧ xsup z) ... | case1 mins = case2 record { sup = z ; isMinSUP = record { as = proj1 mins ; x≤sup = proj2 mins ; minsup = m04 } } where m04 : {sup1 : Ordinal} → odef A sup1 → ({w : Ordinal} → odef B w → (w ≡ sup1) ∨ (w << sup1)) → z o≤ sup1 m04 {s} as lt with trio< z s ... | tri< a ¬b ¬c = o<→≤ a ... | tri≈ ¬a b ¬c = o≤-refl0 b ... | tri> ¬a ¬b s ¬a ¬b c = case2 (o<→≤ c) -- -- 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 = zc41 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 ¬a ¬b px (<-irr (case2 lt1)) (λ eq → <-irr (case1 eq) lt1) lt1 where lt1 : a0 < b0 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt ptotal (case2 a) (case2 b) = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (fcn-cmp (supf0 px) f mf (proj1 a) (proj1 b)) pcha : pchainpx ⊆' A pcha (case1 lt) = proj1 lt pcha (case2 fc) = A∋fc _ f mf (proj1 fc) sup1 : MinSUP A pchainpx sup1 = minsupP pchainpx pcha ptotal sp1 = MinSUP.sup sup1 -- -- supf0 px o≤ sp1 -- sfpx≤sp1 : supf0 px o< x → supf0 px ≤ sp1 sfpx≤sp1 spx ¬a ¬b c = sp1 sf1=sf0 : {z : Ordinal } → z o≤ px → supf1 z ≡ supf0 z sf1=sf0 {z} z≤px with trio< z px ... | tri< a ¬b ¬c = refl ... | tri≈ ¬a b ¬c = refl ... | tri> ¬a ¬b c = ⊥-elim ( o≤> z≤px c ) sf1=sp1 : {z : Ordinal } → px o< z → supf1 z ≡ sp1 sf1=sp1 {z} px px ¬a ¬b c = refl sf=eq : { z : Ordinal } → z o< x → supf0 z ≡ supf1 z sf=eq {z} z ¬a ¬b c = MinSUP.as sup1 supf1-mono : {a b : Ordinal } → a o≤ b → supf1 a o≤ supf1 b supf1-mono {a} {b} a≤b with trio< b px ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (o<→≤ (ordtrans≤-< a≤b a)))) refl ( supf-mono a≤b ) ... | tri≈ ¬a b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (subst (λ k → a o≤ k) b a≤b))) refl ( supf-mono a≤b ) supf1-mono {a} {b} a≤b | tri> ¬a ¬b c with trio< a px ... | tri< a ¬a ¬b c = o≤-refl 0 ¬a ¬b c = ⊥-elim ( o≤> b≤x c ) -- x o< b z51 : FClosure A f (supf1 px) w z51 = subst (λ k → FClosure A f k w) (sym (trans (cong supf1 (sym a=px)) (sf1=sf0 (o≤-refl0 a=px) ))) fc z53 : odef A w z53 = A∋fc {A} _ f mf fc csupf1 : odef (UnionCF A f ay supf1 b) w csupf1 with zc43 px (supf0 px) ... | case2 spx≤px = ⟪ z53 , ch-is-sup (supf0 px) z54 z52 fc1 ⟫ where z54 : supf0 px o< b z54 = subst (λ k → supf0 px o< k ) (trans (Oprev.oprev=x op) (sym b=x) ) spx≤px z52 : supf1 (supf0 px) ≡ supf0 px z52 = trans (sf1=sf0 spx≤px ) ( ZChain.supf-idem zc o≤-refl spx≤px ) fc1 : FClosure A f (supf1 (supf0 px)) w fc1 = subst (λ k → FClosure A f k w ) (trans (cong supf0 a=px) (sym z52) ) fc ... | case1 px ¬a ¬b c = ⊥-elim ( ¬p ¬a ¬b c = ⊥-elim ( ¬p ( IsMinSUP.minsup (is-minsup b≤z) asupf1 (λ {z} uzb → IsMinSUP.x≤sup (is-minsup a≤z) (subst (λ k → odef k z) (sym ua=ub) uzb)) ) sa ¬a ¬b sb ( IsMinSUP.minsup (is-minsup a≤z) asupf1 (λ {z} uza → IsMinSUP.x≤sup (is-minsup b≤z) (subst (λ k → odef k z) ua=ub uza)) ) sb ¬a ¬b c = ⊥-elim ( ¬p ¬a ¬b px ¬a ¬b c = ⊥-elim ( o≤> b≤x c ) -- x o< b ≤ x a≤x : a o≤ x a≤x = o<→≤ ( subst (λ k → a o< k ) b=x a ¬a ¬b 0 ¬a ¬b c = ⊥-elim ( o≤> ( begin ZChain.supf (pzc (pic ia≤ib ic01 ) where ic02 : o∅ o< supfz i ¬a ¬b ib (<-irr (case2 lt1)) (λ eq → <-irr (case1 eq) lt1) lt1 where lt1 : a0 < b0 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt ptotalS (case2 a) (case2 b) = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (fcn-cmp spu f mf (proj1 a) (proj1 b)) S⊆A : pchainS ⊆' A S⊆A (case1 lt) = proj1 lt S⊆A (case2 fc) = A∋fc _ f mf (proj1 fc) ssup : MinSUP A pchainS ssup = minsupP pchainS S⊆A ptotalS sps = MinSUP.sup ssup supf1 : Ordinal → Ordinal supf1 z with trio< z x ... | tri< a ¬b ¬c = ZChain.supf (pzc (ob ¬a ¬b c = sps -- sup of spu which o< x -- if x o< spu, spu is not included in UnionCF x -- the chain pchain : HOD pchain = UnionCF A f ay supf1 x -- pchain ⊆ pchainU ⊆ pchianS sf1=sf : {z : Ordinal } → (a : z o< x ) → supf1 z ≡ ZChain.supf (pzc (ob ¬a ¬b c = ⊥-elim (¬a z ¬a ¬b c = ⊥-elim (¬b (sym eq)) sf1=sps : {z : Ordinal } → (a : x o< z ) → supf1 z ≡ sps sf1=sps {z} x x ¬a ¬b c = refl zc11 : {z : Ordinal } → odef pchain z → odef pchainS z zc11 {z} lt = ? sfpx≤spu : {z : Ordinal } → supf1 z ≤ sps sfpx≤spu {z} with trio< z x ... | tri< a ¬b ¬c = ? -- MinSUP.x≤sup usup ? -- (init (ZChain.asupf (pzc (ob ¬a ¬b c = case1 ? sfpxo≤spu : {z : Ordinal } → supf1 z o≤ sps sfpxo≤spu {z} with trio< z x ... | tri< a ¬b ¬c = ? -- MinSUP.x≤sup usup ? -- (init (ZChain.asupf (pzc (ob ¬a ¬b c = o≤-refl0 ? asupf : {z : Ordinal } → odef A (supf1 z) asupf {z} with trio< z x ... | tri< a ¬b ¬c = ZChain.asupf (pzc (ob ¬a ¬b c = MinSUP.as ssup supf-mono : {z y : Ordinal } → z o≤ y → supf1 z o≤ supf1 y supf-mono {z} {y} z≤y with trio< y x ... | tri< y ¬a ¬b c = zc01 where -- supf1 z o≤ sps zc01 : supf1 z o≤ sps zc01 with trio< z x ... | tri< z ¬a ¬b c = o≤-refl -- (sf1=sps c) 0 ¬a ¬b c | record { eq = eq } = ob ¬a ¬b c = ⊥-elim ( ¬a (ordtrans<-≤ a ¬a ¬b c = ⊥-elim z17 where z15 : (f sp ≡ MinSUP.sup sp1) ∨ (* (f sp) < * (MinSUP.sup sp1) ) z15 = MinSUP.x≤sup sp1 (ZChain.f-next zc z12 ) z17 : ⊥ z17 with z15 ... | case1 eq = ¬b (cong (*) eq) ... | case2 lt = ¬a lt -- ZChain contradicts ¬ Maximal -- -- ZChain forces fix point on any ≤-monotonic function (fixpoint) -- ¬ Maximal create cf which is a <-monotonic function by axiom of choice. This contradicts fix point of ZChain -- z04 : (nmx : ¬ Maximal A ) → (zc : ZChain A (cf nmx) (cf-is-<-monotonic nmx) as0 (& A)) → ⊥ z04 nmx zc = <-irr0 {* (cf nmx c)} {* c} (subst (λ k → odef A k ) (sym &iso) (proj1 (is-cf nmx (MinSUP.as msp1 )))) (subst (λ k → odef A k) (sym &iso) (MinSUP.as msp1) ) (case1 ( cong (*)( fixpoint (cf nmx) (cf-is-≤-monotonic nmx ) (cf-is-<-monotonic nmx ) zc msp1 ))) -- x ≡ f x ̄ (proj1 (cf-is-<-monotonic nmx c (MinSUP.as msp1 ))) where -- x < f x supf = ZChain.supf zc msp1 : MinSUP A (ZChain.chain zc) msp1 = msp0 (cf nmx) (cf-is-<-monotonic nmx) as0 zc c : Ordinal c = MinSUP.sup msp1 zorn00 : Maximal A zorn00 with is-o∅ ( & HasMaximal ) -- we have no Level (suc n) LEM ... | no not = record { maximal = ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ; as = zorn01 ; ¬maximal