Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate ordinal-definable.agda @ 72:f39f1a90d154
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sat, 01 Jun 2019 14:43:05 +0900 |
parents | d088eb66564e |
children | dd430a95610f |
rev | line source |
---|---|
16 | 1 open import Level |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
2 module ordinal-definable where |
3 | 3 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
4 open import zf |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
5 open import ordinal |
3 | 6 |
23 | 7 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ ) |
3 | 8 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
9 open import Relation.Binary.PropositionalEquality |
3 | 10 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
11 open import Data.Nat.Properties |
6 | 12 open import Data.Empty |
13 open import Relation.Nullary | |
14 | |
15 open import Relation.Binary | |
16 open import Relation.Binary.Core | |
17 | |
27 | 18 -- Ordinal Definable Set |
11 | 19 |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
20 record OD {n : Level} : Set (suc n) where |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
21 field |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
22 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
|
23 |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
24 open OD |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
25 open import Data.Unit |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
26 |
44
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
27 open Ordinal |
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
28 |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
29 postulate |
45
33860eb44e47
od∅' {n} = ord→od (o∅ {n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
44
diff
changeset
|
30 od→ord : {n : Level} → OD {n} → Ordinal {n} |
36 | 31 ord→od : {n : Level} → Ordinal {n} → OD {n} |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
32 |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
33 _∋_ : { n : Level } → ( a x : OD {n} ) → Set n |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
34 _∋_ {n} a x = def a ( od→ord x ) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
35 |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
36 _c<_ : { n : Level } → ( a x : OD {n} ) → Set n |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
37 x c< a = a ∋ x |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
38 |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
39 record _==_ {n : Level} ( a b : OD {n} ) : Set n where |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
40 field |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
41 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
|
42 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
|
43 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
44 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
|
45 id x = x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
46 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
47 eq-refl : {n : Level} { x : OD {n} } → x == x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
48 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
|
49 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
50 open _==_ |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
51 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
52 eq-sym : {n : Level} { x y : OD {n} } → x == y → y == x |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
53 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
|
54 |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
55 eq-trans : {n : Level} { x y z : OD {n} } → x == y → y == z → x == z |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
56 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
|
57 |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
58 _c≤_ : {n : Level} → OD {n} → OD {n} → Set (suc n) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
59 a c≤ b = (a ≡ b) ∨ ( b ∋ a ) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
60 |
40 | 61 od∅ : {n : Level} → OD {n} |
62 od∅ {n} = record { def = λ _ → Lift n ⊥ } | |
63 | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
64 postulate |
36 | 65 c<→o< : {n : Level} {x y : OD {n} } → x c< y → od→ord x o< od→ord y |
66 o<→c< : {n : Level} {x y : Ordinal {n} } → x o< y → ord→od x c< ord→od y | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
67 oiso : {n : Level} {x : OD {n}} → ord→od ( od→ord x ) ≡ x |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
68 diso : {n : Level} {x : Ordinal {n}} → od→ord ( ord→od x ) ≡ x |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
69 sup-od : {n : Level } → ( OD {n} → OD {n}) → OD {n} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
70 sup-c< : {n : Level } → ( ψ : OD {n} → OD {n}) → ∀ {x : OD {n}} → ψ x c< sup-od ψ |
40 | 71 ∅-base-def : {n : Level} → def ( ord→od (o∅ {n}) ) ≡ def (od∅ {n}) |
46 | 72 |
60 | 73 congf : {n : Level} {x y : OD {n}} → { f g : OD {n} → OD {n} } → x ≡ y → f ≡ g → f x ≡ g y |
74 congf refl refl = refl | |
75 | |
46 | 76 o∅→od∅ : {n : Level} → ord→od (o∅ {n}) ≡ od∅ {n} |
77 o∅→od∅ {n} = cong ( λ k → record { def = k }) ( ∅-base-def ) | |
78 | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
79 ∅1 : {n : Level} → ( x : OD {n} ) → ¬ ( x c< od∅ {n} ) |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
80 ∅1 {n} x (lift ()) |
28 | 81 |
37 | 82 ∅3 : {n : Level} → { x : Ordinal {n}} → ( ∀(y : Ordinal {n}) → ¬ (y o< x ) ) → x ≡ o∅ {n} |
83 ∅3 {n} {x} = TransFinite {n} c1 c2 c3 x where | |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
84 c0 : Nat → Ordinal {n} → Set n |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
85 c0 lx x = (∀(y : Ordinal {n}) → ¬ (y o< x)) → x ≡ o∅ {n} |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
86 c1 : ∀ (lx : Nat ) → c0 lx (record { lv = Suc lx ; ord = ℵ lx } ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
87 c1 lx 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
|
88 ... | t with t (case1 ≤-refl ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
89 c1 lx not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
90 c2 : (lx : Nat) → c0 lx (record { lv = lx ; ord = Φ lx } ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
91 c2 Zero not = refl |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
92 c2 (Suc lx) 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
|
93 ... | t with t (case1 ≤-refl ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
94 c2 (Suc lx) not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
95 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
|
96 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
|
97 ... | t with t (case2 Φ< ) |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
98 c3 lx (Φ .lx) d not | t | () |
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
99 c3 lx (OSuc .lx x₁) d not with not ( record { lv = lx ; ord = OSuc lx x₁ } ) |
34 | 100 ... | 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
|
101 c3 lx (OSuc .lx x₁) d not | t | () |
34 | 102 c3 (Suc lx) (ℵ lx) d not with not ( record { lv = Suc lx ; ord = OSuc (Suc lx) (Φ (Suc lx)) } ) |
41 | 103 ... | t with t (case2 (s< ℵΦ< )) |
34 | 104 c3 .(Suc lx) (ℵ lx) d not | t | () |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
105 |
36 | 106 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 |
107 def-subst df refl refl = df | |
108 | |
69 | 109 transitive : {n : Level } { z y x : OD {suc n} } → z ∋ y → y ∋ x → z ∋ x |
110 transitive {n} {z} {y} {x} z∋y x∋y with ordtrans ( c<→o< {suc n} {x} {y} x∋y ) ( c<→o< {suc n} {y} {z} z∋y ) | |
36 | 111 ... | t = lemma0 (lemma t) where |
112 lemma : ( od→ord x ) o< ( od→ord z ) → def ( ord→od ( od→ord z )) ( od→ord ( ord→od ( od→ord x ))) | |
113 lemma xo<z = o<→c< xo<z | |
114 lemma0 : def ( ord→od ( od→ord z )) ( od→ord ( ord→od ( od→ord x ))) → def z (od→ord x) | |
69 | 115 lemma0 dz = def-subst {suc n} { ord→od ( od→ord z )} { od→ord ( ord→od ( od→ord x))} dz (oiso) (diso) |
36 | 116 |
41 | 117 record Minimumo {n : Level } (x : Ordinal {n}) : Set (suc n) where |
118 field | |
119 mino : Ordinal {n} | |
120 min<x : mino o< x | |
121 | |
122 ominimal : {n : Level} → (x : Ordinal {n} ) → o∅ o< x → Minimumo {n} x | |
37 | 123 ominimal {n} record { lv = Zero ; ord = (Φ .0) } (case1 ()) |
124 ominimal {n} record { lv = Zero ; ord = (Φ .0) } (case2 ()) | |
125 ominimal {n} record { lv = Zero ; ord = (OSuc .0 ord) } (case1 ()) | |
41 | 126 ominimal {n} record { lv = Zero ; ord = (OSuc .0 ord) } (case2 Φ<) = record { mino = record { lv = Zero ; ord = Φ 0 } ; min<x = case2 Φ< } |
127 ominimal {n} record { lv = (Suc lv) ; ord = (Φ .(Suc lv)) } (case1 (s≤s x)) = record { mino = record { lv = lv ; ord = Φ lv } ; min<x = case1 (s≤s ≤-refl)} | |
37 | 128 ominimal {n} record { lv = (Suc lv) ; ord = (Φ .(Suc lv)) } (case2 ()) |
41 | 129 ominimal {n} record { lv = (Suc lv) ; ord = (OSuc .(Suc lv) ord) } (case1 (s≤s x)) = record { mino = record { lv = (Suc lv) ; ord = ord } ; min<x = case2 s<refl} |
37 | 130 ominimal {n} record { lv = (Suc lv) ; ord = (OSuc .(Suc lv) ord) } (case2 ()) |
44
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
131 ominimal {n} record { lv = (Suc lx) ; ord = (ℵ .lx) } (case1 (s≤s z≤n)) = record { mino = record { lv = Suc lx ; ord = Φ (Suc lx) } ; min<x = case2 ℵΦ< } |
fcac01485f32
od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
43
diff
changeset
|
132 ominimal {n} record { lv = (Suc lx) ; ord = (ℵ .lx) } (case2 ()) |
37 | 133 |
57 | 134 ∅5 : {n : Level} → { x : Ordinal {n} } → ¬ ( x ≡ o∅ {n} ) → o∅ {n} o< x |
135 ∅5 {n} {record { lv = Zero ; ord = (Φ .0) }} not = ⊥-elim (not refl) | |
136 ∅5 {n} {record { lv = Zero ; ord = (OSuc .0 ord) }} not = case2 Φ< | |
137 ∅5 {n} {record { lv = (Suc lv) ; ord = ord }} not = case1 (s≤s z≤n) | |
37 | 138 |
39 | 139 ∅8 : {n : Level} → ( x : Ordinal {n} ) → ¬ x o< o∅ {n} |
140 ∅8 {n} x (case1 ()) | |
141 ∅8 {n} x (case2 ()) | |
142 | |
46 | 143 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 } |
144 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
|
145 |
51 | 146 -- avoiding lv != Zero error |
147 orefl : {n : Level} → { x : OD {n} } → { y : Ordinal {n} } → od→ord x ≡ y → od→ord x ≡ y | |
148 orefl refl = refl | |
149 | |
150 ==-iso : {n : Level} → { x y : OD {n} } → ord→od (od→ord x) == ord→od (od→ord y) → x == y | |
151 ==-iso {n} {x} {y} eq = record { | |
152 eq→ = λ d → lemma ( eq→ eq (def-subst d (sym oiso) refl )) ; | |
153 eq← = λ d → lemma ( eq← eq (def-subst d (sym oiso) refl )) } | |
154 where | |
155 lemma : {x : OD {n} } {z : Ordinal {n}} → def (ord→od (od→ord x)) z → def x z | |
156 lemma {x} {z} d = def-subst d oiso refl | |
157 | |
57 | 158 =-iso : {n : Level } {x y : OD {suc n} } → (x == y) ≡ (ord→od (od→ord x) == y) |
159 =-iso {_} {_} {y} = cong ( λ k → k == y ) (sym oiso) | |
160 | |
51 | 161 ord→== : {n : Level} → { x y : OD {n} } → od→ord x ≡ od→ord y → x == y |
162 ord→== {n} {x} {y} eq = ==-iso (lemma (od→ord x) (od→ord y) (orefl eq)) where | |
163 lemma : ( ox oy : Ordinal {n} ) → ox ≡ oy → (ord→od ox) == (ord→od oy) | |
164 lemma ox ox refl = eq-refl | |
165 | |
166 o≡→== : {n : Level} → { x y : Ordinal {n} } → x ≡ y → ord→od x == ord→od y | |
167 o≡→== {n} {x} {.x} refl = eq-refl | |
168 | |
169 ∅7 : {n : Level} → { x : OD {n} } → od→ord x ≡ o∅ {n} → x == od∅ {n} | |
170 ∅7 {n} {x} eq = record { eq→ = e1 (orefl eq) ; eq← = e2 } where | |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
171 e2 : {y : Ordinal {n}} → def od∅ y → def x y |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
172 e2 {y} (lift ()) |
46 | 173 e1 : {ox y : Ordinal {n}} → ox ≡ o∅ {n} → def x y → def od∅ y |
51 | 174 e1 {o∅} {y} refl x>y = lift ( ∅8 y (o<-subst (c<→o< {n} {ord→od y} {x} (def-subst {n} {x} {y} x>y refl (sym diso))) ord-iso eq )) |
175 | |
176 =→¬< : {x : Nat } → ¬ ( x < x ) | |
177 =→¬< {Zero} () | |
178 =→¬< {Suc x} (s≤s lt) = =→¬< lt | |
179 | |
180 >→¬< : {x y : Nat } → (x < y ) → ¬ ( y < x ) | |
181 >→¬< (s≤s x<y) (s≤s y<x) = >→¬< x<y y<x | |
182 | |
183 c≤-refl : {n : Level} → ( x : OD {n} ) → x c≤ x | |
184 c≤-refl x = case1 refl | |
185 | |
186 o<> : {n : Level } ( ox oy : Ordinal {n}) → ox o< oy → oy o< ox → ⊥ | |
187 o<> ox oy (case1 x<y) (case1 y<x) = >→¬< x<y y<x | |
188 o<> ox oy (case1 x<y) (case2 y<x) with d<→lv y<x | |
189 ... | refl = =→¬< x<y | |
190 o<> ox oy (case2 x<y) (case1 y<x) with d<→lv x<y | |
191 ... | refl = =→¬< y<x | |
192 o<> ox oy (case2 x<y) (case2 y<x) with d<→lv x<y | |
193 ... | refl = trio<> x<y y<x | |
194 | |
54 | 195 o<¬≡ : {n : Level } ( ox oy : Ordinal {n}) → ox ≡ oy → ox o< oy → ⊥ |
51 | 196 o<¬≡ ox ox refl (case1 lt) = =→¬< lt |
197 o<¬≡ ox ox refl (case2 (s< lt)) = trio<≡ refl lt | |
198 | |
54 | 199 o<→o> : {n : Level} → { x y : OD {n} } → (x == y) → (od→ord x ) o< ( od→ord y) → ⊥ |
52 | 200 o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case1 lt) with |
201 yx (def-subst {n} {ord→od (od→ord y)} {od→ord (ord→od (od→ord x))} (o<→c< (case1 lt )) oiso diso ) | |
202 ... | oyx with o<¬≡ (od→ord x) (od→ord x) refl (c<→o< oyx ) | |
203 ... | () | |
204 o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case2 lt) with | |
205 yx (def-subst {n} {ord→od (od→ord y)} {od→ord (ord→od (od→ord x))} (o<→c< (case2 lt )) oiso diso ) | |
206 ... | oyx with o<¬≡ (od→ord x) (od→ord x) refl (c<→o< oyx ) | |
207 ... | () | |
208 | |
51 | 209 o<→¬== : {n : Level} → { x y : OD {n} } → (od→ord x ) o< ( od→ord y) → ¬ (x == y ) |
52 | 210 o<→¬== {n} {x} {y} lt eq = o<→o> eq lt |
51 | 211 |
212 o<→¬c> : {n : Level} → { x y : OD {n} } → (od→ord x ) o< ( od→ord y) → ¬ (y c< x ) | |
213 o<→¬c> {n} {x} {y} olt clt = o<> (od→ord x) (od→ord y) olt (c<→o< clt ) where | |
214 | |
215 o≡→¬c< : {n : Level} → { x y : OD {n} } → (od→ord x ) ≡ ( od→ord y) → ¬ x c< y | |
216 o≡→¬c< {n} {x} {y} oeq lt = o<¬≡ (od→ord x) (od→ord y) (orefl oeq ) (c<→o< lt) | |
217 | |
218 tri-c< : {n : Level} → Trichotomous _==_ (_c<_ {suc n}) | |
219 tri-c< {n} x y with trio< {n} (od→ord x) (od→ord y) | |
220 tri-c< {n} x y | tri< a ¬b ¬c = tri< (def-subst (o<→c< a) oiso diso) (o<→¬== a) ( o<→¬c> a ) | |
221 tri-c< {n} x y | tri≈ ¬a b ¬c = tri≈ (o≡→¬c< b) (ord→== b) (o≡→¬c< (sym b)) | |
222 tri-c< {n} x y | tri> ¬a ¬b c = tri> ( o<→¬c> c) (λ eq → o<→¬== c (eq-sym eq ) ) (def-subst (o<→c< c) oiso diso) | |
223 | |
54 | 224 c<> : {n : Level } { x y : OD {suc n}} → x c< y → y c< x → ⊥ |
225 c<> {n} {x} {y} x<y y<x with tri-c< x y | |
226 c<> {n} {x} {y} x<y y<x | tri< a ¬b ¬c = ¬c y<x | |
227 c<> {n} {x} {y} x<y y<x | tri≈ ¬a b ¬c = o<→o> b ( c<→o< x<y ) | |
228 c<> {n} {x} {y} x<y y<x | tri> ¬a ¬b c = ¬a x<y | |
229 | |
51 | 230 ∅2 : {n : Level} → { x : OD {n} } → o∅ {n} o< od→ord x → ¬ ( x == od∅ {n} ) |
231 ∅2 {n} {x} lt record { eq→ = eq→ ; eq← = eq← } with ominimal (od→ord x ) lt | |
232 ... | min with eq→ ( def-subst (o<→c< (Minimumo.min<x min)) oiso refl ) | |
233 ... | () | |
234 | |
57 | 235 ∅0 : {n : Level} → { x : Ordinal {n} } → o∅ {n} o< x → ¬ ( ord→od x == od∅ {n} ) |
236 ∅0 {n} {x} lt record { eq→ = eq→ ; eq← = eq← } with ominimal x lt | |
237 ... | min with eq→ (o<→c< (Minimumo.min<x min)) | |
238 ... | () | |
60 | 239 |
240 ∅< : {n : Level} → { x y : OD {n} } → def x (od→ord y ) → ¬ ( x == od∅ {n} ) | |
241 ∅< {n} {x} {y} d eq with eq→ eq d | |
242 ∅< {n} {x} {y} d eq | lift () | |
57 | 243 |
51 | 244 |
245 is-od∅ : {n : Level} → ( x : OD {suc n} ) → Dec ( x == od∅ {suc n} ) | |
246 is-od∅ {n} x with trio< {n} (od→ord x) (o∅ {suc n}) | |
247 is-od∅ {n} x | tri≈ ¬a b ¬c = yes ( ∅7 (orefl b) ) | |
248 is-od∅ {n} x | tri< (case1 ()) ¬b ¬c | |
249 is-od∅ {n} x | tri< (case2 ()) ¬b ¬c | |
250 is-od∅ {n} x | tri> ¬a ¬b c = no ( ∅2 c ) | |
46 | 251 |
53 | 252 is-∋ : {n : Level} → ( x y : OD {suc n} ) → Dec ( x ∋ y ) |
253 is-∋ {n} x y with tri-c< x y | |
254 is-∋ {n} x y | tri< a ¬b ¬c = no ¬c | |
255 is-∋ {n} x y | tri≈ ¬a b ¬c = no ¬c | |
256 is-∋ {n} x y | tri> ¬a ¬b c = yes c | |
257 | |
57 | 258 is-o∅ : {n : Level} → ( x : Ordinal {suc n} ) → Dec ( x ≡ o∅ {suc n} ) |
259 is-o∅ {n} record { lv = Zero ; ord = (Φ .0) } = yes refl | |
260 is-o∅ {n} record { lv = Zero ; ord = (OSuc .0 ord₁) } = no ( λ () ) | |
261 is-o∅ {n} record { lv = (Suc lv₁) ; ord = ord } = no (λ()) | |
262 | |
45
33860eb44e47
od∅' {n} = ord→od (o∅ {n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
44
diff
changeset
|
263 open _∧_ |
33860eb44e47
od∅' {n} = ord→od (o∅ {n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
44
diff
changeset
|
264 |
57 | 265 ∅9 : {n : Level} → {x : OD {n} } → ¬ x == od∅ → o∅ o< od→ord x |
266 ∅9 {_} {x} not = ∅5 lemma where | |
38 | 267 lemma : ¬ od→ord x ≡ o∅ |
51 | 268 lemma eq = not ( ∅7 eq ) |
37 | 269 |
66 | 270 ∅10 : {n : Level} → {ox : Ordinal {n}} → (not : ¬ (ord→od ox == od∅)) → o∅ o< ox |
271 ∅10 {n} {ox} not = subst (λ k → o∅ o< k) diso (∅9 not) | |
272 | |
273 ¬∅=→∅∈ : {n : Level} → { x : OD {suc n} } → ¬ ( x == od∅ {suc n} ) → x ∋ od∅ {suc n} | |
274 ¬∅=→∅∈ {n} {x} ne = def-subst (lemma (od→ord x) (subst (λ k → ¬ (k == od∅ {suc n} )) (sym oiso) ne )) oiso refl where | |
275 lemma : (ox : Ordinal {suc n}) → ¬ (ord→od ox == od∅ {suc n} ) → ord→od ox ∋ od∅ {suc n} | |
276 lemma ox ne with is-o∅ ox | |
277 lemma ox ne | yes refl with ne ( ord→== lemma1 ) where | |
278 lemma1 : od→ord (ord→od o∅) ≡ od→ord od∅ | |
279 lemma1 = cong ( λ k → od→ord k ) o∅→od∅ | |
280 lemma o∅ ne | yes refl | () | |
281 lemma ox ne | no ¬p = subst ( λ k → def (ord→od ox) (od→ord k) ) o∅→od∅ (o<→c< (∅5 ¬p)) | |
69 | 282 |
283 -- ∃-set : {n : Level} → ( x : OD {suc n} ) → ( ψ : OD {suc n} → Set (suc n) ) → Set (suc n) | |
284 -- ∃-set = {!!} | |
66 | 285 |
286 -- postulate f-extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality (suc (suc n)) (suc (suc n )) | |
287 | |
288 -- ==∅→≡ : {n : Level} → { x : OD {suc n} } → ( x == od∅ {suc n} ) → x ≡ od∅ {suc n} | |
289 -- ==∅→≡ {n} {x} eq = cong ( λ k → record { def = k } ) (trans {!!} ∅-base-def ) where | |
290 | |
60 | 291 |
59
d13d1351a1fa
lemma = cong₂ (λ x not → minimul x not ) oiso { }6
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
58
diff
changeset
|
292 open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) |
d13d1351a1fa
lemma = cong₂ (λ x not → minimul x not ) oiso { }6
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
58
diff
changeset
|
293 |
54 | 294 OD→ZF : {n : Level} → ZF {suc (suc n)} {suc n} |
40 | 295 OD→ZF {n} = record { |
54 | 296 ZFSet = OD {suc n} |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
297 ; _∋_ = _∋_ |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
298 ; _≈_ = _==_ |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
299 ; ∅ = od∅ |
28 | 300 ; _,_ = _,_ |
301 ; Union = Union | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
302 ; Power = Power |
28 | 303 ; Select = Select |
304 ; Replace = Replace | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
305 ; infinite = record { def = λ x → x ≡ record { lv = Suc Zero ; ord = ℵ Zero } } |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
306 ; isZF = isZF |
28 | 307 } where |
54 | 308 Replace : OD {suc n} → (OD {suc n} → OD {suc n} ) → OD {suc n} |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
309 Replace X ψ = sup-od ψ |
54 | 310 Select : OD {suc n} → (OD {suc n} → Set (suc n) ) → OD {suc n} |
311 Select X ψ = record { def = λ x → ( def X x ∧ ψ ( ord→od x )) } | |
312 _,_ : OD {suc n} → OD {suc n} → OD {suc n} | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
313 x , y = record { def = λ z → ( (z ≡ od→ord x ) ∨ ( z ≡ od→ord y )) } |
54 | 314 Union : OD {suc n} → OD {suc n} |
71
d088eb66564e
add osuc ( next larger element of Ordinal )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
70
diff
changeset
|
315 Union U = record { def = λ y → osuc y o< (od→ord U) } |
54 | 316 Power : OD {suc n} → OD {suc n} |
317 Power x = record { def = λ y → (z : Ordinal {suc n} ) → ( def x y ∧ def (ord→od z) y ) } | |
318 ZFSet = OD {suc n} | |
319 _∈_ : ( 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
|
320 A ∈ B = B ∋ A |
54 | 321 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set (suc n) |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
322 _⊆_ A B {x} = A ∋ x → B ∋ x |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
323 _∩_ : ( A B : ZFSet ) → ZFSet |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
324 A ∩ B = Select (A , B) ( λ x → ( A ∋ x ) ∧ (B ∋ x) ) |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
325 _∪_ : ( A B : ZFSet ) → ZFSet |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
326 A ∪ B = Select (A , B) ( λ x → (A ∋ x) ∨ ( B ∋ x ) ) |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
327 infixr 200 _∈_ |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
328 infixr 230 _∩_ _∪_ |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
329 infixr 220 _⊆_ |
54 | 330 isZF : IsZF (OD {suc n}) _∋_ _==_ od∅ _,_ Union Power Select Replace (record { def = λ x → x ≡ record { lv = Suc Zero ; ord = ℵ Zero } }) |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
331 isZF = record { |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
332 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
|
333 ; pair = pair |
72 | 334 ; union-u = union-u |
335 ; union→ = union→ | |
336 ; union← = union← | |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
337 ; empty = empty |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
338 ; power→ = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
339 ; power← = {!!} |
65
164ad5a703d8
¬∅=→∅∈ : {n : Level} → { x : OD {suc n} } → ¬ ( x == od∅ {suc n} ) → x ∋ od∅ {suc n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
64
diff
changeset
|
340 ; extensionality = {!!} |
30
3b0fdb95618e
problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
341 ; minimul = minimul |
51 | 342 ; regularity = regularity |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
343 ; infinity∅ = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
344 ; infinity = {!!} |
55
9c0a5e28a572
regurality elimination case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
54
diff
changeset
|
345 ; selection = λ {ψ} {X} {y} → selection {ψ} {X} {y} |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
346 ; replacement = {!!} |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
347 } where |
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
28
diff
changeset
|
348 open _∧_ |
41 | 349 open Minimumo |
54 | 350 pair : (A B : OD {suc n} ) → ((A , B) ∋ A) ∧ ((A , B) ∋ B) |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
351 proj1 (pair A B ) = case1 refl |
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
352 proj2 (pair A B ) = case2 refl |
54 | 353 empty : (x : OD {suc n} ) → ¬ (od∅ ∋ x) |
43
0d9b9db14361
equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
42
diff
changeset
|
354 empty x () |
70 | 355 union-u : (X z : OD {suc n}) → Union X ∋ z → OD {suc n} |
72 | 356 union-u X z U>z = ord→od ( osuc ( od→ord z ) ) |
357 union-lemma-u : {X z : OD {suc n}} → (U>z : Union X ∋ z ) → union-u X z U>z ∋ z | |
358 union-lemma-u {X} {z} U>z = lemma <-osuc where | |
359 lemma : {oz ooz : Ordinal {suc n}} → oz o< ooz → def (ord→od ooz) oz | |
360 lemma {oz} {ooz} lt = def-subst {suc n} {ord→od ooz} (o<→c< lt) refl diso | |
361 union→ : (X z u : OD) → (Union X ∋ u) ∧ (u ∋ z) → Union X ∋ z | |
362 union→ X y u xx with trio< ( od→ord u ) ( osuc ( od→ord y )) | |
363 union→ X y u xx | tri< a ¬b ¬c = {!!} | |
364 union→ X y u xx | tri≈ ¬a b ¬c = {!!} | |
365 union→ X y u xx | tri> ¬a ¬b c = {!!} | |
366 -- c<→o< (transitive {n} {X} {x} {y} X∋x x∋y ) | |
367 union← : (X z : OD) (X∋z : Union X ∋ z) → (X ∋ union-u X z X∋z) ∧ (union-u X z X∋z ∋ z ) | |
368 union← X z X∋z = record { proj1 = {!!} ; proj2 = union-lemma-u X∋z } | |
54 | 369 ψiso : {ψ : OD {suc n} → Set (suc n)} {x y : OD {suc n}} → ψ x → x ≡ y → ψ y |
370 ψiso {ψ} t refl = t | |
371 selection : {ψ : OD → Set (suc n)} {X y : OD} → ((X ∋ y) ∧ ψ y) ⇔ (Select X ψ ∋ y) | |
372 selection {ψ} {X} {y} = record { | |
373 proj1 = λ cond → record { proj1 = proj1 cond ; proj2 = ψiso {ψ} (proj2 cond) (sym oiso) } | |
374 ; proj2 = λ select → record { proj1 = proj1 select ; proj2 = ψiso {ψ} (proj2 select) oiso } | |
375 } | |
60 | 376 ∅-iso : {x : OD} → ¬ (x == od∅) → ¬ ((ord→od (od→ord x)) == od∅) |
377 ∅-iso {x} neq = subst (λ k → ¬ k) (=-iso {n} ) neq | |
54 | 378 minimul : (x : OD {suc n} ) → ¬ (x == od∅ )→ OD {suc n} |
68 | 379 minimul x not = od∅ |
57 | 380 regularity : (x : OD) (not : ¬ (x == od∅)) → |
381 (x ∋ minimul x not) ∧ (Select (minimul x not) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅) | |
66 | 382 proj1 (regularity x not ) = ¬∅=→∅∈ not |
383 proj2 (regularity x not ) = record { eq→ = reg ; eq← = λ () } where | |
384 reg : {y : Ordinal} → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) y → def od∅ y | |
385 reg {y} t with proj1 t | |
386 ... | x∈∅ = x∈∅ |