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