annotate OD.agda @ 257:53b7acd63481

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