annotate ordinal-definable.agda @ 46:e584686a1307

== and ∅7
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 25 May 2019 09:09:40 +0900
parents 33860eb44e47
children 264784731a67
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 -- _=='_ : {n : Level} → Set (suc n) -- Rel (OD {n}) (suc n)
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
40 -- _=='_ {n} = ( a b : OD {n} ) → ( ∀ { x : OD {n} } → a ∋ x → b ∋ x ) ∧ ( ∀ { x : OD {n} } → a ∋ x → b ∋ x )
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
41
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
42 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
43 field
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
44 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
45 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
46
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
47 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
48 id x = x
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 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
51 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
52
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
53 open _==_
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-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
56 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
57
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
58 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
59 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
60
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
61 _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
62 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
63
40
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
64 od∅ : {n : Level} → OD {n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
65 od∅ {n} = record { def = λ _ → Lift n ⊥ }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
66
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
67 open import Relation.Binary.HeterogeneousEquality using (_≅_;refl)
45
33860eb44e47 od∅' {n} = ord→od (o∅ {n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 44
diff changeset
68
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
69 postulate
36
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
70 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
71 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
72 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
73 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
74 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
75 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
76 ∅-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
77 ∅-base-lv : {n : Level} → lv ( od→ord (od∅ {n}) ) ≡ lv (o∅ {n})
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
78 ∅-base-d : {n : Level} → ord ( od→ord (od∅ {n}) ) ≅ ord (o∅ {n})
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
79
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
80 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
81 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
82
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
83 od∅→o∅ : {n : Level} → od→ord (od∅ {n} ) ≡ o∅ {n}
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
84 od∅→o∅ = ordinal-cong ∅-base-lv ∅-base-d
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
85
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
86 ∅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
87 ∅1 {n} x (lift ())
28
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
88
37
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
89 ∅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
90 ∅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
91 c0 : Nat → Ordinal {n} → Set n
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
92 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
93 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
94 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
95 ... | t with t (case1 ≤-refl )
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
96 c1 lx not | t | ()
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
97 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
98 c2 Zero not = refl
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
99 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
100 ... | t with t (case1 ≤-refl )
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
101 c2 (Suc lx) not | t | ()
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
102 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
103 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
104 ... | t with t (case2 Φ< )
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
105 c3 lx (Φ .lx) d not | t | ()
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
106 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
107 ... | 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
108 c3 lx (OSuc .lx x₁) d not | t | ()
34
c9ad0d97ce41 fix oridinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 33
diff changeset
109 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
110 ... | t with t (case2 (s< ℵΦ< ))
34
c9ad0d97ce41 fix oridinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 33
diff changeset
111 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
112
37
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
113 -- find : {n : Level} → ( x : Ordinal {n} ) → o∅ o< x → Ordinal {n}
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
114 -- exists : {n : Level} → ( x : Ordinal {n} ) → (0<x : o∅ o< x ) → find x 0<x o< x
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
115
36
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
116 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
117 def-subst df refl refl = df
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
118
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
119 transitive : {n : Level } { x y z : OD {n} } → y ∋ x → z ∋ y → z ∋ x
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
120 transitive {n} {x} {y} {z} x∋y z∋y with ordtrans ( c<→o< {n} {x} {y} x∋y ) ( c<→o< {n} {y} {z} z∋y )
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
121 ... | t = lemma0 (lemma t) where
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
122 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
123 lemma xo<z = o<→c< xo<z
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
124 lemma0 : def ( ord→od ( od→ord z )) ( od→ord ( ord→od ( od→ord x ))) → def z (od→ord x)
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
125 lemma0 dz = def-subst {n} { ord→od ( od→ord z )} { od→ord ( ord→od ( od→ord x))} dz (oiso) (diso)
4d64509067d0 transitive
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 34
diff changeset
126
41
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
127 record Minimumo {n : Level } (x : Ordinal {n}) : Set (suc n) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
128 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
129 mino : Ordinal {n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
130 min<x : mino o< x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
131
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
132 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
133 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
134 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
135 ominimal {n} record { lv = Zero ; ord = (OSuc .0 ord) } (case1 ())
41
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
136 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
137 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
138 ominimal {n} record { lv = (Suc lv) ; ord = (Φ .(Suc lv)) } (case2 ())
41
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
139 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
140 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
141 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
142 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
143
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
144 ∅4 : {n : Level} → ( x : OD {n} ) → x ≡ od∅ {n} → od→ord x ≡ o∅ {n}
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
145 ∅4 {n} x refl = ∅3 lemma1 where
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
146 lemma0 : (y : Ordinal {n}) → def ( od∅ {n} ) y → ⊥
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
147 lemma0 y (lift ())
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
148 lemma1 : (y : Ordinal {n}) → y o< od→ord od∅ → ⊥
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
149 lemma1 y y<o = lemma0 y ( def-subst {n} {ord→od (od→ord od∅ )} {od→ord (ord→od y)} (o<→c< y<o) oiso diso )
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
150
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
151 ∅5 : {n : Level} → ( x : Ordinal {n} ) → ¬ ( x ≡ o∅ {n} ) → o∅ {n} o< x
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
152 ∅5 {n} record { lv = Zero ; ord = (Φ .0) } not = ⊥-elim (not refl)
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
153 ∅5 {n} record { lv = Zero ; ord = (OSuc .0 ord) } not = case2 Φ<
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
154 ∅5 {n} record { lv = (Suc lv) ; ord = ord } not = case1 (s≤s z≤n)
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
155
39
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
156 postulate extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality n (suc n)
37
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
157
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
158 ∅6 : {n : Level } ( x : Ordinal {suc n}) → o∅ o< x → ¬ x ≡ o∅
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
159 ∅6 {n} x lt eq with trio< {n} (o∅ {suc n}) x
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
160 ∅6 {n} x lt refl | tri< a ¬b ¬c = ¬b refl
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
161 ∅6 {n} x lt refl | tri≈ ¬a b ¬c = ¬a lt
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
162 ∅6 {n} x lt refl | tri> ¬a ¬b c = ¬b refl
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
163
39
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
164 ∅8 : {n : Level} → ( x : Ordinal {n} ) → ¬ x o< o∅ {n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
165 ∅8 {n} x (case1 ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
166 ∅8 {n} x (case2 ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
167
40
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
168 -- ∅10 : {n : Level} → (x : OD {n} ) → ¬ ( ( y : OD {n} ) → Lift (suc n) ( x ∋ y)) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
169 -- ∅10 {n} x not = ?
39
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
170
40
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
171 -- ∅77 : {n : Level} → (x : OD {suc n} ) → ¬ ( ord→od (o∅ {suc n}) ∋ x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
172 -- ∅77 {n} x lt = {!!} where
39
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
173
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
174 ∅7' : {n : Level} → ord→od (o∅ {n}) == od∅ {n}
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
175 ∅7' {n} = record { eq→ = e1 ; eq← = e2 } where
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
176 e2 : {y : Ordinal {n}} → def od∅ y → def (ord→od (o∅ {n})) y
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
177 e2 {y} (lift ())
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
178 e1 : {x : Ordinal {n} } → def (ord→od (o∅ {n})) x → def (od∅ {n}) x
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
179 e1 {x} o<x = def-subst {n} {ord→od (o∅ {n})} {x} o<x (cong (λ k → record { def = k }) ∅-base-def ) refl
39
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
180
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
181 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
182 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
183
fcac01485f32 od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 43
diff changeset
184
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
185 ∅7 : {n : Level} → ( x : OD {n} ) → od→ord x ≡ o∅ {n} → x == od∅ {n}
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
186 ∅7 {n} x eq = record { eq→ = e1 eq ; eq← = e2 } where
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
187 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
188 e2 {y} (lift ())
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
189 e1 : {ox y : Ordinal {n}} → ox ≡ o∅ {n} → def x y → def od∅ y
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
190 e1 {ox} {y} eq x>y with lv ox
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
191 e1 {o∅} {y} refl x>y | Zero = 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 ))
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
192 e1 {o∅} {y} refl x>y | Suc lx = 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 ))
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
193
37
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
194
45
33860eb44e47 od∅' {n} = ord→od (o∅ {n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 44
diff changeset
195 open _∧_
33860eb44e47 od∅' {n} = ord→od (o∅ {n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 44
diff changeset
196
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
197 ∅9 : {n : Level} → (x : OD {n} ) → ¬ x == od∅ → o∅ o< od→ord x
38
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
198 ∅9 x not = ∅5 ( od→ord x) lemma where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
199 lemma : ¬ od→ord x ≡ o∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
200 lemma eq = not ( ∅7 x eq )
37
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
201
44
fcac01485f32 od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 43
diff changeset
202
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
203 OD→ZF : {n : Level} → ZF {suc n} {n}
40
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
204 OD→ZF {n} = record {
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
205 ZFSet = OD {n}
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
206 ; _∋_ = _∋_
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
207 ; _≈_ = _==_
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
208 ; ∅ = od∅
28
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
209 ; _,_ = _,_
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
210 ; Union = Union
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
211 ; Power = Power
28
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
212 ; Select = Select
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
213 ; Replace = Replace
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
214 ; 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
215 ; isZF = isZF
28
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
216 } where
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
217 Replace : OD {n} → (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
218 Replace X ψ = sup-od ψ
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
219 Select : OD {n} → (OD {n} → Set n ) → OD {n}
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
220 Select X ψ = record { def = λ x → select ( ord→od x ) } where
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
221 select : OD {n} → Set n
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
222 select x = ψ x
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
223 _,_ : 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
224 x , y = record { def = λ z → ( (z ≡ od→ord x ) ∨ ( z ≡ od→ord y )) }
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
225 Union : OD {n} → OD {n}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
226 Union x = record { def = λ y → {z : Ordinal {n}} → def x z → def (ord→od z) y }
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
227 Power : OD {n} → OD {n}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
228 Power x = record { def = λ y → (z : Ordinal {n} ) → ( def x y ∧ def (ord→od z) y ) }
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
229 ZFSet = OD {n}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
230 _∈_ : ( A B : ZFSet ) → Set n
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
231 A ∈ B = B ∋ A
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
232 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set n
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
233 _⊆_ 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
234 _∩_ : ( A B : ZFSet ) → ZFSet
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
235 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
236 _∪_ : ( A B : ZFSet ) → ZFSet
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
237 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
238 infixr 200 _∈_
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
239 infixr 230 _∩_ _∪_
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
240 infixr 220 _⊆_
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
241 isZF : IsZF (OD {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
242 isZF = record {
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
243 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
244 ; pair = pair
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
245 ; union→ = {!!}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
246 ; union← = {!!}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
247 ; empty = empty
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
248 ; power→ = {!!}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
249 ; power← = {!!}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
250 ; extentionality = {!!}
30
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
251 ; minimul = minimul
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
252 ; regularity = {!!}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
253 ; infinity∅ = {!!}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
254 ; infinity = {!!}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
255 ; selection = {!!}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
256 ; replacement = {!!}
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
257 } where
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
258 open _∧_
41
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
259 open Minimumo
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
260 pair : (A B : OD {n} ) → ((A , B) ∋ A) ∧ ((A , B) ∋ B)
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
261 proj1 (pair A B ) = case1 refl
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
262 proj2 (pair A B ) = case2 refl
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
263 empty : (x : OD {n} ) → ¬ (od∅ ∋ x)
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
264 empty x ()
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
265 union→ : (X x y : OD {n} ) → (X ∋ x) → (x ∋ y) → (Union X ∋ y)
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
266 union→ X x y X∋x x∋y = {!!} where
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
267 lemma : {z : Ordinal {n} } → def X z → z ≡ od→ord y
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
268 lemma {z} X∋z = {!!}
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
269 minord : (x : OD {n} ) → ¬ (x == od∅ )→ Minimumo (od→ord x)
41
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
270 minord x not = ominimal (od→ord x) (∅9 x not)
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
271 minimul : (x : OD {n} ) → ¬ (x == od∅ )→ OD {n}
41
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 40
diff changeset
272 minimul x not = ord→od ( mino (minord x not))
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
273 minimul<x : (x : OD {n} ) → (not : ¬ x == od∅ ) → x ∋ minimul x not
42
4d5fc6381546 regurality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 41
diff changeset
274 minimul<x x not = lemma0 (min<x (minord x not)) where
4d5fc6381546 regurality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 41
diff changeset
275 lemma0 : mino (minord x not) o< (od→ord x) → def x (od→ord (ord→od (mino (minord x not))))
4d5fc6381546 regurality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 41
diff changeset
276 lemma0 m<x = def-subst {n} {ord→od (od→ord x)} {od→ord (ord→od (mino (minord x not)))} (o<→c< m<x) oiso refl
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
277 regularity : (x : OD) (not : ¬ (x == od∅)) → (x ∋ minimul x not) ∧
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
278 (Select (minimul x not , x) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅)
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
279 -- regularity : (x : OD) → (not : ¬ x == od∅ ) →
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
280 -- ((x ∋ minimul x not ) ∧ {!!} ) -- (Select x (λ x₁ → (( minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅)))
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
281 proj1 ( regularity x non ) = minimul<x x non
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
282 proj2 ( regularity x non ) = {!!} where -- cong ( λ k → record { def = k } ) ( extensionality ( λ y → lemma0 y) ) where
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
283 not-min : ( z : Ordinal {n} ) → ¬ ( def (Select x (λ y → (minimul x non ∋ y) ∧ (x ∋ y))) z )
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
284 not-min z = {!!}
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
285 lemma0 : ( z : Ordinal {n} ) → def (Select x (λ y → (minimul x non ∋ y) ∧ (x ∋ y))) z ≡ Lift n ⊥
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
286 lemma0 z = {!!}
42
4d5fc6381546 regurality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 41
diff changeset
287