annotate ordinal-definable.agda @ 104:d92411bed18c

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sun, 16 Jun 2019 02:06:09 +0900
parents c8b79d303867
children ec6235ce0215
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⊔_ )
14
e11e95d5ddee separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 11
diff changeset
8 open import Relation.Binary.PropositionalEquality
e11e95d5ddee separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 11
diff changeset
9 open import Data.Nat.Properties
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
10 open import Data.Empty
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
11 open import Relation.Nullary
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
12 open import Relation.Binary
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
13 open import Relation.Binary.Core
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
14
27
bade0a35fdd9 OD, HOD, TC
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 26
diff changeset
15 -- Ordinal Definable Set
11
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 10
diff changeset
16
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
17 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
18 field
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
19 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
20
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
21 open OD
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
22 open import Data.Unit
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
23
44
fcac01485f32 od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 43
diff changeset
24 open Ordinal
fcac01485f32 od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 43
diff changeset
25
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
26 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
27 field
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
28 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
29 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
30
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
31 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
32 id x = x
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
33
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
34 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
35 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
36
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
37 open _==_
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
38
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
39 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
40 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
41
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
42 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
43 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
44
40
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
45 od∅ : {n : Level} → OD {n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
46 od∅ {n} = record { def = λ _ → Lift n ⊥ }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
47
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
48 postulate
100
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
49 -- OD can be iso to a subset of Ordinal ( by means of Godel Set )
95
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
50 od→ord : {n : Level} → OD {n} → Ordinal {n}
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
51 ord→od : {n : Level} → Ordinal {n} → OD {n}
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
52 oiso : {n : Level} {x : OD {n}} → ord→od ( od→ord x ) ≡ x
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
53 diso : {n : Level} {x : Ordinal {n}} → od→ord ( ord→od x ) ≡ x
100
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
54 -- supermum as Replacement Axiom
95
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
55 sup-o : {n : Level } → ( Ordinal {n} → Ordinal {n}) → Ordinal {n}
98
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 97
diff changeset
56 sup-o< : {n : Level } → { ψ : Ordinal {n} → Ordinal {n}} → ∀ {x : Ordinal {n}} → ψ x o< sup-o ψ
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
57 -- a contra-position of minimality of supermum
98
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 97
diff changeset
58 sup-x : {n : Level } → ( Ordinal {n} → Ordinal {n}) → Ordinal {n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 97
diff changeset
59 sup-lb : {n : Level } → { ψ : Ordinal {n} → Ordinal {n}} → {z : Ordinal {n}} → z o< sup-o ψ → z o< osuc (ψ (sup-x ψ))
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
60 -- sup-min : {n : Level } → ( ψ : Ordinal {n} → Ordinal {n}) → {z : Ordinal {n}} → ψ z o< z → sup-o ψ o< osuc z
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
61 minimul : {n : Level } → (x : OD {suc n} ) → ¬ (x == od∅ )→ OD {suc n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
62 x∋minimul : {n : Level } → (x : OD {suc n} ) → ( ne : ¬ (x == od∅ ) ) → def x ( od→ord ( minimul x ne ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
63 minimul-1 : {n : Level } → (x : OD {suc n} ) → ( ne : ¬ (x == od∅ ) ) → (y : OD {suc n}) → ¬ ( def (minimul x ne) (od→ord y)) ∧ (def x (od→ord y) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
64
95
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
65
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
66 _∋_ : { n : Level } → ( a x : OD {n} ) → Set n
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
67 _∋_ {n} a x = def a ( od→ord x )
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
68
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
69 Ord : { n : Level } → ( a : Ordinal {suc n} ) → OD {suc n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
70 Ord {n} a = record { def = λ y → y o< a }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
71
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
72 _c<_ : { n : Level } → ( x a : Ordinal {n} ) → Set n
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
73 x c< a = Ord a ∋ Ord x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
74
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
75 c<→o< : { n : Level } → { x a : OD {n} } → record { def = λ y → y o< od→ord a } ∋ x → od→ord x o< od→ord a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
76 c<→o< lt = lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
77
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
78 o<→c< : { n : Level } → { x a : OD {n} } → od→ord x o< od→ord a → record { def = λ y → y o< od→ord a } ∋ x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
79 o<→c< lt = lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
80
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
81 ==→o≡' : {n : Level} → { x y : Ordinal {suc n} } → Ord x == Ord y → x ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
82 ==→o≡' {n} {x} {y} eq with trio< {n} x y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
83 ==→o≡' {n} {x} {y} eq | tri< a ¬b ¬c with eq← eq {x} a
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
84 ... | t = ⊥-elim ( o<¬≡ x x refl t )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
85 ==→o≡' {n} {x} {y} eq | tri≈ ¬a refl ¬c = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
86 ==→o≡' {n} {x} {y} eq | tri> ¬a ¬b c with eq→ eq {y} c
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
87 ... | t = ⊥-elim ( o<¬≡ y y refl t )
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
88
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
89 ∅∨ : { n : Level } → { x y : Ordinal {suc n} } → ( Ord {n} x == Ord y ) ∨ ( ¬ ( Ord x == Ord y ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
90 ∅∨ {n} {x} {y} with trio< x y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
91 ∅∨ {n} {x} {y} | tri< a ¬b ¬c = case2 ( λ eq → ¬b ( ==→o≡' eq ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
92 ∅∨ {n} {x} {y} | tri≈ ¬a refl ¬c = case1 ( record { eq→ = id ; eq← = id } )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
93 ∅∨ {n} {x} {y} | tri> ¬a ¬b c = case2 ( λ eq → ¬b ( ==→o≡' eq ) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
94
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
95 ¬x∋x' : { n : Level } → { x : Ordinal {n} } → ¬ ( record { def = λ y → y o< x } ∋ record { def = λ y → y o< x } )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
96 ¬x∋x' {n} {record { lv = Zero ; ord = ord }} (case1 ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
97 ¬x∋x' {n} {record { lv = Suc lx ; ord = Φ .(Suc lx) }} (case1 (s≤s x)) = ¬x∋x' {n} {record { lv = lx ; ord = Φ lx }} (case1 {!!})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
98 ¬x∋x' {n} {record { lv = Suc lx ; ord = OSuc (Suc lx) ox }} (case1 (s≤s x)) = ¬x∋x' {n} {record { lv = Suc lx ; ord = ox}} (case1 {!!})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
99 ¬x∋x' {n} {record { lv = lv ; ord = Φ (lv) }} (case2 ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
100 ¬x∋x' {n} {record { lv = lv ; ord = OSuc (lv) ox }} (case2 x) =
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
101 ¬x∋x' {n} {record { lv = lv ; ord = ox }} (case2 {!!})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
102
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
103 ¬x∋x : { n : Level } → { x : OD {n} } → ¬ x ∋ x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
104 ¬x∋x = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
105
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
106 oc-lemma : { n : Level } → { x : OD {n} } → { oa : Ordinal {n} } → def (record { def = λ y → y o< oa }) oa → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
107 oc-lemma {n} {x} {oa} lt = o<¬≡ oa oa refl lt
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
108
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
109 oc-lemma1 : { n : Level } → { x : OD {n} } → { oa : Ordinal {n} } → od→ord (record { def = λ y → y o< oa }) o< oa → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
110 oc-lemma1 {n} {x} {oa} lt = ¬x∋x' {n} lt -- lt : def (record { def = λ y → y o< oa }) (record { def = λ y → y o< oa })
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
111
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
112 oc-lemma2 : { n : Level } → { x a : OD {n} } → { oa : Ordinal {n} } → oa o< od→ord (record { def = λ y → y o< oa }) → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
113 oc-lemma2 {n} {x} {oa} lt = {!!}
95
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
114
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
115 _c≤_ : {n : Level} → OD {n} → OD {n} → Set (suc n)
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
116 a c≤ b = (a ≡ b) ∨ ( b ∋ a )
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
117
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
118 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
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
119 def-subst df refl refl = df
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
120
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
121 o<-def : {n : Level } {x y : Ordinal {n} } → x o< y → def (record { def = λ x → x o< y }) x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
122 o<-def x<y = x<y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
123
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
124 def-o< : {n : Level } {x y : Ordinal {n} } → def (record { def = λ x → x o< y }) x → x o< y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
125 def-o< x<y = x<y
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
126
95
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
127 sup-od : {n : Level } → ( OD {n} → OD {n}) → OD {n}
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
128 sup-od ψ = ord→od ( sup-o ( λ x → od→ord (ψ (ord→od x ))) )
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
129
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
130 sup-c< : {n : Level } → ( ψ : OD {n} → OD {n}) → ∀ {x : OD {n}} → def ( sup-od ψ ) (od→ord ( ψ x ))
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
131 sup-c< {n} ψ {x} = def-subst {n} {_} {_} {sup-od ψ} {od→ord ( ψ x )}
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
132 {!!} refl (cong ( λ k → od→ord (ψ k) ) oiso)
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
133
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
134 od∅' : {n : Level} → OD {n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
135 od∅' = record { def = λ x → x o< o∅ }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
136
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
137 ∅0 : {n : Level} → od∅ {suc n} == record { def = λ x → x o< o∅ }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
138 eq→ ∅0 {w} (lift ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
139 eq← ∅0 {w} (case1 ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
140 eq← ∅0 {w} (case2 ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
141
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
142 ∅1 : {n : Level} → ( x : Ordinal {n} ) → ¬ ( x c< o∅ {n} )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
143 ∅1 {n} x lt = {!!}
28
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
144
37
f10ceee99d00 ¬ ( y c< x ) → x ≡ od∅
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 36
diff changeset
145 ∅3 : {n : Level} → { x : Ordinal {n}} → ( ∀(y : Ordinal {n}) → ¬ (y o< x ) ) → x ≡ o∅ {n}
81
96c932d0145d simpler ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 80
diff changeset
146 ∅3 {n} {x} = TransFinite {n} c2 c3 x where
30
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
147 c0 : Nat → Ordinal {n} → Set n
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
148 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
149 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
150 c2 Zero not = refl
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
151 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
152 ... | t with t (case1 ≤-refl )
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
153 c2 (Suc lx) not | t | ()
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
154 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
155 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
156 ... | t with t (case2 Φ< )
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
157 c3 lx (Φ .lx) d not | t | ()
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
158 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
159 ... | 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
160 c3 lx (OSuc .lx x₁) d not | t | ()
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
161
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
162 ∅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
163 ∅5 {n} {record { lv = Zero ; ord = (Φ .0) }} not = ⊥-elim (not refl)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
164 ∅5 {n} {record { lv = Zero ; ord = (OSuc .0 ord) }} not = case2 Φ<
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
165 ∅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
166
46
e584686a1307 == and ∅7
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 45
diff changeset
167 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
168 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
169
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
170 -- avoiding lv != Zero error
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
171 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
172 orefl refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
173
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
174 ==-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
175 ==-iso {n} {x} {y} eq = record {
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
176 eq→ = λ d → lemma ( eq→ eq (def-subst d (sym oiso) refl )) ;
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
177 eq← = λ d → lemma ( eq← eq (def-subst d (sym oiso) refl )) }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
178 where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
179 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
180 lemma {x} {z} d = def-subst d oiso refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
181
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
182 =-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
183 =-iso {_} {_} {y} = cong ( λ k → k == y ) (sym oiso)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
184
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
185 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
186 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
187 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
188 lemma ox ox refl = eq-refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
189
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
190 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
191 o≡→== {n} {x} {.x} refl = eq-refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
192
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
193 >→¬< : {x y : Nat } → (x < y ) → ¬ ( y < x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
194 >→¬< (s≤s x<y) (s≤s y<x) = >→¬< x<y y<x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
195
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
196 c≤-refl : {n : Level} → ( x : OD {n} ) → x c≤ x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
197 c≤-refl x = case1 refl
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
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
201 yx (def-subst {n} {ord→od (od→ord y)} {od→ord x} {!!} oiso refl )
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
202 ... | oyx with o<¬≡ (od→ord x) (od→ord x) refl {!!}
52
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
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
205 yx (def-subst {n} {ord→od (od→ord y)} {od→ord x} {!!} oiso refl )
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
206 ... | oyx with o<¬≡ (od→ord x) (od→ord x) refl {!!}
52
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
79
c07c548b2ac1 add some lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 78
diff changeset
209 ==→o≡ : {n : Level} → { x y : Ordinal {suc n} } → ord→od x == ord→od y → x ≡ y
c07c548b2ac1 add some lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 78
diff changeset
210 ==→o≡ {n} {x} {y} eq with trio< {n} x y
c07c548b2ac1 add some lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 78
diff changeset
211 ==→o≡ {n} {x} {y} eq | tri< a ¬b ¬c = ⊥-elim ( o<→o> eq (o<-subst a (sym ord-iso) (sym ord-iso )))
c07c548b2ac1 add some lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 78
diff changeset
212 ==→o≡ {n} {x} {y} eq | tri≈ ¬a b ¬c = b
c07c548b2ac1 add some lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 78
diff changeset
213 ==→o≡ {n} {x} {y} eq | tri> ¬a ¬b c = ⊥-elim ( o<→o> (eq-sym eq) (o<-subst c (sym ord-iso) (sym ord-iso )))
c07c548b2ac1 add some lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 78
diff changeset
214
90
38d139b5edac def ord conversion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 89
diff changeset
215 ≡-def : {n : Level} → { x : Ordinal {suc n} } → x ≡ od→ord (record { def = λ z → z o< x } )
38d139b5edac def ord conversion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 89
diff changeset
216 ≡-def {n} {x} = ==→o≡ {n} (subst (λ k → ord→od x == k ) (sym oiso) lemma ) where
38d139b5edac def ord conversion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 89
diff changeset
217 lemma : ord→od x == record { def = λ z → z o< x }
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
218 eq→ lemma {w} lt = {!!}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
219 -- ?subst₂ (λ k j → k o< j ) diso refl (subst (λ k → (od→ord ( ord→od w)) o< k ) diso t ) where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
220 --t : (od→ord ( ord→od w)) o< (od→ord (ord→od x))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
221 --t = o<-subst lt ? ?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
222 eq← lemma {w} lt = def-subst {suc n} {_} {_} {ord→od x} {w} {!!} refl refl
90
38d139b5edac def ord conversion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 89
diff changeset
223
38d139b5edac def ord conversion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 89
diff changeset
224 od≡-def : {n : Level} → { x : Ordinal {suc n} } → ord→od x ≡ record { def = λ z → z o< x }
38d139b5edac def ord conversion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 89
diff changeset
225 od≡-def {n} {x} = subst (λ k → ord→od x ≡ k ) oiso (cong ( λ k → ord→od k ) (≡-def {n} {x} ))
38d139b5edac def ord conversion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 89
diff changeset
226
99
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
227 ==→o≡1 : {n : Level} → { x y : OD {suc n} } → x == y → od→ord x ≡ od→ord y
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
228 ==→o≡1 eq = ==→o≡ (subst₂ (λ k j → k == j ) (sym oiso) (sym oiso) eq )
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
229
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
230 ==-def-l : {n : Level } {x y : Ordinal {suc n} } { z : OD {suc n} }→ (ord→od x == ord→od y) → def z x → def z y
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
231 ==-def-l {n} {x} {y} {z} eq z>x = subst ( λ k → def z k ) (==→o≡ eq) z>x
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
232
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
233 ==-def-r : {n : Level } {x y : OD {suc n} } { z : Ordinal {suc n} }→ (x == y) → def x z → def y z
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
234 ==-def-r {n} {x} {y} {z} eq z>x = subst (λ k → def k z ) (subst₂ (λ j k → j ≡ k ) oiso oiso (cong (λ k → ord→od k) (==→o≡1 eq))) z>x
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
235
91
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
236 ∋→o< : {n : Level} → { a x : OD {suc n} } → a ∋ x → od→ord x o< od→ord a
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
237 ∋→o< {n} {a} {x} lt = t where
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
238 t : (od→ord x) o< (od→ord a)
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
239 t = {!!}
91
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
240
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
241 o<∋→ : {n : Level} → { a x : OD {suc n} } → od→ord x o< od→ord a → a ∋ x
95
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
242 o<∋→ {n} {a} {x} lt = subst₂ (λ k j → def k j ) oiso refl t where
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
243 t : def (ord→od (od→ord a)) (od→ord x)
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
244 t = {!!}
91
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
245
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
246 o∅≡od∅ : {n : Level} → ord→od (o∅ {suc n}) ≡ od∅' {suc n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
247 o∅≡od∅ {n} with trio< {n} (o∅ {suc n}) (od→ord (od∅' {suc n} ))
80
461690d60d07 remove ∅-base-def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 79
diff changeset
248 o∅≡od∅ {n} | tri< a ¬b ¬c = ⊥-elim (lemma a) where
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
249 lemma : o∅ {suc n } o< (od→ord (od∅' {suc n} )) → ⊥
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
250 lemma lt with def-subst {suc n} {_} {_} {_} {_} ( o<→c< ( o<-subst lt (sym diso) refl ) ) refl diso
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
251 lemma lt | t = {!!}
80
461690d60d07 remove ∅-base-def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 79
diff changeset
252 o∅≡od∅ {n} | tri≈ ¬a b ¬c = trans (cong (λ k → ord→od k ) b ) oiso
461690d60d07 remove ∅-base-def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 79
diff changeset
253 o∅≡od∅ {n} | tri> ¬a ¬b c = ⊥-elim (¬x<0 c)
461690d60d07 remove ∅-base-def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 79
diff changeset
254
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
255 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
256 o<→¬== {n} {x} {y} lt eq = o<→o> eq lt
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
257
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
258 o<→¬c> : {n : Level} → { x y : Ordinal {n} } → x o< y → ¬ (y c< x )
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
259 o<→¬c> {n} {x} {y} olt clt = o<> olt {!!} where
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
260
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
261 o≡→¬c< : {n : Level} → { x y : Ordinal {n} } → x ≡ y → ¬ x c< y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
262 o≡→¬c< {n} {x} {y} oeq lt = o<¬≡ x y {!!} {!!}
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
263
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
264 tri-c< : {n : Level} → Trichotomous _≡_ (_c<_ {suc n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
265 tri-c< {n} x y with trio< {n} x y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
266 tri-c< {n} x y | tri< a ¬b ¬c = tri< (def-subst {!!} oiso refl) {!!} ( o<→¬c> a )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
267 tri-c< {n} x y | tri≈ ¬a b ¬c = tri≈ (o≡→¬c< b) {!!} (o≡→¬c< (sym b))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
268 tri-c< {n} x y | tri> ¬a ¬b c = tri> ( o<→¬c> c) (λ eq → {!!} ) (def-subst {!!} oiso refl)
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
269
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
270 c<> : {n : Level } { x y : Ordinal {suc n}} → x c< y → y c< x → ⊥
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
271 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
272 c<> {n} {x} {y} x<y y<x | tri< a ¬b ¬c = ¬c y<x
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
273 c<> {n} {x} {y} x<y y<x | tri≈ ¬a b ¬c = o<→o> {!!} {!!}
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
274 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
275
60
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
276 ∅< : {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
277 ∅< {n} {x} {y} d eq with eq→ eq d
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
278 ∅< {n} {x} {y} d eq | lift ()
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
279
78
9a7a64b2388c infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 77
diff changeset
280 ∅6 : {n : Level} → { x : OD {suc n} } → ¬ ( x ∋ x ) -- no Russel paradox
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
281 ∅6 {n} {x} x∋x = c<> {n} {{!!}} {{!!}} {!!} {!!}
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
282
76
8e8f54e7a030 extensionality done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 74
diff changeset
283 def-iso : {n : Level} {A B : OD {n}} {x y : Ordinal {n}} → x ≡ y → (def A y → def B y) → def A x → def B x
8e8f54e7a030 extensionality done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 74
diff changeset
284 def-iso refl t = t
8e8f54e7a030 extensionality done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 74
diff changeset
285
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
286 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
287 is-o∅ {n} record { lv = Zero ; ord = (Φ .0) } = yes refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
288 is-o∅ {n} record { lv = Zero ; ord = (OSuc .0 ord₁) } = no ( λ () )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
289 is-o∅ {n} record { lv = (Suc lv₁) ; ord = ord } = no (λ())
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
290
45
33860eb44e47 od∅' {n} = ord→od (o∅ {n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 44
diff changeset
291 open _∧_
33860eb44e47 od∅' {n} = ord→od (o∅ {n})
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 44
diff changeset
292
66
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
293 ¬∅=→∅∈ : {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
294 ¬∅=→∅∈ {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
295 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
296 lemma ox ne with is-o∅ ox
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
297 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
298 lemma1 : od→ord (ord→od o∅) ≡ od→ord od∅
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
299 lemma1 = cong ( λ k → od→ord k ) {!!}
66
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
300 lemma o∅ ne | yes refl | ()
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
301 lemma ox ne | no ¬p = subst ( λ k → def (ord→od ox) (od→ord k) ) {!!} {!!}
69
93abc0133b8a union continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 68
diff changeset
302
79
c07c548b2ac1 add some lemma
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 78
diff changeset
303 -- open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ )
94
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 93
diff changeset
304 -- postulate f-extensionality : { n : Level} → Relation.Binary.PropositionalEquality.Extensionality (suc n) (suc (suc n))
59
d13d1351a1fa lemma = cong₂ (λ x not → minimul x not ) oiso { }6
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 58
diff changeset
305
97
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
306 csuc : {n : Level} → OD {suc n} → OD {suc n}
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
307 csuc x = ord→od ( osuc ( od→ord x ))
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
308
96
f239ffc27fd0 Power Set and L
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 95
diff changeset
309 -- Power Set of X ( or constructible by λ y → def X (od→ord y )
97
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
310
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
311 ZFSubset : {n : Level} → (A x : OD {suc n} ) → OD {suc n}
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
312 ZFSubset A x = record { def = λ y → def A y ∧ def x y }
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
313
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
314 Def : {n : Level} → (A : OD {suc n}) → OD {suc n}
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
315 Def {n} A = record { def = λ y → y o< ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )))) }
96
f239ffc27fd0 Power Set and L
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 95
diff changeset
316
f239ffc27fd0 Power Set and L
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 95
diff changeset
317 -- Constructible Set on α
97
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
318 L : {n : Level} → (α : Ordinal {suc n}) → OD {suc n}
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
319 L {n} record { lv = Zero ; ord = (Φ .0) } = od∅
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
320 L {n} record { lv = lx ; ord = (OSuc lv ox) } = Def ( L {n} ( record { lv = lx ; ord = ox } ) )
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
321 L {n} record { lv = (Suc lx) ; ord = (Φ (Suc lx)) } = -- Union ( L α )
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
322 record { def = λ y → osuc y o< (od→ord (L {n} (record { lv = lx ; ord = Φ lx }) )) }
89
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 87
diff changeset
323
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
324 OD→ZF : {n : Level} → ZF {suc (suc n)} {suc n}
40
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
325 OD→ZF {n} = record {
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
326 ZFSet = OD {suc n}
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
327 ; _∋_ = _∋_
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
328 ; _≈_ = _==_
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
329 ; ∅ = od∅
28
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
330 ; _,_ = _,_
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
331 ; Union = Union
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
332 ; Power = Power
28
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
333 ; Select = Select
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
334 ; Replace = Replace
81
96c932d0145d simpler ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 80
diff changeset
335 ; infinite = ord→od ( record { lv = Suc Zero ; ord = Φ 1 } )
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
336 ; isZF = isZF
28
f36e40d5d2c3 OD continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 27
diff changeset
337 } where
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
338 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
339 Replace X ψ = sup-od ψ
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
340 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
341 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
342 _,_ : OD {suc n} → OD {suc n} → OD {suc n}
84
4b2aff372b7c omax ..
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 83
diff changeset
343 x , y = record { def = λ z → z o< (omax (od→ord x) (od→ord y)) }
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
344 Union : OD {suc n} → OD {suc n}
71
d088eb66564e add osuc ( next larger element of Ordinal )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 70
diff changeset
345 Union U = record { def = λ y → osuc y o< (od→ord U) }
77
75ba8cf64707 Power Set on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 76
diff changeset
346 -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ ( ∀ {x} → t ∋ x → X ∋ x )
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
347 Power : OD {suc n} → OD {suc n}
97
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
348 Power A = Def A
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
349 ZFSet = OD {suc n}
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
350 _∈_ : ( 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
351 A ∈ B = B ∋ A
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
352 _⊆_ : ( 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
353 _⊆_ A B {x} = A ∋ x → B ∋ x
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
354 _∩_ : ( A B : ZFSet ) → ZFSet
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
355 A ∩ B = Select (A , B) ( λ x → ( A ∋ x ) ∧ (B ∋ x) )
96
f239ffc27fd0 Power Set and L
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 95
diff changeset
356 -- _∪_ : ( A B : ZFSet ) → ZFSet
f239ffc27fd0 Power Set and L
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 95
diff changeset
357 -- A ∪ B = Select (A , B) ( λ x → (A ∋ x) ∨ ( B ∋ x ) )
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
358 {_} : ZFSet → ZFSet
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
359 { x } = ( x , x )
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
360 infixr 200 _∈_
96
f239ffc27fd0 Power Set and L
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 95
diff changeset
361 -- infixr 230 _∩_ _∪_
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
362 infixr 220 _⊆_
81
96c932d0145d simpler ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 80
diff changeset
363 isZF : IsZF (OD {suc n}) _∋_ _==_ od∅ _,_ Union Power Select Replace (ord→od ( record { lv = Suc Zero ; ord = Φ 1} ))
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
364 isZF = record {
43
0d9b9db14361 equalitu and internal parametorisity
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 42
diff changeset
365 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
366 ; pair = pair
97
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
367 ; union-u = λ _ z _ → csuc z
72
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 71
diff changeset
368 ; union→ = union→
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 71
diff changeset
369 ; union← = union←
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
370 ; empty = empty
76
8e8f54e7a030 extensionality done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 74
diff changeset
371 ; power→ = power→
8e8f54e7a030 extensionality done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 74
diff changeset
372 ; power← = power←
8e8f54e7a030 extensionality done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 74
diff changeset
373 ; extensionality = extensionality
30
3b0fdb95618e problem on Ordinal ( OSuc ℵ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 29
diff changeset
374 ; minimul = minimul
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
375 ; regularity = regularity
78
9a7a64b2388c infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 77
diff changeset
376 ; infinity∅ = infinity∅
93
d382a7902f5e replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 91
diff changeset
377 ; infinity = λ _ → infinity
55
9c0a5e28a572 regurality elimination case
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 54
diff changeset
378 ; selection = λ {ψ} {X} {y} → selection {ψ} {X} {y}
93
d382a7902f5e replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 91
diff changeset
379 ; replacement = replacement
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
380 } where
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
381 open _∧_
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
382 pair : (A B : OD {suc n} ) → ((A , B) ∋ A) ∧ ((A , B) ∋ B)
87
296388c03358 split omax?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 84
diff changeset
383 proj1 (pair A B ) = omax-x {n} (od→ord A) (od→ord B)
296388c03358 split omax?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 84
diff changeset
384 proj2 (pair A B ) = omax-y {n} (od→ord A) (od→ord B)
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
385 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
386 empty x ()
100
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
387 ---
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
388 --- ZFSubset A x = record { def = λ y → def A y ∧ def x y } subset of A
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
389 --- Power X = ord→od ( sup-o ( λ x → od→ord ( ZFSubset A (ord→od x )) ) ) Power X is a sup of all subset of A
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
390 --
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
391 -- if Power A ∋ t, from a minimulity of sup, there is osuc ZFSubset A ∋ t
100
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
392 -- then ZFSubset A ≡ t or ZFSubset A ∋ t. In the former case ZFSubset A ∋ x implies A ∋ x
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
393 -- In case of later, ZFSubset A ∋ t and t ∋ x implies ZFSubset A ∋ x by transitivity
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
394 --
76
8e8f54e7a030 extensionality done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 74
diff changeset
395 power→ : (A t : OD) → Power A ∋ t → {x : OD} → t ∋ x → A ∋ x
97
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
396 power→ A t P∋t {x} t∋x = proj1 lemma-s where
98
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 97
diff changeset
397 minsup : OD
99
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
398 minsup = ZFSubset A ( ord→od ( sup-x (λ x → od→ord ( ZFSubset A (ord→od x)))))
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
399 lemma-t : csuc minsup ∋ t
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
400 lemma-t = {!!}
98
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 97
diff changeset
401 lemma-s : ZFSubset A ( ord→od ( sup-x (λ x → od→ord ( ZFSubset A (ord→od x))))) ∋ x
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
402 lemma-s = {!!}
100
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
403 --
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
404 -- we have t ∋ x → A ∋ x means t is a subset of A, that is ZFSubset A t == t
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
405 -- Power A is a sup of ZFSubset A t, so Power A ∋ t
a402881cc341 add comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 99
diff changeset
406 --
77
75ba8cf64707 Power Set on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 76
diff changeset
407 power← : (A t : OD) → ({x : OD} → (t ∋ x → A ∋ x)) → Power A ∋ t
99
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
408 power← A t t→A = def-subst {suc n} {_} {_} {Power A} {od→ord t}
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
409 {!!} refl lemma1 where
97
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
410 lemma-eq : ZFSubset A t == t
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
411 eq→ lemma-eq {z} w = proj2 w
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
412 eq← lemma-eq {z} w = record { proj2 = w ;
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
413 proj1 = def-subst {suc n} {_} {_} {A} {z} ( t→A (def-subst {suc n} {_} {_} {t} {od→ord (ord→od z)} w refl (sym diso) )) refl diso }
99
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
414 lemma1 : od→ord (ZFSubset A (ord→od (od→ord t))) ≡ od→ord t
74330d0cdcbc Power Set done with min-sup assumption
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 98
diff changeset
415 lemma1 = subst (λ k → od→ord (ZFSubset A k) ≡ od→ord t ) (sym oiso) (==→o≡1 (lemma-eq))
97
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
416 lemma : od→ord (ZFSubset A (ord→od (od→ord t)) ) o< sup-o (λ x → od→ord (ZFSubset A (ord→od x)))
98
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 97
diff changeset
417 lemma = sup-o<
97
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
418 union-lemma-u : {X z : OD {suc n}} → (U>z : Union X ∋ z ) → csuc z ∋ z
72
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 71
diff changeset
419 union-lemma-u {X} {z} U>z = lemma <-osuc where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 71
diff changeset
420 lemma : {oz ooz : Ordinal {suc n}} → oz o< ooz → def (ord→od ooz) oz
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
421 lemma {oz} {ooz} lt = def-subst {suc n} {ord→od ooz} {!!} refl refl
73
dd430a95610f fix ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 72
diff changeset
422 union→ : (X z u : OD) → (X ∋ u) ∧ (u ∋ z) → Union X ∋ z
72
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 71
diff changeset
423 union→ X y u xx with trio< ( od→ord u ) ( osuc ( od→ord y ))
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
424 union→ X y u xx | tri< a ¬b ¬c with osuc-< a {!!}
74
819da8c08f05 ordinal atomical successor?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 73
diff changeset
425 union→ X y u xx | tri< a ¬b ¬c | ()
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
426 union→ X y u xx | tri≈ ¬a b ¬c = lemma b {!!} where
73
dd430a95610f fix ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 72
diff changeset
427 lemma : {oX ou ooy : Ordinal {suc n}} → ou ≡ ooy → ou o< oX → ooy o< oX
dd430a95610f fix ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 72
diff changeset
428 lemma refl lt = lt
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
429 union→ X y u xx | tri> ¬a ¬b c = ordtrans {suc n} {osuc ( od→ord y )} {od→ord u} {od→ord X} c {!!}
97
f2b579106770 power set using sup on Def
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 96
diff changeset
430 union← : (X z : OD) (X∋z : Union X ∋ z) → (X ∋ csuc z) ∧ (csuc z ∋ z )
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
431 union← X z X∋z = record { proj1 = def-subst {suc n} {_} {_} {X} {od→ord (csuc z )} {!!} oiso (sym diso) ; proj2 = union-lemma-u X∋z }
54
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
432 ψ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
433 ψiso {ψ} t refl = t
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
434 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
435 selection {ψ} {X} {y} = record {
33fb8228ace9 fix selection axiom
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 53
diff changeset
436 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
437 ; 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
438 }
93
d382a7902f5e replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 91
diff changeset
439 replacement : {ψ : OD → OD} (X x : OD) → Replace X ψ ∋ ψ x
d382a7902f5e replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 91
diff changeset
440 replacement {ψ} X x = sup-c< ψ {x}
60
6a1f67a4cc6e dead end
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 59
diff changeset
441 ∅-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
442 ∅-iso {x} neq = subst (λ k → ¬ k) (=-iso {n} ) neq
57
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
443 regularity : (x : OD) (not : ¬ (x == od∅)) →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 56
diff changeset
444 (x ∋ minimul x not) ∧ (Select (minimul x not) (λ x₁ → (minimul x not ∋ x₁) ∧ (x ∋ x₁)) == od∅)
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
445 proj1 (regularity x not ) = x∋minimul x not
66
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
446 proj2 (regularity x not ) = record { eq→ = reg ; eq← = λ () } where
92a11dc6425c regularity done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 65
diff changeset
447 reg : {y : Ordinal} → def (Select (minimul x not) (λ x₂ → (minimul x not ∋ x₂) ∧ (x ∋ x₂))) y → def od∅ y
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
448 reg {y} t with minimul-1 x not (ord→od y) (proj2 t )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
449 ... | t1 = lift t1
76
8e8f54e7a030 extensionality done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 74
diff changeset
450 extensionality : {A B : OD {suc n}} → ((z : OD) → (A ∋ z) ⇔ (B ∋ z)) → A == B
8e8f54e7a030 extensionality done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 74
diff changeset
451 eq→ (extensionality {A} {B} eq ) {x} d = def-iso {suc n} {A} {B} (sym diso) (proj1 (eq (ord→od x))) d
8e8f54e7a030 extensionality done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 74
diff changeset
452 eq← (extensionality {A} {B} eq ) {x} d = def-iso {suc n} {B} {A} (sym diso) (proj2 (eq (ord→od x))) d
89
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 87
diff changeset
453 xx-union : {x : OD {suc n}} → (x , x) ≡ record { def = λ z → z o< osuc (od→ord x) }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 87
diff changeset
454 xx-union {x} = cong ( λ k → record { def = λ z → z o< k } ) (omxx (od→ord x))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 87
diff changeset
455 xxx-union : {x : OD {suc n}} → (x , (x , x)) ≡ record { def = λ z → z o< osuc (osuc (od→ord x))}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 87
diff changeset
456 xxx-union {x} = cong ( λ k → record { def = λ z → z o< k } ) lemma where
91
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
457 lemma1 : {x : OD {suc n}} → od→ord x o< od→ord (x , x)
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
458 lemma1 {x} = {!!}
91
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
459 lemma2 : {x : OD {suc n}} → od→ord (x , x) ≡ osuc (od→ord x)
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
460 lemma2 = trans ( cong ( λ k → od→ord k ) xx-union ) (sym ≡-def)
89
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 87
diff changeset
461 lemma : {x : OD {suc n}} → omax (od→ord x) (od→ord (x , x)) ≡ osuc (osuc (od→ord x))
91
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
462 lemma {x} = trans ( sym ( omax< (od→ord x) (od→ord (x , x)) lemma1 ) ) ( cong ( λ k → osuc k ) lemma2 )
90
38d139b5edac def ord conversion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 89
diff changeset
463 uxxx-union : {x : OD {suc n}} → Union (x , (x , x)) ≡ record { def = λ z → osuc z o< osuc (osuc (od→ord x)) }
38d139b5edac def ord conversion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 89
diff changeset
464 uxxx-union {x} = cong ( λ k → record { def = λ z → osuc z o< k } ) lemma where
38d139b5edac def ord conversion
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 89
diff changeset
465 lemma : od→ord (x , (x , x)) ≡ osuc (osuc (od→ord x))
91
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
466 lemma = trans ( cong ( λ k → od→ord k ) xxx-union ) (sym ≡-def )
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
467 uxxx-2 : {x : OD {suc n}} → record { def = λ z → osuc z o< osuc (osuc (od→ord x)) } == record { def = λ z → z o< osuc (od→ord x) }
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
468 eq→ ( uxxx-2 {x} ) {m} lt = proj1 (osuc2 m (od→ord x)) lt
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
469 eq← ( uxxx-2 {x} ) {m} lt = proj2 (osuc2 m (od→ord x)) lt
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
470 uxxx-ord : {x : OD {suc n}} → od→ord (Union (x , (x , x))) ≡ osuc (od→ord x)
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
471 uxxx-ord {x} = trans (cong (λ k → od→ord k ) uxxx-union) (==→o≡ (subst₂ (λ j k → j == k ) (sym oiso) (sym od≡-def ) uxxx-2 ))
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
472 omega = record { lv = Suc Zero ; ord = Φ 1 }
78
9a7a64b2388c infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 77
diff changeset
473 infinite : OD {suc n}
91
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
474 infinite = ord→od ( omega )
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
475 infinity∅ : ord→od ( omega ) ∋ od∅ {suc n}
95
f3da2c87cee0 clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 94
diff changeset
476 infinity∅ = def-subst {suc n} {_} {o∅} {infinite} {od→ord od∅}
104
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 103
diff changeset
477 {!!} refl (subst ( λ k → ( k ≡ od→ord od∅ )) diso (cong (λ k → od→ord k) {!!} ))
91
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
478 infinite∋x : (x : OD) → infinite ∋ x → od→ord x o< omega
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
479 infinite∋x x lt = subst (λ k → od→ord x o< k ) diso t where
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
480 t : od→ord x o< od→ord (ord→od (omega))
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
481 t = ∋→o< {n} {infinite} {x} lt
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
482 infinite∋uxxx : (x : OD) → osuc (od→ord x) o< omega → infinite ∋ Union (x , (x , x ))
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
483 infinite∋uxxx x lt = o<∋→ t where
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
484 t : od→ord (Union (x , (x , x))) o< od→ord (ord→od (omega))
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
485 t = subst (λ k → od→ord (Union (x , (x , x))) o< k ) (sym diso ) ( subst ( λ k → k o< omega ) ( sym (uxxx-ord {x} ) ) lt )
78
9a7a64b2388c infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 77
diff changeset
486 infinity : (x : OD) → infinite ∋ x → infinite ∋ Union (x , (x , x ))
91
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
487 infinity x lt = infinite∋uxxx x ( lemma (od→ord x) (infinite∋x x lt )) where
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
488 lemma : (ox : Ordinal {suc n} ) → ox o< omega → osuc ox o< omega
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
489 lemma record { lv = Zero ; ord = (Φ .0) } (case1 (s≤s x)) = case1 (s≤s z≤n)
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
490 lemma record { lv = Zero ; ord = (OSuc .0 ord₁) } (case1 (s≤s x)) = case1 (s≤s z≤n)
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
491 lemma record { lv = (Suc lv₁) ; ord = (Φ .(Suc lv₁)) } (case1 (s≤s ()))
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
492 lemma record { lv = (Suc lv₁) ; ord = (OSuc .(Suc lv₁) ord₁) } (case1 (s≤s ()))
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
493 lemma record { lv = 1 ; ord = (Φ 1) } (case2 c2) with d<→lv c2
b4742cf4ef97 infinity axiom done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 90
diff changeset
494 lemma record { lv = (Suc Zero) ; ord = (Φ .1) } (case2 ()) | refl
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
495 -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] -- this form is no good since X is a transitive set
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
496 -- ∀ z [ ∀ x ( x ∈ z → ¬ ( x ≈ ∅ ) ) ∧ ∀ x ∀ y ( x , y ∈ z ∧ ¬ ( x ≈ y ) → x ∩ y ≈ ∅ ) → ∃ u ∀ x ( x ∈ z → ∃ t ( u ∩ x) ≈ { t }) ]
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
497 record Choice (z : OD {suc n}) : Set (suc (suc n)) where
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
498 field
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
499 u : {x : OD {suc n}} ( x∈z : x ∈ z ) → OD {suc n}
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
500 t : {x : OD {suc n}} ( x∈z : x ∈ z ) → (x : OD {suc n} ) → OD {suc n}
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
501 choice : { x : OD {suc n} } → ( x∈z : x ∈ z ) → ( u x∈z ∩ x) == { t x∈z x }
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
502 -- choice : {x : OD {suc n}} ( x ∈ z → ¬ ( x ≈ ∅ ) ) →
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
503 -- axiom-of-choice : { X : OD } → ( ¬x∅ : ¬ ( X == od∅ ) ) → { A : OD } → (A∈X : A ∈ X ) → choice ¬x∅ A∈X ∈ A
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
504 -- axiom-of-choice {X} nx {A} lt = ¬∅=→∅∈ {!!}
78
9a7a64b2388c infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 77
diff changeset
505
103
c8b79d303867 starting over HOD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 100
diff changeset
506