Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate constructible-set.agda @ 21:6d9fdd1dfa79
add transfinite
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Thu, 16 May 2019 16:47:27 +0900 |
parents | 31626f36cd94 |
children | 3da2c00bd24d |
rev | line source |
---|---|
16 | 1 open import Level |
2 module constructible-set (n : Level) where | |
3 | 3 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
4 open import zf |
3 | 5 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
6 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ) |
3 | 7 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
8 open import Relation.Binary.PropositionalEquality |
3 | 9 |
17 | 10 data OrdinalD : (lv : Nat) → Set n where |
11 Φ : {lv : Nat} → OrdinalD lv | |
12 OSuc : {lv : Nat} → OrdinalD lv → OrdinalD lv | |
13 ℵ_ : (lv : Nat) → OrdinalD (Suc lv) | |
3 | 14 |
16 | 15 record Ordinal : Set n where |
16 field | |
17 lv : Nat | |
17 | 18 ord : OrdinalD lv |
16 | 19 |
17 | 20 data _d<_ : {lx ly : Nat} → OrdinalD lx → OrdinalD ly → Set n where |
21 Φ< : {lx : Nat} → {x : OrdinalD lx} → Φ {lx} d< OSuc {lx} x | |
22 s< : {lx : Nat} → {x y : OrdinalD lx} → x d< y → OSuc {lx} x d< OSuc {lx} y | |
23 ℵΦ< : {lx : Nat} → {x : OrdinalD (Suc lx) } → Φ {Suc lx} d< (ℵ lx) | |
24 ℵ< : {lx : Nat} → {x : OrdinalD (Suc lx) } → OSuc {Suc lx} x d< (ℵ lx) | |
25 | |
26 open Ordinal | |
27 | |
28 _o<_ : ( x y : Ordinal ) → Set n | |
29 _o<_ x y = (lv x < lv y ) ∨ ( ord x d< ord y ) | |
3 | 30 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
31 open import Data.Nat.Properties |
6 | 32 open import Data.Empty |
33 open import Relation.Nullary | |
34 | |
35 open import Relation.Binary | |
36 open import Relation.Binary.Core | |
37 | |
21 | 38 o∅ : Ordinal |
39 o∅ = record { lv = Zero ; ord = Φ } | |
40 | |
41 TransFinite : ( ψ : Ordinal → Set n ) | |
42 → ( ∀ (lx : Nat ) → ψ ( record { lv = Suc lx ; ord = ℵ lx } )) | |
43 → ( ∀ (lx : Nat ) → ψ ( record { lv = lx ; ord = Φ } ) ) | |
44 → ( ∀ (lx : Nat ) → (x : OrdinalD lx ) → ψ ( record { lv = lx ; ord = x } ) → ψ ( record { lv = lx ; ord = OSuc x } ) ) | |
45 → ∀ (x : Ordinal) → ψ x | |
46 TransFinite ψ caseℵ caseΦ caseOSuc record { lv = lv ; ord = Φ } = caseΦ lv | |
47 TransFinite ψ caseℵ caseΦ caseOSuc record { lv = lv ; ord = OSuc ord₁ } = caseOSuc lv ord₁ | |
48 ( TransFinite ψ caseℵ caseΦ caseOSuc (record { lv = lv ; ord = ord₁ } )) | |
49 TransFinite ψ caseℵ caseΦ caseOSuc record { lv = Suc lv₁ ; ord = ℵ lv₁ } = caseℵ lv₁ | |
50 | |
51 record SupR (ψ : Ordinal → Ordinal ) : Set n where | |
52 field | |
53 sup : Ordinal | |
54 smax : { x : Ordinal } → ψ x o< sup | |
55 | |
56 open SupR | |
57 | |
58 Sup : (ψ : Ordinal → Ordinal ) → SupR ψ | |
59 sup (Sup ψ) = {!!} | |
60 smax (Sup ψ) = {!!} | |
61 | |
62 | |
17 | 63 ≡→¬d< : {lv : Nat} → {x : OrdinalD lv } → x d< x → ⊥ |
64 ≡→¬d< {lx} {OSuc y} (s< t) = ≡→¬d< t | |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
65 |
17 | 66 trio<> : {lx : Nat} {x : OrdinalD lx } { y : OrdinalD lx } → y d< x → x d< y → ⊥ |
16 | 67 trio<> {lx} {.(OSuc _)} {.(OSuc _)} (s< s) (s< t) = |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
68 trio<> s t |
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
69 |
17 | 70 trio<≡ : {lx : Nat} {x : OrdinalD lx } { y : OrdinalD lx } → x ≡ y → x d< y → ⊥ |
71 trio<≡ refl = ≡→¬d< | |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
72 |
17 | 73 trio>≡ : {lx : Nat} {x : OrdinalD lx } { y : OrdinalD lx } → x ≡ y → y d< x → ⊥ |
74 trio>≡ refl = ≡→¬d< | |
9
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
75 |
17 | 76 triO : {lx ly : Nat} → OrdinalD lx → OrdinalD ly → Tri (lx < ly) ( lx ≡ ly ) ( lx > ly ) |
16 | 77 triO {lx} {ly} x y = <-cmp lx ly |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
78 |
17 | 79 triOrdd : {lx : Nat} → Trichotomous _≡_ ( _d<_ {lx} {lx} ) |
80 triOrdd {lv} Φ Φ = tri≈ ≡→¬d< refl ≡→¬d< | |
81 triOrdd {Suc lv} (ℵ lv) (ℵ lv) = tri≈ ≡→¬d< refl ≡→¬d< | |
82 triOrdd {lv} Φ (OSuc y) = tri< Φ< (λ ()) ( λ lt → trio<> lt Φ< ) | |
83 triOrdd {.(Suc lv)} Φ (ℵ lv) = tri< (ℵΦ< {lv} {Φ} ) (λ ()) ( λ lt → trio<> lt ((ℵΦ< {lv} {Φ} )) ) | |
84 triOrdd {Suc lv} (ℵ lv) Φ = tri> ( λ lt → trio<> lt (ℵΦ< {lv} {Φ} ) ) (λ ()) (ℵΦ< {lv} {Φ} ) | |
85 triOrdd {Suc lv} (ℵ lv) (OSuc y) = tri> ( λ lt → trio<> lt (ℵ< {lv} {y} ) ) (λ ()) (ℵ< {lv} {y} ) | |
86 triOrdd {lv} (OSuc x) Φ = tri> (λ lt → trio<> lt Φ<) (λ ()) Φ< | |
87 triOrdd {.(Suc lv)} (OSuc x) (ℵ lv) = tri< ℵ< (λ ()) (λ lt → trio<> lt ℵ< ) | |
88 triOrdd {lv} (OSuc x) (OSuc y) with triOrdd x y | |
89 triOrdd {lv} (OSuc x) (OSuc y) | tri< a ¬b ¬c = tri< (s< a) (λ tx=ty → trio<≡ tx=ty (s< a) ) ( λ lt → trio<> lt (s< a) ) | |
90 triOrdd {lv} (OSuc x) (OSuc x) | tri≈ ¬a refl ¬c = tri≈ ≡→¬d< refl ≡→¬d< | |
91 triOrdd {lv} (OSuc x) (OSuc y) | tri> ¬a ¬b c = tri> ( λ lt → trio<> lt (s< c) ) (λ tx=ty → trio>≡ tx=ty (s< c) ) (s< c) | |
9
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
92 |
17 | 93 d<→lv : {x y : Ordinal } → ord x d< ord y → lv x ≡ lv y |
94 d<→lv Φ< = refl | |
95 d<→lv (s< lt) = refl | |
96 d<→lv ℵΦ< = refl | |
97 d<→lv ℵ< = refl | |
16 | 98 |
17 | 99 orddtrans : {lx : Nat} {x y z : OrdinalD lx } → x d< y → y d< z → x d< z |
100 orddtrans {lx} {.Φ} {.(OSuc _)} {.(OSuc _)} Φ< (s< y<z) = Φ< | |
101 orddtrans {Suc lx} {Φ {Suc lx}} {OSuc y} {ℵ lx} Φ< ℵ< = ℵΦ< {lx} {y} | |
102 orddtrans {lx} {.(OSuc _)} {.(OSuc _)} {.(OSuc _)} (s< x<y) (s< y<z) = s< ( orddtrans x<y y<z ) | |
103 orddtrans {.(Suc _)} {.(OSuc _)} {.(OSuc _)} {.(ℵ _)} (s< x<y) ℵ< = ℵ< | |
104 orddtrans {.(Suc _)} {.Φ} {.(ℵ _)} {z} ℵΦ< () | |
105 orddtrans {.(Suc _)} {.(OSuc _)} {.(ℵ _)} {z} ℵ< () | |
9
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
106 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
107 max : (x y : Nat) → Nat |
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
108 max Zero Zero = Zero |
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
109 max Zero (Suc x) = (Suc x) |
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
110 max (Suc x) Zero = (Suc x) |
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
111 max (Suc x) (Suc y) = Suc ( max x y ) |
3 | 112 |
17 | 113 maxαd : { lx : Nat } → OrdinalD lx → OrdinalD lx → OrdinalD lx |
114 maxαd x y with triOrdd x y | |
115 maxαd x y | tri< a ¬b ¬c = y | |
116 maxαd x y | tri≈ ¬a b ¬c = x | |
117 maxαd x y | tri> ¬a ¬b c = x | |
6 | 118 |
17 | 119 maxα : Ordinal → Ordinal → Ordinal |
120 maxα x y with <-cmp (lv x) (lv y) | |
121 maxα x y | tri< a ¬b ¬c = x | |
122 maxα x y | tri> ¬a ¬b c = y | |
123 maxα x y | tri≈ ¬a refl ¬c = record { lv = lv x ; ord = maxαd (ord x) (ord y) } | |
7 | 124 |
17 | 125 OrdTrans : Transitive (λ ( a b : Ordinal ) → (a ≡ b) ∨ (Ordinal.lv a < Ordinal.lv b) ∨ (Ordinal.ord a d< Ordinal.ord b) ) |
16 | 126 OrdTrans (case1 refl) (case1 refl) = case1 refl |
127 OrdTrans (case1 refl) (case2 lt2) = case2 lt2 | |
128 OrdTrans (case2 lt1) (case1 refl) = case2 lt1 | |
17 | 129 OrdTrans (case2 (case1 x)) (case2 (case1 y)) = case2 (case1 ( <-trans x y ) ) |
130 OrdTrans (case2 (case1 x)) (case2 (case2 y)) with d<→lv y | |
131 OrdTrans (case2 (case1 x)) (case2 (case2 y)) | refl = case2 (case1 x ) | |
132 OrdTrans (case2 (case2 x)) (case2 (case1 y)) with d<→lv x | |
133 OrdTrans (case2 (case2 x)) (case2 (case1 y)) | refl = case2 (case1 y) | |
134 OrdTrans (case2 (case2 x)) (case2 (case2 y)) with d<→lv x | d<→lv y | |
135 OrdTrans (case2 (case2 x)) (case2 (case2 y)) | refl | refl = case2 (case2 (orddtrans x y )) | |
16 | 136 |
137 OrdPreorder : Preorder n n n | |
138 OrdPreorder = record { Carrier = Ordinal | |
139 ; _≈_ = _≡_ | |
17 | 140 ; _∼_ = λ a b → (a ≡ b) ∨ ( a o< b ) |
16 | 141 ; isPreorder = record { |
142 isEquivalence = record { refl = refl ; sym = sym ; trans = trans } | |
143 ; reflexive = case1 | |
144 ; trans = OrdTrans | |
145 } | |
146 } | |
147 | |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
148 -- X' = { x ∈ X | ψ x } ∪ X , Mα = ( ∪ (β < α) Mβ ) ' |
7 | 149 |
19 | 150 record ConstructibleSet : Set (suc (suc n)) where |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
151 field |
17 | 152 α : Ordinal |
19 | 153 constructible : Ordinal → Set (suc n) |
11 | 154 |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
155 open ConstructibleSet |
11 | 156 |
20 | 157 _∋_ : (ConstructibleSet ) → (ConstructibleSet ) → Set (suc n) |
158 a ∋ x = ((α a ≡ α x) ∨ ( α x o< α a )) | |
159 ∧ constructible a ( α x ) | |
160 | |
11 | 161 |
17 | 162 -- transitiveness : (a b c : ConstructibleSet ) → a ∋ b → b ∋ c → a ∋ c |
163 -- transitiveness a b c a∋b b∋c with constructible a c∋ constructible b | constructible b c∋ constructible c | |
164 -- ... | t1 | t2 = {!!} | |
15 | 165 |
20 | 166 open import Data.Unit |
167 | |
19 | 168 ConstructibleSet→ZF : ZF {suc (suc n)} {suc (suc n)} |
16 | 169 ConstructibleSet→ZF = record { |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
170 ZFSet = ConstructibleSet |
19 | 171 ; _∋_ = λ a b → Lift (suc (suc n)) ( a ∋ b ) |
18 | 172 ; _≈_ = _≡_ |
21 | 173 ; ∅ = record {α = o∅ ; constructible = λ x → Lift (suc n) ⊥ } |
18 | 174 ; _,_ = _,_ |
175 ; Union = Union | |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
176 ; Power = {!!} |
20 | 177 ; Select = Select |
178 ; Replace = Replace | |
14
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
179 ; infinite = {!!} |
e11e95d5ddee
separete constructible set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
11
diff
changeset
|
180 ; isZF = {!!} |
18 | 181 } where |
20 | 182 conv : (ConstructibleSet → Set (suc (suc n))) → ConstructibleSet → Set (suc n) |
183 conv ψ x with ψ x | |
184 ... | t = Lift ( suc n ) ⊤ | |
185 Select : (X : ConstructibleSet) → (ConstructibleSet → Set (suc (suc n))) → ConstructibleSet | |
186 Select X ψ = record { α = α X ; constructible = λ x → (conv ψ) (record { α = x ; constructible = λ x → constructible X x } ) } | |
187 Replace : (X : ConstructibleSet) → (ConstructibleSet → ConstructibleSet) → ConstructibleSet | |
188 Replace X ψ = record { α = α X ; constructible = λ x → {!!} } | |
18 | 189 _,_ : ConstructibleSet → ConstructibleSet → ConstructibleSet |
20 | 190 a , b = record { α = maxα (α a) (α b) ; constructible = λ x → {!!} } |
18 | 191 Union : ConstructibleSet → ConstructibleSet |
192 Union a = {!!} |