annotate OD.agda @ 219:43021d2b8756

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