Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate src/zorn.agda @ 1270:905311ffe7ec
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sun, 26 Mar 2023 17:18:45 +0900 |
parents | d122d0c1b094 |
children | 45cd80181a29 |
rev | line source |
---|---|
478 | 1 {-# OPTIONS --allow-unsolved-metas #-} |
508 | 2 open import Level hiding ( suc ; zero ) |
431 | 3 open import Ordinals |
966 | 4 open import Relation.Binary |
552 | 5 open import Relation.Binary.Core |
6 open import Relation.Binary.PropositionalEquality | |
1096 | 7 import OD |
552 | 8 module zorn {n : Level } (O : Ordinals {n}) (_<_ : (x y : OD.HOD O ) → Set n ) (PO : IsStrictPartialOrder _≡_ _<_ ) where |
431 | 9 |
560 | 10 -- |
966 | 11 -- Zorn-lemma : { A : HOD } |
12 -- → o∅ o< & A | |
560 | 13 -- → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition |
966 | 14 -- → Maximal A |
560 | 15 -- |
16 | |
1091 | 17 open import zf -- hiding ( _⊆_ ) |
477 | 18 open import logic |
19 -- open import partfunc {n} O | |
20 | |
966 | 21 open import Relation.Nullary |
22 open import Data.Empty | |
1124 | 23 import BAlgebra |
431 | 24 |
555 | 25 open import Data.Nat hiding ( _<_ ; _≤_ ) |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
26 open import Data.Nat.Properties |
966 | 27 open import nat |
555 | 28 |
431 | 29 |
30 open inOrdinal O | |
31 open OD O | |
32 open OD.OD | |
33 open ODAxiom odAxiom | |
477 | 34 import OrdUtil |
35 import ODUtil | |
431 | 36 open Ordinals.Ordinals O |
37 open Ordinals.IsOrdinals isOrdinal | |
38 open Ordinals.IsNext isNext | |
39 open OrdUtil O | |
477 | 40 open ODUtil O |
41 | |
42 | |
43 import ODC | |
44 | |
45 open _∧_ | |
46 open _∨_ | |
47 open Bool | |
431 | 48 |
49 open HOD | |
50 | |
560 | 51 -- |
52 -- Partial Order on HOD ( possibly limited in A ) | |
53 -- | |
54 | |
966 | 55 _<<_ : (x y : Ordinal ) → Set n |
570 | 56 x << y = * x < * y |
57 | |
1031 | 58 _≤_ : (x y : Ordinal ) → Set n -- we can't use * x ≡ * y, it is Set (Level.suc n). Level (suc n) troubles Chain |
59 x ≤ y = (x ≡ y ) ∨ ( * x < * y ) | |
765 | 60 |
966 | 61 POO : IsStrictPartialOrder _≡_ _<<_ |
62 POO = record { isEquivalence = record { refl = refl ; sym = sym ; trans = trans } | |
63 ; trans = IsStrictPartialOrder.trans PO | |
570 | 64 ; irrefl = λ x=y x<y → IsStrictPartialOrder.irrefl PO (cong (*) x=y) x<y |
966 | 65 ; <-resp-≈ = record { fst = λ {x} {y} {y1} y=y1 xy1 → subst (λ k → x << k ) y=y1 xy1 ; snd = λ {x} {x1} {y} x=x1 x1y → subst (λ k → k << x ) x=x1 x1y } } |
66 | |
1031 | 67 ≤-ftrans : {x y z : Ordinal } → x ≤ y → y ≤ z → x ≤ z |
554 | 68 ≤-ftrans {x} {y} {z} (case1 refl ) (case1 refl ) = case1 refl |
69 ≤-ftrans {x} {y} {z} (case1 refl ) (case2 y<z) = case2 y<z | |
70 ≤-ftrans {x} {_} {z} (case2 x<y ) (case1 refl ) = case2 x<y | |
71 ≤-ftrans {x} {y} {z} (case2 x<y) (case2 y<z) = case2 ( IsStrictPartialOrder.trans PO x<y y<z ) | |
72 | |
1031 | 73 ftrans≤-< : {x y z : Ordinal } → x ≤ y → y << z → x << z |
74 ftrans≤-< {x} {y} {z} (case1 eq) y<z = subst (λ k → k < * z) (sym (cong (*) eq)) y<z | |
75 ftrans≤-< {x} {y} {z} (case2 lt) y<z = IsStrictPartialOrder.trans PO lt y<z | |
951 | 76 |
1031 | 77 ftrans<-≤ : {x y z : Ordinal } → x << y → y ≤ z → x << z |
78 ftrans<-≤ {x} {y} {z} x<y (case1 eq) = subst (λ k → * x < k ) ((cong (*) eq)) x<y | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
79 ftrans<-≤ {x} {y} {z} x<y (case2 lt) = IsStrictPartialOrder.trans PO x<y lt |
779 | 80 |
556 | 81 <-irr : {a b : HOD} → (a ≡ b ) ∨ (a < b ) → b < a → ⊥ |
82 <-irr {a} {b} (case1 a=b) b<a = IsStrictPartialOrder.irrefl PO (sym a=b) b<a | |
83 <-irr {a} {b} (case2 a<b) b<a = IsStrictPartialOrder.irrefl PO refl | |
84 (IsStrictPartialOrder.trans PO b<a a<b) | |
490 | 85 |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
86 <<-irr : {a b : Ordinal } → a ≤ b → b << a → ⊥ |
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
87 <<-irr {a} {b} (case1 a=b) b<a = IsStrictPartialOrder.irrefl PO (cong (*) (sym a=b)) b<a |
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
88 <<-irr {a} {b} (case2 a<b) b<a = IsStrictPartialOrder.irrefl PO refl (IsStrictPartialOrder.trans PO b<a a<b) |
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
89 |
561 | 90 ptrans = IsStrictPartialOrder.trans PO |
91 | |
492 | 92 open _==_ |
1091 | 93 -- open _⊆_ -- we use different definition |
492 | 94 |
1091 | 95 -- We cannot prove this without Well foundedness of A |
96 -- | |
966 | 97 -- <-TransFinite : {A x : HOD} → {P : HOD → Set n} → x ∈ A |
1091 | 98 -- → ({y : HOD} → A ∋ y → y < x → P y ) → P x |
879 | 99 -- <-TransFinite = ? |
100 | |
530 | 101 -- |
560 | 102 -- Closure of ≤-monotonic function f has total order |
530 | 103 -- |
104 | |
1031 | 105 ≤-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set n |
106 ≤-monotonic-f A f = (x : Ordinal ) → odef A x → ( x ≤ (f x) ) ∧ odef A (f x ) | |
530 | 107 |
992 | 108 <-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set n |
109 <-monotonic-f A f = (x : Ordinal ) → odef A x → ( * x < * (f x) ) ∧ odef A (f x ) | |
110 | |
551 | 111 data FClosure (A : HOD) (f : Ordinal → Ordinal ) (s : Ordinal) : Ordinal → Set n where |
783 | 112 init : {s1 : Ordinal } → odef A s → s ≡ s1 → FClosure A f s s1 |
555 | 113 fsuc : (x : Ordinal) ( p : FClosure A f s x ) → FClosure A f s (f x) |
554 | 114 |
556 | 115 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 |
783 | 116 A∋fc {A} s f mf (init as refl ) = as |
556 | 117 A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {A} s f mf fcy ) ) |
555 | 118 |
714 | 119 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 |
783 | 120 A∋fcs {A} s f mf (init as refl) = as |
966 | 121 A∋fcs {A} s f mf (fsuc y fcy) = A∋fcs {A} s f mf fcy |
714 | 122 |
1031 | 123 s≤fc : {A : HOD} (s : Ordinal ) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → s ≤ y |
783 | 124 s≤fc {A} s {.s} f mf (init x refl ) = case1 refl |
556 | 125 s≤fc {A} s {.(f x)} f mf (fsuc x fcy) with proj1 (mf x (A∋fc s f mf fcy ) ) |
1031 | 126 ... | case1 x=fx = subst₂ (λ j k → j ≤ k ) refl x=fx (s≤fc s f mf fcy) |
966 | 127 ... | case2 x<fx with s≤fc {A} s f mf fcy |
1031 | 128 ... | case1 s≡x = case2 ( subst₂ (λ j k → j < k ) (sym (cong (*) s≡x )) refl x<fx ) |
556 | 129 ... | case2 s<x = case2 ( IsStrictPartialOrder.trans PO s<x x<fx ) |
555 | 130 |
800 | 131 fcn : {A : HOD} (s : Ordinal) { x : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) → FClosure A f s x → ℕ |
132 fcn s mf (init as refl) = zero | |
133 fcn {A} s {x} {f} mf (fsuc y p) with proj1 (mf y (A∋fc s f mf p)) | |
134 ... | case1 eq = fcn s mf p | |
135 ... | case2 y<fy = suc (fcn s mf p ) | |
136 | |
966 | 137 fcn-inject : {A : HOD} (s : Ordinal) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) |
800 | 138 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx ≡ fcn s mf cy → * x ≡ * y |
139 fcn-inject {A} s {x} {y} {f} mf cx cy eq = fc00 (fcn s mf cx) (fcn s mf cy) eq cx cy refl refl where | |
140 fc06 : {y : Ordinal } { as : odef A s } (eq : s ≡ y ) { j : ℕ } → ¬ suc j ≡ fcn {A} s {y} {f} mf (init as eq ) | |
141 fc06 {x} {y} refl {j} not = fc08 not where | |
966 | 142 fc08 : {j : ℕ} → ¬ suc j ≡ 0 |
800 | 143 fc08 () |
144 fc07 : {x : Ordinal } (cx : FClosure A f s x ) → 0 ≡ fcn s mf cx → * s ≡ * x | |
145 fc07 {x} (init as refl) eq = refl | |
146 fc07 {.(f x)} (fsuc x cx) eq with proj1 (mf x (A∋fc s f mf cx ) ) | |
1031 | 147 ... | case1 x=fx = subst (λ k → * s ≡ k ) (cong (*) x=fx) ( fc07 cx eq ) |
800 | 148 -- ... | case2 x<fx = ? |
149 fc00 : (i j : ℕ ) → i ≡ j → {x y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → i ≡ fcn s mf cx → j ≡ fcn s mf cy → * x ≡ * y | |
150 fc00 (suc i) (suc j) x cx (init x₃ x₄) x₁ x₂ = ⊥-elim ( fc06 x₄ x₂ ) | |
151 fc00 (suc i) (suc j) x (init x₃ x₄) (fsuc x₅ cy) x₁ x₂ = ⊥-elim ( fc06 x₄ x₁ ) | |
152 fc00 zero zero refl (init _ refl) (init x₁ refl) i=x i=y = refl | |
153 fc00 zero zero refl (init as refl) (fsuc y cy) i=x i=y with proj1 (mf y (A∋fc s f mf cy ) ) | |
1031 | 154 ... | case1 y=fy = subst (λ k → * s ≡ * k ) y=fy (fc07 cy i=y) -- ( fc00 zero zero refl (init as refl) cy i=x i=y ) |
800 | 155 fc00 zero zero refl (fsuc x cx) (init as refl) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) ) |
1031 | 156 ... | case1 x=fx = subst (λ k → * k ≡ * s ) x=fx (sym (fc07 cx i=x)) -- ( fc00 zero zero refl cx (init as refl) i=x i=y ) |
800 | 157 fc00 zero zero refl (fsuc x cx) (fsuc y cy) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) ) | proj1 (mf y (A∋fc s f mf cy ) ) |
1031 | 158 ... | case1 x=fx | case1 y=fy = subst₂ (λ j k → * j ≡ * k ) x=fx y=fy ( fc00 zero zero refl cx cy i=x i=y ) |
800 | 159 fc00 (suc i) (suc j) i=j {.(f x)} {.(f y)} (fsuc x cx) (fsuc y cy) i=x j=y with proj1 (mf x (A∋fc s f mf cx ) ) | proj1 (mf y (A∋fc s f mf cy ) ) |
1031 | 160 ... | case1 x=fx | case1 y=fy = subst₂ (λ j k → * j ≡ * k ) x=fx y=fy ( fc00 (suc i) (suc j) i=j cx cy i=x j=y ) |
161 ... | case1 x=fx | case2 y<fy = subst (λ k → * k ≡ * (f y)) x=fx (fc02 x cx i=x) where | |
800 | 162 fc02 : (x1 : Ordinal) → (cx1 : FClosure A f s x1 ) → suc i ≡ fcn s mf cx1 → * x1 ≡ * (f y) |
163 fc02 x1 (init x₁ x₂) x = ⊥-elim (fc06 x₂ x) | |
164 fc02 .(f x1) (fsuc x1 cx1) i=x1 with proj1 (mf x1 (A∋fc s f mf cx1 ) ) | |
1031 | 165 ... | case1 eq = trans (sym (cong (*) eq )) ( fc02 x1 cx1 i=x1 ) -- derefence while f x ≡ x |
800 | 166 ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc04) where |
167 fc04 : * x1 ≡ * y | |
168 fc04 = fc00 i j (cong pred i=j) cx1 cy (cong pred i=x1) (cong pred j=y) | |
1031 | 169 ... | case2 x<fx | case1 y=fy = subst (λ k → * (f x) ≡ * k ) y=fy (fc03 y cy j=y) where |
800 | 170 fc03 : (y1 : Ordinal) → (cy1 : FClosure A f s y1 ) → suc j ≡ fcn s mf cy1 → * (f x) ≡ * y1 |
171 fc03 y1 (init x₁ x₂) x = ⊥-elim (fc06 x₂ x) | |
172 fc03 .(f y1) (fsuc y1 cy1) j=y1 with proj1 (mf y1 (A∋fc s f mf cy1 ) ) | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
173 ... | case1 eq = trans ( fc03 y1 cy1 j=y1 ) (cong (*) eq) |
800 | 174 ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc05) where |
175 fc05 : * x ≡ * y1 | |
176 fc05 = fc00 i j (cong pred i=j) cx cy1 (cong pred i=x) (cong pred j=y1) | |
177 ... | case2 x₁ | case2 x₂ = subst₂ (λ j k → * (f j) ≡ * (f k) ) &iso &iso (cong (λ k → * (f (& k))) (fc00 i j (cong pred i=j) cx cy (cong pred i=x) (cong pred j=y))) | |
178 | |
179 | |
180 fcn-< : {A : HOD} (s : Ordinal ) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) | |
181 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx Data.Nat.< fcn s mf cy → * x < * y | |
182 fcn-< {A} s {x} {y} {f} mf cx cy x<y = fc01 (fcn s mf cy) cx cy refl x<y where | |
183 fc06 : {y : Ordinal } { as : odef A s } (eq : s ≡ y ) { j : ℕ } → ¬ suc j ≡ fcn {A} s {y} {f} mf (init as eq ) | |
184 fc06 {x} {y} refl {j} not = fc08 not where | |
966 | 185 fc08 : {j : ℕ} → ¬ suc j ≡ 0 |
800 | 186 fc08 () |
187 fc01 : (i : ℕ ) → {y : Ordinal } → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → (i ≡ fcn s mf cy ) → fcn s mf cx Data.Nat.< i → * x < * y | |
188 fc01 (suc i) cx (init x₁ x₂) x (s≤s x₃) = ⊥-elim (fc06 x₂ x) | |
189 fc01 (suc i) {y} cx (fsuc y1 cy) i=y (s≤s x<i) with proj1 (mf y1 (A∋fc s f mf cy ) ) | |
1031 | 190 ... | case1 y=fy = subst (λ k → * x < k ) (cong (*) y=fy) ( fc01 (suc i) {y1} cx cy i=y (s≤s x<i) ) |
800 | 191 ... | case2 y<fy with <-cmp (fcn s mf cx ) i |
192 ... | tri> ¬a ¬b c = ⊥-elim ( nat-≤> x<i c ) | |
966 | 193 ... | tri≈ ¬a b ¬c = subst (λ k → k < * (f y1) ) (fcn-inject s mf cy cx (sym (trans b (cong pred i=y) ))) y<fy |
800 | 194 ... | tri< a ¬b ¬c = IsStrictPartialOrder.trans PO fc02 y<fy where |
195 fc03 : suc i ≡ suc (fcn s mf cy) → i ≡ fcn s mf cy | |
966 | 196 fc03 eq = cong pred eq |
197 fc02 : * x < * y1 | |
800 | 198 fc02 = fc01 i cx cy (fc03 i=y ) a |
199 | |
557 | 200 |
966 | 201 fcn-cmp : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) |
554 | 202 → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → Tri (* x < * y) (* x ≡ * y) (* y < * x ) |
800 | 203 fcn-cmp {A} s {x} {y} f mf cx cy with <-cmp ( fcn s mf cx ) (fcn s mf cy ) |
204 ... | tri< a ¬b ¬c = tri< fc11 (λ eq → <-irr (case1 (sym eq)) fc11) (λ lt → <-irr (case2 fc11) lt) where | |
205 fc11 : * x < * y | |
206 fc11 = fcn-< {A} s {x} {y} {f} mf cx cy a | |
207 ... | tri≈ ¬a b ¬c = tri≈ (λ lt → <-irr (case1 (sym fc10)) lt) fc10 (λ lt → <-irr (case1 fc10) lt) where | |
208 fc10 : * x ≡ * y | |
209 fc10 = fcn-inject {A} s {x} {y} {f} mf cx cy b | |
966 | 210 ... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where |
800 | 211 fc12 : * y < * x |
212 fc12 = fcn-< {A} s {y} {x} {f} mf cy cx c | |
600 | 213 |
563 | 214 |
729 | 215 |
560 | 216 -- open import Relation.Binary.Properties.Poset as Poset |
217 | |
218 IsTotalOrderSet : ( A : HOD ) → Set (Level.suc n) | |
219 IsTotalOrderSet A = {a b : HOD} → odef A (& a) → odef A (& b) → Tri (a < b) (a ≡ b) (b < a ) | |
220 | |
1091 | 221 ⊆-IsTotalOrderSet : { A B : HOD } → B ⊆ A → IsTotalOrderSet A → IsTotalOrderSet B |
222 ⊆-IsTotalOrderSet {A} {B} B⊆A T ax ay = T (B⊆A ax) (B⊆A ay) | |
554 | 223 |
958
33891adf80ea
IsMinSup contains not HasPrev
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
957
diff
changeset
|
224 record HasPrev (A B : HOD) ( f : Ordinal → Ordinal ) (x : Ordinal ) : Set n where |
533 | 225 field |
836 | 226 ax : odef A x |
534 | 227 y : Ordinal |
541 | 228 ay : odef B y |
966 | 229 x=fy : x ≡ f y |
529 | 230 |
1031 | 231 record IsSUP (A B : HOD) (x : Ordinal ) : Set n where |
962 | 232 field |
1031 | 233 ax : odef A x |
234 x≤sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x ) -- B is Total, use positive | |
568 | 235 |
656 | 236 record SUP ( A B : HOD ) : Set (Level.suc n) where |
237 field | |
238 sup : HOD | |
1031 | 239 isSUP : IsSUP A B (& sup) |
240 ax = IsSUP.ax isSUP | |
241 x≤sup = IsSUP.x≤sup isSUP | |
656 | 242 |
690 | 243 -- |
1091 | 244 -- Our Proof strategy of the Zorn Lemma |
878 | 245 -- |
1097 | 246 -- As ZChain.cfcs closure of supf z is smaller than next supf z1, and supf z o< supfz1, because of mf< |
247 -- if have to be stopped since we have upper bound & A, so there is a Maximul element. | |
248 -- | |
990 | 249 -- f (f ( ... (supf y))) f (f ( ... (supf z1))) |
878 | 250 -- / | / | |
251 -- / | / | | |
990 | 252 -- supf y < supf z1 < supf z2 |
878 | 253 -- o< o< |
990 | 254 -- |
255 -- if sup z1 ≡ sup z2, the chain is stopped at sup z1, then f (sup z1) ≡ sup z1 | |
1039 | 256 -- |
990 | 257 |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
258 fc-stop : ( A : HOD ) ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) { a b : Ordinal } |
990 | 259 → (aa : odef A a ) →( {y : Ordinal} → FClosure A f a y → (y ≡ b ) ∨ (y << b )) → a ≡ b → f a ≡ a |
260 fc-stop A f mf {a} {b} aa x≤sup a=b with x≤sup (fsuc a (init aa refl )) | |
261 ... | case1 eq = trans eq (sym a=b) | |
1031 | 262 ... | case2 lt = ⊥-elim (<-irr (case1 (cong (λ k → * (f k) ) (sym a=b))) (ftrans<-≤ lt fc00 ) ) where |
263 fc00 : b ≤ (f b) | |
990 | 264 fc00 = proj1 (mf _ (subst (λ k → odef A k) a=b aa )) |
265 | |
694 | 266 ∈∧P→o< : {A : HOD } {y : Ordinal} → {P : Set n} → odef A y ∧ P → y o< & A |
267 ∈∧P→o< {A } {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p ))) | |
268 | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
269 -- Union of supf z and FClosure A f y |
662 | 270 |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
271 data UChain { A : HOD } { f : Ordinal → Ordinal } {supf : Ordinal → Ordinal} {y : Ordinal } (ay : odef A y ) |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
272 (x : Ordinal) : (z : Ordinal) → Set n where |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
273 ch-init : {z : Ordinal } (fc : FClosure A f y z) → UChain ay x z |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
274 ch-is-sup : (u : Ordinal) {z : Ordinal } (u<x : u o< x) (supu=u : supf u ≡ u) ( fc : FClosure A f (supf u) z ) → UChain ay x z |
1030 | 275 |
1060 | 276 UnionCF : ( A : HOD ) ( f : Ordinal → Ordinal ) {y : Ordinal } (ay : odef A y ) ( supf : Ordinal → Ordinal ) ( x : Ordinal ) → HOD |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
277 UnionCF A f ay supf x |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
278 = record { od = record { def = λ z → odef A z ∧ UChain {A} {f} {supf} ay x z } ; |
1028 | 279 odmax = & A ; <odmax = λ {y} sy → ∈∧P→o< sy } |
280 | |
1061 | 281 -- Union of chain lower than x |
1060 | 282 |
1078 | 283 data IChain {A : HOD} { f : Ordinal → Ordinal } {y : Ordinal } (ay : odef A y ) |
1063 | 284 {x : Ordinal } (supfz : {z : Ordinal } → z o< x → Ordinal) : (z : Ordinal ) → Set n where |
285 ic-init : {z : Ordinal } (fc : FClosure A f y z) → IChain ay supfz z | |
286 ic-isup : {z : Ordinal} (i : Ordinal) (i<x : i o< x) (s<x : supfz i<x o≤ i ) (fc : FClosure A f (supfz i<x) z) → IChain ay supfz z | |
1060 | 287 |
1063 | 288 UnionIC : ( A : HOD ) ( f : Ordinal → Ordinal ) { x : Ordinal } {y : Ordinal } (ay : odef A y ) (supfz : {z : Ordinal } → z o< x → Ordinal) → HOD |
1078 | 289 UnionIC A f ay supfz |
1063 | 290 = record { od = record { def = λ z → odef A z ∧ IChain {A} {f} ay supfz z } ; |
1060 | 291 odmax = & A ; <odmax = λ {y} sy → ∈∧P→o< sy } |
1028 | 292 |
966 | 293 supf-inject0 : {x y : Ordinal } {supf : Ordinal → Ordinal } → (supf-mono : {x y : Ordinal } → x o≤ y → supf x o≤ supf y ) |
294 → supf x o< supf y → x o< y | |
842 | 295 supf-inject0 {x} {y} {supf} supf-mono sx<sy with trio< x y |
296 ... | tri< a ¬b ¬c = a | |
297 ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy ) | |
298 ... | tri> ¬a ¬b y<x with osuc-≡< (supf-mono (o<→≤ y<x) ) | |
299 ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy ) | |
300 ... | case2 lt = ⊥-elim ( o<> sx<sy lt ) | |
301 | |
1031 | 302 record IsMinSUP ( A B : HOD ) (sup : Ordinal) : Set n where |
303 field | |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
304 as : odef A sup |
1031 | 305 x≤sup : {x : Ordinal } → odef B x → (x ≡ sup ) ∨ (x << sup ) |
306 minsup : { sup1 : Ordinal } → odef A sup1 | |
307 → ( {x : Ordinal } → odef B x → (x ≡ sup1 ) ∨ (x << sup1 )) → sup o≤ sup1 | |
308 | |
879 | 309 record MinSUP ( A B : HOD ) : Set n where |
310 field | |
311 sup : Ordinal | |
1031 | 312 isMinSUP : IsMinSUP A B sup |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
313 as = IsMinSUP.as isMinSUP |
1031 | 314 x≤sup = IsMinSUP.x≤sup isMinSUP |
315 minsup = IsMinSUP.minsup isMinSUP | |
879 | 316 |
966 | 317 chain-mono : {A : HOD} ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) (supf : Ordinal → Ordinal ) |
919 | 318 (supf-mono : {x y : Ordinal } → x o≤ y → supf x o≤ supf y ) {a b c : Ordinal} → a o≤ b |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
319 → odef (UnionCF A f ay supf a) c → odef (UnionCF A f ay supf b) c |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
320 chain-mono f mf ay supf supf-mono {a} {b} {c} a≤b ⟪ ua , ch-init fc ⟫ = ⟪ ua , ch-init fc ⟫ |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
321 chain-mono f mf ay supf supf-mono {a} {b} {c} a≤b ⟪ ua , ch-is-sup u u<x supu=u fc ⟫ = ⟪ ua , ch-is-sup u (ordtrans<-≤ u<x a≤b) supu=u fc ⟫ |
908
d917831fb607
supf (supf x) ≡ supf x is bad
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
907
diff
changeset
|
322 |
1038 | 323 record ZChain ( A : HOD ) ( f : Ordinal → Ordinal ) (mf< : <-monotonic-f A f) |
783 | 324 {y : Ordinal} (ay : odef A y) ( z : Ordinal ) : Set (Level.suc n) where |
655 | 325 field |
966 | 326 supf : Ordinal → Ordinal |
1028 | 327 |
1097 | 328 asupf : {x : Ordinal } → odef A (supf x) |
329 is-minsup : {x : Ordinal } → (x≤z : x o≤ z) → IsMinSUP A (UnionCF A f ay supf x) (supf x) | |
1090 | 330 supf-mono : {a b : Ordinal } → a o≤ b → supf a o≤ supf b |
1046 | 331 cfcs : {a b w : Ordinal } → a o< b → b o≤ z → supf a o< b → FClosure A f (supf a) w → odef (UnionCF A f ay supf b) w |
1097 | 332 zo≤sz : {x : Ordinal } → x o≤ z → x o≤ supf x -- because of mf< |
994 | 333 |
608
6655f03984f9
mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
607
diff
changeset
|
334 chain : HOD |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
335 chain = UnionCF A f ay supf z |
1091 | 336 chain⊆A : chain ⊆ A |
861 | 337 chain⊆A = λ lt → proj1 lt |
934 | 338 |
1058 | 339 chain∋init : {x : Ordinal } → odef (UnionCF A f ay supf x) y |
340 chain∋init {x} = ⟪ ay , ch-init (init ay refl) ⟫ | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
341 |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
342 mf : ≤-monotonic-f A f |
1040 | 343 mf x ax = ⟪ case2 mf00 , proj2 (mf< x ax ) ⟫ where |
344 mf00 : * x < * (f x) | |
345 mf00 = proj1 ( mf< x ax ) | |
1038 | 346 |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
347 f-next : {a z : Ordinal} → odef (UnionCF A f ay supf z) a → odef (UnionCF A f ay supf z) (f a) |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
348 f-next {a} ⟪ ua , ch-init fc ⟫ = ⟪ proj2 ( mf _ ua) , ch-init (fsuc _ fc) ⟫ |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
349 f-next {a} ⟪ ua , ch-is-sup u su<x su=u fc ⟫ = ⟪ proj2 ( mf _ ua) , ch-is-sup u su<x su=u (fsuc _ fc) ⟫ |
861 | 350 |
966 | 351 supf-inject : {x y : Ordinal } → supf x o< supf y → x o< y |
825 | 352 supf-inject {x} {y} sx<sy with trio< x y |
353 ... | tri< a ¬b ¬c = a | |
354 ... | tri≈ ¬a refl ¬c = ⊥-elim ( o<¬≡ (cong supf refl) sx<sy ) | |
355 ... | tri> ¬a ¬b y<x with osuc-≡< (supf-mono (o<→≤ y<x) ) | |
356 ... | case1 eq = ⊥-elim ( o<¬≡ (sym eq) sx<sy ) | |
357 ... | case2 lt = ⊥-elim ( o<> sx<sy lt ) | |
798 | 358 |
1097 | 359 -- another kind of maximality of the chain |
360 -- note that supf z is not an element of this chain | |
361 -- | |
362 csupf : {b : Ordinal } → supf b o< supf z → supf b o< z → odef (UnionCF A f ay supf z) (supf b) | |
1038 | 363 csupf {b} sb<sz sb<z = cfcs (supf-inject sb<sz) o≤-refl sb<z (init asupf refl) |
994 | 364 |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
365 minsup : {x : Ordinal } → x o≤ z → MinSUP A (UnionCF A f ay supf x) |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
366 minsup {x} x≤z = record { sup = supf x ; isMinSUP = is-minsup x≤z } |
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
367 |
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
368 supf-is-minsup : {x : Ordinal } → (x≤z : x o≤ z) → supf x ≡ MinSUP.sup (minsup x≤z) |
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
369 supf-is-minsup _ = refl |
965
1c1c6a6ed4fa
removing ch-init is no good because of initialization
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
964
diff
changeset
|
370 |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
371 -- different from order because y o< supf |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
372 fcy<sup : {u w : Ordinal } → u o≤ z → FClosure A f y w → (w ≡ supf u ) ∨ ( w << supf u ) |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
373 fcy<sup {u} {w} u≤z fc with MinSUP.x≤sup (minsup u≤z) ⟪ subst (λ k → odef A k ) (sym &iso) (A∋fc {A} y f mf fc) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
374 , ch-init (subst (λ k → FClosure A f y k) (sym &iso) fc ) ⟫ |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
375 ... | case1 eq = case1 (subst (λ k → k ≡ supf u ) &iso (trans eq (sym (supf-is-minsup u≤z ) ) )) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
376 ... | case2 lt = case2 (subst₂ (λ j k → j << k ) &iso (sym (supf-is-minsup u≤z )) lt ) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
377 |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
378 initial : {x : Ordinal } → x o≤ z → odef (UnionCF A f ay supf x) x → y ≤ x |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
379 initial {x} x≤z ⟪ aa , ch-init fc ⟫ = s≤fc y f mf fc |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
380 initial {x} x≤z ⟪ aa , ch-is-sup u u<x is-sup fc ⟫ = ≤-ftrans (fcy<sup (ordtrans u<x x≤z) (init ay refl)) (s≤fc _ f mf fc) |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
381 |
1090 | 382 sup=u : {b : Ordinal} → (ab : odef A b) → b o≤ z |
383 → IsSUP A (UnionCF A f ay supf b) b → supf b ≡ b | |
384 sup=u {b} ab b≤z is-sup = z50 where | |
385 z48 : supf b o≤ b | |
386 z48 = IsMinSUP.minsup (is-minsup b≤z) ab (λ ux → IsSUP.x≤sup is-sup ux ) | |
387 z50 : supf b ≡ b | |
388 z50 with trio< (supf b) b | |
389 ... | tri< sb<b ¬b ¬c = ⊥-elim ( o≤> z47 sb<b ) where | |
390 z47 : b o≤ supf b | |
391 z47 = zo≤sz b≤z | |
392 ... | tri≈ ¬a b ¬c = b | |
393 ... | tri> ¬a ¬b b<sb = ⊥-elim ( o≤> z48 b<sb ) | |
394 | |
1097 | 395 -- |
396 -- supf is minsup, so its UniofCF are equal, these are equal | |
397 -- | |
1091 | 398 supfeq : {a b : Ordinal } → a o≤ z → b o≤ z → UnionCF A f ay supf a ≡ UnionCF A f ay supf b → supf a ≡ supf b |
399 supfeq {a} {b} a≤z b≤z ua=ub with trio< (supf a) (supf b) | |
400 ... | tri< sa<sb ¬b ¬c = ⊥-elim ( o≤> ( | |
401 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<sb ) | |
402 ... | tri≈ ¬a b ¬c = b | |
403 ... | tri> ¬a ¬b sb<sa = ⊥-elim ( o≤> ( | |
404 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<sa ) | |
405 | |
1097 | 406 -- |
407 -- supf a over b and supf a is not included in UnionCF a nor UnionCF b, so UnionCF b is equal to the UnionCF a | |
408 -- | |
1090 | 409 union-max : {a b : Ordinal } → b o≤ supf a → b o≤ z → supf a o< supf b → UnionCF A f ay supf a ≡ UnionCF A f ay supf b |
410 union-max {a} {b} b≤sa b≤z sa<sb = ==→o≡ record { eq→ = z47 ; eq← = z48 } where | |
411 z47 : {w : Ordinal } → odef (UnionCF A f ay supf a ) w → odef ( UnionCF A f ay supf b ) w | |
412 z47 {w} ⟪ aw , ch-init fc ⟫ = ⟪ aw , ch-init fc ⟫ | |
413 z47 {w} ⟪ aw , ch-is-sup u u<a su=u fc ⟫ = ⟪ aw , ch-is-sup u u<b su=u fc ⟫ where | |
414 u<b : u o< b | |
415 u<b = ordtrans u<a (supf-inject sa<sb ) | |
416 z48 : {w : Ordinal } → odef (UnionCF A f ay supf b ) w → odef ( UnionCF A f ay supf a ) w | |
417 z48 {w} ⟪ aw , ch-init fc ⟫ = ⟪ aw , ch-init fc ⟫ | |
418 z48 {w} ⟪ aw , ch-is-sup u u<b su=u fc ⟫ = ⟪ aw , ch-is-sup u u<a su=u fc ⟫ where | |
419 u<a : u o< a | |
420 u<a = supf-inject ( osucprev (begin | |
421 osuc (supf u) ≡⟨ cong osuc su=u ⟩ | |
422 osuc u ≤⟨ osucc u<b ⟩ | |
423 b ≤⟨ b≤sa ⟩ | |
424 supf a ∎ )) where open o≤-Reasoning O | |
425 | |
426 x≤supfx→¬sa<sa : {a b : Ordinal } → b o≤ z → b o≤ supf a → ¬ (supf a o< supf b ) | |
427 x≤supfx→¬sa<sa {a} {b} b≤z b≤sa sa<sb = ⊥-elim ( o<¬≡ z27 sa<sb ) where -- x o≤ supf a ∧ supf a o< supf b → ⊥, because it defines the same UnionCF | |
428 z27 : supf a ≡ supf b | |
429 z27 = supfeq (ordtrans (supf-inject sa<sb) b≤z) b≤z ( union-max b≤sa b≤z sa<sb) | |
430 | |
431 order : {a b w : Ordinal } → b o≤ z → supf a o< supf b → FClosure A f (supf a) w → w ≤ supf b | |
432 order {a} {b} {w} b≤z sa<sb fc = IsMinSUP.x≤sup (is-minsup b≤z) (cfcs (supf-inject sa<sb) b≤z sa<b fc) where | |
433 sa<b : supf a o< b | |
434 sa<b with x<y∨y≤x (supf a) b | |
435 ... | case1 lt = lt | |
436 ... | case2 b≤sa = ⊥-elim (x≤supfx→¬sa<sa b≤z b≤sa sa<sb) | |
437 | |
438 supf-idem : {b : Ordinal } → b o≤ z → supf b o≤ z → supf (supf b) ≡ supf b | |
439 supf-idem {b} b≤z sfb≤x = z52 where | |
440 z54 : {w : Ordinal} → odef (UnionCF A f ay supf (supf b)) w → (w ≡ supf b) ∨ (w << supf b) | |
441 z54 {w} ⟪ aw , ch-init fc ⟫ = fcy<sup b≤z fc | |
442 z54 {w} ⟪ aw , ch-is-sup u u<x su=u fc ⟫ = order b≤z (subst (λ k → k o< supf b) (sym su=u) u<x) fc where | |
443 u<b : u o< b | |
444 u<b = supf-inject (subst (λ k → k o< supf b ) (sym (su=u)) u<x ) | |
445 z52 : supf (supf b) ≡ supf b | |
446 z52 = sup=u asupf sfb≤x record { ax = asupf ; x≤sup = z54 } | |
447 | |
448 supf-mono< : {a b : Ordinal } → b o≤ z → supf a o< supf b → supf a << supf b | |
449 supf-mono< {a} {b} b≤z sa<sb with order {a} {b} b≤z sa<sb (init asupf refl) | |
450 ... | case2 lt = lt | |
451 ... | case1 eq = ⊥-elim ( o<¬≡ eq sa<sb ) | |
452 | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
453 f-total : IsTotalOrderSet chain |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
454 f-total {a} {b} ⟪ uaa , ch-is-sup ua sua<x sua=ua fca ⟫ ⟪ uab , ch-is-sup ub sub<x sub=ub fcb ⟫ = |
1035 | 455 subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso fc-total where |
456 fc-total : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) ) | |
457 fc-total with trio< ua ub | |
1046 | 458 ... | tri< a₁ ¬b ¬c with ≤-ftrans (order (o<→≤ sub<x) (subst₂ (λ j k → j o< k) (sym sua=ua) (sym sub=ub) a₁) fca ) (s≤fc (supf ub) f mf fcb ) |
1035 | 459 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where |
460 ct00 : * (& a) ≡ * (& b) | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
461 ct00 = cong (*) eq1 |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
462 ... | case2 a<b = tri< a<b (λ eq → <-irr (case1 (sym eq)) a<b ) (λ lt → <-irr (case2 a<b ) lt) |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
463 fc-total | tri≈ _ refl _ = fcn-cmp _ f mf fca fcb |
1046 | 464 fc-total | tri> ¬a ¬b c with ≤-ftrans (order (o<→≤ sua<x) (subst₂ (λ j k → j o< k) (sym sub=ub) (sym sua=ua) c) fcb ) (s≤fc (supf ua) f mf fca ) |
1035 | 465 ... | case1 eq1 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where |
466 ct00 : * (& a) ≡ * (& b) | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
467 ct00 = cong (*) (sym eq1) |
1035 | 468 ... | case2 b<a = tri> (λ lt → <-irr (case2 b<a ) lt) (λ eq → <-irr (case1 eq) b<a ) b<a |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
469 f-total {a} {b} ⟪ uaa , ch-init fca ⟫ ⟪ uab , ch-is-sup ub sub<x sub=ub fcb ⟫ = ft00 where |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
470 ft01 : (& a) ≤ (& b) → Tri ( a < b) ( a ≡ b) ( b < a ) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
471 ft01 (case1 eq) = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym a=b)) lt)) a=b (λ lt → ⊥-elim (<-irr (case1 a=b) lt)) where |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
472 a=b : a ≡ b |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
473 a=b = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
474 ft01 (case2 lt) = tri< a<b (λ eq → <-irr (case1 (sym eq)) a<b ) (λ lt → <-irr (case2 a<b ) lt) where |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
475 a<b : a < b |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
476 a<b = subst₂ (λ j k → j < k ) *iso *iso lt |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
477 ft00 : Tri ( a < b) ( a ≡ b) ( b < a ) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
478 ft00 = ft01 (≤-ftrans (fcy<sup (o<→≤ sub<x) fca) (s≤fc {A} _ f mf fcb)) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
479 f-total {a} {b} ⟪ uaa , ch-is-sup ua sua<x sua=ua fca ⟫ ⟪ uab , ch-init fcb ⟫ = ft00 where |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
480 ft01 : (& b) ≤ (& a) → Tri ( a < b) ( a ≡ b) ( b < a ) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
481 ft01 (case1 eq) = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym a=b)) lt)) a=b (λ lt → ⊥-elim (<-irr (case1 a=b) lt)) where |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
482 a=b : a ≡ b |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
483 a=b = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) (sym eq)) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
484 ft01 (case2 lt) = tri> (λ lt → <-irr (case2 b<a ) lt) (λ eq → <-irr (case1 eq) b<a ) b<a where |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
485 b<a : b < a |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
486 b<a = subst₂ (λ j k → j < k ) *iso *iso lt |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
487 ft00 : Tri ( a < b) ( a ≡ b) ( b < a ) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
488 ft00 = ft01 (≤-ftrans (fcy<sup (o<→≤ sua<x) fcb) (s≤fc {A} _ f mf fca)) |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
489 f-total {a} {b} ⟪ uaa , ch-init fca ⟫ ⟪ uab , ch-init fcb ⟫ = |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
490 subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso (fcn-cmp y f mf fca fcb ) |
825 | 491 |
1038 | 492 record ZChain1 ( A : HOD ) ( f : Ordinal → Ordinal ) (mf< : <-monotonic-f A f) |
493 {y : Ordinal} (ay : odef A y) (zc : ZChain A f mf< ay (& A)) ( z : Ordinal ) : Set (Level.suc n) where | |
1027 | 494 supf = ZChain.supf zc |
495 field | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
496 is-max : {a b : Ordinal } → (ca : odef (UnionCF A f ay supf z) a ) → b o< z → (ab : odef A b) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
497 → HasPrev A (UnionCF A f ay supf z) f b ∨ IsSUP A (UnionCF A f ay supf b) b |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
498 → * a < * b → odef ((UnionCF A f ay supf z)) b |
1027 | 499 |
500 record Maximal ( A : HOD ) : Set (Level.suc n) where | |
501 field | |
502 maximal : HOD | |
503 as : A ∋ maximal | |
504 ¬maximal<x : {x : HOD} → A ∋ x → ¬ maximal < x -- A is Partial, use negative | |
505 | |
506 -- | |
507 -- supf in TransFinite indution may differ each other, but it is the same because of the minimul sup | |
508 -- | |
1038 | 509 supf-unique : ( A : HOD ) ( f : Ordinal → Ordinal ) (mf< : <-monotonic-f A f) |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
510 {y xa xb : Ordinal} → (ay : odef A y) → (xa o≤ xb ) → (za : ZChain A f mf< ay xa ) (zb : ZChain A f mf< ay xb ) |
1007 | 511 → {z : Ordinal } → z o≤ xa → ZChain.supf za z ≡ ZChain.supf zb z |
1038 | 512 supf-unique A f mf< {y} {xa} {xb} ay xa≤xb za zb {z} z≤xa = TransFinite0 {λ z → z o≤ xa → ZChain.supf za z ≡ ZChain.supf zb z } ind z z≤xa where |
1007 | 513 supfa = ZChain.supf za |
514 supfb = ZChain.supf zb | |
515 ind : (x : Ordinal) → ((w : Ordinal) → w o< x → w o≤ xa → supfa w ≡ supfb w) → x o≤ xa → supfa x ≡ supfb x | |
1008 | 516 ind x prev x≤xa = sxa=sxb where |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
517 ma = ZChain.minsup za x≤xa |
1008 | 518 mb = ZChain.minsup zb (OrdTrans x≤xa xa≤xb ) |
519 spa = MinSUP.sup ma | |
520 spb = MinSUP.sup mb | |
521 sax=spa : supfa x ≡ spa | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
522 sax=spa = ZChain.supf-is-minsup za x≤xa |
1008 | 523 sbx=spb : supfb x ≡ spb |
524 sbx=spb = ZChain.supf-is-minsup zb (OrdTrans x≤xa xa≤xb ) | |
1007 | 525 sxa=sxb : supfa x ≡ supfb x |
526 sxa=sxb with trio< (supfa x) (supfb x) | |
527 ... | tri≈ ¬a b ¬c = b | |
528 ... | tri< a ¬b ¬c = ⊥-elim ( o≤> ( | |
529 begin | |
1008 | 530 supfb x ≡⟨ sbx=spb ⟩ |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
531 spb ≤⟨ MinSUP.minsup mb (MinSUP.as ma) (λ {z} uzb → MinSUP.x≤sup ma (z53 uzb)) ⟩ |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
532 spa ≡⟨ sym sax=spa ⟩ |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
533 supfa x ∎ ) a ) where |
1008 | 534 open o≤-Reasoning O |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
535 z53 : {z : Ordinal } → odef (UnionCF A f ay (ZChain.supf zb) x) z → odef (UnionCF A f ay (ZChain.supf za) x) z |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
536 z53 ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc ⟫ |
1034 | 537 z53 {z} ⟪ as , ch-is-sup u u<x su=u fc ⟫ = ⟪ as , ch-is-sup u u<x (trans ua=ub su=u) z55 ⟫ where |
538 ua=ub : supfa u ≡ supfb u | |
539 ua=ub = prev u u<x (ordtrans u<x x≤xa ) | |
540 z55 : FClosure A f (ZChain.supf za u) z | |
541 z55 = subst (λ k → FClosure A f k z ) (sym ua=ub) fc | |
1007 | 542 ... | tri> ¬a ¬b c = ⊥-elim ( o≤> ( |
543 begin | |
1008 | 544 supfa x ≡⟨ sax=spa ⟩ |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
545 spa ≤⟨ MinSUP.minsup ma (MinSUP.as mb) (λ uza → MinSUP.x≤sup mb (z53 uza)) ⟩ |
1008 | 546 spb ≡⟨ sym sbx=spb ⟩ |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
547 supfb x ∎ ) c ) where |
1009 | 548 open o≤-Reasoning O |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
549 z53 : {z : Ordinal } → odef (UnionCF A f ay (ZChain.supf za) x) z → odef (UnionCF A f ay (ZChain.supf zb) x) z |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
550 z53 ⟪ as , ch-init fc ⟫ = ⟪ as , ch-init fc ⟫ |
1034 | 551 z53 {z} ⟪ as , ch-is-sup u u<x su=u fc ⟫ = ⟪ as , ch-is-sup u u<x (trans ub=ua su=u) z55 ⟫ where |
552 ub=ua : supfb u ≡ supfa u | |
553 ub=ua = sym ( prev u u<x (ordtrans u<x x≤xa )) | |
554 z55 : FClosure A f (ZChain.supf zb u) z | |
555 z55 = subst (λ k → FClosure A f k z ) (sym ub=ua) fc | |
1028 | 556 |
966 | 557 Zorn-lemma : { A : HOD } |
558 → o∅ o< & A | |
1091 | 559 → ( ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → SUP A B ) -- SUP condition |
966 | 560 → Maximal A |
552 | 561 Zorn-lemma {A} 0<A supP = zorn00 where |
571 | 562 <-irr0 : {a b : HOD} → A ∋ a → A ∋ b → (a ≡ b ) ∨ (a < b ) → b < a → ⊥ |
563 <-irr0 {a} {b} A∋a A∋b = <-irr | |
788 | 564 z07 : {y : Ordinal} {A : HOD } → {P : Set n} → odef A y ∧ P → y o< & A |
565 z07 {y} {A} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p ))) | |
530 | 566 s : HOD |
966 | 567 s = ODC.minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A )) |
568 | 568 as : A ∋ * ( & s ) |
569 as = subst (λ k → odef A (& k) ) (sym *iso) ( ODC.x∋minimal O A (λ eq → ¬x<0 ( subst (λ k → o∅ o< k ) (=od∅→≡o∅ eq) 0<A )) ) | |
608
6655f03984f9
mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
607
diff
changeset
|
570 as0 : odef A (& s ) |
6655f03984f9
mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
607
diff
changeset
|
571 as0 = subst (λ k → odef A k ) &iso as |
547 | 572 s<A : & s o< & A |
568 | 573 s<A = c<→o< (subst (λ k → odef A (& k) ) *iso as ) |
530 | 574 HasMaximal : HOD |
966 | 575 HasMaximal = record { od = record { def = λ x → odef A x ∧ ( (m : Ordinal) → odef A m → ¬ (* x < * m)) } ; odmax = & A ; <odmax = z07 } |
537 | 576 no-maximum : HasMaximal =h= od∅ → (x : Ordinal) → odef A x ∧ ((m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m) )) → ⊥ |
966 | 577 no-maximum nomx x P = ¬x<0 (eq→ nomx {x} ⟪ proj1 P , (λ m ma p → proj2 ( proj2 P m ma ) p ) ⟫ ) |
532 | 578 Gtx : { x : HOD} → A ∋ x → HOD |
966 | 579 Gtx {x} ax = record { od = record { def = λ y → odef A y ∧ (x < (* y)) } ; odmax = & A ; <odmax = z07 } |
537 | 580 z08 : ¬ Maximal A → HasMaximal =h= od∅ |
804 | 581 z08 nmx = record { eq→ = λ {x} lt → ⊥-elim ( nmx record {maximal = * x ; as = subst (λ k → odef A k) (sym &iso) (proj1 lt) |
537 | 582 ; ¬maximal<x = λ {y} ay → subst (λ k → ¬ (* x < k)) *iso (proj2 lt (& y) ay) } ) ; eq← = λ {y} lt → ⊥-elim ( ¬x<0 lt )} |
583 x-is-maximal : ¬ Maximal A → {x : Ordinal} → (ax : odef A x) → & (Gtx (subst (λ k → odef A k ) (sym &iso) ax)) ≡ o∅ → (m : Ordinal) → odef A m → odef A x ∧ (¬ (* x < * m)) | |
584 x-is-maximal nmx {x} ax nogt m am = ⟪ subst (λ k → odef A k) &iso (subst (λ k → odef A k ) (sym &iso) ax) , ¬x<m ⟫ where | |
585 ¬x<m : ¬ (* x < * m) | |
966 | 586 ¬x<m x<m = ∅< {Gtx (subst (λ k → odef A k ) (sym &iso) ax)} {* m} ⟪ subst (λ k → odef A k) (sym &iso) am , subst (λ k → * x < k ) (cong (*) (sym &iso)) x<m ⟫ (≡o∅→=od∅ nogt) |
543 | 587 |
1027 | 588 -- |
589 -- we have minsup using LEM, this is similar to the proof of the axiom of choice | |
590 -- | |
1091 | 591 minsupP : ( B : HOD) → (B⊆A : B ⊆ A) → IsTotalOrderSet B → MinSUP A B |
879 | 592 minsupP B B⊆A total = m02 where |
593 xsup : (sup : Ordinal ) → Set n | |
594 xsup sup = {w : Ordinal } → odef B w → (w ≡ sup ) ∨ (w << sup ) | |
595 ∀-imply-or : {A : Ordinal → Set n } {B : Set n } | |
596 → ((x : Ordinal ) → A x ∨ B) → ((x : Ordinal ) → A x) ∨ B | |
597 ∀-imply-or {A} {B} ∀AB with ODC.p∨¬p O ((x : Ordinal ) → A x) -- LEM | |
598 ∀-imply-or {A} {B} ∀AB | case1 t = case1 t | |
599 ∀-imply-or {A} {B} ∀AB | case2 x = case2 (lemma (λ not → x not )) where | |
600 lemma : ¬ ((x : Ordinal ) → A x) → B | |
601 lemma not with ODC.p∨¬p O B | |
602 lemma not | case1 b = b | |
603 lemma not | case2 ¬b = ⊥-elim (not (λ x → dont-orb (∀AB x) ¬b )) | |
604 m00 : (x : Ordinal ) → ( ( z : Ordinal) → z o< x → ¬ (odef A z ∧ xsup z) ) ∨ MinSUP A B | |
605 m00 x = TransFinite0 ind x where | |
606 ind : (x : Ordinal) → ((z : Ordinal) → z o< x → ( ( w : Ordinal) → w o< z → ¬ (odef A w ∧ xsup w )) ∨ MinSUP A B) | |
607 → ( ( w : Ordinal) → w o< x → ¬ (odef A w ∧ xsup w) ) ∨ MinSUP A B | |
608 ind x prev = ∀-imply-or m01 where | |
609 m01 : (z : Ordinal) → (z o< x → ¬ (odef A z ∧ xsup z)) ∨ MinSUP A B | |
610 m01 z with trio< z x | |
611 ... | tri≈ ¬a b ¬c = case1 ( λ lt → ⊥-elim ( ¬a lt ) ) | |
612 ... | tri> ¬a ¬b c = case1 ( λ lt → ⊥-elim ( ¬a lt ) ) | |
613 ... | tri< a ¬b ¬c with prev z a | |
614 ... | case2 mins = case2 mins | |
615 ... | case1 not with ODC.p∨¬p O (odef A z ∧ xsup z) | |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
616 ... | case1 mins = case2 record { sup = z ; isMinSUP = record { as = proj1 mins ; x≤sup = proj2 mins ; minsup = m04 } } where |
879 | 617 m04 : {sup1 : Ordinal} → odef A sup1 → ({w : Ordinal} → odef B w → (w ≡ sup1) ∨ (w << sup1)) → z o≤ sup1 |
618 m04 {s} as lt with trio< z s | |
619 ... | tri< a ¬b ¬c = o<→≤ a | |
620 ... | tri≈ ¬a b ¬c = o≤-refl0 b | |
621 ... | tri> ¬a ¬b s<z = ⊥-elim ( not s s<z ⟪ as , lt ⟫ ) | |
622 ... | case2 notz = case1 (λ _ → notz ) | |
623 m03 : ¬ ((z : Ordinal) → z o< & A → ¬ odef A z ∧ xsup z) | |
1097 | 624 m03 not = ⊥-elim ( not s1 (odef< (SUP.ax S)) ⟪ SUP.ax S , m05 ⟫ ) where |
879 | 625 S : SUP A B |
626 S = supP B B⊆A total | |
627 s1 = & (SUP.sup S) | |
628 m05 : {w : Ordinal } → odef B w → (w ≡ s1 ) ∨ (w << s1 ) | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
629 m05 {w} bw with SUP.x≤sup S bw |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
630 ... | case1 eq = case1 ( subst₂ (λ j k → j ≡ k ) &iso refl (trans &iso eq)) |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
631 ... | case2 lt = case2 lt |
966 | 632 m02 : MinSUP A B |
879 | 633 m02 = dont-or (m00 (& A)) m03 |
634 | |
560 | 635 -- Uncountable ascending chain by axiom of choice |
530 | 636 cf : ¬ Maximal A → Ordinal → Ordinal |
532 | 637 cf nmx x with ODC.∋-p O A (* x) |
638 ... | no _ = o∅ | |
639 ... | yes ax with is-o∅ (& ( Gtx ax )) | |
538 | 640 ... | yes nogt = -- no larger element, so it is maximal |
641 ⊥-elim (no-maximum (z08 nmx) x ⟪ subst (λ k → odef A k) &iso ax , x-is-maximal nmx (subst (λ k → odef A k ) &iso ax) nogt ⟫ ) | |
532 | 642 ... | no not = & (ODC.minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq))) |
537 | 643 is-cf : (nmx : ¬ Maximal A ) → {x : Ordinal} → odef A x → odef A (cf nmx x) ∧ ( * x < * (cf nmx x) ) |
644 is-cf nmx {x} ax with ODC.∋-p O A (* x) | |
645 ... | no not = ⊥-elim ( not (subst (λ k → odef A k ) (sym &iso) ax )) | |
646 ... | yes ax with is-o∅ (& ( Gtx ax )) | |
647 ... | yes nogt = ⊥-elim (no-maximum (z08 nmx) x ⟪ subst (λ k → odef A k) &iso ax , x-is-maximal nmx (subst (λ k → odef A k ) &iso ax) nogt ⟫ ) | |
648 ... | no not = ODC.x∋minimal O (Gtx ax) (λ eq → not (=od∅→≡o∅ eq)) | |
606 | 649 |
650 --- | |
651 --- infintie ascention sequence of f | |
652 --- | |
530 | 653 cf-is-<-monotonic : (nmx : ¬ Maximal A ) → (x : Ordinal) → odef A x → ( * x < * (cf nmx x) ) ∧ odef A (cf nmx x ) |
537 | 654 cf-is-<-monotonic nmx x ax = ⟪ proj2 (is-cf nmx ax ) , proj1 (is-cf nmx ax ) ⟫ |
530 | 655 cf-is-≤-monotonic : (nmx : ¬ Maximal A ) → ≤-monotonic-f A ( cf nmx ) |
532 | 656 cf-is-≤-monotonic nmx x ax = ⟪ case2 (proj1 ( cf-is-<-monotonic nmx x ax )) , proj2 ( cf-is-<-monotonic nmx x ax ) ⟫ |
543 | 657 |
803 | 658 -- |
953 | 659 -- maximality of chain |
660 -- | |
661 -- supf is fixed for z ≡ & A , we can prove order and is-max | |
1016 | 662 -- we have supf-unique now, it is provable in the first Tranfinte induction |
803 | 663 |
992 | 664 SZ1 : ( f : Ordinal → Ordinal ) (mf : ≤-monotonic-f A f) (mf< : <-monotonic-f A f) |
1038 | 665 {init : Ordinal} (ay : odef A init) (zc : ZChain A f mf< ay (& A)) (x : Ordinal) → x o≤ & A → ZChain1 A f mf< ay zc x |
993 | 666 SZ1 f mf mf< {y} ay zc x x≤A = zc1 x x≤A where |
900 | 667 chain-mono1 : {a b c : Ordinal} → a o≤ b |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
668 → odef (UnionCF A f ay (ZChain.supf zc) a) c → odef (UnionCF A f ay (ZChain.supf zc) b) c |
919 | 669 chain-mono1 {a} {b} {c} a≤b = chain-mono f mf ay (ZChain.supf zc) (ZChain.supf-mono zc) a≤b |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
670 is-max-hp : (x : Ordinal) {a : Ordinal} {b : Ordinal} → odef (UnionCF A f ay (ZChain.supf zc) x) a → (ab : odef A b) |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
671 → HasPrev A (UnionCF A f ay (ZChain.supf zc) x) f b |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
672 → * a < * b → odef (UnionCF A f ay (ZChain.supf zc) x) b |
920 | 673 is-max-hp x {a} {b} ua ab has-prev a<b with HasPrev.ay has-prev |
1034 | 674 ... | ⟪ ab0 , ch-init fc ⟫ = ⟪ ab , ch-init ( subst (λ k → FClosure A f y k) (sym (HasPrev.x=fy has-prev)) (fsuc _ fc )) ⟫ |
675 ... | ⟪ ab0 , ch-is-sup u u<x su=u fc ⟫ = ⟪ ab , subst (λ k → UChain ay x k ) | |
676 (sym (HasPrev.x=fy has-prev)) ( ch-is-sup u u<x su=u (fsuc _ fc)) ⟫ | |
677 | |
869 | 678 supf = ZChain.supf zc |
679 | |
1038 | 680 zc1 : (x : Ordinal ) → x o≤ & A → ZChain1 A f mf< ay zc x |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
681 zc1 x x≤A with Oprev-p x |
1024
ab72526316bd
supf-< and ZChain1.order is removed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1023
diff
changeset
|
682 ... | yes op = record { is-max = is-max } where |
988
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
683 px = Oprev.oprev op |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
684 is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f ay supf x) a → |
988
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
685 b o< x → (ab : odef A b) → |
1034 | 686 HasPrev A (UnionCF A f ay supf x) f b ∨ IsSUP A (UnionCF A f ay supf b) b → |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
687 * a < * b → odef (UnionCF A f ay supf x) b |
988
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
688 is-max {a} {b} ua b<x ab P a<b with ODC.or-exclude O P |
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
689 is-max {a} {b} ua b<x ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b |
989 | 690 is-max {a} {b} ua b<x ab P a<b | case2 is-sup with osuc-≡< (ZChain.supf-mono zc (o<→≤ b<x)) |
1024
ab72526316bd
supf-< and ZChain1.order is removed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1023
diff
changeset
|
691 ... | case2 sb<sx = m10 where |
988
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
692 b<A : b o< & A |
1097 | 693 b<A = odef< ab |
988
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
694 m05 : ZChain.supf zc b ≡ b |
1097 | 695 m05 = ZChain.sup=u zc ab (o<→≤ (odef< ab) ) record { ax = ab ; x≤sup = λ {z} uz → IsSUP.x≤sup (proj2 is-sup) uz } |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
696 m10 : odef (UnionCF A f ay supf x) b |
1039 | 697 m10 = ZChain.cfcs zc b<x x≤A (subst (λ k → k o< x) (sym m05) b<x) (init (ZChain.asupf zc) m05) |
992 | 698 ... | case1 sb=sx = ⊥-elim (<-irr (case1 (cong (*) m10)) (proj1 (mf< (supf b) (ZChain.asupf zc)))) where |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
699 m17 : MinSUP A (UnionCF A f ay supf x) -- supf z o< supf ( supf x ) |
992 | 700 m17 = ZChain.minsup zc x≤A |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
701 m18 : supf x ≡ MinSUP.sup m17 |
992 | 702 m18 = ZChain.supf-is-minsup zc x≤A |
990 | 703 m10 : f (supf b) ≡ supf b |
704 m10 = fc-stop A f mf (ZChain.asupf zc) m11 sb=sx where | |
705 m11 : {z : Ordinal} → FClosure A f (supf b) z → (z ≡ ZChain.supf zc x) ∨ (z << ZChain.supf zc x) | |
706 m11 {z} fc = subst (λ k → (z ≡ k) ∨ (z << k)) (sym m18) ( MinSUP.x≤sup m17 m13 ) where | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
707 m04 : ¬ HasPrev A (UnionCF A f ay supf b) f b |
1019 | 708 m04 nhp = proj1 is-sup ( record { ax = HasPrev.ax nhp ; y = HasPrev.y nhp ; ay = |
709 chain-mono1 (o<→≤ b<x) (HasPrev.ay nhp) ; x=fy = HasPrev.x=fy nhp } ) | |
710 m05 : ZChain.supf zc b ≡ b | |
1097 | 711 m05 = ZChain.sup=u zc ab (o<→≤ (odef< ab) ) record { ax = ab ; x≤sup = λ {z} uz → IsSUP.x≤sup (proj2 is-sup) uz } |
1019 | 712 m14 : ZChain.supf zc b o< x |
713 m14 = subst (λ k → k o< x ) (sym m05) b<x | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
714 m13 : odef (UnionCF A f ay supf x) z |
1039 | 715 m13 = ZChain.cfcs zc b<x x≤A m14 fc |
989 | 716 |
1024
ab72526316bd
supf-< and ZChain1.order is removed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1023
diff
changeset
|
717 ... | no lim = record { is-max = is-max } where |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
718 is-max : {a : Ordinal} {b : Ordinal} → odef (UnionCF A f ay supf x) a → |
988
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
719 b o< x → (ab : odef A b) → |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
720 HasPrev A (UnionCF A f ay supf x) f b ∨ IsSUP A (UnionCF A f ay supf b) b → |
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
721 * a < * b → odef (UnionCF A f ay supf x) b |
988
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
722 is-max {a} {b} ua b<x ab P a<b with ODC.or-exclude O P |
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
723 is-max {a} {b} ua b<x ab P a<b | case1 has-prev = is-max-hp x {a} {b} ua ab has-prev a<b |
1058 | 724 is-max {a} {b} ua b<x ab P a<b | case2 is-sup with IsSUP.x≤sup (proj2 is-sup) (ZChain.chain∋init zc ) |
1034 | 725 ... | case1 b=y = ⟪ subst (λ k → odef A k ) b=y ay , ch-init (subst (λ k → FClosure A f y k ) b=y (init ay refl )) ⟫ |
990 | 726 ... | case2 y<b with osuc-≡< (ZChain.supf-mono zc (o<→≤ b<x)) |
1024
ab72526316bd
supf-< and ZChain1.order is removed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1023
diff
changeset
|
727 ... | case2 sb<sx = m10 where |
988
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
728 m09 : b o< & A |
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
729 m09 = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab)) |
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
730 m05 : ZChain.supf zc b ≡ b |
1078 | 731 m05 = ZChain.sup=u zc ab (o<→≤ m09) record { ax = ab ; x≤sup = λ lt → IsSUP.x≤sup (proj2 is-sup) lt } |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
732 m10 : odef (UnionCF A f ay supf x) b |
1039 | 733 m10 = ZChain.cfcs zc b<x x≤A (subst (λ k → k o< x) (sym m05) b<x) (init (ZChain.asupf zc) m05) |
992 | 734 ... | case1 sb=sx = ⊥-elim (<-irr (case1 (cong (*) m10)) (proj1 (mf< (supf b) (ZChain.asupf zc)))) where |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
735 m17 : MinSUP A (UnionCF A f ay supf x) -- supf z o< supf ( supf x ) |
992 | 736 m17 = ZChain.minsup zc x≤A |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
737 m18 : supf x ≡ MinSUP.sup m17 |
992 | 738 m18 = ZChain.supf-is-minsup zc x≤A |
990 | 739 m10 : f (supf b) ≡ supf b |
740 m10 = fc-stop A f mf (ZChain.asupf zc) m11 sb=sx where | |
741 m11 : {z : Ordinal} → FClosure A f (supf b) z → (z ≡ ZChain.supf zc x) ∨ (z << ZChain.supf zc x) | |
742 m11 {z} fc = subst (λ k → (z ≡ k) ∨ (z << k)) (sym m18) ( MinSUP.x≤sup m17 m13 ) where | |
1097 | 743 m05 = ZChain.sup=u zc ab (o<→≤ (odef< ab) ) record { ax = ab ; x≤sup = λ {z} uz → IsSUP.x≤sup (proj2 is-sup) uz } |
1019 | 744 m14 : ZChain.supf zc b o< x |
745 m14 = subst (λ k → k o< x ) (sym m05) b<x | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
746 m13 : odef (UnionCF A f ay supf x) z |
1039 | 747 m13 = ZChain.cfcs zc b<x x≤A m14 fc |
727 | 748 |
757 | 749 uchain : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → HOD |
966 | 750 uchain f mf {y} ay = record { od = record { def = λ x → FClosure A f y x } ; odmax = & A ; <odmax = |
757 | 751 λ {z} cz → subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (A∋fc y f mf cz ))) } |
752 | |
966 | 753 utotal : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) |
757 | 754 → IsTotalOrderSet (uchain f mf ay) |
966 | 755 utotal f mf {y} ay {a} {b} ca cb = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso uz01 where |
757 | 756 uz01 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) ) |
757 uz01 = fcn-cmp y f mf ca cb | |
758 | |
966 | 759 ysup : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) {y : Ordinal} (ay : odef A y) |
928 | 760 → MinSUP A (uchain f mf ay) |
966 | 761 ysup f mf {y} ay = minsupP (uchain f mf ay) (λ lt → A∋fc y f mf lt) (utotal f mf ay) |
757 | 762 |
560 | 763 -- |
547 | 764 -- create all ZChains under o< x |
560 | 765 -- |
608
6655f03984f9
mutual tranfinite in zorn
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
607
diff
changeset
|
766 |
1038 | 767 ind : ( f : Ordinal → Ordinal ) → (mf< : <-monotonic-f A f ) {y : Ordinal} (ay : odef A y) → (x : Ordinal) |
768 → ((z : Ordinal) → z o< x → ZChain A f mf< ay z) → ZChain A f mf< ay x | |
769 ind f mf< {y} ay x prev with Oprev-p x | |
1089
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
770 ... | yes op = zc41 sup1 where |
682 | 771 -- |
772 -- we have previous ordinal to use induction | |
773 -- | |
774 px = Oprev.oprev op | |
1038 | 775 zc : ZChain A f mf< ay (Oprev.oprev op) |
966 | 776 zc = prev px (subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc ) |
682 | 777 px<x : px o< x |
778 px<x = subst (λ k → px o< k) (Oprev.oprev=x op) <-osuc | |
918
4c33f8383d7d
supf px o< px is in csupf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
911
diff
changeset
|
779 opx=x : osuc px ≡ x |
4c33f8383d7d
supf px o< px is in csupf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
911
diff
changeset
|
780 opx=x = Oprev.oprev=x op |
4c33f8383d7d
supf px o< px is in csupf
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
911
diff
changeset
|
781 |
709 | 782 zc-b<x : (b : Ordinal ) → b o< x → b o< osuc px |
966 | 783 zc-b<x b lt = subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) lt |
697 | 784 |
754 | 785 supf0 = ZChain.supf zc |
869 | 786 pchain : HOD |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
787 pchain = UnionCF A f ay supf0 px |
835 | 788 |
857 | 789 supf-mono = ZChain.supf-mono zc |
844 | 790 |
861 | 791 zc04 : {b : Ordinal} → b o≤ x → (b o≤ px ) ∨ (b ≡ x ) |
966 | 792 zc04 {b} b≤x with trio< b px |
861 | 793 ... | tri< a ¬b ¬c = case1 (o<→≤ a) |
794 ... | tri≈ ¬a b ¬c = case1 (o≤-refl0 b) | |
795 ... | tri> ¬a ¬b px<b with osuc-≡< b≤x | |
796 ... | case1 eq = case2 eq | |
966 | 797 ... | case2 b<x = ⊥-elim ( ¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) b<x ⟫ ) |
840 | 798 |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
799 mf : ≤-monotonic-f A f |
1039 | 800 mf x ax = ⟪ case2 mf00 , proj2 (mf< x ax ) ⟫ where |
801 mf00 : * x < * (f x) | |
802 mf00 = proj1 ( mf< x ax ) | |
1038 | 803 |
954 | 804 -- |
805 -- find the next value of supf | |
806 -- | |
807 | |
808 pchainpx : HOD | |
1034 | 809 pchainpx = record { od = record { def = λ z → (odef A z ∧ UChain ay px z ) |
1043 | 810 ∨ (FClosure A f (supf0 px) z ∧ (supf0 px o< x)) } ; odmax = & A ; <odmax = zc00 } where |
811 zc00 : {z : Ordinal } → (odef A z ∧ UChain ay px z ) ∨ (FClosure A f (supf0 px) z ∧ (supf0 px o< x) )→ z o< & A | |
966 | 812 zc00 {z} (case1 lt) = z07 lt |
1097 | 813 zc00 {z} (case2 fc) = odef< ( A∋fc (supf0 px) f mf (proj1 fc) ) |
954 | 814 |
1043 | 815 zc02 : { a b : Ordinal } → odef A a ∧ UChain ay px a → FClosure A f (supf0 px) b ∧ ( supf0 px o< x) → a ≤ b |
816 zc02 {a} {b} ca fb = zc05 (proj1 fb) where | |
1031 | 817 zc05 : {b : Ordinal } → FClosure A f (supf0 px) b → a ≤ b |
954 | 818 zc05 (fsuc b1 fb ) with proj1 ( mf b1 (A∋fc (supf0 px) f mf fb )) |
1034 | 819 ... | case1 eq = subst (λ k → a ≤ k ) eq (zc05 fb) |
1031 | 820 ... | case2 lt = ≤-ftrans (zc05 fb) (case2 lt) |
1034 | 821 zc05 (init b1 refl) = MinSUP.x≤sup (ZChain.minsup zc o≤-refl) ca |
966 | 822 |
954 | 823 ptotal : IsTotalOrderSet pchainpx |
1034 | 824 ptotal (case1 a) (case1 b) = ZChain.f-total zc a b |
954 | 825 ptotal {a0} {b0} (case1 a) (case2 b) with zc02 a b |
826 ... | case1 eq = tri≈ (<-irr (case1 (sym eq1))) eq1 (<-irr (case1 eq1)) where | |
827 eq1 : a0 ≡ b0 | |
828 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq ) | |
829 ... | case2 lt = tri< lt1 (λ eq → <-irr (case1 (sym eq)) lt1) (<-irr (case2 lt1)) where | |
830 lt1 : a0 < b0 | |
831 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt | |
832 ptotal {b0} {a0} (case2 b) (case1 a) with zc02 a b | |
833 ... | case1 eq = tri≈ (<-irr (case1 eq1)) (sym eq1) (<-irr (case1 (sym eq1))) where | |
834 eq1 : a0 ≡ b0 | |
835 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq ) | |
836 ... | case2 lt = tri> (<-irr (case2 lt1)) (λ eq → <-irr (case1 eq) lt1) lt1 where | |
837 lt1 : a0 < b0 | |
838 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt | |
1043 | 839 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)) |
966 | 840 |
1091 | 841 pcha : pchainpx ⊆ A |
954 | 842 pcha (case1 lt) = proj1 lt |
1043 | 843 pcha (case2 fc) = A∋fc _ f mf (proj1 fc) |
966 | 844 |
845 sup1 : MinSUP A pchainpx | |
954 | 846 sup1 = minsupP pchainpx pcha ptotal |
847 | |
848 -- | |
849 -- supf0 px o≤ sp1 | |
966 | 850 -- |
851 | |
1089
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
852 zc41 : MinSUP A pchainpx → ZChain A f mf< ay x |
1090 | 853 zc41 sup1 = record { supf = supf1 ; asupf = asupf1 ; zo≤sz = zo≤sz ; is-minsup = is-minsup ; cfcs = cfcs ; supf-mono = supf1-mono } where |
1034 | 854 |
1089
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
855 sp1 = MinSUP.sup sup1 |
883 | 856 |
871 | 857 supf1 : Ordinal → Ordinal |
966 | 858 supf1 z with trio< z px |
871 | 859 ... | tri< a ¬b ¬c = supf0 z |
966 | 860 ... | tri≈ ¬a b ¬c = supf0 z |
901 | 861 ... | tri> ¬a ¬b c = sp1 |
871 | 862 |
886 | 863 sf1=sf0 : {z : Ordinal } → z o≤ px → supf1 z ≡ supf0 z |
901 | 864 sf1=sf0 {z} z≤px with trio< z px |
874 | 865 ... | tri< a ¬b ¬c = refl |
901 | 866 ... | tri≈ ¬a b ¬c = refl |
867 ... | tri> ¬a ¬b c = ⊥-elim ( o≤> z≤px c ) | |
883 | 868 |
901 | 869 sf1=sp1 : {z : Ordinal } → px o< z → supf1 z ≡ sp1 |
870 sf1=sp1 {z} px<z with trio< z px | |
871 ... | tri< a ¬b ¬c = ⊥-elim ( o<> px<z a ) | |
872 ... | tri≈ ¬a b ¬c = ⊥-elim ( o<¬≡ (sym b) px<z ) | |
873 ... | tri> ¬a ¬b c = refl | |
873 | 874 |
968 | 875 sf=eq : { z : Ordinal } → z o< x → supf0 z ≡ supf1 z |
876 sf=eq {z} z<x = sym (sf1=sf0 (subst (λ k → z o< k ) (sym (Oprev.oprev=x op)) z<x )) | |
877 | |
903 | 878 asupf1 : {z : Ordinal } → odef A (supf1 z) |
879 asupf1 {z} with trio< z px | |
966 | 880 ... | tri< a ¬b ¬c = ZChain.asupf zc |
881 ... | tri≈ ¬a b ¬c = ZChain.asupf zc | |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
882 ... | tri> ¬a ¬b c = MinSUP.as sup1 |
903 | 883 |
966 | 884 supf1-mono : {a b : Ordinal } → a o≤ b → supf1 a o≤ supf1 b |
885 supf1-mono {a} {b} a≤b with trio< b px | |
901 | 886 ... | tri< a ¬b ¬c = subst₂ (λ j k → j o≤ k ) (sym (sf1=sf0 (o<→≤ (ordtrans≤-< a≤b a)))) refl ( supf-mono a≤b ) |
887 ... | 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 ) | |
888 supf1-mono {a} {b} a≤b | tri> ¬a ¬b c with trio< a px | |
889 ... | tri< a<px ¬b ¬c = zc19 where | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
890 zc21 : MinSUP A (UnionCF A f ay supf0 a) |
901 | 891 zc21 = ZChain.minsup zc (o<→≤ a<px) |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
892 zc24 : {x₁ : Ordinal} → odef (UnionCF A f ay supf0 a) x₁ → (x₁ ≡ sp1) ∨ (x₁ << sp1) |
950 | 893 zc24 {x₁} ux = MinSUP.x≤sup sup1 (case1 (chain-mono f mf ay supf0 (ZChain.supf-mono zc) (o<→≤ a<px) ux ) ) |
966 | 894 zc19 : supf0 a o≤ sp1 |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
895 zc19 = subst (λ k → k o≤ sp1) (sym (ZChain.supf-is-minsup zc (o<→≤ a<px))) ( MinSUP.minsup zc21 (MinSUP.as sup1) zc24 ) |
901 | 896 ... | tri≈ ¬a b ¬c = zc18 where |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
897 zc21 : MinSUP A (UnionCF A f ay supf0 a) |
901 | 898 zc21 = ZChain.minsup zc (o≤-refl0 b) |
899 zc20 : MinSUP.sup zc21 ≡ supf0 a | |
966 | 900 zc20 = sym (ZChain.supf-is-minsup zc (o≤-refl0 b)) |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
901 zc24 : {x₁ : Ordinal} → odef (UnionCF A f ay supf0 a) x₁ → (x₁ ≡ sp1) ∨ (x₁ << sp1) |
950 | 902 zc24 {x₁} ux = MinSUP.x≤sup sup1 (case1 (chain-mono f mf ay supf0 (ZChain.supf-mono zc) (o≤-refl0 b) ux ) ) |
966 | 903 zc18 : supf0 a o≤ sp1 |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
904 zc18 = subst (λ k → k o≤ sp1) zc20( MinSUP.minsup zc21 (MinSUP.as sup1) zc24 ) |
901 | 905 ... | tri> ¬a ¬b c = o≤-refl |
885 | 906 |
966 | 907 fcup : {u z : Ordinal } → FClosure A f (supf1 u) z → u o≤ px → FClosure A f (supf0 u) z |
903 | 908 fcup {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sf1=sf0 u≤px) fc |
966 | 909 fcpu : {u z : Ordinal } → FClosure A f (supf0 u) z → u o≤ px → FClosure A f (supf1 u) z |
903 | 910 fcpu {u} {z} fc u≤px = subst (λ k → FClosure A f k z ) (sym (sf1=sf0 u≤px)) fc |
967 | 911 |
999
3ffbdd53d1ea
fcs<sup requires <-monotonicity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
998
diff
changeset
|
912 -- this is a kind of maximality, so we cannot prove this without <-monotonicity |
3ffbdd53d1ea
fcs<sup requires <-monotonicity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
998
diff
changeset
|
913 -- |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
914 cfcs : {a b w : Ordinal } |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
915 → a o< b → b o≤ x → supf1 a o< b → FClosure A f (supf1 a) w → odef (UnionCF A f ay supf1 b) w |
1090 | 916 cfcs {a} {b} {w} a<b b≤x sa<b fc with x<y∨y≤x (supf0 a) px |
1012 | 917 ... | case2 px≤sa = z50 where |
1023 | 918 a<x : a o< x |
919 a<x = ordtrans<-≤ a<b b≤x | |
1012 | 920 a≤px : a o≤ px |
921 a≤px = subst (λ k → a o< k) (sym (Oprev.oprev=x op)) (ordtrans<-≤ a<b b≤x) | |
922 -- supf0 a ≡ px we cannot use previous cfcs, it is in the chain because | |
923 -- supf0 a ≡ supf0 (supf0 a) ≡ supf0 px o< x | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
924 z50 : odef (UnionCF A f ay supf1 b) w |
1012 | 925 z50 with osuc-≡< px≤sa |
1034 | 926 ... | case1 px=sa = ⟪ A∋fc {A} _ f mf fc , cp ⟫ where |
1023 | 927 sa≤px : supf0 a o≤ px |
928 sa≤px = subst₂ (λ j k → j o< k) px=sa (sym (Oprev.oprev=x op)) px<x | |
1026 | 929 spx=sa : supf0 px ≡ supf0 a |
930 spx=sa = begin | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
931 supf0 px ≡⟨ cong supf0 px=sa ⟩ |
1039 | 932 supf0 (supf0 a) ≡⟨ ZChain.supf-idem zc a≤px sa≤px ⟩ |
1026 | 933 supf0 a ∎ where open ≡-Reasoning |
1020 | 934 z51 : supf0 px o< b |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
935 z51 = subst (λ k → k o< b ) (sym ( begin supf0 px ≡⟨ spx=sa ⟩ |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
936 supf0 a ≡⟨ sym (sf1=sf0 a≤px) ⟩ |
1025 | 937 supf1 a ∎ )) sa<b where open ≡-Reasoning |
1020 | 938 z52 : supf1 a ≡ supf1 (supf0 px) |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
939 z52 = begin supf1 a ≡⟨ sf1=sf0 a≤px ⟩ |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
940 supf0 a ≡⟨ sym (ZChain.supf-idem zc a≤px sa≤px ) ⟩ |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
941 supf0 (supf0 a) ≡⟨ sym (sf1=sf0 sa≤px) ⟩ |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
942 supf1 (supf0 a) ≡⟨ cong supf1 (sym spx=sa) ⟩ |
1025 | 943 supf1 (supf0 px) ∎ where open ≡-Reasoning |
944 z53 : supf1 (supf0 px) ≡ supf0 px | |
945 z53 = begin | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
946 supf1 (supf0 px) ≡⟨ cong supf1 spx=sa ⟩ |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
947 supf1 (supf0 a) ≡⟨ sf1=sf0 sa≤px ⟩ |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
948 supf0 (supf0 a) ≡⟨ sym ( cong supf0 px=sa ) ⟩ |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
949 supf0 px ∎ where open ≡-Reasoning |
1034 | 950 cp : UChain ay b w |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
951 cp = ch-is-sup (supf0 px) z51 z53 (subst (λ k → FClosure A f k w) z52 fc) |
1020 | 952 ... | case2 px<sa = ⊥-elim ( ¬p<x<op ⟪ px<sa , subst₂ (λ j k → j o< k ) (sf1=sf0 a≤px) (sym (Oprev.oprev=x op)) z53 ⟫ ) where |
953 z53 : supf1 a o< x | |
954 z53 = ordtrans<-≤ sa<b b≤x | |
1012 | 955 ... | case1 sa<px with trio< a px |
996 | 956 ... | tri< a<px ¬b ¬c = z50 where |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
957 z50 : odef (UnionCF A f ay supf1 b) w |
997 | 958 z50 with osuc-≡< b≤x |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
959 ... | case2 lt with ZChain.cfcs zc a<b (subst (λ k → b o< k) (sym (Oprev.oprev=x op)) lt) sa<b fc |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
960 ... | ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫ |
1034 | 961 ... | ⟪ az , ch-is-sup u u<b su=u fc ⟫ = ⟪ az , ch-is-sup u u<b (trans (sym (sf=eq u<x)) su=u) (fcpu fc u≤px ) ⟫ where |
962 u≤px : u o≤ px | |
963 u≤px = subst (λ k → u o< k) (sym (Oprev.oprev=x op)) (ordtrans<-≤ u<b b≤x ) | |
964 u<x : u o< x | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
965 u<x = ordtrans<-≤ u<b b≤x |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
966 z50 | case1 eq with ZChain.cfcs zc a<px o≤-refl sa<px fc |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
967 ... | ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫ |
1034 | 968 ... | ⟪ az , ch-is-sup u u<px su=u fc ⟫ = ⟪ az , ch-is-sup u u<b (trans (sym (sf=eq u<x)) su=u) (fcpu fc (o<→≤ u<px)) ⟫ where -- u o< px → u o< b ? |
969 u<b : u o< b | |
970 u<b = subst (λ k → u o< k ) (trans (Oprev.oprev=x op) (sym eq) ) (ordtrans u<px <-osuc ) | |
971 u<x : u o< x | |
972 u<x = subst (λ k → u o< k ) (Oprev.oprev=x op) ( ordtrans u<px <-osuc ) | |
1000 | 973 ... | tri≈ ¬a a=px ¬c = csupf1 where |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
974 -- a ≡ px , b ≡ x, sp o≤ x |
995 | 975 px<b : px o< b |
976 px<b = subst₂ (λ j k → j o< k) a=px refl a<b | |
977 b=x : b ≡ x | |
978 b=x with trio< b x | |
996 | 979 ... | tri< a ¬b ¬c = ⊥-elim ( ¬p<x<op ⟪ px<b , subst (λ k → b o< k ) (sym (Oprev.oprev=x op)) a ⟫ ) -- px o< b o< x |
995 | 980 ... | tri≈ ¬a b ¬c = b |
996 | 981 ... | tri> ¬a ¬b c = ⊥-elim ( o≤> b≤x c ) -- x o< b |
997 | 982 z51 : FClosure A f (supf1 px) w |
983 z51 = subst (λ k → FClosure A f k w) (sym (trans (cong supf1 (sym a=px)) (sf1=sf0 (o≤-refl0 a=px) ))) fc | |
1001 | 984 z53 : odef A w |
985 z53 = A∋fc {A} _ f mf fc | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
986 csupf1 : odef (UnionCF A f ay supf1 b) w |
1090 | 987 csupf1 with x<y∨y≤x px (supf0 px) |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
988 ... | case2 spx≤px = ⟪ z53 , ch-is-sup (supf0 px) z54 z52 fc1 ⟫ where |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
989 z54 : supf0 px o< b |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
990 z54 = subst (λ k → supf0 px o< k ) (trans (Oprev.oprev=x op) (sym b=x) ) spx≤px |
1003 | 991 z52 : supf1 (supf0 px) ≡ supf0 px |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
992 z52 = trans (sf1=sf0 spx≤px ) ( ZChain.supf-idem zc o≤-refl spx≤px ) |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
993 fc1 : FClosure A f (supf1 (supf0 px)) w |
1004 | 994 fc1 = subst (λ k → FClosure A f k w ) (trans (cong supf0 a=px) (sym z52) ) fc |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
995 ... | case1 px<spx = ⊥-elim (¬p<x<op ⟪ px<spx , z54 ⟫ ) where -- supf1 px o≤ spuf1 x → supf1 px ≡ x o< x |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
996 z54 : supf0 px o≤ px |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
997 z54 = subst₂ (λ j k → supf0 j o< k ) a=px (trans b=x (sym (Oprev.oprev=x op))) sa<b |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
998 |
996 | 999 ... | tri> ¬a ¬b c = ⊥-elim ( ¬p<x<op ⟪ c , subst (λ k → a o< k ) (sym (Oprev.oprev=x op)) ( ordtrans<-≤ a<b b≤x) ⟫ ) -- px o< a o< b o≤ x |
994 | 1000 |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1001 zc11 : {z : Ordinal} → odef (UnionCF A f ay supf1 x) z → odef pchainpx z |
1034 | 1002 zc11 {z} ⟪ az , ch-init fc ⟫ = case1 ⟪ az , ch-init fc ⟫ |
1003 zc11 {z} ⟪ az , ch-is-sup u u<x su=u fc ⟫ = zc21 fc where | |
1004 zc21 : {z1 : Ordinal } → FClosure A f (supf1 u) z1 → odef pchainpx z1 | |
903 | 1005 zc21 {z1} (fsuc z2 fc ) with zc21 fc |
1034 | 1006 ... | case1 ⟪ ua1 , ch-init fc₁ ⟫ = case1 ⟪ proj2 ( mf _ ua1) , ch-init (fsuc _ fc₁) ⟫ |
1007 ... | case1 ⟪ ua1 , ch-is-sup u u<x su=u fc₁ ⟫ = case1 ⟪ proj2 ( mf _ ua1) , ch-is-sup u u<x su=u (fsuc _ fc₁) ⟫ | |
1043 | 1008 ... | case2 fc = case2 ⟪ fsuc _ (proj1 fc) , proj2 fc ⟫ |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1009 zc21 (init asp refl ) with trio< (supf0 u) (supf0 px) |
1036 | 1010 ... | tri< a ¬b ¬c = case1 ⟪ asp , ch-is-sup u u<px (trans (sym (sf1=sf0 (o<→≤ u<px))) su=u )(init asp0 (sym (sf1=sf0 (o<→≤ u<px))) ) ⟫ where |
1011 u<px : u o< px | |
1012 u<px = ZChain.supf-inject zc a | |
1013 asp0 : odef A (supf0 u) | |
1014 asp0 = ZChain.asupf zc | |
1043 | 1015 ... | tri≈ ¬a b ¬c = case2 ⟪ (init (subst (λ k → odef A k) b (ZChain.asupf zc) ) |
1044 | 1016 (sym (trans (sf1=sf0 (zc-b<x _ u<x)) b ))) , spx<x ⟫ where |
1017 spx<x : supf0 px o< x | |
1018 spx<x = osucprev ( begin | |
1019 osuc (supf0 px) ≡⟨ cong osuc (sym b) ⟩ | |
1020 osuc (supf0 u) ≡⟨ cong osuc (sym (sf1=sf0 (zc-b<x _ u<x) )) ⟩ | |
1021 osuc (supf1 u) ≡⟨ cong osuc su=u ⟩ | |
1022 osuc u ≤⟨ osucc u<x ⟩ | |
1023 x ∎ ) where open o≤-Reasoning O | |
1036 | 1024 ... | tri> ¬a ¬b c = ⊥-elim ( ¬p<x<op ⟪ ZChain.supf-inject zc c , subst (λ k → u o< k ) (sym (Oprev.oprev=x op)) u<x ⟫ ) |
967 | 1025 |
1035 | 1026 is-minsup : {z : Ordinal} → z o≤ x → IsMinSUP A (UnionCF A f ay supf1 z) (supf1 z) |
1027 is-minsup {z} z≤x with osuc-≡< z≤x | |
1036 | 1028 ... | case1 z=x = record { as = zc22 ; x≤sup = z23 ; minsup = z24 } where |
1029 px<z : px o< z | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1030 px<z = subst (λ k → px o< k) (sym z=x) px<x |
1036 | 1031 zc22 : odef A (supf1 z) |
1032 zc22 = subst (λ k → odef A k ) (sym (sf1=sp1 px<z )) ( MinSUP.as sup1 ) | |
1033 z23 : {w : Ordinal } → odef ( UnionCF A f ay supf1 z ) w → w ≤ supf1 z | |
1045
022d2ef3f20b
is-minsup in px case done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1044
diff
changeset
|
1034 z23 {w} uz = subst (λ k → w ≤ k ) (sym (sf1=sp1 px<z)) ( MinSUP.x≤sup sup1 ( |
1078 | 1035 zc11 (subst (λ k → odef (UnionCF A f ay supf1 k) w) z=x uz ))) |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1036 z24 : {s : Ordinal } → odef A s → ( {w : Ordinal } → odef ( UnionCF A f ay supf1 z ) w → w ≤ s ) |
1036 | 1037 → supf1 z o≤ s |
1038 z24 {s} as sup = subst (λ k → k o≤ s ) (sym (sf1=sp1 px<z)) ( MinSUP.minsup sup1 as z25 ) where | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1039 z25 : {w : Ordinal } → odef pchainpx w → w ≤ s |
1043 | 1040 z25 {w} (case2 fc) = sup ⟪ A∋fc _ f mf (proj1 fc) , ch-is-sup (supf0 px) z28 z27 fc1 ⟫ where |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1041 -- z=x , supf0 px o< x |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1042 z28 : supf0 px o< z -- supf0 px ≡ supf1 px o≤ supf1 x ≡ sp1 o≤ x ≡ z |
1044 | 1043 z28 = subst (λ k → supf0 px o< k) (sym z=x) (proj2 fc) |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1044 z29 : supf0 px o≤ px |
1044 | 1045 z29 = zc-b<x _ (proj2 fc) |
1040 | 1046 z27 : supf1 (supf0 px) ≡ supf0 px |
1047 z27 = trans (sf1=sf0 z29) ( ZChain.supf-idem zc o≤-refl z29 ) | |
1048 fc1 : FClosure A f (supf1 (supf0 px)) w | |
1043 | 1049 fc1 = subst (λ k → FClosure A f k w) (sym z27) (proj1 fc) |
1036 | 1050 z25 {w} (case1 ⟪ ua , ch-init fc ⟫) = sup ⟪ ua , ch-init fc ⟫ |
1051 z25 {w} (case1 ⟪ ua , ch-is-sup u u<x su=u fc ⟫) = sup ⟪ ua , ch-is-sup u u<z | |
1052 (trans (sf1=sf0 u≤px) su=u) (fcpu fc u≤px) ⟫ where | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1053 u≤px : u o< osuc px |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1054 u≤px = ordtrans u<x <-osuc |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1055 u<z : u o< z |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1056 u<z = ordtrans u<x (subst (λ k → px o< k ) (sym z=x) px<x ) |
1036 | 1057 ... | case2 z<x = record { as = zc22 ; x≤sup = z23 ; minsup = z24 } where |
1058 z≤px = zc-b<x _ z<x | |
1059 m = ZChain.is-minsup zc z≤px | |
1060 zc22 : odef A (supf1 z) | |
1061 zc22 = subst (λ k → odef A k ) (sym (sf1=sf0 z≤px)) ( IsMinSUP.as m ) | |
1062 z23 : {w : Ordinal } → odef ( UnionCF A f ay supf1 z ) w → w ≤ supf1 z | |
1063 z23 {w} ⟪ ua , ch-init fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=sf0 z≤px)) ( ZChain.fcy<sup zc z≤px fc ) | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1064 z23 {w} ⟪ ua , ch-is-sup u u<x su=u fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=sf0 z≤px)) |
1036 | 1065 (IsMinSUP.x≤sup m ⟪ ua , ch-is-sup u u<x (trans (sym (sf1=sf0 u≤px )) su=u) (fcup fc u≤px ) ⟫ ) where |
1066 u≤px : u o≤ px | |
1067 u≤px = ordtrans u<x z≤px | |
1068 z24 : {s : Ordinal } → odef A s → ( {w : Ordinal } → odef ( UnionCF A f ay supf1 z ) w → w ≤ s ) | |
1069 → supf1 z o≤ s | |
1070 z24 {s} as sup = subst (λ k → k o≤ s ) (sym (sf1=sf0 z≤px)) ( IsMinSUP.minsup m as z25 ) where | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1071 z25 : {w : Ordinal } → odef ( UnionCF A f ay supf0 z ) w → w ≤ s |
1036 | 1072 z25 {w} ⟪ ua , ch-init fc ⟫ = sup ⟪ ua , ch-init fc ⟫ |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1073 z25 {w} ⟪ ua , ch-is-sup u u<x su=u fc ⟫ = sup ⟪ ua , ch-is-sup u u<x |
1036 | 1074 (trans (sf1=sf0 u≤px) su=u) (fcpu fc u≤px) ⟫ where |
1075 u≤px : u o≤ px | |
1076 u≤px = ordtrans u<x z≤px | |
885 | 1077 |
1059 | 1078 zo≤sz : {z : Ordinal} → z o≤ x → z o≤ supf1 z |
1079 zo≤sz {z} z≤x with osuc-≡< z≤x | |
1080 ... | case2 z<x = subst (λ k → z o≤ k) (sym (sf1=sf0 (zc-b<x _ z<x ))) (ZChain.zo≤sz zc (zc-b<x _ z<x )) | |
1081 ... | case1 refl with osuc-≡< (supf1-mono (o<→≤ (px<x))) -- px o≤ supf1 px o≤ supf1 x ≡ sp1 → x o≤ sp1 | |
1078 | 1082 ... | case2 lt = begin |
1083 x ≡⟨ sym (Oprev.oprev=x op) ⟩ | |
1084 osuc px ≤⟨ osucc (ZChain.zo≤sz zc o≤-refl) ⟩ | |
1085 osuc (supf0 px) ≡⟨ sym (cong osuc (sf1=sf0 o≤-refl )) ⟩ | |
1086 osuc (supf1 px) ≤⟨ osucc lt ⟩ | |
1059 | 1087 supf1 x ∎ where open o≤-Reasoning O |
1088 ... | case1 spx=sx with osuc-≡< ( ZChain.zo≤sz zc o≤-refl ) | |
1089 ... | case2 lt = begin | |
1078 | 1090 x ≡⟨ sym (Oprev.oprev=x op) ⟩ |
1091 osuc px ≤⟨ osucc lt ⟩ | |
1092 supf0 px ≡⟨ sym (sf1=sf0 o≤-refl) ⟩ | |
1093 supf1 px ≤⟨ supf1-mono (o<→≤ px<x) ⟩ | |
1059 | 1094 supf1 x ∎ where open o≤-Reasoning O |
1095 ... | case1 px=spx = ⊥-elim ( <<-irr zc40 (proj1 ( mf< (supf0 px) (ZChain.asupf zc))) ) where | |
1096 zc37 : supf0 px ≡ px | |
1097 zc37 = sym px=spx | |
1098 zc39 : supf0 px ≡ sp1 | |
1099 zc39 = begin | |
1100 supf0 px ≡⟨ sym (sf1=sf0 o≤-refl) ⟩ | |
1101 supf1 px ≡⟨ spx=sx ⟩ | |
1102 supf1 x ≡⟨ sf1=sp1 px<x ⟩ | |
1103 sp1 ∎ where open ≡-Reasoning | |
1078 | 1104 zc40 : f (supf0 px) ≤ supf0 px |
1105 zc40 = subst (λ k → f (supf0 px) ≤ k ) (sym zc39) | |
1059 | 1106 ( MinSUP.x≤sup sup1 (case2 ⟪ fsuc _ (init (ZChain.asupf zc) refl) , subst (λ k → k o< x) (sym zc37) px<x ⟫ )) |
1107 | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1108 ... | no lim with trio< x o∅ |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1109 ... | tri< a ¬b ¬c = ⊥-elim ( ¬x<0 a ) |
1090 | 1110 ... | tri≈ ¬a x=0 ¬c = record { supf = λ _ → MinSUP.sup (ysup f mf ay) ; asupf = MinSUP.as (ysup f mf ay) ; supf-mono = λ _ → o≤-refl |
1089
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1111 ; zo≤sz = zo≤sz ; is-minsup = is-minsup ; cfcs = λ a<b b≤0 → ⊥-elim ( ¬x<0 (subst (λ k → _ o< k ) x=0 (ordtrans<-≤ a<b b≤0))) } where |
1101 | 1112 -- initial case |
1089
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1113 |
1081 | 1114 mf : ≤-monotonic-f A f |
1115 mf x ax = ⟪ case2 mf00 , proj2 (mf< x ax ) ⟫ where | |
1116 mf00 : * x < * (f x) | |
1117 mf00 = proj1 ( mf< x ax ) | |
1118 ym = MinSUP.sup (ysup f mf ay) | |
1083 | 1119 |
1081 | 1120 zo≤sz : {z : Ordinal} → z o≤ x → z o≤ MinSUP.sup (ysup f mf ay) |
1121 zo≤sz {z} z≤x with osuc-≡< z≤x | |
1083 | 1122 ... | case1 refl = subst (λ k → k o≤ _) (sym x=0) o∅≤z |
1081 | 1123 ... | case2 lt = ⊥-elim ( ¬x<0 (subst (λ k → z o< k ) x=0 lt ) ) |
1083 | 1124 |
1081 | 1125 is-minsup : {z : Ordinal} → z o≤ x → |
1126 IsMinSUP A (UnionCF A f ay (λ _ → MinSUP.sup (ysup f mf ay)) z) (MinSUP.sup (ysup f mf ay)) | |
1127 is-minsup {z} z≤x with osuc-≡< z≤x | |
1128 ... | case1 refl = record { as = MinSUP.as (ysup f mf ay) ; x≤sup = λ {w} uw → is00 uw ; minsup = λ {s} as sup → is01 as sup } where | |
1129 is00 : {w : Ordinal } → odef (UnionCF A f ay (λ _ → MinSUP.sup (ysup f mf ay)) x ) w → w ≤ MinSUP.sup (ysup f mf ay) | |
1130 is00 {w} ⟪ aw , ch-init fc ⟫ = MinSUP.x≤sup (ysup f mf ay) fc | |
1131 is00 {w} ⟪ aw , ch-is-sup u u<z su=u fc ⟫ = ⊥-elim (¬x<0 (subst (λ k → u o< k ) x=0 u<z )) | |
1132 is01 : { s : Ordinal } → odef A s → ( {w : Ordinal } → odef (UnionCF A f ay (λ _ → MinSUP.sup (ysup f mf ay)) x ) w → w ≤ s ) | |
1133 → ym o≤ s | |
1134 is01 {s} as sup = MinSUP.minsup (ysup f mf ay) as is02 where | |
1135 is02 : {w : Ordinal } → odef (uchain f mf ay) w → (w ≡ s) ∨ (w << s) | |
1136 is02 fc = sup ⟪ A∋fc _ f mf fc , ch-init fc ⟫ | |
1137 ... | case2 lt = ⊥-elim ( ¬x<0 (subst (λ k → z o< k ) x=0 lt ) ) | |
1089
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1138 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1139 ... | tri> ¬a ¬b 0<x = zc400 usup ssup where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1140 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1141 mf : ≤-monotonic-f A f |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1142 mf x ax = ⟪ case2 mf00 , proj2 (mf< x ax ) ⟫ where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1143 mf00 : * x < * (f x) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1144 mf00 = proj1 ( mf< x ax ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1145 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1146 pzc : {z : Ordinal} → z o< x → ZChain A f mf< ay z |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1147 pzc {z} z<x = prev z z<x |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1148 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1149 ysp = MinSUP.sup (ysup f mf ay) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1150 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1151 supfz : {z : Ordinal } → z o< x → Ordinal |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1152 supfz {z} z<x = ZChain.supf (pzc (ob<x lim z<x)) z |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1153 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1154 pchainU : HOD |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1155 pchainU = UnionIC A f ay supfz |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1156 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1157 zeq : {xa xb z : Ordinal } |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1158 → (xa<x : xa o< x) → (xb<x : xb o< x) → xa o≤ xb → z o≤ xa |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1159 → ZChain.supf (pzc xa<x) z ≡ ZChain.supf (pzc xb<x) z |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1160 zeq {xa} {xb} {z} xa<x xb<x xa≤xb z≤xa = supf-unique A f mf< ay xa≤xb |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1161 (pzc xa<x) (pzc xb<x) z≤xa |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1162 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1163 iceq : {ix iy : Ordinal } → ix ≡ iy → {i<x : ix o< x } {i<y : iy o< x } → supfz i<x ≡ supfz i<y |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1164 iceq refl = cong supfz o<-irr |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1165 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1166 IChain-i : {z : Ordinal } → IChain ay supfz z → Ordinal |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1167 IChain-i (ic-init fc) = o∅ |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1168 IChain-i (ic-isup ia ia<x sa<x fca) = ia |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1169 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1170 pic<x : {z : Ordinal } → (ic : IChain ay supfz z ) → osuc (IChain-i ic) o< x |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1171 pic<x {z} (ic-init fc) = ob<x lim 0<x -- 0<x ∧ lim x → osuc o∅ o< x |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1172 pic<x {z} (ic-isup ia ia<x sa<x fca) = ob<x lim ia<x |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1173 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1174 pchainU⊆chain : {z : Ordinal } → (pz : odef pchainU z) → odef (ZChain.chain (pzc (pic<x (proj2 pz)))) z |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1175 pchainU⊆chain {z} ⟪ aw , ic-init fc ⟫ = ⟪ aw , ch-init fc ⟫ |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1176 pchainU⊆chain {z} ⟪ aw , (ic-isup ia ia<x sa<x fca) ⟫ = ZChain.cfcs (pzc (ob<x lim ia<x) ) <-osuc o≤-refl uz03 fca where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1177 uz02 : FClosure A f (ZChain.supf (pzc (ob<x lim ia<x)) ia ) z |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1178 uz02 = fca |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1179 uz03 : ZChain.supf (pzc (ob<x lim ia<x)) ia o≤ ia |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1180 uz03 = sa<x |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1181 |
1091 | 1182 chain⊆pchainU : {z w : Ordinal } → (z<x : z o< x) → odef (UnionCF A f ay (ZChain.supf (pzc (ob<x lim z<x))) z) w → odef pchainU w |
1183 chain⊆pchainU {z} {w} z<x ⟪ aw , ch-init fc ⟫ = ⟪ aw , ic-init fc ⟫ | |
1184 chain⊆pchainU {z} {w} z<x ⟪ aw , ch-is-sup u u<oz su=u fc ⟫ | |
1089
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1185 = ⟪ aw , ic-isup u u<x (o≤-refl0 su≡u) (subst (λ k → FClosure A f k w ) su=su fc) ⟫ where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1186 u<x : u o< x |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1187 u<x = ordtrans u<oz z<x |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1188 su=su : ZChain.supf (pzc (ob<x lim z<x)) u ≡ supfz u<x |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1189 su=su = sym ( zeq _ _ (o<→≤ (osucc u<oz)) (o<→≤ <-osuc) ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1190 su≡u : supfz u<x ≡ u |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1191 su≡u = begin |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1192 ZChain.supf (pzc (ob<x lim u<x )) u ≡⟨ sym su=su ⟩ |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1193 ZChain.supf (pzc (ob<x lim z<x)) u ≡⟨ su=u ⟩ |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1194 u ∎ where open ≡-Reasoning |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1195 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1196 IC⊆ : {a b : Ordinal } (ia : IChain ay supfz a ) (ib : IChain ay supfz b ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1197 → IChain-i ia o< IChain-i ib → odef (ZChain.chain (pzc (pic<x ib))) a |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1198 IC⊆ {a} {b} (ic-init fc ) ib ia<ib = ⟪ A∋fc _ f mf fc , ch-init fc ⟫ |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1199 IC⊆ {a} {b} (ic-isup i i<x sa<x fc ) (ic-init fcb ) ia<ib = ⊥-elim ( ¬x<0 ia<ib ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1200 IC⊆ {a} {b} (ic-isup i i<x sa<x fc ) (ic-isup j j<x sb<x fcb ) ia<ib |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1201 = ZChain.cfcs (pzc (ob<x lim j<x) ) (o<→≤ ia<ib) o≤-refl (OrdTrans (ZChain.supf-mono (pzc (ob<x lim j<x)) (o<→≤ ia<ib)) sb<x) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1202 (subst (λ k → FClosure A f k a) (zeq _ _ (osucc (o<→≤ ia<ib)) (o<→≤ <-osuc)) fc ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1203 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1204 ptotalU : IsTotalOrderSet pchainU |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1205 ptotalU {a} {b} ia ib with trio< (IChain-i (proj2 ia)) (IChain-i (proj2 ib)) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1206 ... | tri< ia<ib ¬b ¬c = ZChain.f-total (pzc (pic<x (proj2 ib))) (IC⊆ (proj2 ia) (proj2 ib) ia<ib) (pchainU⊆chain ib) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1207 ... | tri≈ ¬a ia=ib ¬c = subst₂ (λ j k → Tri (j < k) (j ≡ k) (k < j)) *iso *iso ( pcmp (proj2 ia) (proj2 ib) ia=ib ) where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1208 pcmp : (ia : IChain ay supfz (& a)) → (ib : IChain ay supfz (& b)) → IChain-i ia ≡ IChain-i ib |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1209 → Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1210 pcmp (ic-init fca) (ic-init fcb) eq = fcn-cmp _ f mf fca fcb |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1211 pcmp (ic-init fca) (ic-isup i i<x s<x fcb) eq with ZChain.fcy<sup (pzc i<x) o≤-refl fca |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1212 ... | case1 eq1 = ct22 where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1213 ct22 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1214 ct22 with subst (λ k → k ≤ & b) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl )) (s≤fc _ f mf fcb ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1215 ... | case1 eq2 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1216 ct00 : * (& a) ≡ * (& b) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1217 ct00 = cong (*) (trans eq1 eq2) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1218 ... | case2 lt = tri< fc11 (λ eq → <-irr (case1 (sym eq)) fc11) (λ lt → <-irr (case2 fc11) lt) where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1219 fc11 : * (& a) < * (& b) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1220 fc11 = subst (λ k → k < * (& b) ) (cong (*) (sym eq1)) lt |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1221 ... | case2 lt = tri< fc11 (λ eq → <-irr (case1 (sym eq)) fc11) (λ lt → <-irr (case2 fc11) lt) where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1222 fc11 : * (& a) < * (& b) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1223 fc11 = ftrans<-≤ lt (subst (λ k → k ≤ & b) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl )) (s≤fc _ f mf fcb ) ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1224 pcmp (ic-isup i i<x s<x fca) (ic-init fcb) eq with ZChain.fcy<sup (pzc i<x) o≤-refl fcb |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1225 ... | case1 eq1 = ct22 where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1226 ct22 : Tri (* (& a) < * (& b)) (* (& a) ≡ * (& b)) (* (& b) < * (& a) ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1227 ct22 with subst (λ k → k ≤ & a) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl )) (s≤fc _ f mf fca ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1228 ... | case1 eq2 = tri≈ (λ lt → ⊥-elim (<-irr (case1 (sym ct00)) lt)) ct00 (λ lt → ⊥-elim (<-irr (case1 ct00) lt)) where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1229 ct00 : * (& a) ≡ * (& b) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1230 ct00 = cong (*) (sym (trans eq1 eq2)) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1231 ... | case2 lt = tri> (λ lt → <-irr (case2 fc11) lt) (λ eq → <-irr (case1 eq) fc11) fc11 where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1232 fc11 : * (& b) < * (& a) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1233 fc11 = subst (λ k → k < * (& a) ) (cong (*) (sym eq1)) lt |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1234 ... | case2 lt = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1235 fc12 : * (& b) < * (& a) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1236 fc12 = ftrans<-≤ lt (subst (λ k → k ≤ & a) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl )) (s≤fc _ f mf fca ) ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1237 pcmp (ic-isup i i<x s<x fca) (ic-isup i i<y s<y fcb) refl = fcn-cmp _ f mf fca (subst (λ k → FClosure A f k (& b)) pc01 fcb ) where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1238 pc01 : supfz i<y ≡ supfz i<x |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1239 pc01 = cong supfz o<-irr |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1240 ... | tri> ¬a ¬b ib<ia = ZChain.f-total (pzc (pic<x (proj2 ia))) (pchainU⊆chain ia) (IC⊆ (proj2 ib) (proj2 ia) ib<ia) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1241 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1242 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1243 usup : MinSUP A pchainU |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1244 usup = minsupP pchainU (λ ic → proj1 ic ) ptotalU |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1245 spu0 = MinSUP.sup usup |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1246 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1247 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1248 pchainS : HOD |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1249 pchainS = record { od = record { def = λ z → (odef A z ∧ IChain ay supfz z ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1250 ∨ (FClosure A f spu0 z ∧ (spu0 o< x)) } ; odmax = & A ; <odmax = zc00 } where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1251 zc00 : {z : Ordinal } → (odef A z ∧ IChain ay supfz z ) ∨ (FClosure A f spu0 z ∧ (spu0 o< x) )→ z o< & A |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1252 zc00 {z} (case1 lt) = z07 lt |
1097 | 1253 zc00 {z} (case2 fc) = odef< ( A∋fc spu0 f mf (proj1 fc) ) |
1089
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1254 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1255 zc02 : { a b : Ordinal } → odef A a ∧ IChain ay supfz a → FClosure A f spu0 b ∧ ( spu0 o< x) → a ≤ b |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1256 zc02 {a} {b} ca fb = zc05 (proj1 fb) where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1257 zc05 : {b : Ordinal } → FClosure A f spu0 b → a ≤ b |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1258 zc05 (fsuc b1 fb ) with proj1 ( mf b1 (A∋fc spu0 f mf fb )) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1259 ... | case1 eq = subst (λ k → a ≤ k ) eq (zc05 fb) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1260 ... | case2 lt = ≤-ftrans (zc05 fb) (case2 lt) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1261 zc05 (init b1 refl) = MinSUP.x≤sup usup ca |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1262 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1263 ptotalS : IsTotalOrderSet pchainS |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1264 ptotalS (case1 a) (case1 b) = ptotalU a b |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1265 ptotalS {a0} {b0} (case1 a) (case2 b) with zc02 a b |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1266 ... | case1 eq = tri≈ (<-irr (case1 (sym eq1))) eq1 (<-irr (case1 eq1)) where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1267 eq1 : a0 ≡ b0 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1268 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1269 ... | case2 lt = tri< lt1 (λ eq → <-irr (case1 (sym eq)) lt1) (<-irr (case2 lt1)) where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1270 lt1 : a0 < b0 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1271 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1272 ptotalS {b0} {a0} (case2 b) (case1 a) with zc02 a b |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1273 ... | case1 eq = tri≈ (<-irr (case1 eq1)) (sym eq1) (<-irr (case1 (sym eq1))) where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1274 eq1 : a0 ≡ b0 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1275 eq1 = subst₂ (λ j k → j ≡ k ) *iso *iso (cong (*) eq ) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1276 ... | case2 lt = tri> (<-irr (case2 lt1)) (λ eq → <-irr (case1 eq) lt1) lt1 where |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1277 lt1 : a0 < b0 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1278 lt1 = subst₂ (λ j k → j < k ) *iso *iso lt |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1279 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)) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1280 |
1091 | 1281 S⊆A : pchainS ⊆ A |
1089
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1282 S⊆A (case1 lt) = proj1 lt |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1283 S⊆A (case2 fc) = A∋fc _ f mf (proj1 fc) |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1284 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1285 ssup : MinSUP A pchainS |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1286 ssup = minsupP pchainS S⊆A ptotalS |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1287 |
b627e3ea7266
try to hide spu from source
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1084
diff
changeset
|
1288 zc400 : MinSUP A pchainU → MinSUP A pchainS → ZChain A f mf< ay x |
1090 | 1289 zc400 usup ssup = record { supf = supf1 ; asupf = asupf ; zo≤sz = zo≤sz ; is-minsup = is-minsup ; cfcs = cfcs ; supf-mono = supf-mono } where |
726
b2e2cd12e38f
psupf-mono and is-max conflict
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
725
diff
changeset
|
1290 |
1069 | 1291 spu = MinSUP.sup usup |
1292 sps = MinSUP.sup ssup | |
834 | 1293 |
794 | 1294 supf1 : Ordinal → Ordinal |
835 | 1295 supf1 z with trio< z x |
1069 | 1296 ... | tri< a ¬b ¬c = ZChain.supf (pzc (ob<x lim a)) z -- each sup o< x |
1297 ... | tri≈ ¬a b ¬c = spu -- sup of all sup o< x | |
1078 | 1298 ... | tri> ¬a ¬b c = sps -- sup of spu which o< x |
1069 | 1299 -- if x o< spu, spu is not included in UnionCF x |
1300 -- the chain | |
755 | 1301 |
1010
f80d525e6a6b
Recursive record IChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1009
diff
changeset
|
1302 pchain : HOD |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1303 pchain = UnionCF A f ay supf1 x |
1010
f80d525e6a6b
Recursive record IChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1009
diff
changeset
|
1304 |
1069 | 1305 -- pchain ⊆ pchainU ⊆ pchianS |
1009 | 1306 |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1307 sf1=sf : {z : Ordinal } → (a : z o< x ) → supf1 z ≡ ZChain.supf (pzc (ob<x lim a)) z |
1009 | 1308 sf1=sf {z} z<x with trio< z x |
1309 ... | tri< a ¬b ¬c = cong ( λ k → ZChain.supf (pzc (ob<x lim k)) z) o<-irr | |
1310 ... | tri≈ ¬a b ¬c = ⊥-elim (¬a z<x) | |
1311 ... | tri> ¬a ¬b c = ⊥-elim (¬a z<x) | |
1312 | |
1071 | 1313 sf1=spu : {z : Ordinal } → x ≡ z → supf1 z ≡ spu |
1314 sf1=spu {z} eq with trio< z x | |
1078 | 1315 ... | tri< a ¬b ¬c = ⊥-elim (¬b (sym eq)) |
1009 | 1316 ... | tri≈ ¬a b ¬c = refl |
1078 | 1317 ... | tri> ¬a ¬b c = ⊥-elim (¬b (sym eq)) |
1009 | 1318 |
1071 | 1319 sf1=sps : {z : Ordinal } → (a : x o< z ) → supf1 z ≡ sps |
1320 sf1=sps {z} x<z with trio< z x | |
1321 ... | tri< a ¬b ¬c = ⊥-elim (o<> x<z a) | |
1322 ... | tri≈ ¬a b ¬c = ⊥-elim (¬c x<z ) | |
1323 ... | tri> ¬a ¬b c = refl | |
1069 | 1324 |
1057 | 1325 asupf : {z : Ordinal } → odef A (supf1 z) |
1326 asupf {z} with trio< z x | |
1078 | 1327 ... | tri< a ¬b ¬c = ZChain.asupf (pzc (ob<x lim a)) |
1057 | 1328 ... | tri≈ ¬a b ¬c = MinSUP.as usup |
1069 | 1329 ... | tri> ¬a ¬b c = MinSUP.as ssup |
1057 | 1330 |
1331 supf-mono : {z y : Ordinal } → z o≤ y → supf1 z o≤ supf1 y | |
1332 supf-mono {z} {y} z≤y with trio< y x | |
1071 | 1333 ... | tri< y<x ¬b ¬c = zc01 where |
1057 | 1334 open o≤-Reasoning O |
1071 | 1335 zc01 : supf1 z o≤ ZChain.supf (pzc (ob<x lim y<x)) y |
1072 | 1336 zc01 = begin |
1337 supf1 z ≡⟨ sf1=sf (ordtrans≤-< z≤y y<x) ⟩ | |
1338 ZChain.supf (pzc (ob<x lim (ordtrans≤-< z≤y y<x))) z ≡⟨ zeq _ _ (osucc z≤y) (o<→≤ <-osuc) ⟩ | |
1070 | 1339 ZChain.supf (pzc (ob<x lim y<x)) z ≤⟨ ZChain.supf-mono (pzc (ob<x lim y<x)) z≤y ⟩ |
1078 | 1340 ZChain.supf (pzc (ob<x lim y<x)) y ∎ |
1072 | 1341 ... | tri≈ ¬a b ¬c = zc01 where -- supf1 z o≤ spu |
1342 open o≤-Reasoning O | |
1343 zc01 : supf1 z o≤ spu | |
1344 zc01 with osuc-≡< (subst (λ k → z o≤ k) b z≤y) | |
1345 ... | case1 z=x = o≤-refl0 (sf1=spu (sym z=x)) | |
1078 | 1346 ... | case2 z<x = subst (λ k → k o≤ spu ) (sym (sf1=sf z<x)) ( IsMinSUP.minsup (ZChain.is-minsup (pzc (ob<x lim z<x)) (o<→≤ <-osuc) ) |
1091 | 1347 (MinSUP.as usup) (λ uw → MinSUP.x≤sup usup (chain⊆pchainU z<x uw)) ) |
1072 | 1348 ... | tri> ¬a ¬b c = zc01 where -- supf1 z o≤ sps |
1349 zc01 : supf1 z o≤ sps | |
1350 zc01 with trio< z x | |
1078 | 1351 ... | tri< z<x ¬b ¬c = IsMinSUP.minsup (ZChain.is-minsup (pzc (ob<x lim z<x)) (o<→≤ <-osuc) ) |
1091 | 1352 (MinSUP.as ssup) (λ uw → MinSUP.x≤sup ssup (case1 (chain⊆pchainU z<x uw)) ) |
1072 | 1353 ... | tri≈ ¬a z=x ¬c = MinSUP.minsup usup (MinSUP.as ssup) (λ uw → MinSUP.x≤sup ssup (case1 uw) ) |
1354 ... | tri> ¬a ¬b c = o≤-refl -- (sf1=sps c) | |
797 | 1355 |
1035 | 1356 is-minsup : {z : Ordinal} → z o≤ x → IsMinSUP A (UnionCF A f ay supf1 z) (supf1 z) |
1073 | 1357 is-minsup {z} z≤x with osuc-≡< z≤x |
1075 | 1358 ... | case1 z=x = record { as = asupf ; x≤sup = zm00 ; minsup = zm01 } where |
1078 | 1359 zm00 : {w : Ordinal } → odef (UnionCF A f ay supf1 z) w → w ≤ supf1 z |
1360 zm00 {w} ⟪ az , ch-init fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=spu (sym z=x))) ( MinSUP.x≤sup usup ⟪ az , ic-init fc ⟫ ) | |
1361 zm00 {w} ⟪ az , ch-is-sup u u<b su=u fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=spu (sym z=x))) | |
1076 | 1362 ( MinSUP.x≤sup usup ⟪ az , ic-isup u u<x (o≤-refl0 zm05) (subst (λ k → FClosure A f k w) (sym zm06) fc) ⟫ ) where |
1363 u<x : u o< x | |
1364 u<x = subst (λ k → u o< k) z=x u<b | |
1365 zm06 : supfz (subst (λ k → u o< k) z=x u<b) ≡ supf1 u | |
1366 zm06 = trans (zeq _ _ o≤-refl (o<→≤ <-osuc) ) (sym (sf1=sf u<x )) | |
1078 | 1367 zm05 : supfz (subst (λ k → u o< k) z=x u<b) ≡ u |
1076 | 1368 zm05 = trans zm06 su=u |
1075 | 1369 zm01 : { s : Ordinal } → odef A s → ( {x : Ordinal } → odef (UnionCF A f ay supf1 z) x → x ≤ s ) → supf1 z o≤ s |
1078 | 1370 zm01 {s} as sup = subst (λ k → k o≤ s ) (sym (sf1=spu (sym z=x))) ( MinSUP.minsup usup as zm02 ) where |
1075 | 1371 zm02 : {w : Ordinal } → odef pchainU w → w ≤ s |
1076 | 1372 zm02 {w} uw with pchainU⊆chain uw |
1078 | 1373 ... | ⟪ az , ch-init fc ⟫ = sup ⟪ az , ch-init fc ⟫ |
1076 | 1374 ... | ⟪ az , ch-is-sup u1 u<b su=u fc ⟫ = sup ⟪ az , ch-is-sup u1 (ordtrans u<b zm05) (trans zm03 su=u) zm04 ⟫ where |
1078 | 1375 zm05 : osuc (IChain-i (proj2 uw)) o< z |
1076 | 1376 zm05 = subst (λ k → osuc (IChain-i (proj2 uw)) o< k) (sym z=x) ( pic<x (proj2 uw) ) |
1377 u<x : u1 o< x | |
1378 u<x = subst (λ k → u1 o< k) z=x ( ordtrans u<b zm05 ) | |
1379 zm03 : supf1 u1 ≡ ZChain.supf (prev (osuc (IChain-i (proj2 uw))) (pic<x (proj2 uw))) u1 | |
1380 zm03 = trans (sf1=sf u<x) (zeq _ _ (osucc u<b) (o<→≤ <-osuc) ) | |
1381 zm04 : FClosure A f (supf1 u1) w | |
1382 zm04 = subst (λ k → FClosure A f k w) (sym zm03) fc | |
1073 | 1383 ... | case2 z<x = record { as = asupf ; x≤sup = zm00 ; minsup = zm01 } where |
1078 | 1384 supf0 = ZChain.supf (pzc (ob<x lim z<x)) |
1073 | 1385 msup : IsMinSUP A (UnionCF A f ay supf0 z) (supf0 z) |
1386 msup = ZChain.is-minsup (pzc (ob<x lim z<x)) (o<→≤ <-osuc) | |
1074 | 1387 s1=0 : {u : Ordinal } → u o< z → supf1 u ≡ supf0 u |
1078 | 1388 s1=0 {u} u<z = trans (sf1=sf (ordtrans u<z z<x)) (zeq _ _ (o<→≤ (osucc u<z)) (o<→≤ <-osuc) ) |
1389 zm00 : {w : Ordinal } → odef (UnionCF A f ay supf1 z) w → w ≤ supf1 z | |
1390 zm00 {w} ⟪ az , ch-init fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=sf z<x)) ( IsMinSUP.x≤sup msup ⟪ az , ch-init fc ⟫ ) | |
1391 zm00 {w} ⟪ az , ch-is-sup u u<b su=u fc ⟫ = subst (λ k → w ≤ k ) (sym (sf1=sf z<x)) | |
1074 | 1392 ( IsMinSUP.x≤sup msup ⟪ az , ch-is-sup u u<b (trans (sym (s1=0 u<b)) su=u) (subst (λ k → FClosure A f k w) (s1=0 u<b) fc) ⟫ ) |
1073 | 1393 zm01 : { s : Ordinal } → odef A s → ( {x : Ordinal } → odef (UnionCF A f ay supf1 z) x → x ≤ s ) → supf1 z o≤ s |
1078 | 1394 zm01 {s} as sup = subst (λ k → k o≤ s ) (sym (sf1=sf z<x)) ( IsMinSUP.minsup msup as zm02 ) where |
1073 | 1395 zm02 : {w : Ordinal } → odef (UnionCF A f ay supf0 z) w → w ≤ s |
1396 zm02 {w} ⟪ az , ch-init fc ⟫ = sup ⟪ az , ch-init fc ⟫ | |
1078 | 1397 zm02 {w} ⟪ az , ch-is-sup u u<b su=u fc ⟫ = sup |
1398 ⟪ az , ch-is-sup u u<b (trans (s1=0 u<b) su=u) (subst (λ k → FClosure A f k w) (sym (s1=0 u<b)) fc) ⟫ | |
1073 | 1399 |
1400 | |
1039 | 1401 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 |
1402 cfcs {a} {b} {w} a<b b≤x sa<b fc with osuc-≡< b≤x | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1403 ... | case1 b=x with trio< a x |
1016 | 1404 ... | tri< a<x ¬b ¬c = zc40 where |
1405 sa = ZChain.supf (pzc (ob<x lim a<x)) a | |
1020 | 1406 m = omax a sa -- x is limit ordinal, so we have sa o< m o< x |
1016 | 1407 m<x : m o< x |
1408 m<x with trio< a sa | inspect (omax a) sa | |
1020 | 1409 ... | tri< a<sa ¬b ¬c | record { eq = eq } = ob<x lim (ordtrans<-≤ sa<b b≤x ) |
1016 | 1410 ... | tri≈ ¬a a=sa ¬c | record { eq = eq } = subst (λ k → k o< x) eq zc41 where |
1411 zc41 : omax a sa o< x | |
1412 zc41 = osucprev ( begin | |
1413 osuc ( omax a sa ) ≡⟨ cong (λ k → osuc (omax a k)) (sym a=sa) ⟩ | |
1414 osuc ( omax a a ) ≡⟨ cong osuc (omxx _) ⟩ | |
1415 osuc ( osuc a ) ≤⟨ o<→≤ (ob<x lim (ob<x lim a<x)) ⟩ | |
1416 x ∎ ) where open o≤-Reasoning O | |
1417 ... | tri> ¬a ¬b c | record { eq = eq } = ob<x lim a<x | |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1418 sam = ZChain.supf (pzc (ob<x lim m<x)) a |
1016 | 1419 zc42 : osuc a o≤ osuc m |
1420 zc42 = osucc (o<→≤ ( omax-x _ _ ) ) | |
1421 sam<m : sam o< m | |
1039 | 1422 sam<m = subst (λ k → k o< m ) (supf-unique A f mf< ay zc42 (pzc (ob<x lim a<x)) (pzc (ob<x lim m<x)) (o<→≤ <-osuc)) ( omax-y _ _ ) |
1016 | 1423 fcm : FClosure A f (ZChain.supf (pzc (ob<x lim m<x)) a) w |
1424 fcm = subst (λ k → FClosure A f k w ) (zeq (ob<x lim a<x) (ob<x lim m<x) zc42 (o<→≤ <-osuc) ) fc | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1425 zcm : odef (UnionCF A f ay (ZChain.supf (pzc (ob<x lim m<x))) (osuc (omax a sa))) w |
1039 | 1426 zcm = ZChain.cfcs (pzc (ob<x lim m<x)) (o<→≤ (omax-x _ _)) o≤-refl (o<→≤ sam<m) fcm |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1427 zc40 : odef (UnionCF A f ay supf1 b) w |
1039 | 1428 zc40 with ZChain.cfcs (pzc (ob<x lim m<x)) (o<→≤ (omax-x _ _)) o≤-refl (o<→≤ sam<m) fcm |
1058 | 1429 ... | ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫ |
1070 | 1430 ... | ⟪ az , ch-is-sup u u<x su=u fc ⟫ = ⟪ az , ch-is-sup u u<b (trans zc45 su=u) zc44 ⟫ where |
1431 u<b : u o< b | |
1432 u<b = osucprev ( begin | |
1433 osuc u ≤⟨ osucc u<x ⟩ | |
1434 osuc m ≤⟨ osucc m<x ⟩ | |
1435 x ≡⟨ sym b=x ⟩ | |
1436 b ∎ ) where open o≤-Reasoning O | |
1078 | 1437 zc45 : supf1 u ≡ ZChain.supf (pzc (ob<x lim m<x)) u |
1070 | 1438 zc45 = begin |
1439 supf1 u ≡⟨ sf1=sf (subst (λ k → u o< k) b=x u<b ) ⟩ | |
1440 ZChain.supf (pzc (ob<x lim (subst (λ k → u o< k) b=x u<b ))) u ≡⟨ zeq _ _ (osucc u<x) (o<→≤ <-osuc) ⟩ | |
1078 | 1441 ZChain.supf (pzc (ob<x lim m<x)) u ∎ where open ≡-Reasoning |
1070 | 1442 zc44 : FClosure A f (supf1 u) w |
1443 zc44 = subst (λ k → FClosure A f k w ) (sym zc45) fc | |
1016 | 1444 ... | tri≈ ¬a b ¬c = ⊥-elim ( ¬a (ordtrans<-≤ a<b b≤x)) |
1445 ... | tri> ¬a ¬b c = ⊥-elim ( ¬a (ordtrans<-≤ a<b b≤x)) | |
1039 | 1446 cfcs {a} {b} {w} a<b b≤x sa<b fc | case2 b<x = zc40 where |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1447 supfb = ZChain.supf (pzc (ob<x lim b<x)) |
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1448 sb=sa : {a : Ordinal } → a o< b → supf1 a ≡ ZChain.supf (pzc (ob<x lim b<x)) a |
1020 | 1449 sb=sa {a} a<b = trans (sf1=sf (ordtrans<-≤ a<b b≤x)) (zeq _ _ (o<→≤ (osucc a<b)) (o<→≤ <-osuc) ) |
1017
ffdfd8d1303a
trying cscf as odef (UnionCF A f mf ay supf z) w
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1016
diff
changeset
|
1450 fcb : FClosure A f (supfb a) w |
1042
912ca4abe806
pxhainx conditon is requied?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1041
diff
changeset
|
1451 fcb = subst (λ k → FClosure A f k w) (sb=sa a<b) fc |
1020 | 1452 -- supfb a o< b assures it is in Union b |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1453 zcb : odef (UnionCF A f ay supfb b) w |
1039 | 1454 zcb = ZChain.cfcs (pzc (ob<x lim b<x)) a<b (o<→≤ <-osuc) (subst (λ k → k o< b) (sb=sa a<b) sa<b) fcb |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1455 zc40 : odef (UnionCF A f ay supf1 b) w |
1016 | 1456 zc40 with zcb |
1058 | 1457 ... | ⟪ az , ch-init fc ⟫ = ⟪ az , ch-init fc ⟫ |
1070 | 1458 ... | ⟪ az , ch-is-sup u u<x su=u fc ⟫ = ⟪ az , ch-is-sup u u<x (trans zc45 su=u) zc44 ⟫ where |
1078 | 1459 zc45 : supf1 u ≡ ZChain.supf (pzc (ob<x lim b<x)) u |
1070 | 1460 zc45 = begin |
1461 supf1 u ≡⟨ sf1=sf (ordtrans u<x b<x) ⟩ | |
1462 ZChain.supf (pzc (ob<x lim (ordtrans u<x b<x) )) u ≡⟨ zeq _ _ (o<→≤ (osucc u<x)) (o<→≤ <-osuc) ⟩ | |
1078 | 1463 ZChain.supf (pzc (ob<x lim b<x )) u ∎ where open ≡-Reasoning |
1070 | 1464 zc44 : FClosure A f (supf1 u) w |
1465 zc44 = subst (λ k → FClosure A f k w ) (sym zc45) fc | |
1028 | 1466 |
1079 | 1467 zo≤sz : {z : Ordinal} → z o≤ x → z o≤ supf1 z |
1080 | 1468 zo≤sz {z} z≤x with osuc-≡< z≤x |
1469 ... | case2 z<x = subst (λ k → z o≤ k) (sym (trans (sf1=sf z<x) (sym (zeq _ _ (o<→≤ <-osuc) o≤-refl)))) ( ZChain.zo≤sz (pzc z<x) o≤-refl ) | |
1083 | 1470 ... | case1 refl with x<y∨y≤x (supf1 spu) x |
1471 ... | case2 x≤ssp = z40 where | |
1472 z40 : z o≤ supf1 z | |
1473 z40 with x<y∨y≤x z spu | |
1474 ... | case1 z<spu = o<→≤ ( subst (λ k → z o< k ) (sym (sf1=spu refl)) z<spu ) | |
1475 ... | case2 spu≤z = begin -- x ≡ supf1 spu ≡ spu ≡ supf1 x | |
1476 x ≤⟨ x≤ssp ⟩ | |
1477 supf1 spu ≤⟨ supf-mono spu≤z ⟩ | |
1478 supf1 x ∎ where open o≤-Reasoning O | |
1479 ... | case1 ssp<x = subst (λ k → x o≤ k) (sym (sf1=spu refl)) z47 where | |
1080 | 1480 z47 : x o≤ spu |
1481 z47 with x<y∨y≤x spu x | |
1482 ... | case2 lt = lt | |
1483 ... | case1 spu<x = ⊥-elim ( <<-irr (MinSUP.x≤sup usup z48) (proj1 ( mf< spu (MinSUP.as usup)))) where | |
1083 | 1484 z70 : odef (UnionCF A f ay supf1 z) (supf1 spu) |
1485 z70 = cfcs spu<x o≤-refl ssp<x (init asupf refl ) | |
1486 z73 : IsSUP A (UnionCF A f ay (ZChain.supf (pzc (ob<x lim spu<x))) spu) spu | |
1091 | 1487 z73 = record { ax = MinSUP.as usup ; x≤sup = λ uw → MinSUP.x≤sup usup (chain⊆pchainU spu<x uw ) } |
1080 | 1488 z49 : supfz spu<x ≡ spu |
1083 | 1489 z49 = begin |
1490 supfz spu<x ≡⟨ ZChain.sup=u (pzc (ob<x lim spu<x)) (MinSUP.as usup) (o<→≤ <-osuc) z73 ⟩ | |
1491 spu ∎ where open ≡-Reasoning | |
1082 | 1492 z50 : supfz spu<x o≤ spu |
1493 z50 = o≤-refl0 z49 | |
1080 | 1494 z48 : odef pchainU (f spu) |
1082 | 1495 z48 = ⟪ proj2 (mf _ (MinSUP.as usup) ) , ic-isup _ (subst (λ k → k o< x) refl spu<x) z50 |
1080 | 1496 (fsuc _ (init (ZChain.asupf (pzc (ob<x lim spu<x))) z49)) ⟫ |
1079 | 1497 |
1077 | 1498 |
921 | 1499 --- |
1500 --- the maximum chain has fix point of any ≤-monotonic function | |
1501 --- | |
1502 | |
1039 | 1503 SZ : ( f : Ordinal → Ordinal ) → (mf< : <-monotonic-f A f ) → {y : Ordinal} (ay : odef A y) → (x : Ordinal) → ZChain A f mf< ay x |
1504 SZ f mf< {y} ay x = TransFinite {λ z → ZChain A f mf< ay z } (λ x → ind f mf< ay x ) x | |
921 | 1505 |
1039 | 1506 msp0 : ( f : Ordinal → Ordinal ) → (mf< : <-monotonic-f A f ) {x y : Ordinal} (ay : odef A y) |
1507 → (zc : ZChain A f mf< ay x ) | |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1508 → MinSUP A (UnionCF A f ay (ZChain.supf zc) x) |
1039 | 1509 msp0 f mf< {x} ay zc = minsupP (UnionCF A f ay (ZChain.supf zc) x) (ZChain.chain⊆A zc) (ZChain.f-total zc) |
922 | 1510 |
1090 | 1511 -- f eventualy stop |
1512 -- we can prove contradict here, it is here for a historical reason | |
1513 -- | |
1039 | 1514 fixpoint : ( f : Ordinal → Ordinal ) → (mf : ≤-monotonic-f A f ) (mf< : <-monotonic-f A f ) (zc : ZChain A f mf< as0 (& A) ) |
966 | 1515 → (sp1 : MinSUP A (ZChain.chain zc)) |
959 | 1516 → f (MinSUP.sup sp1) ≡ MinSUP.sup sp1 |
992 | 1517 fixpoint f mf mf< zc sp1 = z14 where |
924
a48dc906796c
supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
923
diff
changeset
|
1518 chain = ZChain.chain zc |
a48dc906796c
supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
923
diff
changeset
|
1519 supf = ZChain.supf zc |
935
ed711d7be191
mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
934
diff
changeset
|
1520 sp : Ordinal |
959 | 1521 sp = MinSUP.sup sp1 |
935
ed711d7be191
mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
934
diff
changeset
|
1522 asp : odef A sp |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
1523 asp = MinSUP.as sp1 |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1524 ay = as0 |
988
9a85233384f7
is-max and supf b = supf x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
987
diff
changeset
|
1525 z10 : {a b : Ordinal } → (ca : odef chain a ) → b o< (& A) → (ab : odef A b ) |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1526 → HasPrev A chain f b ∨ IsSUP A (UnionCF A f ay (ZChain.supf zc) b) b |
921 | 1527 → * a < * b → odef chain b |
993 | 1528 z10 = ZChain1.is-max (SZ1 f mf mf< as0 zc (& A) o≤-refl ) |
935
ed711d7be191
mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
934
diff
changeset
|
1529 z22 : sp o< & A |
1097 | 1530 z22 = odef< asp |
935
ed711d7be191
mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
934
diff
changeset
|
1531 z12 : odef chain sp |
ed711d7be191
mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
934
diff
changeset
|
1532 z12 with o≡? (& s) sp |
1058 | 1533 ... | yes eq = subst (λ k → odef chain k) eq ( ZChain.chain∋init zc ) |
1097 | 1534 ... | no ne = ZChain1.is-max (SZ1 f mf mf< as0 zc (& A) o≤-refl) {& s} {sp} ( ZChain.chain∋init zc ) (odef< asp) asp (case2 z19 ) z13 where |
935
ed711d7be191
mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
934
diff
changeset
|
1535 z13 : * (& s) < * sp |
1058 | 1536 z13 with MinSUP.x≤sup sp1 ( ZChain.chain∋init zc ) |
960 | 1537 ... | case1 eq = ⊥-elim ( ne eq ) |
966 | 1538 ... | case2 lt = lt |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1539 z19 : IsSUP A (UnionCF A f ay (ZChain.supf zc) sp) sp |
1069 | 1540 z19 = record { ax = asp ; x≤sup = z20 } where |
1033
2da8dcbb0825
ch-init again, because ch-is-sup require u<x which is not valid supf o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1032
diff
changeset
|
1541 z20 : {y : Ordinal} → odef (UnionCF A f ay (ZChain.supf zc) sp) y → (y ≡ sp) ∨ (y << sp) |
966 | 1542 z20 {y} zy with MinSUP.x≤sup sp1 |
961 | 1543 (subst (λ k → odef chain k ) (sym &iso) (chain-mono f mf as0 supf (ZChain.supf-mono zc) (o<→≤ z22) zy )) |
966 | 1544 ... | case1 y=p = case1 (subst (λ k → k ≡ _ ) &iso y=p ) |
960 | 1545 ... | case2 y<p = case2 (subst (λ k → * k < _ ) &iso y<p ) |
935
ed711d7be191
mem exhaust fix on fixpoint
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
934
diff
changeset
|
1546 z14 : f sp ≡ sp |
960 | 1547 z14 with ZChain.f-total zc (subst (λ k → odef chain k) (sym &iso) (ZChain.f-next zc z12 )) (subst (λ k → odef chain k) (sym &iso) z12 ) |
924
a48dc906796c
supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
923
diff
changeset
|
1548 ... | tri< a ¬b ¬c = ⊥-elim z16 where |
a48dc906796c
supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
923
diff
changeset
|
1549 z16 : ⊥ |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
1550 z16 with proj1 (mf (( MinSUP.sup sp1)) ( MinSUP.as sp1 )) |
1031 | 1551 ... | case1 eq = ⊥-elim (¬b (sym (cong (*) eq ) )) |
966 | 1552 ... | case2 lt = ⊥-elim (¬c lt ) |
1553 ... | tri≈ ¬a b ¬c = subst₂ (λ j k → j ≡ k ) &iso &iso ( cong (&) b ) | |
924
a48dc906796c
supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
923
diff
changeset
|
1554 ... | tri> ¬a ¬b c = ⊥-elim z17 where |
959 | 1555 z15 : (f sp ≡ MinSUP.sup sp1) ∨ (* (f sp) < * (MinSUP.sup sp1) ) |
960 | 1556 z15 = MinSUP.x≤sup sp1 (ZChain.f-next zc z12 ) |
924
a48dc906796c
supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
923
diff
changeset
|
1557 z17 : ⊥ |
a48dc906796c
supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
923
diff
changeset
|
1558 z17 with z15 |
960 | 1559 ... | case1 eq = ¬b (cong (*) eq) |
1560 ... | case2 lt = ¬a lt | |
924
a48dc906796c
supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
923
diff
changeset
|
1561 |
921 | 1562 -- ZChain contradicts ¬ Maximal |
1563 -- | |
1564 -- ZChain forces fix point on any ≤-monotonic function (fixpoint) | |
1565 -- ¬ Maximal create cf which is a <-monotonic function by axiom of choice. This contradicts fix point of ZChain | |
1566 -- | |
924
a48dc906796c
supf usp0 instead of supf (& A) ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
923
diff
changeset
|
1567 |
1090 | 1568 ¬Maximal→¬cf-mono : (nmx : ¬ Maximal A ) → (zc : ZChain A (cf nmx) (cf-is-<-monotonic nmx) as0 (& A)) → ⊥ |
1569 ¬Maximal→¬cf-mono nmx zc = <-irr0 {* (cf nmx c)} {* c} | |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
1570 (subst (λ k → odef A k ) (sym &iso) (proj1 (is-cf nmx (MinSUP.as msp1 )))) |
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
1571 (subst (λ k → odef A k) (sym &iso) (MinSUP.as msp1) ) |
992 | 1572 (case1 ( cong (*)( fixpoint (cf nmx) (cf-is-≤-monotonic nmx ) (cf-is-<-monotonic nmx ) zc msp1 ))) -- x ≡ f x ̄ |
1032
382680c3e9be
minsup is not obvious in ZChain
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1031
diff
changeset
|
1573 (proj1 (cf-is-<-monotonic nmx c (MinSUP.as msp1 ))) where -- x < f x |
937 | 1574 |
927 | 1575 supf = ZChain.supf zc |
934 | 1576 msp1 : MinSUP A (ZChain.chain zc) |
1039 | 1577 msp1 = msp0 (cf nmx) (cf-is-<-monotonic nmx) as0 zc |
966 | 1578 c : Ordinal |
1579 c = MinSUP.sup msp1 | |
934 | 1580 |
966 | 1581 zorn00 : Maximal A |
1090 | 1582 zorn00 with is-o∅ ( & HasMaximal ) |
804 | 1583 ... | no not = record { maximal = ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ; as = zorn01 ; ¬maximal<x = zorn02 } where |
1090 | 1584 -- yes we have the maximal because of the axiom of choice |
551 | 1585 zorn03 : odef HasMaximal ( & ( ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) ) ) |
606 | 1586 zorn03 = ODC.x∋minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) -- Axiom of choice |
551 | 1587 zorn01 : A ∋ ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) |
966 | 1588 zorn01 = proj1 zorn03 |
551 | 1589 zorn02 : {x : HOD} → A ∋ x → ¬ (ODC.minimal O HasMaximal (λ eq → not (=od∅→≡o∅ eq)) < x) |
1590 zorn02 {x} ax m<x = proj2 zorn03 (& x) ax (subst₂ (λ j k → j < k) (sym *iso) (sym *iso) m<x ) | |
1090 | 1591 ... | yes ¬Maximal = ⊥-elim ( ¬Maximal→¬cf-mono nmx (SZ (cf nmx) (cf-is-<-monotonic nmx) as0 (& A) )) where |
551 | 1592 -- if we have no maximal, make ZChain, which contradict SUP condition |
966 | 1593 nmx : ¬ Maximal A |
551 | 1594 nmx mx = ∅< {HasMaximal} zc5 ( ≡o∅→=od∅ ¬Maximal ) where |
966 | 1595 zc5 : odef A (& (Maximal.maximal mx)) ∧ (( y : Ordinal ) → odef A y → ¬ (* (& (Maximal.maximal mx)) < * y)) |
804 | 1596 zc5 = ⟪ Maximal.as mx , (λ y ay mx<y → Maximal.¬maximal<x mx (subst (λ k → odef A k ) (sym &iso) ay) (subst (λ k → k < * y) *iso mx<y) ) ⟫ |
551 | 1597 |
516 | 1598 -- usage (see filter.agda ) |
1599 -- | |
1091 | 1600 -- import OD hiding ( _⊆_ ) |
1601 -- _⊆_ : ( A B : HOD ) → Set n | |
1602 -- _⊆_ A B = {x : Ordinal } → odef A x → odef B x | |
1603 -- | |
1604 -- import zorn | |
1605 -- open zorn O _⊆_ -- Zorn on Set inclusion order | |
1606 -- | |
1607 -- open import Relation.Binary.Structures | |
482 | 1608 |
966 | 1609 -- MaximumSubset : {L P : HOD} |
497 | 1610 -- → o∅ o< & L → o∅ o< & P → P ⊆ L |
1091 | 1611 -- → IsPartialOrderSet P _⊆_ |
1612 -- → ( (B : HOD) → B ⊆ P → IsTotalOrderSet B _⊆_ → SUP P B _⊆_ ) | |
1613 -- → Maximal P (_⊆_) | |
1614 -- MaximumSubset {L} {P} 0<L 0<P P⊆L PO SP = Zorn-lemma {P} {_⊆_} 0<P PO SP | |
1615 -- |