Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate OD.agda @ 324:fbabb20f222e
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sat, 04 Jul 2020 18:18:17 +0900 |
parents | e228e96965f0 |
children | 1a54dbe1ea4c |
rev | line source |
---|---|
16 | 1 open import Level |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
2 open import Ordinals |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
3 module OD {n : Level } (O : Ordinals {n} ) where |
3 | 4 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
5 open import zf |
23 | 6 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
7 open import Relation.Binary.PropositionalEquality |
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
8 open import Data.Nat.Properties |
6 | 9 open import Data.Empty |
10 open import Relation.Nullary | |
11 open import Relation.Binary | |
12 open import Relation.Binary.Core | |
13 | |
213
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
14 open import logic |
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
15 open import nat |
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
16 |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
17 open inOrdinal O |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
18 |
27 | 19 -- Ordinal Definable Set |
11 | 20 |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
21 record OD : Set (suc n ) where |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
22 field |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
23 def : (x : Ordinal ) → Set n |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
24 |
141 | 25 open OD |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
26 |
120 | 27 open _∧_ |
213
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
28 open _∨_ |
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
29 open Bool |
44
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
30 |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
31 record _==_ ( a b : OD ) : Set n where |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
32 field |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
33 eq→ : ∀ { x : Ordinal } → def a x → def b x |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
34 eq← : ∀ { x : Ordinal } → def b x → def a x |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
35 |
234
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
228
diff
changeset
|
36 id : {A : Set n} → A → A |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
37 id x = x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
38 |
271 | 39 ==-refl : { x : OD } → x == x |
40 ==-refl {x} = record { eq→ = id ; eq← = id } | |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
41 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
42 open _==_ |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
43 |
271 | 44 ==-trans : { x y z : OD } → x == y → y == z → x == z |
45 ==-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) } | |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
46 |
271 | 47 ==-sym : { x y : OD } → x == y → y == x |
48 ==-sym x=y = record { eq→ = λ {m} t → eq← x=y t ; eq← = λ {m} t → eq→ x=y t } | |
49 | |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
50 |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
51 ⇔→== : { x y : OD } → ( {z : Ordinal } → def x z ⇔ def y z) → x == y |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
52 eq→ ( ⇔→== {x} {y} eq ) {z} m = proj1 eq m |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
53 eq← ( ⇔→== {x} {y} eq ) {z} m = proj2 eq m |
120 | 54 |
277
d9d3654baee1
seperate choice from LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
276
diff
changeset
|
55 -- next assumptions are our axiom |
322 | 56 -- |
57 -- OD is an equation on Ordinals, so it contains Ordinals. If these Ordinals have one-to-one | |
58 -- correspondence to the OD then the OD looks like a ZF Set. | |
59 -- | |
60 -- If all ZF Set have supreme upper bound, the solutions of OD have to be bounded, i.e. | |
61 -- bbounded ODs are ZF Set. Unbounded ODs are classes. | |
62 -- | |
290 | 63 -- In classical Set Theory, HOD is used, as a subset of OD, |
277
d9d3654baee1
seperate choice from LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
276
diff
changeset
|
64 -- HOD = { x | TC x ⊆ OD } |
290 | 65 -- where TC x is a transitive clusure of x, i.e. Union of all elemnts of all subset of x. |
322 | 66 -- This is not possible because we don't have V yet. So we assumes HODs are bounded OD. |
290 | 67 -- |
309 | 68 -- We also assumes HODs are isomorphic to Ordinals, which is ususally proved by Goedel number tricks. |
322 | 69 -- There two contraints on the HOD order, one is ∋, the other one is ⊂. |
70 -- ODs have an ovbious maximum, but Ordinals are not, but HOD has no maximum because there is an aribtrary | |
71 -- bound on each HOD. | |
290 | 72 -- |
322 | 73 -- In classical Set Theory, sup is defined by Uion, since we are working on constructive logic, |
290 | 74 -- we need explict assumption on sup. |
309 | 75 -- |
76 -- ==→o≡ is necessary to prove axiom of extensionality. | |
277
d9d3654baee1
seperate choice from LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
276
diff
changeset
|
77 |
303 | 78 data One : Set n where |
79 OneObj : One | |
80 | |
81 -- Ordinals in OD , the maximum | |
82 Ords : OD | |
83 Ords = record { def = λ x → One } | |
84 | |
85 record HOD : Set (suc n) where | |
302
304c271b3d47
HOD and reduction mapping of Ordinals
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
301
diff
changeset
|
86 field |
303 | 87 od : OD |
304 | 88 odmax : Ordinal |
308 | 89 <odmax : {y : Ordinal} → def od y → y o< odmax |
302
304c271b3d47
HOD and reduction mapping of Ordinals
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
301
diff
changeset
|
90 |
304c271b3d47
HOD and reduction mapping of Ordinals
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
301
diff
changeset
|
91 open HOD |
304c271b3d47
HOD and reduction mapping of Ordinals
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
301
diff
changeset
|
92 |
277
d9d3654baee1
seperate choice from LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
276
diff
changeset
|
93 record ODAxiom : Set (suc n) where |
d9d3654baee1
seperate choice from LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
276
diff
changeset
|
94 field |
304 | 95 -- HOD is isomorphic to Ordinal (by means of Goedel number) |
303 | 96 od→ord : HOD → Ordinal |
97 ord→od : Ordinal → HOD | |
98 c<→o< : {x y : HOD } → def (od y) ( od→ord x ) → od→ord x o< od→ord y | |
312 | 99 ⊆→o≤ : {y z : HOD } → ({x : Ordinal} → def (od y) x → def (od z) x ) → od→ord y o< osuc (od→ord z) |
303 | 100 oiso : {x : HOD } → ord→od ( od→ord x ) ≡ x |
322 | 101 diso : {x : Ordinal } → od→ord ( ord→od x ) ≡ x |
102 ==→o≡ : { x y : HOD } → (od x == od y) → x ≡ y | |
306 | 103 sup-o : (A : HOD) → (( x : Ordinal ) → def (od A) x → Ordinal ) → Ordinal |
104 sup-o< : (A : HOD) → { ψ : ( x : Ordinal ) → def (od A) x → Ordinal } → ∀ {x : Ordinal } → (lt : def (od A) x ) → ψ x lt o< sup-o A ψ | |
302
304c271b3d47
HOD and reduction mapping of Ordinals
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
301
diff
changeset
|
105 |
277
d9d3654baee1
seperate choice from LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
276
diff
changeset
|
106 postulate odAxiom : ODAxiom |
d9d3654baee1
seperate choice from LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
276
diff
changeset
|
107 open ODAxiom odAxiom |
d9d3654baee1
seperate choice from LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
276
diff
changeset
|
108 |
303 | 109 -- maxod : {x : OD} → od→ord x o< od→ord Ords |
110 -- maxod {x} = c<→o< OneObj | |
277
d9d3654baee1
seperate choice from LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
276
diff
changeset
|
111 |
304 | 112 -- we have not this contradiction |
303 | 113 -- bad-bad : ⊥ |
304 | 114 -- bad-bad = osuc-< <-osuc (c<→o< { record { od = record { def = λ x → One }; <odmax = {!!} } } OneObj) |
301
b012a915bbb5
contradiction found ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
115 |
179 | 116 -- Ordinal in OD ( and ZFSet ) Transitive Set |
303 | 117 Ord : ( a : Ordinal ) → HOD |
304 | 118 Ord a = record { od = record { def = λ y → y o< a } ; odmax = a ; <odmax = lemma } where |
119 lemma : {x : Ordinal} → x o< a → x o< a | |
120 lemma {x} lt = lt | |
109
dab56d357fa3
remove o<→c< and add otrans in OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
103
diff
changeset
|
121 |
303 | 122 od∅ : HOD |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
123 od∅ = Ord o∅ |
40 | 124 |
303 | 125 odef : HOD → Ordinal → Set n |
126 odef A x = def ( od A ) x | |
123 | 127 |
303 | 128 o<→c<→HOD=Ord : ( {x y : Ordinal } → x o< y → odef (ord→od y) x ) → {x : HOD } → x ≡ Ord (od→ord x) |
129 o<→c<→HOD=Ord o<→c< {x} = ==→o≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where | |
130 lemma1 : {y : Ordinal} → odef x y → odef (Ord (od→ord x)) y | |
131 lemma1 {y} lt = subst ( λ k → k o< od→ord x ) diso (c<→o< {ord→od y} {x} (subst (λ k → odef x k ) (sym diso) lt)) | |
132 lemma2 : {y : Ordinal} → odef (Ord (od→ord x)) y → odef x y | |
133 lemma2 {y} lt = subst (λ k → odef k y ) oiso (o<→c< {y} {od→ord x} lt ) | |
95 | 134 |
324 | 135 |
303 | 136 _∋_ : ( a x : HOD ) → Set n |
137 _∋_ a x = odef a ( od→ord x ) | |
138 | |
139 _c<_ : ( x a : HOD ) → Set n | |
109
dab56d357fa3
remove o<→c< and add otrans in OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
103
diff
changeset
|
140 x c< a = a ∋ x |
103 | 141 |
303 | 142 cseq : {n : Level} → HOD → HOD |
308 | 143 cseq x = record { od = record { def = λ y → odef x (osuc y) } ; odmax = osuc (odmax x) ; <odmax = lemma } where |
144 lemma : {y : Ordinal} → def (od x) (osuc y) → y o< osuc (odmax x) | |
145 lemma {y} lt = ordtrans <-osuc (ordtrans (<odmax x lt) <-osuc ) | |
95 | 146 |
303 | 147 odef-subst : {Z : HOD } {X : Ordinal }{z : HOD } {x : Ordinal }→ odef Z X → Z ≡ z → X ≡ x → odef z x |
148 odef-subst df refl refl = df | |
95 | 149 |
303 | 150 otrans : {n : Level} {a x y : Ordinal } → odef (Ord a) x → odef (Ord x) y → odef (Ord a) y |
187 | 151 otrans x<a y<x = ordtrans y<x x<a |
123 | 152 |
303 | 153 odef→o< : {X : HOD } → {x : Ordinal } → odef X x → x o< od→ord X |
154 odef→o< {X} {x} lt = o<-subst {_} {_} {x} {od→ord X} ( c<→o< ( odef-subst {X} {x} lt (sym oiso) (sym diso) )) diso diso | |
44
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
155 |
51 | 156 -- avoiding lv != Zero error |
303 | 157 orefl : { x : HOD } → { y : Ordinal } → od→ord x ≡ y → od→ord x ≡ y |
51 | 158 orefl refl = refl |
159 | |
303 | 160 ==-iso : { x y : HOD } → od (ord→od (od→ord x)) == od (ord→od (od→ord y)) → od x == od y |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
161 ==-iso {x} {y} eq = record { |
303 | 162 eq→ = λ d → lemma ( eq→ eq (odef-subst d (sym oiso) refl )) ; |
163 eq← = λ d → lemma ( eq← eq (odef-subst d (sym oiso) refl )) } | |
51 | 164 where |
303 | 165 lemma : {x : HOD } {z : Ordinal } → odef (ord→od (od→ord x)) z → odef x z |
166 lemma {x} {z} d = odef-subst d oiso refl | |
51 | 167 |
303 | 168 =-iso : {x y : HOD } → (od x == od y) ≡ (od (ord→od (od→ord x)) == od y) |
169 =-iso {_} {y} = cong ( λ k → od k == od y ) (sym oiso) | |
57 | 170 |
303 | 171 ord→== : { x y : HOD } → od→ord x ≡ od→ord y → od x == od y |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
172 ord→== {x} {y} eq = ==-iso (lemma (od→ord x) (od→ord y) (orefl eq)) where |
303 | 173 lemma : ( ox oy : Ordinal ) → ox ≡ oy → od (ord→od ox) == od (ord→od oy) |
271 | 174 lemma ox ox refl = ==-refl |
51 | 175 |
303 | 176 o≡→== : { x y : Ordinal } → x ≡ y → od (ord→od x) == od (ord→od y) |
271 | 177 o≡→== {x} {.x} refl = ==-refl |
51 | 178 |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
179 o∅≡od∅ : ord→od (o∅ ) ≡ od∅ |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
180 o∅≡od∅ = ==→o≡ lemma where |
303 | 181 lemma0 : {x : Ordinal} → odef (ord→od o∅) x → odef od∅ x |
182 lemma0 {x} lt = o<-subst (c<→o< {ord→od x} {ord→od o∅} (odef-subst {ord→od o∅} {x} lt refl (sym diso)) ) diso diso | |
183 lemma1 : {x : Ordinal} → odef od∅ x → odef (ord→od o∅) x | |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
184 lemma1 {x} lt = ⊥-elim (¬x<0 lt) |
303 | 185 lemma : od (ord→od o∅) == od od∅ |
150 | 186 lemma = record { eq→ = lemma0 ; eq← = lemma1 } |
187 | |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
188 ord-od∅ : od→ord (od∅ ) ≡ o∅ |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
189 ord-od∅ = sym ( subst (λ k → k ≡ od→ord (od∅ ) ) diso (cong ( λ k → od→ord k ) o∅≡od∅ ) ) |
80 | 190 |
303 | 191 ∅0 : record { def = λ x → Lift n ⊥ } == od od∅ |
109
dab56d357fa3
remove o<→c< and add otrans in OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
103
diff
changeset
|
192 eq→ ∅0 {w} (lift ()) |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
193 eq← ∅0 {w} lt = lift (¬x<0 lt) |
109
dab56d357fa3
remove o<→c< and add otrans in OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
103
diff
changeset
|
194 |
303 | 195 ∅< : { x y : HOD } → odef x (od→ord y ) → ¬ ( od x == od od∅ ) |
271 | 196 ∅< {x} {y} d eq with eq→ (==-trans eq (==-sym ∅0) ) d |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
197 ∅< {x} {y} d eq | lift () |
57 | 198 |
303 | 199 ∅6 : { x : HOD } → ¬ ( x ∋ x ) -- no Russel paradox |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
200 ∅6 {x} x∋x = o<¬≡ refl ( c<→o< {x} {x} x∋x ) |
51 | 201 |
303 | 202 odef-iso : {A B : HOD } {x y : Ordinal } → x ≡ y → (odef A y → odef B y) → odef A x → odef B x |
203 odef-iso refl t = t | |
76 | 204 |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
205 is-o∅ : ( x : Ordinal ) → Dec ( x ≡ o∅ ) |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
206 is-o∅ x with trio< x o∅ |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
207 is-o∅ x | tri< a ¬b ¬c = no ¬b |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
208 is-o∅ x | tri≈ ¬a b ¬c = yes b |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
209 is-o∅ x | tri> ¬a ¬b c = no ¬b |
57 | 210 |
303 | 211 _,_ : HOD → HOD → HOD |
308 | 212 x , y = record { od = record { def = λ t → (t ≡ od→ord x ) ∨ ( t ≡ od→ord y ) } ; odmax = omax (od→ord x) (od→ord y) ; <odmax = lemma } where |
213 lemma : {t : Ordinal} → (t ≡ od→ord x) ∨ (t ≡ od→ord y) → t o< omax (od→ord x) (od→ord y) | |
214 lemma {t} (case1 refl) = omax-x _ _ | |
215 lemma {t} (case2 refl) = omax-y _ _ | |
216 | |
188
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
187
diff
changeset
|
217 |
79 | 218 -- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
219 -- postulate f-extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality n (suc n) |
59
d13d1351a1fa
lemma = cong₂ (λ x not → minimul x not ) oiso { }6
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
58
diff
changeset
|
220 |
318 | 221 in-codomain : (X : HOD ) → ( ψ : HOD → HOD ) → OD |
222 in-codomain X ψ = record { def = λ x → ¬ ( (y : Ordinal ) → ¬ ( odef X y ∧ ( x ≡ od→ord (ψ (ord→od y ))))) } | |
141 | 223 |
303 | 224 -- Power Set of X ( or constructible by λ y → odef X (od→ord y ) |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
225 |
303 | 226 ZFSubset : (A x : HOD ) → HOD |
308 | 227 ZFSubset A x = record { od = record { def = λ y → odef A y ∧ odef x y } ; odmax = omin (odmax A) (odmax x) ; <odmax = lemma } where -- roughly x = A → Set |
228 lemma : {y : Ordinal} → def (od A) y ∧ def (od x) y → y o< omin (odmax A) (odmax x) | |
229 lemma {y} and = min1 (<odmax A (proj1 and)) (<odmax x (proj2 and)) | |
230 | |
303 | 231 record _⊆_ ( A B : HOD ) : Set (suc n) where |
271 | 232 field |
303 | 233 incl : { x : HOD } → A ∋ x → B ∋ x |
271 | 234 |
235 open _⊆_ | |
190 | 236 infixr 220 _⊆_ |
237 | |
303 | 238 subset-lemma : {A x : HOD } → ( {y : HOD } → x ∋ y → ZFSubset A x ∋ y ) ⇔ ( x ⊆ A ) |
271 | 239 subset-lemma {A} {x} = record { |
240 proj1 = λ lt → record { incl = λ x∋z → proj1 (lt x∋z) } | |
241 ; proj2 = λ x⊆A lt → record { proj1 = incl x⊆A lt ; proj2 = lt } | |
190 | 242 } |
243 | |
324 | 244 od⊆→o≤ : {x y : HOD } → x ⊆ y → od→ord x o< osuc (od→ord y) |
245 od⊆→o≤ {x} {y} lt = ⊆→o≤ {x} {y} (λ {z} x>z → subst (λ k → def (od y) k ) diso (incl lt (subst (λ k → def (od x) k ) (sym diso) x>z ))) | |
246 | |
312 | 247 power< : {A x : HOD } → x ⊆ A → Ord (osuc (od→ord A)) ∋ x |
248 power< {A} {x} x⊆A = ⊆→o≤ (λ {y} x∋y → subst (λ k → def (od A) k) diso (lemma y x∋y ) ) where | |
249 lemma : (y : Ordinal) → def (od x) y → def (od A) (od→ord (ord→od y)) | |
250 lemma y x∋y = incl x⊆A (subst (λ k → def (od x) k ) (sym diso) x∋y ) | |
251 | |
261
d9d178d1457c
ε-induction from TransFinite induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
260
diff
changeset
|
252 open import Data.Unit |
d9d178d1457c
ε-induction from TransFinite induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
260
diff
changeset
|
253 |
324 | 254 ε-induction : { ψ : HOD → Set n} |
303 | 255 → ( {x : HOD } → ({ y : HOD } → x ∋ y → ψ y ) → ψ x ) |
256 → (x : HOD ) → ψ x | |
261
d9d178d1457c
ε-induction from TransFinite induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
260
diff
changeset
|
257 ε-induction {ψ} ind x = subst (λ k → ψ k ) oiso (ε-induction-ord (osuc (od→ord x)) <-osuc ) where |
d9d178d1457c
ε-induction from TransFinite induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
260
diff
changeset
|
258 induction : (ox : Ordinal) → ((oy : Ordinal) → oy o< ox → ψ (ord→od oy)) → ψ (ord→od ox) |
d9d178d1457c
ε-induction from TransFinite induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
260
diff
changeset
|
259 induction ox prev = ind ( λ {y} lt → subst (λ k → ψ k ) oiso (prev (od→ord y) (o<-subst (c<→o< lt) refl diso ))) |
d9d178d1457c
ε-induction from TransFinite induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
260
diff
changeset
|
260 ε-induction-ord : (ox : Ordinal) { oy : Ordinal } → oy o< ox → ψ (ord→od oy) |
d9d178d1457c
ε-induction from TransFinite induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
260
diff
changeset
|
261 ε-induction-ord ox {oy} lt = TransFinite {λ oy → ψ (ord→od oy)} induction oy |
d9d178d1457c
ε-induction from TransFinite induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
260
diff
changeset
|
262 |
303 | 263 HOD→ZF : ZF |
264 HOD→ZF = record { | |
265 ZFSet = HOD | |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
266 ; _∋_ = _∋_ |
303 | 267 ; _≈_ = _=h=_ |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
268 ; ∅ = od∅ |
28 | 269 ; _,_ = _,_ |
270 ; Union = Union | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
271 ; Power = Power |
28 | 272 ; Select = Select |
273 ; Replace = Replace | |
161 | 274 ; infinite = infinite |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
275 ; isZF = isZF |
28 | 276 } where |
303 | 277 ZFSet = HOD -- is less than Ords because of maxod |
278 Select : (X : HOD ) → ((x : HOD ) → Set n ) → HOD | |
308 | 279 Select X ψ = record { od = record { def = λ x → ( odef X x ∧ ψ ( ord→od x )) } ; odmax = odmax X ; <odmax = λ y → <odmax X (proj1 y) } |
310 | 280 Replace : HOD → (HOD → HOD) → HOD |
318 | 281 Replace X ψ = record { od = record { def = λ x → (x o< sup-o X (λ y X∋y → od→ord (ψ (ord→od y)))) ∧ def (in-codomain X ψ) x } |
282 ; odmax = rmax ; <odmax = rmax<} where | |
283 rmax : Ordinal | |
284 rmax = sup-o X (λ y X∋y → od→ord (ψ (ord→od y))) | |
285 rmax< : {y : Ordinal} → (y o< rmax) ∧ def (in-codomain X ψ) y → y o< rmax | |
286 rmax< lt = proj1 lt | |
144 | 287 _∩_ : ( A B : ZFSet ) → ZFSet |
318 | 288 A ∩ B = record { od = record { def = λ x → odef A x ∧ odef B x } |
289 ; odmax = omin (odmax A) (odmax B) ; <odmax = λ y → min1 (<odmax A (proj1 y)) (<odmax B (proj2 y))} | |
303 | 290 Union : HOD → HOD |
318 | 291 Union U = record { od = record { def = λ x → ¬ (∀ (u : Ordinal ) → ¬ ((odef U u) ∧ (odef (ord→od u) x))) } |
292 ; odmax = osuc (od→ord U) ; <odmax = umax< } where | |
293 umax< : {y : Ordinal} → ¬ ((u : Ordinal) → ¬ def (od U) u ∧ def (od (ord→od u)) y) → y o< osuc (od→ord U) | |
319 | 294 umax< {y} not = lemma (FExists _ lemma1 not ) where |
295 lemma0 : {x : Ordinal} → def (od (ord→od x)) y → y o< x | |
296 lemma0 {x} x<y = subst₂ (λ j k → j o< k ) diso diso (c<→o< (subst (λ k → def (od (ord→od x)) k) (sym diso) x<y)) | |
297 lemma2 : {x : Ordinal} → def (od U) x → x o< od→ord U | |
298 lemma2 {x} x<U = subst (λ k → k o< od→ord U ) diso (c<→o< (subst (λ k → def (od U) k) (sym diso) x<U)) | |
299 lemma1 : {x : Ordinal} → def (od U) x ∧ def (od (ord→od x)) y → ¬ (od→ord U o< y) | |
300 lemma1 {x} lt u<y = o<> u<y (ordtrans (lemma0 (proj2 lt)) (lemma2 (proj1 lt)) ) | |
301 lemma : ¬ ((od→ord U) o< y ) → y o< osuc (od→ord U) | |
302 lemma not with trio< y (od→ord U) | |
303 lemma not | tri< a ¬b ¬c = ordtrans a <-osuc | |
304 lemma not | tri≈ ¬a refl ¬c = <-osuc | |
305 lemma not | tri> ¬a ¬b c = ⊥-elim (not c) | |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
306 _∈_ : ( A B : ZFSet ) → Set n |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
307 A ∈ B = B ∋ A |
312 | 308 |
309 OPwr : (A : HOD ) → HOD | |
310 OPwr A = Ord ( sup-o (Ord (osuc (od→ord A))) ( λ x A∋x → od→ord ( ZFSubset A (ord→od x)) ) ) | |
311 | |
303 | 312 Power : HOD → HOD |
300
e70980bd80c7
-- the set of finite partial functions from ω to 2
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
291
diff
changeset
|
313 Power A = Replace (OPwr (Ord (od→ord A))) ( λ x → A ∩ x ) |
277
d9d3654baee1
seperate choice from LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
276
diff
changeset
|
314 -- {_} : ZFSet → ZFSet |
287 | 315 -- { x } = ( x , x ) -- it works but we don't use |
109
dab56d357fa3
remove o<→c< and add otrans in OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
103
diff
changeset
|
316 |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
317 data infinite-d : ( x : Ordinal ) → Set n where |
161 | 318 iφ : infinite-d o∅ |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
319 isuc : {x : Ordinal } → infinite-d x → |
161 | 320 infinite-d (od→ord ( Union (ord→od x , (ord→od x , ord→od x ) ) )) |
321 | |
303 | 322 infinite : HOD |
324 | 323 infinite = record { od = record { def = λ x → infinite-d x } ; odmax = next o∅ ; <odmax = lemma } where |
324 u : HOD → HOD | |
325 u x = Union (x , (x , x)) | |
326 lemma1 : {x : HOD} → u x ⊆ Union (u x , (u x , u x)) | |
327 lemma1 {x} = record { incl = λ {y} lt → lemma2 y lt } where | |
328 lemma2 : (y : HOD) → u x ∋ y → ((z : Ordinal) → (z ≡ od→ord (u x)) ∨ (z ≡ od→ord (u x , u x)) ∧ def (od (ord→od z)) (od→ord y) → ⊥) → ⊥ | |
329 lemma2 y lt not = not (od→ord (u x)) record { proj1 = case1 refl ; proj2 = subst (λ k → def (od k) (od→ord y) ) (sym oiso) lt } | |
330 lemma3 : {x : HOD} → od→ord (u x) o< osuc (od→ord ( Union (u x , (u x , u x)) )) | |
331 lemma3 {x} = od⊆→o≤ lemma1 | |
332 lemma : {y : Ordinal} → infinite-d y → y o< next o∅ | |
333 lemma {y} = TransFinite {λ x → infinite-d x → x o< next o∅ } ind y where | |
334 ind : (x : Ordinal) → ((z : Ordinal) → z o< x → infinite-d z → z o< (next o∅)) → | |
335 infinite-d x → x o< (next o∅) | |
336 ind o∅ prev iφ = proj1 next-limit | |
337 ind x prev (isuc lt) = lemma0 where | |
338 lemma0 : {x : Ordinal} → od→ord (Union (ord→od x , (ord→od x , ord→od x))) o< next o∅ | |
339 lemma0 {x} = {!!} | |
303 | 340 |
341 _=h=_ : (x y : HOD) → Set n | |
342 x =h= y = od x == od y | |
161 | 343 |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
344 infixr 200 _∈_ |
96 | 345 -- infixr 230 _∩_ _∪_ |
303 | 346 isZF : IsZF (HOD ) _∋_ _=h=_ od∅ _,_ Union Power Select Replace infinite |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
347 isZF = record { |
271 | 348 isEquivalence = record { refl = ==-refl ; sym = ==-sym; trans = ==-trans } |
247 | 349 ; pair→ = pair→ |
350 ; pair← = pair← | |
72 | 351 ; union→ = union→ |
352 ; union← = union← | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
353 ; empty = empty |
129 | 354 ; power→ = power→ |
76 | 355 ; power← = power← |
186 | 356 ; extensionality = λ {A} {B} {w} → extensionality {A} {B} {w} |
274 | 357 ; ε-induction = ε-induction |
78
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
358 ; infinity∅ = infinity∅ |
160 | 359 ; infinity = infinity |
116 | 360 ; selection = λ {X} {ψ} {y} → selection {X} {ψ} {y} |
135 | 361 ; replacement← = replacement← |
317 | 362 ; replacement→ = λ {ψ} → replacement→ {ψ} |
274 | 363 -- ; choice-func = choice-func |
364 -- ; choice = choice | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
365 } where |
129 | 366 |
303 | 367 pair→ : ( x y t : ZFSet ) → (x , y) ∋ t → ( t =h= x ) ∨ ( t =h= y ) |
368 pair→ x y t (case1 t≡x ) = case1 (subst₂ (λ j k → j =h= k ) oiso oiso (o≡→== t≡x )) | |
369 pair→ x y t (case2 t≡y ) = case2 (subst₂ (λ j k → j =h= k ) oiso oiso (o≡→== t≡y )) | |
247 | 370 |
303 | 371 pair← : ( x y t : ZFSet ) → ( t =h= x ) ∨ ( t =h= y ) → (x , y) ∋ t |
372 pair← x y t (case1 t=h=x) = case1 (cong (λ k → od→ord k ) (==→o≡ t=h=x)) | |
373 pair← x y t (case2 t=h=y) = case2 (cong (λ k → od→ord k ) (==→o≡ t=h=y)) | |
247 | 374 |
303 | 375 empty : (x : HOD ) → ¬ (od∅ ∋ x) |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
376 empty x = ¬x<0 |
129 | 377 |
271 | 378 o<→c< : {x y : Ordinal } → x o< y → (Ord x) ⊆ (Ord y) |
379 o<→c< lt = record { incl = λ z → ordtrans z lt } | |
155 | 380 |
271 | 381 ⊆→o< : {x y : Ordinal } → (Ord x) ⊆ (Ord y) → x o< osuc y |
155 | 382 ⊆→o< {x} {y} lt with trio< x y |
383 ⊆→o< {x} {y} lt | tri< a ¬b ¬c = ordtrans a <-osuc | |
384 ⊆→o< {x} {y} lt | tri≈ ¬a b ¬c = subst ( λ k → k o< osuc y) (sym b) <-osuc | |
271 | 385 ⊆→o< {x} {y} lt | tri> ¬a ¬b c with (incl lt) (o<-subst c (sym diso) refl ) |
155 | 386 ... | ttt = ⊥-elim ( o<¬≡ refl (o<-subst ttt diso refl )) |
151 | 387 |
303 | 388 union→ : (X z u : HOD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z |
157
afc030b7c8d0
explict logical definition of Union failed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
156
diff
changeset
|
389 union→ X z u xx not = ⊥-elim ( not (od→ord u) ( record { proj1 = proj1 xx |
303 | 390 ; proj2 = subst ( λ k → odef k (od→ord z)) (sym oiso) (proj2 xx) } )) |
391 union← : (X z : HOD) (X∋z : Union X ∋ z) → ¬ ( (u : HOD ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) | |
258 | 392 union← X z UX∋z = FExists _ lemma UX∋z where |
303 | 393 lemma : {y : Ordinal} → odef X y ∧ odef (ord→od y) (od→ord z) → ¬ ((u : HOD) → ¬ (X ∋ u) ∧ (u ∋ z)) |
394 lemma {y} xx not = not (ord→od y) record { proj1 = subst ( λ k → odef X k ) (sym diso) (proj1 xx ) ; proj2 = proj2 xx } | |
144 | 395 |
303 | 396 ψiso : {ψ : HOD → Set n} {x y : HOD } → ψ x → x ≡ y → ψ y |
144 | 397 ψiso {ψ} t refl = t |
303 | 398 selection : {ψ : HOD → Set n} {X y : HOD} → ((X ∋ y) ∧ ψ y) ⇔ (Select X ψ ∋ y) |
144 | 399 selection {ψ} {X} {y} = record { |
400 proj1 = λ cond → record { proj1 = proj1 cond ; proj2 = ψiso {ψ} (proj2 cond) (sym oiso) } | |
401 ; proj2 = λ select → record { proj1 = proj1 select ; proj2 = ψiso {ψ} (proj2 select) oiso } | |
402 } | |
311 | 403 sup-c< : (ψ : HOD → HOD) → {X x : HOD} → X ∋ x → od→ord (ψ x) o< (sup-o X (λ y X∋y → od→ord (ψ (ord→od y)))) |
404 sup-c< ψ {X} {x} lt = subst (λ k → od→ord (ψ k) o< _ ) oiso (sup-o< X lt ) | |
303 | 405 replacement← : {ψ : HOD → HOD} (X x : HOD) → X ∋ x → Replace X ψ ∋ ψ x |
311 | 406 replacement← {ψ} X x lt = record { proj1 = sup-c< ψ {X} {x} lt ; proj2 = lemma } where |
318 | 407 lemma : def (in-codomain X ψ) (od→ord (ψ x)) |
150 | 408 lemma not = ⊥-elim ( not ( od→ord x ) (record { proj1 = lt ; proj2 = cong (λ k → od→ord (ψ k)) (sym oiso)} )) |
303 | 409 replacement→ : {ψ : HOD → HOD} (X x : HOD) → (lt : Replace X ψ ∋ x) → ¬ ( (y : HOD) → ¬ (x =h= ψ y)) |
150 | 410 replacement→ {ψ} X x lt = contra-position lemma (lemma2 (proj2 lt)) where |
303 | 411 lemma2 : ¬ ((y : Ordinal) → ¬ odef X y ∧ ((od→ord x) ≡ od→ord (ψ (ord→od y)))) |
412 → ¬ ((y : Ordinal) → ¬ odef X y ∧ (ord→od (od→ord x) =h= ψ (ord→od y))) | |
144 | 413 lemma2 not not2 = not ( λ y d → not2 y (record { proj1 = proj1 d ; proj2 = lemma3 (proj2 d)})) where |
303 | 414 lemma3 : {y : Ordinal } → (od→ord x ≡ od→ord (ψ (ord→od y))) → (ord→od (od→ord x) =h= ψ (ord→od y)) |
415 lemma3 {y} eq = subst (λ k → ord→od (od→ord x) =h= k ) oiso (o≡→== eq ) | |
416 lemma : ( (y : HOD) → ¬ (x =h= ψ y)) → ( (y : Ordinal) → ¬ odef X y ∧ (ord→od (od→ord x) =h= ψ (ord→od y)) ) | |
417 lemma not y not2 = not (ord→od y) (subst (λ k → k =h= ψ (ord→od y)) oiso ( proj2 not2 )) | |
144 | 418 |
419 --- | |
420 --- Power Set | |
421 --- | |
303 | 422 --- First consider ordinals in HOD |
100 | 423 --- |
303 | 424 --- ZFSubset A x = record { def = λ y → odef A y ∧ odef x y } subset of A |
100 | 425 -- |
426 -- | |
303 | 427 ∩-≡ : { a b : HOD } → ({x : HOD } → (a ∋ x → b ∋ x)) → a =h= ( b ∩ a ) |
142 | 428 ∩-≡ {a} {b} inc = record { |
429 eq→ = λ {x} x<a → record { proj2 = x<a ; | |
303 | 430 proj1 = odef-subst {_} {_} {b} {x} (inc (odef-subst {_} {_} {a} {_} x<a refl (sym diso) )) refl diso } ; |
142 | 431 eq← = λ {x} x<a∩b → proj2 x<a∩b } |
100 | 432 -- |
258 | 433 -- Transitive Set case |
303 | 434 -- we have t ∋ x → Ord a ∋ x means t is a subset of Ord a, that is ZFSubset (Ord a) t =h= t |
300
e70980bd80c7
-- the set of finite partial functions from ω to 2
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
291
diff
changeset
|
435 -- OPwr (Ord a) is a sup of ZFSubset (Ord a) t, so OPwr (Ord a) ∋ t |
e70980bd80c7
-- the set of finite partial functions from ω to 2
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
291
diff
changeset
|
436 -- OPwr A = Ord ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) ) |
100 | 437 -- |
303 | 438 ord-power← : (a : Ordinal ) (t : HOD) → ({x : HOD} → (t ∋ x → (Ord a) ∋ x)) → OPwr (Ord a) ∋ t |
439 ord-power← a t t→A = odef-subst {_} {_} {OPwr (Ord a)} {od→ord t} | |
127 | 440 lemma refl (lemma1 lemma-eq )where |
303 | 441 lemma-eq : ZFSubset (Ord a) t =h= t |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
442 eq→ lemma-eq {z} w = proj2 w |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
443 eq← lemma-eq {z} w = record { proj2 = w ; |
303 | 444 proj1 = odef-subst {_} {_} {(Ord a)} {z} |
445 ( t→A (odef-subst {_} {_} {t} {od→ord (ord→od z)} w refl (sym diso) )) refl diso } | |
446 lemma1 : {a : Ordinal } { t : HOD } | |
447 → (eq : ZFSubset (Ord a) t =h= t) → od→ord (ZFSubset (Ord a) (ord→od (od→ord t))) ≡ od→ord t | |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
448 lemma1 {a} {t} eq = subst (λ k → od→ord (ZFSubset (Ord a) k) ≡ od→ord t ) (sym oiso) (cong (λ k → od→ord k ) (==→o≡ eq )) |
312 | 449 lemma2 : (od→ord t) o< (osuc (od→ord (Ord a))) |
450 lemma2 = ⊆→o≤ {t} {Ord a} (λ {x} x<t → subst (λ k → def (od (Ord a)) k) diso (t→A (subst (λ k → def (od t) k) (sym diso) x<t))) | |
451 lemma : od→ord (ZFSubset (Ord a) (ord→od (od→ord t)) ) o< sup-o (Ord (osuc (od→ord (Ord a)))) (λ x lt → od→ord (ZFSubset (Ord a) (ord→od x))) | |
311 | 452 lemma = sup-o< _ lemma2 |
129 | 453 |
144 | 454 -- |
303 | 455 -- Every set in HOD is a subset of Ordinals, so make OPwr (Ord (od→ord A)) first |
258 | 456 -- then replace of all elements of the Power set by A ∩ y |
144 | 457 -- |
300
e70980bd80c7
-- the set of finite partial functions from ω to 2
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
291
diff
changeset
|
458 -- Power A = Replace (OPwr (Ord (od→ord A))) ( λ y → A ∩ y ) |
166 | 459 |
460 -- we have oly double negation form because of the replacement axiom | |
461 -- | |
303 | 462 power→ : ( A t : HOD) → Power A ∋ t → {x : HOD} → t ∋ x → ¬ ¬ (A ∋ x) |
258 | 463 power→ A t P∋t {x} t∋x = FExists _ lemma5 lemma4 where |
142 | 464 a = od→ord A |
303 | 465 lemma2 : ¬ ( (y : HOD) → ¬ (t =h= (A ∩ y))) |
317 | 466 lemma2 = replacement→ {λ x → A ∩ x} (OPwr (Ord (od→ord A))) t P∋t |
303 | 467 lemma3 : (y : HOD) → t =h= ( A ∩ y ) → ¬ ¬ (A ∋ x) |
166 | 468 lemma3 y eq not = not (proj1 (eq→ eq t∋x)) |
303 | 469 lemma4 : ¬ ((y : Ordinal) → ¬ (t =h= (A ∩ ord→od y))) |
470 lemma4 not = lemma2 ( λ y not1 → not (od→ord y) (subst (λ k → t =h= ( A ∩ k )) (sym oiso) not1 )) | |
471 lemma5 : {y : Ordinal} → t =h= (A ∩ ord→od y) → ¬ ¬ (odef A (od→ord x)) | |
166 | 472 lemma5 {y} eq not = (lemma3 (ord→od y) eq) not |
473 | |
303 | 474 power← : (A t : HOD) → ({x : HOD} → (t ∋ x → A ∋ x)) → Power A ∋ t |
311 | 475 power← A t t→A = record { proj1 = lemma1 ; proj2 = lemma2 } where |
142 | 476 a = od→ord A |
303 | 477 lemma0 : {x : HOD} → t ∋ x → Ord a ∋ x |
142 | 478 lemma0 {x} t∋x = c<→o< (t→A t∋x) |
300
e70980bd80c7
-- the set of finite partial functions from ω to 2
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
291
diff
changeset
|
479 lemma3 : OPwr (Ord a) ∋ t |
142 | 480 lemma3 = ord-power← a t lemma0 |
152 | 481 lemma4 : (A ∩ ord→od (od→ord t)) ≡ t |
482 lemma4 = let open ≡-Reasoning in begin | |
483 A ∩ ord→od (od→ord t) | |
484 ≡⟨ cong (λ k → A ∩ k) oiso ⟩ | |
485 A ∩ t | |
317 | 486 ≡⟨ sym (==→o≡ ( ∩-≡ {t} {A} t→A )) ⟩ |
152 | 487 t |
488 ∎ | |
317 | 489 sup1 : Ordinal |
314 | 490 sup1 = sup-o (Ord (osuc (od→ord (Ord (od→ord A))))) (λ x A∋x → od→ord (ZFSubset (Ord (od→ord A)) (ord→od x))) |
313 | 491 lemma9 : def (od (Ord (Ordinals.osuc O (od→ord (Ord (od→ord A)))))) (od→ord (Ord (od→ord A))) |
492 lemma9 = <-osuc | |
315 | 493 lemmab : od→ord (ZFSubset (Ord (od→ord A)) (ord→od (od→ord (Ord (od→ord A) )))) o< sup1 |
494 lemmab = sup-o< (Ord (osuc (od→ord (Ord (od→ord A))))) lemma9 | |
495 lemmad : Ord (osuc (od→ord A)) ∋ t | |
317 | 496 lemmad = ⊆→o≤ (λ {x} lt → subst (λ k → def (od A) k ) diso (t→A (subst (λ k → def (od t) k ) (sym diso) lt))) |
315 | 497 lemmac : ZFSubset (Ord (od→ord A)) (ord→od (od→ord (Ord (od→ord A) ))) =h= Ord (od→ord A) |
317 | 498 lemmac = record { eq→ = lemmaf ; eq← = lemmag } where |
499 lemmaf : {x : Ordinal} → def (od (ZFSubset (Ord (od→ord A)) (ord→od (od→ord (Ord (od→ord A)))))) x → def (od (Ord (od→ord A))) x | |
500 lemmaf {x} lt = proj1 lt | |
501 lemmag : {x : Ordinal} → def (od (Ord (od→ord A))) x → def (od (ZFSubset (Ord (od→ord A)) (ord→od (od→ord (Ord (od→ord A)))))) x | |
502 lemmag {x} lt = record { proj1 = lt ; proj2 = subst (λ k → def (od k) x) (sym oiso) lt } | |
315 | 503 lemmae : od→ord (ZFSubset (Ord (od→ord A)) (ord→od (od→ord (Ord (od→ord A))))) ≡ od→ord (Ord (od→ord A)) |
504 lemmae = cong (λ k → od→ord k ) ( ==→o≡ lemmac) | |
311 | 505 lemma7 : def (od (OPwr (Ord (od→ord A)))) (od→ord t) |
315 | 506 lemma7 with osuc-≡< lemmad |
507 lemma7 | case2 lt = ordtrans (c<→o< lt) (subst (λ k → k o< sup1) lemmae lemmab ) | |
317 | 508 lemma7 | case1 eq with osuc-≡< (⊆→o≤ {ord→od (od→ord t)} {ord→od (od→ord (Ord (od→ord t)))} (λ {x} lt → lemmah lt )) where |
509 lemmah : {x : Ordinal } → def (od (ord→od (od→ord t))) x → def (od (ord→od (od→ord (Ord (od→ord t))))) x | |
510 lemmah {x} lt = subst (λ k → def (od k) x ) (sym oiso) (subst (λ k → k o< (od→ord t)) | |
511 diso | |
512 (c<→o< (subst₂ (λ j k → def (od j) k) oiso (sym diso) lt ))) | |
513 lemma7 | case1 eq | case1 eq1 = subst (λ k → k o< sup1) (trans lemmae lemmai) lemmab where | |
514 lemmai : od→ord (Ord (od→ord A)) ≡ od→ord t | |
515 lemmai = let open ≡-Reasoning in begin | |
516 od→ord (Ord (od→ord A)) | |
517 ≡⟨ sym (cong (λ k → od→ord (Ord k)) eq) ⟩ | |
518 od→ord (Ord (od→ord t)) | |
519 ≡⟨ sym diso ⟩ | |
520 od→ord (ord→od (od→ord (Ord (od→ord t)))) | |
521 ≡⟨ sym eq1 ⟩ | |
522 od→ord (ord→od (od→ord t)) | |
523 ≡⟨ diso ⟩ | |
524 od→ord t | |
525 ∎ | |
526 lemma7 | case1 eq | case2 lt = ordtrans lemmaj (subst (λ k → k o< sup1) lemmae lemmab ) where | |
527 lemmak : od→ord (ord→od (od→ord (Ord (od→ord t)))) ≡ od→ord (Ord (od→ord A)) | |
528 lemmak = let open ≡-Reasoning in begin | |
529 od→ord (ord→od (od→ord (Ord (od→ord t)))) | |
530 ≡⟨ diso ⟩ | |
531 od→ord (Ord (od→ord t)) | |
532 ≡⟨ cong (λ k → od→ord (Ord k)) eq ⟩ | |
533 od→ord (Ord (od→ord A)) | |
534 ∎ | |
535 lemmaj : od→ord t o< od→ord (Ord (od→ord A)) | |
536 lemmaj = subst₂ (λ j k → j o< k ) diso lemmak lt | |
310 | 537 lemma1 : od→ord t o< sup-o (OPwr (Ord (od→ord A))) (λ x lt → od→ord (A ∩ (ord→od x))) |
538 lemma1 = subst (λ k → od→ord k o< sup-o (OPwr (Ord (od→ord A))) (λ x lt → od→ord (A ∩ (ord→od x)))) | |
311 | 539 lemma4 (sup-o< (OPwr (Ord (od→ord A))) lemma7 ) |
318 | 540 lemma2 : def (in-codomain (OPwr (Ord (od→ord A))) (_∩_ A)) (od→ord t) |
151 | 541 lemma2 not = ⊥-elim ( not (od→ord t) (record { proj1 = lemma3 ; proj2 = lemma6 }) ) where |
542 lemma6 : od→ord t ≡ od→ord (A ∩ ord→od (od→ord t)) | |
317 | 543 lemma6 = cong ( λ k → od→ord k ) (==→o≡ (subst (λ k → t =h= (A ∩ k)) (sym oiso) ( ∩-≡ {t} {A} t→A ))) |
142 | 544 |
311 | 545 |
271 | 546 ord⊆power : (a : Ordinal) → (Ord (osuc a)) ⊆ (Power (Ord a)) |
547 ord⊆power a = record { incl = λ {x} lt → power← (Ord a) x (lemma lt) } where | |
303 | 548 lemma : {x y : HOD} → od→ord x o< osuc a → x ∋ y → Ord a ∋ y |
271 | 549 lemma lt y<x with osuc-≡< lt |
550 lemma lt y<x | case1 refl = c<→o< y<x | |
551 lemma lt y<x | case2 x<a = ordtrans (c<→o< y<x) x<a | |
262 | 552 |
276 | 553 continuum-hyphotheis : (a : Ordinal) → Set (suc n) |
554 continuum-hyphotheis a = Power (Ord a) ⊆ Ord (osuc a) | |
129 | 555 |
303 | 556 extensionality0 : {A B : HOD } → ((z : HOD) → (A ∋ z) ⇔ (B ∋ z)) → A =h= B |
557 eq→ (extensionality0 {A} {B} eq ) {x} d = odef-iso {A} {B} (sym diso) (proj1 (eq (ord→od x))) d | |
558 eq← (extensionality0 {A} {B} eq ) {x} d = odef-iso {B} {A} (sym diso) (proj2 (eq (ord→od x))) d | |
186 | 559 |
303 | 560 extensionality : {A B w : HOD } → ((z : HOD ) → (A ∋ z) ⇔ (B ∋ z)) → (w ∋ A) ⇔ (w ∋ B) |
186 | 561 proj1 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) ( ==→o≡ (extensionality0 {A} {B} eq) ) d |
562 proj2 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) (sym ( ==→o≡ (extensionality0 {A} {B} eq) )) d | |
129 | 563 |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
219
diff
changeset
|
564 infinity∅ : infinite ∋ od∅ |
303 | 565 infinity∅ = odef-subst {_} {_} {infinite} {od→ord (od∅ )} iφ refl lemma where |
161 | 566 lemma : o∅ ≡ od→ord od∅ |
567 lemma = let open ≡-Reasoning in begin | |
568 o∅ | |
569 ≡⟨ sym diso ⟩ | |
570 od→ord ( ord→od o∅ ) | |
571 ≡⟨ cong ( λ k → od→ord k ) o∅≡od∅ ⟩ | |
572 od→ord od∅ | |
573 ∎ | |
303 | 574 infinity : (x : HOD) → infinite ∋ x → infinite ∋ Union (x , (x , x )) |
575 infinity x lt = odef-subst {_} {_} {infinite} {od→ord (Union (x , (x , x )))} ( isuc {od→ord x} lt ) refl lemma where | |
161 | 576 lemma : od→ord (Union (ord→od (od→ord x) , (ord→od (od→ord x) , ord→od (od→ord x)))) |
577 ≡ od→ord (Union (x , (x , x))) | |
578 lemma = cong (λ k → od→ord (Union ( k , ( k , k ) ))) oiso | |
579 | |
234
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
228
diff
changeset
|
580 |
303 | 581 Union = ZF.Union HOD→ZF |
582 Power = ZF.Power HOD→ZF | |
583 Select = ZF.Select HOD→ZF | |
584 Replace = ZF.Replace HOD→ZF | |
585 isZF = ZF.isZF HOD→ZF |