{-# OPTIONS --allow-unsolved-metas #-} open import Level hiding ( suc ; zero ) open import Ordinals import OD module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) where open import zf open import logic -- open import partfunc {n} O open import Relation.Nullary open import Relation.Binary open import Data.Empty open import Relation.Binary open import Relation.Binary.Core open import Relation.Binary.PropositionalEquality import BAlgbra 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 record Element (A : HOD) : Set (Level.suc n) where field elm : HOD is-elm : A ∋ elm open Element _ & x record OSup> (A : HOD) {x : Ordinal} (ax : A ∋ * x) : Set n where field y : Ordinal icy : odef (IChainSet {A} (me ax)) y y>x : x o< y record IChainSup> (A : HOD) {x : Ordinal} (ax : A ∋ * x) : Set n where field y : Ordinal A∋y : odef A y y>x : * x < * y -- finite IChain ic→A∋y : (A : HOD) {x y : Ordinal} (ax : A ∋ * x) → odef (IChainSet {A} (me ax)) y → A ∋ * y ic→A∋y A {x} {y} ax ⟪ ay , _ ⟫ = subst (λ k → odef A k) (sym &iso) ay record InFiniteIChain (A : HOD) {x : Ordinal} (ax : A ∋ * x) : Set n where field chain A (ic→A∋y A ax cy) open import Data.Nat hiding (_<_) import Data.Nat.Properties as NP open import nat data Chain (A : HOD) (s : Ordinal) (next : Ordinal → Ordinal ) : ( x : Ordinal ) → Set n where cfirst : odef A s → Chain A s next s csuc : (x : Ordinal ) → (ax : odef A x ) → Chain A s next x → odef A (next x) → Chain A s next (next x ) ct∈A : (A : HOD ) (s : Ordinal) → (next : Ordinal → Ordinal ) → {x : Ordinal} → Chain A s next x → odef A x ct∈A A s next {x} (cfirst x₁) = x₁ ct∈A A s next {.(next x )} (csuc x ax t anx) = anx -- -- extract single chain from countable infinite chains -- ChainClosure : (A : HOD) (s : Ordinal) → (next : Ordinal → Ordinal ) → HOD ChainClosure A s next = record { od = record { def = λ x → Chain A s next x } ; odmax = & A ; .y ( InFiniteIChain.c-infinite ifc y (subst (λ k → odef (IChainSet (me ax)) k) &iso ics-y )) ... | no _ = o∅ InFCSet : (A : HOD) → {x : Ordinal} (ax : A ∋ * x) → (ifc : InFiniteIChain A ax ) → HOD InFCSet A {x} ax ifc = ChainClosure (IChainSet {A} (me ax)) x (cinext A ax ifc ) InFCSet⊆A : (A : HOD) → {x : Ordinal} (ax : A ∋ * x) → (ifc : InFiniteIChain A ax ) → InFCSet A ax ifc ⊆ A InFCSet⊆A A {x} ax ifc = record { incl = λ {y} lt → incl (IChainSet⊆A (me ax)) ( ct∈A (IChainSet {A} (me ax)) x (cinext A ax ifc) lt ) } ChainClosure-is-total : (A : HOD) → {x : Ordinal} (ax : A ∋ * x) → IsPartialOrderSet A → (ifc : InFiniteIChain A ax ) → IsTotalOrderSet ( InFCSet A ax ifc ) ChainClosure-is-total A {x} ax PO ifc = record { isEquivalence = IsStrictPartialOrder.isEquivalence IPO ; trans = λ {x} {y} {z} → IsStrictPartialOrder.trans IPO {x} {y} {z} ; compare = cmp } where IPO : IsPartialOrderSet (InFCSet A ax ifc ) IPO = ⊆-IsPartialOrderSet record { incl = λ {y} lt → incl (InFCSet⊆A A {x} ax ifc) lt} PO B = IChainSet {A} (me ax) cnext = cinext A ax ifc ct02 : {oy : Ordinal} → (y : Chain B x cnext oy ) → A ∋ * oy ct02 y = incl (IChainSet⊆A {A} (me ax)) (subst (λ k → odef (IChainSet (me ax)) k) (sym &iso) (ct∈A B x cnext y) ) ct-inject : {ox oy : Ordinal} → (x1 : Chain B x cnext ox ) → (y : Chain B x cnext oy ) → (cton0 B x cnext x1) ≡ (cton0 B x cnext y) → ox ≡ oy ct-inject {ox} {ox} (cfirst x) (cfirst x₁) refl = refl ct-inject {.(cnext x₀ )} {.(cnext x₃ )} (csuc x₀ ax x₁ x₂) (csuc x₃ ax₁ y x₄) eq = cong cnext ct05 where ct06 : {x y : ℕ} → suc x ≡ suc y → x ≡ y ct06 refl = refl ct05 : x₀ ≡ x₃ ct05 = ct-inject x₁ y (ct06 eq) ct-monotonic : {ox oy : Ordinal} → (x1 : Chain B x cnext ox ) → (y : Chain B x cnext oy ) → (cton0 B x cnext x1) Data.Nat.< (cton0 B x cnext y) → * ox < * oy ct-monotonic {ox} {oy} x1 (csuc oy1 ay y _) (s≤s lt) with NP.<-cmp ( cton0 B x cnext x1 ) ( cton0 B x cnext y ) ... | tri< a ¬b ¬c = ct07 where ct07 : * ox < * (cnext oy1) ct07 with ODC.∋-p O (IChainSet {A} (me ax)) (* oy1) ... | no not = ⊥-elim ( not (subst (λ k → odef (IChainSet {A} (me ax)) k ) (sym &iso) ay ) ) ... | yes ay1 = IsStrictPartialOrder.trans PO {me (ct02 x1) } {me (ct02 y)} {me ct031 } (ct-monotonic x1 y a ) ct011 where ct031 : A ∋ * (IChainSup>.y (InFiniteIChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet {A} (me ax)) k) &iso ay1 ) )) ct031 = subst (λ k → odef A k ) (sym &iso) ( IChainSup>.A∋y (InFiniteIChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet {A} (me ax)) k) &iso ay1 )) ) ct011 : * oy1 < * ( IChainSup>.y (InFiniteIChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet {A} (me ax)) k) &iso ay1 )) ) ct011 = IChainSup>.y>x (InFiniteIChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet {A} (me ax)) k) &iso ay1 )) ... | tri≈ ¬a b ¬c = ct11 where ct11 : * ox < * (cnext oy1) ct11 with ODC.∋-p O (IChainSet {A} (me ax)) (* oy1) ... | no not = ⊥-elim ( not (subst (λ k → odef (IChainSet {A} (me ax)) k ) (sym &iso) ay ) ) ... | yes ay1 = subst (λ k → * k < _) (sym (ct-inject _ _ b)) ct011 where ct011 : * oy1 < * ( IChainSup>.y (InFiniteIChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet {A} (me ax)) k) &iso ay1 )) ) ct011 = IChainSup>.y>x (InFiniteIChain.c-infinite ifc oy1 (subst (λ k → odef (IChainSet {A} (me ax)) k) &iso ay1 )) ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> lt c ) ct12 : {y z : Element (ChainClosure B x cnext) } → elm y ≡ elm z → elm y < elm z → ⊥ ct12 {y} {z} y=z yz = IsStrictPartialOrder.irrefl IPO {y} {y} refl ( IsStrictPartialOrder.trans IPO {y} {z} {y} yz ) ct17 : (x1 : Element (ChainClosure B x cnext)) → Chain B x cnext (& (elm x1)) ct17 x1 = is-elm x1 cmp : Trichotomous _ _ cmp x1 y with NP.<-cmp (cton B x cnext x1) (cton B x cnext y) ... | tri< a ¬b ¬c = tri< ct04 ct14 ct15 where ct04 : elm x1 < elm y ct04 = subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm x1) (is-elm y) a) ct14 : ¬ elm x1 ≡ elm y ct14 eq = ct12 {x1} {y} eq (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm x1) (is-elm y) a ) ) ct15 : ¬ (elm y < elm x1) ct15 lt = ct13 {y} {x1} lt (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm x1) (is-elm y) a ) ) ... | tri≈ ¬a b ¬c = tri≈ (ct12 {x1} {y} ct16) ct16 (ct12 {y} {x1} (sym ct16)) where ct16 : elm x1 ≡ elm y ct16 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) (ct-inject {& (elm x1)} {& (elm y)} (is-elm x1) (is-elm y) b )) ... | tri> ¬a ¬b c = tri> ct15 ct14 ct04 where ct04 : elm y < elm x1 ct04 = subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm y) (is-elm x1) c) ct14 : ¬ elm x1 ≡ elm y ct14 eq = ct12 {y} {x1} (sym eq) (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm y) (is-elm x1) c ) ) ct15 : ¬ (elm x1 < elm y) ct15 lt = ct13 {x1} {y} lt (subst₂ (λ j k → j < k ) *iso *iso (ct-monotonic (is-elm y) (is-elm x1) c ) ) record IsFC (A : HOD) {x : Ordinal} (ax : A ∋ * x) (y : Ordinal) : Set n where field icy : odef (IChainSet {A} (me ax)) y c-finite : ¬ IChainSup> A (subst (λ k → odef A k ) (sym &iso) (proj1 icy) ) record Maximal ( A : HOD ) : Set (Level.suc n) where field maximal : HOD A∋maximal : A ∋ maximal ¬maximal x (will contradic in the transfinite induction ) -- case 2) no > x in some chain ( maximal ) -- case 3) countably infinite chain below x (will be prohibited by sup condtion ) -- Zorn-lemma-3case : { A : HOD } → o∅ o< & A → IsPartialOrderSet A → (x : Element A) → OSup> A (d→∋ A (is-elm x)) ∨ Maximal A ∨ InFiniteIChain A (d→∋ A (is-elm x)) Zorn-lemma-3case {A} 0 A (d→∋ A (is-elm x)) ∨ Maximal A ∨ InFiniteIChain A (d→∋ A (is-elm x)) zc2 with is-o∅ (& Gtx) ... | no not = case1 record { y = & y ; icy = zc4 ; y>x = proj2 zc3 } where y : HOD y = ODC.minimal O Gtx (λ eq → not (=od∅→≡o∅ eq)) zc3 : odef ( IChainSet x ) (& y) ∧ ( & (elm x) o< (& y )) zc3 = ODC.x∋minimal O Gtx (λ eq → not (=od∅→≡o∅ eq)) zc4 : odef (IChainSet (me (subst (OD.def (od A)) (sym &iso) (is-elm x)))) (& y) zc4 = ⟪ proj1 (proj1 zc3) , subst (λ k → IChained A (& k) (& y) ) (sym *iso) (proj2 (proj1 zc3)) ⟫ ... | yes nogt with is-o∅ (& HG) ... | no finite-chain = case2 (case1 record { maximal = y ; A∋maximal = proj1 zc3 ; ¬maximalx = subst₂ (λ j k → j < k ) (sym *iso) (sym *iso) y A (ic→A∋y A (subst (OD.def (od A)) (sym &iso) (is-elm x)) cy) zc9 y cy with ODC.∋-p O B (* y) ... | no not = ⊥-elim (not (subst (λ k → odef B k ) (sym &iso) cy)) ... | yes cy1 with is-o∅ (& (Nx y (proj1 cy) )) ... | yes no-next = ⊥-elim zc16 where -- cy : OD.def (od A) y ∧ IChained A (& (* (& (elm x)))) y zc18 : ¬ IChainSup> A (subst (odef A) (sym &iso) (proj1 (subst (λ k → odef (IChainSet (me (d→∋ A (is-elm x)))) k) (sym &iso) cy))) zc18 ics = ¬A∋x→A≡od∅ (Nx y (proj1 cy) ) ⟪ subst (λ k → odef A k ) (sym &iso) (IChainSup>.A∋y ics) , subst₂ (λ j k → j < k ) *iso (cong (*) (sym &iso))( IChainSup>.y>x ics) ⟫ no-next zc17 : IsFC A {& (elm x)} (d→∋ A (is-elm x)) (& (* y)) zc17 = record { icy = subst (λ k → odef (IChainSet (me (d→∋ A (is-elm x)))) k ) (sym &iso) cy ; c-finite = zc18 } zc16 : ⊥ zc16 = ¬A∋x→A≡od∅ HG ⟪ subst (λ k → odef A k ) (sym &iso) (proj1 cy ) , zc17 ⟫ inifite ... | no not = record { y = & zc13 ; A∋y = proj1 zc12 ; y>x = proj2 zc12 } where zc13 = ODC.minimal O (Nx y (proj1 cy)) (λ eq → not (=od∅→≡o∅ eq )) zc12 : odef A (& zc13 ) ∧ ( * y < * ( & zc13 )) zc12 = ODC.x∋minimal O (Nx y (proj1 cy)) (λ eq → not (=od∅→≡o∅ eq )) all-climb-case : { A : HOD } → (0 A (d→∋ A (is-elm x) )) → InFiniteIChain A (d→∋ A (ODC.x∋minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0 ¬a ¬b c = {!!} -- can't happen ... | yes ax = zc4 where -- we have previous ordinal and A ∋ x px = Oprev.oprev op zc1 : OSup> A (subst (OD.def (od A)) (sym &iso) ax) → ZChain A x ∨ Maximal A zc1 os with prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc) ... | case2 x = case2 x ... | case1 x = {!!} zc4 : ZChain A x ∨ Maximal A zc4 with Zorn-lemma-3case 0x = zc1 y>x ... | case2 (case1 x) = case2 x ... | case2 (case2 x) = ⊥-elim (zc5 x) where FC : HOD FC = IChainSet {A} (me ax) B : InFiniteIChain A ax → HOD B ifc = InFCSet A ax ifc zc6 : (ifc : InFiniteIChain A ax ) → ¬ SUP A (B ifc) zc6 = {!!} FC-is-total : (ifc : InFiniteIChain A ax) → IsTotalOrderSet (B ifc) FC-is-total ifc = ChainClosure-is-total A ax PO ifc B⊆A : (ifc : InFiniteIChain A ax) → B ifc ⊆ A B⊆A = {!!} ifc : InFiniteIChain A (subst (OD.def (od A)) (sym &iso) ax) → InFiniteIChain A ax ifc record { c-infinite = c-infinite } = record { c-infinite = {!!} } where ifc01 : {!!} -- me (subst (OD.def (od A)) (sym &iso) ax) ifc01 = {!!} -- (y : Ordinal) → odef (IChainSet (me (subst (OD.def (od A)) (sym &iso) ax))) y → y o< & (* x₁) -- (y : Ordinal) → odef (IChainSet (me ax)) y → y o< x₁ zc5 : InFiniteIChain A (subst (OD.def (od A)) (sym &iso) ax) → ⊥ zc5 x = zc6 (ifc x) ( supP (B (ifc x)) (B⊆A (ifc x)) (FC-is-total (ifc x) )) ind x prev | no ¬ox with trio< (& A) x --- limit ordinal case ... | tri< a ¬b ¬c = {!!} where zc1 : ZChain A (& A) zc1 with prev (& A) a ... | t = {!!} ... | tri≈ ¬a b ¬c = {!!} where ... | tri> ¬a ¬b c with ODC.∋-p O A (* x) ... | no ¬Ax = {!!} where ... | yes ax = {!!} zorn03 : (x : Ordinal) → ZChain A x ∨ Maximal A zorn03 x = TransFinite ind x zorn04 : Maximal A zorn04 with zorn03 (& A) ... | case1 chain = ⊥-elim ( o<> (c<→o< {ZChain.max chain} {A} (ZChain.A∋max chain)) (ZChain.y