annotate cardinal.agda @ 253:0446b6c5e7bc

proudct uniquness done
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Thu, 29 Aug 2019 16:08:46 +0900
parents 8a58e2cd1f55
children 2ea2a19f9cd6
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
224
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
2 open import Ordinals
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
3 module cardinal {n : Level } (O : Ordinals {n}) where
3
e7990ff544bf reocrd ZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
4
14
e11e95d5ddee separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 11
diff changeset
5 open import zf
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
6 open import logic
224
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
7 import OD
23
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 22
diff changeset
8 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ; _⊔_ to _n⊔_ )
224
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
9 open import Relation.Binary.PropositionalEquality
14
e11e95d5ddee separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 11
diff changeset
10 open import Data.Nat.Properties
6
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
11 open import Data.Empty
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
12 open import Relation.Nullary
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
13 open import Relation.Binary
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
14 open import Relation.Binary.Core
d9b704508281 isEquiv and isZF
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 4
diff changeset
15
224
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
16 open inOrdinal O
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
17 open OD O
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
18 open OD.OD
29
fce60b99dc55 posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 28
diff changeset
19
120
ac214eab1c3c inifinite done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 119
diff changeset
20 open _∧_
213
22d435172d1a separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 210
diff changeset
21 open _∨_
22d435172d1a separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 210
diff changeset
22 open Bool
44
fcac01485f32 od→lv : {n : Level} → OD {n} → Nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 43
diff changeset
23
230
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
24 -- we have to work on Ordinal to keep OD Level n
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
25 -- since we use p∨¬p which works only on Level n
225
5f48299929ac does not work
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 224
diff changeset
26
233
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
27 <_,_> : (x y : OD) → OD
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
28 < x , y > = (x , x ) , (x , y )
226
176ff97547b4 set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 225
diff changeset
29
243
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 242
diff changeset
30 open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 242
diff changeset
31
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 242
diff changeset
32
248
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
33 open _==_
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
34
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
35 exg-pair : { x y : OD } → (x , y ) == ( y , x )
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
36 exg-pair {x} {y} = record { eq→ = left ; eq← = right } where
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
37 left : {z : Ordinal} → def (x , y) z → def (y , x) z
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
38 left (case1 t) = case2 t
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
39 left (case2 t) = case1 t
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
40 right : {z : Ordinal} → def (y , x) z → def (x , y) z
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
41 right (case1 t) = case2 t
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
42 right (case2 t) = case1 t
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
43
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
44 ==-trans : { x y z : OD } → x == y → y == z → x == z
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
45 ==-trans x=y y=z = record { eq→ = λ {m} t → eq→ y=z (eq→ x=y t) ; eq← = λ {m} t → eq← x=y (eq← y=z t) }
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
46
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
47 ==-sym : { x y : OD } → x == y → y == x
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
48 ==-sym x=y = record { eq→ = λ {m} t → eq← x=y t ; eq← = λ {m} t → eq→ x=y t }
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
49
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
50 ord≡→≡ : { x y : OD } → od→ord x ≡ od→ord y → x ≡ y
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
51 ord≡→≡ eq = subst₂ (λ j k → j ≡ k ) oiso oiso ( cong ( λ k → ord→od k ) eq )
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
52
251
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 250
diff changeset
53 od≡→≡ : { x y : Ordinal } → ord→od x ≡ ord→od y → x ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 250
diff changeset
54 od≡→≡ eq = subst₂ (λ j k → j ≡ k ) diso diso ( cong ( λ k → od→ord k ) eq )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 250
diff changeset
55
249
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
56 eq-prod : { x x' y y' : OD } → x ≡ x' → y ≡ y' → < x , y > ≡ < x' , y' >
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
57 eq-prod refl refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
58
248
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
59 prod-eq : { x x' y y' : OD } → < x , y > == < x' , y' > → (x ≡ x' ) ∧ ( y ≡ y' )
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
60 prod-eq {x} {x'} {y} {y'} eq = record { proj1 = lemmax ; proj2 = lemmay } where
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
61 lemma0 : {x y z : OD } → ( x , x ) == ( z , y ) → x ≡ y
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
62 lemma0 {x} {y} eq with trio< (od→ord x) (od→ord y)
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
63 lemma0 {x} {y} eq | tri< a ¬b ¬c with eq← eq {od→ord y} (case2 refl)
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
64 lemma0 {x} {y} eq | tri< a ¬b ¬c | case1 s = ⊥-elim ( o<¬≡ (sym s) a )
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
65 lemma0 {x} {y} eq | tri< a ¬b ¬c | case2 s = ⊥-elim ( o<¬≡ (sym s) a )
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
66 lemma0 {x} {y} eq | tri≈ ¬a b ¬c = ord≡→≡ b
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
67 lemma0 {x} {y} eq | tri> ¬a ¬b c with eq← eq {od→ord y} (case2 refl)
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
68 lemma0 {x} {y} eq | tri> ¬a ¬b c | case1 s = ⊥-elim ( o<¬≡ s c )
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
69 lemma0 {x} {y} eq | tri> ¬a ¬b c | case2 s = ⊥-elim ( o<¬≡ s c )
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
70 lemma2 : {x y z : OD } → ( x , x ) == ( z , y ) → z ≡ y
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
71 lemma2 {x} {y} {z} eq = trans (sym (lemma0 lemma3 )) ( lemma0 eq ) where
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
72 lemma3 : ( x , x ) == ( y , z )
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
73 lemma3 = ==-trans eq exg-pair
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
74 lemma1 : {x y : OD } → ( x , x ) == ( y , y ) → x ≡ y
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
75 lemma1 {x} {y} eq with eq← eq {od→ord y} (case2 refl)
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
76 lemma1 {x} {y} eq | case1 s = ord≡→≡ (sym s)
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
77 lemma1 {x} {y} eq | case2 s = ord≡→≡ (sym s)
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
78 lemma4 : {x y z : OD } → ( x , y ) == ( x , z ) → y ≡ z
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
79 lemma4 {x} {y} {z} eq with eq← eq {od→ord z} (case2 refl)
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
80 lemma4 {x} {y} {z} eq | case1 s with ord≡→≡ s -- x ≡ z
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
81 ... | refl with lemma2 (==-sym eq )
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
82 ... | refl = refl
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
83 lemma4 {x} {y} {z} eq | case2 s = ord≡→≡ (sym s) -- y ≡ z
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
84 lemmax : x ≡ x'
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
85 lemmax with eq→ eq {od→ord (x , x)} (case1 refl)
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
86 lemmax | case1 s = lemma1 (ord→== s ) -- (x,x)≡(x',x')
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
87 lemmax | case2 s with lemma2 (ord→== s ) -- (x,x)≡(x',y') with x'≡y'
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
88 ... | refl = lemma1 (ord→== s )
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
89 lemmay : y ≡ y'
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
90 lemmay with lemmax
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
91 ... | refl with lemma4 eq -- with (x,y)≡(x,y')
9fd85b954477 prod-eq done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 247
diff changeset
92 ... | eq1 = lemma4 (ord→== (cong (λ k → od→ord k ) eq1 ))
247
d09437fcfc7c fix pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 246
diff changeset
93
250
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 249
diff changeset
94
249
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
95 data ord-pair : (p : Ordinal) → Set n where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
96 pair : (x y : Ordinal ) → ord-pair ( od→ord ( < ord→od x , ord→od y > ) )
247
d09437fcfc7c fix pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 246
diff changeset
97
249
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
98 ZFProduct : OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
99 ZFProduct = record { def = λ x → ord-pair x }
247
d09437fcfc7c fix pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 246
diff changeset
100
249
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
101 eq-pair : { x x' y y' : Ordinal } → x ≡ x' → y ≡ y' → pair x y ≅ pair x' y'
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
102 eq-pair refl refl = HE.refl
247
d09437fcfc7c fix pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 246
diff changeset
103
249
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
104 pi1 : { p : Ordinal } → ord-pair p → Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
105 pi1 ( pair x y) = x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
106
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
107 π1 : { p : OD } → ZFProduct ∋ p → Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
108 π1 lt = pi1 lt
247
d09437fcfc7c fix pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 246
diff changeset
109
249
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
110 pi2 : { p : Ordinal } → ord-pair p → Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
111 pi2 ( pair x y ) = y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
112
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
113 π2 : { p : OD } → ZFProduct ∋ p → Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
114 π2 lt = pi2 lt
247
d09437fcfc7c fix pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 246
diff changeset
115
249
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
116 p-cons : ( x y : OD ) → ZFProduct ∋ < x , y >
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
117 p-cons x y = def-subst {_} {_} {ZFProduct} {od→ord (< x , y >)} (pair (od→ord x) ( od→ord y )) refl (
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
118 let open ≡-Reasoning in begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
119 od→ord < ord→od (od→ord x) , ord→od (od→ord y) >
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
120 ≡⟨ cong₂ (λ j k → od→ord < j , k >) oiso oiso ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 248
diff changeset
121 od→ord < x , y >
252
8a58e2cd1f55 give up product uniquness
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 251
diff changeset
122 ∎ )
250
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 249
diff changeset
123
251
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 250
diff changeset
124
252
8a58e2cd1f55 give up product uniquness
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 251
diff changeset
125 p-iso1 : { ox oy : Ordinal } → ZFProduct ∋ < ord→od ox , ord→od oy >
8a58e2cd1f55 give up product uniquness
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 251
diff changeset
126 p-iso1 {ox} {oy} = pair ox oy
251
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 250
diff changeset
127
253
0446b6c5e7bc proudct uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 252
diff changeset
128 p-iso : { x : OD } → (p : ZFProduct ∋ x ) → < ord→od (π1 p) , ord→od (π2 p) > ≡ x
0446b6c5e7bc proudct uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 252
diff changeset
129 p-iso {x} p = ord≡→≡ (lemma2 p) where
0446b6c5e7bc proudct uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 252
diff changeset
130 lemma : { op : Ordinal } → (q : ord-pair op ) → od→ord < ord→od (pi1 q) , ord→od (pi2 q) > ≡ op
0446b6c5e7bc proudct uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 252
diff changeset
131 lemma (pair ox oy) = refl
0446b6c5e7bc proudct uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 252
diff changeset
132 lemma2 : { x : OD } → (p : ZFProduct ∋ x ) → od→ord < ord→od (π1 p) , ord→od (π2 p) > ≡ od→ord x
0446b6c5e7bc proudct uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 252
diff changeset
133 lemma2 {x} p = lemma p
250
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 249
diff changeset
134
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 249
diff changeset
135
234
e06b76e5b682 ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 233
diff changeset
136 ∋-p : (A x : OD ) → Dec ( A ∋ x )
e06b76e5b682 ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 233
diff changeset
137 ∋-p A x with p∨¬p ( A ∋ x )
e06b76e5b682 ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 233
diff changeset
138 ∋-p A x | case1 t = yes t
e06b76e5b682 ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 233
diff changeset
139 ∋-p A x | case2 t = no t
e06b76e5b682 ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 233
diff changeset
140
233
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
141 _⊗_ : (A B : OD) → OD
239
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 238
diff changeset
142 A ⊗ B = record { def = λ x → def ZFProduct x ∧ ( { x : Ordinal } → (p : def ZFProduct x ) → checkAB p ) } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 238
diff changeset
143 checkAB : { p : Ordinal } → def ZFProduct p → Set n
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 238
diff changeset
144 checkAB (pair x y) = def A x ∧ def B y
233
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
145
242
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 241
diff changeset
146 func→od0 : (f : Ordinal → Ordinal ) → OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 241
diff changeset
147 func→od0 f = record { def = λ x → def ZFProduct x ∧ ( { x : Ordinal } → (p : def ZFProduct x ) → checkfunc p ) } where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 241
diff changeset
148 checkfunc : { p : Ordinal } → def ZFProduct p → Set n
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 241
diff changeset
149 checkfunc (pair x y) = f x ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 241
diff changeset
150
233
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
151 -- Power (Power ( A ∪ B )) ∋ ( A ⊗ B )
225
5f48299929ac does not work
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 224
diff changeset
152
233
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
153 Func : ( A B : OD ) → OD
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
154 Func A B = record { def = λ x → def (Power (A ⊗ B)) x }
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
155
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
156 -- power→ : ( A t : OD) → Power A ∋ t → {x : OD} → t ∋ x → ¬ ¬ (A ∋ x)
226
176ff97547b4 set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 225
diff changeset
157
236
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 235
diff changeset
158
233
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
159 func→od : (f : Ordinal → Ordinal ) → ( dom : OD ) → OD
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
160 func→od f dom = Replace dom ( λ x → < x , ord→od (f (od→ord x)) > )
af60c40298a4 function continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 230
diff changeset
161
242
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 241
diff changeset
162 record Func←cd { dom cod : OD } {f : Ordinal } : Set n where
236
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 235
diff changeset
163 field
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 235
diff changeset
164 func-1 : Ordinal → Ordinal
242
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 241
diff changeset
165 func→od∈Func-1 : Func dom cod ∋ func→od func-1 dom
236
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 235
diff changeset
166
242
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 241
diff changeset
167 od→func : { dom cod : OD } → {f : Ordinal } → def (Func dom cod ) f → Func←cd {dom} {cod} {f}
240
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
168 od→func {dom} {cod} {f} lt = record { func-1 = λ x → sup-o ( λ y → lemma x y ) ; func→od∈Func-1 = record { proj1 = {!!} ; proj2 = {!!} } } where
236
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 235
diff changeset
169 lemma : Ordinal → Ordinal → Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 235
diff changeset
170 lemma x y with IsZF.power→ isZF (dom ⊗ cod) (ord→od f) (subst (λ k → def (Power (dom ⊗ cod)) k ) (sym diso) lt ) | ∋-p (ord→od f) (ord→od y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 235
diff changeset
171 lemma x y | p | no n = o∅
240
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
172 lemma x y | p | yes f∋y = lemma2 (proj1 (double-neg-eilm ( p {ord→od y} f∋y ))) where -- p : {y : OD} → f ∋ y → ¬ ¬ (dom ⊗ cod ∋ y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
173 lemma2 : {p : Ordinal} → ord-pair p → Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
174 lemma2 (pair x1 y1) with decp ( x1 ≡ x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
175 lemma2 (pair x1 y1) | yes p = y1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
176 lemma2 (pair x1 y1) | no ¬p = o∅
242
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 241
diff changeset
177 fod : OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 241
diff changeset
178 fod = Replace dom ( λ x → < x , ord→od (sup-o ( λ y → lemma (od→ord x) y )) > )
240
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
179
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
180
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
181 open Func←cd
236
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 235
diff changeset
182
227
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 226
diff changeset
183 -- contra position of sup-o<
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 226
diff changeset
184 --
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 226
diff changeset
185
235
846e0926bb89 fix cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 234
diff changeset
186 -- postulate
846e0926bb89 fix cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 234
diff changeset
187 -- -- contra-position of mimimulity of supermum required in Cardinal
846e0926bb89 fix cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 234
diff changeset
188 -- sup-x : ( Ordinal → Ordinal ) → Ordinal
846e0926bb89 fix cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 234
diff changeset
189 -- sup-lb : { ψ : Ordinal → Ordinal } → {z : Ordinal } → z o< sup-o ψ → z o< osuc (ψ (sup-x ψ))
227
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 226
diff changeset
190
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
191 ------------
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
192 --
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
193 -- Onto map
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
194 -- def X x -> xmap
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
195 -- X ---------------------------> Y
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
196 -- ymap <- def Y y
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
197 --
224
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
198 record Onto (X Y : OD ) : Set n where
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
199 field
226
176ff97547b4 set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 225
diff changeset
200 xmap : Ordinal
176ff97547b4 set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 225
diff changeset
201 ymap : Ordinal
176ff97547b4 set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 225
diff changeset
202 xfunc : def (Func X Y) xmap
176ff97547b4 set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 225
diff changeset
203 yfunc : def (Func Y X) ymap
234
e06b76e5b682 ac from LEM in abstract ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 233
diff changeset
204 onto-iso : {y : Ordinal } → (lty : def Y y ) →
240
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
205 func-1 ( od→func {X} {Y} {xmap} xfunc ) ( func-1 (od→func yfunc) y ) ≡ y
230
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
206
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
207 open Onto
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
208
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
209 onto-restrict : {X Y Z : OD} → Onto X Y → ({x : OD} → _⊆_ Z Y {x}) → Onto X Z
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
210 onto-restrict {X} {Y} {Z} onto Z⊆Y = record {
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
211 xmap = xmap1
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
212 ; ymap = zmap
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
213 ; xfunc = xfunc1
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
214 ; yfunc = zfunc
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
215 ; onto-iso = onto-iso1
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
216 } where
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
217 xmap1 : Ordinal
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
218 xmap1 = od→ord (Select (ord→od (xmap onto)) {!!} )
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
219 zmap : Ordinal
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
220 zmap = {!!}
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
221 xfunc1 : def (Func X Z) xmap1
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
222 xfunc1 = {!!}
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
223 zfunc : def (Func Z X) zmap
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
224 zfunc = {!!}
240
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 239
diff changeset
225 onto-iso1 : {z : Ordinal } → (ltz : def Z z ) → func-1 (od→func xfunc1 ) (func-1 (od→func zfunc ) z ) ≡ z
230
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
226 onto-iso1 = {!!}
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
227
51
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 50
diff changeset
228
224
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
229 record Cardinal (X : OD ) : Set n where
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
230 field
224
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
231 cardinal : Ordinal
230
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
232 conto : Onto X (Ord cardinal)
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
233 cmax : ( y : Ordinal ) → cardinal o< y → ¬ Onto X (Ord y)
151
b5a337fb7a6d recovering...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 150
diff changeset
234
224
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
235 cardinal : (X : OD ) → Cardinal X
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
236 cardinal X = record {
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
237 cardinal = sup-o ( λ x → proj1 ( cardinal-p x) )
226
176ff97547b4 set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 225
diff changeset
238 ; conto = onto
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
239 ; cmax = cmax
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
240 } where
230
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
241 cardinal-p : (x : Ordinal ) → ( Ordinal ∧ Dec (Onto X (Ord x) ) )
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
242 cardinal-p x with p∨¬p ( Onto X (Ord x) )
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
243 cardinal-p x | case1 True = record { proj1 = x ; proj2 = yes True }
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
244 cardinal-p x | case2 False = record { proj1 = o∅ ; proj2 = no False }
229
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 228
diff changeset
245 S = sup-o (λ x → proj1 (cardinal-p x))
230
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
246 lemma1 : (x : Ordinal) → ((y : Ordinal) → y o< x → Lift (suc n) (y o< (osuc S) → Onto X (Ord y))) →
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
247 Lift (suc n) (x o< (osuc S) → Onto X (Ord x) )
229
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 228
diff changeset
248 lemma1 x prev with trio< x (osuc S)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 228
diff changeset
249 lemma1 x prev | tri< a ¬b ¬c with osuc-≡< a
230
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
250 lemma1 x prev | tri< a ¬b ¬c | case1 x=S = lift ( λ lt → {!!} )
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
251 lemma1 x prev | tri< a ¬b ¬c | case2 x<S = lift ( λ lt → lemma2 ) where
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
252 lemma2 : Onto X (Ord x)
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
253 lemma2 with prev {!!} {!!}
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
254 ... | lift t = t {!!}
229
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 228
diff changeset
255 lemma1 x prev | tri≈ ¬a b ¬c = lift ( λ lt → ⊥-elim ( o<¬≡ b lt ))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 228
diff changeset
256 lemma1 x prev | tri> ¬a ¬b c = lift ( λ lt → ⊥-elim ( o<> c lt ))
230
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
257 onto : Onto X (Ord S)
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
258 onto with TransFinite {λ x → Lift (suc n) ( x o< osuc S → Onto X (Ord x) ) } lemma1 S
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
259 ... | lift t = t <-osuc
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
260 cmax : (y : Ordinal) → S o< y → ¬ Onto X (Ord y)
229
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 228
diff changeset
261 cmax y lt ontoy = o<> lt (o<-subst {_} {_} {y} {S}
224
afc864169325 recover ε-induction
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 219
diff changeset
262 (sup-o< {λ x → proj1 ( cardinal-p x)}{y} ) lemma refl ) where
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
263 lemma : proj1 (cardinal-p y) ≡ y
230
1b1620e2053c we need ordered pair
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 229
diff changeset
264 lemma with p∨¬p ( Onto X (Ord y) )
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
265 lemma | case1 x = refl
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
266 lemma | case2 not = ⊥-elim ( not ontoy )
217
d5668179ee69 cardinal continue
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 216
diff changeset
267
226
176ff97547b4 set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 225
diff changeset
268
176ff97547b4 set theortic function definition using sup
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 225
diff changeset
269 -----
219
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
270 -- All cardinal is ℵ0, since we are working on Countable Ordinal,
43021d2b8756 separate cardinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 218
diff changeset
271 -- Power ω is larger than ℵ0, so it has no cardinal.
218
eee983e4b402 try func
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 217
diff changeset
272
eee983e4b402 try func
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 217
diff changeset
273
eee983e4b402 try func
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 217
diff changeset
274