# HG changeset patch # User Shinji KONO # Date 1667622102 -32400 # Node ID 39c6801887385104c531239754a390fca17b264a # Parent 1c1c6a6ed4facc52d14619a08c24eb4b897bc115 ... diff -r 1c1c6a6ed4fa -r 39c680188738 src/zorn.agda --- a/src/zorn.agda Sat Nov 05 10:23:57 2022 +0900 +++ b/src/zorn.agda Sat Nov 05 13:21:42 2022 +0900 @@ -1,30 +1,30 @@ {-# OPTIONS --allow-unsolved-metas #-} open import Level hiding ( suc ; zero ) open import Ordinals -open import Relation.Binary +open import Relation.Binary open import Relation.Binary.Core open import Relation.Binary.PropositionalEquality -import OD +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 +-- Zorn-lemma : { A : HOD } +-- → o∅ o< & A -- → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition --- → Maximal A +-- → 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 Relation.Nullary +open import Data.Empty +import BAlgbra open import Data.Nat hiding ( _<_ ; _≤_ ) -open import Data.Nat.Properties -open import nat +open import Data.Nat.Properties +open import nat open inOrdinal O @@ -52,42 +52,42 @@ -- Partial Order on HOD ( possibly limited in A ) -- -_<<_ : (x y : Ordinal ) → Set n +_<<_ : (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 +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 +... | tri> ¬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 @@ -238,7 +238,7 @@ ax : odef A x y : Ordinal ay : odef B y - x=fy : x ≡ f y + x=fy : x ≡ f y record IsSUP (A B : HOD) {x : Ordinal } (xa : odef A x) : Set n where field @@ -247,7 +247,7 @@ 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 + 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 ) @@ -261,18 +261,19 @@ -- 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 ϕ +-- 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 + 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 : * 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 : * 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 @@ -326,13 +358,13 @@ -- Union of supf z which o< x -- -UnionCF : ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) {y : Ordinal } (ay : odef A y ) +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 c = ⊥-elim (<-irr (case2 sx sx uc02 s ¬a ¬b c = ⊥-elim ( o≤> uc03 s ¬a ¬b px ¬a ¬b c = sp1 sf1=sf0 : {z : Ordinal } → z o≤ px → supf1 z ≡ supf0 z @@ -819,12 +891,12 @@ asupf1 : {z : Ordinal } → odef A (supf1 z) asupf1 {z} with trio< z px - ... | tri< a ¬b ¬c = ZChain.asupf zc - ... | tri≈ ¬a b ¬c = ZChain.asupf zc + ... | tri< a ¬b ¬c = ZChain.asupf zc + ... | tri≈ ¬a b ¬c = ZChain.asupf zc ... | tri> ¬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 + 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 @@ -833,25 +905,26 @@ zc21 = ZChain.minsup zc (o<→≤ 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 : 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 : 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 zc11 : {z : Ordinal} → odef (UnionCF A f mf ay supf1 x) z → odef pchainpx z + zc11 {z} ⟪ az , ch-init fc ⟫ = case1 ⟪ az , ch-init fc ⟫ zc11 {z} ⟪ az , ch-is-sup u su ¬a ¬b c | _ = ⊥-elim ( ¬p ¬a ¬b px ¬a ¬b px ¬a ¬b c = supf0 px @@ -1069,7 +1147,7 @@ zc17 = ? -- px o< z, px o< supf0 px supf-mono1 : {z w : Ordinal } → z o≤ w → supf1 z o≤ supf1 w - supf-mono1 {z} {w} z≤w with trio< w px + supf-mono1 {z} {w} z≤w with trio< w px ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (ordtrans≤-< z≤w a))) refl ( supf-mono z≤w ) ... | tri≈ ¬a refl ¬c with trio< z px ... | tri< a ¬b ¬c = zc17 @@ -1084,22 +1162,25 @@ pchain1 = UnionCF A f mf ay supf1 x zc10 : {z : Ordinal} → OD.def (od pchain) z → OD.def (od pchain1) z + zc10 {z} ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫ zc10 {z} ⟪ az , ch-is-sup u1 u1 ¬a ¬b px ¬a ¬b px ¬a ¬b px ¬a ¬b px ¬a ¬b c = ? -- ⊥-elim ( o≤> u ¬a ¬b x ¬a ¬b x ¬a ¬b x ¬a ¬b x ¬a ¬b x ¬a ¬b x ¬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 ) @@ -1367,37 +1454,37 @@ -- 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 )))) + 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 ) zc msp1 ss ( subst (λ k → k o≤ supf mc) (ChainP.supu=u is-sup1) z51) sc