{-# 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 IsPartialOrderSet : ( A : HOD ) → Set (Level.suc n) IsPartialOrderSet A = IsStrictPartialOrder _≡A_ _ ¬a ¬b c = ⊥-elim ( nat-≤> (s≤s le) c ) IC-is-total : {A : HOD} → (ax : Element A ) → IsPartialOrderSet A → IsTotalOrderSet (IChainSet {A} ax) IC-is-total {A} ax PO = record { isEquivalence = IsStrictPartialOrder.isEquivalence IPO ; trans = λ {x} {y} {z} → IsStrictPartialOrder.trans IPO {x} {y} {z} ; compare = cmp } where IPO : IsPartialOrderSet (IChainSet {A} ax) IPO = ⊆-IsPartialOrderSet record { incl = λ {x} lt → proj1 lt } PO cmp : Trichotomous _ _ cmp x y with iclen A (IChained.iy (proj2 (is-elm x))) | iclen A (IChained.iy (proj2 (is-elm y))) ... | zero | zero = {!!} ... | zero | suc t = {!!} ... | suc s | zero = {!!} ... | suc s | suc t = {!!} -- there is a y, & y > & 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 record InFiniteIChain (A : HOD) {x : Ordinal} (ax : A ∋ * x) : Set n where field chain A ax 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 ax 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) zc6 } where zc8 : ic-connect (& (* (& (elm x)))) (IChained.iy (proj2 zc5)) zc8 = IChained.ic (proj2 zc5) zc7 : elm x < y zc7 = subst₂ (λ j k → j < k ) *iso *iso ( ic→< {A} PO (& (elm x)) (is-elm x) (IChained.iy (proj2 zc5)) (subst (λ k → ic-connect (& k) (IChained.iy (proj2 zc5)) ) (me-elm-refl A x) (IChained.ic (proj2 zc5)) ) ) zc6 : elm x < z zc6 = IsStrictPartialOrder.trans PO {x} {me (proj1 zc3)} {me az} zc7 y max cannot happen ) -- ¬ A ∋ y use previous chain -- A ∋ y is use oridinaly min of y or previous -- y is limit ordinal -- has maximal in some lower use this -- no maximal in all lower -- & A < y A is a counter example of assumption -- A ∋ y is maximal use this -- ¬ A ∋ y use previous chain -- check some y ≤ max -- if none A < y is the counter example -- else use the ordinaly smallest max as next chain ind x prev with Oprev-p x ... | yes op with ODC.∋-p O A (* x) ... | no ¬Ax = zc1 where -- we have previous ordinal and ¬ A ∋ x, use previous Zchain px = Oprev.oprev op zc1 : ZChain A x ∨ Maximal A zc1 with prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc) ... | case2 x = case2 x -- we have the Maximal ... | case1 z with trio< x (& (ZChain.max z)) ... | tri< a ¬b ¬c = case1 record { max = ZChain.max z ; y ¬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) zc6 : InFiniteIChain A (subst (OD.def (od A)) (sym &iso) ax) → ¬ SUP A FC zc6 = {!!} FC-is-total : IsTotalOrderSet FC FC-is-total = {!!} FC⊆A : FC ⊆ A FC⊆A = {!!} zc5 : InFiniteIChain A (subst (OD.def (od A)) (sym &iso) ax) → ⊥ zc5 x = zc6 x ( supP FC FC⊆A FC-is-total ) 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 = {!!} ... | no not = {!!} where 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