{-# 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 _<=_ : (x y : Ordinal ) → Set n -- Set n order 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) {x : Ordinal } (xa : odef A x) ( f : Ordinal → Ordinal ) : Set n where field 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 (λ 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≤x supa fca) (ch-init fcb) | case2 lt = tri> (λ 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 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 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 refl) ⟫ Zorn-lemma : { A : HOD } → o∅ o< & A → ( ( B : HOD) → (B⊆A : B ⊆' A) → IsTotalOrderSet B → SUP A B ) -- SUP condition → Maximal A Zorn-lemma {A} 0 ¬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 = sp1 pchain1 : HOD pchain1 = UnionCF A f mf ay supf1 x pcy1 : odef pchain1 y pcy1 = ⟪ ay , ch-init (init ay refl) ⟫ pinit1 : {y₁ : Ordinal} → odef pchain1 y₁ → * y ≤ * y₁ pinit1 {a} ⟪ aa , ua ⟫ with ua ... | ch-init fc = s≤fc y f mf fc ... | ch-is-sup u u≤x is-sup fc = ≤-ftrans (<=to≤ zc7) (s≤fc _ f mf fc) where zc7 : y <= supf1 u zc7 = ChainP.fcy ¬a ¬b c = SUP⊆ (λ lt → chain-mono f mf ay _ ? (UnionCF⊆ ? lt )) sup1 sup=u : {b : Ordinal} (ab : odef A b) → b o≤ x → IsSup A (UnionCF A f mf ay supf1 (osuc b)) ab → supf1 b ≡ b sup=u {b} ab b ¬a ¬b c = ? zc4 : ZChain A f mf ay x zc4 with ODC.∋-p O A (* x) ... | no noax = no-extension {!!} -- ¬ A ∋ p, just skip ... | yes ax with ODC.p∨¬p O ( HasPrev A (ZChain.chain zc ) ax f ) -- we have to check adding x preserve is-max ZChain A y f mf x ... | case1 pr = no-extension {!!} -- we have previous A ∋ z < x , f z ≡ x, so chain ∋ f z ≡ x because of f-next ... | case2 ¬fy ¬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 = ysp 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) supf1 : Ordinal → Ordinal supf1 z with trio< z x ... | tri< a ¬b ¬c = ZChain.supf (pzc (osuc z) (ob ¬a ¬b c = spu pchain : HOD pchain = UnionCF A f mf ay supf1 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 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 (<=to≤ zc7) (s≤fc _ f mf fc) where zc7 : y <= supf1 _ zc7 = ChainP.fcy ¬a ¬b c = SUP⊆ (UnionCF0⊆ ?) usup sis : {z : Ordinal} (x≤z : z o≤ x) → supf1 z ≡ & (SUP.sup (sup x≤z)) sis {z} z≤x with trio< z x ... | tri< a ¬b ¬c = {!!} where zc8 = ZChain.supf-is-sup (pzc z a) o≤-refl ... | tri≈ ¬a b ¬c = refl ... | tri> ¬a ¬b c with osuc-≡< z≤x ... | case1 eq = ⊥-elim ( ¬b eq ) ... | case2 lt = ⊥-elim ( ¬a lt ) sup=u : {b : Ordinal} (ab : odef A b) → b o≤ x → IsSup A (UnionCF A f mf ay supf1 (osuc b)) ab → supf1 b ≡ b sup=u {b} ab b ¬a ¬b c = {!!} csupf : {z : Ordinal} → z o≤ x → odef (UnionCF A f mf ay supf1 z) (supf1 z) csupf {z} z ¬a ¬b c = {!!} -- ⊥-elim (¬a z