Mercurial > hg > Members > kono > Proof > ZF-in-agda
annotate zf.agda @ 224:afc864169325
recover ε-induction
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sat, 10 Aug 2019 12:31:25 +0900 |
parents | 2e1f19c949dc |
children | 650bdad56729 |
rev | line source |
---|---|
3 | 1 module zf where |
2 | |
3 open import Level | |
4 | |
213
22d435172d1a
separate logic and nat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
209
diff
changeset
|
5 open import logic |
123 | 6 |
6 | 7 open import Relation.Nullary |
8 open import Relation.Binary | |
188
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
186
diff
changeset
|
9 open import Data.Empty |
1f2c8b094908
axiom of choice → p ∨ ¬ p
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
186
diff
changeset
|
10 |
6 | 11 record IsZF {n m : Level } |
12 (ZFSet : Set n) | |
13 (_∋_ : ( 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
|
14 (_≈_ : Rel ZFSet m) |
6 | 15 (∅ : ZFSet) |
18 | 16 (_,_ : ( A B : ZFSet ) → ZFSet) |
6 | 17 (Union : ( A : ZFSet ) → ZFSet) |
18 (Power : ( A : ZFSet ) → ZFSet) | |
115 | 19 (Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet ) |
18 | 20 (Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet ) |
6 | 21 (infinite : ZFSet) |
22 : Set (suc (n ⊔ m)) where | |
3 | 23 field |
29
fce60b99dc55
posturate OD is isomorphic to Ordinal
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
23
diff
changeset
|
24 isEquivalence : IsEquivalence {n} {m} {ZFSet} _≈_ |
3 | 25 -- ∀ x ∀ y ∃ z(x ∈ z ∧ y ∈ z) |
18 | 26 pair : ( A B : ZFSet ) → ( (A , B) ∋ A ) ∧ ( (A , B) ∋ B ) |
69 | 27 -- ∀ x ∃ y ∀ z (z ∈ y ⇔ ∃ u ∈ x ∧ (z ∈ u)) |
73 | 28 union→ : ( X z u : ZFSet ) → ( X ∋ u ) ∧ (u ∋ z ) → Union X ∋ z |
159 | 29 union← : ( X z : ZFSet ) → (X∋z : Union X ∋ z ) → ¬ ( (u : ZFSet ) → ¬ ((X ∋ u) ∧ (u ∋ z ))) |
3 | 30 _∈_ : ( A B : ZFSet ) → Set m |
31 A ∈ B = B ∋ A | |
23 | 32 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set m |
33 _⊆_ A B {x} = A ∋ x → B ∋ x | |
3 | 34 _∩_ : ( A B : ZFSet ) → ZFSet |
115 | 35 A ∩ B = Select A ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) |
3 | 36 _∪_ : ( A B : ZFSet ) → ZFSet |
103 | 37 A ∪ B = Union (A , B) -- Select A ( λ x → ( A ∋ x ) ∨ ( B ∋ x ) ) is easer |
78
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
38 {_} : ZFSet → ZFSet |
9a7a64b2388c
infinite and replacement begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
77
diff
changeset
|
39 { x } = ( x , x ) |
3 | 40 infixr 200 _∈_ |
41 infixr 230 _∩_ _∪_ | |
42 infixr 220 _⊆_ | |
43 field | |
4 | 44 empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x ) |
3 | 45 -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ t ⊆ X ) ) |
166 | 46 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → t ∋ x → ¬ ¬ ( A ∋ x ) -- _⊆_ t A {x} |
77 | 47 power← : ∀( A t : ZFSet ) → ( ∀ {x} → _⊆_ t A {x}) → Power A ∋ t |
65
164ad5a703d8
¬∅=→∅∈ : {n : Level} → { x : OD {suc n} } → ¬ ( x == od∅ {suc n} ) → x ∋ od∅ {suc n}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
54
diff
changeset
|
48 -- extensionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w ) |
186 | 49 extensionality : { A B w : ZFSet } → ( (z : ZFSet) → ( A ∋ z ) ⇔ (B ∋ z) ) → ( A ∈ w ⇔ B ∈ w ) |
183 | 50 -- This form of regurality forces choice function |
3 | 51 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) ) |
183 | 52 -- minimul : (x : ZFSet ) → ¬ (x ≈ ∅) → ZFSet |
53 -- regularity : ∀( x : ZFSet ) → (not : ¬ (x ≈ ∅)) → ( minimul x not ∈ x ∧ ( minimul x not ∩ x ≈ ∅ ) ) | |
54 -- another form of regularity | |
223
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
213
diff
changeset
|
55 -- ε-induction : { ψ : ZFSet → Set m} |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
213
diff
changeset
|
56 -- → ( {x : ZFSet } → ({ y : ZFSet } → x ∋ y → ψ y ) → ψ x ) |
2e1f19c949dc
sepration of ordinal from OD
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
213
diff
changeset
|
57 -- → (x : ZFSet ) → ψ x |
3 | 58 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) ) |
59 infinity∅ : ∅ ∈ infinite | |
160 | 60 infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ { x }) ∈ infinite |
140
312e27aa3cb5
remove otrans again. start over
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
138
diff
changeset
|
61 selection : { ψ : ZFSet → Set m } → ∀ { X y : ZFSet } → ( ( y ∈ X ) ∧ ψ y ) ⇔ (y ∈ Select X ψ ) |
3 | 62 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) ) |
130 | 63 replacement← : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → x ∈ X → ψ x ∈ Replace X ψ |
138 | 64 replacement→ : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → ( lt : x ∈ Replace X ψ ) → ¬ ( ∀ (y : ZFSet) → ¬ ( x ≈ ψ y ) ) |
183 | 65 -- ∀ X [ ∅ ∉ X → (∃ f : X → ⋃ X ) → ∀ A ∈ X ( f ( A ) ∈ A ) ] |
66 choice-func : (X : ZFSet ) → {x : ZFSet } → ¬ ( x ≈ ∅ ) → ( X ∋ x ) → ZFSet | |
67 choice : (X : ZFSet ) → {A : ZFSet } → ( X∋A : X ∋ A ) → (not : ¬ ( A ≈ ∅ )) → A ∋ choice-func X not X∋A | |
3 | 68 |
6 | 69 record ZF {n m : Level } : Set (suc (n ⊔ m)) where |
18 | 70 infixr 210 _,_ |
6 | 71 infixl 200 _∋_ |
72 infixr 220 _≈_ | |
73 field | |
74 ZFSet : Set n | |
75 _∋_ : ( A x : ZFSet ) → Set m | |
76 _≈_ : ( A B : ZFSet ) → Set m | |
77 -- ZF Set constructor | |
78 ∅ : ZFSet | |
18 | 79 _,_ : ( A B : ZFSet ) → ZFSet |
6 | 80 Union : ( A : ZFSet ) → ZFSet |
81 Power : ( A : ZFSet ) → ZFSet | |
115 | 82 Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet |
18 | 83 Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet |
6 | 84 infinite : ZFSet |
18 | 85 isZF : IsZF ZFSet _∋_ _≈_ ∅ _,_ Union Power Select Replace infinite |
6 | 86 |