changeset 478:c6346d92f1a1

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sun, 03 Apr 2022 07:59:55 +0900
parents 24b4b854b310
children fea0c2454b85
files src/filter.agda src/zorn.agda
diffstat 2 files changed, 53 insertions(+), 11 deletions(-) [+]
line wrap: on
line diff
--- a/src/filter.agda	Sat Apr 02 08:37:17 2022 +0900
+++ b/src/filter.agda	Sun Apr 03 07:59:55 2022 +0900
@@ -173,3 +173,15 @@
        genf : Filter LP
        generic : (D : Dense LP ) → M ∋ Dense.dense D → ¬ ( (Dense.dense D ∩ Filter.filter genf ) ≡ od∅ )
 
+open import zorn
+
+record MaximumFilter {L P : HOD} (LP : L ⊆ Power P) : Set (suc n) where
+    field
+       mf : Filter LP
+       proper  : ¬ (filter mf ∋ od∅)
+       is-maximum : ( f : Filter LP ) →  filter f  ⊆ filter mf
+
+max→ultra : {L P : HOD} (LP : L ⊆ Power P) → (mx : MaximumFilter LP ) → ultra-filter ( MaximumFilter.mf mx )
+max→ultra {L} {P} LP mx = record { proper = {!!} ; ultra = {!!} }
+
+ 
--- a/src/zorn.agda	Sat Apr 02 08:37:17 2022 +0900
+++ b/src/zorn.agda	Sun Apr 03 07:59:55 2022 +0900
@@ -1,3 +1,4 @@
+{-# OPTIONS --allow-unsolved-metas #-}
 open import Level
 open import Ordinals
 module zorn {n : Level } (O : Ordinals {n})   where
@@ -60,13 +61,13 @@
    field
       sup : HOD
       A∋maximal : A ∋ sup
-      x≤sup : {x : HOD} → B ∋ x → (x ≡ sup ) ∨ (x < sup )
+      x≤sup : {x : HOD} → B ∋ x → (x ≡ sup ) ∨ (x < sup )   -- B is Total
 
 record Maximal ( A : HOD ) (_<_ : (x y : HOD) → Set n ) : Set (suc n) where
    field
       maximal : HOD
       A∋maximal : A ∋ maximal
-      ¬maximal<x : {x : HOD} → A ∋ x  → ¬ maximal < x
+      ¬maximal<x : {x : HOD} → A ∋ x  → ¬ maximal < x       -- A is Partial, use negative
 
 open _==_
 open _⊆_
@@ -82,11 +83,10 @@
 
 Zorn-lemma : { A : HOD } → { _<_ : (x y : HOD) → Set n }
     → o∅ o< & A 
-    → ( {a b c : HOD} → a < b → b < c → a < c )
     → PartialOrderSet A _<_
-    → ( ( B : HOD) → (B⊆A : B ⊆ A) → TotalOrderSet B _<_ → SUP A B  _<_  )
+    → ( ( B : HOD) → (B⊆A : B ⊆ A) → TotalOrderSet B _<_ → SUP A B  _<_  ) -- SUP condition
     → Maximal A _<_ 
-Zorn-lemma {A} {_<_} 0<A TR PO supP = zorn00 where
+Zorn-lemma {A} {_<_} 0<A PO supP = zorn00 where
      someA : HOD
      someA = ODC.minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A )) 
      isSomeA : A ∋ someA
@@ -102,6 +102,7 @@
      z01 : {a b : HOD} → A ∋ a → A ∋ b  → (a ≡ b ) ∨ (a < b ) → b < a → ⊥
      z01 {a} {b} A∋a A∋b (case1 a=b) b<a = proj1 (proj2 (PO (me  A∋b) (me A∋a)) (sym a=b)) b<a
      z01 {a} {b} A∋a A∋b (case2 a<b) b<a = proj1 (PO (me  A∋b) (me A∋a)) b<a ⟪ a<b , (λ b=a → proj1 (proj2 (PO (me  A∋b) (me A∋a)) b=a ) b<a ) ⟫
