Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate zf.agda @ 11:2df90eb0896c
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Mon, 13 May 2019 20:51:45 +0900 |
parents | 8022e14fce74 |
children | b531d2b417ad e11e95d5ddee |
rev | line source |
---|---|
3 | 1 module zf where |
2 | |
3 open import Level | |
4 | |
5 | |
6 record _∧_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where | |
7 field | |
8 proj1 : A | |
9 proj2 : B | |
10 | |
11 open _∧_ | |
12 | |
13 | |
14 data _∨_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where | |
15 case1 : A → A ∨ B | |
16 case2 : B → A ∨ B | |
17 | |
6 | 18 -- open import Relation.Binary.PropositionalEquality |
3 | 19 |
20 _⇔_ : {n : Level } → ( A B : Set n ) → Set n | |
21 _⇔_ A B = ( A → B ) ∧ ( B → A ) | |
22 | |
6 | 23 open import Data.Empty |
24 open import Relation.Nullary | |
25 | |
26 open import Relation.Binary | |
27 open import Relation.Binary.Core | |
28 | |
3 | 29 infixr 130 _∧_ |
30 infixr 140 _∨_ | |
31 infixr 150 _⇔_ | |
32 | |
9
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
33 record Func {n m : Level } (ZFSet : Set n) (_≈_ : Rel ZFSet m) : Set (n ⊔ suc m) where |
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
34 field |
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
35 rel : Rel ZFSet m |
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
36 dom : ( y : ZFSet ) → ∀ { x : ZFSet } → rel x y |
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
37 fun-eq : { x y z : ZFSet } → ( rel x y ∧ rel x z ) → y ≈ z |
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
38 |
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
39 open Func |
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
40 |
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
41 |
6 | 42 record IsZF {n m : Level } |
43 (ZFSet : Set n) | |
44 (_∋_ : ( A x : ZFSet ) → Set m) | |
9
5ed16e2d8eb7
try to fix axiom of replacement
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
8
diff
changeset
|
45 (_≈_ : Rel ZFSet m) |
6 | 46 (∅ : ZFSet) |
47 (_×_ : ( A B : ZFSet ) → ZFSet) | |
48 (Union : ( A : ZFSet ) → ZFSet) | |
49 (Power : ( A : ZFSet ) → ZFSet) | |
10 | 50 (Select : ( ZFSet → Set m ) → ZFSet ) |
51 (Replace : ( ZFSet → ZFSet ) → ZFSet ) | |
6 | 52 (infinite : ZFSet) |
53 : Set (suc (n ⊔ m)) where | |
3 | 54 field |
6 | 55 isEquivalence : {A B : ZFSet} → IsEquivalence {n} {m} {ZFSet} _≈_ |
3 | 56 -- ∀ x ∀ y ∃ z(x ∈ z ∧ y ∈ z) |
6 | 57 pair : ( A B : ZFSet ) → ( (A × B) ∋ A ) ∧ ( (A × B) ∋ B ) |
3 | 58 -- ∀ X ∃ A∀ t(t ∈ A ⇔ ∃ x ∈ X(t ∈ x)) |
59 union→ : ( X x y : ZFSet ) → X ∋ x → x ∋ y → Union X ∋ y | |
60 union← : ( X x y : ZFSet ) → Union X ∋ y → X ∋ x → x ∋ y | |
61 _∈_ : ( A B : ZFSet ) → Set m | |
62 A ∈ B = B ∋ A | |
7 | 63 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → ∀{ A∋x : Set m } → Set m |
3 | 64 _⊆_ A B {x} {A∋x} = B ∋ x |
65 _∩_ : ( A B : ZFSet ) → ZFSet | |
10 | 66 A ∩ B = Select ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) |
3 | 67 _∪_ : ( A B : ZFSet ) → ZFSet |
10 | 68 A ∪ B = Select ( λ x → ( A ∋ x ) ∨ ( B ∋ x ) ) |
3 | 69 infixr 200 _∈_ |
70 infixr 230 _∩_ _∪_ | |
71 infixr 220 _⊆_ | |
72 field | |
4 | 73 empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x ) |
3 | 74 -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ t ⊆ X ) ) |
8 | 75 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} {y} → _⊆_ t A {x} {y} |
76 power← : ∀( A t : ZFSet ) → ∀ {x} {y} → _⊆_ t A {x} {y} → Power A ∋ t | |
3 | 77 -- extentionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w ) |
6 | 78 extentionality : ( A B z : ZFSet ) → (( A ∋ z ) ⇔ (B ∋ z) ) → A ≈ B |
3 | 79 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) ) |
10 | 80 minimul : ZFSet → ZFSet |
81 regularity : ∀( x : ZFSet ) → ¬ (x ≈ ∅) → ( minimul x ∈ x ∧ ( minimul x ∩ x ≈ ∅ ) ) | |
3 | 82 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) ) |
83 infinity∅ : ∅ ∈ infinite | |
10 | 84 infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ Select ( λ y → x ≈ y )) ∈ infinite |
85 selection : { ψ : ZFSet → Set m } → ∀ ( y : ZFSet ) → ( y ∈ Select ψ ) → ψ y | |
3 | 86 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) ) |
10 | 87 replacement : {ψ : ZFSet → ZFSet} → ∀ ( x : ZFSet ) → ( ψ x ∈ Replace ψ ) |
3 | 88 |
6 | 89 record ZF {n m : Level } : Set (suc (n ⊔ m)) where |
90 infixr 210 _×_ | |
91 infixl 200 _∋_ | |
92 infixr 220 _≈_ | |
93 field | |
94 ZFSet : Set n | |
95 _∋_ : ( A x : ZFSet ) → Set m | |
96 _≈_ : ( A B : ZFSet ) → Set m | |
97 -- ZF Set constructor | |
98 ∅ : ZFSet | |
99 _×_ : ( A B : ZFSet ) → ZFSet | |
100 Union : ( A : ZFSet ) → ZFSet | |
101 Power : ( A : ZFSet ) → ZFSet | |
10 | 102 Select : ( ZFSet → Set m ) → ZFSet |
103 Replace : ( ZFSet → ZFSet ) → ZFSet | |
6 | 104 infinite : ZFSet |
10 | 105 isZF : IsZF ZFSet _∋_ _≈_ ∅ _×_ Union Power Select Replace infinite |
6 | 106 |
10 | 107 module zf-exapmle {n m : Level } ( zf : ZF {m} {n} ) where |
7 | 108 |
10 | 109 _≈_ = ZF._≈_ zf |
110 ZFSet = ZF.ZFSet zf | |
111 Select = ZF.Select zf | |
112 ∅ = ZF.∅ zf | |
113 _∩_ = ( IsZF._∩_ ) (ZF.isZF zf) | |
114 _∋_ = ZF._∋_ zf | |
115 replacement = IsZF.replacement ( ZF.isZF zf ) | |
116 selection = IsZF.selection ( ZF.isZF zf ) | |
117 minimul = IsZF.minimul ( ZF.isZF zf ) | |
118 regularity = IsZF.regularity ( ZF.isZF zf ) | |
7 | 119 |
11 | 120 -- russel : Select ( λ x → x ∋ x ) ≈ ∅ |
121 -- russel with Select ( λ x → x ∋ x ) | |
122 -- ... | s = {!!} | |
7 | 123 |
10 | 124 module constructible-set where |
7 | 125 |
11 | 126 open import Data.Nat renaming ( zero to Zero ; suc to Suc ; ℕ to Nat ) |
10 | 127 |
128 open import Relation.Binary.PropositionalEquality | |
129 | |
11 | 130 data Ordinal {n : Level } : Set n where |
131 Φ : {lv : Nat} → Ordinal {n} lv | |
132 T-suc : {lv : Nat} → Ordinal {n} lv → Ordinal lv | |
133 ℵ_ : (lv : Nat) → Ordinal (Suc lv) | |
134 | |
135 data _o<_ {n : Level } : Ordinal {n} → Ordinal {n} → Set n where | |
136 l< : {lx ly : Nat } → {x : Ordinal {n} lx } → {y : Ordinal {n} ly } → lx < ly → x o< y | |
137 Φ< : {lx : Nat} → {x : Ordinal {n} lx} → Φ {n} {lx} o< T-suc {n} {lx} x | |
138 s< : {lx : Nat} → {x : Ordinal {n} lx} → x o< T-suc {n} {lx} x | |
139 ℵΦ< : {lx : Nat} → {x : Ordinal {n} lx } → Φ {n} {lx} o< (ℵ lx) | |
140 ℵ< : {lx : Nat} → {x : Ordinal {n} lx } → T-suc {n} {lx} x o< (ℵ lx) | |
141 | |
142 _o≈_ : {n : Level } {lv : Nat } → Rel ( Ordinal {n} lv ) n | |
143 _o≈_ = {!!} | |
144 | |
145 triO : {n : Level } → {lx ly : Nat} → Trichotomous _o≈_ ( _o<_ {n} {lx} {ly} ) | |
146 triO {n} {lv} Φ y = {!!} | |
147 triO {n} {lv} (T-suc x) y = {!!} | |
148 triO {n} {.(Suc lv)} (ℵ lv) y = {!!} | |
149 | |
150 | |
151 max = Data.Nat._⊔_ | |
10 | 152 |
11 | 153 maxα : {n : Level } → { lx ly : Nat } → Ordinal {n} lx → Ordinal {n} ly → Ordinal {n} (max lx ly) |
154 maxα x y with x o< y | |
155 ... | t = {!!} | |
156 | |
157 -- X' = { x ∈ X | ψ x } ∪ X , Mα = ( ∪ (β < α) Mβ ) ' | |
158 | |
159 data Constructible {n : Level } {lv : Nat} ( α : Ordinal {n} lv ) : Set (suc n) where | |
160 fsub : ( ψ : Ordinal {n} lv → Set n ) → Constructible α | |
161 xself : Ordinal {n} lv → Constructible α | |
10 | 162 |
163 record ConstructibleSet {n : Level } : Set (suc n) where | |
164 field | |
165 level : Nat | |
11 | 166 α : Ordinal {n} level |
10 | 167 constructible : Constructible α |
168 | |
169 open ConstructibleSet | |
170 | |
11 | 171 data _c∋_ {n : Level } : {lv lv' : Nat} {α : Ordinal {n} lv } {α' : Ordinal {n} lv' } → |
10 | 172 Constructible {n} {lv} α → Constructible {n} {lv'} α' → Set n where |
11 | 173 c> : {lv lv' : Nat} {α : Ordinal {n} lv } {α' : Ordinal {n} lv' } |
174 (ta : Constructible {n} {lv} α ) ( tx : Constructible {n} {lv'} α' ) → α' o< α → ta c∋ tx | |
175 xself-fsub : {lv : Nat} {α : Ordinal {n} lv } | |
176 (ta : Ordinal {n} lv ) ( ψ : Ordinal {n} lv → Set n ) → _c∋_ {n} {_} {_} {α} {α} (xself ta ) ( fsub ψ) | |
177 fsub-fsub : {lv lv' : Nat} {α : Ordinal {n} lv } | |
178 ( ψ : Ordinal {n} lv → Set n ) ( ψ₁ : Ordinal {n} lv → Set n ) → | |
179 ( ∀ ( x : Ordinal {n} lv ) → ψ x → ψ₁ x ) → _c∋_ {n} {_} {_} {α} {α} ( fsub ψ ) ( fsub ψ₁) | |
7 | 180 |
11 | 181 _∋_ : {n : Level} → (ConstructibleSet {n}) → (ConstructibleSet {n} ) → Set n |
182 a ∋ x = constructible a c∋ constructible x | |
183 | |
184 data _c≈_ {n : Level } : {lv lv' : Nat} {α : Ordinal {n} lv } {α' : Ordinal {n} lv' } → | |
185 Constructible {n} {lv} α → Constructible {n} {lv'} α' → Set n where | |
186 crefl : {lv : Nat} {α : Ordinal {n} lv } → _c≈_ {n} {_} {_} {α} {α} (xself α ) (xself α ) | |
187 feq : {lv : Nat} {α : Ordinal {n} lv } | |
188 → ( ψ : Ordinal {n} lv → Set n ) ( ψ₁ : Ordinal {n} lv → Set n ) | |
189 → (∀ ( x : Ordinal {n} lv ) → ψ x ⇔ ψ₁ x ) → _c≈_ {n} {_} {_} {α} {α} ( fsub ψ ) ( fsub ψ₁) | |
190 | |
191 _≈_ : {n : Level} → (ConstructibleSet {n}) → (ConstructibleSet {n} ) → Set n | |
192 a ≈ x = constructible a c≈ constructible x | |
10 | 193 |
194 | |
11 | 195 ConstructibleSet→ZF : {n : Level } → ZF {suc n} {n} |
196 ConstructibleSet→ZF {n} = record { | |
10 | 197 ZFSet = ConstructibleSet |
198 ; _∋_ = _∋_ | |
11 | 199 ; _≈_ = _≈_ |
10 | 200 ; ∅ = record { level = Zero ; α = Φ ; constructible = xself Φ } |
201 ; _×_ = {!!} | |
202 ; Union = {!!} | |
203 ; Power = {!!} | |
204 ; Select = {!!} | |
205 ; Replace = {!!} | |
206 ; infinite = {!!} | |
207 ; isZF = {!!} | |
208 } where |