Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate ordinal-definable.agda @ 103:c8b79d303867
starting over HOD
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Wed, 12 Jun 2019 10:45:00 +0900 |
parents | a402881cc341 |
children | d92411bed18c dab56d357fa3 |
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⊔_ ) |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
8 open import Relation.Binary.PropositionalEquality |
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
9 open import Data.Nat.Properties |
6 | 10 open import Data.Empty |
11 open import Relation.Nullary | |
12 open import Relation.Binary | |
13 open import Relation.Binary.Core | |
14 | |
27 | 15 -- Ordinal Definable Set |
11 | 16 |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
17 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
|
18 field |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
19 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
|
20 |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
21 open OD |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
22 open import Data.Unit |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
23 |
44
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
24 open Ordinal |
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
25 |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
26 record _==_ {n : Level} ( a b : OD {n} ) : Set n where |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
27 field |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
28 eq→ : ∀ { x : Ordinal {n} } → def a x → def b x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
29 eq← : ∀ { x : Ordinal {n} } → def b x → def a x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
30 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
31 id : {n : Level} {A : Set n} → A → A |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
32 id x = x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
33 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
34 eq-refl : {n : Level} { x : OD {n} } → x == x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
35 eq-refl {n} {x} = record { eq→ = id ; eq← = id } |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
36 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
37 open _==_ |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
38 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
39 eq-sym : {n : Level} { x y : OD {n} } → x == y → y == x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
40 eq-sym eq = record { eq→ = eq← eq ; eq← = eq→ eq } |
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 eq-trans : {n : Level} { x y z : OD {n} } → x == y → y == z → x == z |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
43 eq-trans x=y y=z = record { eq→ = λ t → eq→ y=z ( eq→ x=y t) ; eq← = λ t → eq← x=y ( eq← y=z t) } |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
44 |
40 | 45 od∅ : {n : Level} → OD {n} |
46 od∅ {n} = record { def = λ _ → Lift n ⊥ } | |
47 | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
48 postulate |
100 | 49 -- OD can be iso to a subset of Ordinal ( by means of Godel Set ) |
95 | 50 od→ord : {n : Level} → OD {n} → Ordinal {n} |
51 ord→od : {n : Level} → Ordinal {n} → OD {n} | |
52 oiso : {n : Level} {x : OD {n}} → ord→od ( od→ord x ) ≡ x | |
53 diso : {n : Level} {x : Ordinal {n}} → od→ord ( ord→od x ) ≡ x | |
100 | 54 -- supermum as Replacement Axiom |
95 | 55 sup-o : {n : Level } → ( Ordinal {n} → Ordinal {n}) → Ordinal {n} |
98 | 56 sup-o< : {n : Level } → { ψ : Ordinal {n} → Ordinal {n}} → ∀ {x : Ordinal {n}} → ψ x o< sup-o ψ |
103 | 57 -- a contra-position of minimality of supermum |
98 | 58 sup-x : {n : Level } → ( Ordinal {n} → Ordinal {n}) → Ordinal {n} |
59 sup-lb : {n : Level } → { ψ : Ordinal {n} → Ordinal {n}} → {z : Ordinal {n}} → z o< sup-o ψ → z o< osuc (ψ (sup-x ψ)) | |
95 | 60 |
61 _∋_ : { n : Level } → ( a x : OD {n} ) → Set n | |
62 _∋_ {n} a x = def a ( od→ord x ) | |
63 | |
64 _c<_ : { n : Level } → ( x a : OD {n} ) → Set n | |
103 | 65 x c< a = od→ord x o< od→ord a |
66 | |
67 postulate | |
68 o<→c< : {n : Level} {x y : Ordinal {n} } → x o< y → ord→od x c< ord→od y | |
95 | 69 |
70 _c≤_ : {n : Level} → OD {n} → OD {n} → Set (suc n) | |
71 a c≤ b = (a ≡ b) ∨ ( b ∋ a ) | |
72 | |
73 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 | |
74 def-subst df refl refl = df | |
75 | |
103 | 76 -- sup-min : {n : Level } → ( ψ : Ordinal {n} → Ordinal {n}) → {z : Ordinal {n}} → ψ z o< z → sup-o ψ o< osuc z |
77 | |
95 | 78 sup-od : {n : Level } → ( OD {n} → OD {n}) → OD {n} |
79 sup-od ψ = ord→od ( sup-o ( λ x → od→ord (ψ (ord→od x ))) ) | |
80 | |
81 sup-c< : {n : Level } → ( ψ : OD {n} → OD {n}) → ∀ {x : OD {n}} → def ( sup-od ψ ) (od→ord ( ψ x )) | |
82 sup-c< {n} ψ {x} = def-subst {n} {_} {_} {sup-od ψ} {od→ord ( ψ x )} | |
103 | 83 {!!} refl (cong ( λ k → od→ord (ψ k) ) oiso) |
46 | 84 |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
85 ∅1 : {n : Level} → ( x : OD {n} ) → ¬ ( x c< od∅ {n} ) |
103 | 86 ∅1 {n} x = {!!} |
28 | 87 |
37 | 88 ∅3 : {n : Level} → { x : Ordinal {n}} → ( ∀(y : Ordinal {n}) → ¬ (y o< x ) ) → x ≡ o∅ {n} |
81 | 89 ∅3 {n} {x} = TransFinite {n} c2 c3 x where |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
90 c0 : Nat → Ordinal {n} → Set n |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
91 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
|
92 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
|
93 c2 Zero not = refl |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
94 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
|
95 ... | t with t (case1 ≤-refl ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
96 c2 (Suc lx) not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
97 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
|
98 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
|
99 ... | t with t (case2 Φ< ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
100 c3 lx (Φ .lx) d not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
101 c3 lx (OSuc .lx x₁) d not with not ( record { lv = lx ; ord = OSuc lx x₁ } ) |
34 | 102 ... | 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
|
103 c3 lx (OSuc .lx x₁) d not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
104 |
69 | 105 transitive : {n : Level } { z y x : OD {suc n} } → z ∋ y → y ∋ x → z ∋ x |
103 | 106 transitive {n} {z} {y} {x} z∋y x∋y with ordtrans {!!} {!!} |
36 | 107 ... | t = lemma0 (lemma t) where |
95 | 108 lemma : ( od→ord x ) o< ( od→ord z ) → def ( ord→od ( od→ord z )) ( od→ord x) |
103 | 109 lemma xo<z = {!!} |
95 | 110 lemma0 : def ( ord→od ( od→ord z )) ( od→ord x) → def z (od→ord x) |
111 lemma0 dz = def-subst {suc n} { ord→od ( od→ord z )} { od→ord x} dz (oiso) refl | |
36 | 112 |
57 | 113 ∅5 : {n : Level} → { x : Ordinal {n} } → ¬ ( x ≡ o∅ {n} ) → o∅ {n} o< x |
114 ∅5 {n} {record { lv = Zero ; ord = (Φ .0) }} not = ⊥-elim (not refl) | |
115 ∅5 {n} {record { lv = Zero ; ord = (OSuc .0 ord) }} not = case2 Φ< | |
116 ∅5 {n} {record { lv = (Suc lv) ; ord = ord }} not = case1 (s≤s z≤n) | |
37 | 117 |
46 | 118 ord-iso : {n : Level} {y : Ordinal {n} } → record { lv = lv (od→ord (ord→od y)) ; ord = ord (od→ord (ord→od y)) } ≡ record { lv = lv y ; ord = ord y } |
119 ord-iso = cong ( λ k → record { lv = lv k ; ord = ord k } ) diso | |
44
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
120 |
51 | 121 -- avoiding lv != Zero error |
122 orefl : {n : Level} → { x : OD {n} } → { y : Ordinal {n} } → od→ord x ≡ y → od→ord x ≡ y | |
123 orefl refl = refl | |
124 | |
125 ==-iso : {n : Level} → { x y : OD {n} } → ord→od (od→ord x) == ord→od (od→ord y) → x == y | |
126 ==-iso {n} {x} {y} eq = record { | |
127 eq→ = λ d → lemma ( eq→ eq (def-subst d (sym oiso) refl )) ; | |
128 eq← = λ d → lemma ( eq← eq (def-subst d (sym oiso) refl )) } | |
129 where | |
130 lemma : {x : OD {n} } {z : Ordinal {n}} → def (ord→od (od→ord x)) z → def x z | |
131 lemma {x} {z} d = def-subst d oiso refl | |
132 | |
57 | 133 =-iso : {n : Level } {x y : OD {suc n} } → (x == y) ≡ (ord→od (od→ord x) == y) |
134 =-iso {_} {_} {y} = cong ( λ k → k == y ) (sym oiso) | |
135 | |
51 | 136 ord→== : {n : Level} → { x y : OD {n} } → od→ord x ≡ od→ord y → x == y |
137 ord→== {n} {x} {y} eq = ==-iso (lemma (od→ord x) (od→ord y) (orefl eq)) where | |
138 lemma : ( ox oy : Ordinal {n} ) → ox ≡ oy → (ord→od ox) == (ord→od oy) | |
139 lemma ox ox refl = eq-refl | |
140 | |
141 o≡→== : {n : Level} → { x y : Ordinal {n} } → x ≡ y → ord→od x == ord→od y | |
142 o≡→== {n} {x} {.x} refl = eq-refl | |
143 | |
144 >→¬< : {x y : Nat } → (x < y ) → ¬ ( y < x ) | |
145 >→¬< (s≤s x<y) (s≤s y<x) = >→¬< x<y y<x | |
146 | |
147 c≤-refl : {n : Level} → ( x : OD {n} ) → x c≤ x | |
148 c≤-refl x = case1 refl | |
149 | |
54 | 150 o<→o> : {n : Level} → { x y : OD {n} } → (x == y) → (od→ord x ) o< ( od→ord y) → ⊥ |
52 | 151 o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case1 lt) with |
103 | 152 yx (def-subst {n} {ord→od (od→ord y)} {od→ord x} {!!} oiso refl ) |
153 ... | oyx with o<¬≡ (od→ord x) (od→ord x) refl {!!} | |
52 | 154 ... | () |
155 o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case2 lt) with | |
103 | 156 yx (def-subst {n} {ord→od (od→ord y)} {od→ord x} {!!} oiso refl ) |
157 ... | oyx with o<¬≡ (od→ord x) (od→ord x) refl {!!} | |
52 | 158 ... | () |
159 | |
79 | 160 ==→o≡ : {n : Level} → { x y : Ordinal {suc n} } → ord→od x == ord→od y → x ≡ y |
161 ==→o≡ {n} {x} {y} eq with trio< {n} x y | |
162 ==→o≡ {n} {x} {y} eq | tri< a ¬b ¬c = ⊥-elim ( o<→o> eq (o<-subst a (sym ord-iso) (sym ord-iso ))) | |
163 ==→o≡ {n} {x} {y} eq | tri≈ ¬a b ¬c = b | |
164 ==→o≡ {n} {x} {y} eq | tri> ¬a ¬b c = ⊥-elim ( o<→o> (eq-sym eq) (o<-subst c (sym ord-iso) (sym ord-iso ))) | |
165 | |
90 | 166 ≡-def : {n : Level} → { x : Ordinal {suc n} } → x ≡ od→ord (record { def = λ z → z o< x } ) |
167 ≡-def {n} {x} = ==→o≡ {n} (subst (λ k → ord→od x == k ) (sym oiso) lemma ) where | |
168 lemma : ord→od x == record { def = λ z → z o< x } | |
95 | 169 eq→ lemma {w} z = subst₂ (λ k j → k o< j ) diso refl (subst (λ k → (od→ord ( ord→od w)) o< k ) diso t ) where |
170 t : (od→ord ( ord→od w)) o< (od→ord (ord→od x)) | |
103 | 171 t = {!!} |
172 eq← lemma {w} z = def-subst {suc n} {_} {_} {ord→od x} {w} {!!} refl refl | |
90 | 173 |
174 od≡-def : {n : Level} → { x : Ordinal {suc n} } → ord→od x ≡ record { def = λ z → z o< x } | |
175 od≡-def {n} {x} = subst (λ k → ord→od x ≡ k ) oiso (cong ( λ k → ord→od k ) (≡-def {n} {x} )) | |
176 | |
99
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
177 ==→o≡1 : {n : Level} → { x y : OD {suc n} } → x == y → od→ord x ≡ od→ord y |
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
178 ==→o≡1 eq = ==→o≡ (subst₂ (λ k j → k == j ) (sym oiso) (sym oiso) eq ) |
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
179 |
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
180 ==-def-l : {n : Level } {x y : Ordinal {suc n} } { z : OD {suc n} }→ (ord→od x == ord→od y) → def z x → def z y |
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
181 ==-def-l {n} {x} {y} {z} eq z>x = subst ( λ k → def z k ) (==→o≡ eq) z>x |
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
182 |
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
183 ==-def-r : {n : Level } {x y : OD {suc n} } { z : Ordinal {suc n} }→ (x == y) → def x z → def y z |
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
184 ==-def-r {n} {x} {y} {z} eq z>x = subst (λ k → def k z ) (subst₂ (λ j k → j ≡ k ) oiso oiso (cong (λ k → ord→od k) (==→o≡1 eq))) z>x |
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
185 |
91 | 186 ∋→o< : {n : Level} → { a x : OD {suc n} } → a ∋ x → od→ord x o< od→ord a |
187 ∋→o< {n} {a} {x} lt = t where | |
188 t : (od→ord x) o< (od→ord a) | |
103 | 189 t = {!!} |
91 | 190 |
191 o<∋→ : {n : Level} → { a x : OD {suc n} } → od→ord x o< od→ord a → a ∋ x | |
95 | 192 o<∋→ {n} {a} {x} lt = subst₂ (λ k j → def k j ) oiso refl t where |
193 t : def (ord→od (od→ord a)) (od→ord x) | |
103 | 194 t = {!!} |
91 | 195 |
80 | 196 o∅≡od∅ : {n : Level} → ord→od (o∅ {suc n}) ≡ od∅ {suc n} |
197 o∅≡od∅ {n} with trio< {n} (o∅ {suc n}) (od→ord (od∅ {suc n} )) | |
198 o∅≡od∅ {n} | tri< a ¬b ¬c = ⊥-elim (lemma a) where | |
199 lemma : o∅ {suc n } o< (od→ord (od∅ {suc n} )) → ⊥ | |
103 | 200 lemma lt with def-subst {!!} oiso refl |
201 lemma lt | t = {!!} | |
80 | 202 o∅≡od∅ {n} | tri≈ ¬a b ¬c = trans (cong (λ k → ord→od k ) b ) oiso |
203 o∅≡od∅ {n} | tri> ¬a ¬b c = ⊥-elim (¬x<0 c) | |
204 | |
51 | 205 o<→¬== : {n : Level} → { x y : OD {n} } → (od→ord x ) o< ( od→ord y) → ¬ (x == y ) |
52 | 206 o<→¬== {n} {x} {y} lt eq = o<→o> eq lt |
51 | 207 |
208 o<→¬c> : {n : Level} → { x y : OD {n} } → (od→ord x ) o< ( od→ord y) → ¬ (y c< x ) | |
103 | 209 o<→¬c> {n} {x} {y} olt clt = o<> olt {!!} where |
51 | 210 |
211 o≡→¬c< : {n : Level} → { x y : OD {n} } → (od→ord x ) ≡ ( od→ord y) → ¬ x c< y | |
103 | 212 o≡→¬c< {n} {x} {y} oeq lt = o<¬≡ (od→ord x) (od→ord y) (orefl oeq ) lt |
51 | 213 |
214 tri-c< : {n : Level} → Trichotomous _==_ (_c<_ {suc n}) | |
215 tri-c< {n} x y with trio< {n} (od→ord x) (od→ord y) | |
103 | 216 tri-c< {n} x y | tri< a ¬b ¬c = tri< (def-subst {!!} oiso refl) (o<→¬== a) ( o<→¬c> a ) |
51 | 217 tri-c< {n} x y | tri≈ ¬a b ¬c = tri≈ (o≡→¬c< b) (ord→== b) (o≡→¬c< (sym b)) |
103 | 218 tri-c< {n} x y | tri> ¬a ¬b c = tri> ( o<→¬c> c) (λ eq → o<→¬== c (eq-sym eq ) ) (def-subst {!!} oiso refl) |
51 | 219 |
54 | 220 c<> : {n : Level } { x y : OD {suc n}} → x c< y → y c< x → ⊥ |
221 c<> {n} {x} {y} x<y y<x with tri-c< x y | |
222 c<> {n} {x} {y} x<y y<x | tri< a ¬b ¬c = ¬c y<x | |
103 | 223 c<> {n} {x} {y} x<y y<x | tri≈ ¬a b ¬c = o<→o> b x<y |
54 | 224 c<> {n} {x} {y} x<y y<x | tri> ¬a ¬b c = ¬a x<y |
225 | |
60 | 226 ∅< : {n : Level} → { x y : OD {n} } → def x (od→ord y ) → ¬ ( x == od∅ {n} ) |
227 ∅< {n} {x} {y} d eq with eq→ eq d | |
228 ∅< {n} {x} {y} d eq | lift () | |
57 | 229 |
78
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
230 ∅6 : {n : Level} → { x : OD {suc n} } → ¬ ( x ∋ x ) -- no Russel paradox |
103 | 231 ∅6 {n} {x} x∋x = c<> {n} {x} {x} {!!} {!!} |
51 | 232 |
76 | 233 def-iso : {n : Level} {A B : OD {n}} {x y : Ordinal {n}} → x ≡ y → (def A y → def B y) → def A x → def B x |
234 def-iso refl t = t | |
235 | |
53 | 236 is-∋ : {n : Level} → ( x y : OD {suc n} ) → Dec ( x ∋ y ) |
237 is-∋ {n} x y with tri-c< x y | |
103 | 238 is-∋ {n} x y | tri< a ¬b ¬c = no {!!} |
239 is-∋ {n} x y | tri≈ ¬a b ¬c = no {!!} | |
240 is-∋ {n} x y | tri> ¬a ¬b c = yes {!!} | |
53 | 241 |
57 | 242 is-o∅ : {n : Level} → ( x : Ordinal {suc n} ) → Dec ( x ≡ o∅ {suc n} ) |
243 is-o∅ {n} record { lv = Zero ; ord = (Φ .0) } = yes refl | |
244 is-o∅ {n} record { lv = Zero ; ord = (OSuc .0 ord₁) } = no ( λ () ) | |
245 is-o∅ {n} record { lv = (Suc lv₁) ; ord = ord } = no (λ()) | |
246 | |
45
33860eb44e47
od∅' {n} = ord→od (o∅ {n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
44
diff
changeset
|
247 open _∧_ |
33860eb44e47
od∅' {n} = ord→od (o∅ {n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
44
diff
changeset
|
248 |
66 | 249 ¬∅=→∅∈ : {n : Level} → { x : OD {suc n} } → ¬ ( x == od∅ {suc n} ) → x ∋ od∅ {suc n} |
250 ¬∅=→∅∈ {n} {x} ne = def-subst (lemma (od→ord x) (subst (λ k → ¬ (k == od∅ {suc n} )) (sym oiso) ne )) oiso refl where | |
251 lemma : (ox : Ordinal {suc n}) → ¬ (ord→od ox == od∅ {suc n} ) → ord→od ox ∋ od∅ {suc n} | |
252 lemma ox ne with is-o∅ ox | |
253 lemma ox ne | yes refl with ne ( ord→== lemma1 ) where | |
254 lemma1 : od→ord (ord→od o∅) ≡ od→ord od∅ | |
80 | 255 lemma1 = cong ( λ k → od→ord k ) o∅≡od∅ |
66 | 256 lemma o∅ ne | yes refl | () |
103 | 257 lemma ox ne | no ¬p = subst ( λ k → def (ord→od ox) (od→ord k) ) o∅≡od∅ {!!} |
69 | 258 |
79 | 259 -- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) |
94 | 260 -- postulate f-extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality (suc n) (suc (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
|
261 |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
262 csuc : {n : Level} → OD {suc n} → OD {suc n} |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
263 csuc x = ord→od ( osuc ( od→ord x )) |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
264 |
96 | 265 -- Power Set of X ( or constructible by λ y → def X (od→ord y ) |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
266 |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
267 ZFSubset : {n : Level} → (A x : OD {suc n} ) → OD {suc n} |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
268 ZFSubset A x = record { def = λ y → def A y ∧ def x y } |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
269 |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
270 Def : {n : Level} → (A : OD {suc n}) → OD {suc n} |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
271 Def {n} A = ord→od ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) ) |
96 | 272 |
273 -- Constructible Set on α | |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
274 L : {n : Level} → (α : Ordinal {suc n}) → OD {suc n} |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
275 L {n} record { lv = Zero ; ord = (Φ .0) } = od∅ |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
276 L {n} record { lv = lx ; ord = (OSuc lv ox) } = Def ( L {n} ( record { lv = lx ; ord = ox } ) ) |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
277 L {n} record { lv = (Suc lx) ; ord = (Φ (Suc lx)) } = -- Union ( L α ) |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
278 record { def = λ y → osuc y o< (od→ord (L {n} (record { lv = lx ; ord = Φ lx }) )) } |
89 | 279 |
54 | 280 OD→ZF : {n : Level} → ZF {suc (suc n)} {suc n} |
40 | 281 OD→ZF {n} = record { |
54 | 282 ZFSet = OD {suc n} |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
283 ; _∋_ = _∋_ |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
284 ; _≈_ = _==_ |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
285 ; ∅ = od∅ |
28 | 286 ; _,_ = _,_ |
287 ; Union = Union | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
288 ; Power = Power |
28 | 289 ; Select = Select |
290 ; Replace = Replace | |
81 | 291 ; infinite = ord→od ( record { lv = Suc Zero ; ord = Φ 1 } ) |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
292 ; isZF = isZF |
28 | 293 } where |
54 | 294 Replace : OD {suc n} → (OD {suc n} → OD {suc n} ) → OD {suc n} |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
295 Replace X ψ = sup-od ψ |
54 | 296 Select : OD {suc n} → (OD {suc n} → Set (suc n) ) → OD {suc n} |
297 Select X ψ = record { def = λ x → ( def X x ∧ ψ ( ord→od x )) } | |
298 _,_ : OD {suc n} → OD {suc n} → OD {suc n} | |
84 | 299 x , y = record { def = λ z → z o< (omax (od→ord x) (od→ord y)) } |
54 | 300 Union : OD {suc n} → OD {suc n} |
71
d088eb66564e
add osuc ( next larger element of Ordinal )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
70
diff
changeset
|
301 Union U = record { def = λ y → osuc y o< (od→ord U) } |
77 | 302 -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ ( ∀ {x} → t ∋ x → X ∋ x ) |
54 | 303 Power : OD {suc n} → OD {suc n} |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
304 Power A = Def A |
54 | 305 ZFSet = OD {suc n} |
306 _∈_ : ( A B : ZFSet ) → Set (suc 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 |
54 | 308 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set (suc n) |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
309 _⊆_ A B {x} = A ∋ x → B ∋ x |
103 | 310 _∩_ : ( A B : ZFSet ) → ZFSet |
311 A ∩ B = Select (A , B) ( λ x → ( A ∋ x ) ∧ (B ∋ x) ) | |
96 | 312 -- _∪_ : ( A B : ZFSet ) → ZFSet |
313 -- A ∪ B = Select (A , B) ( λ x → (A ∋ x) ∨ ( B ∋ x ) ) | |
103 | 314 {_} : ZFSet → ZFSet |
315 { x } = ( x , x ) | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
316 infixr 200 _∈_ |
96 | 317 -- infixr 230 _∩_ _∪_ |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
318 infixr 220 _⊆_ |
81 | 319 isZF : IsZF (OD {suc n}) _∋_ _==_ od∅ _,_ Union Power Select Replace (ord→od ( record { lv = Suc Zero ; ord = Φ 1} )) |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
320 isZF = record { |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
321 isEquivalence = record { refl = eq-refl ; sym = eq-sym; trans = eq-trans } |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
322 ; pair = pair |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
323 ; union-u = λ _ z _ → csuc z |
72 | 324 ; union→ = union→ |
325 ; union← = union← | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
326 ; empty = empty |
76 | 327 ; power→ = power→ |
328 ; power← = power← | |
329 ; extensionality = extensionality | |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
330 ; minimul = minimul |
51 | 331 ; regularity = regularity |
78
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
332 ; infinity∅ = infinity∅ |
93 | 333 ; infinity = λ _ → infinity |
55
9c0a5e28a572
regurality elimination case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
54
diff
changeset
|
334 ; selection = λ {ψ} {X} {y} → selection {ψ} {X} {y} |
93 | 335 ; replacement = replacement |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
336 } where |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
337 open _∧_ |
54 | 338 pair : (A B : OD {suc n} ) → ((A , B) ∋ A) ∧ ((A , B) ∋ B) |
87 | 339 proj1 (pair A B ) = omax-x {n} (od→ord A) (od→ord B) |
340 proj2 (pair A B ) = omax-y {n} (od→ord A) (od→ord B) | |
54 | 341 empty : (x : OD {suc n} ) → ¬ (od∅ ∋ x) |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
342 empty x () |
100 | 343 --- |
344 --- ZFSubset A x = record { def = λ y → def A y ∧ def x y } subset of A | |
345 --- Power X = ord→od ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) ) Power X is a sup of all subset of A | |
346 -- | |
103 | 347 -- if Power A ∋ t, from a minimulity of sup, there is osuc ZFSubset A ∋ t |
100 | 348 -- then ZFSubset A ≡ t or ZFSubset A ∋ t. In the former case ZFSubset A ∋ x implies A ∋ x |
349 -- In case of later, ZFSubset A ∋ t and t ∋ x implies ZFSubset A ∋ x by transitivity | |
350 -- | |
76 | 351 power→ : (A t : OD) → Power A ∋ t → {x : OD} → t ∋ x → A ∋ x |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
352 power→ A t P∋t {x} t∋x = proj1 lemma-s where |
98 | 353 minsup : OD |
99
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
354 minsup = ZFSubset A ( ord→od ( sup-x (λ x → od→ord ( ZFSubset A (ord→od x))))) |
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
355 lemma-t : csuc minsup ∋ t |
103 | 356 lemma-t = {!!} |
98 | 357 lemma-s : ZFSubset A ( ord→od ( sup-x (λ x → od→ord ( ZFSubset A (ord→od x))))) ∋ x |
103 | 358 lemma-s = {!!} |
100 | 359 -- |
360 -- we have t ∋ x → A ∋ x means t is a subset of A, that is ZFSubset A t == t | |
361 -- Power A is a sup of ZFSubset A t, so Power A ∋ t | |
362 -- | |
77 | 363 power← : (A t : OD) → ({x : OD} → (t ∋ x → A ∋ x)) → Power A ∋ t |
99
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
364 power← A t t→A = def-subst {suc n} {_} {_} {Power A} {od→ord t} |
103 | 365 {!!} refl lemma1 where |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
366 lemma-eq : ZFSubset A t == t |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
367 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
|
368 eq← lemma-eq {z} w = record { proj2 = w ; |
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
369 proj1 = def-subst {suc n} {_} {_} {A} {z} ( t→A (def-subst {suc n} {_} {_} {t} {od→ord (ord→od z)} w refl (sym diso) )) refl diso } |
99
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
370 lemma1 : od→ord (ZFSubset A (ord→od (od→ord t))) ≡ od→ord t |
74330d0cdcbc
Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
98
diff
changeset
|
371 lemma1 = subst (λ k → od→ord (ZFSubset A k) ≡ od→ord t ) (sym oiso) (==→o≡1 (lemma-eq)) |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
372 lemma : od→ord (ZFSubset A (ord→od (od→ord t)) ) o< sup-o (λ x → od→ord (ZFSubset A (ord→od x))) |
98 | 373 lemma = sup-o< |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
374 union-lemma-u : {X z : OD {suc n}} → (U>z : Union X ∋ z ) → csuc z ∋ z |
72 | 375 union-lemma-u {X} {z} U>z = lemma <-osuc where |
376 lemma : {oz ooz : Ordinal {suc n}} → oz o< ooz → def (ord→od ooz) oz | |
103 | 377 lemma {oz} {ooz} lt = def-subst {suc n} {ord→od ooz} {!!} refl refl |
73 | 378 union→ : (X z u : OD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z |
72 | 379 union→ X y u xx with trio< ( od→ord u ) ( osuc ( od→ord y )) |
103 | 380 union→ X y u xx | tri< a ¬b ¬c with osuc-< a {!!} |
74
819da8c08f05
ordinal atomical successor?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
73
diff
changeset
|
381 union→ X y u xx | tri< a ¬b ¬c | () |
103 | 382 union→ X y u xx | tri≈ ¬a b ¬c = lemma b {!!} where |
73 | 383 lemma : {oX ou ooy : Ordinal {suc n}} → ou ≡ ooy → ou o< oX → ooy o< oX |
384 lemma refl lt = lt | |
103 | 385 union→ X y u xx | tri> ¬a ¬b c = ordtrans {suc n} {osuc ( od→ord y )} {od→ord u} {od→ord X} c {!!} |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
386 union← : (X z : OD) (X∋z : Union X ∋ z) → (X ∋ csuc z) ∧ (csuc z ∋ z ) |
103 | 387 union← X z X∋z = record { proj1 = def-subst {suc n} {_} {_} {X} {od→ord (csuc z )} {!!} oiso (sym diso) ; proj2 = union-lemma-u X∋z } |
54 | 388 ψiso : {ψ : OD {suc n} → Set (suc n)} {x y : OD {suc n}} → ψ x → x ≡ y → ψ y |
389 ψiso {ψ} t refl = t | |
390 selection : {ψ : OD → Set (suc n)} {X y : OD} → ((X ∋ y) ∧ ψ y) ⇔ (Select X ψ ∋ y) | |
391 selection {ψ} {X} {y} = record { | |
392 proj1 = λ cond → record { proj1 = proj1 cond ; proj2 = ψiso {ψ} (proj2 cond) (sym oiso) } | |
393 ; proj2 = λ select → record { proj1 = proj1 select ; proj2 = ψiso {ψ} (proj2 select) oiso } | |
394 } | |
93 | 395 replacement : {ψ : OD → OD} (X x : OD) → Replace X ψ ∋ ψ x |
396 replacement {ψ} X x = sup-c< ψ {x} | |
60 | 397 ∅-iso : {x : OD} → ¬ (x == od∅) → ¬ ((ord→od (od→ord x)) == od∅) |
398 ∅-iso {x} neq = subst (λ k → ¬ k) (=-iso {n} ) neq | |
54 | 399 minimul : (x : OD {suc n} ) → ¬ (x == od∅ )→ OD {suc n} |
68 | 400 minimul x not = od∅ |
57 | 401 regularity : (x : OD) (not : ¬ (x == od∅)) → |
402 (x ∋ minimul x not) ∧ (Select (minimul x not) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅) | |
66 | 403 proj1 (regularity x not ) = ¬∅=→∅∈ not |
404 proj2 (regularity x not ) = record { eq→ = reg ; eq← = λ () } where | |
405 reg : {y : Ordinal} → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) y → def od∅ y | |
406 reg {y} t with proj1 t | |
407 ... | x∈∅ = x∈∅ | |
76 | 408 extensionality : {A B : OD {suc n}} → ((z : OD) → (A ∋ z) ⇔ (B ∋ z)) → A == B |
409 eq→ (extensionality {A} {B} eq ) {x} d = def-iso {suc n} {A} {B} (sym diso) (proj1 (eq (ord→od x))) d | |
410 eq← (extensionality {A} {B} eq ) {x} d = def-iso {suc n} {B} {A} (sym diso) (proj2 (eq (ord→od x))) d | |
89 | 411 xx-union : {x : OD {suc n}} → (x , x) ≡ record { def = λ z → z o< osuc (od→ord x) } |
412 xx-union {x} = cong ( λ k → record { def = λ z → z o< k } ) (omxx (od→ord x)) | |
413 xxx-union : {x : OD {suc n}} → (x , (x , x)) ≡ record { def = λ z → z o< osuc (osuc (od→ord x))} | |
414 xxx-union {x} = cong ( λ k → record { def = λ z → z o< k } ) lemma where | |
91 | 415 lemma1 : {x : OD {suc n}} → od→ord x o< od→ord (x , x) |
103 | 416 lemma1 {x} = {!!} |
91 | 417 lemma2 : {x : OD {suc n}} → od→ord (x , x) ≡ osuc (od→ord x) |
418 lemma2 = trans ( cong ( λ k → od→ord k ) xx-union ) (sym ≡-def) | |
89 | 419 lemma : {x : OD {suc n}} → omax (od→ord x) (od→ord (x , x)) ≡ osuc (osuc (od→ord x)) |
91 | 420 lemma {x} = trans ( sym ( omax< (od→ord x) (od→ord (x , x)) lemma1 ) ) ( cong ( λ k → osuc k ) lemma2 ) |
90 | 421 uxxx-union : {x : OD {suc n}} → Union (x , (x , x)) ≡ record { def = λ z → osuc z o< osuc (osuc (od→ord x)) } |
422 uxxx-union {x} = cong ( λ k → record { def = λ z → osuc z o< k } ) lemma where | |
423 lemma : od→ord (x , (x , x)) ≡ osuc (osuc (od→ord x)) | |
91 | 424 lemma = trans ( cong ( λ k → od→ord k ) xxx-union ) (sym ≡-def ) |
425 uxxx-2 : {x : OD {suc n}} → record { def = λ z → osuc z o< osuc (osuc (od→ord x)) } == record { def = λ z → z o< osuc (od→ord x) } | |
426 eq→ ( uxxx-2 {x} ) {m} lt = proj1 (osuc2 m (od→ord x)) lt | |
427 eq← ( uxxx-2 {x} ) {m} lt = proj2 (osuc2 m (od→ord x)) lt | |
428 uxxx-ord : {x : OD {suc n}} → od→ord (Union (x , (x , x))) ≡ osuc (od→ord x) | |
429 uxxx-ord {x} = trans (cong (λ k → od→ord k ) uxxx-union) (==→o≡ (subst₂ (λ j k → j == k ) (sym oiso) (sym od≡-def ) uxxx-2 )) | |
430 omega = record { lv = Suc Zero ; ord = Φ 1 } | |
78
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
431 infinite : OD {suc n} |
91 | 432 infinite = ord→od ( omega ) |
433 infinity∅ : ord→od ( omega ) ∋ od∅ {suc n} | |
95 | 434 infinity∅ = def-subst {suc n} {_} {o∅} {infinite} {od→ord od∅} |
103 | 435 {!!} refl (subst ( λ k → ( k ≡ od→ord od∅ )) diso (cong (λ k → od→ord k) o∅≡od∅ )) |
91 | 436 infinite∋x : (x : OD) → infinite ∋ x → od→ord x o< omega |
437 infinite∋x x lt = subst (λ k → od→ord x o< k ) diso t where | |
438 t : od→ord x o< od→ord (ord→od (omega)) | |
439 t = ∋→o< {n} {infinite} {x} lt | |
440 infinite∋uxxx : (x : OD) → osuc (od→ord x) o< omega → infinite ∋ Union (x , (x , x )) | |
441 infinite∋uxxx x lt = o<∋→ t where | |
442 t : od→ord (Union (x , (x , x))) o< od→ord (ord→od (omega)) | |
443 t = subst (λ k → od→ord (Union (x , (x , x))) o< k ) (sym diso ) ( subst ( λ k → k o< omega ) ( sym (uxxx-ord {x} ) ) lt ) | |
78
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
444 infinity : (x : OD) → infinite ∋ x → infinite ∋ Union (x , (x , x )) |
91 | 445 infinity x lt = infinite∋uxxx x ( lemma (od→ord x) (infinite∋x x lt )) where |
446 lemma : (ox : Ordinal {suc n} ) → ox o< omega → osuc ox o< omega | |
447 lemma record { lv = Zero ; ord = (Φ .0) } (case1 (s≤s x)) = case1 (s≤s z≤n) | |
448 lemma record { lv = Zero ; ord = (OSuc .0 ord₁) } (case1 (s≤s x)) = case1 (s≤s z≤n) | |
449 lemma record { lv = (Suc lv₁) ; ord = (Φ .(Suc lv₁)) } (case1 (s≤s ())) | |
450 lemma record { lv = (Suc lv₁) ; ord = (OSuc .(Suc lv₁) ord₁) } (case1 (s≤s ())) | |
451 lemma record { lv = 1 ; ord = (Φ 1) } (case2 c2) with d<→lv c2 | |
452 lemma record { lv = (Suc Zero) ; ord = (Φ .1) } (case2 ()) | refl | |
103 | 453 -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] -- this form is no good since X is a transitive set |
454 -- ∀ z [ ∀ x ( x ∈ z → ¬ ( x ≈ ∅ ) ) ∧ ∀ x ∀ y ( x , y ∈ z ∧ ¬ ( x ≈ y ) → x ∩ y ≈ ∅ ) → ∃ u ∀ x ( x ∈ z → ∃ t ( u ∩ x) ≈ { t }) ] | |
455 record Choice (z : OD {suc n}) : Set (suc (suc n)) where | |
456 field | |
457 u : {x : OD {suc n}} ( x∈z : x ∈ z ) → OD {suc n} | |
458 t : {x : OD {suc n}} ( x∈z : x ∈ z ) → (x : OD {suc n} ) → OD {suc n} | |
459 choice : { x : OD {suc n} } → ( x∈z : x ∈ z ) → ( u x∈z ∩ x) == { t x∈z x } | |
460 -- choice : {x : OD {suc n}} ( x ∈ z → ¬ ( x ≈ ∅ ) ) → | |
461 -- axiom-of-choice : { X : OD } → ( ¬x∅ : ¬ ( X == od∅ ) ) → { A : OD } → (A∈X : A ∈ X ) → choice ¬x∅ A∈X ∈ A | |
462 -- axiom-of-choice {X} nx {A} lt = ¬∅=→∅∈ {!!} | |
78
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
463 |
103 | 464 |