+     -- ZChain is not compatible with the SUP condition
      ZChain→¬SUP :  (z : ZChain A (& A) _<_ ) →  ¬ (SUP A (ZChain.B z) _<_ )
      ZChain→¬SUP z sp = ⊥-elim (z02 (ZChain.fb z (SUP.A∋maximal sp)) (ZChain.B∋fb z  _ (SUP.A∋maximal sp)) (ZChain.¬x≤sup z _  (SUP.A∋maximal sp) z03 )) where
          z03 : & (SUP.sup sp) o< osuc (& A)
@@ -137,29 +138,57 @@
                  ¬x<m :  ¬ (* x < * m)
                  ¬x<m x<m = ∅< {Gtx ax} {* m} ⟪ subst (λ k → odef A k) (sym &iso) am , subst (λ k → * x < k ) (cong (*) (sym &iso)) x<m ⟫  (≡o∅→=od∅ nogt) 
           ... | no not = record { B = Bx     --  we have larger element, let's create ZChain
-              ; B⊆A = B⊆A ; total = {!!} ; fb = {!!} ; B∋fb = {!!} ; ¬x≤sup = {!!} } where
+              ; B⊆A = B⊆A ; total = total ; fb = fb ; B∋fb = {!!} ; ¬x≤sup = {!!} } where
                  B = ZChain.B zc1 
                  Bx : HOD
-                 Bx = record { od = record { def = λ y → (x ≡ y) ∨ odef B y } ; odmax = & A ; <odmax = {!!}  } 
+                 Bx = record { od = record { def = λ y → (x ≡ y) ∨ odef B y } ; odmax = & A ; <odmax = {!!}  }  -- Union (B , x)
                  B⊆A : Bx ⊆ A
                  B⊆A = record { incl = λ {y} by → z07 y by  } where
                      z07 : (y : HOD) → Bx ∋ y → A ∋ y
                      z07 y (case1 x=y) = subst (λ k → odef A k ) (trans &iso x=y) ax
                      z07 y (case2 by) = incl (ZChain.B⊆A zc1 ) by
                  m = ODC.minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq))
