Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate OD.agda @ 216:5b9b6ef971dd
Cardinal start
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sun, 04 Aug 2019 18:09:00 +0900 |
parents | f15eaa7c5932 |
children | d5668179ee69 |
rev | line source |
---|---|
16 | 1 open import Level |
182
9f3c0e0b2bc9
remove ordinal-definable
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
180
diff
changeset
|
2 module OD 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 | |
213
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
14 open import logic |
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
15 open import nat |
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
16 |
27 | 17 -- Ordinal Definable Set |
11 | 18 |
141 | 19 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
|
20 field |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
21 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
|
22 |
141 | 23 open OD |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
24 |
44
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
25 open Ordinal |
120 | 26 open _∧_ |
213
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
27 open _∨_ |
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
210
diff
changeset
|
28 open Bool |
44
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
29 |
141 | 30 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
|
31 field |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
32 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
|
33 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
|
34 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
35 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
|
36 id x = x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
37 |
141 | 38 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
|
39 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
|
40 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
41 open _==_ |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
42 |
141 | 43 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
|
44 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
|
45 |
141 | 46 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
|
47 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
|
48 |
141 | 49 ⇔→== : {n : Level} { x y : OD {suc n} } → ( {z : Ordinal {suc n}} → def x z ⇔ def y z) → x == y |
120 | 50 eq→ ( ⇔→== {n} {x} {y} eq ) {z} m = proj1 eq m |
51 eq← ( ⇔→== {n} {x} {y} eq ) {z} m = proj2 eq m | |
52 | |
179 | 53 -- Ordinal in OD ( and ZFSet ) Transitive Set |
141 | 54 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
|
55 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
|
56 |
141 | 57 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
|
58 od∅ {n} = Ord o∅ |
40 | 59 |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
60 postulate |
141 | 61 -- OD can be iso to a subset of Ordinal ( by means of Godel Set ) |
62 od→ord : {n : Level} → OD {n} → Ordinal {n} | |
63 ord→od : {n : Level} → Ordinal {n} → OD {n} | |
166 | 64 c<→o< : {n : Level} {x y : OD {n} } → def y ( od→ord x ) → od→ord x o< od→ord y |
65 oiso : {n : Level} {x : OD {n}} → ord→od ( od→ord x ) ≡ x | |
113 | 66 diso : {n : Level} {x : Ordinal {n}} → od→ord ( ord→od x ) ≡ x |
150 | 67 -- we should prove this in agda, but simply put here |
141 | 68 ==→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
|
69 -- next assumption causes ∀ x ∋ ∅ . It menas only an ordinal becomes a set |
159 | 70 -- o<→c< : {n : Level} {x y : Ordinal {n} } → x o< y → def (ord→od y) x |
71 -- ord→od x ≡ Ord x results the same | |
100 | 72 -- supermum as Replacement Axiom |
95 | 73 sup-o : {n : Level } → ( Ordinal {n} → Ordinal {n}) → Ordinal {n} |
98 | 74 sup-o< : {n : Level } → { ψ : Ordinal {n} → Ordinal {n}} → ∀ {x : Ordinal {n}} → ψ x o< sup-o ψ |
111 | 75 -- contra-position of mimimulity of supermum required in Power Set Axiom |
165 | 76 -- sup-x : {n : Level } → ( Ordinal {n} → Ordinal {n}) → Ordinal {n} |
77 -- sup-lb : {n : Level } → { ψ : Ordinal {n} → Ordinal {n}} → {z : Ordinal {n}} → z o< sup-o ψ → z o< osuc (ψ (sup-x ψ)) | |
183 | 78 -- mimimul and x∋minimul is an Axiom of choice |
141 | 79 minimul : {n : Level } → (x : OD {suc n} ) → ¬ (x == od∅ )→ OD {suc n} |
117 | 80 -- this should be ¬ (x == od∅ )→ ∃ ox → x ∋ Ord ox ( minimum of x ) |
141 | 81 x∋minimul : {n : Level } → (x : OD {suc n} ) → ( ne : ¬ (x == od∅ ) ) → def x ( od→ord ( minimul x ne ) ) |
187 | 82 -- minimulity (may proved by ε-induction ) |
141 | 83 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 | 84 |
141 | 85 _∋_ : { n : Level } → ( a x : OD {n} ) → Set n |
95 | 86 _∋_ {n} a x = def a ( od→ord x ) |
87 | |
141 | 88 _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
|
89 x c< a = a ∋ x |
103 | 90 |
141 | 91 _c≤_ : {n : Level} → OD {n} → OD {n} → Set (suc n) |
95 | 92 a c≤ b = (a ≡ b) ∨ ( b ∋ a ) |
93 | |
141 | 94 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
|
95 cseq x = record { def = λ y → def x (osuc y) } where |
113 | 96 |
141 | 97 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 | 98 def-subst df refl refl = df |
99 | |
141 | 100 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
|
101 sup-od ψ = Ord ( sup-o ( λ x → od→ord (ψ (ord→od x ))) ) |
95 | 102 |
141 | 103 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
|
104 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
|
105 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
|
106 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
|
107 lemma = subst₂ (λ j k → j o< k ) refl diso (o<-subst sup-o< refl (sym diso) ) |
28 | 108 |
187 | 109 otrans : {n : Level} {a x y : Ordinal {n} } → def (Ord a) x → def (Ord x) y → def (Ord a) y |
110 otrans x<a y<x = ordtrans y<x x<a | |
123 | 111 |
205 | 112 def→o< : {n : Level } {X : OD {suc n}} → {x : Ordinal {suc n}} → def X x → x o< od→ord X |
113 def→o< {n} {X} {x} lt = o<-subst {suc n} {_} {_} {x} {od→ord X} ( c<→o< ( def-subst {suc n} {X} {x} lt (sym oiso) (sym diso) )) diso diso | |
114 | |
204
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
115 ∅3 : {n : Level} → { x : Ordinal {suc n}} → ( ∀(y : Ordinal {suc n}) → ¬ (y o< x ) ) → x ≡ o∅ {suc n} |
81 | 116 ∅3 {n} {x} = TransFinite {n} c2 c3 x where |
204
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
117 c0 : Nat → Ordinal {suc n} → Set (suc n) |
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
118 c0 lx x = (∀(y : Ordinal {suc n}) → ¬ (y o< x)) → x ≡ o∅ {suc n} |
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
119 c2 : (lx : Nat) → ((x₁ : Ordinal) → x₁ o< ordinal lx (Φ lx) → c0 (lv x₁) (record { lv = lv x₁ ; ord = ord x₁ }))→ c0 lx (record { lv = lx ; ord = Φ lx } ) |
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
120 c2 Zero _ not = refl |
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
121 c2 (Suc lx) _ not with not ( record { lv = lx ; ord = Φ lx } ) |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
122 ... | t with t (case1 ≤-refl ) |
204
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
123 c2 (Suc lx) _ not | t | () |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
124 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
|
125 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
|
126 ... | t with t (case2 Φ< ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
127 c3 lx (Φ .lx) d not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
128 c3 lx (OSuc .lx x₁) d not with not ( record { lv = lx ; ord = OSuc lx x₁ } ) |
34 | 129 ... | 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
|
130 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
|
131 |
57 | 132 ∅5 : {n : Level} → { x : Ordinal {n} } → ¬ ( x ≡ o∅ {n} ) → o∅ {n} o< x |
133 ∅5 {n} {record { lv = Zero ; ord = (Φ .0) }} not = ⊥-elim (not refl) | |
134 ∅5 {n} {record { lv = Zero ; ord = (OSuc .0 ord) }} not = case2 Φ< | |
135 ∅5 {n} {record { lv = (Suc lv) ; ord = ord }} not = case1 (s≤s z≤n) | |
37 | 136 |
46 | 137 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 } |
138 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
|
139 |
51 | 140 -- avoiding lv != Zero error |
141 | 141 orefl : {n : Level} → { x : OD {n} } → { y : Ordinal {n} } → od→ord x ≡ y → od→ord x ≡ y |
51 | 142 orefl refl = refl |
143 | |
141 | 144 ==-iso : {n : Level} → { x y : OD {n} } → ord→od (od→ord x) == ord→od (od→ord y) → x == y |
51 | 145 ==-iso {n} {x} {y} eq = record { |
146 eq→ = λ d → lemma ( eq→ eq (def-subst d (sym oiso) refl )) ; | |
147 eq← = λ d → lemma ( eq← eq (def-subst d (sym oiso) refl )) } | |
148 where | |
141 | 149 lemma : {x : OD {n} } {z : Ordinal {n}} → def (ord→od (od→ord x)) z → def x z |
51 | 150 lemma {x} {z} d = def-subst d oiso refl |
151 | |
141 | 152 =-iso : {n : Level } {x y : OD {suc n} } → (x == y) ≡ (ord→od (od→ord x) == y) |
57 | 153 =-iso {_} {_} {y} = cong ( λ k → k == y ) (sym oiso) |
154 | |
141 | 155 ord→== : {n : Level} → { x y : OD {n} } → od→ord x ≡ od→ord y → x == y |
51 | 156 ord→== {n} {x} {y} eq = ==-iso (lemma (od→ord x) (od→ord y) (orefl eq)) where |
157 lemma : ( ox oy : Ordinal {n} ) → ox ≡ oy → (ord→od ox) == (ord→od oy) | |
158 lemma ox ox refl = eq-refl | |
159 | |
160 o≡→== : {n : Level} → { x y : Ordinal {n} } → x ≡ y → ord→od x == ord→od y | |
161 o≡→== {n} {x} {.x} refl = eq-refl | |
162 | |
141 | 163 c≤-refl : {n : Level} → ( x : OD {n} ) → x c≤ x |
51 | 164 c≤-refl x = case1 refl |
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 |
140
312e27aa3cb5
remove otrans again. start over
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
139
diff
changeset
|
179 ∅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
|
180 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
|
181 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
|
182 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
|
183 |
141 | 184 ∅< : {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
|
185 ∅< {n} {x} {y} d eq with eq→ (eq-trans eq (eq-sym ∅0) ) d |
60 | 186 ∅< {n} {x} {y} d eq | lift () |
57 | 187 |
141 | 188 ∅6 : {n : Level} → { x : OD {suc n} } → ¬ ( x ∋ x ) -- no Russel paradox |
120 | 189 ∅6 {n} {x} x∋x = o<¬≡ refl ( c<→o< {suc n} {x} {x} x∋x ) |
51 | 190 |
141 | 191 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 | 192 def-iso refl t = t |
193 | |
57 | 194 is-o∅ : {n : Level} → ( x : Ordinal {suc n} ) → Dec ( x ≡ o∅ {suc n} ) |
195 is-o∅ {n} record { lv = Zero ; ord = (Φ .0) } = yes refl | |
196 is-o∅ {n} record { lv = Zero ; ord = (OSuc .0 ord₁) } = no ( λ () ) | |
197 is-o∅ {n} record { lv = (Suc lv₁) ; ord = ord } = no (λ()) | |
198 | |
188
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
187
diff
changeset
|
199 ppp : { n : Level } → { p : Set (suc n) } { a : OD {suc n} } → record { def = λ x → p } ∋ a → p |
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
187
diff
changeset
|
200 ppp {n} {p} {a} d = d |
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
187
diff
changeset
|
201 |
189
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
202 -- |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
203 -- Axiom of choice in intutionistic logic implies the exclude middle |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
204 -- https://plato.stanford.edu/entries/axiom-choice/#AxiChoLog |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
205 -- |
190 | 206 p∨¬p : { n : Level } → ( p : Set (suc n) ) → p ∨ ( ¬ p ) -- assuming axiom of choice |
189
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
207 p∨¬p {n} p with is-o∅ ( od→ord ( record { def = λ x → p } )) |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
208 p∨¬p {n} p | yes eq = case2 (¬p eq) where |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
209 ps = record { def = λ x → p } |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
210 lemma : ps == od∅ → p → ⊥ |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
211 lemma eq p0 = ¬x<0 {n} {od→ord ps} (eq→ eq p0 ) |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
212 ¬p : (od→ord ps ≡ o∅) → p → ⊥ |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
213 ¬p eq = lemma ( subst₂ (λ j k → j == k ) oiso o∅≡od∅ ( o≡→== eq )) |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
214 p∨¬p {n} p | no ¬p = case1 (ppp {n} {p} {minimul ps (λ eq → ¬p (eqo∅ eq))} lemma) where |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
215 ps = record { def = λ x → p } |
188
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
187
diff
changeset
|
216 eqo∅ : ps == od∅ {suc n} → od→ord ps ≡ o∅ {suc n} |
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
187
diff
changeset
|
217 eqo∅ eq = subst (λ k → od→ord ps ≡ k) ord-od∅ ( cong (λ k → od→ord k ) (==→o≡ eq)) |
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
187
diff
changeset
|
218 lemma : ps ∋ minimul ps (λ eq → ¬p (eqo∅ eq)) |
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
187
diff
changeset
|
219 lemma = x∋minimul ps (λ eq → ¬p (eqo∅ eq)) |
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
187
diff
changeset
|
220 |
190 | 221 ∋-p : { n : Level } → ( p : Set (suc n) ) → Dec p -- assuming axiom of choice |
189
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
222 ∋-p {n} p with p∨¬p p |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
223 ∋-p {n} p | case1 x = yes x |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
224 ∋-p {n} p | case2 x = no x |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
225 |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
226 double-neg-eilm : {n : Level } {A : Set (suc n)} → ¬ ¬ A → A -- we don't have this in intutionistic logic |
190 | 227 double-neg-eilm {n} {A} notnot with ∋-p A -- assuming axiom of choice |
189
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
228 ... | yes p = p |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
229 ... | no ¬p = ⊥-elim ( notnot ¬p ) |
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
230 |
167 | 231 OrdP : {n : Level} → ( x : Ordinal {suc n} ) ( y : OD {suc n} ) → Dec ( Ord x ∋ y ) |
232 OrdP {n} x y with trio< x (od→ord y) | |
233 OrdP {n} x y | tri< a ¬b ¬c = no ¬c | |
234 OrdP {n} x y | tri≈ ¬a refl ¬c = no ( o<¬≡ refl ) | |
235 OrdP {n} x y | tri> ¬a ¬b c = yes c | |
119 | 236 |
79 | 237 -- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) |
94 | 238 -- 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
|
239 |
141 | 240 in-codomain : {n : Level} → (X : OD {suc n} ) → ( ψ : OD {suc n} → OD {suc n} ) → OD {suc n} |
148 | 241 in-codomain {n} X ψ = record { def = λ x → ¬ ( (y : Ordinal {suc n}) → ¬ ( def X y ∧ ( x ≡ od→ord (ψ (ord→od y ))))) } |
141 | 242 |
96 | 243 -- 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
|
244 |
141 | 245 ZFSubset : {n : Level} → (A x : OD {suc n} ) → OD {suc n} |
191
9eb6a8691f02
choice function cannot jump between ordinal level
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
190
diff
changeset
|
246 ZFSubset A x = record { def = λ y → def A y ∧ def x y } -- roughly x = A → Set |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
96
diff
changeset
|
247 |
141 | 248 Def : {n : Level} → (A : OD {suc n}) → OD {suc n} |
191
9eb6a8691f02
choice function cannot jump between ordinal level
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
190
diff
changeset
|
249 Def {n} A = Ord ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) ) |
96 | 250 |
190 | 251 |
252 _⊆_ : {n : Level} ( A B : OD {suc n} ) → ∀{ x : OD {suc n} } → Set (suc n) | |
253 _⊆_ A B {x} = A ∋ x → B ∋ x | |
254 | |
255 infixr 220 _⊆_ | |
256 | |
257 subset-lemma : {n : Level} → {A x y : OD {suc n} } → ( x ∋ y → ZFSubset A x ∋ y ) ⇔ ( _⊆_ x A {y} ) | |
258 subset-lemma {n} {A} {x} {y} = record { | |
259 proj1 = λ z lt → proj1 (z lt) | |
260 ; proj2 = λ x⊆A lt → record { proj1 = x⊆A lt ; proj2 = lt } | |
261 } | |
262 | |
96 | 263 -- Constructible Set on α |
122 | 264 -- Def X = record { def = λ y → ∀ (x : OD ) → y < X ∧ y < od→ord x } |
265 -- L (Φ 0) = Φ | |
266 -- L (OSuc lv n) = { Def ( L n ) } | |
267 -- L (Φ (Suc n)) = Union (L α) (α < Φ (Suc n) ) | |
141 | 268 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
|
269 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
|
270 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
|
271 L {n} record { lv = (Suc lx) ; ord = (Φ (Suc lx)) } = -- Union ( L α ) |
121 | 272 cseq ( Ord (od→ord (L {n} (record { lv = lx ; ord = Φ lx })))) |
89 | 273 |
167 | 274 -- L0 : {n : Level} → (α : Ordinal {suc n}) → L (osuc α) ∋ L α |
141 | 275 -- L1 : {n : Level} → (α β : Ordinal {suc n}) → α o< β → ∀ (x : OD {suc n}) → L α ∋ x → L β ∋ x |
122 | 276 |
170 | 277 |
141 | 278 OD→ZF : {n : Level} → ZF {suc (suc n)} {suc n} |
279 OD→ZF {n} = record { | |
280 ZFSet = OD {suc n} | |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
281 ; _∋_ = _∋_ |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
282 ; _≈_ = _==_ |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
283 ; ∅ = od∅ |
28 | 284 ; _,_ = _,_ |
285 ; Union = Union | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
286 ; Power = Power |
28 | 287 ; Select = Select |
288 ; Replace = Replace | |
161 | 289 ; infinite = infinite |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
290 ; isZF = isZF |
28 | 291 } where |
144 | 292 ZFSet = OD {suc n} |
141 | 293 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
|
294 Select X ψ = record { def = λ x → ( def X x ∧ ψ ( ord→od x )) } |
141 | 295 Replace : OD {suc n} → (OD {suc n} → OD {suc n} ) → OD {suc n} |
296 Replace X ψ = record { def = λ x → (x o< sup-o ( λ x → od→ord (ψ (ord→od x )))) ∧ def (in-codomain X ψ) x } | |
297 _,_ : 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
|
298 x , y = Ord (omax (od→ord x) (od→ord y)) |
144 | 299 _∩_ : ( A B : ZFSet ) → ZFSet |
145 | 300 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
|
301 Union : OD {suc n} → OD {suc n} |
3e7475fb28db
differeent Union approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
155
diff
changeset
|
302 Union U = record { def = λ x → ¬ (∀ (u : Ordinal ) → ¬ ((def U u) ∧ (def (ord→od u) x))) } |
54 | 303 _∈_ : ( 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
|
304 A ∈ B = B ∋ A |
141 | 305 Power : OD {suc n} → OD {suc n} |
129 | 306 Power A = Replace (Def (Ord (od→ord A))) ( λ x → A ∩ x ) |
103 | 307 {_} : ZFSet → ZFSet |
308 { 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
|
309 |
161 | 310 data infinite-d : ( x : Ordinal {suc n} ) → Set (suc n) where |
311 iφ : infinite-d o∅ | |
312 isuc : {x : Ordinal {suc n} } → infinite-d x → | |
313 infinite-d (od→ord ( Union (ord→od x , (ord→od x , ord→od x ) ) )) | |
314 | |
315 infinite : OD {suc n} | |
316 infinite = record { def = λ x → infinite-d x } | |
317 | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
318 infixr 200 _∈_ |
96 | 319 -- infixr 230 _∩_ _∪_ |
161 | 320 isZF : IsZF (OD {suc n}) _∋_ _==_ od∅ _,_ Union Power Select Replace infinite |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
321 isZF = record { |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
322 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
|
323 ; pair = pair |
72 | 324 ; union→ = union→ |
325 ; union← = union← | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
326 ; empty = empty |
129 | 327 ; power→ = power→ |
76 | 328 ; power← = power← |
186 | 329 ; extensionality = λ {A} {B} {w} → extensionality {A} {B} {w} |
183 | 330 ; ε-induction = ε-induction |
78
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
331 ; infinity∅ = infinity∅ |
160 | 332 ; infinity = infinity |
116 | 333 ; selection = λ {X} {ψ} {y} → selection {X} {ψ} {y} |
135 | 334 ; replacement← = replacement← |
335 ; replacement→ = replacement→ | |
183 | 336 ; choice-func = choice-func |
337 ; choice = choice | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
338 } where |
129 | 339 |
141 | 340 pair : (A B : OD {suc n} ) → ((A , B) ∋ A) ∧ ((A , B) ∋ B) |
87 | 341 proj1 (pair A B ) = omax-x {n} (od→ord A) (od→ord B) |
342 proj2 (pair A B ) = omax-y {n} (od→ord A) (od→ord B) | |
129 | 343 |
167 | 344 empty : {n : Level } (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
|
345 empty x (case1 ()) |
dab56d357fa3
remove o<→c< and add otrans in OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
103
diff
changeset
|
346 empty x (case2 ()) |
129 | 347 |
154 | 348 o<→c< : {x y : Ordinal {suc n}} {z : OD {suc n}}→ x o< y → _⊆_ (Ord x) (Ord y) {z} |
155 | 349 o<→c< lt lt1 = ordtrans lt1 lt |
350 | |
351 ⊆→o< : {x y : Ordinal {suc n}} → (∀ (z : OD) → _⊆_ (Ord x) (Ord y) {z} ) → x o< osuc y | |
352 ⊆→o< {x} {y} lt with trio< x y | |
353 ⊆→o< {x} {y} lt | tri< a ¬b ¬c = ordtrans a <-osuc | |
354 ⊆→o< {x} {y} lt | tri≈ ¬a b ¬c = subst ( λ k → k o< osuc y) (sym b) <-osuc | |
355 ⊆→o< {x} {y} lt | tri> ¬a ¬b c with lt (ord→od y) (o<-subst c (sym diso) refl ) | |
356 ... | ttt = ⊥-elim ( o<¬≡ refl (o<-subst ttt diso refl )) | |
151 | 357 |
144 | 358 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
|
359 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
|
360 ; proj2 = subst ( λ k → def k (od→ord z)) (sym oiso) (proj2 xx) } )) |
159 | 361 union← : (X z : OD) (X∋z : Union X ∋ z) → ¬ ( (u : OD ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) |
166 | 362 union← X z UX∋z = TransFiniteExists _ lemma UX∋z where |
165 | 363 lemma : {y : Ordinal} → def X y ∧ def (ord→od y) (od→ord z) → ¬ ((u : OD) → ¬ (X ∋ u) ∧ (u ∋ z)) |
364 lemma {y} xx not = not (ord→od y) record { proj1 = subst ( λ k → def X k ) (sym diso) (proj1 xx ) ; proj2 = proj2 xx } | |
144 | 365 |
366 ψiso : {ψ : OD {suc n} → Set (suc n)} {x y : OD {suc n}} → ψ x → x ≡ y → ψ y | |
367 ψiso {ψ} t refl = t | |
368 selection : {ψ : OD → Set (suc n)} {X y : OD} → ((X ∋ y) ∧ ψ y) ⇔ (Select X ψ ∋ y) | |
369 selection {ψ} {X} {y} = record { | |
370 proj1 = λ cond → record { proj1 = proj1 cond ; proj2 = ψiso {ψ} (proj2 cond) (sym oiso) } | |
371 ; proj2 = λ select → record { proj1 = proj1 select ; proj2 = ψiso {ψ} (proj2 select) oiso } | |
372 } | |
373 replacement← : {ψ : OD → OD} (X x : OD) → X ∋ x → Replace X ψ ∋ ψ x | |
374 replacement← {ψ} X x lt = record { proj1 = sup-c< ψ {x} ; proj2 = lemma } where | |
375 lemma : def (in-codomain X ψ) (od→ord (ψ x)) | |
150 | 376 lemma not = ⊥-elim ( not ( od→ord x ) (record { proj1 = lt ; proj2 = cong (λ k → od→ord (ψ k)) (sym oiso)} )) |
144 | 377 replacement→ : {ψ : OD → OD} (X x : OD) → (lt : Replace X ψ ∋ x) → ¬ ( (y : OD) → ¬ (x == ψ y)) |
150 | 378 replacement→ {ψ} X x lt = contra-position lemma (lemma2 (proj2 lt)) where |
379 lemma2 : ¬ ((y : Ordinal) → ¬ def X y ∧ ((od→ord x) ≡ od→ord (ψ (ord→od y)))) | |
380 → ¬ ((y : Ordinal) → ¬ def X y ∧ (ord→od (od→ord x) == ψ (ord→od y))) | |
144 | 381 lemma2 not not2 = not ( λ y d → not2 y (record { proj1 = proj1 d ; proj2 = lemma3 (proj2 d)})) where |
150 | 382 lemma3 : {y : Ordinal } → (od→ord x ≡ od→ord (ψ (ord→od y))) → (ord→od (od→ord x) == ψ (ord→od y)) |
144 | 383 lemma3 {y} eq = subst (λ k → ord→od (od→ord x) == k ) oiso (o≡→== eq ) |
150 | 384 lemma : ( (y : OD) → ¬ (x == ψ y)) → ( (y : Ordinal) → ¬ def X y ∧ (ord→od (od→ord x) == ψ (ord→od y)) ) |
385 lemma not y not2 = not (ord→od y) (subst (λ k → k == ψ (ord→od y)) oiso ( proj2 not2 )) | |
144 | 386 |
387 --- | |
388 --- Power Set | |
389 --- | |
390 --- First consider ordinals in OD | |
100 | 391 --- |
392 --- ZFSubset A x = record { def = λ y → def A y ∧ def x y } subset of A | |
393 --- Power X = ord→od ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) ) Power X is a sup of all subset of A | |
394 -- | |
395 -- | |
142 | 396 ∩-≡ : { a b : OD {suc n} } → ({x : OD {suc n} } → (a ∋ x → b ∋ x)) → a == ( b ∩ a ) |
397 ∩-≡ {a} {b} inc = record { | |
398 eq→ = λ {x} x<a → record { proj2 = x<a ; | |
399 proj1 = def-subst {suc n} {_} {_} {b} {x} (inc (def-subst {suc n} {_} {_} {a} {_} x<a refl (sym diso) )) refl diso } ; | |
400 eq← = λ {x} x<a∩b → proj2 x<a∩b } | |
100 | 401 -- |
402 -- we have t ∋ x → A ∋ x means t is a subset of A, that is ZFSubset A t == t | |
403 -- Power A is a sup of ZFSubset A t, so Power A ∋ t | |
404 -- | |
141 | 405 ord-power← : (a : Ordinal ) (t : OD) → ({x : OD} → (t ∋ x → (Ord a) ∋ x)) → Def (Ord a) ∋ t |
129 | 406 ord-power← a t t→A = def-subst {suc n} {_} {_} {Def (Ord a)} {od→ord t} |
127 | 407 lemma refl (lemma1 lemma-eq )where |
129 | 408 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
|
409 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
|
410 eq← lemma-eq {z} w = record { proj2 = w ; |
129 | 411 proj1 = def-subst {suc n} {_} {_} {(Ord a)} {z} |
126 | 412 ( t→A (def-subst {suc n} {_} {_} {t} {od→ord (ord→od z)} w refl (sym diso) )) refl diso } |
141 | 413 lemma1 : {n : Level } {a : Ordinal {suc n}} { t : OD {suc n}} |
129 | 414 → (eq : ZFSubset (Ord a) t == t) → od→ord (ZFSubset (Ord a) (ord→od (od→ord t))) ≡ od→ord t |
150 | 415 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 | 416 lemma : od→ord (ZFSubset (Ord a) (ord→od (od→ord t)) ) o< sup-o (λ x → od→ord (ZFSubset (Ord a) (ord→od x))) |
98 | 417 lemma = sup-o< |
129 | 418 |
144 | 419 -- |
420 -- Every set in OD is a subset of Ordinals | |
421 -- | |
142 | 422 -- Power A = Replace (Def (Ord (od→ord A))) ( λ y → A ∩ y ) |
166 | 423 |
424 -- we have oly double negation form because of the replacement axiom | |
425 -- | |
426 power→ : ( A t : OD) → Power A ∋ t → {x : OD} → t ∋ x → ¬ ¬ (A ∋ x) | |
427 power→ A t P∋t {x} t∋x = TransFiniteExists _ lemma5 lemma4 where | |
142 | 428 a = od→ord A |
429 lemma2 : ¬ ( (y : OD) → ¬ (t == (A ∩ y))) | |
430 lemma2 = replacement→ (Def (Ord (od→ord A))) t P∋t | |
166 | 431 lemma3 : (y : OD) → t == ( A ∩ y ) → ¬ ¬ (A ∋ x) |
432 lemma3 y eq not = not (proj1 (eq→ eq t∋x)) | |
142 | 433 lemma4 : ¬ ((y : Ordinal) → ¬ (t == (A ∩ ord→od y))) |
434 lemma4 not = lemma2 ( λ y not1 → not (od→ord y) (subst (λ k → t == ( A ∩ k )) (sym oiso) not1 )) | |
166 | 435 lemma5 : {y : Ordinal} → t == (A ∩ ord→od y) → ¬ ¬ (def A (od→ord x)) |
436 lemma5 {y} eq not = (lemma3 (ord→od y) eq) not | |
437 | |
142 | 438 power← : (A t : OD) → ({x : OD} → (t ∋ x → A ∋ x)) → Power A ∋ t |
439 power← A t t→A = record { proj1 = lemma1 ; proj2 = lemma2 } where | |
440 a = od→ord A | |
441 lemma0 : {x : OD} → t ∋ x → Ord a ∋ x | |
442 lemma0 {x} t∋x = c<→o< (t→A t∋x) | |
443 lemma3 : Def (Ord a) ∋ t | |
444 lemma3 = ord-power← a t lemma0 | |
152 | 445 lt1 : od→ord (A ∩ ord→od (od→ord t)) o< sup-o (λ x → od→ord (A ∩ ord→od x)) |
446 lt1 = sup-o< {suc n} {λ x → od→ord (A ∩ ord→od x)} {od→ord t} | |
447 lemma4 : (A ∩ ord→od (od→ord t)) ≡ t | |
448 lemma4 = let open ≡-Reasoning in begin | |
449 A ∩ ord→od (od→ord t) | |
450 ≡⟨ cong (λ k → A ∩ k) oiso ⟩ | |
451 A ∩ t | |
452 ≡⟨ sym (==→o≡ ( ∩-≡ t→A )) ⟩ | |
453 t | |
454 ∎ | |
142 | 455 lemma1 : od→ord t o< sup-o (λ x → od→ord (A ∩ ord→od x)) |
152 | 456 lemma1 = subst (λ k → od→ord k o< sup-o (λ x → od→ord (A ∩ ord→od x))) |
457 lemma4 (sup-o< {suc n} {λ x → od→ord (A ∩ ord→od x)} {od→ord t}) | |
142 | 458 lemma2 : def (in-codomain (Def (Ord (od→ord A))) (_∩_ A)) (od→ord t) |
151 | 459 lemma2 not = ⊥-elim ( not (od→ord t) (record { proj1 = lemma3 ; proj2 = lemma6 }) ) where |
460 lemma6 : od→ord t ≡ od→ord (A ∩ ord→od (od→ord t)) | |
461 lemma6 = cong ( λ k → od→ord k ) (==→o≡ (subst (λ k → t == (A ∩ k)) (sym oiso) ( ∩-≡ t→A ))) | |
142 | 462 |
190 | 463 -- assuming axiom of choice |
141 | 464 regularity : (x : OD) (not : ¬ (x == od∅)) → |
115 | 465 (x ∋ minimul x not) ∧ (Select (minimul x not) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅) |
117 | 466 proj1 (regularity x not ) = x∋minimul x not |
467 proj2 (regularity x not ) = record { eq→ = lemma1 ; eq← = λ {y} d → lemma2 {y} d } where | |
468 lemma1 : {x₁ : Ordinal} → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) x₁ → def od∅ x₁ | |
469 lemma1 {x₁} s = ⊥-elim ( minimul-1 x not (ord→od x₁) lemma3 ) where | |
470 lemma3 : def (minimul x not) (od→ord (ord→od x₁)) ∧ def x (od→ord (ord→od x₁)) | |
142 | 471 lemma3 = record { proj1 = def-subst {suc n} {_} {_} {minimul x not} {_} (proj1 s) refl (sym diso) |
472 ; proj2 = proj2 (proj2 s) } | |
117 | 473 lemma2 : {x₁ : Ordinal} → def od∅ x₁ → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) x₁ |
474 lemma2 {y} d = ⊥-elim (empty (ord→od y) (def-subst {suc n} {_} {_} {od∅} {od→ord (ord→od y)} d refl (sym diso) )) | |
129 | 475 |
186 | 476 extensionality0 : {A B : OD {suc n}} → ((z : OD) → (A ∋ z) ⇔ (B ∋ z)) → A == B |
477 eq→ (extensionality0 {A} {B} eq ) {x} d = def-iso {suc n} {A} {B} (sym diso) (proj1 (eq (ord→od x))) d | |
478 eq← (extensionality0 {A} {B} eq ) {x} d = def-iso {suc n} {B} {A} (sym diso) (proj2 (eq (ord→od x))) d | |
479 | |
480 extensionality : {A B w : OD {suc n} } → ((z : OD {suc n}) → (A ∋ z) ⇔ (B ∋ z)) → (w ∋ A) ⇔ (w ∋ B) | |
481 proj1 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) ( ==→o≡ (extensionality0 {A} {B} eq) ) d | |
482 proj2 (extensionality {A} {B} {w} eq ) d = subst (λ k → w ∋ k) (sym ( ==→o≡ (extensionality0 {A} {B} eq) )) d | |
129 | 483 |
161 | 484 infinity∅ : infinite ∋ od∅ {suc n} |
485 infinity∅ = def-subst {suc n} {_} {_} {infinite} {od→ord (od∅ {suc n})} iφ refl lemma where | |
486 lemma : o∅ ≡ od→ord od∅ | |
487 lemma = let open ≡-Reasoning in begin | |
488 o∅ | |
489 ≡⟨ sym diso ⟩ | |
490 od→ord ( ord→od o∅ ) | |
491 ≡⟨ cong ( λ k → od→ord k ) o∅≡od∅ ⟩ | |
492 od→ord od∅ | |
493 ∎ | |
494 infinity : (x : OD) → infinite ∋ x → infinite ∋ Union (x , (x , x )) | |
495 infinity x lt = def-subst {suc n} {_} {_} {infinite} {od→ord (Union (x , (x , x )))} ( isuc {od→ord x} lt ) refl lemma where | |
496 lemma : od→ord (Union (ord→od (od→ord x) , (ord→od (od→ord x) , ord→od (od→ord x)))) | |
497 ≡ od→ord (Union (x , (x , x))) | |
498 lemma = cong (λ k → od→ord (Union ( k , ( k , k ) ))) oiso | |
499 | |
179 | 500 -- Axiom of choice ( is equivalent to the existence of minimul in our case ) |
162 | 501 -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] |
502 choice-func : (X : OD {suc n} ) → {x : OD } → ¬ ( x == od∅ ) → ( X ∋ x ) → OD | |
503 choice-func X {x} not X∋x = minimul x not | |
504 choice : (X : OD {suc n} ) → {A : OD } → ( X∋A : X ∋ A ) → (not : ¬ ( A == od∅ )) → A ∋ choice-func X not X∋A | |
505 choice X {A} X∋A not = x∋minimul A not | |
78
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
506 |
189
540b845ea2de
Axiom of choies implies p ∨ ( ¬ p )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
188
diff
changeset
|
507 -- |
176 | 508 -- another form of regularity |
509 -- | |
510 ε-induction : {n m : Level} { ψ : OD {suc n} → Set m} | |
511 → ( {x : OD {suc n} } → ({ y : OD {suc n} } → x ∋ y → ψ y ) → ψ x ) | |
512 → (x : OD {suc n} ) → ψ x | |
513 ε-induction {n} {m} {ψ} ind x = subst (λ k → ψ k ) oiso (ε-induction-ord (lv (osuc (od→ord x))) (ord (osuc (od→ord x))) <-osuc) where | |
514 ε-induction-ord : (lx : Nat) ( ox : OrdinalD {suc n} lx ) {ly : Nat} {oy : OrdinalD {suc n} ly } | |
515 → (ly < lx) ∨ (oy d< ox ) → ψ (ord→od (record { lv = ly ; ord = oy } ) ) | |
516 ε-induction-ord lx (OSuc lx ox) {ly} {oy} y<x = | |
517 ind {ord→od (record { lv = ly ; ord = oy })} ( λ {y} lt → subst (λ k → ψ k ) oiso (ε-induction-ord lx ox (lemma y lt ))) where | |
197 | 518 lemma : (z : OD) → ord→od record { lv = ly ; ord = oy } ∋ z → od→ord z o< record { lv = lx ; ord = ox } |
519 lemma z lt with osuc-≡< y<x | |
520 lemma z lt | case1 refl = o<-subst (c<→o< lt) refl diso | |
521 lemma z lt | case2 lt1 = ordtrans (o<-subst (c<→o< lt) refl diso) lt1 | |
176 | 522 ε-induction-ord (Suc lx) (Φ (Suc lx)) {ly} {oy} (case1 y<sox ) = |
523 ind {ord→od (record { lv = ly ; ord = oy })} ( λ {y} lt → lemma y lt ) where | |
179 | 524 -- |
525 -- if lv of z if less than x Ok | |
526 -- else lv z = lv x. We can create OSuc of z which is larger than z and less than x in lemma | |
527 -- | |
528 -- lx Suc lx (1) lz(a) <lx by case1 | |
529 -- ly(1) ly(2) (2) lz(b) <lx by case1 | |
530 -- lz(a) lz(b) lz(c) lz(c) <lx by case2 ( ly==lz==lx) | |
531 -- | |
176 | 532 lemma0 : { lx ly : Nat } → ly < Suc lx → lx < ly → ⊥ |
533 lemma0 {Suc lx} {Suc ly} (s≤s lt1) (s≤s lt2) = lemma0 lt1 lt2 | |
534 lemma1 : {n : Level } {ly : Nat} {oy : OrdinalD {suc n} ly} → lv (od→ord (ord→od (record { lv = ly ; ord = oy }))) ≡ ly | |
535 lemma1 {n} {ly} {oy} = let open ≡-Reasoning in begin | |
536 lv (od→ord (ord→od (record { lv = ly ; ord = oy }))) | |
537 ≡⟨ cong ( λ k → lv k ) diso ⟩ | |
538 lv (record { lv = ly ; ord = oy }) | |
539 ≡⟨⟩ | |
540 ly | |
541 ∎ | |
542 lemma : (z : OD) → ord→od record { lv = ly ; ord = oy } ∋ z → ψ z | |
543 lemma z lt with c<→o< {suc n} {z} {ord→od (record { lv = ly ; ord = oy })} lt | |
544 lemma z lt | case1 lz<ly with <-cmp lx ly | |
545 lemma z lt | case1 lz<ly | tri< a ¬b ¬c = ⊥-elim ( lemma0 y<sox a) -- can't happen | |
180 | 546 lemma z lt | case1 lz<ly | tri≈ ¬a refl ¬c = -- ly(1) |
176 | 547 subst (λ k → ψ k ) oiso (ε-induction-ord lx (Φ lx) {_} {ord (od→ord z)} (case1 (subst (λ k → lv (od→ord z) < k ) lemma1 lz<ly ) )) |
180 | 548 lemma z lt | case1 lz<ly | tri> ¬a ¬b c = -- lz(a) |
176 | 549 subst (λ k → ψ k ) oiso (ε-induction-ord lx (Φ lx) {_} {ord (od→ord z)} (case1 (<-trans lz<ly (subst (λ k → k < lx ) (sym lemma1) c)))) |
550 lemma z lt | case2 lz=ly with <-cmp lx ly | |
551 lemma z lt | case2 lz=ly | tri< a ¬b ¬c = ⊥-elim ( lemma0 y<sox a) -- can't happen | |
180 | 552 lemma z lt | case2 lz=ly | tri> ¬a ¬b c with d<→lv lz=ly -- lz(b) |
179 | 553 ... | eq = subst (λ k → ψ k ) oiso |
554 (ε-induction-ord lx (Φ lx) {_} {ord (od→ord z)} (case1 (subst (λ k → k < lx ) (trans (sym lemma1)(sym eq) ) c ))) | |
180 | 555 lemma z lt | case2 lz=ly | tri≈ ¬a lx=ly ¬c with d<→lv lz=ly -- lz(c) |
179 | 556 ... | eq = lemma6 {ly} {Φ lx} {oy} lx=ly (sym (subst (λ k → lv (od→ord z) ≡ k) lemma1 eq)) where |
176 | 557 lemma5 : (ox : OrdinalD lx) → (lv (od→ord z) < lx) ∨ (ord (od→ord z) d< ox) → ψ z |
558 lemma5 ox lt = subst (λ k → ψ k ) oiso (ε-induction-ord lx ox lt ) | |
559 lemma6 : { ly : Nat } { ox : OrdinalD {suc n} lx } { oy : OrdinalD {suc n} ly } → | |
179 | 560 lx ≡ ly → ly ≡ lv (od→ord z) → ψ z |
176 | 561 lemma6 {ly} {ox} {oy} refl refl = lemma5 (OSuc lx (ord (od→ord z)) ) (case2 s<refl) |
562 | |
210
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
563 --- |
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
564 --- With assuption of OD is ordered, p ∨ ( ¬ p ) <=> axiom of choice |
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
565 --- |
196 | 566 record choiced {n : Level} ( X : OD {suc n}) : Set (suc (suc n)) where |
567 field | |
568 a-choice : OD {suc n} | |
569 is-in : X ∋ a-choice | |
209 | 570 choice-func' : (X : OD {suc n} ) → (p∨¬p : { n : Level } → ( p : Set (suc n) ) → p ∨ ( ¬ p )) → ¬ ( X == od∅ ) → choiced X |
571 choice-func' X p∨¬p not = have_to_find | |
205 | 572 where |
207 | 573 ψ : ( ox : Ordinal {suc n}) → Set (suc (suc n)) |
209 | 574 ψ ox = (( x : Ordinal {suc n}) → x o< ox → ( ¬ def X x )) ∨ choiced X |
207 | 575 lemma-ord : ( ox : Ordinal {suc n} ) → ψ ox |
576 lemma-ord ox = TransFinite {n} {suc (suc n)} {ψ} caseΦ caseOSuc ox where | |
209 | 577 ∋-p' : (A x : OD {suc n} ) → Dec ( A ∋ x ) |
578 ∋-p' A x with p∨¬p ( A ∋ x ) | |
579 ∋-p' A x | case1 t = yes t | |
580 ∋-p' A x | case2 t = no t | |
581 ∀-imply-or : {n : Level} {A : Ordinal {suc n} → Set (suc n) } {B : Set (suc (suc n)) } | |
582 → ((x : Ordinal {suc n}) → A x ∨ B) → ((x : Ordinal {suc n}) → A x) ∨ B | |
583 ∀-imply-or {n} {A} {B} ∀AB with p∨¬p ((x : Ordinal {suc n}) → A x) | |
584 ∀-imply-or {n} {A} {B} ∀AB | case1 t = case1 t | |
585 ∀-imply-or {n} {A} {B} ∀AB | case2 x = case2 (lemma x) where | |
586 lemma : ¬ ((x : Ordinal {suc n}) → A x) → B | |
587 lemma not with p∨¬p B | |
588 lemma not | case1 b = b | |
589 lemma not | case2 ¬b = ⊥-elim (not (λ x → dont-orb (∀AB x) ¬b )) | |
207 | 590 caseΦ : (lx : Nat) → ( (x : Ordinal {suc n} ) → x o< ordinal lx (Φ lx) → ψ x) → ψ (ordinal lx (Φ lx) ) |
209 | 591 caseΦ lx prev with ∋-p' X ( ord→od (ordinal lx (Φ lx) )) |
592 caseΦ lx prev | yes p = case2 ( record { a-choice = ord→od (ordinal lx (Φ lx)) ; is-in = p } ) | |
593 caseΦ lx prev | no ¬p = lemma where | |
594 lemma1 : (x : Ordinal) → (((Suc (lv x) ≤ lx) ∨ (ord x d< Φ lx) → def X x → ⊥) ∨ choiced X) | |
595 lemma1 x with trio< x (ordinal lx (Φ lx)) | |
210
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
596 lemma1 x | tri< a ¬b ¬c with prev (osuc x) (lemma2 a) where |
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
597 lemma2 : x o< (ordinal lx (Φ lx)) → osuc x o< ordinal lx (Φ lx) |
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
598 lemma2 (case1 lt) = case1 lt |
209 | 599 lemma1 x | tri< a ¬b ¬c | case2 found = case2 found |
210
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
600 lemma1 x | tri< a ¬b ¬c | case1 not_found = case1 ( λ lt df → not_found x <-osuc df ) |
209 | 601 lemma1 x | tri≈ ¬a refl ¬c = case1 ( λ lt → ⊥-elim (o<¬≡ refl lt )) |
602 lemma1 x | tri> ¬a ¬b c = case1 ( λ lt → ⊥-elim (o<> lt c )) | |
603 lemma : ((x : Ordinal) → (Suc (lv x) ≤ lx) ∨ (ord x d< Φ lx) → def X x → ⊥) ∨ choiced X | |
604 lemma = ∀-imply-or lemma1 | |
207 | 605 caseOSuc : (lx : Nat) (x : OrdinalD lx) → ψ ( ordinal lx x ) → ψ ( ordinal lx (OSuc lx x) ) |
209 | 606 caseOSuc lx x prev with ∋-p' X (ord→od record { lv = lx ; ord = x } ) |
607 caseOSuc lx x prev | yes p = case2 (record { a-choice = ord→od record { lv = lx ; ord = x } ; is-in = p }) | |
608 caseOSuc lx x (case1 not_found) | no ¬p = case1 lemma where | |
609 lemma : (y : Ordinal) → (Suc (lv y) ≤ lx) ∨ (ord y d< OSuc lx x) → def X y → ⊥ | |
610 lemma y lt with trio< y (ordinal lx x ) | |
611 lemma y lt | tri< a ¬b ¬c = not_found y a | |
612 lemma y lt | tri≈ ¬a refl ¬c = subst (λ k → def X k → ⊥ ) diso ¬p | |
613 lemma y lt | tri> ¬a ¬b c with osuc-≡< lt | |
614 lemma y lt | tri> ¬a ¬b c | case1 refl = ⊥-elim ( ¬b refl ) | |
615 lemma y lt | tri> ¬a ¬b c | case2 lt1 = ⊥-elim (o<> c lt1 ) | |
616 caseOSuc lx x (case2 found) | no ¬p = case2 found | |
206 | 617 have_to_find : choiced X |
618 have_to_find with lemma-ord (od→ord X ) | |
210
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
619 have_to_find | t = dont-or t ¬¬X∋x where |
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
620 ¬¬X∋x : ¬ ((x : Ordinal) → (Suc (lv x) ≤ lv (od→ord X)) ∨ (ord x d< ord (od→ord X)) → def X x → ⊥) |
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
621 ¬¬X∋x nn = not record { |
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
622 eq→ = λ {x} lt → ⊥-elim (nn x (def→o< lt) lt) |
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
623 ; eq← = λ {x} lt → ⊥-elim ( ¬x<0 lt ) |
2c7d45734e3b
axiom of choice from exclusive middle done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
624 } |
205 | 625 |
216 | 626 ------------ |
627 -- | |
628 -- Onto map | |
629 -- def X x -> xmap | |
630 -- X ---------------------------> Y | |
631 -- ymap <- def Y y | |
632 -- | |
633 record Onto {n : Level } (X Y : OD {suc n}) : Set (suc (suc n)) where | |
634 field | |
635 xmap : (x : Ordinal {suc n}) → Ordinal {suc n} | |
636 ymap : (y : Ordinal {suc n}) → Ordinal {suc n} | |
637 xmap-on-Y : (x : Ordinal {suc n} ) → def X x → def Y (xmap x) | |
638 ymap-on-X : (y : Ordinal {suc n} ) → def Y y → def X (ymap y) | |
639 onto-iso : (y : Ordinal {suc n} ) → def Y y → xmap ( ymap y ) ≡ y | |
640 | |
641 record Cardinal {n : Level } (X : OD {suc n}) : Set (suc (suc n)) where | |
642 field | |
643 cardinal : Ordinal {suc n} | |
644 conto : Onto (Ord cardinal) X | |
645 cmax : ( y : Ordinal {suc n} ) → cardinal o< y → ¬ Onto (Ord y) X | |
646 | |
647 cardinal : {n : Level } (X : OD {suc n}) → Cardinal X | |
648 cardinal {n} X = record { | |
649 cardinal = sup-o ( λ x → cardinal-p x ) | |
650 ; conto = {!!} | |
651 ; cmax = {!!} | |
652 } where | |
653 cardinal-p : (x : Ordinal {suc n}) → Ordinal {suc n} | |
654 cardinal-p x with p∨¬p ( Onto (Ord x) X ) | |
655 cardinal-p x | case1 True = x | |
656 cardinal-p x | case2 False = o∅ |