{-# 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 maxmum 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 } (xa : odef A x) : Set n where field x≤sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x ) record IsMinSUP (A B : HOD) ( f : Ordinal → Ordinal ) {x : Ordinal } (xa : odef A x) : Set n where field x≤sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x ) minsup : { sup1 : Ordinal } → odef A sup1 → ( {z : Ordinal } → odef B z → (z ≡ sup1 ) ∨ (z << sup1 )) → x o≤ sup1 not-hp : ¬ ( HasPrev A B f x ) record SUP ( A B : HOD ) : Set (Level.suc n) where field sup : HOD as : 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 sup of y not ϕ -- record ChainP (A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal) (u : Ordinal) : Set n where field fcy (λ 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 (λ lt → <-irr (case2 ct01) lt) (λ eq → <-irr (case1 eq) ct01) ct01 where ct00 : * b < * (supf ua) ct00 = lt 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 ¬a ¬b c with ChainP.order supa (subst₂ (λ j k → j o< k ) (sym (ChainP.supu=u supb )) (sym (ChainP.supu=u supa )) c) 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 ct00 = sym (trans (cong (*) eq) eq1) ... | case2 lt = tri> (λ lt → <-irr (case2 ct02) lt) (λ eq → <-irr (case1 eq) ct02) ct02 where ct02 : * b < * a ct02 = subst (λ k → * k < * a ) (sym eq) lt ct-ind xa xb {a} {b} (ch-is-sup ua ua ¬a ¬b c | case2 lt = tri> (λ lt → <-irr (case2 ct04) lt) (λ eq → <-irr (case1 (eq)) ct04) ct04 where ct05 : * b < * (supf ua) ct05 = lt 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 ∈∧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 which o< x -- 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 ; ¬a ¬b y sx ¬a ¬b c = ⊥-elim (<-irr (case2 sx ¬a ¬b c = ⊥-elim (<-irr (case2 sx ¬a ¬b y sx ( begin supfb x ≡⟨ sbx=spb ⟩ spb ≤⟨ MinSUP.minsup mb (MinSUP.asm 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 mf ay (ZChain.supf zb) x) z → odef (UnionCF A f mf 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.asm 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 mf ay (ZChain.supf za) x) z → odef (UnionCF A f mf ay (ZChain.supf zb) x) z z53 ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc ⟫ z53 {z} ⟪ as , ch-is-sup u u uc02 s ¬a ¬b c = ⊥-elim ( o≤> uc03 s ¬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 ; asm = 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 a b) pcha : pchainpx ⊆' A pcha (case1 lt) = proj1 lt pcha (case2 fc) = A∋fc _ f mf fc sup1 : MinSUP A pchainpx sup1 = minsupP pchainpx pcha ptotal sp1 = MinSUP.sup sup1 sfpx<=sp1 : supf0 px <= sp1 sfpx<=sp1 = MinSUP.x≤sup sup1 (case2 (init (ZChain.asupf zc {px}) refl )) sfpx≤sp1 : supf0 px o≤ sp1 sfpx≤sp1 = subst ( λ k → k o≤ sp1) (sym (ZChain.supf-is-minsup zc o≤-refl )) ( MinSUP.minsup (ZChain.minsup zc o≤-refl) (MinSUP.asm sup1) (λ {x} ux → MinSUP.x≤sup sup1 (case1 ux)) ) -- -- supf0 px o≤ sp1 -- zc41 : ZChain A f mf ay x zc41 with zc43 x sp1 zc41 | (case2 sp≤x ) = record { supf = supf1 ; sup=u = ? ; asupf = ? ; supf-mono = supf1-mono ; supf-< = ? ; supfmax = ? ; minsup = ? ; supf-is-minsup = ? ; cfcs = ? } where supf1 : Ordinal → Ordinal supf1 z with trio< z px ... | tri< a ¬b ¬c = supf0 z ... | tri≈ ¬a b ¬c = supf0 z ... | tri> ¬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.asm 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 sf≤ : { z : Ordinal } → x o≤ z → supf0 x o≤ supf1 z sf≤ {z} x≤z with trio< z px ... | tri< a ¬b ¬c = ⊥-elim ( o<> (osucc a) (subst (λ k → k o≤ z) (sym (Oprev.oprev=x op)) x≤z ) ) ... | tri≈ ¬a b ¬c = ⊥-elim ( o≤> x≤z (subst (λ k → k o< x ) (sym b) px ¬a ¬b c = subst₂ (λ j k → j o≤ k ) (trans (sf1=sf0 o≤-refl ) (sym (ZChain.supfmax zc px ¬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 mf ay supf1 b) w csupf1 with trio< (supf0 px) x ... | tri< sfpx sp≤x (subst (λ k → k o< sp1) spx=x lt )) -- supf0 px o< sp1 , x o< sp1 m10 : f (supf0 px) ≡ supf0 px m10 = fc-stop A f mf (ZChain.asupf zc) m11 m12 where m11 : {z : Ordinal} → FClosure A f (supf0 px) z → (z ≡ sp1) ∨ (z << sp1) m11 {z} fc = MinSUP.x≤sup sup1 (case2 fc) ... | tri> ¬a ¬b c = ⊥-elim ( o<¬≡ refl (ordtrans<-≤ c (OrdTrans sfpx≤sp1 sp≤x))) -- x o< supf0 px o≤ sp1 ≤ x ... | tri> ¬a ¬b c = ⊥-elim ( ¬p ¬a ¬b c | _ = ⊥-elim ( ¬p ¬a ¬b px ¬a ¬b px ¬a ¬b c = ⊥-elim ( o≤> ? c ) ... | tri≈ ¬a refl ¬c = ? -- supf0 px o< x → odef (UnionCF A f mf ay supf0 x) (supf0 px) -- x o≤ supf0 px o≤ sp → zc17 : {z : Ordinal } → supf0 z o≤ supf0 px zc17 {z} with trio< z px ... | tri< a ¬b ¬c = ZChain.supf-mono zc (o<→≤ a) ... | tri≈ ¬a b ¬c = o≤-refl0 (cong supf0 b) ... | tri> ¬a ¬b px ¬a ¬b px ¬a ¬b c = zc36 ¬b ... | tri≈ ¬a b ¬c = record { tsup = ? ; tsup=sup = ? } where zc37 : MinSUP A (UnionCF A f mf ay supf0 z) zc37 = record { sup = ? ; asm = ? ; x≤sup = ? } sup=u : {b : Ordinal} (ab : odef A b) → b o≤ x → IsMinSUP A (UnionCF A f mf ay supf0 b) supf0 ab ∧ (¬ HasPrev A (UnionCF A f mf ay supf0 b) f b ) → supf0 b ≡ b sup=u {b} ab b≤x is-sup with trio< b px ... | tri< a ¬b ¬c = ZChain.sup=u zc ab (o<→≤ a) ⟪ record { x≤sup = λ lt → IsMinSUP.x≤sup (proj1 is-sup) lt } , proj2 is-sup ⟫ ... | tri≈ ¬a b ¬c = ZChain.sup=u zc ab (o≤-refl0 b) ⟪ record { x≤sup = λ lt → IsMinSUP.x≤sup (proj1 is-sup) lt } , proj2 is-sup ⟫ ... | tri> ¬a ¬b px ¬a ¬b c = ? usup : MinSUP A pchainx usup = minsupP pchainx (λ lt → ? ) ptotalx spu = MinSUP.sup usup supf1 : Ordinal → Ordinal supf1 z with trio< z x ... | tri< a ¬b ¬c = ZChain.supf (pzc (ob ¬a ¬b c = spu pchain : HOD pchain = UnionCF A f mf ay supf1 x -- pchain ⊆ pchainx 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 supf1 ( (proj2 ca)) ( (proj2 cb)) sf1=sf : {z : Ordinal } → (a : z o< x ) → supf1 z ≡ ZChain.supf (pzc (ob ¬a ¬b c = ⊥-elim (¬a z x≤z a) ... | tri≈ ¬a b ¬c = refl ... | tri> ¬a ¬b c = refl -- zc11 : {z : Ordinal } → (a : z o< x ) → odef pchain (ZChain.supf (pzc (ob ¬a ¬b c = case1 refl sfpx≤spu : {z : Ordinal } → supf1 z o≤ spu sfpx≤spu {z} with trio< z x ... | tri< a ¬b ¬c = subst ( λ k → k o≤ spu) ? ( MinSUP.minsup (ZChain.minsup ? o≤-refl) ? (λ {x} ux → MinSUP.x≤sup ? ?) ) ... | tri≈ ¬a b ¬c = ? ... | tri> ¬a ¬b c = ? supf-mono : {x y : Ordinal } → x o≤ y → supf1 x o≤ supf1 y supf-mono {x} {y} x≤y with trio< y x ... | tri< a ¬b ¬c = ? ... | tri≈ ¬a b ¬c = ? ... | tri> ¬a ¬b c = ? zc5 : ZChain A f mf ay x zc5 = ? where cfcs : (mf< : <-monotonic-f A f) {a b w : Ordinal } → a o< b → b o≤ x → supf1 a o< x → FClosure A f (supf1 a) w → odef (UnionCF A f mf ay supf1 b) w cfcs mf< {a} {b} {w} a ¬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 tri : {n : Level} (u w : Ordinal ) { R : Set n } → ( u o< w → R ) → ( u ≡ w → R ) → ( w o< u → R ) → R tri {_} u w p q r with trio< u w ... | tri< a ¬b ¬c = p a ... | tri≈ ¬a b ¬c = q b ... | tri> ¬a ¬b c = r c or : {n m r : Level } {P : Set n } {Q : Set m} {R : Set r} → P ∨ Q → ( P → R ) → (Q → R ) → R or (case1 p) p→r q→r = p→r p or (case2 q) p→r q→r = q→r q -- 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.asm msp1 )))) (subst (λ k → odef A k) (sym &iso) (MinSUP.asm 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.asm 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