+                 p : odef A (& m) ∧ (* x < (* (& m)))
+                 p = ODC.x∋minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq))
+                 fb :  {y : HOD} → A ∋ y → HOD
+                 fb {y} ay with trio< (& y) x
+                 ... | tri< a ¬b ¬c = ZChain.fb zc1 ay
+                 ... | tri≈ ¬a b ¬c = m
+                 ... | tri> ¬a ¬b c = od∅
+                 total : TotalOrderSet Bx _<_
+                 total ex ey with is-elm ex | is-elm ey
+                 ... | case1 eq | case1 eq1 = tri≈ {!!} {!!} {!!} 
+                 ... | case1 x | case2 x₁ = tri< {!!} {!!} {!!} 
+                 ... | case2 x | case1 x₁ = {!!}
+                 ... | case2 x | case2 x₁ = ZChain.total zc1 (me x) (me x₁)
      ind nomx x prev | no ¬ox with trio< (& A) x   --- limit ordinal case
      ... | tri< a ¬b ¬c = record { B = ZChain.B zc1
               ; B⊆A = ZChain.B⊆A zc1 ; total = ZChain.total zc1 ; fb = ZChain.fb zc1 ; B∋fb = ZChain.B∋fb zc1 ; ¬x≤sup = {!!} } where
           zc1 : ZChain A (& A) _<_
           zc1 = prev (& A) a 
      ... | tri≈ ¬a b ¬c = record { B = B
-              ; B⊆A = {!!} ; total = {!!} ; fb = {!!} ; B∋fb = {!!} ; ¬x≤sup = {!!} } where
-          B : HOD
-          B = record { od = record { def = λ y → (y<x : y o< x ) → odef (ZChain.B (prev y y<x)) y } ; odmax = & A ; <odmax = {!!} } 
-     ... | tri> ¬a ¬b c = {!!}
+              ; B⊆A = B⊆A ; total = {!!} ; fb = {!!} ; B∋fb = {!!} ; ¬x≤sup = {!!} } where
+          B : HOD  -- Union (previous B)
+          B = record { od = record { def = λ y → (y o< x) ∧ ((y<x : y o< x ) → odef (ZChain.B (prev y y<x)) y) } ; odmax = & A ; <odmax = {!!} } 
+          B⊆A : B ⊆ A
+          B⊆A = record { incl = λ {y} bx → incl (ZChain.B⊆A (prev (& y) (proj1 bx))) (proj2 bx (proj1 bx)) } 
+     ... | tri> ¬a ¬b c with ODC.∋-p O A (* x)
+     ... | no ¬Ax = {!!} where
+          B : HOD  -- Union (previous B)
+          B = record { od = record { def = λ y → (y o< x) ∧ ((y<x : y o< x ) → odef (ZChain.B (prev y y<x)) y) } ; odmax = & A ; <odmax = {!!} } 
+     ... | yes ax with is-o∅ (& (Gtx ax))
+     ... | yes nogt = ⊥-elim (no-maximum nomx x x-is-maximal ) where -- no larger element, so it is maximal
+              x-is-maximal :  (m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m))
+              x-is-maximal m am  = ⟪ subst (λ k → odef A k) &iso ax ,  ¬x<m  ⟫ where
+                 ¬x<m :  ¬ (* x < * m)
+                 ¬x<m x<m = ∅< {Gtx ax} {* m} ⟪ subst (λ k → odef A k) (sym &iso) am , subst (λ k → * x < k ) (cong (*) (sym &iso)) x<m ⟫  (≡o∅→=od∅ nogt) 
+     ... | no not = {!!} where
+          B : HOD  -- Union (x , previous B)
+          B = record { od = record { def = λ y → (y o< osuc x) ∧ ((y<x : y o< x ) → odef (ZChain.B (prev y y<x)) y) } ; odmax = & A ; <odmax = {!!} } 
      zorn00 : Maximal A _<_
      zorn00 with is-o∅ ( & HasMaximal )
      ... | no not = record { maximal = ODC.minimal O HasMaximal  (λ eq → not (=od∅→≡o∅ eq)) ; A∋maximal = zorn01 ; ¬maximal<x  = zorn02 } where
+         -- yes we have the maximal
          zorn03 :  odef HasMaximal ( & ( ODC.minimal O HasMaximal  (λ eq → not (=od∅→≡o∅ eq)) ) )
          zorn03 =  ODC.x∋minimal  O HasMaximal  (λ eq → not (=od∅→≡o∅ eq))
          zorn01 :  A ∋ ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq))
@@ -167,6 +196,7 @@
          zorn02 : {x : HOD} → A ∋ x → ¬ (ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) < x)
          zorn02 {x} ax m<x = proj2 (zorn03 (& x) ax) (subst₂ (λ j k → j < k) (sym *iso) (sym *iso) m<x )
      ... | yes ¬Maximal = ⊥-elim ( ZChain→¬SUP (z (& A) (≡o∅→=od∅ ¬Maximal)) ( supP B (ZChain.B⊆A (z (& A) (≡o∅→=od∅ ¬Maximal))) (ZChain.total (z (& A) (≡o∅→=od∅ ¬Maximal))) )) where
+         -- if we have no maximal, make ZChain, which contradict SUP condition
          z : (x : Ordinal) → HasMaximal =h= od∅  → ZChain A x _<_ 
          z x nomx = TransFinite (ind nomx) x
          B = ZChain.B (z (& A) (≡o∅→=od∅ ¬Maximal)  )