comparison src/OD.agda @ 1091:63c1167b2343

fix comments
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Tue, 20 Dec 2022 11:20:52 +0900
parents 88fae58f89f5
children 08b6aa6870d9
comparison
equal deleted inserted replaced
1090:2cf182f0f583 1091:63c1167b2343
1 {-# OPTIONS --allow-unsolved-metas #-} 1 {-# OPTIONS --allow-unsolved-metas #-}
2 open import Level 2 open import Level
3 open import Ordinals 3 open import Ordinals
4 module OD {n : Level } (O : Ordinals {n} ) where 4 module OD {n : Level } (O : Ordinals {n} ) where
5 5
6 open import zf 6 open import zf
7 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) 7 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ )
8 open import Relation.Binary.PropositionalEquality hiding ( [_] ) 8 open import Relation.Binary.PropositionalEquality hiding ( [_] )
9 open import Data.Nat.Properties 9 open import Data.Nat.Properties
10 open import Data.Empty 10 open import Data.Empty
11 open import Relation.Nullary 11 open import Relation.Nullary
12 open import Relation.Binary hiding (_⇔_) 12 open import Relation.Binary hiding (_⇔_)
13 open import Relation.Binary.Core hiding (_⇔_) 13 open import Relation.Binary.Core hiding (_⇔_)
14 14
15 open import logic 15 open import logic
16 import OrdUtil 16 import OrdUtil
17 open import nat 17 open import nat
18 18
19 open Ordinals.Ordinals O 19 open Ordinals.Ordinals O
20 open Ordinals.IsOrdinals isOrdinal 20 open Ordinals.IsOrdinals isOrdinal
21 open Ordinals.IsNext isNext 21 open Ordinals.IsNext isNext
22 open OrdUtil O 22 open OrdUtil O
23 23
24 -- Ordinal Definable Set 24 -- Ordinal Definable Set
25 25
26 record OD : Set (suc n ) where 26 record OD : Set (suc n ) where
33 open _∨_ 33 open _∨_
34 open Bool 34 open Bool
35 35
36 record _==_ ( a b : OD ) : Set n where 36 record _==_ ( a b : OD ) : Set n where
37 field 37 field
38 eq→ : ∀ { x : Ordinal } → def a x → def b x 38 eq→ : ∀ { x : Ordinal } → def a x → def b x
39 eq← : ∀ { x : Ordinal } → def b x → def a x 39 eq← : ∀ { x : Ordinal } → def b x → def a x
40 40
41 ==-refl : { x : OD } → x == x 41 ==-refl : { x : OD } → x == x
42 ==-refl {x} = record { eq→ = λ x → x ; eq← = λ x → x } 42 ==-refl {x} = record { eq→ = λ x → x ; eq← = λ x → x }
43 43
44 open _==_ 44 open _==_
45 45
46 ==-trans : { x y z : OD } → x == y → y == z → x == z 46 ==-trans : { x y z : OD } → x == y → y == z → x == z
47 ==-trans x=y y=z = record { eq→ = λ {m} t → eq→ y=z (eq→ x=y t) ; eq← = λ {m} t → eq← x=y (eq← y=z t) } 47 ==-trans x=y y=z = record { eq→ = λ {m} t → eq→ y=z (eq→ x=y t) ; eq← = λ {m} t → eq← x=y (eq← y=z t) }
48 48
49 ==-sym : { x y : OD } → x == y → y == x 49 ==-sym : { x y : OD } → x == y → y == x
50 ==-sym x=y = record { eq→ = λ {m} t → eq← x=y t ; eq← = λ {m} t → eq→ x=y t } 50 ==-sym x=y = record { eq→ = λ {m} t → eq← x=y t ; eq← = λ {m} t → eq→ x=y t }
51 51
52 52
53 ⇔→== : { x y : OD } → ( {z : Ordinal } → (def x z ⇔ def y z)) → x == y 53 ⇔→== : { x y : OD } → ( {z : Ordinal } → (def x z ⇔ def y z)) → x == y
54 eq→ ( ⇔→== {x} {y} eq ) {z} m = proj1 eq m 54 eq→ ( ⇔→== {x} {y} eq ) {z} m = proj1 eq m
55 eq← ( ⇔→== {x} {y} eq ) {z} m = proj2 eq m 55 eq← ( ⇔→== {x} {y} eq ) {z} m = proj2 eq m
56 56
57 -- next assumptions are our axiom 57 -- next assumptions are our axiom
58 -- 58 --
59 -- OD is an equation on Ordinals, so it contains Ordinals. If these Ordinals have one-to-one 59 -- OD is an equation on Ordinals, so it contains Ordinals. If these Ordinals have one-to-one
60 -- correspondence to the OD then the OD looks like a ZF Set. 60 -- correspondence to the OD then the OD looks like a ZF Set.
61 -- 61 --
62 -- If all ZF Set have supreme upper bound, the solutions of OD have to be bounded, i.e. 62 -- If all ZF Set have supreme upper bound, the solutions of OD have to be bounded, i.e.
63 -- bbounded ODs are ZF Set. Unbounded ODs are classes. 63 -- bbounded ODs are ZF Set. Unbounded ODs are classes.
64 -- 64 --
65 -- In classical Set Theory, HOD is used, as a subset of OD, 65 -- In classical Set Theory, HOD is used, as a subset of OD,
66 -- HOD = { x | TC x ⊆ OD } 66 -- HOD = { x | TC x ⊆ OD }
67 -- where TC x is a transitive clusure of x, i.e. Union of all elemnts of all subset of x. 67 -- where TC x is a transitive clusure of x, i.e. Union of all elemnts of all subset of x.
68 -- This is not possible because we don't have V yet. So we assumes HODs are bounded OD. 68 -- This is not possible because we don't have V yet. So we assumes HODs are bounded OD.
69 -- 69 --
70 -- We also assumes HODs are isomorphic to Ordinals, which is ususally proved by Goedel number tricks. 70 -- We also assumes HODs are isomorphic to Ordinals, which is ususally proved by Goedel number tricks.
87 odmax : Ordinal 87 odmax : Ordinal
88 <odmax : {y : Ordinal} → def od y → y o< odmax 88 <odmax : {y : Ordinal} → def od y → y o< odmax
89 89
90 open HOD 90 open HOD
91 91
92 record ODAxiom : Set (suc n) where 92 record ODAxiom : Set (suc n) where
93 field 93 field
94 -- HOD is isomorphic to Ordinal (by means of Goedel number) 94 -- HOD is isomorphic to Ordinal (by means of Goedel number)
95 & : HOD → Ordinal 95 & : HOD → Ordinal
96 * : Ordinal → HOD 96 * : Ordinal → HOD
97 c<→o< : {x y : HOD } → def (od y) ( & x ) → & x o< & y 97 c<→o< : {x y : HOD } → def (od y) ( & x ) → & x o< & y
98 ⊆→o≤ : {y z : HOD } → ({x : Ordinal} → def (od y) x → def (od z) x ) → & y o< osuc (& z) 98 ⊆→o≤ : {y z : HOD } → ({x : Ordinal} → def (od y) x → def (od z) x ) → & y o< osuc (& z)
99 *iso : {x : HOD } → * ( & x ) ≡ x 99 *iso : {x : HOD } → * ( & x ) ≡ x
100 &iso : {x : Ordinal } → & ( * x ) ≡ x 100 &iso : {x : Ordinal } → & ( * x ) ≡ x
101 ==→o≡ : {x y : HOD } → (od x == od y) → x ≡ y 101 ==→o≡ : {x y : HOD } → (od x == od y) → x ≡ y
102 sup-o : (A : HOD) → ( ( x : Ordinal ) → def (od A) x → Ordinal ) → Ordinal 102 sup-o : (A : HOD) → ( ( x : Ordinal ) → def (od A) x → Ordinal ) → Ordinal
103 sup-o≤ : (A : HOD) → { ψ : ( x : Ordinal ) → def (od A) x → Ordinal } → ∀ {x : Ordinal } → (lt : def (od A) x ) → ψ x lt o≤ sup-o A ψ 103 sup-o≤ : (A : HOD) → { ψ : ( x : Ordinal ) → def (od A) x → Ordinal } → ∀ {x : Ordinal } → (lt : def (od A) x ) → ψ x lt o≤ sup-o A ψ
104 -- possible order restriction 104 -- possible order restriction
105 ho< : {x : HOD} → & x o< next (odmax x) 105 ho< : {x : HOD} → & x o< next (odmax x)
106 106
107 107
108 postulate odAxiom : ODAxiom 108 postulate odAxiom : ODAxiom
110 110
111 -- odmax minimality 111 -- odmax minimality
112 -- 112 --
113 -- since we have ==→o≡ , so odmax have to be unique. We should have odmaxmin in HOD. 113 -- since we have ==→o≡ , so odmax have to be unique. We should have odmaxmin in HOD.
114 -- We can calculate the minimum using sup but it is tedius. 114 -- We can calculate the minimum using sup but it is tedius.
115 -- Only Select has non minimum odmax. 115 -- Only Select has non minimum odmax.
116 -- We have the same problem on 'def' itself, but we leave it. 116 -- We have the same problem on 'def' itself, but we leave it.
117 117
118 odmaxmin : Set (suc n) 118 odmaxmin : Set (suc n)
119 odmaxmin = (y : HOD) (z : Ordinal) → ((x : Ordinal)→ def (od y) x → x o< z) → odmax y o< osuc z 119 odmaxmin = (y : HOD) (z : Ordinal) → ((x : Ordinal)→ def (od y) x → x o< z) → odmax y o< osuc z
120 120
121 -- OD ⇔ Ordinal leads a contradiction, so we need bounded HOD 121 -- OD ⇔ Ordinal leads a contradiction, so we need bounded HOD
122 ¬OD-order : ( & : OD → Ordinal ) → ( * : Ordinal → OD ) → ( { x y : OD } → def y ( & x ) → & x o< & y) → ⊥ 122 ¬OD-order : ( & : OD → Ordinal ) → ( * : Ordinal → OD ) → ( { x y : OD } → def y ( & x ) → & x o< & y) → ⊥
123 ¬OD-order & * c<→o< = o≤> <-osuc (c<→o< {Ords} OneObj ) 123 ¬OD-order & * c<→o< = o≤> <-osuc (c<→o< {Ords} OneObj )
124 124
125 -- Open supreme upper bound leads a contradition, so we use domain restriction on sup
126 ¬open-sup : ( sup-o : (Ordinal → Ordinal ) → Ordinal) → ((ψ : Ordinal → Ordinal ) → (x : Ordinal) → ψ x o< sup-o ψ ) → ⊥
127 ¬open-sup sup-o sup-o< = o<> <-osuc (sup-o< next-ord (sup-o next-ord)) where
128 next-ord : Ordinal → Ordinal
129 next-ord x = osuc x
130
131 -- Ordinal in OD ( and ZFSet ) Transitive Set 125 -- Ordinal in OD ( and ZFSet ) Transitive Set
132 Ord : ( a : Ordinal ) → HOD 126 Ord : ( a : Ordinal ) → HOD
133 Ord a = record { od = record { def = λ y → y o< a } ; odmax = a ; <odmax = lemma } where 127 Ord a = record { od = record { def = λ y → y o< a } ; odmax = a ; <odmax = lemma } where
134 lemma : {x : Ordinal} → x o< a → x o< a 128 lemma : {x : Ordinal} → x o< a → x o< a
135 lemma {x} lt = lt 129 lemma {x} lt = lt
136 130
137 od∅ : HOD 131 od∅ : HOD
138 od∅ = Ord o∅ 132 od∅ = Ord o∅
139 133
140 odef : HOD → Ordinal → Set n 134 odef : HOD → Ordinal → Set n
141 odef A x = def ( od A ) x 135 odef A x = def ( od A ) x
142 136
143 _∋_ : ( a x : HOD ) → Set n 137 _∋_ : ( a x : HOD ) → Set n
144 _∋_ a x = odef a ( & x ) 138 _∋_ a x = odef a ( & x )
145 139
146 -- _c<_ : ( x a : HOD ) → Set n 140 -- _c<_ : ( x a : HOD ) → Set n
147 -- x c< a = a ∋ x 141 -- x c< a = a ∋ x
148 142
149 d→∋ : ( a : HOD ) { x : Ordinal} → odef a x → a ∋ (* x) 143 d→∋ : ( a : HOD ) { x : Ordinal} → odef a x → a ∋ (* x)
150 d→∋ a lt = subst (λ k → odef a k ) (sym &iso) lt 144 d→∋ a lt = subst (λ k → odef a k ) (sym &iso) lt
151 145
152 -- odef-subst : {Z : HOD } {X : Ordinal }{z : HOD } {x : Ordinal }→ odef Z X → Z ≡ z → X ≡ x → odef z x 146 -- odef-subst : {Z : HOD } {X : Ordinal }{z : HOD } {x : Ordinal }→ odef Z X → Z ≡ z → X ≡ x → odef z x
167 orefl : { x : HOD } → { y : Ordinal } → & x ≡ y → & x ≡ y 161 orefl : { x : HOD } → { y : Ordinal } → & x ≡ y → & x ≡ y
168 orefl refl = refl 162 orefl refl = refl
169 163
170 ==-iso : { x y : HOD } → od (* (& x)) == od (* (& y)) → od x == od y 164 ==-iso : { x y : HOD } → od (* (& x)) == od (* (& y)) → od x == od y
171 ==-iso {x} {y} eq = record { 165 ==-iso {x} {y} eq = record {
172 eq→ = λ {z} d → lemma ( eq→ eq (subst (λ k → odef k z ) (sym *iso) d )) ; 166 eq→ = λ {z} d → lemma ( eq→ eq (subst (λ k → odef k z ) (sym *iso) d )) ;
173 eq← = λ {z} d → lemma ( eq← eq (subst (λ k → odef k z ) (sym *iso) d )) } 167 eq← = λ {z} d → lemma ( eq← eq (subst (λ k → odef k z ) (sym *iso) d )) }
174 where 168 where
175 lemma : {x : HOD } {z : Ordinal } → odef (* (& x)) z → odef x z 169 lemma : {x : HOD } {z : Ordinal } → odef (* (& x)) z → odef x z
176 lemma {x} {z} d = subst (λ k → odef k z) (*iso) d 170 lemma {x} {z} d = subst (λ k → odef k z) (*iso) d
177 171
178 =-iso : {x y : HOD } → (od x == od y) ≡ (od (* (& x)) == od y) 172 =-iso : {x y : HOD } → (od x == od y) ≡ (od (* (& x)) == od y)
179 =-iso {_} {y} = cong ( λ k → od k == od y ) (sym *iso) 173 =-iso {_} {y} = cong ( λ k → od k == od y ) (sym *iso)
180 174
181 ord→== : { x y : HOD } → & x ≡ & y → od x == od y 175 ord→== : { x y : HOD } → & x ≡ & y → od x == od y
190 *≡*→≡ eq = subst₂ (λ j k → j ≡ k ) &iso &iso ( cong (&) eq ) 184 *≡*→≡ eq = subst₂ (λ j k → j ≡ k ) &iso &iso ( cong (&) eq )
191 185
192 &≡&→≡ : { x y : HOD } → & x ≡ & y → x ≡ y 186 &≡&→≡ : { x y : HOD } → & x ≡ & y → x ≡ y
193 &≡&→≡ eq = subst₂ (λ j k → j ≡ k ) *iso *iso ( cong (*) eq ) 187 &≡&→≡ eq = subst₂ (λ j k → j ≡ k ) *iso *iso ( cong (*) eq )
194 188
195 o∅≡od∅ : * (o∅ ) ≡ od∅ 189 o∅≡od∅ : * (o∅ ) ≡ od∅
196 o∅≡od∅ = ==→o≡ lemma where 190 o∅≡od∅ = ==→o≡ lemma where
197 lemma0 : {x : Ordinal} → odef (* o∅) x → odef od∅ x 191 lemma0 : {x : Ordinal} → odef (* o∅) x → odef od∅ x
198 lemma0 {x} lt with c<→o< {* x} {* o∅} (subst (λ k → odef (* o∅) k ) (sym &iso) lt) 192 lemma0 {x} lt with c<→o< {* x} {* o∅} (subst (λ k → odef (* o∅) k ) (sym &iso) lt)
199 ... | t = subst₂ (λ j k → j o< k ) &iso &iso t 193 ... | t = subst₂ (λ j k → j o< k ) &iso &iso t
200 lemma1 : {x : Ordinal} → odef od∅ x → odef (* o∅) x 194 lemma1 : {x : Ordinal} → odef od∅ x → odef (* o∅) x
201 lemma1 {x} lt = ⊥-elim (¬x<0 lt) 195 lemma1 {x} lt = ⊥-elim (¬x<0 lt)
202 lemma : od (* o∅) == od od∅ 196 lemma : od (* o∅) == od od∅
203 lemma = record { eq→ = lemma0 ; eq← = lemma1 } 197 lemma = record { eq→ = lemma0 ; eq← = lemma1 }
204 198
205 ord-od∅ : & (od∅ ) ≡ o∅ 199 ord-od∅ : & (od∅ ) ≡ o∅
206 ord-od∅ = sym ( subst (λ k → k ≡ & (od∅ ) ) &iso (cong ( λ k → & k ) o∅≡od∅ ) ) 200 ord-od∅ = sym ( subst (λ k → k ≡ & (od∅ ) ) &iso (cong ( λ k → & k ) o∅≡od∅ ) )
207 201
208 ≡o∅→=od∅ : {x : HOD} → & x ≡ o∅ → od x == od od∅ 202 ≡o∅→=od∅ : {x : HOD} → & x ≡ o∅ → od x == od od∅
209 ≡o∅→=od∅ {x} eq = record { eq→ = λ {y} lt → ⊥-elim ( ¬x<0 {y} (subst₂ (λ j k → j o< k ) &iso eq ( c<→o< {* y} {x} (d→∋ x lt)))) 203 ≡o∅→=od∅ {x} eq = record { eq→ = λ {y} lt → ⊥-elim ( ¬x<0 {y} (subst₂ (λ j k → j o< k ) &iso eq ( c<→o< {* y} {x} (d→∋ x lt))))
210 ; eq← = λ {y} lt → ⊥-elim ( ¬x<0 lt )} 204 ; eq← = λ {y} lt → ⊥-elim ( ¬x<0 lt )}
211 205
212 =od∅→≡o∅ : {x : HOD} → od x == od od∅ → & x ≡ o∅ 206 =od∅→≡o∅ : {x : HOD} → od x == od od∅ → & x ≡ o∅
213 =od∅→≡o∅ {x} eq = trans (cong (λ k → & k ) (==→o≡ {x} {od∅} eq)) ord-od∅ 207 =od∅→≡o∅ {x} eq = trans (cong (λ k → & k ) (==→o≡ {x} {od∅} eq)) ord-od∅
214 208
215 ≡od∅→=od∅ : {x : HOD} → x ≡ od∅ → od x == od od∅ 209 ≡od∅→=od∅ : {x : HOD} → x ≡ od∅ → od x == od od∅
216 ≡od∅→=od∅ {x} eq = ≡o∅→=od∅ (subst (λ k → & x ≡ k ) ord-od∅ ( cong & eq ) ) 210 ≡od∅→=od∅ {x} eq = ≡o∅→=od∅ (subst (λ k → & x ≡ k ) ord-od∅ ( cong & eq ) )
217 211
218 ∅0 : record { def = λ x → Lift n ⊥ } == od od∅ 212 ∅0 : record { def = λ x → Lift n ⊥ } == od od∅
219 eq→ ∅0 {w} (lift ()) 213 eq→ ∅0 {w} (lift ())
220 eq← ∅0 {w} lt = lift (¬x<0 lt) 214 eq← ∅0 {w} lt = lift (¬x<0 lt)
221 215
222 ∅< : { x y : HOD } → odef x (& y ) → ¬ ( od x == od od∅ ) 216 ∅< : { x y : HOD } → odef x (& y ) → ¬ ( od x == od od∅ )
223 ∅< {x} {y} d eq with eq→ (==-trans eq (==-sym ∅0) ) d 217 ∅< {x} {y} d eq with eq→ (==-trans eq (==-sym ∅0) ) d
239 is-o∅ x with trio< x o∅ 233 is-o∅ x with trio< x o∅
240 is-o∅ x | tri< a ¬b ¬c = no ¬b 234 is-o∅ x | tri< a ¬b ¬c = no ¬b
241 is-o∅ x | tri≈ ¬a b ¬c = yes b 235 is-o∅ x | tri≈ ¬a b ¬c = yes b
242 is-o∅ x | tri> ¬a ¬b c = no ¬b 236 is-o∅ x | tri> ¬a ¬b c = no ¬b
243 237
238 odef< : {b : Ordinal } { A : HOD } → odef A b → b o< & A
239 odef< {b} {A} ab = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) ab))
240
241 odef∧< : {A : HOD } {y : Ordinal} {n : Level } → {P : Set n} → odef A y ∧ P → y o< & A
242 odef∧< {A } {y} p = subst (λ k → k o< & A) &iso ( c<→o< (subst (λ k → odef A k ) (sym &iso ) (proj1 p )))
243
244 -- the pair 244 -- the pair
245 _,_ : HOD → HOD → HOD 245 _,_ : HOD → HOD → HOD
246 x , y = record { od = record { def = λ t → (t ≡ & x ) ∨ ( t ≡ & y ) } ; odmax = omax (& x) (& y) ; <odmax = lemma } where 246 x , y = record { od = record { def = λ t → (t ≡ & x ) ∨ ( t ≡ & y ) } ; odmax = omax (& x) (& y) ; <odmax = lemma } where
247 lemma : {t : Ordinal} → (t ≡ & x) ∨ (t ≡ & y) → t o< omax (& x) (& y) 247 lemma : {t : Ordinal} → (t ≡ & x) ∨ (t ≡ & y) → t o< omax (& x) (& y)
248 lemma {t} (case1 refl) = omax-x _ _ 248 lemma {t} (case1 refl) = omax-x _ _
249 lemma {t} (case2 refl) = omax-y _ _ 249 lemma {t} (case2 refl) = omax-y _ _
250 250
256 256
257 -- another possible restriction. We require no minimality on odmax, so it may arbitrary larger. 257 -- another possible restriction. We require no minimality on odmax, so it may arbitrary larger.
258 odmax<& : { x y : HOD } → x ∋ y → Set n 258 odmax<& : { x y : HOD } → x ∋ y → Set n
259 odmax<& {x} {y} x∋y = odmax x o< & x 259 odmax<& {x} {y} x∋y = odmax x o< & x
260 260
261 in-codomain : (X : HOD ) → ( ψ : HOD → HOD ) → OD 261 in-codomain : (X : HOD ) → ( ψ : HOD → HOD ) → OD
262 in-codomain X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( odef X y ∧ ( x ≡ & (ψ (* y ))))) } 262 in-codomain X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( odef X y ∧ ( x ≡ & (ψ (* y ))))) }
263 263
264 _∩_ : ( A B : HOD ) → HOD 264 _∩_ : ( A B : HOD ) → HOD
265 A ∩ B = record { od = record { def = λ x → odef A x ∧ odef B x } 265 A ∩ B = record { od = record { def = λ x → odef A x ∧ odef B x }
266 ; odmax = omin (odmax A) (odmax B) ; <odmax = λ y → min1 (<odmax A (proj1 y)) (<odmax B (proj2 y))} 266 ; odmax = omin (odmax A) (odmax B) ; <odmax = λ y → min1 (<odmax A (proj1 y)) (<odmax B (proj2 y))}
267 267
268 record _⊆_ ( A B : HOD ) : Set (suc n) where 268 record _⊆_ ( A B : HOD ) : Set (suc n) where
269 field 269 field
270 incl : { x : HOD } → A ∋ x → B ∋ x 270 incl : { x : HOD } → A ∋ x → B ∋ x
271 271
272 open _⊆_ 272 open _⊆_
273 infixr 220 _⊆_ 273 infixr 220 _⊆_
274 274
275 -- if we have & (x , x) ≡ osuc (& x), ⊆→o≤ → c<→o< 275 -- if we have & (x , x) ≡ osuc (& x), ⊆→o≤ → c<→o<
276 ⊆→o≤→c<→o< : ({x : HOD} → & (x , x) ≡ osuc (& x) ) 276 ⊆→o≤→c<→o< : ({x : HOD} → & (x , x) ≡ osuc (& x) )
277 → ({y z : HOD } → ({x : Ordinal} → def (od y) x → def (od z) x ) → & y o< osuc (& z) ) 277 → ({y z : HOD } → ({x : Ordinal} → def (od y) x → def (od z) x ) → & y o< osuc (& z) )
278 → {x y : HOD } → def (od y) ( & x ) → & x o< & y 278 → {x y : HOD } → def (od y) ( & x ) → & x o< & y
279 ⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x with trio< (& x) (& y) 279 ⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x with trio< (& x) (& y)
280 ⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri< a ¬b ¬c = a 280 ⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri< a ¬b ¬c = a
281 ⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri≈ ¬a b ¬c = ⊥-elim ( o<¬≡ (peq {x}) (pair<y (subst (λ k → k ∋ x) (sym ( ==→o≡ {x} {y} (ord→== b))) y∋x ))) 281 ⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri≈ ¬a b ¬c = ⊥-elim ( o<¬≡ (peq {x}) (pair<y (subst (λ k → k ∋ x) (sym ( ==→o≡ {x} {y} (ord→== b))) y∋x )))
282 ⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri> ¬a ¬b c = 282 ⊆→o≤→c<→o< peq ⊆→o≤ {x} {y} y∋x | tri> ¬a ¬b c =
283 ⊥-elim ( o<> (⊆→o≤ {x , x} {y} y⊆x,x ) lemma1 ) where 283 ⊥-elim ( o<> (⊆→o≤ {x , x} {y} y⊆x,x ) lemma1 ) where
284 lemma : {z : Ordinal} → (z ≡ & x) ∨ (z ≡ & x) → & x ≡ z 284 lemma : {z : Ordinal} → (z ≡ & x) ∨ (z ≡ & x) → & x ≡ z
285 lemma (case1 refl) = refl 285 lemma (case1 refl) = refl
286 lemma (case2 refl) = refl 286 lemma (case2 refl) = refl
287 y⊆x,x : {z : Ordinal} → def (od (x , x)) z → def (od y) z 287 y⊆x,x : {z : Ordinal} → def (od (x , x)) z → def (od y) z
288 y⊆x,x {z} lt = subst (λ k → def (od y) k ) (lemma lt) y∋x 288 y⊆x,x {z} lt = subst (λ k → def (od y) k ) (lemma lt) y∋x
289 lemma1 : osuc (& y) o< & (x , x) 289 lemma1 : osuc (& y) o< & (x , x)
290 lemma1 = subst (λ k → osuc (& y) o< k ) (sym (peq {x})) (osucc c ) 290 lemma1 = subst (λ k → osuc (& y) o< k ) (sym (peq {x})) (osucc c )
291 291
292 ε-induction : { ψ : HOD → Set (suc n)} 292 ε-induction : { ψ : HOD → Set (suc n)}
293 → ( {x : HOD } → ({ y : HOD } → x ∋ y → ψ y ) → ψ x ) 293 → ( {x : HOD } → ({ y : HOD } → x ∋ y → ψ y ) → ψ x )
294 → (x : HOD ) → ψ x 294 → (x : HOD ) → ψ x
295 ε-induction {ψ} ind x = subst (λ k → ψ k ) *iso (ε-induction-ord (osuc (& x)) <-osuc ) where 295 ε-induction {ψ} ind x = subst (λ k → ψ k ) *iso (ε-induction-ord (osuc (& x)) <-osuc ) where
296 induction : (ox : Ordinal) → ((oy : Ordinal) → oy o< ox → ψ (* oy)) → ψ (* ox) 296 induction : (ox : Ordinal) → ((oy : Ordinal) → oy o< ox → ψ (* oy)) → ψ (* ox)
297 induction ox prev = ind ( λ {y} lt → subst (λ k → ψ k ) *iso (prev (& y) (o<-subst (c<→o< lt) refl &iso ))) 297 induction ox prev = ind ( λ {y} lt → subst (λ k → ψ k ) *iso (prev (& y) (o<-subst (c<→o< lt) refl &iso )))
298 ε-induction-ord : (ox : Ordinal) { oy : Ordinal } → oy o< ox → ψ (* oy) 298 ε-induction-ord : (ox : Ordinal) { oy : Ordinal } → oy o< ox → ψ (* oy)
299 ε-induction-ord ox {oy} lt = TransFinite {λ oy → ψ (* oy)} induction oy 299 ε-induction-ord ox {oy} lt = TransFinite {λ oy → ψ (* oy)} induction oy
300 300
301 Select : (X : HOD ) → ((x : HOD ) → Set n ) → HOD 301 record OSUP (A x : Ordinal ) ( ψ : ( x : Ordinal ) → odef (* A) x → Ordinal ) : Set n where
302 field
303 sup-o' : Ordinal
304 sup-o≤' : {z : Ordinal } → z o< x → (lt : odef (* A) z ) → ψ z lt o≤ sup-o'
305
306 -- o<-sup : ( A x : Ordinal) { ψ : ( x : Ordinal ) → odef (* A) x → Ordinal } → OSUP A x ψ
307 -- o<-sup A x {ψ} = ? where
308 -- m00 : (x : Ordinal ) → OSUP A x ψ
309 -- m00 x = Ordinals.inOrdinal.TransFinite0 O ind x where
310 -- ind : (x : Ordinal) → ((z : Ordinal) → z o< x → OSUP A z ψ ) → OSUP A x ψ
311 -- ind x prev = ? where
312 -- if has prev , od01 is empty use prev (cannot be odef (* A) x)
313 -- not empty, take larger
314 -- limit ordinal, take address of Union
315 --
316 -- od01 : HOD
317 -- od01 = record { od = record { def = λ z → odef A z ∧ ( z ≡ & x ) } ; odmax = & A ; <odmax = odef∧< }
318 -- m01 : OSUP A x ψ
319 -- m01 with is-o∅ (& od01 )
320 -- ... | case1 t=0 = record { sup-o' = prevjjk
321 -- ... | case2 t>0 = ?
322
323 -- Open supreme upper bound leads a contradition, so we use domain restriction on sup
324 ¬open-sup : ( sup-o : (Ordinal → Ordinal ) → Ordinal) → ((ψ : Ordinal → Ordinal ) → (x : Ordinal) → ψ x o< sup-o ψ ) → ⊥
325 ¬open-sup sup-o sup-o< = o<> <-osuc (sup-o< next-ord (sup-o next-ord)) where
326 next-ord : Ordinal → Ordinal
327 next-ord x = osuc x
328
329 Select : (X : HOD ) → ((x : HOD ) → Set n ) → HOD
302 Select X ψ = record { od = record { def = λ x → ( odef X x ∧ ψ ( * x )) } ; odmax = odmax X ; <odmax = λ y → <odmax X (proj1 y) } 330 Select X ψ = record { od = record { def = λ x → ( odef X x ∧ ψ ( * x )) } ; odmax = odmax X ; <odmax = λ y → <odmax X (proj1 y) }
303 331
304 Replace : HOD → (HOD → HOD) → HOD 332 Replace : HOD → (HOD → HOD) → HOD
305 Replace X ψ = record { od = record { def = λ x → (x o≤ sup-o X (λ y X∋y → & (ψ (* y)))) ∧ def (in-codomain X ψ) x } 333 Replace X ψ = record { od = record { def = λ x → (x o≤ sup-o X (λ y X∋y → & (ψ (* y)))) ∧ def (in-codomain X ψ) x }
306 ; odmax = rmax ; <odmax = rmax<} where 334 ; odmax = rmax ; <odmax = rmax<} where
307 rmax : Ordinal 335 rmax : Ordinal
308 rmax = osuc ( sup-o X (λ y X∋y → & (ψ (* y)))) 336 rmax = osuc ( sup-o X (λ y X∋y → & (ψ (* y))))
309 rmax< : {y : Ordinal} → (y o< rmax) ∧ def (in-codomain X ψ) y → y o< rmax 337 rmax< : {y : Ordinal} → (y o< rmax) ∧ def (in-codomain X ψ) y → y o< rmax
310 rmax< lt = proj1 lt 338 rmax< lt = proj1 lt
311 339
312 -- 340 --
313 -- If we have LEM, Replace' is equivalent to Replace 341 -- If we have LEM, Replace' is equivalent to Replace
314 -- 342 --
315 in-codomain' : (X : HOD ) → ((x : HOD) → X ∋ x → HOD) → OD 343 in-codomain' : (X : HOD ) → ((x : HOD) → X ∋ x → HOD) → OD
316 in-codomain' X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( odef X y ∧ ((lt : odef X y) → x ≡ & (ψ (* y ) (d→∋ X lt) )))) } 344 in-codomain' X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( odef X y ∧ ((lt : odef X y) → x ≡ & (ψ (* y ) (d→∋ X lt) )))) }
317 Replace' : (X : HOD) → ((x : HOD) → X ∋ x → HOD) → HOD 345 Replace' : (X : HOD) → ((x : HOD) → X ∋ x → HOD) → HOD
318 Replace' X ψ = record { od = record { def = λ x → (x o< sup-o X (λ y X∋y → & (ψ (* y) (d→∋ X X∋y) ))) ∧ def (in-codomain' X ψ) x } 346 Replace' X ψ = record { od = record { def = λ x → (x o< sup-o X (λ y X∋y → & (ψ (* y) (d→∋ X X∋y) ))) ∧ def (in-codomain' X ψ) x }
319 ; odmax = rmax ; <odmax = rmax< } where 347 ; odmax = rmax ; <odmax = rmax< } where
320 rmax : Ordinal 348 rmax : Ordinal
321 rmax = sup-o X (λ y X∋y → & (ψ (* y) (d→∋ X X∋y))) 349 rmax = sup-o X (λ y X∋y → & (ψ (* y) (d→∋ X X∋y)))
322 rmax< : {y : Ordinal} → (y o< rmax) ∧ def (in-codomain' X ψ) y → y o< rmax 350 rmax< : {y : Ordinal} → (y o< rmax) ∧ def (in-codomain' X ψ) y → y o< rmax
323 rmax< lt = proj1 lt 351 rmax< lt = proj1 lt
324 352
325 Union : HOD → HOD 353 Union : HOD → HOD
326 Union U = record { od = record { def = λ x → ¬ (∀ (u : Ordinal ) → ¬ ((odef U u) ∧ (odef (* u) x))) } 354 Union U = record { od = record { def = λ x → ¬ (∀ (u : Ordinal ) → ¬ ((odef U u) ∧ (odef (* u) x))) }
327 ; odmax = osuc (& U) ; <odmax = umax< } where 355 ; odmax = osuc (& U) ; <odmax = umax< } where
328 umax< : {y : Ordinal} → ¬ ((u : Ordinal) → ¬ def (od U) u ∧ def (od (* u)) y) → y o< osuc (& U) 356 umax< : {y : Ordinal} → ¬ ((u : Ordinal) → ¬ def (od U) u ∧ def (od (* u)) y) → y o< osuc (& U)
329 umax< {y} not = lemma (FExists _ lemma1 not ) where 357 umax< {y} not = lemma (FExists _ lemma1 not ) where
330 lemma0 : {x : Ordinal} → def (od (* x)) y → y o< x 358 lemma0 : {x : Ordinal} → def (od (* x)) y → y o< x
331 lemma0 {x} x<y = subst₂ (λ j k → j o< k ) &iso &iso (c<→o< (d→∋ (* x) x<y )) 359 lemma0 {x} x<y = subst₂ (λ j k → j o< k ) &iso &iso (c<→o< (d→∋ (* x) x<y ))
332 lemma2 : {x : Ordinal} → def (od U) x → x o< & U 360 lemma2 : {x : Ordinal} → def (od U) x → x o< & U
333 lemma2 {x} x<U = subst (λ k → k o< & U ) &iso (c<→o< (d→∋ U x<U)) 361 lemma2 {x} x<U = subst (λ k → k o< & U ) &iso (c<→o< (d→∋ U x<U))
334 lemma1 : {x : Ordinal} → def (od U) x ∧ def (od (* x)) y → ¬ (& U o< y) 362 lemma1 : {x : Ordinal} → def (od U) x ∧ def (od (* x)) y → ¬ (& U o< y)
335 lemma1 {x} lt u<y = o<> u<y (ordtrans (lemma0 (proj2 lt)) (lemma2 (proj1 lt)) ) 363 lemma1 {x} lt u<y = o<> u<y (ordtrans (lemma0 (proj2 lt)) (lemma2 (proj1 lt)) )
336 lemma : ¬ ((& U) o< y ) → y o< osuc (& U) 364 lemma : ¬ ((& U) o< y ) → y o< osuc (& U)
337 lemma not with trio< y (& U) 365 lemma not with trio< y (& U)
338 lemma not | tri< a ¬b ¬c = ordtrans a <-osuc 366 lemma not | tri< a ¬b ¬c = ordtrans a <-osuc
339 lemma not | tri≈ ¬a refl ¬c = <-osuc 367 lemma not | tri≈ ¬a refl ¬c = <-osuc
340 lemma not | tri> ¬a ¬b c = ⊥-elim (not c) 368 lemma not | tri> ¬a ¬b c = ⊥-elim (not c)
341 _∈_ : ( A B : HOD ) → Set n 369 _∈_ : ( A B : HOD ) → Set n
342 A ∈ B = B ∋ A 370 A ∈ B = B ∋ A
343 371
344 OPwr : (A : HOD ) → HOD 372 OPwr : (A : HOD ) → HOD
345 OPwr A = Ord ( osuc ( sup-o (Ord (osuc (& A))) ( λ x A∋x → & ( A ∩ (* x)) ) ) ) 373 OPwr A = Ord ( osuc ( sup-o (Ord (osuc (& A))) ( λ x A∋x → & ( A ∩ (* x)) ) ) )
346 374
347 Power : HOD → HOD 375 Power : HOD → HOD
348 Power A = Replace (OPwr (Ord (& A))) ( λ x → A ∩ x ) 376 Power A = Replace (OPwr (Ord (& A))) ( λ x → A ∩ x )
349 -- {_} : ZFSet → ZFSet 377 -- {_} : ZFSet → ZFSet
350 -- { x } = ( x , x ) -- better to use (x , x) directly 378 -- { x } = ( x , x ) -- better to use (x , x) directly
351 379
352 union→ : (X z u : HOD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z 380 union→ : (X z u : HOD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z
362 isuc : {x : Ordinal } → infinite-d x → 390 isuc : {x : Ordinal } → infinite-d x →
363 infinite-d (& ( Union (* x , (* x , * x ) ) )) 391 infinite-d (& ( Union (* x , (* x , * x ) ) ))
364 392
365 -- ω can be diverged in our case, since we have no restriction on the corresponding ordinal of a pair. 393 -- ω can be diverged in our case, since we have no restriction on the corresponding ordinal of a pair.
366 -- We simply assumes infinite-d y has a maximum. 394 -- We simply assumes infinite-d y has a maximum.
367 -- 395 --
368 -- This means that many of OD may not be HODs because of the & mapping divergence. 396 -- This means that many of OD may not be HODs because of the & mapping divergence.
369 -- We should have some axioms to prevent this such as & x o< next (odmax x). 397 -- We should have some axioms to prevent this such as & x o< next (odmax x).
370 -- 398 --
371 -- postulate 399 -- postulate
372 -- ωmax : Ordinal 400 -- ωmax : Ordinal
373 -- <ωmax : {y : Ordinal} → infinite-d y → y o< ωmax 401 -- <ωmax : {y : Ordinal} → infinite-d y → y o< ωmax
374 -- 402 --
375 -- infinite : HOD 403 -- infinite : HOD
376 -- infinite = record { od = record { def = λ x → infinite-d x } ; odmax = ωmax ; <odmax = <ωmax } 404 -- infinite = record { od = record { def = λ x → infinite-d x } ; odmax = ωmax ; <odmax = <ωmax }
377 405
378 infinite : HOD 406 infinite : HOD
379 infinite = record { od = record { def = λ x → infinite-d x } ; odmax = next o∅ ; <odmax = lemma } where 407 infinite = record { od = record { def = λ x → infinite-d x } ; odmax = next o∅ ; <odmax = lemma } where
380 u : (y : Ordinal ) → HOD 408 u : (y : Ordinal ) → HOD
381 u y = Union (* y , (* y , * y)) 409 u y = Union (* y , (* y , * y))
382 -- next< : {x y z : Ordinal} → x o< next z → y o< next x → y o< next z 410 -- next< : {x y z : Ordinal} → x o< next z → y o< next x → y o< next z
383 lemma8 : {y : Ordinal} → & (* y , * y) o< next (odmax (* y , * y)) 411 lemma8 : {y : Ordinal} → & (* y , * y) o< next (odmax (* y , * y))
384 lemma8 = ho< 412 lemma8 = ho<
385 --- (x,y) < next (omax x y) < next (osuc y) = next y 413 --- (x,y) < next (omax x y) < next (osuc y) = next y
386 lemmaa : {x y : HOD} → & x o< & y → & (x , y) o< next (& y) 414 lemmaa : {x y : HOD} → & x o< & y → & (x , y) o< next (& y)
387 lemmaa {x} {y} x<y = subst (λ k → & (x , y) o< k ) (sym nexto≡) (subst (λ k → & (x , y) o< next k ) (sym (omax< _ _ x<y)) ho< ) 415 lemmaa {x} {y} x<y = subst (λ k → & (x , y) o< k ) (sym nexto≡) (subst (λ k → & (x , y) o< next k ) (sym (omax< _ _ x<y)) ho< )
388 lemma81 : {y : Ordinal} → & (* y , * y) o< next (& (* y)) 416 lemma81 : {y : Ordinal} → & (* y , * y) o< next (& (* y))
389 lemma81 {y} = nexto=n (subst (λ k → & (* y , * y) o< k ) (cong (λ k → next k) (omxx _)) lemma8) 417 lemma81 {y} = nexto=n (subst (λ k → & (* y , * y) o< k ) (cong (λ k → next k) (omxx _)) lemma8)
390 lemma9 : {y : Ordinal} → & (* y , (* y , * y)) o< next (& (* y , * y)) 418 lemma9 : {y : Ordinal} → & (* y , (* y , * y)) o< next (& (* y , * y))
397 lemma : {y : Ordinal} → infinite-d y → y o< next o∅ 425 lemma : {y : Ordinal} → infinite-d y → y o< next o∅
398 lemma {o∅} iφ = x<nx 426 lemma {o∅} iφ = x<nx
399 lemma (isuc {y} x) = next< (lemma x) (next< (subst (λ k → & (* y , (* y , * y)) o< next k) &iso lemma71 ) (nexto=n lemma1)) 427 lemma (isuc {y} x) = next< (lemma x) (next< (subst (λ k → & (* y , (* y , * y)) o< next k) &iso lemma71 ) (nexto=n lemma1))
400 428
401 empty : (x : HOD ) → ¬ (od∅ ∋ x) 429 empty : (x : HOD ) → ¬ (od∅ ∋ x)
402 empty x = ¬x<0 430 empty x = ¬x<0
403 431
404 _=h=_ : (x y : HOD) → Set n 432 _=h=_ : (x y : HOD) → Set n
405 x =h= y = od x == od y 433 x =h= y = od x == od y
406 434
407 pair→ : ( x y t : HOD ) → (x , y) ∋ t → ( t =h= x ) ∨ ( t =h= y ) 435 pair→ : ( x y t : HOD ) → (x , y) ∋ t → ( t =h= x ) ∨ ( t =h= y )
408 pair→ x y t (case1 t≡x ) = case1 (subst₂ (λ j k → j =h= k ) *iso *iso (o≡→== t≡x )) 436 pair→ x y t (case1 t≡x ) = case1 (subst₂ (λ j k → j =h= k ) *iso *iso (o≡→== t≡x ))
409 pair→ x y t (case2 t≡y ) = case2 (subst₂ (λ j k → j =h= k ) *iso *iso (o≡→== t≡y )) 437 pair→ x y t (case2 t≡y ) = case2 (subst₂ (λ j k → j =h= k ) *iso *iso (o≡→== t≡y ))
410 438
411 pair← : ( x y t : HOD ) → ( t =h= x ) ∨ ( t =h= y ) → (x , y) ∋ t 439 pair← : ( x y t : HOD ) → ( t =h= x ) ∨ ( t =h= y ) → (x , y) ∋ t
412 pair← x y t (case1 t=h=x) = case1 (cong (λ k → & k ) (==→o≡ t=h=x)) 440 pair← x y t (case1 t=h=x) = case1 (cong (λ k → & k ) (==→o≡ t=h=x))
413 pair← x y t (case2 t=h=y) = case2 (cong (λ k → & k ) (==→o≡ t=h=y)) 441 pair← x y t (case2 t=h=y) = case2 (cong (λ k → & k ) (==→o≡ t=h=y))
414 442
415 o<→c< : {x y : Ordinal } → x o< y → (Ord x) ⊆ (Ord y) 443 o<→c< : {x y : Ordinal } → x o< y → (Ord x) ⊆ (Ord y)
416 o<→c< lt = record { incl = λ z → ordtrans z lt } 444 o<→c< lt = record { incl = λ z → ordtrans z lt }
417 445
418 ⊆→o< : {x y : Ordinal } → (Ord x) ⊆ (Ord y) → x o< osuc y 446 ⊆→o< : {x y : Ordinal } → (Ord x) ⊆ (Ord y) → x o< osuc y
419 ⊆→o< {x} {y} lt with trio< x y 447 ⊆→o< {x} {y} lt with trio< x y
420 ⊆→o< {x} {y} lt | tri< a ¬b ¬c = ordtrans a <-osuc 448 ⊆→o< {x} {y} lt | tri< a ¬b ¬c = ordtrans a <-osuc
421 ⊆→o< {x} {y} lt | tri≈ ¬a b ¬c = subst ( λ k → k o< osuc y) (sym b) <-osuc 449 ⊆→o< {x} {y} lt | tri≈ ¬a b ¬c = subst ( λ k → k o< osuc y) (sym b) <-osuc
422 ⊆→o< {x} {y} lt | tri> ¬a ¬b c with (incl lt) (o<-subst c (sym &iso) refl ) 450 ⊆→o< {x} {y} lt | tri> ¬a ¬b c with (incl lt) (o<-subst c (sym &iso) refl )
423 ... | ttt = ⊥-elim ( o<¬≡ refl (o<-subst ttt &iso refl )) 451 ... | ttt = ⊥-elim ( o<¬≡ refl (o<-subst ttt &iso refl ))
424 452
428 selection {ψ} {X} {y} = ⟪ 456 selection {ψ} {X} {y} = ⟪
429 ( λ cond → ⟪ proj1 cond , ψiso {ψ} (proj2 cond) (sym *iso) ⟫ ) 457 ( λ cond → ⟪ proj1 cond , ψiso {ψ} (proj2 cond) (sym *iso) ⟫ )
430 , ( λ select → ⟪ proj1 select , ψiso {ψ} (proj2 select) *iso ⟫ ) 458 , ( λ select → ⟪ proj1 select , ψiso {ψ} (proj2 select) *iso ⟫ )
431 459
432 460
433 selection-in-domain : {ψ : HOD → Set n} {X y : HOD} → Select X ψ ∋ y → X ∋ y 461 selection-in-domain : {ψ : HOD → Set n} {X y : HOD} → Select X ψ ∋ y → X ∋ y
434 selection-in-domain {ψ} {X} {y} lt = proj1 ((proj2 (selection {ψ} {X} )) lt) 462 selection-in-domain {ψ} {X} {y} lt = proj1 ((proj2 (selection {ψ} {X} )) lt)
435 463
436 sup-c≤ : (ψ : HOD → HOD) → {X x : HOD} → X ∋ x → & (ψ x) o≤ (sup-o X (λ y X∋y → & (ψ (* y)))) 464 sup-c≤ : (ψ : HOD → HOD) → {X x : HOD} → X ∋ x → & (ψ x) o≤ (sup-o X (λ y X∋y → & (ψ (* y))))
437 sup-c≤ ψ {X} {x} lt = subst (λ k → & (ψ k) o< _ ) *iso (sup-o≤ X lt ) 465 sup-c≤ ψ {X} {x} lt = subst (λ k → & (ψ k) o< _ ) *iso (sup-o≤ X lt )
438 466
439 replacement← : {ψ : HOD → HOD} (X x : HOD) → X ∋ x → Replace X ψ ∋ ψ x 467 replacement← : {ψ : HOD → HOD} (X x : HOD) → X ∋ x → Replace X ψ ∋ ψ x
440 replacement← {ψ} X x lt = ⟪ sup-c≤ ψ {X} {x} lt , lemma ⟫ where 468 replacement← {ψ} X x lt = ⟪ sup-c≤ ψ {X} {x} lt , lemma ⟫ where
441 lemma : def (in-codomain X ψ) (& (ψ x)) 469 lemma : def (in-codomain X ψ) (& (ψ x))
442 lemma not = ⊥-elim ( not ( & x ) ⟪ lt , cong (λ k → & (ψ k)) (sym *iso)⟫ ) 470 lemma not = ⊥-elim ( not ( & x ) ⟪ lt , cong (λ k → & (ψ k)) (sym *iso)⟫ )
443 replacement→ : {ψ : HOD → HOD} (X x : HOD) → (lt : Replace X ψ ∋ x) → ¬ ( (y : HOD) → ¬ (x =h= ψ y)) 471 replacement→ : {ψ : HOD → HOD} (X x : HOD) → (lt : Replace X ψ ∋ x) → ¬ ( (y : HOD) → ¬ (x =h= ψ y))
444 replacement→ {ψ} X x lt = contra-position lemma (lemma2 (proj2 lt)) where 472 replacement→ {ψ} X x lt = contra-position lemma (lemma2 (proj2 lt)) where
445 lemma2 : ¬ ((y : Ordinal) → ¬ odef X y ∧ ((& x) ≡ & (ψ (* y)))) 473 lemma2 : ¬ ((y : Ordinal) → ¬ odef X y ∧ ((& x) ≡ & (ψ (* y))))
446 → ¬ ((y : Ordinal) → ¬ odef X y ∧ (* (& x) =h= ψ (* y))) 474 → ¬ ((y : Ordinal) → ¬ odef X y ∧ (* (& x) =h= ψ (* y)))
447 lemma2 not not2 = not ( λ y d → not2 y ⟪ proj1 d , lemma3 (proj2 d)⟫) where 475 lemma2 not not2 = not ( λ y d → not2 y ⟪ proj1 d , lemma3 (proj2 d)⟫) where
448 lemma3 : {y : Ordinal } → (& x ≡ & (ψ (* y))) → (* (& x) =h= ψ (* y)) 476 lemma3 : {y : Ordinal } → (& x ≡ & (ψ (* y))) → (* (& x) =h= ψ (* y))
449 lemma3 {y} eq = subst (λ k → * (& x) =h= k ) *iso (o≡→== eq ) 477 lemma3 {y} eq = subst (λ k → * (& x) =h= k ) *iso (o≡→== eq )
450 lemma : ( (y : HOD) → ¬ (x =h= ψ y)) → ( (y : Ordinal) → ¬ odef X y ∧ (* (& x) =h= ψ (* y)) ) 478 lemma : ( (y : HOD) → ¬ (x =h= ψ y)) → ( (y : Ordinal) → ¬ odef X y ∧ (* (& x) =h= ψ (* y)) )
451 lemma not y not2 = not (* y) (subst (λ k → k =h= ψ (* y)) *iso ( proj2 not2 )) 479 lemma not y not2 = not (* y) (subst (λ k → k =h= ψ (* y)) *iso ( proj2 not2 ))
452 480
453 --- 481 ---
460 -- 488 --
461 ∩-≡ : { a b : HOD } → ({x : HOD } → (a ∋ x → b ∋ x)) → a =h= ( b ∩ a ) 489 ∩-≡ : { a b : HOD } → ({x : HOD } → (a ∋ x → b ∋ x)) → a =h= ( b ∩ a )
462 ∩-≡ {a} {b} inc = record { 490 ∩-≡ {a} {b} inc = record {
463 eq→ = λ {x} x<a → ⟪ (subst (λ k → odef b k ) &iso (inc (d→∋ a x<a))) , x<a ⟫ ; 491 eq→ = λ {x} x<a → ⟪ (subst (λ k → odef b k ) &iso (inc (d→∋ a x<a))) , x<a ⟫ ;
464 eq← = λ {x} x<a∩b → proj2 x<a∩b } 492 eq← = λ {x} x<a∩b → proj2 x<a∩b }
465 -- 493 --
466 -- Transitive Set case 494 -- Transitive Set case
467 -- we have t ∋ x → Ord a ∋ x means t is a subset of Ord a, that is (Ord a) ∩ t =h= t 495 -- we have t ∋ x → Ord a ∋ x means t is a subset of Ord a, that is (Ord a) ∩ t =h= t
468 -- OPwr (Ord a) is a sup of (Ord a) ∩ t, so OPwr (Ord a) ∋ t 496 -- OPwr (Ord a) is a sup of (Ord a) ∩ t, so OPwr (Ord a) ∋ t
469 -- OPwr A = Ord ( sup-o ( λ x → & ( A ∩ (* x )) ) ) 497 -- OPwr A = Ord ( sup-o ( λ x → & ( A ∩ (* x )) ) )
470 -- 498 --
471 ord-power← : (a : Ordinal ) (t : HOD) → ({x : HOD} → (t ∋ x → (Ord a) ∋ x)) → OPwr (Ord a) ∋ t 499 ord-power← : (a : Ordinal ) (t : HOD) → ({x : HOD} → (t ∋ x → (Ord a) ∋ x)) → OPwr (Ord a) ∋ t
472 ord-power← a t t→A = subst (λ k → odef (OPwr (Ord a)) k ) (lemma1 lemma-eq) lemma where 500 ord-power← a t t→A = subst (λ k → odef (OPwr (Ord a)) k ) (lemma1 lemma-eq) lemma where
473 lemma-eq : ((Ord a) ∩ t) =h= t 501 lemma-eq : ((Ord a) ∩ t) =h= t
474 eq→ lemma-eq {z} w = proj2 w 502 eq→ lemma-eq {z} w = proj2 w
475 eq← lemma-eq {z} w = ⟪ subst (λ k → odef (Ord a) k ) &iso ( t→A (d→∋ t w)) , w ⟫ 503 eq← lemma-eq {z} w = ⟪ subst (λ k → odef (Ord a) k ) &iso ( t→A (d→∋ t w)) , w ⟫
476 lemma1 : {a : Ordinal } { t : HOD } 504 lemma1 : {a : Ordinal } { t : HOD }
477 → (eq : ((Ord a) ∩ t) =h= t) → & ((Ord a) ∩ (* (& t))) ≡ & t 505 → (eq : ((Ord a) ∩ t) =h= t) → & ((Ord a) ∩ (* (& t))) ≡ & t
478 lemma1 {a} {t} eq = subst (λ k → & ((Ord a) ∩ k) ≡ & t ) (sym *iso) (cong (λ k → & k ) (==→o≡ eq )) 506 lemma1 {a} {t} eq = subst (λ k → & ((Ord a) ∩ k) ≡ & t ) (sym *iso) (cong (λ k → & k ) (==→o≡ eq ))
479 lemma2 : (& t) o< (osuc (& (Ord a))) 507 lemma2 : (& t) o< (osuc (& (Ord a)))
480 lemma2 = ⊆→o≤ {t} {Ord a} (λ {x} x<t → subst (λ k → def (od (Ord a)) k) &iso (t→A (d→∋ t x<t))) 508 lemma2 = ⊆→o≤ {t} {Ord a} (λ {x} x<t → subst (λ k → def (od (Ord a)) k) &iso (t→A (d→∋ t x<t)))
481 lemma : & ((Ord a) ∩ (* (& t)) ) o≤ sup-o (Ord (osuc (& (Ord a)))) (λ x lt → & ((Ord a) ∩ (* x))) 509 lemma : & ((Ord a) ∩ (* (& t)) ) o≤ sup-o (Ord (osuc (& (Ord a)))) (λ x lt → & ((Ord a) ∩ (* x)))
482 lemma = sup-o≤ _ lemma2 510 lemma = sup-o≤ _ lemma2
483 511
484 -- 512 --
485 -- Every set in HOD is a subset of Ordinals, so make OPwr (Ord (& A)) first 513 -- Every set in HOD is a subset of Ordinals, so make OPwr (Ord (& A)) first
486 -- then replace of all elements of the Power set by A ∩ y 514 -- then replace of all elements of the Power set by A ∩ y
487 -- 515 --
488 -- Power A = Replace (OPwr (Ord (& A))) ( λ y → A ∩ y ) 516 -- Power A = Replace (OPwr (Ord (& A))) ( λ y → A ∩ y )
489 517
490 -- we have oly double negation form because of the replacement axiom 518 -- we have oly double negation form because of the replacement axiom
491 -- 519 --
492 power→ : ( A t : HOD) → Power A ∋ t → {x : HOD} → t ∋ x → ¬ ¬ (A ∋ x) 520 power→ : ( A t : HOD) → Power A ∋ t → {x : HOD} → t ∋ x → ¬ ¬ (A ∋ x)
500 lemma4 not = lemma2 ( λ y not1 → not (& y) (subst (λ k → t =h= ( A ∩ k )) (sym *iso) not1 )) 528 lemma4 not = lemma2 ( λ y not1 → not (& y) (subst (λ k → t =h= ( A ∩ k )) (sym *iso) not1 ))
501 lemma5 : {y : Ordinal} → t =h= (A ∩ * y) → ¬ ¬ (odef A (& x)) 529 lemma5 : {y : Ordinal} → t =h= (A ∩ * y) → ¬ ¬ (odef A (& x))
502 lemma5 {y} eq not = (lemma3 (* y) eq) not 530 lemma5 {y} eq not = (lemma3 (* y) eq) not
503 531
504 power← : (A t : HOD) → ({x : HOD} → (t ∋ x → A ∋ x)) → Power A ∋ t 532 power← : (A t : HOD) → ({x : HOD} → (t ∋ x → A ∋ x)) → Power A ∋ t
505 power← A t t→A = ⟪ lemma1 , lemma2 ⟫ where 533 power← A t t→A = ⟪ lemma1 , lemma2 ⟫ where
506 a = & A 534 a = & A
507 lemma0 : {x : HOD} → t ∋ x → Ord a ∋ x 535 lemma0 : {x : HOD} → t ∋ x → Ord a ∋ x
508 lemma0 {x} t∋x = c<→o< (t→A t∋x) 536 lemma0 {x} t∋x = c<→o< (t→A t∋x)
509 lemma3 : OPwr (Ord a) ∋ t 537 lemma3 : OPwr (Ord a) ∋ t
510 lemma3 = ord-power← a t lemma0 538 lemma3 = ord-power← a t lemma0
517 t 545 t
518 546
519 sup1 : Ordinal 547 sup1 : Ordinal
520 sup1 = sup-o (Ord (osuc (& (Ord (& A))))) (λ x A∋x → & ((Ord (& A)) ∩ (* x))) 548 sup1 = sup-o (Ord (osuc (& (Ord (& A))))) (λ x A∋x → & ((Ord (& A)) ∩ (* x)))
521 lemma9 : def (od (Ord (Ordinals.osuc O (& (Ord (& A)))))) (& (Ord (& A))) 549 lemma9 : def (od (Ord (Ordinals.osuc O (& (Ord (& A)))))) (& (Ord (& A)))
522 lemma9 = <-osuc 550 lemma9 = <-osuc
523 lemmab : & ((Ord (& A)) ∩ (* (& (Ord (& A) )))) o≤ sup1 551 lemmab : & ((Ord (& A)) ∩ (* (& (Ord (& A) )))) o≤ sup1
524 lemmab = sup-o≤ (Ord (osuc (& (Ord (& A))))) lemma9 552 lemmab = sup-o≤ (Ord (osuc (& (Ord (& A))))) lemma9
525 lemmad : Ord (osuc (& A)) ∋ t 553 lemmad : Ord (osuc (& A)) ∋ t
526 lemmad = ⊆→o≤ (λ {x} lt → subst (λ k → def (od A) k ) &iso (t→A (d→∋ t lt))) 554 lemmad = ⊆→o≤ (λ {x} lt → subst (λ k → def (od A) k ) &iso (t→A (d→∋ t lt)))
527 lemmac : ((Ord (& A)) ∩ (* (& (Ord (& A) )))) =h= Ord (& A) 555 lemmac : ((Ord (& A)) ∩ (* (& (Ord (& A) )))) =h= Ord (& A)
528 lemmac = record { eq→ = lemmaf ; eq← = lemmag } where 556 lemmac = record { eq→ = lemmaf ; eq← = lemmag } where
529 lemmaf : {x : Ordinal} → def (od ((Ord (& A)) ∩ (* (& (Ord (& A)))))) x → def (od (Ord (& A))) x 557 lemmaf : {x : Ordinal} → def (od ((Ord (& A)) ∩ (* (& (Ord (& A)))))) x → def (od (Ord (& A))) x
530 lemmaf {x} lt = proj1 lt 558 lemmaf {x} lt = proj1 lt
531 lemmag : {x : Ordinal} → def (od (Ord (& A))) x → def (od ((Ord (& A)) ∩ (* (& (Ord (& A)))))) x 559 lemmag : {x : Ordinal} → def (od (Ord (& A))) x → def (od ((Ord (& A)) ∩ (* (& (Ord (& A)))))) x
532 lemmag {x} lt = ⟪ lt , subst (λ k → def (od k) x) (sym *iso) lt ⟫ 560 lemmag {x} lt = ⟪ lt , subst (λ k → def (od k) x) (sym *iso) lt ⟫
533 lemmae : & ((Ord (& A)) ∩ (* (& (Ord (& A))))) ≡ & (Ord (& A)) 561 lemmae : & ((Ord (& A)) ∩ (* (& (Ord (& A))))) ≡ & (Ord (& A))
534 lemmae = cong (λ k → & k ) ( ==→o≡ lemmac) 562 lemmae = cong (λ k → & k ) ( ==→o≡ lemmac)
535 lemma7 : def (od (OPwr (Ord (& A)))) (& t) 563 lemma7 : def (od (OPwr (Ord (& A)))) (& t)
536 lemma7 with osuc-≡< lemmad 564 lemma7 with osuc-≡< lemmad
537 lemma7 | case2 lt = ordtrans (c<→o< lt) (subst (λ k → k o≤ sup1) lemmae lemmab ) 565 lemma7 | case2 lt = ordtrans (c<→o< lt) (subst (λ k → k o≤ sup1) lemmae lemmab )
538 lemma7 | case1 eq with osuc-≡< (⊆→o≤ {* (& t)} {* (& (Ord (& t)))} (λ {x} lt → lemmah lt )) where 566 lemma7 | case1 eq with osuc-≡< (⊆→o≤ {* (& t)} {* (& (Ord (& t)))} (λ {x} lt → lemmah lt )) where
539 lemmah : {x : Ordinal } → def (od (* (& t))) x → def (od (* (& (Ord (& t))))) x 567 lemmah : {x : Ordinal } → def (od (* (& t))) x → def (od (* (& (Ord (& t))))) x
540 lemmah {x} lt = subst (λ k → def (od k) x ) (sym *iso) (subst (λ k → k o< (& t)) 568 lemmah {x} lt = subst (λ k → def (od k) x ) (sym *iso) (subst (λ k → k o< (& t))
541 &iso 569 &iso
542 (c<→o< (subst₂ (λ j k → def (od j) k) *iso (sym &iso) lt ))) 570 (c<→o< (subst₂ (λ j k → def (od j) k) *iso (sym &iso) lt )))
543 lemma7 | case1 eq | case1 eq1 = subst (λ k → k o≤ sup1) (trans lemmae lemmai) lemmab where 571 lemma7 | case1 eq | case1 eq1 = subst (λ k → k o≤ sup1) (trans lemmae lemmai) lemmab where
544 lemmai : & (Ord (& A)) ≡ & t 572 lemmai : & (Ord (& A)) ≡ & t
545 lemmai = let open ≡-Reasoning in begin 573 lemmai = let open ≡-Reasoning in begin
546 & (Ord (& A)) 574 & (Ord (& A))
547 ≡⟨ sym (cong (λ k → & (Ord k)) eq) ⟩ 575 ≡⟨ sym (cong (λ k → & (Ord k)) eq) ⟩
548 & (Ord (& t)) 576 & (Ord (& t))
549 ≡⟨ sym &iso ⟩ 577 ≡⟨ sym &iso ⟩
550 & (* (& (Ord (& t)))) 578 & (* (& (Ord (& t))))
551 ≡⟨ sym eq1 ⟩ 579 ≡⟨ sym eq1 ⟩
552 & (* (& t)) 580 & (* (& t))
553 ≡⟨ &iso ⟩ 581 ≡⟨ &iso ⟩
554 & t 582 & t
555 583
556 lemma7 | case1 eq | case2 lt = ordtrans lemmaj (subst (λ k → k o≤ sup1) lemmae lemmab ) where 584 lemma7 | case1 eq | case2 lt = ordtrans lemmaj (subst (λ k → k o≤ sup1) lemmae lemmab ) where
557 lemmak : & (* (& (Ord (& t)))) ≡ & (Ord (& A)) 585 lemmak : & (* (& (Ord (& t)))) ≡ & (Ord (& A))
558 lemmak = let open ≡-Reasoning in begin 586 lemmak = let open ≡-Reasoning in begin
559 & (* (& (Ord (& t)))) 587 & (* (& (Ord (& t))))
561 & (Ord (& t)) 589 & (Ord (& t))
562 ≡⟨ cong (λ k → & (Ord k)) eq ⟩ 590 ≡⟨ cong (λ k → & (Ord k)) eq ⟩
563 & (Ord (& A)) 591 & (Ord (& A))
564 592
565 lemmaj : & t o< & (Ord (& A)) 593 lemmaj : & t o< & (Ord (& A))
566 lemmaj = subst₂ (λ j k → j o< k ) &iso lemmak lt 594 lemmaj = subst₂ (λ j k → j o< k ) &iso lemmak lt
567 lemma1 : & t o≤ sup-o (OPwr (Ord (& A))) (λ x lt → & (A ∩ (* x))) 595 lemma1 : & t o≤ sup-o (OPwr (Ord (& A))) (λ x lt → & (A ∩ (* x)))
568 lemma1 = subst (λ k → & k o≤ sup-o (OPwr (Ord (& A))) (λ x lt → & (A ∩ (* x)))) 596 lemma1 = subst (λ k → & k o≤ sup-o (OPwr (Ord (& A))) (λ x lt → & (A ∩ (* x))))
569 lemma4 (sup-o≤ (OPwr (Ord (& A))) lemma7 ) 597 lemma4 (sup-o≤ (OPwr (Ord (& A))) lemma7 )
570 lemma2 : def (in-codomain (OPwr (Ord (& A))) (_∩_ A)) (& t) 598 lemma2 : def (in-codomain (OPwr (Ord (& A))) (_∩_ A)) (& t)
571 lemma2 not = ⊥-elim ( not (& t) ⟪ lemma3 , lemma6 ⟫ ) where 599 lemma2 not = ⊥-elim ( not (& t) ⟪ lemma3 , lemma6 ⟫ ) where
572 lemma6 : & t ≡ & (A ∩ * (& t)) 600 lemma6 : & t ≡ & (A ∩ * (& t))
573 lemma6 = cong ( λ k → & k ) (==→o≡ (subst (λ k → t =h= (A ∩ k)) (sym *iso) ( ∩-≡ {t} {A} t→A ))) 601 lemma6 = cong ( λ k → & k ) (==→o≡ (subst (λ k → t =h= (A ∩ k)) (sym *iso) ( ∩-≡ {t} {A} t→A )))
574 602
575 603
576 extensionality0 : {A B : HOD } → ((z : HOD) → (A ∋ z) ⇔ (B ∋ z)) → A =h= B 604 extensionality0 : {A B : HOD } → ((z : HOD) → (A ∋ z) ⇔ (B ∋ z)) → A =h= B
577 eq→ (extensionality0 {A} {B} eq ) {x} d = odef-iso {A} {B} (sym &iso) (proj1 (eq (* x))) d 605 eq→ (extensionality0 {A} {B} eq ) {x} d = odef-iso {A} {B} (sym &iso) (proj1 (eq (* x))) d
578 eq← (extensionality0 {A} {B} eq ) {x} d = odef-iso {B} {A} (sym &iso) (proj2 (eq (* x))) d 606 eq← (extensionality0 {A} {B} eq ) {x} d = odef-iso {B} {A} (sym &iso) (proj2 (eq (* x))) d
579 607
580 extensionality : {A B w : HOD } → ((z : HOD ) → (A ∋ z) ⇔ (B ∋ z)) → (w ∋ A) ⇔ (w ∋ B) 608 extensionality : {A B w : HOD } → ((z : HOD ) → (A ∋ z) ⇔ (B ∋ z)) → (w ∋ A) ⇔ (w ∋ B)
581 proj1 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) ( ==→o≡ (extensionality0 {A} {B} eq) ) d 609 proj1 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) ( ==→o≡ (extensionality0 {A} {B} eq) ) d
582 proj2 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) (sym ( ==→o≡ (extensionality0 {A} {B} eq) )) d 610 proj2 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) (sym ( ==→o≡ (extensionality0 {A} {B} eq) )) d
583 611
584 infinity∅ : infinite ∋ od∅ 612 infinity∅ : infinite ∋ od∅
585 infinity∅ = subst (λ k → odef infinite k ) lemma iφ where 613 infinity∅ = subst (λ k → odef infinite k ) lemma iφ where
586 lemma : o∅ ≡ & od∅ 614 lemma : o∅ ≡ & od∅
587 lemma = let open ≡-Reasoning in begin 615 lemma = let open ≡-Reasoning in begin
588 o∅ 616 o∅
589 ≡⟨ sym &iso ⟩ 617 ≡⟨ sym &iso ⟩
590 & ( * o∅ ) 618 & ( * o∅ )
593 621
594 infinity : (x : HOD) → infinite ∋ x → infinite ∋ Union (x , (x , x )) 622 infinity : (x : HOD) → infinite ∋ x → infinite ∋ Union (x , (x , x ))
595 infinity x lt = subst (λ k → odef infinite k ) lemma (isuc {& x} lt) where 623 infinity x lt = subst (λ k → odef infinite k ) lemma (isuc {& x} lt) where
596 lemma : & (Union (* (& x) , (* (& x) , * (& x)))) 624 lemma : & (Union (* (& x) , (* (& x) , * (& x))))
597 ≡ & (Union (x , (x , x))) 625 ≡ & (Union (x , (x , x)))
598 lemma = cong (λ k → & (Union ( k , ( k , k ) ))) *iso 626 lemma = cong (λ k → & (Union ( k , ( k , k ) ))) *iso
599 627
600 isZF : IsZF (HOD ) _∋_ _=h=_ od∅ _,_ Union Power Select Replace infinite 628 isZF : IsZF (HOD ) _∋_ _=h=_ od∅ _,_ Union Power Select Replace infinite
601 isZF = record { 629 isZF = record {
602 isEquivalence = record { refl = ==-refl ; sym = ==-sym; trans = ==-trans } 630 isEquivalence = record { refl = ==-refl ; sym = ==-sym; trans = ==-trans }
603 ; pair→ = pair→ 631 ; pair→ = pair→
604 ; pair← = pair← 632 ; pair← = pair←
605 ; union→ = union→ 633 ; union→ = union→
606 ; union← = union← 634 ; union← = union←
607 ; empty = empty 635 ; empty = empty
608 ; power→ = power→ 636 ; power→ = power→
609 ; power← = power← 637 ; power← = power←
610 ; extensionality = λ {A} {B} {w} → extensionality {A} {B} {w} 638 ; extensionality = λ {A} {B} {w} → extensionality {A} {B} {w}
611 ; ε-induction = ε-induction 639 ; ε-induction = ε-induction
612 ; infinity∅ = infinity∅ 640 ; infinity∅ = infinity∅
613 ; infinity = infinity 641 ; infinity = infinity
614 ; selection = λ {X} {ψ} {y} → selection {X} {ψ} {y} 642 ; selection = λ {X} {ψ} {y} → selection {X} {ψ} {y}
615 ; replacement← = replacement← 643 ; replacement← = replacement←
616 ; replacement→ = λ {ψ} → replacement→ {ψ} 644 ; replacement→ = λ {ψ} → replacement→ {ψ}
617 } 645 }
618 646
619 HOD→ZF : ZF 647 HOD→ZF : ZF
620 HOD→ZF = record { 648 HOD→ZF = record {
621 ZFSet = HOD 649 ZFSet = HOD
622 ; _∋_ = _∋_ 650 ; _∋_ = _∋_
623 ; _≈_ = _=h=_ 651 ; _≈_ = _=h=_
624 ; ∅ = od∅ 652 ; ∅ = od∅
625 ; _,_ = _,_ 653 ; _,_ = _,_
626 ; Union = Union 654 ; Union = Union
627 ; Power = Power 655 ; Power = Power
628 ; Select = Select 656 ; Select = Select
629 ; Replace = Replace 657 ; Replace = Replace
630 ; infinite = infinite 658 ; infinite = infinite
631 ; isZF = isZF 659 ; isZF = isZF
632 } 660 }
633 661
634 662