# HG changeset patch # User Shinji KONO # Date 1671575418 -32400 # Node ID 87c2da3811c3b8c64a357c4e0942e211b1080b45 # Parent 63c1167b234376e0d042a79ac80e5e70d3985497 index version diff -r 63c1167b2343 -r 87c2da3811c3 src/zorn.agda --- a/src/zorn.agda Tue Dec 20 11:20:52 2022 +0900 +++ b/src/zorn.agda Wed Dec 21 07:30:18 2022 +0900 @@ -5,7 +5,7 @@ open import Relation.Binary.Core open import Relation.Binary.PropositionalEquality import OD hiding ( _⊆_ ) -module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) (PO : IsStrictPartialOrder _≡_ _<_ ) where +module zorn1 {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) (PO : IsStrictPartialOrder _≡_ _<_ ) where -- -- Zorn-lemma : { A : HOD } @@ -108,113 +108,6 @@ <-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set n <-monotonic-f A f = (x : Ordinal ) → odef A x → ( * x < * (f x) ) ∧ odef A (f x ) -data FClosure (A : HOD) (f : Ordinal → Ordinal ) (s : Ordinal) : Ordinal → Set n where - init : {s1 : Ordinal } → odef A s → s ≡ s1 → FClosure A f s s1 - fsuc : (x : Ordinal) ( p : FClosure A f s x ) → FClosure A f s (f x) - -A∋fc : {A : HOD} (s : Ordinal) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → odef A y -A∋fc {A} s f mf (init as refl ) = as -A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {A} s f mf fcy ) ) - -A∋fcs : {A : HOD} (s : Ordinal) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → odef A s -A∋fcs {A} s f mf (init as refl) = as -A∋fcs {A} s f mf (fsuc y fcy) = A∋fcs {A} s f mf fcy - -s≤fc : {A : HOD} (s : Ordinal ) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → s ≤ y -s≤fc {A} s {.s} f mf (init x refl ) = case1 refl -s≤fc {A} s {.(f x)} f mf (fsuc x fcy) with proj1 (mf x (A∋fc s f mf fcy ) ) -... | case1 x=fx = subst₂ (λ j k → j ≤ k ) refl x=fx (s≤fc s f mf fcy) -... | case2 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 ) @@ -243,63 +136,6 @@ ax = IsSUP.ax isSUP x≤sup = IsSUP.x≤sup isSUP --- --- Our Proof strategy of the Zorn Lemma --- --- f (f ( ... (supf y))) f (f ( ... (supf z1))) --- / | / | --- / | / | --- supf y < supf z1 < supf z2 --- o< o< --- --- if sup z1 ≡ sup z2, the chain is stopped at sup z1, then f (sup z1) ≡ sup z1 --- this means sup z1 is the Maximal, so f is <-monotonic if we have no Maximal. --- - -fc-stop : ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) { a b : Ordinal } - → (aa : odef A a ) →( {y : Ordinal} → FClosure A f a y → (y ≡ b ) ∨ (y << b )) → a ≡ b → f a ≡ a -fc-stop A f mf {a} {b} aa x≤sup a=b with x≤sup (fsuc a (init aa refl )) -... | case1 eq = trans eq (sym a=b) -... | case2 lt = ⊥-elim (<-irr (case1 (cong (λ k → * (f k) ) (sym a=b))) (ftrans<-≤ lt fc00 ) ) where - fc00 : b ≤ (f b) - fc00 = proj1 (mf _ (subst (λ k → odef A k) a=b aa )) - -∈∧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 and FClosure A f y - -data UChain { A : HOD } { f : Ordinal → Ordinal } {supf : Ordinal → Ordinal} {y : Ordinal } (ay : odef A y ) - (x : Ordinal) : (z : Ordinal) → Set n where - ch-init : {z : Ordinal } (fc : FClosure A f y z) → UChain ay x z - ch-is-sup : (u : Ordinal) {z : Ordinal } (u ¬a ¬b y sx ¬a ¬b y sx z47 sb ¬a ¬b b z48 b ( - IsMinSUP.minsup (is-minsup b≤z) asupf (λ {z} uzb → IsMinSUP.x≤sup (is-minsup a≤z) (subst (λ k → odef k z) (sym ua=ub) uzb)) ) sa ¬a ¬b sb ( - IsMinSUP.minsup (is-minsup a≤z) asupf (λ {z} uza → IsMinSUP.x≤sup (is-minsup b≤z) (subst (λ k → odef k z) ua=ub uza)) ) sb ¬a ¬b c with ≤-ftrans (order (o<→≤ sua (λ lt → <-irr (case2 b (λ lt → <-irr (case2 b ( - begin - supfb x ≡⟨ sbx=spb ⟩ - spb ≤⟨ MinSUP.minsup mb (MinSUP.as 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 ay (ZChain.supf zb) x) z → odef (UnionCF A f 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.as 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 ay (ZChain.supf za) x) z → odef (UnionCF A f ay (ZChain.supf zb) x) z - z53 ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc ⟫ - z53 {z} ⟪ as , ch-is-sup u u ¬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 (proj1 a) (proj1 b)) - - pcha : pchainpx ⊆ A - pcha (case1 lt) = proj1 lt - pcha (case2 fc) = A∋fc _ f mf (proj1 fc) - - sup1 : MinSUP A pchainpx - sup1 = minsupP pchainpx pcha ptotal - - -- - -- supf0 px o≤ sp1 - -- - - zc41 : MinSUP A pchainpx → ZChain A f mf< ay x - zc41 sup1 = record { supf = supf1 ; asupf = asupf1 ; zo≤sz = zo≤sz ; is-minsup = is-minsup ; cfcs = cfcs ; supf-mono = supf1-mono } where - - sp1 = MinSUP.sup sup1 - - 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.as 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 - - fcup : {u z : Ordinal } → FClosure A f (supf1 u) z → u o≤ px → FClosure A f (supf0 u) z - fcup {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sf1=sf0 u≤px) fc - fcpu : {u z : Ordinal } → FClosure A f (supf0 u) z → u o≤ px → FClosure A f (supf1 u) z - fcpu {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sym (sf1=sf0 u≤px)) fc - - -- this is a kind of maximality, so we cannot prove this without <-monotonicity - -- - cfcs : {a b w : Ordinal } - → a o< b → b o≤ x → supf1 a o< b → FClosure A f (supf1 a) w → odef (UnionCF A f ay supf1 b) w - cfcs {a} {b} {w} a ¬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 ay supf1 b) w - csupf1 with x ¬a ¬b c = ⊥-elim ( ¬p ¬a ¬b c = ⊥-elim ( ¬p ¬a ¬b 0 (λ lt → <-irr (case2 fc11) lt) (λ eq → <-irr (case1 eq) fc11) fc11 where - fc11 : * (& b) < * (& a) - fc11 = subst (λ k → k < * (& a) ) (cong (*) (sym eq1)) lt - ... | case2 lt = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where - fc12 : * (& b) < * (& a) - fc12 = ftrans<-≤ lt (subst (λ k → k ≤ & a) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl )) (s≤fc _ f mf fca ) ) - pcmp (ic-isup i i ¬a ¬b ib (<-irr (case2 lt1)) (λ eq → <-irr (case1 eq) lt1) lt1 where - lt1 : a0 < b0 - lt1 = subst₂ (λ j k → j < k ) *iso *iso lt - ptotalS (case2 a) (case2 b) = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (fcn-cmp spu0 f mf (proj1 a) (proj1 b)) - - S⊆A : pchainS ⊆ A - S⊆A (case1 lt) = proj1 lt - S⊆A (case2 fc) = A∋fc _ f mf (proj1 fc) - - ssup : MinSUP A pchainS - ssup = minsupP pchainS S⊆A ptotalS - - zc400 : MinSUP A pchainU → MinSUP A pchainS → ZChain A f mf< ay x - zc400 usup ssup = record { supf = supf1 ; asupf = asupf ; zo≤sz = zo≤sz ; is-minsup = is-minsup ; cfcs = cfcs ; supf-mono = supf-mono } where - - spu = MinSUP.sup usup - sps = MinSUP.sup ssup - - supf1 : Ordinal → Ordinal - supf1 z with trio< z x - ... | tri< a ¬b ¬c = ZChain.supf (pzc (ob ¬a ¬b c = sps -- sup of spu which o< x - -- if x o< spu, spu is not included in UnionCF x - -- the chain - - pchain : HOD - pchain = UnionCF A f ay supf1 x - - -- pchain ⊆ pchainU ⊆ pchianS - - sf1=sf : {z : Ordinal } → (a : z o< x ) → supf1 z ≡ ZChain.supf (pzc (ob ¬a ¬b c = ⊥-elim (¬a z ¬a ¬b c = ⊥-elim (¬b (sym eq)) - - sf1=sps : {z : Ordinal } → (a : x o< z ) → supf1 z ≡ sps - sf1=sps {z} x x ¬a ¬b c = refl - - asupf : {z : Ordinal } → odef A (supf1 z) - asupf {z} with trio< z x - ... | tri< a ¬b ¬c = ZChain.asupf (pzc (ob ¬a ¬b c = MinSUP.as ssup - - supf-mono : {z y : Ordinal } → z o≤ y → supf1 z o≤ supf1 y - supf-mono {z} {y} z≤y with trio< y x - ... | tri< y ¬a ¬b c = zc01 where -- supf1 z o≤ sps - zc01 : supf1 z o≤ sps - zc01 with trio< z x - ... | tri< z ¬a ¬b c = o≤-refl -- (sf1=sps c) - - is-minsup : {z : Ordinal} → z o≤ x → IsMinSUP A (UnionCF A f ay supf1 z) (supf1 z) - is-minsup {z} z≤x with osuc-≡< z≤x - ... | case1 z=x = record { as = asupf ; x≤sup = zm00 ; minsup = zm01 } where - zm00 : {w : Ordinal } → odef (UnionCF A f ay supf1 z) w → w ≤ supf1 z - zm00 {w} ⟪ az , ch-init fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=spu (sym z=x))) ( MinSUP.x≤sup usup ⟪ az , ic-init fc ⟫ ) - zm00 {w} ⟪ az , ch-is-sup u u ¬a ¬b c | record { eq = eq } = ob ¬a ¬b c = ⊥-elim ( ¬a (ordtrans<-≤ a