comparison src/zorn.agda @ 600:71a1ed72cd21

not yet ...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Tue, 14 Jun 2022 06:17:24 +0900
parents d041941a8866
children 8b2a4af84e25
comparison
equal deleted inserted replaced
599:d041941a8866 600:71a1ed72cd21
87 -- 87 --
88 88
89 ≤-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set (Level.suc n) 89 ≤-monotonic-f : (A : HOD) → ( Ordinal → Ordinal ) → Set (Level.suc n)
90 ≤-monotonic-f A f = (x : Ordinal ) → odef A x → ( * x ≤ * (f x) ) ∧ odef A (f x ) 90 ≤-monotonic-f A f = (x : Ordinal ) → odef A x → ( * x ≤ * (f x) ) ∧ odef A (f x )
91 91
92 -- immieate-f : (A : HOD) → ( f : Ordinal → Ordinal ) → Set n
93 -- immieate-f A f = { x y : Ordinal } → odef A x → odef A y → ¬ ( ( * x < * y ) ∧ ( * y < * (f x )) )
94
95 data FClosure (A : HOD) (f : Ordinal → Ordinal ) (s : Ordinal) : Ordinal → Set n where 92 data FClosure (A : HOD) (f : Ordinal → Ordinal ) (s : Ordinal) : Ordinal → Set n where
96 init : {x : Ordinal} → odef A s → x ≡ s → FClosure A f s x 93 init : odef A s → FClosure A f s s
97 fsuc : (x : Ordinal) ( p : FClosure A f s x ) → FClosure A f s (f x) 94 fsuc : (x : Ordinal) ( p : FClosure A f s x ) → FClosure A f s (f x)
98 95
99 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 96 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
100 A∋fc {A} s f mf (init as refl ) = as 97 A∋fc {A} s f mf (init as) = as
101 A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {A} s f mf fcy ) ) 98 A∋fc {A} s f mf (fsuc y fcy) = proj2 (mf y ( A∋fc {A} s f mf fcy ) )
102 99
103 s≤fc : {A : HOD} (s : Ordinal ) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → * s ≤ * y 100 s≤fc : {A : HOD} (s : Ordinal ) {y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) → (fcy : FClosure A f s y ) → * s ≤ * y
104 s≤fc {A} s {.s} f mf (init x refl ) = case1 refl 101 s≤fc {A} s {.s} f mf (init x) = case1 refl
105 s≤fc {A} s {.(f x)} f mf (fsuc x fcy) with proj1 (mf x (A∋fc s f mf fcy ) ) 102 s≤fc {A} s {.(f x)} f mf (fsuc x fcy) with proj1 (mf x (A∋fc s f mf fcy ) )
106 ... | case1 x=fx = subst (λ k → * s ≤ * k ) (*≡*→≡ x=fx) ( s≤fc {A} s f mf fcy ) 103 ... | case1 x=fx = subst (λ k → * s ≤ * k ) (*≡*→≡ x=fx) ( s≤fc {A} s f mf fcy )
107 ... | case2 x<fx with s≤fc {A} s f mf fcy 104 ... | case2 x<fx with s≤fc {A} s f mf fcy
108 ... | case1 s≡x = case2 ( subst₂ (λ j k → j < k ) (sym s≡x) refl x<fx ) 105 ... | case1 s≡x = case2 ( subst₂ (λ j k → j < k ) (sym s≡x) refl x<fx )
109 ... | case2 s<x = case2 ( IsStrictPartialOrder.trans PO s<x x<fx ) 106 ... | case2 s<x = case2 ( IsStrictPartialOrder.trans PO s<x x<fx )
110 107
111 fcn : {A : HOD} (s : Ordinal) { x : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) → FClosure A f s x → ℕ 108 fcn : {A : HOD} (s : Ordinal) { x : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) → FClosure A f s x → ℕ
112 fcn s mf (init as refl ) = zero 109 fcn s mf (init as) = zero
113 fcn {A} s {x} {f} mf (fsuc y p) with proj1 (mf y (A∋fc s f mf p)) 110 fcn {A} s {x} {f} mf (fsuc y p) with proj1 (mf y (A∋fc s f mf p))
114 ... | case1 eq = fcn s mf p 111 ... | case1 eq = fcn s mf p
115 ... | case2 y<fy = suc (fcn s mf p ) 112 ... | case2 y<fy = suc (fcn s mf p )
116 113
117 fcn-inject : {A : HOD} (s : Ordinal) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) 114 fcn-inject : {A : HOD} (s : Ordinal) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f)
118 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx ≡ fcn s mf cy → * x ≡ * y 115 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx ≡ fcn s mf cy → * x ≡ * y
119 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 116 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
120 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 117 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
121 fc00 zero zero refl (init _ refl ) (init x₁ refl ) i=x i=y = refl 118 fc00 zero zero refl (init _) (init x₁) i=x i=y = refl
122 fc00 zero zero refl (init as refl ) (fsuc y cy) i=x i=y with proj1 (mf y (A∋fc s f mf cy ) ) 119 fc00 zero zero refl (init as) (fsuc y cy) i=x i=y with proj1 (mf y (A∋fc s f mf cy ) )
123 ... | case1 y=fy = subst (λ k → * s ≡ k ) y=fy ( fc00 zero zero refl (init as refl ) cy i=x i=y ) 120 ... | case1 y=fy = subst (λ k → * s ≡ k ) y=fy ( fc00 zero zero refl (init as) cy i=x i=y )
124 fc00 zero zero refl (fsuc x cx) (init as refl ) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) ) 121 fc00 zero zero refl (fsuc x cx) (init as) i=x i=y with proj1 (mf x (A∋fc s f mf cx ) )
125 ... | case1 x=fx = subst (λ k → k ≡ * s ) x=fx ( fc00 zero zero refl cx (init as refl ) i=x i=y ) 122 ... | case1 x=fx = subst (λ k → k ≡ * s ) x=fx ( fc00 zero zero refl cx (init as) i=x i=y )
126 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 ) ) 123 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 ) )
127 ... | 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 ) 124 ... | 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 )
128 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 ) ) 125 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 ) )
129 ... | 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 ) 126 ... | 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 )
130 ... | case1 x=fx | case2 y<fy = subst (λ k → k ≡ * (f y)) x=fx (fc02 x cx i=x) where 127 ... | case1 x=fx | case2 y<fy = subst (λ k → k ≡ * (f y)) x=fx (fc02 x cx i=x) where
141 ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc05) where 138 ... | case2 lt = subst₂ (λ j k → * (f j) ≡ * (f k )) &iso &iso ( cong (λ k → * ( f (& k ))) fc05) where
142 fc05 : * x ≡ * y1 139 fc05 : * x ≡ * y1
143 fc05 = fc00 i j (cong pred i=j) cx cy1 (cong pred i=x) (cong pred j=y1) 140 fc05 = fc00 i j (cong pred i=j) cx cy1 (cong pred i=x) (cong pred j=y1)
144 ... | 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))) 141 ... | 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)))
145 142
143
146 fcn-< : {A : HOD} (s : Ordinal ) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f) 144 fcn-< : {A : HOD} (s : Ordinal ) { x y : Ordinal} {f : Ordinal → Ordinal} → (mf : ≤-monotonic-f A f)
147 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx Data.Nat.< fcn s mf cy → * x < * y 145 → (cx : FClosure A f s x ) (cy : FClosure A f s y ) → fcn s mf cx Data.Nat.< fcn s mf cy → * x < * y
148 fcn-< {A} s {x} {y} {f} mf cx cy x<y = fc01 (fcn s mf cy) cx cy refl x<y where 146 fcn-< {A} s {x} {y} {f} mf cx cy x<y = fc01 (fcn s mf cy) cx cy refl x<y where
149 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 147 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
150 fc01 (suc i) {y} cx (fsuc y1 cy) i=y (s≤s x<i) with proj1 (mf y1 (A∋fc s f mf cy ) ) 148 fc01 (suc i) {y} cx (fsuc y1 cy) i=y (s≤s x<i) with proj1 (mf y1 (A∋fc s f mf cy ) )
168 fc10 : * x ≡ * y 166 fc10 : * x ≡ * y
169 fc10 = fcn-inject {A} s {x} {y} {f} mf cx cy b 167 fc10 = fcn-inject {A} s {x} {y} {f} mf cx cy b
170 ... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where 168 ... | tri> ¬a ¬b c = tri> (λ lt → <-irr (case2 fc12) lt) (λ eq → <-irr (case1 eq) fc12) fc12 where
171 fc12 : * y < * x 169 fc12 : * y < * x
172 fc12 = fcn-< {A} s {y} {x} {f} mf cy cx c 170 fc12 = fcn-< {A} s {y} {x} {f} mf cy cx c
171
173 172
174 fcn-imm : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f) 173 fcn-imm : {A : HOD} (s : Ordinal) { x y : Ordinal } (f : Ordinal → Ordinal) (mf : ≤-monotonic-f A f)
175 → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → ¬ ( ( * x < * y ) ∧ ( * y < * (f x )) ) 174 → (cx : FClosure A f s x) → (cy : FClosure A f s y ) → ¬ ( ( * x < * y ) ∧ ( * y < * (f x )) )
176 fcn-imm {A} s {x} {y} f mf cx cy ⟪ x<y , y<fx ⟫ = fc21 where 175 fcn-imm {A} s {x} {y} f mf cx cy ⟪ x<y , y<fx ⟫ = fc21 where
177 fc20 : fcn s mf cy Data.Nat.< suc (fcn s mf cx) → (fcn s mf cy ≡ fcn s mf cx) ∨ ( fcn s mf cy Data.Nat.< fcn s mf cx ) 176 fc20 : fcn s mf cy Data.Nat.< suc (fcn s mf cx) → (fcn s mf cy ≡ fcn s mf cx) ∨ ( fcn s mf cy Data.Nat.< fcn s mf cx )
191 fc21 with <-cmp (suc ( fcn s mf cx )) (fcn s mf cy ) 190 fc21 with <-cmp (suc ( fcn s mf cx )) (fcn s mf cy )
192 ... | tri< a ¬b ¬c = <-irr (case2 y<fx) (fc22 a) where -- suc ncx < ncy 191 ... | tri< a ¬b ¬c = <-irr (case2 y<fx) (fc22 a) where -- suc ncx < ncy
193 cxx : FClosure A f s (f x) 192 cxx : FClosure A f s (f x)
194 cxx = fsuc x cx 193 cxx = fsuc x cx
195 fc16 : (x : Ordinal ) → (cx : FClosure A f s x) → (fcn s mf cx ≡ fcn s mf (fsuc x cx)) ∨ ( suc (fcn s mf cx ) ≡ fcn s mf (fsuc x cx)) 194 fc16 : (x : Ordinal ) → (cx : FClosure A f s x) → (fcn s mf cx ≡ fcn s mf (fsuc x cx)) ∨ ( suc (fcn s mf cx ) ≡ fcn s mf (fsuc x cx))
196 fc16 x (init as refl ) with proj1 (mf s as ) 195 fc16 x (init as) with proj1 (mf s as )
197 ... | case1 _ = case1 refl 196 ... | case1 _ = case1 refl
198 ... | case2 _ = case2 refl 197 ... | case2 _ = case2 refl
199 fc16 .(f x) (fsuc x cx ) with proj1 (mf (f x) (A∋fc s f mf (fsuc x cx)) ) 198 fc16 .(f x) (fsuc x cx ) with proj1 (mf (f x) (A∋fc s f mf (fsuc x cx)) )
200 ... | case1 _ = case1 refl 199 ... | case1 _ = case1 refl
201 ... | case2 _ = case2 refl 200 ... | case2 _ = case2 refl
206 ... | tri≈ ¬a b ¬c = <-irr (case1 (fc17 cx cy b)) y<fx 205 ... | tri≈ ¬a b ¬c = <-irr (case1 (fc17 cx cy b)) y<fx
207 ... | tri> ¬a ¬b c with fc20 c -- ncy < suc ncx 206 ... | tri> ¬a ¬b c with fc20 c -- ncy < suc ncx
208 ... | case1 y=x = <-irr (case1 ( fcn-inject s mf cy cx y=x )) x<y 207 ... | case1 y=x = <-irr (case1 ( fcn-inject s mf cy cx y=x )) x<y
209 ... | case2 y<x = <-irr (case2 x<y) (fcn-< s mf cy cx y<x ) 208 ... | case2 y<x = <-irr (case2 x<y) (fcn-< s mf cy cx y<x )
210 209
211
212 -- open import Relation.Binary.Properties.Poset as Poset 210 -- open import Relation.Binary.Properties.Poset as Poset
213 211
214 IsTotalOrderSet : ( A : HOD ) → Set (Level.suc n) 212 IsTotalOrderSet : ( A : HOD ) → Set (Level.suc n)
215 IsTotalOrderSet A = {a b : HOD} → odef A (& a) → odef A (& b) → Tri (a < b) (a ≡ b) (b < a ) 213 IsTotalOrderSet A = {a b : HOD} → odef A (& a) → odef A (& b) → Tri (a < b) (a ≡ b) (b < a )
216 214
233 231
234 record IsSup (A B : HOD) {x : Ordinal } (xa : odef A x) : Set n where 232 record IsSup (A B : HOD) {x : Ordinal } (xa : odef A x) : Set n where
235 field 233 field
236 x<sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x ) 234 x<sup : {y : Ordinal} → odef B y → (y ≡ x ) ∨ (y << x )
237 235
238 y1 : (A : HOD) → (y : Ordinal) → odef A y → * (& (* y , * y)) ⊆' A
239 y1 A y ay {x} lt with subst (λ k → odef k x) *iso lt
240 ... | case1 eq = subst (λ k → odef A k ) (sym (trans eq &iso)) ay
241 ... | case2 eq = subst (λ k → odef A k ) (sym (trans eq &iso)) ay
242
243 record FChain ( A : HOD ) ( f : Ordinal → Ordinal ) (p c : Ordinal) ( x : Ordinal ) : Set n where 236 record FChain ( A : HOD ) ( f : Ordinal → Ordinal ) (p c : Ordinal) ( x : Ordinal ) : Set n where
244 field 237 field
245 fc∨sup : FClosure A f p x 238 fc∨sup : FClosure A f p x
246 chain∋p : odef (* c) p 239 chain∋p : odef (* c) p
247 240
241 record FSup ( A : HOD ) ( f : Ordinal → Ordinal ) (p c : Ordinal) ( x : Ordinal ) : Set n where
242 field
243 sup : (z : Ordinal) → FClosure A f p z → * z < * x
244 min : ( x1 : Ordinal) → ((z : Ordinal) → FClosure A f p z → * z < * x1 ) → ( x ≡ x1 ) ∨ ( * x < * x1 )
245 chain∋x : odef (* c) x
246 chain∋p : odef (* c) p
247
248 data Fc∨sup (A : HOD) {y : Ordinal} (ay : odef A y) ( f : Ordinal → Ordinal ) (c : Ordinal) : (x : Ordinal) → Set n where 248 data Fc∨sup (A : HOD) {y : Ordinal} (ay : odef A y) ( f : Ordinal → Ordinal ) (c : Ordinal) : (x : Ordinal) → Set n where
249 Finit : {z : Ordinal} → z ≡ y → Fc∨sup A ay f c z 249 Finit : {z : Ordinal} → z ≡ y → Fc∨sup A ay f c z
250 Fc : {p x : Ordinal} → p o< x → Fc∨sup A ay f c p → FChain A f p c x → Fc∨sup A ay f c x 250 Fsup : {p x : Ordinal} → p o< x → Fc∨sup A ay f c p → FSup A f p c x → Fc∨sup A ay f c x
251 Fc : {p x : Ordinal} → p o< x → Fc∨sup A ay f c p → FChain A f p c x → Fc∨sup A ay f c x
251 252
252 record ZChain ( A : HOD ) (x : Ordinal) ( f : Ordinal → Ordinal ) ( z : Ordinal ) : Set (Level.suc n) where 253 record ZChain ( A : HOD ) (x : Ordinal) ( f : Ordinal → Ordinal ) ( z : Ordinal ) : Set (Level.suc n) where
253 field 254 field
254 chain : HOD 255 chain : HOD
255 chain⊆A : chain ⊆' A 256 chain⊆A : chain ⊆' A
421 zc4 : ZChain A y f x 422 zc4 : ZChain A y f x
422 zc4 with ODC.∋-p O A (* x) 423 zc4 with ODC.∋-p O A (* x)
423 ... | no noax = -- ¬ A ∋ p, just skip 424 ... | no noax = -- ¬ A ∋ p, just skip
424 record { chain = ZChain.chain zc0 ; chain⊆A = ZChain.chain⊆A zc0 ; initial = ZChain.initial zc0 425 record { chain = ZChain.chain zc0 ; chain⊆A = ZChain.chain⊆A zc0 ; initial = ZChain.initial zc0
425 ; f-total = ZChain.f-total zc0 ; f-next = ZChain.f-next zc0 426 ; f-total = ZChain.f-total zc0 ; f-next = ZChain.f-next zc0
426 ; f-immediate = ZChain.f-immediate zc0 ; chain∋x = ZChain.chain∋x zc0 ; is-max = zc11 ; fc∨sup = {!!} } where -- no extention 427 ; f-immediate = ZChain.f-immediate zc0 ; chain∋x = ZChain.chain∋x zc0 ; is-max = zc11 ; fc∨sup = ZChain.fc∨sup zc0 } where -- no extention
427 zc11 : {a b : Ordinal} → odef (ZChain.chain zc0) a → b o< osuc x → (ab : odef A b) → 428 zc11 : {a b : Ordinal} → odef (ZChain.chain zc0) a → b o< osuc x → (ab : odef A b) →
428 HasPrev A (ZChain.chain zc0) ab f ∨ IsSup A (ZChain.chain zc0) ab → 429 HasPrev A (ZChain.chain zc0) ab f ∨ IsSup A (ZChain.chain zc0) ab →
429 * a < * b → odef (ZChain.chain zc0) b 430 * a < * b → odef (ZChain.chain zc0) b
430 zc11 {a} {b} za b<ox ab P a<b with osuc-≡< b<ox 431 zc11 {a} {b} za b<ox ab P a<b with osuc-≡< b<ox
431 ... | case1 eq = ⊥-elim ( noax (subst (λ k → odef A k) (trans eq (sym &iso)) ab ) ) 432 ... | case1 eq = ⊥-elim ( noax (subst (λ k → odef A k) (trans eq (sym &iso)) ab ) )
439 zc17 {a} {b} za b<ox ab P a<b with osuc-≡< b<ox 440 zc17 {a} {b} za b<ox ab P a<b with osuc-≡< b<ox
440 ... | case2 lt = ZChain.is-max zc0 za (zc0-b<x b lt) ab P a<b 441 ... | case2 lt = ZChain.is-max zc0 za (zc0-b<x b lt) ab P a<b
441 ... | case1 b=x = subst (λ k → odef chain k ) (trans (sym (HasPrev.x=fy pr )) (trans &iso (sym b=x)) ) ( ZChain.f-next zc0 (HasPrev.ay pr)) 442 ... | case1 b=x = subst (λ k → odef chain k ) (trans (sym (HasPrev.x=fy pr )) (trans &iso (sym b=x)) ) ( ZChain.f-next zc0 (HasPrev.ay pr))
442 zc9 : ZChain A y f x 443 zc9 : ZChain A y f x
443 zc9 = record { chain = ZChain.chain zc0 ; chain⊆A = ZChain.chain⊆A zc0 ; f-total = ZChain.f-total zc0 ; f-next = ZChain.f-next zc0 -- no extention 444 zc9 = record { chain = ZChain.chain zc0 ; chain⊆A = ZChain.chain⊆A zc0 ; f-total = ZChain.f-total zc0 ; f-next = ZChain.f-next zc0 -- no extention
444 ; initial = ZChain.initial zc0 ; f-immediate = ZChain.f-immediate zc0 ; chain∋x = ZChain.chain∋x zc0 ; is-max = zc17 ; fc∨sup = {!!}} 445 ; initial = ZChain.initial zc0 ; f-immediate = ZChain.f-immediate zc0 ; chain∋x = ZChain.chain∋x zc0 ; is-max = zc17 ; fc∨sup = ZChain.fc∨sup zc0 }
445 ... | case2 ¬fy<x with ODC.p∨¬p O (IsSup A (ZChain.chain zc0) ax ) 446 ... | case2 ¬fy<x with ODC.p∨¬p O (IsSup A (ZChain.chain zc0) ax )
446 ... | case1 is-sup = -- x is a sup of zc0 447 ... | case1 is-sup = -- x is a sup of zc0
447 record { chain = schain ; chain⊆A = s⊆A ; f-total = scmp ; f-next = scnext 448 record { chain = schain ; chain⊆A = s⊆A ; f-total = scmp ; f-next = scnext
448 ; initial = scinit ; f-immediate = simm ; chain∋x = case1 (ZChain.chain∋x zc0) ; is-max = s-ismax ; fc∨sup = {!!}} where 449 ; initial = scinit ; f-immediate = simm ; chain∋x = case1 (ZChain.chain∋x zc0) ; is-max = s-ismax ; fc∨sup = s-fc∨sup} where
449 sup0 : SUP A (ZChain.chain zc0) 450 sup0 : SUP A (ZChain.chain zc0)
450 sup0 = record { sup = * x ; A∋maximal = ax ; x<sup = x21 } where 451 sup0 = record { sup = * x ; A∋maximal = ax ; x<sup = x21 } where
451 x21 : {y : HOD} → ZChain.chain zc0 ∋ y → (y ≡ * x) ∨ (y < * x) 452 x21 : {y : HOD} → ZChain.chain zc0 ∋ y → (y ≡ * x) ∨ (y < * x)
452 x21 {y} zy with IsSup.x<sup is-sup zy 453 x21 {y} zy with IsSup.x<sup is-sup zy
453 ... | case1 y=x = case1 ( subst₂ (λ j k → j ≡ * k ) *iso &iso ( cong (*) y=x) ) 454 ... | case1 y=x = case1 ( subst₂ (λ j k → j ≡ * k ) *iso &iso ( cong (*) y=x) )
466 A∋schain (case1 zx ) = ZChain.chain⊆A zc0 zx 467 A∋schain (case1 zx ) = ZChain.chain⊆A zc0 zx
467 A∋schain {y} (case2 fx ) = A∋fc (& sp) f mf fx 468 A∋schain {y} (case2 fx ) = A∋fc (& sp) f mf fx
468 s⊆A : schain ⊆' A 469 s⊆A : schain ⊆' A
469 s⊆A {x} (case1 zx) = ZChain.chain⊆A zc0 zx 470 s⊆A {x} (case1 zx) = ZChain.chain⊆A zc0 zx
470 s⊆A {x} (case2 fx) = A∋fc (& sp) f mf fx 471 s⊆A {x} (case2 fx) = A∋fc (& sp) f mf fx
472 s-fc∨sup : {c : Ordinal} → odef schain c → Fc∨sup A (s⊆A (case1 (ZChain.chain∋x zc0))) f (& schain) c
473 s-fc∨sup {c} (case1 cx) = {!!}
474 s-fc∨sup {c} (case2 fc) = {!!}
471 cmp : {a b : HOD} (za : odef chain (& a)) (fb : FClosure A f (& sp) (& b)) → Tri (a < b) (a ≡ b) (b < a ) 475 cmp : {a b : HOD} (za : odef chain (& a)) (fb : FClosure A f (& sp) (& b)) → Tri (a < b) (a ≡ b) (b < a )
472 cmp {a} {b} za fb with SUP.x<sup sup0 za | s≤fc (& sp) f mf fb 476 cmp {a} {b} za fb with SUP.x<sup sup0 za | s≤fc (& sp) f mf fb
473 ... | case1 sp=a | case1 sp=b = tri≈ (λ lt → <-irr (case1 (sym eq)) lt ) eq (λ lt → <-irr (case1 eq) lt ) where 477 ... | case1 sp=a | case1 sp=b = tri≈ (λ lt → <-irr (case1 (sym eq)) lt ) eq (λ lt → <-irr (case1 eq) lt ) where
474 eq : a ≡ b 478 eq : a ≡ b
475 eq = trans sp=a (subst₂ (λ j k → j ≡ k ) *iso *iso sp=b ) 479 eq = trans sp=a (subst₂ (λ j k → j ≡ k ) *iso *iso sp=b )
523 simm {a} {b} (case2 sa) (case2 sb) p = fcn-imm {A} (& sp) {a} {b} f mf sa sb p 527 simm {a} {b} (case2 sa) (case2 sb) p = fcn-imm {A} (& sp) {a} {b} f mf sa sb p
524 s-ismax : {a b : Ordinal} → odef schain a → b o< osuc x → (ab : odef A b) 528 s-ismax : {a b : Ordinal} → odef schain a → b o< osuc x → (ab : odef A b)
525 → HasPrev A schain ab f ∨ IsSup A schain ab 529 → HasPrev A schain ab f ∨ IsSup A schain ab
526 → * a < * b → odef schain b 530 → * a < * b → odef schain b
527 s-ismax {a} {b} sa b<ox ab p a<b with osuc-≡< b<ox -- b is x? 531 s-ismax {a} {b} sa b<ox ab p a<b with osuc-≡< b<ox -- b is x?
528 ... | case1 b=x = case2 (subst (λ k → FClosure A f (& sp) k ) (sym (trans b=x x=sup )) (init (SUP.A∋maximal sup0) refl )) 532 ... | case1 b=x = case2 (subst (λ k → FClosure A f (& sp) k ) (sym (trans b=x x=sup )) (init (SUP.A∋maximal sup0) ))
529 s-ismax {a} {b} (case1 za) b<ox ab (case1 p) a<b | case2 b<x = z21 p where -- has previous 533 s-ismax {a} {b} (case1 za) b<ox ab (case1 p) a<b | case2 b<x = z21 p where -- has previous
530 z21 : HasPrev A schain ab f → odef schain b 534 z21 : HasPrev A schain ab f → odef schain b
531 z21 record { y = y ; ay = (case1 zy) ; x=fy = x=fy } = 535 z21 record { y = y ; ay = (case1 zy) ; x=fy = x=fy } =
532 case1 (ZChain.is-max zc0 za (zc0-b<x b b<x) ab (case1 record { y = y ; ay = zy ; x=fy = x=fy }) a<b ) 536 case1 (ZChain.is-max zc0 za (zc0-b<x b b<x) ab (case1 record { y = y ; ay = zy ; x=fy = x=fy }) a<b )
533 z21 record { y = y ; ay = (case2 sy) ; x=fy = x=fy } = subst (λ k → odef schain k) (sym x=fy) (case2 (fsuc y sy) ) 537 z21 record { y = y ; ay = (case2 sy) ; x=fy = x=fy } = subst (λ k → odef schain k) (sym x=fy) (case2 (fsuc y sy) )
544 z23 with IsSup.x<sup (z24 p) ( ZChain.chain∋x zc0 ) 548 z23 with IsSup.x<sup (z24 p) ( ZChain.chain∋x zc0 )
545 ... | case1 y=b = subst (λ k → odef chain k ) y=b ( ZChain.chain∋x zc0 ) 549 ... | case1 y=b = subst (λ k → odef chain k ) y=b ( ZChain.chain∋x zc0 )
546 ... | case2 y<b = ZChain.is-max zc0 (ZChain.chain∋x zc0 ) (zc0-b<x b b<x) ab (case2 (z24 p)) y<b 550 ... | case2 y<b = ZChain.is-max zc0 (ZChain.chain∋x zc0 ) (zc0-b<x b b<x) ab (case2 (z24 p)) y<b
547 ... | case2 ¬x=sup = -- x is not f y' nor sup of former ZChain from y 551 ... | case2 ¬x=sup = -- x is not f y' nor sup of former ZChain from y
548 record { chain = ZChain.chain zc0 ; chain⊆A = ZChain.chain⊆A zc0 ; f-total = ZChain.f-total zc0 ; f-next = ZChain.f-next zc0 552 record { chain = ZChain.chain zc0 ; chain⊆A = ZChain.chain⊆A zc0 ; f-total = ZChain.f-total zc0 ; f-next = ZChain.f-next zc0
549 ; initial = ZChain.initial zc0 ; f-immediate = ZChain.f-immediate zc0 ; chain∋x = ZChain.chain∋x zc0 ; is-max = z18 ; fc∨sup = {!!} } where 553 ; initial = ZChain.initial zc0 ; f-immediate = ZChain.f-immediate zc0 ; chain∋x = ZChain.chain∋x zc0 ; is-max = z18 ; fc∨sup = ZChain.fc∨sup zc0 } where
550 -- no extention 554 -- no extention
551 z18 : {a b : Ordinal} → odef (ZChain.chain zc0) a → b o< osuc x → (ab : odef A b) → 555 z18 : {a b : Ordinal} → odef (ZChain.chain zc0) a → b o< osuc x → (ab : odef A b) →
552 HasPrev A (ZChain.chain zc0) ab f ∨ IsSup A (ZChain.chain zc0) ab → 556 HasPrev A (ZChain.chain zc0) ab f ∨ IsSup A (ZChain.chain zc0) ab →
553 * a < * b → odef (ZChain.chain zc0) b 557 * a < * b → odef (ZChain.chain zc0) b
554 z18 {a} {b} za b<x ab p a<b with osuc-≡< b<x 558 z18 {a} {b} za b<x ab p a<b with osuc-≡< b<x
559 x<sup = λ {y} zy → subst (λ k → (y ≡ k) ∨ (y << k)) (trans b=x (sym &iso)) (IsSup.x<sup b=sup zy) } ) 563 x<sup = λ {y} zy → subst (λ k → (y ≡ k) ∨ (y << k)) (trans b=x (sym &iso)) (IsSup.x<sup b=sup zy) } )
560 ... | no ¬ox with trio< x y 564 ... | no ¬ox with trio< x y
561 ... | tri< a ¬b ¬c = record { chain = {!!} ; chain⊆A = {!!} ; f-total = {!!} ; f-next = {!!} 565 ... | tri< a ¬b ¬c = record { chain = {!!} ; chain⊆A = {!!} ; f-total = {!!} ; f-next = {!!}
562 ; initial = {!!} ; f-immediate = {!!} ; chain∋x = {!!} ; is-max = {!!} ; fc∨sup = {!!} } 566 ; initial = {!!} ; f-immediate = {!!} ; chain∋x = {!!} ; is-max = {!!} ; fc∨sup = {!!} }
563 ... | tri≈ ¬a b ¬c = {!!} 567 ... | tri≈ ¬a b ¬c = {!!}
564 ... | tri> ¬a ¬b y<x = {!!} where 568 ... | tri> ¬a ¬b y<x = UnionZ where
565 UnionZ : ZChain A y f x 569 UnionZ : ZChain A y f x
566 UnionZ = record { chain = Uz ; chain⊆A = Uz⊆A ; f-total = u-total ; f-next = u-next 570 UnionZ = record { chain = Uz ; chain⊆A = Uz⊆A ; f-total = u-total ; f-next = u-next
567 ; initial = u-initial ; f-immediate = {!!} ; chain∋x = {!!} ; is-max = {!!} ; fc∨sup = {!!} } where --- limit ordinal case 571 ; initial = u-initial ; f-immediate = {!!} ; chain∋x = {!!} ; is-max = {!!} ; fc∨sup = {!!} } where --- limit ordinal case
568 record UZFChain (z : Ordinal) : Set n where -- Union of ZFChain from y which has maximality o< x 572 record UZFChain (z : Ordinal) : Set n where -- Union of ZFChain from y which has maximality o< x
569 field 573 field
593 FC : Fc∨sup A (chain⊆A za (chain∋x za)) f (& (chain za)) i 597 FC : Fc∨sup A (chain⊆A za (chain∋x za)) f (& (chain za)) i
594 FC = fc∨sup za zai 598 FC = fc∨sup za zai
595 um01 : odef (chain zb) i 599 um01 : odef (chain zb) i
596 um01 with FC 600 um01 with FC
597 ... | Finit i=y = subst (λ k → odef (chain zb) k ) (sym i=y) ( chain∋x zb ) 601 ... | Finit i=y = subst (λ k → odef (chain zb) k ) (sym i=y) ( chain∋x zb )
602 ... | Fsup {p} {i} p<i pFC sup = ?
598 ... | Fc {p} {i} p<i pFC FC with (FChain.fc∨sup FC) 603 ... | Fc {p} {i} p<i pFC FC with (FChain.fc∨sup FC)
599 ... | fc = um02 i fc where 604 ... | fc = um02 i fc where
600 um02 : (i2 : Ordinal) → FClosure A f p i2 → odef (chain zb) i2 605 um02 : (i2 : Ordinal) → FClosure A f p i2 → odef (chain zb) i2
601 um02 i2 (init ap i2=p ) = subst (λ k → odef (chain zb) k ) (sym i2=p) (previ p p<i um04 ) where 606 um02 i2 (init ap ) = previ p p<i um04 where
602 um04 : odef (chain za) p 607 um04 : odef (chain za) p
603 um04 = subst (λ k → odef k p ) *iso ( FChain.chain∋p FC ) 608 um04 = subst (λ k → odef k p ) *iso ( FChain.chain∋p FC )
604 um02 i (fsuc j fc) = f-next zb ( um02 j fc ) 609 um02 i (fsuc j fc) = f-next zb ( um02 j fc )
605 u-total : IsTotalOrderSet Uz 610 u-total : IsTotalOrderSet Uz
606 u-total {x} {y} ux uy with trio< (UZFChain.u ux) (UZFChain.u uy) 611 u-total {x} {y} ux uy with trio< (UZFChain.u ux) (UZFChain.u uy)