431
|
1 module zf where
|
|
2
|
|
3 open import Level
|
|
4
|
|
5 open import logic
|
|
6
|
|
7 open import Relation.Nullary
|
|
8 open import Relation.Binary hiding (_⇔_)
|
|
9 open import Data.Empty
|
|
10
|
|
11 record IsZF {n m : Level }
|
|
12 (ZFSet : Set n)
|
|
13 (_∋_ : ( A x : ZFSet ) → Set m)
|
|
14 (_≈_ : Rel ZFSet m)
|
|
15 (∅ : ZFSet)
|
|
16 (_,_ : ( A B : ZFSet ) → ZFSet)
|
|
17 (Union : ( A : ZFSet ) → ZFSet)
|
|
18 (Power : ( A : ZFSet ) → ZFSet)
|
|
19 (Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet )
|
|
20 (Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet )
|
|
21 (infinite : ZFSet)
|
|
22 : Set (suc (n ⊔ suc m)) where
|
|
23 field
|
|
24 isEquivalence : IsEquivalence {n} {m} {ZFSet} _≈_
|
|
25 -- ∀ x ∀ y ∃ z ∀ t ( z ∋ t → t ≈ x ∨ t ≈ y)
|
|
26 pair→ : ( x y t : ZFSet ) → (x , y) ∋ t → ( t ≈ x ) ∨ ( t ≈ y )
|
|
27 pair← : ( x y t : ZFSet ) → ( t ≈ x ) ∨ ( t ≈ y ) → (x , y) ∋ t
|
|
28 -- ∀ x ∃ y ∀ z (z ∈ y ⇔ ∃ u ∈ x ∧ (z ∈ u))
|
|
29 union→ : ( X z u : ZFSet ) → ( X ∋ u ) ∧ (u ∋ z ) → Union X ∋ z
|
|
30 union← : ( X z : ZFSet ) → (X∋z : Union X ∋ z ) → ¬ ( (u : ZFSet ) → ¬ ((X ∋ u) ∧ (u ∋ z )))
|
|
31 _∈_ : ( A B : ZFSet ) → Set m
|
|
32 A ∈ B = B ∋ A
|
|
33 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set m
|
|
34 _⊆_ A B {x} = A ∋ x → B ∋ x
|
|
35 _∩_ : ( A B : ZFSet ) → ZFSet
|
|
36 A ∩ B = Select A ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) )
|
|
37 _∪_ : ( A B : ZFSet ) → ZFSet
|
|
38 A ∪ B = Union (A , B)
|
|
39 {_} : ZFSet → ZFSet
|
|
40 { x } = ( x , x )
|
|
41 infixr 200 _∈_
|
|
42 infixr 230 _∩_ _∪_
|
|
43 infixr 220 _⊆_
|
|
44 field
|
|
45 empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x )
|
|
46 -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ t ⊆ X ) )
|
1095
|
47 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → t ∋ x → A ∋ x -- _⊆_ t A {x}
|
431
|
48 power← : ∀( A t : ZFSet ) → ( ∀ {x} → _⊆_ t A {x}) → Power A ∋ t
|
|
49 -- extensionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w )
|
|
50 extensionality : { A B w : ZFSet } → ( (z : ZFSet) → ( A ∋ z ) ⇔ (B ∋ z) ) → ( A ∈ w ⇔ B ∈ w )
|
|
51 -- regularity without minimum
|
437
|
52 ε-induction : { ψ : ZFSet → Set (suc m Level.⊔ n)}
|
431
|
53 → ( {x : ZFSet } → ({ y : ZFSet } → x ∋ y → ψ y ) → ψ x )
|
|
54 → (x : ZFSet ) → ψ x
|
|
55 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) )
|
|
56 infinity∅ : ∅ ∈ infinite
|
|
57 infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ { x }) ∈ infinite
|
|
58 selection : { ψ : ZFSet → Set m } → ∀ { X y : ZFSet } → ( ( y ∈ X ) ∧ ψ y ) ⇔ (y ∈ Select X ψ )
|
|
59 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) )
|
|
60 replacement← : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → x ∈ X → ψ x ∈ Replace X ψ
|
|
61 replacement→ : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet ) → ( lt : x ∈ Replace X ψ ) → ¬ ( ∀ (y : ZFSet) → ¬ ( x ≈ ψ y ) )
|
|
62
|
|
63 record ZF {n m : Level } : Set (suc (n ⊔ suc m)) where
|
|
64 infixr 210 _,_
|
|
65 infixl 200 _∋_
|
|
66 infixr 220 _≈_
|
|
67 field
|
|
68 ZFSet : Set n
|
|
69 _∋_ : ( A x : ZFSet ) → Set m
|
|
70 _≈_ : ( A B : ZFSet ) → Set m
|
|
71 -- ZF Set constructor
|
|
72 ∅ : ZFSet
|
|
73 _,_ : ( A B : ZFSet ) → ZFSet
|
|
74 Union : ( A : ZFSet ) → ZFSet
|
|
75 Power : ( A : ZFSet ) → ZFSet
|
|
76 Select : (X : ZFSet ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet
|
|
77 Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet
|
|
78 infinite : ZFSet
|
|
79 isZF : IsZF ZFSet _∋_ _≈_ ∅ _,_ Union Power Select Replace infinite
|
|
80
|