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