annotate ordinal-definable.agda @ 70:cd9cf8b09610

Union needs +1 space
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 01 Jun 2019 10:01:38 +0900
parents 93abc0133b8a
children d088eb66564e
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
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
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
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
6
23
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 22
diff changeset
7 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ )
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
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
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
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
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
12 open import Data.Empty
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
13 open import Relation.Nullary
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
14
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
15 open import Relation.Binary
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
16 open import Relation.Binary.Core
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
17
27
bade0a35fdd9 OD, HOD, TC
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 26
diff changeset
18 -- Ordinal Definable Set
11
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 10
diff changeset
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
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
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
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
61 od∅ : {n : Level} → OD {n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
62 od∅ {n} = record { def = λ _ → Lift n ⊥ }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
63
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
64 postulate
36
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
65 c<→o< : {n : Level} {x y : OD {n} } → x c< y → od→ord x o< od→ord y
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
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
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
71 ∅-base-def : {n : Level} → def ( ord→od (o∅ {n}) ) ≡ def (od∅ {n})
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
72
60
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
73 congf : {n : Level} {x y : OD {n}} → { f g : OD {n} → OD {n} } → x ≡ y → f ≡ g → f x ≡ g y
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
74 congf refl refl = refl
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
75
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
76 o∅→od∅ : {n : Level} → ord→od (o∅ {n}) ≡ od∅ {n}
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
77 o∅→od∅ {n} = cong ( λ k → record { def = k }) ( ∅-base-def )
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
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
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
81
37
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
82 ∅3 : {n : Level} → { x : Ordinal {n}} → ( ∀(y : Ordinal {n}) → ¬ (y o< x ) ) → x ≡ o∅ {n}
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
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
c9ad0d97ce41 fix oridinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 33
diff changeset
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
c9ad0d97ce41 fix oridinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 33
diff changeset
102 c3 (Suc lx) (ℵ lx) d not with not ( record { lv = Suc lx ; ord = OSuc (Suc lx) (Φ (Suc lx)) } )
41
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
103 ... | t with t (case2 (s< ℵΦ< ))
34
c9ad0d97ce41 fix oridinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 33
diff changeset
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
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
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
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
107 def-subst df refl refl = df
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
108
69
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
109 transitive : {n : Level } { z y x : OD {suc n} } → z ∋ y → y ∋ x → z ∋ x
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
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
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
111 ... | t = lemma0 (lemma t) where
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
112 lemma : ( od→ord x ) o< ( od→ord z ) → def ( ord→od ( od→ord z )) ( od→ord ( ord→od ( od→ord x )))
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
113 lemma xo<z = o<→c< xo<z
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
114 lemma0 : def ( ord→od ( od→ord z )) ( od→ord ( ord→od ( od→ord x ))) → def z (od→ord x)
69
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
115 lemma0 dz = def-subst {suc n} { ord→od ( od→ord z )} { od→ord ( ord→od ( od→ord x))} dz (oiso) (diso)
36
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
116
41
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
117 record Minimumo {n : Level } (x : Ordinal {n}) : Set (suc n) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
118 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
119 mino : Ordinal {n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
120 min<x : mino o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
121
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
122 ominimal : {n : Level} → (x : Ordinal {n} ) → o∅ o< x → Minimumo {n} x
37
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
123 ominimal {n} record { lv = Zero ; ord = (Φ .0) } (case1 ())
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
124 ominimal {n} record { lv = Zero ; ord = (Φ .0) } (case2 ())
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
125 ominimal {n} record { lv = Zero ; ord = (OSuc .0 ord) } (case1 ())
41
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
126 ominimal {n} record { lv = Zero ; ord = (OSuc .0 ord) } (case2 Φ<) = record { mino = record { lv = Zero ; ord = Φ 0 } ; min<x = case2 Φ< }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
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
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
128 ominimal {n} record { lv = (Suc lv) ; ord = (Φ .(Suc lv)) } (case2 ())
41
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
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
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
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
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
133
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
134 ∅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
135 ∅5 {n} {record { lv = Zero ; ord = (Φ .0) }} not = ⊥-elim (not refl)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
136 ∅5 {n} {record { lv = Zero ; ord = (OSuc .0 ord) }} not = case2 Φ<
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
137 ∅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
138
39
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
139 ∅8 : {n : Level} → ( x : Ordinal {n} ) → ¬ x o< o∅ {n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
140 ∅8 {n} x (case1 ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
141 ∅8 {n} x (case2 ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
142
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
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 }
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
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
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
146 -- avoiding lv != Zero error
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
147 orefl : {n : Level} → { x : OD {n} } → { y : Ordinal {n} } → od→ord x ≡ y → od→ord x ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
148 orefl refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
149
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
150 ==-iso : {n : Level} → { x y : OD {n} } → ord→od (od→ord x) == ord→od (od→ord y) → x == y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
151 ==-iso {n} {x} {y} eq = record {
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
152 eq→ = λ d → lemma ( eq→ eq (def-subst d (sym oiso) refl )) ;
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
153 eq← = λ d → lemma ( eq← eq (def-subst d (sym oiso) refl )) }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
154 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
155 lemma : {x : OD {n} } {z : Ordinal {n}} → def (ord→od (od→ord x)) z → def x z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
156 lemma {x} {z} d = def-subst d oiso refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
157
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
158 =-iso : {n : Level } {x y : OD {suc n} } → (x == y) ≡ (ord→od (od→ord x) == y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
159 =-iso {_} {_} {y} = cong ( λ k → k == y ) (sym oiso)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
160
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
161 ord→== : {n : Level} → { x y : OD {n} } → od→ord x ≡ od→ord y → x == y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
162 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
163 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
164 lemma ox ox refl = eq-refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
165
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
166 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
167 o≡→== {n} {x} {.x} refl = eq-refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
168
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
169 ∅7 : {n : Level} → { x : OD {n} } → od→ord x ≡ o∅ {n} → x == od∅ {n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
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
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
173 e1 : {ox y : Ordinal {n}} → ox ≡ o∅ {n} → def x y → def od∅ y
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
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 ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
175
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
176 =→¬< : {x : Nat } → ¬ ( x < x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
177 =→¬< {Zero} ()
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
178 =→¬< {Suc x} (s≤s lt) = =→¬< lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
179
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
180 >→¬< : {x y : Nat } → (x < y ) → ¬ ( y < x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
181 >→¬< (s≤s x<y) (s≤s y<x) = >→¬< x<y y<x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
182
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
183 c≤-refl : {n : Level} → ( x : OD {n} ) → x c≤ x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
184 c≤-refl x = case1 refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
185
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
186 o<> : {n : Level } ( ox oy : Ordinal {n}) → ox o< oy → oy o< ox → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
187 o<> ox oy (case1 x<y) (case1 y<x) = >→¬< x<y y<x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
188 o<> ox oy (case1 x<y) (case2 y<x) with d<→lv y<x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
189 ... | refl = =→¬< x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
190 o<> ox oy (case2 x<y) (case1 y<x) with d<→lv x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
191 ... | refl = =→¬< y<x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
192 o<> ox oy (case2 x<y) (case2 y<x) with d<→lv x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
193 ... | refl = trio<> x<y y<x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
194
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
195 o<¬≡ : {n : Level } ( ox oy : Ordinal {n}) → ox ≡ oy → ox o< oy → ⊥
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
196 o<¬≡ ox ox refl (case1 lt) = =→¬< lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
197 o<¬≡ ox ox refl (case2 (s< lt)) = trio<≡ refl lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
198
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
199 o<→o> : {n : Level} → { x y : OD {n} } → (x == y) → (od→ord x ) o< ( od→ord y) → ⊥
52
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 51
diff changeset
200 o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case1 lt) with
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 51
diff changeset
201 yx (def-subst {n} {ord→od (od→ord y)} {od→ord (ord→od (od→ord x))} (o<→c< (case1 lt )) oiso diso )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 51
diff changeset
202 ... | oyx with o<¬≡ (od→ord x) (od→ord x) refl (c<→o< oyx )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 51
diff changeset
203 ... | ()
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 51
diff changeset
204 o<→o> {n} {x} {y} record { eq→ = xy ; eq← = yx } (case2 lt) with
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 51
diff changeset
205 yx (def-subst {n} {ord→od (od→ord y)} {od→ord (ord→od (od→ord x))} (o<→c< (case2 lt )) oiso diso )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 51
diff changeset
206 ... | oyx with o<¬≡ (od→ord x) (od→ord x) refl (c<→o< oyx )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 51
diff changeset
207 ... | ()
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 51
diff changeset
208
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
209 o<→¬== : {n : Level} → { x y : OD {n} } → (od→ord x ) o< ( od→ord y) → ¬ (x == y )
52
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 51
diff changeset
210 o<→¬== {n} {x} {y} lt eq = o<→o> eq lt
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
211
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
212 o<→¬c> : {n : Level} → { x y : OD {n} } → (od→ord x ) o< ( od→ord y) → ¬ (y c< x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
213 o<→¬c> {n} {x} {y} olt clt = o<> (od→ord x) (od→ord y) olt (c<→o< clt ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
214
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
215 o≡→¬c< : {n : Level} → { x y : OD {n} } → (od→ord x ) ≡ ( od→ord y) → ¬ x c< y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
216 o≡→¬c< {n} {x} {y} oeq lt = o<¬≡ (od→ord x) (od→ord y) (orefl oeq ) (c<→o< lt)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
217
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
218 tri-c< : {n : Level} → Trichotomous _==_ (_c<_ {suc n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
219 tri-c< {n} x y with trio< {n} (od→ord x) (od→ord y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
220 tri-c< {n} x y | tri< a ¬b ¬c = tri< (def-subst (o<→c< a) oiso diso) (o<→¬== a) ( o<→¬c> a )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
221 tri-c< {n} x y | tri≈ ¬a b ¬c = tri≈ (o≡→¬c< b) (ord→== b) (o≡→¬c< (sym b))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
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)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
223
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
224 c<> : {n : Level } { x y : OD {suc n}} → x c< y → y c< x → ⊥
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
225 c<> {n} {x} {y} x<y y<x with tri-c< x y
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
226 c<> {n} {x} {y} x<y y<x | tri< a ¬b ¬c = ¬c y<x
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
227 c<> {n} {x} {y} x<y y<x | tri≈ ¬a b ¬c = o<→o> b ( c<→o< x<y )
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
228 c<> {n} {x} {y} x<y y<x | tri> ¬a ¬b c = ¬a x<y
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
229
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
230 ∅2 : {n : Level} → { x : OD {n} } → o∅ {n} o< od→ord x → ¬ ( x == od∅ {n} )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
231 ∅2 {n} {x} lt record { eq→ = eq→ ; eq← = eq← } with ominimal (od→ord x ) lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
232 ... | min with eq→ ( def-subst (o<→c< (Minimumo.min<x min)) oiso refl )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
233 ... | ()
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
234
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
235 ∅0 : {n : Level} → { x : Ordinal {n} } → o∅ {n} o< x → ¬ ( ord→od x == od∅ {n} )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
236 ∅0 {n} {x} lt record { eq→ = eq→ ; eq← = eq← } with ominimal x lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
237 ... | min with eq→ (o<→c< (Minimumo.min<x min))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
238 ... | ()
60
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
239
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
240 ∅< : {n : Level} → { x y : OD {n} } → def x (od→ord y ) → ¬ ( x == od∅ {n} )
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
241 ∅< {n} {x} {y} d eq with eq→ eq d
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
242 ∅< {n} {x} {y} d eq | lift ()
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
243
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
244
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
245 is-od∅ : {n : Level} → ( x : OD {suc n} ) → Dec ( x == od∅ {suc n} )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
246 is-od∅ {n} x with trio< {n} (od→ord x) (o∅ {suc n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
247 is-od∅ {n} x | tri≈ ¬a b ¬c = yes ( ∅7 (orefl b) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
248 is-od∅ {n} x | tri< (case1 ()) ¬b ¬c
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
249 is-od∅ {n} x | tri< (case2 ()) ¬b ¬c
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
250 is-od∅ {n} x | tri> ¬a ¬b c = no ( ∅2 c )
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
251
53
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 52
diff changeset
252 is-∋ : {n : Level} → ( x y : OD {suc n} ) → Dec ( x ∋ y )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 52
diff changeset
253 is-∋ {n} x y with tri-c< x y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 52
diff changeset
254 is-∋ {n} x y | tri< a ¬b ¬c = no ¬c
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 52
diff changeset
255 is-∋ {n} x y | tri≈ ¬a b ¬c = no ¬c
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 52
diff changeset
256 is-∋ {n} x y | tri> ¬a ¬b c = yes c
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 52
diff changeset
257
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
258 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
259 is-o∅ {n} record { lv = Zero ; ord = (Φ .0) } = yes refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
260 is-o∅ {n} record { lv = Zero ; ord = (OSuc .0 ord₁) } = no ( λ () )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
261 is-o∅ {n} record { lv = (Suc lv₁) ; ord = ord } = no (λ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
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
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
265 ∅9 : {n : Level} → {x : OD {n} } → ¬ x == od∅ → o∅ o< od→ord x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
266 ∅9 {_} {x} not = ∅5 lemma where
38
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
267 lemma : ¬ od→ord x ≡ o∅
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
268 lemma eq = not ( ∅7 eq )
37
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
269
66
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
270 ∅10 : {n : Level} → {ox : Ordinal {n}} → (not : ¬ (ord→od ox == od∅)) → o∅ o< ox
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
271 ∅10 {n} {ox} not = subst (λ k → o∅ o< k) diso (∅9 not)
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
272
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
273 ¬∅=→∅∈ : {n : Level} → { x : OD {suc n} } → ¬ ( x == od∅ {suc n} ) → x ∋ od∅ {suc n}
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
274 ¬∅=→∅∈ {n} {x} ne = def-subst (lemma (od→ord x) (subst (λ k → ¬ (k == od∅ {suc n} )) (sym oiso) ne )) oiso refl where
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
275 lemma : (ox : Ordinal {suc n}) → ¬ (ord→od ox == od∅ {suc n} ) → ord→od ox ∋ od∅ {suc n}
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
276 lemma ox ne with is-o∅ ox
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
277 lemma ox ne | yes refl with ne ( ord→== lemma1 ) where
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
278 lemma1 : od→ord (ord→od o∅) ≡ od→ord od∅
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
279 lemma1 = cong ( λ k → od→ord k ) o∅→od∅
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
280 lemma o∅ ne | yes refl | ()
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
281 lemma ox ne | no ¬p = subst ( λ k → def (ord→od ox) (od→ord k) ) o∅→od∅ (o<→c< (∅5 ¬p))
69
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
282
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
283 -- ∃-set : {n : Level} → ( x : OD {suc n} ) → ( ψ : OD {suc n} → Set (suc n) ) → Set (suc n)
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
284 -- ∃-set = {!!}
66
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
285
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
286 -- postulate f-extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality (suc (suc n)) (suc (suc n ))
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
287
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
288 -- ==∅→≡ : {n : Level} → { x : OD {suc n} } → ( x == od∅ {suc n} ) → x ≡ od∅ {suc n}
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
289 -- ==∅→≡ {n} {x} eq = cong ( λ k → record { def = k } ) (trans {!!} ∅-base-def ) where
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
290
60
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
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
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
294 OD→ZF : {n : Level} → ZF {suc (suc n)} {suc n}
40
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
295 OD→ZF {n} = record {
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
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
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
300 ; _,_ = _,_
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
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
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
303 ; Select = Select
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
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
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
307 } where
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
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
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
310 Select : OD {suc n} → (OD {suc n} → Set (suc n) ) → OD {suc n}
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
311 Select X ψ = record { def = λ x → ( def X x ∧ ψ ( ord→od x )) }
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
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
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
314 Union : OD {suc n} → OD {suc n}
69
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
315 Union U = record { def = λ y → y o< (od→ord U) }
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
316 Power : OD {suc n} → OD {suc n}
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
317 Power x = record { def = λ y → (z : Ordinal {suc n} ) → ( def x y ∧ def (ord→od z) y ) }
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
318 ZFSet = OD {suc n}
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
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
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
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
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
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
69
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
334 ; union-u = {!!}
70
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
335 ; union→ = {!!}
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
336 ; 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
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
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
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
349 open Minimumo
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
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
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
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
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
355 union-u : (X z : OD {suc n}) → Union X ∋ z → OD {suc n}
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
356 union-u X z U>z with od→ord X | od→ord z
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
357 union-u X z (case1 (s≤s x)) | record { lv = Suc lX ; ord = oX } | record { lv = lz ; ord = oz } = ord→od (mid (Suc lX) lz oX oz (s≤s x) ) where
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
358 mid : (lX lz : Nat) → (oX : OrdinalD {suc n} lX ) → (oz : OrdinalD {suc n} lz ) → lX > lz → Ordinal {suc n}
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
359 mid (Suc lX) 0 oX oz (s≤s z≤n) = record { lv = 0 ; ord = OSuc 0 oz }
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
360 mid (Suc (Suc lX)) (Suc lz) (Φ (Suc (Suc lX))) (Φ (Suc lz)) (s≤s (s≤s lt)) = record { lv = Suc lz ; ord = OSuc (Suc lz) (Φ (Suc lz) ) }
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
361 mid (Suc (Suc lX)) (Suc lz) (Φ (Suc (Suc lX))) (OSuc (Suc lz) oz) (s≤s (s≤s lt)) = record { lv = Suc lz ; ord = OSuc (Suc lz) (OSuc (Suc lz) oz) }
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
362 mid (Suc (Suc lX)) (Suc lz) (Φ (Suc (Suc lX))) (ℵ lz) (s≤s (s≤s lt)) = record { lv = (Suc (Suc lX)) ; ord = OSuc (Suc (Suc lX)) (ℵ (Suc lX)) }
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
363 mid (Suc (Suc lX)) (Suc lz) (OSuc (Suc (Suc lX)) oX) oz (s≤s (s≤s lt)) = record { lv = Suc (Suc lX) ; ord = oX }
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
364 mid (Suc (Suc lX)) (Suc lz) (ℵ (Suc lX)) oz (s≤s (s≤s lt)) = record { lv = Suc (Suc lX) ; ord = OSuc (Suc (Suc lX)) (Φ (Suc (Suc lX))) }
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
365 union-u X z (case2 Φ<) | record { lv = Zero ; ord = (OSuc 0 oX) } | record { lv = .0 ; ord = .(Φ 0) } =
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
366 ord→od (record { lv = Zero ; ord = oX }) -- this case contains X = odΦ
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
367 union-u X z (case2 Φ<) | record { lv = Suc lv₁ ; ord = (OSuc (Suc lv₁) oX) } | record { lv = .(Suc lv₁) ; ord = .(Φ (Suc lv₁)) } = ord→od (record { lv = Suc lv₁ ; ord = oX } )
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
368 union-u X z (case2 (s< x)) | record { lv = lv₁ ; ord = .(OSuc lv₁ _) } | record { lv = .lv₁ ; ord = .(OSuc lv₁ _) } = {!!}
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
369 union-u X z (case2 ℵΦ<) | record { lv = Suc lX ; ord = .(ℵ _) } | record { lv = .(Suc _) ; ord = .(Φ (Suc _)) } = ord→od ( record { lv = Suc lX ; ord = OSuc (Suc lX ) (Φ (Suc lX))} )
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
370 union-u X z (case2 (ℵ< x)) | record { lv = Suc lX ; ord = .(ℵ _) } | record { lv = .(Suc _) ; ord = OSuc (Suc _) oz } =
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
371 ord→od ( record { lv = Suc lX ; ord = OSuc (Suc lX ) (OSuc (Suc lX) (OSuc (Suc lX) oz ))} )
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
372 union-u X z (case2 ℵs<) | record { lv = Suc lX ; ord = OSuc (Suc lX) (ℵ lX) } | record { lv = Suc lX ; ord = ℵ lX } = {!!}
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
373 union-u X z (case2 (ℵss< x)) | record { lv = .(Suc _) ; ord = .(OSuc (Suc _) _) } | record { lv = .(Suc _) ; ord = .(ℵ _) } = {!!}
69
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
374 union→ : (X x y : OD {suc n}) → X ∋ x → x ∋ y → Union X ∋ y
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
375 union→ X x y X∋x x∋y = c<→o< (transitive {n} {X} {x} {y} X∋x x∋y )
70
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
376 union← : (X z : OD {suc n}) → (X∋z : Union X ∋ z) → ((Union X ∋ union-u X z X∋z) ∧ (union-u X z X∋z ∋ z))
cd9cf8b09610 Union needs +1 space
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 69
diff changeset
377 union← = {!!}
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
378 ψiso : {ψ : OD {suc n} → Set (suc n)} {x y : OD {suc n}} → ψ x → x ≡ y → ψ y
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
379 ψiso {ψ} t refl = t
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
380 selection : {ψ : OD → Set (suc n)} {X y : OD} → ((X ∋ y) ∧ ψ y) ⇔ (Select X ψ ∋ y)
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
381 selection {ψ} {X} {y} = record {
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
382 proj1 = λ cond → record { proj1 = proj1 cond ; proj2 = ψiso {ψ} (proj2 cond) (sym oiso) }
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
383 ; proj2 = λ select → record { proj1 = proj1 select ; proj2 = ψiso {ψ} (proj2 select) oiso }
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
384 }
60
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
385 ∅-iso : {x : OD} → ¬ (x == od∅) → ¬ ((ord→od (od→ord x)) == od∅)
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
386 ∅-iso {x} neq = subst (λ k → ¬ k) (=-iso {n} ) neq
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
387 minimul : (x : OD {suc n} ) → ¬ (x == od∅ )→ OD {suc n}
68
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 66
diff changeset
388 minimul x not = od∅
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
389 regularity : (x : OD) (not : ¬ (x == od∅)) →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
390 (x ∋ minimul x not) ∧ (Select (minimul x not) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅)
66
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
391 proj1 (regularity x not ) = ¬∅=→∅∈ not
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
392 proj2 (regularity x not ) = record { eq→ = reg ; eq← = λ () } where
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
393 reg : {y : Ordinal} → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) y → def od∅ y
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
394 reg {y} t with proj1 t
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
395 ... | x∈∅ = x∈∅