{-# 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 ¬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 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 ¬a ¬b c = ⊥-elim ( nat-≤> y ¬a ¬b c with fc20 c -- ncy < suc ncx ... | case1 y=x = <-irr (case1 ( fcn-inject s mf cy cx y=x )) x (λ lt → <-irr (case2 ct01) lt) (λ eq → <-irr (case1 eq) ct01) ct01 where ct00 : * b < * (supf ua) ct00 = ChainP.fcy ¬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 { supfu=u = ChainP.supfu=u cp ; fcy ¬a ¬b c = ⊥-elim (¬p ¬a ¬b c = ? -- u = x → u = sup → (b o< x → b < a ) → ⊥ m04 : odef (UnionCF A f mf ay (ZChain.supf zc) px) b m04 = ZChain1.is-max (prev px px ¬a ¬b c = ⊥-elim z17 where z15 : (* (f ( & ( SUP.sup sp1 ))) ≡ SUP.sup sp1) ∨ (* (f ( & ( SUP.sup sp1 ))) < SUP.sup sp1) z15 = SUP.x ¬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 ¬a ¬b c = & A -- pchain0 : HOD pchain0 = UnionCF A f mf ay psupf0 x ptotal0 : IsTotalOrderSet pchain0 ptotal0 {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)) usup : SUP A pchain0 usup = supP pchain0 (λ lt → proj1 lt) ptotal0 spu = & (SUP.sup usup) psupf : Ordinal → Ordinal psupf z with trio< z x ... | tri< a ¬b ¬c = ZChain.supf (pzc (osuc z) (ob ¬a ¬b c = spu psupf>z : {z : Ordinal } → x o< z → spu ≡ psupf z psupf>z {z} x ¬a ¬b c = refl psupf=x : spu ≡ psupf x psupf=x = zc20 refl where zc20 : {z : Ordinal } → z ≡ x → spu ≡ psupf x zc20 {z} z=x with trio< z x | inspect psupf z ... | tri< a ¬b ¬c | _ = ⊥-elim ( ¬b z=x) ... | tri≈ ¬a b ¬c | record { eq = eq1 } = subst (λ k → spu ≡ psupf k) b (sym eq1) ... | tri> ¬a ¬b c | _ = ⊥-elim ( ¬b z=x) csupf :{z : Ordinal} → odef (UnionCF A f mf ay psupf x) (psupf z) csupf {z} with trio< z x | inspect psupf z ... | tri< z ¬a ¬b c | record { eq = eq1 } = ⟪ SUP.A∋maximal usup , ch-is-sup x <-osuc zc16 (subst (λ k → FClosure A f k spu) psupf=x (init (SUP.A∋maximal usup))) ⟫ where zc16 : ChainP A f mf ay psupf x spu zc16 = ? pchain : HOD pchain = UnionCF A f mf ay psupf 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 , 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 << psupf ? zc7 = ChainP.fcy ¬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