comparison OD.agda @ 276:6f10c47e4e7a

separate choice fix sup-o
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 09 May 2020 09:02:52 +0900
parents 29a85a427ed2
children d9d3654baee1
comparison
equal deleted inserted replaced
275:455792eaa611 276:6f10c47e4e7a
73 diso : {x : Ordinal } → od→ord ( ord→od x ) ≡ x 73 diso : {x : Ordinal } → od→ord ( ord→od x ) ≡ x
74 ==→o≡ : { x y : OD } → (x == y) → x ≡ y 74 ==→o≡ : { x y : OD } → (x == y) → x ≡ y
75 -- next assumption causes ∀ x ∋ ∅ . It menas only an ordinal is allowed as OD 75 -- next assumption causes ∀ x ∋ ∅ . It menas only an ordinal is allowed as OD
76 -- o<→c< : {n : Level} {x y : Ordinal } → x o< y → def (ord→od y) x 76 -- o<→c< : {n : Level} {x y : Ordinal } → x o< y → def (ord→od y) x
77 -- ord→od x ≡ Ord x results the same 77 -- ord→od x ≡ Ord x results the same
78 -- supermum as Replacement Axiom ( this assumes Ordinal has some upper bound ) 78 -- supermum as Replacement Axiom ( corresponding Ordinal of OD has maximum )
79 sup-o : ( Ordinal → Ordinal ) → Ordinal 79 sup-o : ( OD → Ordinal ) → Ordinal
80 sup-o< : { ψ : Ordinal → Ordinal } → ∀ {x : Ordinal } → ψ x o< sup-o ψ 80 sup-o< : { ψ : OD → Ordinal } → ∀ {x : OD } → ψ x o< sup-o ψ
81 -- contra-position of mimimulity of supermum required in Power Set Axiom 81 -- contra-position of mimimulity of supermum required in Power Set Axiom
82 -- sup-x : {n : Level } → ( Ordinal → Ordinal ) → Ordinal 82 -- sup-x : {n : Level } → ( Ordinal → Ordinal ) → Ordinal
83 -- sup-lb : {n : Level } → { ψ : Ordinal → Ordinal } → {z : Ordinal } → z o< sup-o ψ → z o< osuc (ψ (sup-x ψ)) 83 -- sup-lb : {n : Level } → { ψ : Ordinal → Ordinal } → {z : Ordinal } → z o< sup-o ψ → z o< osuc (ψ (sup-x ψ))
84 -- mimimul and x∋minimal is an Axiom of choice 84
85 minimal : (x : OD ) → ¬ (x == od∅ )→ OD 85 data One : Set n where
86 -- this should be ¬ (x == od∅ )→ ∃ ox → x ∋ Ord ox ( minimum of x ) 86 OneObj : One
87 x∋minimal : (x : OD ) → ( ne : ¬ (x == od∅ ) ) → def x ( od→ord ( minimal x ne ) ) 87
88 -- minimality (may proved by ε-induction ) 88 Ords : OD
89 minimal-1 : (x : OD ) → ( ne : ¬ (x == od∅ ) ) → (y : OD ) → ¬ ( def (minimal x ne) (od→ord y)) ∧ (def x (od→ord y) ) 89 Ords = record { def = λ x → One }
90
91 maxod : {x : OD} → od→ord x o< od→ord Ords
92 maxod {x} = c<→o< OneObj
90 93
91 o<→c<→OD=Ord : ( {x y : Ordinal } → x o< y → def (ord→od y) x ) → {x : OD } → x ≡ Ord (od→ord x) 94 o<→c<→OD=Ord : ( {x y : Ordinal } → x o< y → def (ord→od y) x ) → {x : OD } → x ≡ Ord (od→ord x)
92 o<→c<→OD=Ord o<→c< {x} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where 95 o<→c<→OD=Ord o<→c< {x} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where
93 lemma1 : {y : Ordinal} → def x y → def (Ord (od→ord x)) y 96 lemma1 : {y : Ordinal} → def x y → def (Ord (od→ord x)) y
94 lemma1 {y} lt = subst ( λ k → k o< od→ord x ) diso (c<→o< {ord→od y} {x} (subst (λ k → def x k ) (sym diso) lt)) 97 lemma1 {y} lt = subst ( λ k → k o< od→ord x ) diso (c<→o< {ord→od y} {x} (subst (λ k → def x k ) (sym diso) lt))
106 109
107 def-subst : {Z : OD } {X : Ordinal }{z : OD } {x : Ordinal }→ def Z X → Z ≡ z → X ≡ x → def z x 110 def-subst : {Z : OD } {X : Ordinal }{z : OD } {x : Ordinal }→ def Z X → Z ≡ z → X ≡ x → def z x
108 def-subst df refl refl = df 111 def-subst df refl refl = df
109 112
110 sup-od : ( OD → OD ) → OD 113 sup-od : ( OD → OD ) → OD
111 sup-od ψ = Ord ( sup-o ( λ x → od→ord (ψ (ord→od x ))) ) 114 sup-od ψ = Ord ( sup-o ( λ x → od→ord (ψ x)) )
112 115
113 sup-c< : ( ψ : OD → OD ) → ∀ {x : OD } → def ( sup-od ψ ) (od→ord ( ψ x )) 116 sup-c< : ( ψ : OD → OD ) → ∀ {x : OD } → def ( sup-od ψ ) (od→ord ( ψ x ))
114 sup-c< ψ {x} = def-subst {_} {_} {Ord ( sup-o ( λ x → od→ord (ψ (ord→od x ))) )} {od→ord ( ψ x )} 117 sup-c< ψ {x} = def-subst {_} {_} {Ord ( sup-o ( λ x → od→ord (ψ x)) )} {od→ord ( ψ x )}
115 lemma refl (cong ( λ k → od→ord (ψ k) ) oiso) where 118 lemma refl (cong ( λ k → od→ord (ψ k) ) oiso) where
116 lemma : od→ord (ψ (ord→od (od→ord x))) o< sup-o (λ x → od→ord (ψ (ord→od x))) 119 lemma : od→ord (ψ (ord→od (od→ord x))) o< sup-o (λ x → od→ord (ψ x))
117 lemma = subst₂ (λ j k → j o< k ) refl diso (o<-subst (sup-o< ) refl (sym diso) ) 120 lemma = subst₂ (λ j k → j o< k ) refl diso (o<-subst (sup-o< ) refl (sym diso) )
118 121
119 otrans : {n : Level} {a x y : Ordinal } → def (Ord a) x → def (Ord x) y → def (Ord a) y 122 otrans : {n : Level} {a x y : Ordinal } → def (Ord a) x → def (Ord x) y → def (Ord a) y
120 otrans x<a y<x = ordtrans y<x x<a 123 otrans x<a y<x = ordtrans y<x x<a
121 124
179 is-o∅ x | tri> ¬a ¬b c = no ¬b 182 is-o∅ x | tri> ¬a ¬b c = no ¬b
180 183
181 _,_ : OD → OD → OD 184 _,_ : OD → OD → OD
182 x , y = record { def = λ t → (t ≡ od→ord x ) ∨ ( t ≡ od→ord y ) } -- Ord (omax (od→ord x) (od→ord y)) 185 x , y = record { def = λ t → (t ≡ od→ord x ) ∨ ( t ≡ od→ord y ) } -- Ord (omax (od→ord x) (od→ord y))
183 186
184 --
185 -- Axiom of choice in intutionistic logic implies the exclude middle
186 -- https://plato.stanford.edu/entries/axiom-choice/#AxiChoLog
187 --
188
189 ppp : { p : Set n } { a : OD } → record { def = λ x → p } ∋ a → p
190 ppp {p} {a} d = d
191
192 p∨¬p : ( p : Set n ) → p ∨ ( ¬ p ) -- assuming axiom of choice
193 p∨¬p p with is-o∅ ( od→ord ( record { def = λ x → p } ))
194 p∨¬p p | yes eq = case2 (¬p eq) where
195 ps = record { def = λ x → p }
196 lemma : ps == od∅ → p → ⊥
197 lemma eq p0 = ¬x<0 {od→ord ps} (eq→ eq p0 )
198 ¬p : (od→ord ps ≡ o∅) → p → ⊥
199 ¬p eq = lemma ( subst₂ (λ j k → j == k ) oiso o∅≡od∅ ( o≡→== eq ))
200 p∨¬p p | no ¬p = case1 (ppp {p} {minimal ps (λ eq → ¬p (eqo∅ eq))} lemma) where
201 ps = record { def = λ x → p }
202 eqo∅ : ps == od∅ → od→ord ps ≡ o∅
203 eqo∅ eq = subst (λ k → od→ord ps ≡ k) ord-od∅ ( cong (λ k → od→ord k ) (==→o≡ eq))
204 lemma : ps ∋ minimal ps (λ eq → ¬p (eqo∅ eq))
205 lemma = x∋minimal ps (λ eq → ¬p (eqo∅ eq))
206
207 decp : ( p : Set n ) → Dec p -- assuming axiom of choice
208 decp p with p∨¬p p
209 decp p | case1 x = yes x
210 decp p | case2 x = no x
211
212 double-neg-eilm : {A : Set n} → ¬ ¬ A → A -- we don't have this in intutionistic logic
213 double-neg-eilm {A} notnot with decp A -- assuming axiom of choice
214 ... | yes p = p
215 ... | no ¬p = ⊥-elim ( notnot ¬p )
216
217 OrdP : ( x : Ordinal ) ( y : OD ) → Dec ( Ord x ∋ y )
218 OrdP x y with trio< x (od→ord y)
219 OrdP x y | tri< a ¬b ¬c = no ¬c
220 OrdP x y | tri≈ ¬a refl ¬c = no ( o<¬≡ refl )
221 OrdP x y | tri> ¬a ¬b c = yes c
222
223 -- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) 187 -- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ )
224 -- postulate f-extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality n (suc n) 188 -- postulate f-extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality n (suc n)
225 189
226 in-codomain : (X : OD ) → ( ψ : OD → OD ) → OD 190 in-codomain : (X : OD ) → ( ψ : OD → OD ) → OD
227 in-codomain X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( def X y ∧ ( x ≡ od→ord (ψ (ord→od y ))))) } 191 in-codomain X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( def X y ∧ ( x ≡ od→ord (ψ (ord→od y ))))) }
230 194
231 ZFSubset : (A x : OD ) → OD 195 ZFSubset : (A x : OD ) → OD
232 ZFSubset A x = record { def = λ y → def A y ∧ def x y } -- roughly x = A → Set 196 ZFSubset A x = record { def = λ y → def A y ∧ def x y } -- roughly x = A → Set
233 197
234 Def : (A : OD ) → OD 198 Def : (A : OD ) → OD
235 Def A = Ord ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) ) 199 Def A = Ord ( sup-o ( λ x → od→ord ( ZFSubset A x) ) )
236 200
237 -- _⊆_ : ( A B : OD ) → ∀{ x : OD } → Set n 201 -- _⊆_ : ( A B : OD ) → ∀{ x : OD } → Set n
238 -- _⊆_ A B {x} = A ∋ x → B ∋ x 202 -- _⊆_ A B {x} = A ∋ x → B ∋ x
239 203
240 record _⊆_ ( A B : OD ) : Set (suc n) where 204 record _⊆_ ( A B : OD ) : Set (suc n) where
281 } where 245 } where
282 ZFSet = OD 246 ZFSet = OD
283 Select : (X : OD ) → ((x : OD ) → Set n ) → OD 247 Select : (X : OD ) → ((x : OD ) → Set n ) → OD
284 Select X ψ = record { def = λ x → ( def X x ∧ ψ ( ord→od x )) } 248 Select X ψ = record { def = λ x → ( def X x ∧ ψ ( ord→od x )) }
285 Replace : OD → (OD → OD ) → OD 249 Replace : OD → (OD → OD ) → OD
286 Replace X ψ = record { def = λ x → (x o< sup-o ( λ x → od→ord (ψ (ord→od x )))) ∧ def (in-codomain X ψ) x } 250 Replace X ψ = record { def = λ x → (x o< sup-o ( λ x → od→ord (ψ x))) ∧ def (in-codomain X ψ) x }
287 _∩_ : ( A B : ZFSet ) → ZFSet 251 _∩_ : ( A B : ZFSet ) → ZFSet
288 A ∩ B = record { def = λ x → def A x ∧ def B x } 252 A ∩ B = record { def = λ x → def A x ∧ def B x }
289 Union : OD → OD 253 Union : OD → OD
290 Union U = record { def = λ x → ¬ (∀ (u : Ordinal ) → ¬ ((def U u) ∧ (def (ord→od u) x))) } 254 Union U = record { def = λ x → ¬ (∀ (u : Ordinal ) → ¬ ((def U u) ∧ (def (ord→od u) x))) }
291 _∈_ : ( A B : ZFSet ) → Set n 255 _∈_ : ( A B : ZFSet ) → Set n
404 proj1 = def-subst {_} {_} {(Ord a)} {z} 368 proj1 = def-subst {_} {_} {(Ord a)} {z}
405 ( t→A (def-subst {_} {_} {t} {od→ord (ord→od z)} w refl (sym diso) )) refl diso } 369 ( t→A (def-subst {_} {_} {t} {od→ord (ord→od z)} w refl (sym diso) )) refl diso }
406 lemma1 : {a : Ordinal } { t : OD } 370 lemma1 : {a : Ordinal } { t : OD }
407 → (eq : ZFSubset (Ord a) t == t) → od→ord (ZFSubset (Ord a) (ord→od (od→ord t))) ≡ od→ord t 371 → (eq : ZFSubset (Ord a) t == t) → od→ord (ZFSubset (Ord a) (ord→od (od→ord t))) ≡ od→ord t
408 lemma1 {a} {t} eq = subst (λ k → od→ord (ZFSubset (Ord a) k) ≡ od→ord t ) (sym oiso) (cong (λ k → od→ord k ) (==→o≡ eq )) 372 lemma1 {a} {t} eq = subst (λ k → od→ord (ZFSubset (Ord a) k) ≡ od→ord t ) (sym oiso) (cong (λ k → od→ord k ) (==→o≡ eq ))
409 lemma : od→ord (ZFSubset (Ord a) (ord→od (od→ord t)) ) o< sup-o (λ x → od→ord (ZFSubset (Ord a) (ord→od x))) 373 lemma : od→ord (ZFSubset (Ord a) (ord→od (od→ord t)) ) o< sup-o (λ x → od→ord (ZFSubset (Ord a) x))
410 lemma = sup-o< 374 lemma = sup-o<
411 375
412 -- 376 --
413 -- Every set in OD is a subset of Ordinals, so make Def (Ord (od→ord A)) first 377 -- Every set in OD is a subset of Ordinals, so make Def (Ord (od→ord A)) first
414 -- then replace of all elements of the Power set by A ∩ y 378 -- then replace of all elements of the Power set by A ∩ y
442 ≡⟨ cong (λ k → A ∩ k) oiso ⟩ 406 ≡⟨ cong (λ k → A ∩ k) oiso ⟩
443 A ∩ t 407 A ∩ t
444 ≡⟨ sym (==→o≡ ( ∩-≡ t→A )) ⟩ 408 ≡⟨ sym (==→o≡ ( ∩-≡ t→A )) ⟩
445 t 409 t
446 410
447 lemma1 : od→ord t o< sup-o (λ x → od→ord (A ∩ ord→od x)) 411 lemma1 : od→ord t o< sup-o (λ x → od→ord (A ∩ x))
448 lemma1 = subst (λ k → od→ord k o< sup-o (λ x → od→ord (A ∩ ord→od x))) 412 lemma1 = subst (λ k → od→ord k o< sup-o (λ x → od→ord (A ∩ x)))
449 lemma4 (sup-o< {λ x → od→ord (A ∩ ord→od x)} {od→ord t} ) 413 lemma4 (sup-o< {λ x → od→ord (A ∩ x)} )
450 lemma2 : def (in-codomain (Def (Ord (od→ord A))) (_∩_ A)) (od→ord t) 414 lemma2 : def (in-codomain (Def (Ord (od→ord A))) (_∩_ A)) (od→ord t)
451 lemma2 not = ⊥-elim ( not (od→ord t) (record { proj1 = lemma3 ; proj2 = lemma6 }) ) where 415 lemma2 not = ⊥-elim ( not (od→ord t) (record { proj1 = lemma3 ; proj2 = lemma6 }) ) where
452 lemma6 : od→ord t ≡ od→ord (A ∩ ord→od (od→ord t)) 416 lemma6 : od→ord t ≡ od→ord (A ∩ ord→od (od→ord t))
453 lemma6 = cong ( λ k → od→ord k ) (==→o≡ (subst (λ k → t == (A ∩ k)) (sym oiso) ( ∩-≡ t→A ))) 417 lemma6 = cong ( λ k → od→ord k ) (==→o≡ (subst (λ k → t == (A ∩ k)) (sym oiso) ( ∩-≡ t→A )))
454 418
457 lemma : {x y : OD} → od→ord x o< osuc a → x ∋ y → Ord a ∋ y 421 lemma : {x y : OD} → od→ord x o< osuc a → x ∋ y → Ord a ∋ y
458 lemma lt y<x with osuc-≡< lt 422 lemma lt y<x with osuc-≡< lt
459 lemma lt y<x | case1 refl = c<→o< y<x 423 lemma lt y<x | case1 refl = c<→o< y<x
460 lemma lt y<x | case2 x<a = ordtrans (c<→o< y<x) x<a 424 lemma lt y<x | case2 x<a = ordtrans (c<→o< y<x) x<a
461 425
462 -- continuum-hyphotheis : (a : Ordinal) → Power (Ord a) ⊆ Ord (osuc a) 426 continuum-hyphotheis : (a : Ordinal) → Set (suc n)
463 -- continuum-hyphotheis a = ? 427 continuum-hyphotheis a = Power (Ord a) ⊆ Ord (osuc a)
464
465 -- assuming axiom of choice
466 regularity : (x : OD) (not : ¬ (x == od∅)) →
467 (x ∋ minimal x not) ∧ (Select (minimal x not) (λ x₁ → (minimal x not ∋ x₁) ∧ (x ∋ x₁)) == od∅)
468 proj1 (regularity x not ) = x∋minimal x not
469 proj2 (regularity x not ) = record { eq→ = lemma1 ; eq← = λ {y} d → lemma2 {y} d } where
470 lemma1 : {x₁ : Ordinal} → def (Select (minimal x not) (λ x₂ → (minimal x not ∋ x₂) ∧ (x ∋ x₂))) x₁ → def od∅ x₁
471 lemma1 {x₁} s = ⊥-elim ( minimal-1 x not (ord→od x₁) lemma3 ) where
472 lemma3 : def (minimal x not) (od→ord (ord→od x₁)) ∧ def x (od→ord (ord→od x₁))
473 lemma3 = record { proj1 = def-subst {_} {_} {minimal x not} {_} (proj1 s) refl (sym diso)
474 ; proj2 = proj2 (proj2 s) }
475 lemma2 : {x₁ : Ordinal} → def od∅ x₁ → def (Select (minimal x not) (λ x₂ → (minimal x not ∋ x₂) ∧ (x ∋ x₂))) x₁
476 lemma2 {y} d = ⊥-elim (empty (ord→od y) (def-subst {_} {_} {od∅} {od→ord (ord→od y)} d refl (sym diso) ))
477 428
478 extensionality0 : {A B : OD } → ((z : OD) → (A ∋ z) ⇔ (B ∋ z)) → A == B 429 extensionality0 : {A B : OD } → ((z : OD) → (A ∋ z) ⇔ (B ∋ z)) → A == B
479 eq→ (extensionality0 {A} {B} eq ) {x} d = def-iso {A} {B} (sym diso) (proj1 (eq (ord→od x))) d 430 eq→ (extensionality0 {A} {B} eq ) {x} d = def-iso {A} {B} (sym diso) (proj1 (eq (ord→od x))) d
480 eq← (extensionality0 {A} {B} eq ) {x} d = def-iso {B} {A} (sym diso) (proj2 (eq (ord→od x))) d 431 eq← (extensionality0 {A} {B} eq ) {x} d = def-iso {B} {A} (sym diso) (proj2 (eq (ord→od x))) d
481 432
497 infinity x lt = def-subst {_} {_} {infinite} {od→ord (Union (x , (x , x )))} ( isuc {od→ord x} lt ) refl lemma where 448 infinity x lt = def-subst {_} {_} {infinite} {od→ord (Union (x , (x , x )))} ( isuc {od→ord x} lt ) refl lemma where
498 lemma : od→ord (Union (ord→od (od→ord x) , (ord→od (od→ord x) , ord→od (od→ord x)))) 449 lemma : od→ord (Union (ord→od (od→ord x) , (ord→od (od→ord x) , ord→od (od→ord x))))
499 ≡ od→ord (Union (x , (x , x))) 450 ≡ od→ord (Union (x , (x , x)))
500 lemma = cong (λ k → od→ord (Union ( k , ( k , k ) ))) oiso 451 lemma = cong (λ k → od→ord (Union ( k , ( k , k ) ))) oiso
501 452
502 -- Axiom of choice ( is equivalent to the existence of minimal in our case )
503 -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ]
504 choice-func : (X : OD ) → {x : OD } → ¬ ( x == od∅ ) → ( X ∋ x ) → OD
505 choice-func X {x} not X∋x = minimal x not
506 choice : (X : OD ) → {A : OD } → ( X∋A : X ∋ A ) → (not : ¬ ( A == od∅ )) → A ∋ choice-func X not X∋A
507 choice X {A} X∋A not = x∋minimal A not
508
509 ---
510 --- With assuption of OD is ordered, p ∨ ( ¬ p ) <=> axiom of choice
511 ---
512 record choiced ( X : OD) : Set (suc n) where
513 field
514 a-choice : OD
515 is-in : X ∋ a-choice
516
517 choice-func' : (X : OD ) → (p∨¬p : ( p : Set (suc n)) → p ∨ ( ¬ p )) → ¬ ( X == od∅ ) → choiced X
518 choice-func' X p∨¬p not = have_to_find where
519 ψ : ( ox : Ordinal ) → Set (suc n)
520 ψ ox = (( x : Ordinal ) → x o< ox → ( ¬ def X x )) ∨ choiced X
521 lemma-ord : ( ox : Ordinal ) → ψ ox
522 lemma-ord ox = TransFinite {ψ} induction ox where
523 ∋-p : (A x : OD ) → Dec ( A ∋ x )
524 ∋-p A x with p∨¬p (Lift (suc n) ( A ∋ x )) -- LEM
525 ∋-p A x | case1 (lift t) = yes t
526 ∋-p A x | case2 t = no (λ x → t (lift x ))
527 ∀-imply-or : {A : Ordinal → Set n } {B : Set (suc n) }
528 → ((x : Ordinal ) → A x ∨ B) → ((x : Ordinal ) → A x) ∨ B
529 ∀-imply-or {A} {B} ∀AB with p∨¬p (Lift ( suc n ) ((x : Ordinal ) → A x)) -- LEM
530 ∀-imply-or {A} {B} ∀AB | case1 (lift t) = case1 t
531 ∀-imply-or {A} {B} ∀AB | case2 x = case2 (lemma (λ not → x (lift not ))) where
532 lemma : ¬ ((x : Ordinal ) → A x) → B
533 lemma not with p∨¬p B
534 lemma not | case1 b = b
535 lemma not | case2 ¬b = ⊥-elim (not (λ x → dont-orb (∀AB x) ¬b ))
536 induction : (x : Ordinal) → ((y : Ordinal) → y o< x → ψ y) → ψ x
537 induction x prev with ∋-p X ( ord→od x)
538 ... | yes p = case2 ( record { a-choice = ord→od x ; is-in = p } )
539 ... | no ¬p = lemma where
540 lemma1 : (y : Ordinal) → (y o< x → def X y → ⊥) ∨ choiced X
541 lemma1 y with ∋-p X (ord→od y)
542 lemma1 y | yes y<X = case2 ( record { a-choice = ord→od y ; is-in = y<X } )
543 lemma1 y | no ¬y<X = case1 ( λ lt y<X → ¬y<X (subst (λ k → def X k ) (sym diso) y<X ) )
544 lemma : ((y : Ordinals.ord O) → (O Ordinals.o< y) x → def X y → ⊥) ∨ choiced X
545 lemma = ∀-imply-or lemma1
546 have_to_find : choiced X
547 have_to_find = dont-or (lemma-ord (od→ord X )) ¬¬X∋x where
548 ¬¬X∋x : ¬ ((x : Ordinal) → x o< (od→ord X) → def X x → ⊥)
549 ¬¬X∋x nn = not record {
550 eq→ = λ {x} lt → ⊥-elim (nn x (def→o< lt) lt)
551 ; eq← = λ {x} lt → ⊥-elim ( ¬x<0 lt )
552 }
553
554 453
555 Union = ZF.Union OD→ZF 454 Union = ZF.Union OD→ZF
556 Power = ZF.Power OD→ZF 455 Power = ZF.Power OD→ZF
557 Select = ZF.Select OD→ZF 456 Select = ZF.Select OD→ZF
558 Replace = ZF.Replace OD→ZF 457 Replace = ZF.Replace OD→ZF