Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate ordinal.agda @ 224:afc864169325
recover ε-induction
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sat, 10 Aug 2019 12:31:25 +0900 |
parents | 59771eb07df0 |
children | e06b76e5b682 |
rev | line source |
---|---|
34 | 1 {-# OPTIONS --allow-unsolved-metas #-} |
16 | 2 open import Level |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
3 module ordinal where |
3 | 4 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
5 open import zf |
3 | 6 |
23 | 7 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) |
75 | 8 open import Data.Empty |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
9 open import Relation.Binary.PropositionalEquality |
213
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
204
diff
changeset
|
10 open import logic |
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
204
diff
changeset
|
11 open import nat |
3 | 12 |
24 | 13 data OrdinalD {n : Level} : (lv : Nat) → Set n where |
14 Φ : (lv : Nat) → OrdinalD lv | |
15 OSuc : (lv : Nat) → OrdinalD {n} lv → OrdinalD lv | |
3 | 16 |
24 | 17 record Ordinal {n : Level} : Set n where |
202
ed88384b5102
ε-induction like loop again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
184
diff
changeset
|
18 constructor ordinal |
16 | 19 field |
20 lv : Nat | |
24 | 21 ord : OrdinalD {n} lv |
16 | 22 |
24 | 23 data _d<_ {n : Level} : {lx ly : Nat} → OrdinalD {n} lx → OrdinalD {n} ly → Set n where |
24 Φ< : {lx : Nat} → {x : OrdinalD {n} lx} → Φ lx d< OSuc lx x | |
25 s< : {lx : Nat} → {x y : OrdinalD {n} lx} → x d< y → OSuc lx x d< OSuc lx y | |
17 | 26 |
27 open Ordinal | |
28 | |
27 | 29 _o<_ : {n : Level} ( x y : Ordinal ) → Set n |
17 | 30 _o<_ x y = (lv x < lv y ) ∨ ( ord x d< ord y ) |
3 | 31 |
75 | 32 s<refl : {n : Level } {lx : Nat } { x : OrdinalD {n} lx } → x d< OSuc lx x |
33 s<refl {n} {lv} {Φ lv} = Φ< | |
34 s<refl {n} {lv} {OSuc lv x} = s< s<refl | |
35 | |
36 trio<> : {n : Level} → {lx : Nat} {x : OrdinalD {n} lx } { y : OrdinalD lx } → y d< x → x d< y → ⊥ | |
37 trio<> {n} {lx} {.(OSuc lx _)} {.(OSuc lx _)} (s< s) (s< t) = trio<> s t | |
38 trio<> {n} {lx} {.(OSuc lx _)} {.(Φ lx)} Φ< () | |
39 | |
40 d<→lv : {n : Level} {x y : Ordinal {n}} → ord x d< ord y → lv x ≡ lv y | |
41 d<→lv Φ< = refl | |
42 d<→lv (s< lt) = refl | |
43 | |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
41
diff
changeset
|
44 o<-subst : {n : Level } {Z X z x : Ordinal {n}} → Z o< X → Z ≡ z → X ≡ x → z o< x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
41
diff
changeset
|
45 o<-subst df refl refl = df |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
41
diff
changeset
|
46 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
47 open import Data.Nat.Properties |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
48 open import Data.Unit using ( ⊤ ) |
6 | 49 open import Relation.Nullary |
50 | |
51 open import Relation.Binary | |
52 open import Relation.Binary.Core | |
53 | |
24 | 54 o∅ : {n : Level} → Ordinal {n} |
55 o∅ = record { lv = Zero ; ord = Φ Zero } | |
21 | 56 |
39 | 57 open import Relation.Binary.HeterogeneousEquality using (_≅_;refl) |
58 | |
59 ordinal-cong : {n : Level} {x y : Ordinal {n}} → | |
60 lv x ≡ lv y → ord x ≅ ord y → x ≡ y | |
61 ordinal-cong refl refl = refl | |
21 | 62 |
24 | 63 ≡→¬d< : {n : Level} → {lv : Nat} → {x : OrdinalD {n} lv } → x d< x → ⊥ |
64 ≡→¬d< {n} {lx} {OSuc lx y} (s< t) = ≡→¬d< t | |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
65 |
24 | 66 trio<≡ : {n : Level} → {lx : Nat} {x : OrdinalD {n} lx } { y : OrdinalD lx } → x ≡ y → x d< y → ⊥ |
17 | 67 trio<≡ refl = ≡→¬d< |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
68 |
24 | 69 trio>≡ : {n : Level} → {lx : Nat} {x : OrdinalD {n} lx } { y : OrdinalD lx } → x ≡ y → y d< x → ⊥ |
17 | 70 trio>≡ refl = ≡→¬d< |
9
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
71 |
24 | 72 triOrdd : {n : Level} → {lx : Nat} → Trichotomous _≡_ ( _d<_ {n} {lx} {lx} ) |
73 triOrdd {_} {lv} (Φ lv) (Φ lv) = tri≈ ≡→¬d< refl ≡→¬d< | |
74 triOrdd {_} {lv} (Φ lv) (OSuc lv y) = tri< Φ< (λ ()) ( λ lt → trio<> lt Φ< ) | |
75 triOrdd {_} {lv} (OSuc lv x) (Φ lv) = tri> (λ lt → trio<> lt Φ<) (λ ()) Φ< | |
76 triOrdd {_} {lv} (OSuc lv x) (OSuc lv y) with triOrdd x y | |
77 triOrdd {_} {lv} (OSuc lv x) (OSuc lv y) | tri< a ¬b ¬c = tri< (s< a) (λ tx=ty → trio<≡ tx=ty (s< a) ) ( λ lt → trio<> lt (s< a) ) | |
78 triOrdd {_} {lv} (OSuc lv x) (OSuc lv x) | tri≈ ¬a refl ¬c = tri≈ ≡→¬d< refl ≡→¬d< | |
79 triOrdd {_} {lv} (OSuc lv x) (OSuc lv y) | tri> ¬a ¬b c = tri> ( λ lt → trio<> lt (s< c) ) (λ tx=ty → trio>≡ tx=ty (s< c) ) (s< c) | |
9
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
80 |
74
819da8c08f05
ordinal atomical successor?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
73
diff
changeset
|
81 osuc : {n : Level} ( x : Ordinal {n} ) → Ordinal {n} |
75 | 82 osuc record { lv = lx ; ord = ox } = record { lv = lx ; ord = OSuc lx ox } |
74
819da8c08f05
ordinal atomical successor?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
73
diff
changeset
|
83 |
819da8c08f05
ordinal atomical successor?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
73
diff
changeset
|
84 <-osuc : {n : Level} { x : Ordinal {n} } → x o< osuc x |
75 | 85 <-osuc {n} {record { lv = lx ; ord = Φ .lx }} = case2 Φ< |
86 <-osuc {n} {record { lv = lx ; ord = OSuc .lx ox }} = case2 ( s< s<refl ) | |
74
819da8c08f05
ordinal atomical successor?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
73
diff
changeset
|
87 |
203
8edd2a13a7f3
fixing transfinte induction...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
88 o<¬≡ : {n : Level } { ox oy : Ordinal {suc n}} → ox ≡ oy → ox o< oy → ⊥ |
111 | 89 o<¬≡ {_} {ox} {ox} refl (case1 lt) = =→¬< lt |
90 o<¬≡ {_} {ox} {ox} refl (case2 (s< lt)) = trio<≡ refl lt | |
94 | 91 |
92 ¬x<0 : {n : Level} → { x : Ordinal {suc n} } → ¬ ( x o< o∅ {suc n} ) | |
93 ¬x<0 {n} {x} (case1 ()) | |
94 ¬x<0 {n} {x} (case2 ()) | |
95 | |
81 | 96 o<> : {n : Level} → {x y : Ordinal {n} } → y o< x → x o< y → ⊥ |
97 o<> {n} {x} {y} (case1 x₁) (case1 x₂) = nat-<> x₁ x₂ | |
98 o<> {n} {x} {y} (case1 x₁) (case2 x₂) = nat-≡< (sym (d<→lv x₂)) x₁ | |
99 o<> {n} {x} {y} (case2 x₁) (case1 x₂) = nat-≡< (sym (d<→lv x₁)) x₂ | |
100 o<> {n} {record { lv = lv₁ ; ord = .(OSuc lv₁ _) }} {record { lv = .lv₁ ; ord = .(Φ lv₁) }} (case2 Φ<) (case2 ()) | |
101 o<> {n} {record { lv = lv₁ ; ord = .(OSuc lv₁ _) }} {record { lv = .lv₁ ; ord = .(OSuc lv₁ _) }} (case2 (s< y<x)) (case2 (s< x<y)) = | |
102 o<> (case2 y<x) (case2 x<y) | |
16 | 103 |
24 | 104 orddtrans : {n : Level} {lx : Nat} {x y z : OrdinalD {n} lx } → x d< y → y d< z → x d< z |
105 orddtrans {_} {lx} {.(Φ lx)} {.(OSuc lx _)} {.(OSuc lx _)} Φ< (s< y<z) = Φ< | |
106 orddtrans {_} {lx} {.(OSuc lx _)} {.(OSuc lx _)} {.(OSuc lx _)} (s< x<y) (s< y<z) = s< ( orddtrans x<y y<z ) | |
9
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
107 |
75 | 108 osuc-≡< : {n : Level} { a x : Ordinal {n} } → x o< osuc a → (x ≡ a ) ∨ (x o< a) |
109 osuc-≡< {n} {a} {x} (case1 lt) = case2 (case1 lt) | |
110 osuc-≡< {n} {record { lv = lv₁ ; ord = Φ .lv₁ }} {record { lv = .lv₁ ; ord = .(Φ lv₁) }} (case2 Φ<) = case1 refl | |
111 osuc-≡< {n} {record { lv = lv₁ ; ord = OSuc .lv₁ ord₁ }} {record { lv = .lv₁ ; ord = .(Φ lv₁) }} (case2 Φ<) = case2 (case2 Φ<) | |
112 osuc-≡< {n} {record { lv = lv₁ ; ord = Φ .lv₁ }} {record { lv = .lv₁ ; ord = .(OSuc lv₁ _) }} (case2 (s< ())) | |
113 osuc-≡< {n} {record { lv = la ; ord = OSuc la oa }} {record { lv = la ; ord = (OSuc la ox) }} (case2 (s< lt)) with | |
114 osuc-≡< {n} {record { lv = la ; ord = oa }} {record { lv = la ; ord = ox }} (case2 lt ) | |
115 ... | case1 refl = case1 refl | |
116 ... | case2 (case2 x) = case2 (case2( s< x) ) | |
117 ... | case2 (case1 x) = ⊥-elim (¬a≤a x) where | |
118 | |
119 osuc-< : {n : Level} { x y : Ordinal {n} } → y o< osuc x → x o< y → ⊥ | |
120 osuc-< {n} {x} {y} y<ox x<y with osuc-≡< y<ox | |
121 osuc-< {n} {x} {x} y<ox (case1 x₁) | case1 refl = ⊥-elim (¬a≤a x₁) | |
122 osuc-< {n} {x} {x} (case1 x₂) (case2 x₁) | case1 refl = ⊥-elim (¬a≤a x₂) | |
123 osuc-< {n} {x} {x} (case2 x₂) (case2 x₁) | case1 refl = ≡→¬d< x₁ | |
81 | 124 osuc-< {n} {x} {y} y<ox (case1 x₂) | case2 (case1 x₁) = nat-<> x₂ x₁ |
125 osuc-< {n} {x} {y} y<ox (case1 x₂) | case2 (case2 x₁) = nat-≡< (sym (d<→lv x₁)) x₂ | |
126 osuc-< {n} {x} {y} y<ox (case2 x<y) | case2 y<x = o<> (case2 x<y) y<x | |
75 | 127 |
23 | 128 |
27 | 129 ordtrans : {n : Level} {x y z : Ordinal {n} } → x o< y → y o< z → x o< z |
130 ordtrans {n} {x} {y} {z} (case1 x₁) (case1 x₂) = case1 ( <-trans x₁ x₂ ) | |
81 | 131 ordtrans {n} {x} {y} {z} (case1 x₁) (case2 x₂) with d<→lv x₂ |
27 | 132 ... | refl = case1 x₁ |
81 | 133 ordtrans {n} {x} {y} {z} (case2 x₁) (case1 x₂) with d<→lv x₁ |
27 | 134 ... | refl = case1 x₂ |
135 ordtrans {n} {x} {y} {z} (case2 x₁) (case2 x₂) with d<→lv x₁ | d<→lv x₂ | |
136 ... | refl | refl = case2 ( orddtrans x₁ x₂ ) | |
137 | |
24 | 138 trio< : {n : Level } → Trichotomous {suc n} _≡_ _o<_ |
23 | 139 trio< a b with <-cmp (lv a) (lv b) |
24 | 140 trio< a b | tri< a₁ ¬b ¬c = tri< (case1 a₁) (λ refl → ¬b (cong ( λ x → lv x ) refl ) ) lemma1 where |
141 lemma1 : ¬ (Suc (lv b) ≤ lv a) ∨ (ord b d< ord a) | |
142 lemma1 (case1 x) = ¬c x | |
81 | 143 lemma1 (case2 x) = ⊥-elim (nat-≡< (sym ( d<→lv x )) a₁ ) |
24 | 144 trio< a b | tri> ¬a ¬b c = tri> lemma1 (λ refl → ¬b (cong ( λ x → lv x ) refl ) ) (case1 c) where |
145 lemma1 : ¬ (Suc (lv a) ≤ lv b) ∨ (ord a d< ord b) | |
146 lemma1 (case1 x) = ¬a x | |
81 | 147 lemma1 (case2 x) = ⊥-elim (nat-≡< (sym ( d<→lv x )) c ) |
23 | 148 trio< a b | tri≈ ¬a refl ¬c with triOrdd ( ord a ) ( ord b ) |
24 | 149 trio< record { lv = .(lv b) ; ord = x } b | tri≈ ¬a refl ¬c | tri< a ¬b ¬c₁ = tri< (case2 a) (λ refl → ¬b (lemma1 refl )) lemma2 where |
150 lemma1 : (record { lv = _ ; ord = x }) ≡ b → x ≡ ord b | |
151 lemma1 refl = refl | |
152 lemma2 : ¬ (Suc (lv b) ≤ lv b) ∨ (ord b d< x) | |
153 lemma2 (case1 x) = ¬a x | |
154 lemma2 (case2 x) = trio<> x a | |
155 trio< record { lv = .(lv b) ; ord = x } b | tri≈ ¬a refl ¬c | tri> ¬a₁ ¬b c = tri> lemma2 (λ refl → ¬b (lemma1 refl )) (case2 c) where | |
156 lemma1 : (record { lv = _ ; ord = x }) ≡ b → x ≡ ord b | |
157 lemma1 refl = refl | |
158 lemma2 : ¬ (Suc (lv b) ≤ lv b) ∨ (x d< ord b) | |
159 lemma2 (case1 x) = ¬a x | |
160 lemma2 (case2 x) = trio<> x c | |
161 trio< record { lv = .(lv b) ; ord = x } b | tri≈ ¬a refl ¬c | tri≈ ¬a₁ refl ¬c₁ = tri≈ lemma1 refl lemma1 where | |
162 lemma1 : ¬ (Suc (lv b) ≤ lv b) ∨ (ord b d< ord b) | |
163 lemma1 (case1 x) = ¬a x | |
164 lemma1 (case2 x) = ≡→¬d< x | |
23 | 165 |
86 | 166 |
91 | 167 open _∧_ |
168 | |
203
8edd2a13a7f3
fixing transfinte induction...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
169 TransFinite : {n m : Level} → { ψ : Ordinal {suc n} → Set m } |
8edd2a13a7f3
fixing transfinte induction...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
170 → ( ∀ (lx : Nat ) → ( (x : Ordinal {suc n} ) → x o< ordinal lx (Φ lx) → ψ x ) → ψ ( record { lv = lx ; ord = Φ lx } ) ) |
222
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
171 → ( ∀ (lx : Nat ) → (x : OrdinalD lx ) → ( (y : Ordinal {suc n} ) → y o< ordinal lx (OSuc lx x) → ψ y ) → ψ ( record { lv = lx ; ord = OSuc lx x } ) ) |
22 | 172 → ∀ (x : Ordinal) → ψ x |
204
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
203
diff
changeset
|
173 TransFinite {n} {m} {ψ} caseΦ caseOSuc x = proj1 (TransFinite1 (lv x) (ord x) ) where |
222
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
174 TransFinite1 : (lx : Nat) (ox : OrdinalD lx ) → ψ (ordinal lx ox) ∧ ( ( (x : Ordinal {suc n} ) → x o< ordinal lx ox → ψ x ) ) |
204
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
203
diff
changeset
|
175 TransFinite1 Zero (Φ 0) = record { proj1 = caseΦ Zero lemma ; proj2 = lemma1 } where |
203
8edd2a13a7f3
fixing transfinte induction...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
176 lemma : (x : Ordinal) → x o< ordinal Zero (Φ Zero) → ψ x |
8edd2a13a7f3
fixing transfinte induction...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
177 lemma x (case1 ()) |
8edd2a13a7f3
fixing transfinte induction...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
202
diff
changeset
|
178 lemma x (case2 ()) |
204
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
203
diff
changeset
|
179 lemma1 : (x : Ordinal) → x o< ordinal Zero (Φ Zero) → ψ x |
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
203
diff
changeset
|
180 lemma1 x (case1 ()) |
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
203
diff
changeset
|
181 lemma1 x (case2 ()) |
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
203
diff
changeset
|
182 TransFinite1 (Suc lx) (Φ (Suc lx)) = record { proj1 = caseΦ (Suc lx) (λ x → lemma (lv x) (ord x)) ; proj2 = (λ x → lemma (lv x) (ord x)) } where |
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
203
diff
changeset
|
183 lemma0 : (ly : Nat) (oy : OrdinalD ly ) → ordinal ly oy o< ordinal lx (Φ lx) → ψ (ordinal ly oy) |
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
203
diff
changeset
|
184 lemma0 ly oy lt = proj2 ( TransFinite1 lx (Φ lx) ) (ordinal ly oy) lt |
d4802eb159ff
Transfinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
203
diff
changeset
|
185 lemma : (ly : Nat) (oy : OrdinalD ly ) → ordinal ly oy o< ordinal (Suc lx) (Φ (Suc lx)) → ψ (ordinal ly oy) |
213
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
204
diff
changeset
|
186 lemma lx1 ox1 (case1 lt) with <-∨ lt |
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
204
diff
changeset
|
187 lemma lx (Φ lx) (case1 lt) | case1 refl = proj1 ( TransFinite1 lx (Φ lx) ) |
222
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
188 lemma lx (Φ lx) (case1 lt) | case2 lt1 = lemma0 lx (Φ lx) (case1 lt1) |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
189 lemma lx (OSuc lx ox1) (case1 lt) | case1 refl = caseOSuc lx ox1 lemma2 where |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
190 lemma2 : (y : Ordinal) → (Suc (lv y) ≤ lx) ∨ (ord y d< OSuc lx ox1) → ψ y |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
191 lemma2 y lt1 with osuc-≡< lt1 |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
192 lemma2 y lt1 | case1 refl = lemma lx ox1 (case1 a<sa) |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
193 lemma2 y lt1 | case2 t = proj2 (TransFinite1 lx ox1) y t |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
194 lemma lx1 (OSuc lx1 ox1) (case1 lt) | case2 lt1 = caseOSuc lx1 ox1 lemma2 where |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
195 lemma2 : (y : Ordinal) → (Suc (lv y) ≤ lx1) ∨ (ord y d< OSuc lx1 ox1) → ψ y |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
196 lemma2 y lt2 with osuc-≡< lt2 |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
197 lemma2 y lt2 | case1 refl = lemma lx1 ox1 (ordtrans lt2 (case1 lt)) |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
198 lemma2 y lt2 | case2 (case1 lt3) = proj2 (TransFinite1 lx (Φ lx)) y (case1 (<-trans lt3 lt1 )) |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
199 lemma2 y lt2 | case2 (case2 lt3) with d<→lv lt3 |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
200 ... | refl = proj2 (TransFinite1 lx (Φ lx)) y (case1 lt1) |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
201 TransFinite1 lx (OSuc lx ox) = record { proj1 = caseOSuc lx ox lemma ; proj2 = lemma } where |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
202 lemma : (y : Ordinal) → y o< ordinal lx (OSuc lx ox) → ψ y |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
203 lemma y lt with osuc-≡< lt |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
204 lemma y lt | case1 refl = proj1 ( TransFinite1 lx ox ) |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
205 lemma y lt | case2 lt1 = proj2 ( TransFinite1 lx ox ) y lt1 |
97
f2b579106770
power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
94
diff
changeset
|
206 |
184 | 207 -- we cannot prove this in intutionistic logic |
142 | 208 -- (¬ (∀ y → ¬ ( ψ y ))) → (ψ y → p ) → p |
166 | 209 -- postulate |
210 -- TransFiniteExists : {n m l : Level} → ( ψ : Ordinal {n} → Set m ) | |
211 -- → (exists : ¬ (∀ y → ¬ ( ψ y ) )) | |
212 -- → {p : Set l} ( P : { y : Ordinal {n} } → ψ y → p ) | |
213 -- → p | |
214 -- | |
215 -- Instead we prove | |
216 -- | |
217 TransFiniteExists : {n m l : Level} → ( ψ : Ordinal {n} → Set m ) | |
165 | 218 → {p : Set l} ( P : { y : Ordinal {n} } → ψ y → ¬ p ) |
219 → (exists : ¬ (∀ y → ¬ ( ψ y ) )) | |
220 → ¬ p | |
166 | 221 TransFiniteExists {n} {m} {l} ψ {p} P = contra-position ( λ p y ψy → P {y} ψy p ) |
165 | 222 |
224 | 223 open import Ordinals |
222
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
224 |
224 | 225 C-Ordinal : {n : Level} → Ordinals {suc n} |
222
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
226 C-Ordinal {n} = record { |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
227 ord = Ordinal {suc n} |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
228 ; o∅ = o∅ |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
229 ; osuc = osuc |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
230 ; _o<_ = _o<_ |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
231 ; isOrdinal = record { |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
232 Otrans = ordtrans |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
233 ; OTri = trio< |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
234 ; ¬x<0 = ¬x<0 |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
235 ; <-osuc = <-osuc |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
236 ; osuc-≡< = osuc-≡< |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
237 ; TransFinite = TransFinite1 |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
238 } |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
239 } where |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
240 ord1 : Set (suc n) |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
241 ord1 = Ordinal {suc n} |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
242 TransFinite1 : { ψ : ord1 → Set (suc (suc n)) } |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
243 → ( (x : ord1) → ( (y : ord1 ) → y o< x → ψ y ) → ψ x ) |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
244 → ∀ (x : ord1) → ψ x |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
245 TransFinite1 {ψ} lt x = TransFinite {n} {suc (suc n)} {ψ} caseΦ caseOSuc x where |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
246 caseΦ : (lx : Nat) → ((x₁ : Ordinal) → x₁ o< ordinal lx (Φ lx) → ψ x₁) → |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
247 ψ (record { lv = lx ; ord = Φ lx }) |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
248 caseΦ lx prev = lt (ordinal lx (Φ lx) ) prev |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
249 caseOSuc : (lx : Nat) (x₁ : OrdinalD lx) → ((y : Ordinal) → y o< ordinal lx (OSuc lx x₁) → ψ y) → |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
250 ψ (record { lv = lx ; ord = OSuc lx x₁ }) |
59771eb07df0
TransFinite induction fixed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
220
diff
changeset
|
251 caseOSuc lx ox prev = lt (ordinal lx (OSuc lx ox)) prev |
224 | 252 |
253 module C-Ordinal-with-choice {n : Level} where | |
254 | |
255 import OD | |
256 -- open inOrdinal C-Ordinal | |
257 open OD (C-Ordinal {n}) | |
258 open OD.OD | |
259 | |
260 -- | |
261 -- another form of regularity | |
262 -- | |
263 ε-induction : {m : Level} { ψ : OD → Set m} | |
264 → ( {x : OD } → ({ y : OD } → x ∋ y → ψ y ) → ψ x ) | |
265 → (x : OD ) → ψ x | |
266 ε-induction {m} {ψ} ind x = subst (λ k → ψ k ) oiso (ε-induction-ord (lv (osuc (od→ord x))) (ord (osuc (od→ord x))) <-osuc) where | |
267 ε-induction-ord : (lx : Nat) ( ox : OrdinalD {suc n} lx ) {ly : Nat} {oy : OrdinalD {suc n} ly } | |
268 → (ly < lx) ∨ (oy d< ox ) → ψ (ord→od (record { lv = ly ; ord = oy } ) ) | |
269 ε-induction-ord lx (OSuc lx ox) {ly} {oy} y<x = | |
270 ind {ord→od (record { lv = ly ; ord = oy })} ( λ {y} lt → subst (λ k → ψ k ) oiso (ε-induction-ord lx ox (lemma y lt ))) where | |
271 lemma : (z : OD) → ord→od record { lv = ly ; ord = oy } ∋ z → od→ord z o< record { lv = lx ; ord = ox } | |
272 lemma z lt with osuc-≡< y<x | |
273 lemma z lt | case1 refl = o<-subst (c<→o< lt) refl diso | |
274 lemma z lt | case2 lt1 = ordtrans (o<-subst (c<→o< lt) refl diso) lt1 | |
275 ε-induction-ord (Suc lx) (Φ (Suc lx)) {ly} {oy} (case1 y<sox ) = | |
276 ind {ord→od (record { lv = ly ; ord = oy })} ( λ {y} lt → lemma y lt ) where | |
277 -- | |
278 -- if lv of z if less than x Ok | |
279 -- else lv z = lv x. We can create OSuc of z which is larger than z and less than x in lemma | |
280 -- | |
281 -- lx Suc lx (1) lz(a) <lx by case1 | |
282 -- ly(1) ly(2) (2) lz(b) <lx by case1 | |
283 -- lz(a) lz(b) lz(c) lz(c) <lx by case2 ( ly==lz==lx) | |
284 -- | |
285 lemma0 : { lx ly : Nat } → ly < Suc lx → lx < ly → ⊥ | |
286 lemma0 {Suc lx} {Suc ly} (s≤s lt1) (s≤s lt2) = lemma0 lt1 lt2 | |
287 lemma1 : {ly : Nat} {oy : OrdinalD {suc n} ly} → lv (od→ord (ord→od (record { lv = ly ; ord = oy }))) ≡ ly | |
288 lemma1 {ly} {oy} = let open ≡-Reasoning in begin | |
289 lv (od→ord (ord→od (record { lv = ly ; ord = oy }))) | |
290 ≡⟨ cong ( λ k → lv k ) diso ⟩ | |
291 lv (record { lv = ly ; ord = oy }) | |
292 ≡⟨⟩ | |
293 ly | |
294 ∎ | |
295 lemma : (z : OD) → ord→od record { lv = ly ; ord = oy } ∋ z → ψ z | |
296 lemma z lt with c<→o< {z} {ord→od (record { lv = ly ; ord = oy })} lt | |
297 lemma z lt | case1 lz<ly with <-cmp lx ly | |
298 lemma z lt | case1 lz<ly | tri< a ¬b ¬c = ⊥-elim ( lemma0 y<sox a) -- can't happen | |
299 lemma z lt | case1 lz<ly | tri≈ ¬a refl ¬c = -- ly(1) | |
300 subst (λ k → ψ k ) oiso (ε-induction-ord lx (Φ lx) {_} {ord (od→ord z)} (case1 (subst (λ k → lv (od→ord z) < k ) lemma1 lz<ly ) )) | |
301 lemma z lt | case1 lz<ly | tri> ¬a ¬b c = -- lz(a) | |
302 subst (λ k → ψ k ) oiso (ε-induction-ord lx (Φ lx) {_} {ord (od→ord z)} (case1 (<-trans lz<ly (subst (λ k → k < lx ) (sym lemma1) c)))) | |
303 lemma z lt | case2 lz=ly with <-cmp lx ly | |
304 lemma z lt | case2 lz=ly | tri< a ¬b ¬c = ⊥-elim ( lemma0 y<sox a) -- can't happen | |
305 lemma z lt | case2 lz=ly | tri> ¬a ¬b c with d<→lv lz=ly -- lz(b) | |
306 ... | eq = subst (λ k → ψ k ) oiso | |
307 (ε-induction-ord lx (Φ lx) {_} {ord (od→ord z)} (case1 (subst (λ k → k < lx ) (trans (sym lemma1)(sym eq) ) c ))) | |
308 lemma z lt | case2 lz=ly | tri≈ ¬a lx=ly ¬c with d<→lv lz=ly -- lz(c) | |
309 ... | eq = subst (λ k → ψ k ) oiso ( ε-induction-ord lx (dz oz=lx) {lv (od→ord z)} {ord (od→ord z)} (case2 (dz<dz oz=lx) )) where | |
310 oz=lx : lv (od→ord z) ≡ lx | |
311 oz=lx = let open ≡-Reasoning in begin | |
312 lv (od→ord z) | |
313 ≡⟨ eq ⟩ | |
314 lv (od→ord (ord→od (ordinal ly oy))) | |
315 ≡⟨ cong ( λ k → lv k ) diso ⟩ | |
316 lv (ordinal ly oy) | |
317 ≡⟨ sym lx=ly ⟩ | |
318 lx | |
319 ∎ | |
320 dz : lv (od→ord z) ≡ lx → OrdinalD lx | |
321 dz refl = OSuc lx (ord (od→ord z)) | |
322 dz<dz : (z=x : lv (od→ord z) ≡ lx ) → ord (od→ord z) d< dz z=x | |
323 dz<dz refl = s<refl | |
324 | |
325 --- | |
326 --- With assuption of OD is ordered, p ∨ ( ¬ p ) <=> axiom of choice | |
327 --- | |
328 record choiced ( X : OD) : Set (suc (suc n)) where | |
329 field | |
330 a-choice : OD | |
331 is-in : X ∋ a-choice | |
332 | |
333 choice-func' : (X : OD ) → (p∨¬p : { n : Level } → ( p : Set (suc n) ) → p ∨ ( ¬ p )) → ¬ ( X == od∅ ) → choiced X | |
334 choice-func' X p∨¬p not = have_to_find where | |
335 ψ : ( ox : Ordinal {suc n}) → Set (suc (suc n)) | |
336 ψ ox = (( x : Ordinal {suc n}) → x o< ox → ( ¬ def X x )) ∨ choiced X | |
337 lemma-ord : ( ox : Ordinal {suc n} ) → ψ ox | |
338 lemma-ord ox = TransFinite {n} {suc (suc n)} {ψ} caseΦ caseOSuc1 ox where | |
339 ∋-p' : (A x : OD ) → Dec ( A ∋ x ) | |
340 ∋-p' A x with p∨¬p ( A ∋ x ) | |
341 ∋-p' A x | case1 t = yes t | |
342 ∋-p' A x | case2 t = no t | |
343 ∀-imply-or : {n : Level} {A : Ordinal {suc n} → Set (suc n) } {B : Set (suc (suc n)) } | |
344 → ((x : Ordinal {suc n}) → A x ∨ B) → ((x : Ordinal {suc n}) → A x) ∨ B | |
345 ∀-imply-or {n} {A} {B} ∀AB with p∨¬p ((x : Ordinal {suc n}) → A x) | |
346 ∀-imply-or {n} {A} {B} ∀AB | case1 t = case1 t | |
347 ∀-imply-or {n} {A} {B} ∀AB | case2 x = case2 (lemma x) where | |
348 lemma : ¬ ((x : Ordinal {suc n}) → A x) → B | |
349 lemma not with p∨¬p B | |
350 lemma not | case1 b = b | |
351 lemma not | case2 ¬b = ⊥-elim (not (λ x → dont-orb (∀AB x) ¬b )) | |
352 caseΦ : (lx : Nat) → ( (x : Ordinal {suc n} ) → x o< ordinal lx (Φ lx) → ψ x) → ψ (ordinal lx (Φ lx) ) | |
353 caseΦ lx prev with ∋-p' X ( ord→od (ordinal lx (Φ lx) )) | |
354 caseΦ lx prev | yes p = case2 ( record { a-choice = ord→od (ordinal lx (Φ lx)) ; is-in = p } ) | |
355 caseΦ lx prev | no ¬p = lemma where | |
356 lemma1 : (x : Ordinal) → (((Suc (lv x) ≤ lx) ∨ (ord x d< Φ lx) → def X x → ⊥) ∨ choiced X) | |
357 lemma1 x with trio< x (ordinal lx (Φ lx)) | |
358 lemma1 x | tri< a ¬b ¬c with prev (osuc x) (lemma2 a) where | |
359 lemma2 : x o< (ordinal lx (Φ lx)) → osuc x o< ordinal lx (Φ lx) | |
360 lemma2 (case1 lt) = case1 lt | |
361 lemma1 x | tri< a ¬b ¬c | case2 found = case2 found | |
362 lemma1 x | tri< a ¬b ¬c | case1 not_found = case1 ( λ lt df → not_found x <-osuc df ) | |
363 lemma1 x | tri≈ ¬a refl ¬c = case1 ( λ lt → ⊥-elim (o<¬≡ refl lt )) | |
364 lemma1 x | tri> ¬a ¬b c = case1 ( λ lt → ⊥-elim (o<> lt c )) | |
365 lemma : ((x : Ordinal) → (Suc (lv x) ≤ lx) ∨ (ord x d< Φ lx) → def X x → ⊥) ∨ choiced X | |
366 lemma = ∀-imply-or lemma1 | |
367 caseOSuc : (lx : Nat) (x : OrdinalD lx) → ψ ( ordinal lx x ) → ψ ( ordinal lx (OSuc lx x) ) | |
368 caseOSuc lx x prev with ∋-p' X (ord→od record { lv = lx ; ord = x } ) | |
369 caseOSuc lx x prev | yes p = case2 (record { a-choice = ord→od record { lv = lx ; ord = x } ; is-in = p }) | |
370 caseOSuc lx x (case1 not_found) | no ¬p = case1 lemma where | |
371 lemma : (y : Ordinal) → (Suc (lv y) ≤ lx) ∨ (ord y d< OSuc lx x) → def X y → ⊥ | |
372 lemma y lt with trio< y (ordinal lx x ) | |
373 lemma y lt | tri< a ¬b ¬c = not_found y a | |
374 lemma y lt | tri≈ ¬a refl ¬c = subst (λ k → def X k → ⊥ ) diso ¬p | |
375 lemma y lt | tri> ¬a ¬b c with osuc-≡< lt | |
376 lemma y lt | tri> ¬a ¬b c | case1 refl = ⊥-elim ( ¬b refl ) | |
377 lemma y lt | tri> ¬a ¬b c | case2 lt1 = ⊥-elim (o<> c lt1 ) | |
378 caseOSuc lx x (case2 found) | no ¬p = case2 found | |
379 caseOSuc1 : (lx : Nat) (x : OrdinalD lx) → ((y : Ordinal) → y o< ordinal lx (OSuc lx x) → ψ y) → | |
380 ψ (record { lv = lx ; ord = OSuc lx x }) | |
381 caseOSuc1 lx x lt = caseOSuc lx x (lt ( ordinal lx x ) (case2 s<refl)) | |
382 have_to_find : choiced X | |
383 have_to_find with lemma-ord (od→ord X ) | |
384 have_to_find | t = dont-or t ¬¬X∋x where | |
385 ¬¬X∋x : ¬ ((x : Ordinal) → (Suc (lv x) ≤ lv (od→ord X)) ∨ (ord x d< ord (od→ord X)) → def X x → ⊥) | |
386 ¬¬X∋x nn = not record { | |
387 eq→ = λ {x} lt → ⊥-elim (nn x (def→o< lt) lt) | |
388 ; eq← = λ {x} lt → ⊥-elim ( ¬x<0 lt ) | |
389 } | |
390 |