Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate ordinal-definable.agda @ 40:9439ff020cbd
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Thu, 23 May 2019 20:24:15 +0900 |
parents | 457e6626e0b1 |
children | b60db5903f01 |
rev | line source |
---|---|
16 | 1 open import Level |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
2 module ordinal-definable where |
3 | 3 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
4 open import zf |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
5 open import ordinal |
3 | 6 |
23 | 7 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) |
3 | 8 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
9 open import Relation.Binary.PropositionalEquality |
3 | 10 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
11 open import Data.Nat.Properties |
6 | 12 open import Data.Empty |
13 open import Relation.Nullary | |
14 | |
15 open import Relation.Binary | |
16 open import Relation.Binary.Core | |
17 | |
27 | 18 -- Ordinal Definable Set |
11 | 19 |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
20 record OD {n : Level} : Set (suc n) where |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
21 field |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
22 def : (x : Ordinal {n} ) → Set n |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
23 |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
24 open OD |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
25 open import Data.Unit |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
26 |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
27 postulate |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
28 od→ord : {n : Level} → OD {n} → Ordinal {n} |
36 | 29 ord→od : {n : Level} → Ordinal {n} → OD {n} |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
30 |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
31 _∋_ : { n : Level } → ( a x : OD {n} ) → Set n |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
32 _∋_ {n} a x = def a ( od→ord x ) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
33 |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
34 _c<_ : { n : Level } → ( a x : OD {n} ) → Set n |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
35 x c< a = a ∋ x |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
36 |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
37 _c≤_ : {n : Level} → OD {n} → OD {n} → Set (suc n) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
38 a c≤ b = (a ≡ b) ∨ ( b ∋ a ) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
39 |
40 | 40 od∅ : {n : Level} → OD {n} |
41 od∅ {n} = record { def = λ _ → Lift n ⊥ } | |
42 | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
43 postulate |
36 | 44 c<→o< : {n : Level} {x y : OD {n} } → x c< y → od→ord x o< od→ord y |
45 o<→c< : {n : Level} {x y : Ordinal {n} } → x o< y → ord→od x c< ord→od y | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
46 oiso : {n : Level} {x : OD {n}} → ord→od ( od→ord x ) ≡ x |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
47 diso : {n : Level} {x : Ordinal {n}} → od→ord ( ord→od x ) ≡ x |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
48 sup-od : {n : Level } → ( OD {n} → OD {n}) → OD {n} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
49 sup-c< : {n : Level } → ( ψ : OD {n} → OD {n}) → ∀ {x : OD {n}} → ψ x c< sup-od ψ |
40 | 50 ∅-base-def : {n : Level} → def ( ord→od (o∅ {n}) ) ≡ def (od∅ {n}) |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
51 |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
52 ∅1 : {n : Level} → ( x : OD {n} ) → ¬ ( x c< od∅ {n} ) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
53 ∅1 {n} x (lift ()) |
28 | 54 |
37 | 55 ∅3 : {n : Level} → { x : Ordinal {n}} → ( ∀(y : Ordinal {n}) → ¬ (y o< x ) ) → x ≡ o∅ {n} |
56 ∅3 {n} {x} = TransFinite {n} c1 c2 c3 x where | |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
57 c0 : Nat → Ordinal {n} → Set n |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
58 c0 lx x = (∀(y : Ordinal {n}) → ¬ (y o< x)) → x ≡ o∅ {n} |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
59 c1 : ∀ (lx : Nat ) → c0 lx (record { lv = Suc lx ; ord = ℵ lx } ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
60 c1 lx not with not ( record { lv = lx ; ord = Φ lx } ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
61 ... | t with t (case1 ≤-refl ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
62 c1 lx not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
63 c2 : (lx : Nat) → c0 lx (record { lv = lx ; ord = Φ lx } ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
64 c2 Zero not = refl |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
65 c2 (Suc lx) not with not ( record { lv = lx ; ord = Φ lx } ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
66 ... | t with t (case1 ≤-refl ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
67 c2 (Suc lx) not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
68 c3 : (lx : Nat) (x₁ : OrdinalD lx) → c0 lx (record { lv = lx ; ord = x₁ }) → c0 lx (record { lv = lx ; ord = OSuc lx x₁ }) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
69 c3 lx (Φ .lx) d not with not ( record { lv = lx ; ord = Φ lx } ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
70 ... | t with t (case2 Φ< ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
71 c3 lx (Φ .lx) d not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
72 c3 lx (OSuc .lx x₁) d not with not ( record { lv = lx ; ord = OSuc lx x₁ } ) |
34 | 73 ... | t with t (case2 (s< s<refl ) ) |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
74 c3 lx (OSuc .lx x₁) d not | t | () |
34 | 75 c3 (Suc lx) (ℵ lx) d not with not ( record { lv = Suc lx ; ord = OSuc (Suc lx) (Φ (Suc lx)) } ) |
76 ... | t with t (case2 (s< (ℵΦ< {_} {_} {Φ (Suc lx)}))) | |
77 c3 .(Suc lx) (ℵ lx) d not | t | () | |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
78 |
37 | 79 -- find : {n : Level} → ( x : Ordinal {n} ) → o∅ o< x → Ordinal {n} |
80 -- exists : {n : Level} → ( x : Ordinal {n} ) → (0<x : o∅ o< x ) → find x 0<x o< x | |
81 | |
36 | 82 def-subst : {n : Level } {Z : OD {n}} {X : Ordinal {n} }{z : OD {n}} {x : Ordinal {n} }→ def Z X → Z ≡ z → X ≡ x → def z x |
83 def-subst df refl refl = df | |
84 | |
85 transitive : {n : Level } { x y z : OD {n} } → y ∋ x → z ∋ y → z ∋ x | |
86 transitive {n} {x} {y} {z} x∋y z∋y with ordtrans ( c<→o< {n} {x} {y} x∋y ) ( c<→o< {n} {y} {z} z∋y ) | |
87 ... | t = lemma0 (lemma t) where | |
88 lemma : ( od→ord x ) o< ( od→ord z ) → def ( ord→od ( od→ord z )) ( od→ord ( ord→od ( od→ord x ))) | |
89 lemma xo<z = o<→c< xo<z | |
90 lemma0 : def ( ord→od ( od→ord z )) ( od→ord ( ord→od ( od→ord x ))) → def z (od→ord x) | |
91 lemma0 dz = def-subst {n} { ord→od ( od→ord z )} { od→ord ( ord→od ( od→ord x))} dz (oiso) (diso) | |
92 | |
37 | 93 ominimal : {n : Level} → (x : Ordinal {n} ) → o∅ o< x → Ordinal {n} |
94 ominimal {n} record { lv = Zero ; ord = (Φ .0) } (case1 ()) | |
95 ominimal {n} record { lv = Zero ; ord = (Φ .0) } (case2 ()) | |
96 ominimal {n} record { lv = Zero ; ord = (OSuc .0 ord) } (case1 ()) | |
97 ominimal {n} record { lv = Zero ; ord = (OSuc .0 ord) } (case2 Φ<) = record { lv = Zero ; ord = Φ 0 } | |
98 ominimal {n} record { lv = (Suc lv) ; ord = (Φ .(Suc lv)) } (case1 (s≤s x)) = record { lv = lv ; ord = Φ lv } | |
99 ominimal {n} record { lv = (Suc lv) ; ord = (Φ .(Suc lv)) } (case2 ()) | |
100 ominimal {n} record { lv = (Suc lv) ; ord = (OSuc .(Suc lv) ord) } (case1 (s≤s x)) = record { lv = (Suc lv) ; ord = ord } | |
101 ominimal {n} record { lv = (Suc lv) ; ord = (OSuc .(Suc lv) ord) } (case2 ()) | |
102 ominimal {n} record { lv = (Suc lv) ; ord = (ℵ .lv) } (case1 (s≤s z≤n)) = record { lv = lv ; ord = Φ lv } | |
103 ominimal {n} record { lv = (Suc lv) ; ord = (ℵ .lv) } (case2 ()) | |
104 | |
105 ∅4 : {n : Level} → ( x : OD {n} ) → x ≡ od∅ {n} → od→ord x ≡ o∅ {n} | |
106 ∅4 {n} x refl = ∅3 lemma1 where | |
107 lemma0 : (y : Ordinal {n}) → def ( od∅ {n} ) y → ⊥ | |
108 lemma0 y (lift ()) | |
109 lemma1 : (y : Ordinal {n}) → y o< od→ord od∅ → ⊥ | |
110 lemma1 y y<o = lemma0 y ( def-subst {n} {ord→od (od→ord od∅ )} {od→ord (ord→od y)} (o<→c< y<o) oiso diso ) | |
111 | |
112 ∅5 : {n : Level} → ( x : Ordinal {n} ) → ¬ ( x ≡ o∅ {n} ) → o∅ {n} o< x | |
113 ∅5 {n} record { lv = Zero ; ord = (Φ .0) } not = ⊥-elim (not refl) | |
114 ∅5 {n} record { lv = Zero ; ord = (OSuc .0 ord) } not = case2 Φ< | |
115 ∅5 {n} record { lv = (Suc lv) ; ord = ord } not = case1 (s≤s z≤n) | |
116 | |
39 | 117 postulate extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality n (suc n) |
37 | 118 |
119 ∅6 : {n : Level } ( x : Ordinal {suc n}) → o∅ o< x → ¬ x ≡ o∅ | |
120 ∅6 {n} x lt eq with trio< {n} (o∅ {suc n}) x | |
121 ∅6 {n} x lt refl | tri< a ¬b ¬c = ¬b refl | |
122 ∅6 {n} x lt refl | tri≈ ¬a b ¬c = ¬a lt | |
123 ∅6 {n} x lt refl | tri> ¬a ¬b c = ¬b refl | |
124 | |
39 | 125 ∅8 : {n : Level} → ( x : Ordinal {n} ) → ¬ x o< o∅ {n} |
126 ∅8 {n} x (case1 ()) | |
127 ∅8 {n} x (case2 ()) | |
128 | |
40 | 129 -- ∅10 : {n : Level} → (x : OD {n} ) → ¬ ( ( y : OD {n} ) → Lift (suc n) ( x ∋ y)) → x ≡ od∅ |
130 -- ∅10 {n} x not = ? | |
39 | 131 |
132 open Ordinal | |
133 | |
40 | 134 ∋-subst : {n : Level} {X Y x y : OD {suc n} } → X ≡ x → Y ≡ y → X ∋ Y → x ∋ y |
135 ∋-subst refl refl x = x | |
39 | 136 |
40 | 137 -- ∅77 : {n : Level} → (x : OD {suc n} ) → ¬ ( ord→od (o∅ {suc n}) ∋ x ) |
138 -- ∅77 {n} x lt = {!!} where | |
39 | 139 |
140 ∅7' : {n : Level} → ord→od (o∅ {n}) ≡ od∅ {n} | |
40 | 141 ∅7' {n} = cong ( λ k → record { def = k }) ( ∅-base-def ) where |
39 | 142 |
38 | 143 ∅7 : {n : Level} → ( x : OD {n} ) → od→ord x ≡ o∅ {n} → x ≡ od∅ {n} |
39 | 144 ∅7 {n} x eq = trans ( trans (sym oiso)( cong ( λ k → ord→od k ) eq ) ) ∅7' |
37 | 145 |
38 | 146 ∅9 : {n : Level} → (x : OD {n} ) → ¬ x ≡ od∅ → o∅ o< od→ord x |
147 ∅9 x not = ∅5 ( od→ord x) lemma where | |
148 lemma : ¬ od→ord x ≡ o∅ | |
149 lemma eq = not ( ∅7 x eq ) | |
37 | 150 |
40 | 151 OD→ZF : {n : Level} → ZF {suc n} {suc n} |
152 OD→ZF {n} = record { | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
153 ZFSet = OD {n} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
154 ; _∋_ = λ a x → Lift (suc n) ( a ∋ x ) |
28 | 155 ; _≈_ = _≡_ |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
156 ; ∅ = od∅ |
28 | 157 ; _,_ = _,_ |
158 ; Union = Union | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
159 ; Power = Power |
28 | 160 ; Select = Select |
161 ; Replace = Replace | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
162 ; infinite = record { def = λ x → x ≡ record { lv = Suc Zero ; ord = ℵ Zero } } |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
163 ; isZF = isZF |
28 | 164 } where |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
165 Replace : OD {n} → (OD {n} → OD {n} ) → OD {n} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
166 Replace X ψ = sup-od ψ |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
167 Select : OD {n} → (OD {n} → Set (suc n) ) → OD {n} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
168 Select X ψ = record { def = λ x → select ( ord→od x ) } where |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
169 select : OD {n} → Set n |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
170 select x with ψ x |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
171 ... | t = Lift n ⊤ |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
172 _,_ : OD {n} → OD {n} → OD {n} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
173 x , y = record { def = λ z → ( (z ≡ od→ord x ) ∨ ( z ≡ od→ord y )) } |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
174 Union : OD {n} → OD {n} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
175 Union x = record { def = λ y → {z : Ordinal {n}} → def x z → def (ord→od z) y } |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
176 Power : OD {n} → OD {n} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
177 Power x = record { def = λ y → (z : Ordinal {n} ) → ( def x y ∧ def (ord→od z) y ) } |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
178 ZFSet = OD {n} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
179 _∈_ : ( A B : ZFSet ) → Set n |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
180 A ∈ B = B ∋ A |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
181 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set n |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
182 _⊆_ A B {x} = A ∋ x → B ∋ x |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
183 _∩_ : ( A B : ZFSet ) → ZFSet |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
184 A ∩ B = Select (A , B) ( λ x → (Lift (suc n) ( A ∋ x )) ∧ (Lift (suc n) ( B ∋ x ) )) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
185 _∪_ : ( A B : ZFSet ) → ZFSet |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
186 A ∪ B = Select (A , B) ( λ x → (Lift (suc n) ( A ∋ x )) ∨ (Lift (suc n) ( B ∋ x ) )) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
187 infixr 200 _∈_ |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
188 infixr 230 _∩_ _∪_ |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
189 infixr 220 _⊆_ |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
190 isZF : IsZF (OD {n}) (λ a x → Lift (suc n) ( a ∋ x )) _≡_ od∅ _,_ Union Power Select Replace (record { def = λ x → x ≡ record { lv = Suc Zero ; ord = ℵ Zero } }) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
191 isZF = record { |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
192 isEquivalence = record { refl = refl ; sym = sym ; trans = trans } |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
193 ; pair = pair |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
194 ; union→ = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
195 ; union← = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
196 ; empty = empty |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
197 ; power→ = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
198 ; power← = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
199 ; extentionality = {!!} |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
200 ; minimul = minimul |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
201 ; regularity = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
202 ; infinity∅ = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
203 ; infinity = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
204 ; selection = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
205 ; replacement = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
206 } where |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
207 open _∧_ |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
208 pair : (A B : OD {n} ) → Lift (suc n) ((A , B) ∋ A) ∧ Lift (suc n) ((A , B) ∋ B) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
209 proj1 (pair A B ) = lift ( case1 refl ) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
210 proj2 (pair A B ) = lift ( case2 refl ) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
211 empty : (x : OD {n} ) → ¬ Lift (suc n) (od∅ ∋ x) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
212 empty x (lift (lift ())) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
213 union→ : (X x y : OD {n} ) → Lift (suc n) (X ∋ x) → Lift (suc n) (x ∋ y) → Lift (suc n) (Union X ∋ y) |
36 | 214 union→ X x y (lift X∋x) (lift x∋y) = lift {!!} where |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
215 lemma : {z : Ordinal {n} } → def X z → z ≡ od→ord y |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
216 lemma {z} X∋z = {!!} |
37 | 217 minimul : (x : OD {n} ) → ¬ x ≡ od∅ → OD {n} |
38 | 218 minimul x not = ord→od ( ominimal (od→ord x) (∅9 x not) ) |
37 | 219 regularity : (x : OD) → (not : ¬ x ≡ od∅ ) → |
220 Lift (suc n) (x ∋ minimul x not ) ∧ | |
221 (Select x (λ x₁ → Lift (suc n) ( minimul x not ∋ x₁) ∧ Lift (suc n) (x ∋ x₁)) ≡ od∅) | |
222 proj1 ( regularity x non ) = {!!} | |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
223 proj2 ( regularity x non ) = {!!} |