Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate HOD.agda @ 134:b52683497a27
replacement in HOD
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sun, 07 Jul 2019 00:31:15 +0900 |
parents | 3849614bef18 |
children | b60b6e8a57b0 |
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 |
120 | 25 open _∧_ |
44
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
26 |
112 | 27 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
|
28 field |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
29 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
|
30 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
|
31 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
32 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
|
33 id x = x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
34 |
112 | 35 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
|
36 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
|
37 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
38 open _==_ |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
39 |
112 | 40 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
|
41 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
|
42 |
112 | 43 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
|
44 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
|
45 |
120 | 46 ⇔→== : {n : Level} { x y : HOD {suc n} } → ( {z : Ordinal {suc n}} → def x z ⇔ def y z) → x == y |
47 eq→ ( ⇔→== {n} {x} {y} eq ) {z} m = proj1 eq m | |
48 eq← ( ⇔→== {n} {x} {y} eq ) {z} m = proj2 eq m | |
49 | |
112 | 50 -- Ordinal in HOD ( and ZFSet ) |
51 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
|
52 Ord {n} a = record { def = λ y → y o< a ; otrans = lemma } where |
114 | 53 lemma : {x : Ordinal} → x o< a → {y : Ordinal} → y o< x → y o< a |
54 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
|
55 |
112 | 56 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
|
57 od∅ {n} = Ord o∅ |
40 | 58 |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
59 postulate |
112 | 60 -- HOD can be iso to a subset of Ordinal ( by means of Godel Set ) |
61 od→ord : {n : Level} → HOD {n} → Ordinal {n} | |
113 | 62 ord→od : {n : Level} → Ordinal {n} → HOD {n} |
129 | 63 c<→o< : {n : Level} {x y : HOD {n} } → def y ( od→ord x ) → od→ord x o< od→ord y |
113 | 64 oiso : {n : Level} {x : HOD {n}} → ord→od ( od→ord x ) ≡ x |
65 diso : {n : Level} {x : Ordinal {n}} → od→ord ( ord→od x ) ≡ x | |
116 | 66 ord-Ord :{n : Level} {x : Ordinal {n}} → x ≡ od→ord (Ord x) |
120 | 67 ==→o≡ : {n : Level} → { x y : HOD {suc n} } → (x == y) → x ≡ y |
109
dab56d357fa3
remove o<→c< and add otrans in OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
103
diff
changeset
|
68 -- 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
|
69 -- o<→c< : {n : Level} {x y : Ordinal {n} } → x o< y → def (ord→od y) x |
100 | 70 -- supermum as Replacement Axiom |
95 | 71 sup-o : {n : Level } → ( Ordinal {n} → Ordinal {n}) → Ordinal {n} |
98 | 72 sup-o< : {n : Level } → { ψ : Ordinal {n} → Ordinal {n}} → ∀ {x : Ordinal {n}} → ψ x o< sup-o ψ |
111 | 73 -- contra-position of mimimulity of supermum required in Power Set Axiom |
98 | 74 sup-x : {n : Level } → ( Ordinal {n} → Ordinal {n}) → Ordinal {n} |
75 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
|
76 -- sup-lb : {n : Level } → ( ψ : Ordinal {n} → Ordinal {n}) → ( ∀ {x : Ordinal {n}} → ψx o< z ) → z o< osuc ( sup-o ψ ) |
117 | 77 minimul : {n : Level } → (x : HOD {suc n} ) → ¬ (x == od∅ )→ HOD {suc n} |
78 -- this should be ¬ (x == od∅ )→ ∃ ox → x ∋ Ord ox ( minimum of x ) | |
79 x∋minimul : {n : Level } → (x : HOD {suc n} ) → ( ne : ¬ (x == od∅ ) ) → def x ( od→ord ( minimul x ne ) ) | |
80 minimul-1 : {n : Level } → (x : HOD {suc n} ) → ( ne : ¬ (x == od∅ ) ) → (y : HOD {suc n}) → ¬ ( def (minimul x ne) (od→ord y)) ∧ (def x (od→ord y) ) | |
123 | 81 -- we should prove this in agda, but simply put here |
82 ===-≡ : {n : Level} { x y : HOD {suc n}} → x == y → x ≡ y | |
83 | |
84 Ord-ord : {n : Level } {ox : Ordinal {suc n}} → Ord ox ≡ ord→od ox | |
85 Ord-ord {n} {px} = trans (sym oiso) (cong ( λ k → ord→od k ) (sym ord-Ord)) | |
95 | 86 |
112 | 87 _∋_ : { n : Level } → ( a x : HOD {n} ) → Set n |
95 | 88 _∋_ {n} a x = def a ( od→ord x ) |
89 | |
112 | 90 _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
|
91 x c< a = a ∋ x |
103 | 92 |
112 | 93 _c≤_ : {n : Level} → HOD {n} → HOD {n} → Set (suc n) |
95 | 94 a c≤ b = (a ≡ b) ∨ ( b ∋ a ) |
95 | |
113 | 96 cseq : {n : Level} → HOD {n} → HOD {n} |
118 | 97 cseq x = record { def = λ y → def x (osuc y) ; otrans = lemma } where |
98 lemma : {ox : Ordinal} → def x (osuc ox) → { oy : Ordinal}→ oy o< ox → def x (osuc oy) | |
99 lemma {ox} oox<Ox {oy} oy<ox = otrans x oox<Ox (osucc oy<ox ) | |
113 | 100 |
112 | 101 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 | 102 def-subst df refl refl = df |
103 | |
113 | 104 o<→c< : {n : Level} {x y : Ordinal {n} } → x o< y → Ord y ∋ Ord x |
105 o<→c< {n} {x} {y} lt = subst ( λ k → k o< y ) ord-Ord lt | |
106 | |
112 | 107 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
|
108 sup-od ψ = Ord ( sup-o ( λ x → od→ord (ψ (ord→od x ))) ) |
95 | 109 |
112 | 110 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
|
111 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
|
112 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
|
113 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
|
114 lemma = subst₂ (λ j k → j o< k ) refl diso (o<-subst sup-o< refl (sym diso) ) |
28 | 115 |
123 | 116 o<→o> : {n : Level} → { x y : Ordinal {n} } → (Ord x == Ord y) → x o< y → ⊥ |
117 o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case1 lt) with o<-subst (yx (case1 lt)) ord-Ord refl | |
118 ... | oyx with o<¬≡ refl (c<→o< {n} {Ord x} oyx ) | |
119 ... | () | |
120 o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case2 lt) with o<-subst (yx (case2 lt)) ord-Ord refl | |
121 ... | oyx with o<¬≡ refl (c<→o< {n} {Ord x} oyx ) | |
122 ... | () | |
123 | |
124 Ord==→≡ : {n : Level} { x y : Ordinal {suc n}} → Ord x == Ord y → x ≡ y | |
125 Ord==→≡ {n} {x} {y} eq with trio< x y | |
126 Ord==→≡ {n} {x} {y} eq | tri< a ¬b ¬c = ⊥-elim ( o<→o> eq a ) | |
127 Ord==→≡ {n} {x} {y} eq | tri≈ ¬a b ¬c = b | |
128 Ord==→≡ {n} {x} {y} eq | tri> ¬a ¬b c = ⊥-elim ( o<→o> (eq-sym eq) c ) | |
129 | |
130 | |
37 | 131 ∅3 : {n : Level} → { x : Ordinal {n}} → ( ∀(y : Ordinal {n}) → ¬ (y o< x ) ) → x ≡ o∅ {n} |
81 | 132 ∅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
|
133 c0 : Nat → Ordinal {n} → Set n |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
134 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
|
135 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
|
136 c2 Zero not = refl |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
137 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
|
138 ... | t with t (case1 ≤-refl ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
139 c2 (Suc lx) not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
140 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
|
141 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
|
142 ... | t with t (case2 Φ< ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
143 c3 lx (Φ .lx) d not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
144 c3 lx (OSuc .lx x₁) d not with not ( record { lv = lx ; ord = OSuc lx x₁ } ) |
34 | 145 ... | 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
|
146 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
|
147 |
112 | 148 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
|
149 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 | 150 ... | t = otrans z z∋y (c<→o< {suc n} {x} {y} x∋y ) |
36 | 151 |
57 | 152 ∅5 : {n : Level} → { x : Ordinal {n} } → ¬ ( x ≡ o∅ {n} ) → o∅ {n} o< x |
153 ∅5 {n} {record { lv = Zero ; ord = (Φ .0) }} not = ⊥-elim (not refl) | |
154 ∅5 {n} {record { lv = Zero ; ord = (OSuc .0 ord) }} not = case2 Φ< | |
155 ∅5 {n} {record { lv = (Suc lv) ; ord = ord }} not = case1 (s≤s z≤n) | |
37 | 156 |
46 | 157 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 } |
158 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
|
159 |
51 | 160 -- avoiding lv != Zero error |
112 | 161 orefl : {n : Level} → { x : HOD {n} } → { y : Ordinal {n} } → od→ord x ≡ y → od→ord x ≡ y |
51 | 162 orefl refl = refl |
163 | |
112 | 164 ==-iso : {n : Level} → { x y : HOD {n} } → ord→od (od→ord x) == ord→od (od→ord y) → x == y |
51 | 165 ==-iso {n} {x} {y} eq = record { |
166 eq→ = λ d → lemma ( eq→ eq (def-subst d (sym oiso) refl )) ; | |
167 eq← = λ d → lemma ( eq← eq (def-subst d (sym oiso) refl )) } | |
168 where | |
112 | 169 lemma : {x : HOD {n} } {z : Ordinal {n}} → def (ord→od (od→ord x)) z → def x z |
51 | 170 lemma {x} {z} d = def-subst d oiso refl |
171 | |
112 | 172 =-iso : {n : Level } {x y : HOD {suc n} } → (x == y) ≡ (ord→od (od→ord x) == y) |
57 | 173 =-iso {_} {_} {y} = cong ( λ k → k == y ) (sym oiso) |
174 | |
112 | 175 ord→== : {n : Level} → { x y : HOD {n} } → od→ord x ≡ od→ord y → x == y |
51 | 176 ord→== {n} {x} {y} eq = ==-iso (lemma (od→ord x) (od→ord y) (orefl eq)) where |
177 lemma : ( ox oy : Ordinal {n} ) → ox ≡ oy → (ord→od ox) == (ord→od oy) | |
178 lemma ox ox refl = eq-refl | |
179 | |
180 o≡→== : {n : Level} → { x y : Ordinal {n} } → x ≡ y → ord→od x == ord→od y | |
181 o≡→== {n} {x} {.x} refl = eq-refl | |
182 | |
183 >→¬< : {x y : Nat } → (x < y ) → ¬ ( y < x ) | |
184 >→¬< (s≤s x<y) (s≤s y<x) = >→¬< x<y y<x | |
185 | |
112 | 186 c≤-refl : {n : Level} → ( x : HOD {n} ) → x c≤ x |
51 | 187 c≤-refl x = case1 refl |
188 | |
112 | 189 ∋→o< : {n : Level} → { a x : HOD {suc n} } → a ∋ x → od→ord x o< od→ord a |
91 | 190 ∋→o< {n} {a} {x} lt = t where |
191 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
|
192 t = c<→o< {suc n} {x} {a} lt |
91 | 193 |
80 | 194 o∅≡od∅ : {n : Level} → ord→od (o∅ {suc n}) ≡ od∅ {suc n} |
195 o∅≡od∅ {n} with trio< {n} (o∅ {suc n}) (od→ord (od∅ {suc n} )) | |
196 o∅≡od∅ {n} | tri< a ¬b ¬c = ⊥-elim (lemma a) where | |
197 lemma : o∅ {suc n } o< (od→ord (od∅ {suc n} )) → ⊥ | |
113 | 198 lemma lt with o<→c< lt |
199 lemma lt | t = o<¬≡ refl t | |
80 | 200 o∅≡od∅ {n} | tri≈ ¬a b ¬c = trans (cong (λ k → ord→od k ) b ) oiso |
201 o∅≡od∅ {n} | tri> ¬a ¬b c = ⊥-elim (¬x<0 c) | |
202 | |
112 | 203 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
|
204 o<→¬c> {n} {x} {y} olt clt = o<> olt (c<→o< clt ) where |
51 | 205 |
112 | 206 o≡→¬c< : {n : Level} → { x y : HOD {n} } → (od→ord x ) ≡ ( od→ord y) → ¬ x c< y |
111 | 207 o≡→¬c< {n} {x} {y} oeq lt = o<¬≡ (orefl oeq ) (c<→o< lt) |
54 | 208 |
109
dab56d357fa3
remove o<→c< and add otrans in OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
103
diff
changeset
|
209 ∅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
|
210 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
|
211 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
|
212 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
|
213 |
112 | 214 ∅< : {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
|
215 ∅< {n} {x} {y} d eq with eq→ (eq-trans eq (eq-sym ∅0) ) d |
60 | 216 ∅< {n} {x} {y} d eq | lift () |
57 | 217 |
120 | 218 ∅6 : {n : Level} → { x : HOD {suc n} } → ¬ ( x ∋ x ) -- no Russel paradox |
219 ∅6 {n} {x} x∋x = o<¬≡ refl ( c<→o< {suc n} {x} {x} x∋x ) | |
51 | 220 |
112 | 221 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 | 222 def-iso refl t = t |
223 | |
57 | 224 is-o∅ : {n : Level} → ( x : Ordinal {suc n} ) → Dec ( x ≡ o∅ {suc n} ) |
225 is-o∅ {n} record { lv = Zero ; ord = (Φ .0) } = yes refl | |
226 is-o∅ {n} record { lv = Zero ; ord = (OSuc .0 ord₁) } = no ( λ () ) | |
227 is-o∅ {n} record { lv = (Suc lv₁) ; ord = ord } = no (λ()) | |
228 | |
119 | 229 |
79 | 230 -- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) |
94 | 231 -- 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
|
232 |
112 | 233 csuc : {n : Level} → HOD {suc n} → HOD {suc n} |
122 | 234 csuc x = Ord ( osuc ( od→ord x )) |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
235 |
134 | 236 in-codomain : {n : Level} → (X : HOD {suc n} ) → ( ψ : HOD {suc n} → HOD {suc n} ) → HOD {suc n} |
237 in-codomain {n} X ψ = record { def = λ x → ¬ ( (y : Ordinal {suc n}) → ¬ ( def X y ∧ ( x ≡ od→ord (ψ (Ord y ))))) } | |
238 | |
96 | 239 -- 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
|
240 |
112 | 241 ZFSubset : {n : Level} → (A x : HOD {suc n} ) → HOD {suc n} |
121 | 242 ZFSubset A x = record { def = λ y → def A y ∧ def x y ; otrans = lemma } where |
243 lemma : {z : Ordinal} → def A z ∧ def x z → {y : Ordinal} → y o< z → def A y ∧ def x y | |
244 lemma {z} d {y} y<z = record { proj1 = otrans A (proj1 d) y<z ; proj2 = otrans x (proj2 d) y<z } | |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
245 |
112 | 246 Def : {n : Level} → (A : HOD {suc n}) → HOD {suc n} |
121 | 247 Def {n} A = Ord ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) ) |
96 | 248 |
129 | 249 OrdSubset : {n : Level} → (A x : Ordinal {suc n} ) → ZFSubset (Ord A) (Ord x) ≡ Ord ( minα A x ) |
250 OrdSubset {n} A x = ===-≡ ( record { eq→ = lemma1 ; eq← = lemma2 } ) where | |
251 lemma1 : {y : Ordinal} → def (ZFSubset (Ord A) (Ord x)) y → def (Ord (minα A x)) y | |
252 lemma1 {y} s with trio< A x | |
253 lemma1 {y} s | tri< a ¬b ¬c = proj1 s | |
254 lemma1 {y} s | tri≈ ¬a refl ¬c = proj1 s | |
255 lemma1 {y} s | tri> ¬a ¬b c = proj2 s | |
256 lemma2 : {y : Ordinal} → def (Ord (minα A x)) y → def (ZFSubset (Ord A) (Ord x)) y | |
257 lemma2 {y} lt with trio< A x | |
258 lemma2 {y} lt | tri< a ¬b ¬c = record { proj1 = lt ; proj2 = ordtrans lt a } | |
259 lemma2 {y} lt | tri≈ ¬a refl ¬c = record { proj1 = lt ; proj2 = lt } | |
260 lemma2 {y} lt | tri> ¬a ¬b c = record { proj1 = ordtrans lt c ; proj2 = lt } | |
261 | |
96 | 262 -- Constructible Set on α |
122 | 263 -- Def X = record { def = λ y → ∀ (x : OD ) → y < X ∧ y < od→ord x } |
264 -- L (Φ 0) = Φ | |
265 -- L (OSuc lv n) = { Def ( L n ) } | |
266 -- L (Φ (Suc n)) = Union (L α) (α < Φ (Suc n) ) | |
112 | 267 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
|
268 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
|
269 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
|
270 L {n} record { lv = (Suc lx) ; ord = (Φ (Suc lx)) } = -- Union ( L α ) |
121 | 271 cseq ( Ord (od→ord (L {n} (record { lv = lx ; ord = Φ lx })))) |
89 | 272 |
123 | 273 L00 : {n : Level} → (ox : Ordinal {suc n}) → ox o< sup-o ( λ x → od→ord ( ZFSubset (Ord ox) (ord→od x ))) |
274 L00 {n} ox = o<-subst {suc n} {_} {_} {ox} {sup-o ( λ x → od→ord ( ZFSubset (Ord ox) (ord→od x )))} | |
275 (sup-o< {suc n} {λ x → od→ord ( ZFSubset (Ord ox) (ord→od x ))} {ox} ) (lemma0 ox) refl where | |
276 lemma1 : {n : Level } {ox z : Ordinal {suc n}} → ( def (Ord ox) z ∧ def (ord→od ox) z ) ⇔ def ( Ord ox ) z | |
277 lemma1 {n} {ox} {z} = record { proj1 = proj1 ; proj2 = λ t → record { proj1 = t ; proj2 = subst (λ k → def k z ) Ord-ord t }} | |
122 | 278 lemma0 : {n : Level} → (ox : Ordinal {suc n}) → od→ord (ZFSubset (Ord ox) (ord→od ox)) ≡ ox |
123 | 279 lemma0 {n} ox = trans (cong (λ k → od→ord k) (===-≡ (⇔→== lemma1) )) (sym ord-Ord) |
122 | 280 |
123 | 281 -- L0 : {n : Level} → (α : Ordinal {suc n}) → α o< β → L (osuc α) ∋ L α |
282 -- L1 : {n : Level} → (α β : Ordinal {suc n}) → α o< β → ∀ (x : HOD {suc n}) → L α ∋ x → L β ∋ x | |
122 | 283 |
111 | 284 omega : { n : Level } → Ordinal {n} |
285 omega = record { lv = Suc Zero ; ord = Φ 1 } | |
286 | |
112 | 287 HOD→ZF : {n : Level} → ZF {suc (suc n)} {suc n} |
288 HOD→ZF {n} = record { | |
289 ZFSet = HOD {suc n} | |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
290 ; _∋_ = _∋_ |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
291 ; _≈_ = _==_ |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
292 ; ∅ = od∅ |
28 | 293 ; _,_ = _,_ |
294 ; Union = Union | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
295 ; Power = Power |
28 | 296 ; Select = Select |
297 ; Replace = Replace | |
111 | 298 ; infinite = Ord omega |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
299 ; isZF = isZF |
28 | 300 } where |
112 | 301 Replace : HOD {suc n} → (HOD {suc n} → HOD {suc n} ) → HOD {suc n} |
134 | 302 Replace X ψ = record { def = λ x → (x o< sup-o ( λ x → od→ord (ψ (ord→od x )))) ∧ def (in-codomain X ψ) x } |
115 | 303 Select : (X : HOD {suc n} ) → ((x : HOD {suc n} ) → Set (suc n) ) → HOD {suc n} |
116 | 304 Select X ψ = record { def = λ x → ((y : Ordinal {suc n} ) → X ∋ ord→od y → ψ (ord→od y)) ∧ (X ∋ ord→od x ) ; otrans = lemma } where |
305 lemma : {x : Ordinal} → ((y : Ordinal) → X ∋ ord→od y → ψ (ord→od y)) ∧ (X ∋ ord→od x) → | |
306 {y : Ordinal} → y o< x → ((y₁ : Ordinal) → X ∋ ord→od y₁ → ψ (ord→od y₁)) ∧ (X ∋ ord→od y) | |
115 | 307 lemma {x} select {y} y<x = record { proj1 = proj1 select ; proj2 = y<X } where |
308 y<X : X ∋ ord→od y | |
309 y<X = otrans X (proj2 select) (o<-subst y<x (sym diso) (sym diso) ) | |
112 | 310 _,_ : 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
|
311 x , y = Ord (omax (od→ord x) (od→ord y)) |
112 | 312 Union : HOD {suc n} → HOD {suc n} |
113 | 313 Union U = cseq U |
77 | 314 -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ ( ∀ {x} → t ∋ x → X ∋ x ) |
112 | 315 ZFSet = HOD {suc n} |
54 | 316 _∈_ : ( 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
|
317 A ∈ B = B ∋ A |
54 | 318 _⊆_ : ( 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
|
319 _⊆_ A B {x} = A ∋ x → B ∋ x |
103 | 320 _∩_ : ( A B : ZFSet ) → ZFSet |
115 | 321 A ∩ B = Select (A , B) ( λ x → ( A ∋ x ) ∧ (B ∋ x) ) |
129 | 322 Power : HOD {suc n} → HOD {suc n} |
323 Power A = Replace (Def (Ord (od→ord A))) ( λ x → A ∩ x ) | |
96 | 324 -- _∪_ : ( A B : ZFSet ) → ZFSet |
325 -- A ∪ B = Select (A , B) ( λ x → (A ∋ x) ∨ ( B ∋ x ) ) | |
103 | 326 {_} : ZFSet → ZFSet |
327 { 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
|
328 |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
329 infixr 200 _∈_ |
96 | 330 -- infixr 230 _∩_ _∪_ |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
331 infixr 220 _⊆_ |
112 | 332 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
|
333 isZF = record { |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
334 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
|
335 ; pair = pair |
118 | 336 ; union-u = λ X z UX∋z → union-u {X} {z} UX∋z |
72 | 337 ; union→ = union→ |
338 ; union← = union← | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
339 ; empty = empty |
129 | 340 ; power→ = power→ |
76 | 341 ; power← = power← |
342 ; extensionality = extensionality | |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
343 ; minimul = minimul |
51 | 344 ; regularity = regularity |
78
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
345 ; infinity∅ = infinity∅ |
93 | 346 ; infinity = λ _ → infinity |
116 | 347 ; selection = λ {X} {ψ} {y} → selection {X} {ψ} {y} |
134 | 348 ; replacement← = {!!} |
349 ; replacement→ = {!!} | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
350 } where |
129 | 351 |
112 | 352 pair : (A B : HOD {suc n} ) → ((A , B) ∋ A) ∧ ((A , B) ∋ B) |
87 | 353 proj1 (pair A B ) = omax-x {n} (od→ord A) (od→ord B) |
354 proj2 (pair A B ) = omax-y {n} (od→ord A) (od→ord B) | |
129 | 355 |
112 | 356 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
|
357 empty x (case1 ()) |
dab56d357fa3
remove o<→c< and add otrans in OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
103
diff
changeset
|
358 empty x (case2 ()) |
129 | 359 |
100 | 360 --- |
361 --- ZFSubset A x = record { def = λ y → def A y ∧ def x y } subset of A | |
362 --- Power X = ord→od ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) ) Power X is a sup of all subset of A | |
363 -- | |
109
dab56d357fa3
remove o<→c< and add otrans in OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
103
diff
changeset
|
364 -- if Power A ∋ t, from a propertiy of minimum sup there is osuc ZFSubset A ∋ t |
100 | 365 -- then ZFSubset A ≡ t or ZFSubset A ∋ t. In the former case ZFSubset A ∋ x implies A ∋ x |
128 | 366 -- In case of later, ZFSubset A ∋ t and t ∋ x implies A ∋ x by transitivity |
100 | 367 -- |
129 | 368 POrd : {a : Ordinal } {t : HOD} → Def (Ord a) ∋ t → Def (Ord a) ∋ Ord (od→ord t) |
369 POrd {a} {t} P∋t = o<→c< P∋t | |
370 ord-power→ : (a : Ordinal ) ( t : HOD) → Def (Ord a) ∋ t → {x : HOD} → t ∋ x → Ord a ∋ x | |
371 ord-power→ a t P∋t {x} t∋x with osuc-≡< (sup-lb (POrd P∋t)) | |
127 | 372 ... | case1 eq = proj1 (def-subst (Ltx t∋x) (sym (subst₂ (λ j k → j ≡ k ) oiso oiso ( cong (λ k → ord→od k) (sym eq) ))) refl ) where |
373 Ltx : {n : Level} → {x t : HOD {suc n}} → t ∋ x → Ord (od→ord t) ∋ x | |
374 Ltx {n} {x} {t} lt = c<→o< lt | |
129 | 375 ... | case2 lt = otrans (Ord a) (proj1 (lemma1 lt) ) |
127 | 376 (c<→o< {suc n} {x} {Ord (od→ord t)} (Ltx t∋x) ) where |
129 | 377 sp = sup-x (λ x → od→ord ( ZFSubset (Ord a) (ord→od x))) |
112 | 378 minsup : HOD |
129 | 379 minsup = ZFSubset (Ord a) ( ord→od ( sup-x (λ x → od→ord ( ZFSubset (Ord a) (ord→od x))))) |
127 | 380 Ltx : {n : Level} → {x t : HOD {suc n}} → t ∋ x → Ord (od→ord t) ∋ x |
381 Ltx {n} {x} {t} lt = c<→o< lt | |
130 | 382 -- lemma1 hold because minsup is Ord (minα a sp) |
127 | 383 lemma1 : od→ord (Ord (od→ord t)) o< od→ord minsup → minsup ∋ Ord (od→ord t) |
129 | 384 lemma1 lt with OrdSubset a ( sup-x (λ x → od→ord ( ZFSubset (Ord a) (ord→od x)))) |
385 ... | eq with subst ( λ k → ZFSubset (Ord a) k ≡ Ord (minα a sp)) Ord-ord eq | |
386 ... | eq1 = def-subst {suc n} {_} {_} {minsup} {od→ord (Ord (od→ord t))} (o<→c< lt) lemma2 (sym ord-Ord) where | |
387 lemma2 : Ord (od→ord (ZFSubset (Ord a) (ord→od sp))) ≡ minsup | |
388 lemma2 = let open ≡-Reasoning in begin | |
389 Ord (od→ord (ZFSubset (Ord a) (ord→od sp))) | |
390 ≡⟨ cong (λ k → Ord (od→ord k)) eq1 ⟩ | |
391 Ord (od→ord (Ord (minα a sp))) | |
392 ≡⟨ cong (λ k → Ord (od→ord k)) Ord-ord ⟩ | |
393 Ord (od→ord (ord→od (minα a sp))) | |
394 ≡⟨ cong (λ k → Ord k) diso ⟩ | |
395 Ord (minα a sp) | |
396 ≡⟨ sym eq1 ⟩ | |
397 minsup | |
398 ∎ | |
100 | 399 -- |
400 -- we have t ∋ x → A ∋ x means t is a subset of A, that is ZFSubset A t == t | |
401 -- Power A is a sup of ZFSubset A t, so Power A ∋ t | |
402 -- | |
129 | 403 ord-power← : (a : Ordinal ) (t : HOD) → ({x : HOD} → (t ∋ x → (Ord a) ∋ x)) → Def (Ord a) ∋ t |
404 ord-power← a t t→A = def-subst {suc n} {_} {_} {Def (Ord a)} {od→ord t} | |
127 | 405 lemma refl (lemma1 lemma-eq )where |
129 | 406 lemma-eq : ZFSubset (Ord a) t == t |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
407 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
|
408 eq← lemma-eq {z} w = record { proj2 = w ; |
129 | 409 proj1 = def-subst {suc n} {_} {_} {(Ord a)} {z} |
126 | 410 ( t→A (def-subst {suc n} {_} {_} {t} {od→ord (ord→od z)} w refl (sym diso) )) refl diso } |
129 | 411 lemma1 : {n : Level } {a : Ordinal {suc n}} { t : HOD {suc n}} |
412 → (eq : ZFSubset (Ord a) t == t) → od→ord (ZFSubset (Ord a) (ord→od (od→ord t))) ≡ od→ord t | |
413 lemma1 {n} {a} {t} eq = subst (λ k → od→ord (ZFSubset (Ord a) k) ≡ od→ord t ) (sym oiso) (cong (λ k → od→ord k ) (===-≡ eq )) | |
414 lemma : od→ord (ZFSubset (Ord a) (ord→od (od→ord t)) ) o< sup-o (λ x → od→ord (ZFSubset (Ord a) (ord→od x))) | |
98 | 415 lemma = sup-o< |
129 | 416 |
130 | 417 -- Power A = Replace (Def (Ord (od→ord A))) ( λ x → A ∩ x ) |
129 | 418 power→ : ( A t : HOD) → Power A ∋ t → {x : HOD} → t ∋ x → A ∋ x |
419 power→ = {!!} | |
420 power← : (A t : HOD) → ({x : HOD} → (t ∋ x → A ∋ x)) → Power A ∋ t | |
134 | 421 power← A t t→A = def-subst {suc n} {Replace (Def (Ord a)) ψ} {_} {Power A} {od→ord t} {!!} {!!} lemma1 where |
130 | 422 a = od→ord A |
423 ψ : HOD → HOD | |
424 ψ y = Def (Ord a) ∩ y | |
425 lemma1 : od→ord (Def (Ord a) ∩ t) ≡ od→ord t | |
426 lemma1 = {!!} | |
427 lemma2 : Ord ( sup-o ( λ x → od→ord (ψ (ord→od x )))) ≡ Power A | |
428 lemma2 = {!!} | |
129 | 429 |
118 | 430 union-u : {X z : HOD {suc n}} → (U>z : Union X ∋ z ) → HOD {suc n} |
431 union-u {X} {z} U>z = Ord ( osuc ( od→ord z ) ) | |
112 | 432 union→ : (X z u : HOD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z |
118 | 433 union→ X z u xx with trio< ( od→ord u ) ( osuc ( od→ord z )) |
434 union→ X z u xx | tri< a ¬b ¬c with osuc-< a (c<→o< (proj2 xx)) | |
435 union→ X z u xx | tri< a ¬b ¬c | () | |
122 | 436 union→ X z u xx | tri≈ ¬a b ¬c = def-subst {suc n} {_} {_} {X} {osuc (od→ord z)} (proj1 xx) refl b |
118 | 437 union→ X z u xx | tri> ¬a ¬b c = otrans X (proj1 xx) c |
438 union← : (X z : HOD) (X∋z : Union X ∋ z) → (X ∋ union-u {X} {z} X∋z ) ∧ (union-u {X} {z} X∋z ∋ z ) | |
439 union← X z X∋z = record { proj1 = lemma ; proj2 = <-osuc } where | |
440 lemma : X ∋ union-u {X} {z} X∋z | |
441 lemma = def-subst {suc n} {_} {_} {X} {od→ord (Ord (osuc ( od→ord z )))} X∋z refl ord-Ord | |
129 | 442 |
112 | 443 ψiso : {ψ : HOD {suc n} → Set (suc n)} {x y : HOD {suc n}} → ψ x → x ≡ y → ψ y |
54 | 444 ψiso {ψ} t refl = t |
116 | 445 selection : {X : HOD } {ψ : (x : HOD ) → Set (suc n)} {y : HOD} → (((y₁ : HOD) → X ∋ y₁ → ψ y₁) ∧ (X ∋ y)) ⇔ (Select X ψ ∋ y) |
115 | 446 selection {X} {ψ} {y} = record { proj1 = λ s → record { |
116 | 447 proj1 = λ y1 y1<X → proj1 s (ord→od y1) y1<X ; proj2 = subst (λ k → def X k ) (sym diso) (proj2 s) } |
448 ; proj2 = λ s → record { proj1 = λ y1 dy1 → subst (λ k → ψ k ) oiso ((proj1 s) (od→ord y1) (def-subst {suc n} {_} {_} {X} {_} dy1 refl (sym diso ))) | |
449 ; proj2 = def-subst {suc n} {_} {_} {X} {od→ord y} (proj2 s ) refl diso } } where | |
134 | 450 replacement← : {ψ : HOD → HOD} (X x : HOD) → X ∋ x → Replace X ψ ∋ ψ x |
451 replacement← {ψ} X x lt = record { proj1 = sup-c< ψ {x} ; proj2 = lemma } where | |
452 lemma : def (in-codomain X ψ) (od→ord (ψ x)) | |
453 lemma not = ⊥-elim ( not ( od→ord x ) (record { proj1 = lt ; proj2 = cong (λ k → od→ord (ψ k)) (sym (trans Ord-ord oiso ))} )) | |
454 replacement→ : {ψ : HOD → HOD} (X x : HOD) → (lt : Replace X ψ ∋ x) → ¬ ( (y : HOD) → ¬ (ψ x == y)) | |
455 replacement→ {ψ} X x lt not = ⊥-elim ( not (ψ x) (ord→== refl ) ) | |
129 | 456 |
112 | 457 ∅-iso : {x : HOD} → ¬ (x == od∅) → ¬ ((ord→od (od→ord x)) == od∅) |
60 | 458 ∅-iso {x} neq = subst (λ k → ¬ k) (=-iso {n} ) neq |
112 | 459 regularity : (x : HOD) (not : ¬ (x == od∅)) → |
115 | 460 (x ∋ minimul x not) ∧ (Select (minimul x not) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅) |
117 | 461 proj1 (regularity x not ) = x∋minimul x not |
462 proj2 (regularity x not ) = record { eq→ = lemma1 ; eq← = λ {y} d → lemma2 {y} d } where | |
463 lemma1 : {x₁ : Ordinal} → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) x₁ → def od∅ x₁ | |
464 lemma1 {x₁} s = ⊥-elim ( minimul-1 x not (ord→od x₁) lemma3 ) where | |
465 lemma3 : def (minimul x not) (od→ord (ord→od x₁)) ∧ def x (od→ord (ord→od x₁)) | |
466 lemma3 = proj1 s x₁ (proj2 s) | |
467 lemma2 : {x₁ : Ordinal} → def od∅ x₁ → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) x₁ | |
468 lemma2 {y} d = ⊥-elim (empty (ord→od y) (def-subst {suc n} {_} {_} {od∅} {od→ord (ord→od y)} d refl (sym diso) )) | |
129 | 469 |
112 | 470 extensionality : {A B : HOD {suc n}} → ((z : HOD) → (A ∋ z) ⇔ (B ∋ z)) → A == B |
76 | 471 eq→ (extensionality {A} {B} eq ) {x} d = def-iso {suc n} {A} {B} (sym diso) (proj1 (eq (ord→od x))) d |
472 eq← (extensionality {A} {B} eq ) {x} d = def-iso {suc n} {B} {A} (sym diso) (proj2 (eq (ord→od x))) d | |
129 | 473 |
119 | 474 open import Relation.Binary.PropositionalEquality |
475 uxxx-ord : {x : HOD {suc n}} → {y : Ordinal {suc n}} → def (Union (x , (x , x))) y ⇔ ( y o< osuc (od→ord x) ) | |
476 uxxx-ord {x} {y} = subst (λ k → k ⇔ ( y o< osuc (od→ord x) )) (sym lemma) ( osuc2 y (od→ord x)) where | |
477 lemma : {y : Ordinal {suc n}} → def (Union (x , (x , x))) y ≡ osuc y o< osuc (osuc (od→ord x)) | |
478 lemma {y} = let open ≡-Reasoning in begin | |
479 def (Union (x , (x , x))) y | |
480 ≡⟨⟩ | |
481 def ( Ord ( omax (od→ord x) (od→ord (Ord (omax (od→ord x) (od→ord x) )) ))) ( osuc y ) | |
482 ≡⟨⟩ | |
483 osuc y o< omax (od→ord x) (od→ord (Ord (omax (od→ord x) (od→ord x) )) ) | |
484 ≡⟨ cong (λ k → osuc y o< omax (od→ord x) k ) (sym ord-Ord) ⟩ | |
485 osuc y o< omax (od→ord x) (omax (od→ord x) (od→ord x) ) | |
486 ≡⟨ cong (λ k → osuc y o< k ) (omxxx (od→ord x) ) ⟩ | |
487 osuc y o< osuc (osuc (od→ord x)) | |
488 ∎ | |
112 | 489 infinite : HOD {suc n} |
111 | 490 infinite = Ord omega |
491 infinity∅ : Ord omega ∋ od∅ {suc n} | |
119 | 492 infinity∅ = o<-subst (case1 (s≤s z≤n) ) ord-Ord refl |
112 | 493 infinity : (x : HOD) → infinite ∋ x → infinite ∋ Union (x , (x , x )) |
120 | 494 infinity x lt = o<-subst ( lemma (od→ord x) lt ) eq refl where |
495 eq : osuc (od→ord x) ≡ od→ord (Union (x , (x , x))) | |
496 eq = let open ≡-Reasoning in begin | |
497 osuc (od→ord x) | |
498 ≡⟨ ord-Ord ⟩ | |
499 od→ord (Ord (osuc (od→ord x))) | |
500 ≡⟨ cong ( λ k → od→ord k ) ( sym (==→o≡ ( ⇔→== uxxx-ord ))) ⟩ | |
501 od→ord (Union (x , (x , x))) | |
502 ∎ | |
91 | 503 lemma : (ox : Ordinal {suc n} ) → ox o< omega → osuc ox o< omega |
504 lemma record { lv = Zero ; ord = (Φ .0) } (case1 (s≤s x)) = case1 (s≤s z≤n) | |
505 lemma record { lv = Zero ; ord = (OSuc .0 ord₁) } (case1 (s≤s x)) = case1 (s≤s z≤n) | |
506 lemma record { lv = (Suc lv₁) ; ord = (Φ .(Suc lv₁)) } (case1 (s≤s ())) | |
507 lemma record { lv = (Suc lv₁) ; ord = (OSuc .(Suc lv₁) ord₁) } (case1 (s≤s ())) | |
508 lemma record { lv = 1 ; ord = (Φ 1) } (case2 c2) with d<→lv c2 | |
509 lemma record { lv = (Suc Zero) ; ord = (Φ .1) } (case2 ()) | refl | |
103 | 510 -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] -- this form is no good since X is a transitive set |
511 -- ∀ z [ ∀ x ( x ∈ z → ¬ ( x ≈ ∅ ) ) ∧ ∀ x ∀ y ( x , y ∈ z ∧ ¬ ( x ≈ y ) → x ∩ y ≈ ∅ ) → ∃ u ∀ x ( x ∈ z → ∃ t ( u ∩ x) ≈ { t }) ] | |
112 | 512 record Choice (z : HOD {suc n}) : Set (suc (suc n)) where |
103 | 513 field |
112 | 514 u : {x : HOD {suc n}} ( x∈z : x ∈ z ) → HOD {suc n} |
515 t : {x : HOD {suc n}} ( x∈z : x ∈ z ) → (x : HOD {suc n} ) → HOD {suc n} | |
516 choice : { x : HOD {suc n} } → ( x∈z : x ∈ z ) → ( u x∈z ∩ x) == { t x∈z x } | |
517 -- choice : {x : HOD {suc n}} ( x ∈ z → ¬ ( x ≈ ∅ ) ) → | |
518 -- axiom-of-choice : { X : HOD } → ( ¬x∅ : ¬ ( X == od∅ ) ) → { A : HOD } → (A∈X : A ∈ X ) → choice ¬x∅ A∈X ∈ A | |
103 | 519 -- 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
|
520 |