comparison src/zorn.agda @ 1092:87c2da3811c3

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