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