Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate cardinal.agda @ 238:a8c6239176db
ZFProduct
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Mon, 19 Aug 2019 11:39:46 +0900 |
parents | 521290e85527 |
children | b6d80eec5f9e |
rev | line source |
---|---|
16 | 1 open import Level |
224 | 2 open import Ordinals |
3 module cardinal {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 |
219 | 6 open import logic |
224 | 7 import OD |
23 | 8 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) |
224 | 9 open import Relation.Binary.PropositionalEquality |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
10 open import Data.Nat.Properties |
6 | 11 open import Data.Empty |
12 open import Relation.Nullary | |
13 open import Relation.Binary | |
14 open import Relation.Binary.Core | |
15 | |
224 | 16 open inOrdinal O |
17 open OD O | |
219 | 18 open OD.OD |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
19 |
120 | 20 open _∧_ |
213
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
21 open _∨_ |
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
22 open Bool |
44
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
23 |
230 | 24 -- we have to work on Ordinal to keep OD Level n |
25 -- since we use p∨¬p which works only on Level n | |
225 | 26 |
233 | 27 <_,_> : (x y : OD) → OD |
28 < x , y > = (x , x ) , (x , y ) | |
226
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
29 |
238 | 30 data ord-pair : (p : Ordinal) → Set n where |
31 pair : (x y : Ordinal ) → ord-pair ( od→ord ( < ord→od x , ord→od y > ) ) | |
32 | |
33 ZFProduct : OD | |
34 ZFProduct = record { def = λ x → ord-pair x } | |
35 | |
36 π1' : { p : OD } → ZFProduct ∋ p → OD | |
37 π1' lt = ord→od (pi1 lt) where | |
38 pi1 : { p : Ordinal } → ord-pair p → Ordinal | |
39 pi1 ( pair x y ) = x | |
237 | 40 |
238 | 41 π2' : { p : OD } → ZFProduct ∋ p → OD |
42 π2' lt = ord→od (pi2 lt) where | |
43 pi2 : { p : Ordinal } → ord-pair p → Ordinal | |
44 pi2 ( pair x y ) = y | |
237 | 45 |
238 | 46 p-cons : { x y : OD } → ZFProduct ∋ < x , y > |
47 p-cons {x} {y} = def-subst {_} {_} {ZFProduct} {od→ord (< x , y >)} (pair (od→ord x) ( od→ord y )) refl ( | |
48 let open ≡-Reasoning in begin | |
49 od→ord < ord→od (od→ord x) , ord→od (od→ord y) > | |
50 ≡⟨ cong₂ (λ j k → od→ord < j , k >) oiso oiso ⟩ | |
51 od→ord < x , y > | |
52 ∎ ) | |
53 | |
54 | |
237 | 55 |
236 | 56 record SetProduct ( A B : OD ) : Set n where |
226
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
57 field |
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
58 π1 : Ordinal |
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
59 π2 : Ordinal |
233 | 60 A∋π1 : def A π1 |
61 B∋π2 : def B π2 | |
62 -- opair : x ≡ od→ord (Ord ( omax (omax π1 π1) (omax π1 π2) )) -- < π1 , π2 > | |
63 | |
64 open SetProduct | |
226
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
65 |
234
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
66 ∋-p : (A x : OD ) → Dec ( A ∋ x ) |
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
67 ∋-p A x with p∨¬p ( A ∋ x ) |
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
68 ∋-p A x | case1 t = yes t |
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
69 ∋-p A x | case2 t = no t |
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
70 |
233 | 71 _⊗_ : (A B : OD) → OD |
236 | 72 A ⊗ B = record { def = λ x → SetProduct A B } |
233 | 73 -- A ⊗ B = record { def = λ x → (y z : Ordinal) → def A y ∧ def B z ∧ ( x ≡ od→ord (< ord→od y , ord→od z >) ) } |
74 | |
75 -- Power (Power ( A ∪ B )) ∋ ( A ⊗ B ) | |
225 | 76 |
233 | 77 Func : ( A B : OD ) → OD |
78 Func A B = record { def = λ x → def (Power (A ⊗ B)) x } | |
79 | |
80 -- power→ : ( A t : OD) → Power A ∋ t → {x : OD} → t ∋ x → ¬ ¬ (A ∋ x) | |
226
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
81 |
234
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
82 func←od : { dom cod : OD } → {f : Ordinal } → def (Func dom cod ) f → (Ordinal → Ordinal ) |
233 | 83 func←od {dom} {cod} {f} lt x = sup-o ( λ y → lemma y ) where |
226
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
84 lemma : Ordinal → Ordinal |
235 | 85 lemma y with IsZF.power→ isZF (dom ⊗ cod) (ord→od f) (subst (λ k → def (Power (dom ⊗ cod)) k ) (sym diso) lt ) | ∋-p (ord→od f) (ord→od y) |
234
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
86 lemma y | p | no n = o∅ |
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
87 lemma y | p | yes f∋y with double-neg-eilm ( p {ord→od y} f∋y ) -- p : {x : OD} → f ∋ x → ¬ ¬ (dom ⊗ cod ∋ x) |
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
88 ... | t with decp ( x ≡ π1 t ) |
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
89 ... | yes _ = π2 t |
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
90 ... | no _ = o∅ |
233 | 91 |
236 | 92 |
233 | 93 func→od : (f : Ordinal → Ordinal ) → ( dom : OD ) → OD |
94 func→od f dom = Replace dom ( λ x → < x , ord→od (f (od→ord x)) > ) | |
95 | |
236 | 96 record Func←cd { dom cod : OD } {f : Ordinal } (f<F : def (Func dom cod ) f) : Set n where |
97 field | |
98 func-1 : Ordinal → Ordinal | |
99 func→od∈Func-1 : (Func dom (Ord (sup-o (λ x → func-1 x)) )) ∋ func→od func-1 dom | |
100 | |
101 func←od1 : { dom cod : OD } → {f : Ordinal } → (f<F : def (Func dom cod ) f ) → Func←cd {dom} {cod} {f} f<F | |
102 func←od1 {dom} {cod} {f} lt = record { func-1 = λ x → sup-o ( λ y → lemma x y ) ; func→od∈Func-1 = {!!} } where | |
103 lemma : Ordinal → Ordinal → Ordinal | |
104 lemma x y with IsZF.power→ isZF (dom ⊗ cod) (ord→od f) (subst (λ k → def (Power (dom ⊗ cod)) k ) (sym diso) lt ) | ∋-p (ord→od f) (ord→od y) | |
105 lemma x y | p | no n = o∅ | |
106 lemma x y | p | yes f∋y with double-neg-eilm ( p {ord→od y} f∋y ) -- p : {x : OD} → f ∋ x → ¬ ¬ (dom ⊗ cod ∋ x) | |
107 ... | t with decp ( x ≡ π1 t ) | |
108 ... | yes _ = π2 t | |
109 ... | no _ = o∅ | |
110 | |
111 func→od∈Func : (f : Ordinal → Ordinal ) ( dom : OD ) → (Func dom (Ord (sup-o (λ x → f x)) )) ∋ func→od f dom | |
112 func→od∈Func f dom = record { proj1 = {!!} ; proj2 = {!!} } | |
225 | 113 |
227 | 114 -- contra position of sup-o< |
115 -- | |
116 | |
235 | 117 -- postulate |
118 -- -- contra-position of mimimulity of supermum required in Cardinal | |
119 -- sup-x : ( Ordinal → Ordinal ) → Ordinal | |
120 -- sup-lb : { ψ : Ordinal → Ordinal } → {z : Ordinal } → z o< sup-o ψ → z o< osuc (ψ (sup-x ψ)) | |
227 | 121 |
219 | 122 ------------ |
123 -- | |
124 -- Onto map | |
125 -- def X x -> xmap | |
126 -- X ---------------------------> Y | |
127 -- ymap <- def Y y | |
128 -- | |
224 | 129 record Onto (X Y : OD ) : Set n where |
219 | 130 field |
226
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
131 xmap : Ordinal |
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
132 ymap : Ordinal |
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
133 xfunc : def (Func X Y) xmap |
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
134 yfunc : def (Func Y X) ymap |
234
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
135 onto-iso : {y : Ordinal } → (lty : def Y y ) → |
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
136 func←od {X} {Y} {xmap} xfunc ( func←od yfunc y ) ≡ y |
230 | 137 |
138 open Onto | |
139 | |
140 onto-restrict : {X Y Z : OD} → Onto X Y → ({x : OD} → _⊆_ Z Y {x}) → Onto X Z | |
141 onto-restrict {X} {Y} {Z} onto Z⊆Y = record { | |
142 xmap = xmap1 | |
143 ; ymap = zmap | |
144 ; xfunc = xfunc1 | |
145 ; yfunc = zfunc | |
146 ; onto-iso = onto-iso1 | |
147 } where | |
148 xmap1 : Ordinal | |
149 xmap1 = od→ord (Select (ord→od (xmap onto)) {!!} ) | |
150 zmap : Ordinal | |
151 zmap = {!!} | |
152 xfunc1 : def (Func X Z) xmap1 | |
153 xfunc1 = {!!} | |
154 zfunc : def (Func Z X) zmap | |
155 zfunc = {!!} | |
234
e06b76e5b682
ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
233
diff
changeset
|
156 onto-iso1 : {z : Ordinal } → (ltz : def Z z ) → func←od xfunc1 ( func←od zfunc z ) ≡ z |
230 | 157 onto-iso1 = {!!} |
158 | |
51 | 159 |
224 | 160 record Cardinal (X : OD ) : Set n where |
219 | 161 field |
224 | 162 cardinal : Ordinal |
230 | 163 conto : Onto X (Ord cardinal) |
164 cmax : ( y : Ordinal ) → cardinal o< y → ¬ Onto X (Ord y) | |
151 | 165 |
224 | 166 cardinal : (X : OD ) → Cardinal X |
167 cardinal X = record { | |
219 | 168 cardinal = sup-o ( λ x → proj1 ( cardinal-p x) ) |
226
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
169 ; conto = onto |
219 | 170 ; cmax = cmax |
171 } where | |
230 | 172 cardinal-p : (x : Ordinal ) → ( Ordinal ∧ Dec (Onto X (Ord x) ) ) |
173 cardinal-p x with p∨¬p ( Onto X (Ord x) ) | |
174 cardinal-p x | case1 True = record { proj1 = x ; proj2 = yes True } | |
219 | 175 cardinal-p x | case2 False = record { proj1 = o∅ ; proj2 = no False } |
229 | 176 S = sup-o (λ x → proj1 (cardinal-p x)) |
230 | 177 lemma1 : (x : Ordinal) → ((y : Ordinal) → y o< x → Lift (suc n) (y o< (osuc S) → Onto X (Ord y))) → |
178 Lift (suc n) (x o< (osuc S) → Onto X (Ord x) ) | |
229 | 179 lemma1 x prev with trio< x (osuc S) |
180 lemma1 x prev | tri< a ¬b ¬c with osuc-≡< a | |
230 | 181 lemma1 x prev | tri< a ¬b ¬c | case1 x=S = lift ( λ lt → {!!} ) |
182 lemma1 x prev | tri< a ¬b ¬c | case2 x<S = lift ( λ lt → lemma2 ) where | |
183 lemma2 : Onto X (Ord x) | |
184 lemma2 with prev {!!} {!!} | |
185 ... | lift t = t {!!} | |
229 | 186 lemma1 x prev | tri≈ ¬a b ¬c = lift ( λ lt → ⊥-elim ( o<¬≡ b lt )) |
187 lemma1 x prev | tri> ¬a ¬b c = lift ( λ lt → ⊥-elim ( o<> c lt )) | |
230 | 188 onto : Onto X (Ord S) |
189 onto with TransFinite {λ x → Lift (suc n) ( x o< osuc S → Onto X (Ord x) ) } lemma1 S | |
190 ... | lift t = t <-osuc | |
191 cmax : (y : Ordinal) → S o< y → ¬ Onto X (Ord y) | |
229 | 192 cmax y lt ontoy = o<> lt (o<-subst {_} {_} {y} {S} |
224 | 193 (sup-o< {λ x → proj1 ( cardinal-p x)}{y} ) lemma refl ) where |
219 | 194 lemma : proj1 (cardinal-p y) ≡ y |
230 | 195 lemma with p∨¬p ( Onto X (Ord y) ) |
219 | 196 lemma | case1 x = refl |
197 lemma | case2 not = ⊥-elim ( not ontoy ) | |
217 | 198 |
226
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
199 |
176ff97547b4
set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
225
diff
changeset
|
200 ----- |
219 | 201 -- All cardinal is ℵ0, since we are working on Countable Ordinal, |
202 -- Power ω is larger than ℵ0, so it has no cardinal. | |
218 | 203 |
204 | |
205 |