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
|
6
|
33 record IsZF {n m : Level }
|
|
34 (ZFSet : Set n)
|
|
35 (_∋_ : ( A x : ZFSet ) → Set m)
|
|
36 (_≈_ : ( A B : ZFSet ) → Set m)
|
|
37 (∅ : ZFSet)
|
|
38 (_×_ : ( A B : ZFSet ) → ZFSet)
|
|
39 (Union : ( A : ZFSet ) → ZFSet)
|
|
40 (Power : ( A : ZFSet ) → ZFSet)
|
|
41 (Restrict : ( ZFSet → Set m ) → ZFSet)
|
|
42 (infinite : ZFSet)
|
|
43 : Set (suc (n ⊔ m)) where
|
3
|
44 field
|
6
|
45 isEquivalence : {A B : ZFSet} → IsEquivalence {n} {m} {ZFSet} _≈_
|
3
|
46 -- ∀ x ∀ y ∃ z(x ∈ z ∧ y ∈ z)
|
6
|
47 pair : ( A B : ZFSet ) → ( (A × B) ∋ A ) ∧ ( (A × B) ∋ B )
|
3
|
48 -- ∀ X ∃ A∀ t(t ∈ A ⇔ ∃ x ∈ X(t ∈ x))
|
|
49 union→ : ( X x y : ZFSet ) → X ∋ x → x ∋ y → Union X ∋ y
|
|
50 union← : ( X x y : ZFSet ) → Union X ∋ y → X ∋ x → x ∋ y
|
|
51 _∈_ : ( A B : ZFSet ) → Set m
|
|
52 A ∈ B = B ∋ A
|
7
|
53 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → ∀{ A∋x : Set m } → Set m
|
3
|
54 _⊆_ A B {x} {A∋x} = B ∋ x
|
|
55 _∩_ : ( A B : ZFSet ) → ZFSet
|
|
56 A ∩ B = Restrict ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) )
|
|
57 _∪_ : ( A B : ZFSet ) → ZFSet
|
|
58 A ∪ B = Restrict ( λ x → ( A ∋ x ) ∨ ( B ∋ x ) )
|
|
59 infixr 200 _∈_
|
|
60 infixr 230 _∩_ _∪_
|
|
61 infixr 220 _⊆_
|
|
62 field
|
4
|
63 empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x )
|
3
|
64 -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ t ⊆ X ) )
|
8
|
65 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} {y} → _⊆_ t A {x} {y}
|
|
66 power← : ∀( A t : ZFSet ) → ∀ {x} {y} → _⊆_ t A {x} {y} → Power A ∋ t
|
3
|
67 -- extentionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w )
|
6
|
68 extentionality : ( A B z : ZFSet ) → (( A ∋ z ) ⇔ (B ∋ z) ) → A ≈ B
|
3
|
69 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) )
|
7
|
70 smaller : ZFSet → ZFSet
|
8
|
71 regularity : ∀( x : ZFSet ) → ¬ (x ≈ ∅) → ( smaller x ∈ x ∧ ( smaller x ∩ x ≈ ∅ ) )
|
3
|
72 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) )
|
|
73 infinity∅ : ∅ ∈ infinite
|
8
|
74 infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ Restrict ( λ y → x ≈ y )) ∈ infinite
|
3
|
75 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) )
|
8
|
76 -- this form looks like specification
|
|
77 replacement : ( ψ : ZFSet → Set m ) → ( x : ZFSet ) → x ∈ Restrict ψ → ψ x
|
3
|
78
|
6
|
79 record ZF {n m : Level } : Set (suc (n ⊔ m)) where
|
|
80 infixr 210 _×_
|
|
81 infixl 200 _∋_
|
|
82 infixr 220 _≈_
|
|
83 field
|
|
84 ZFSet : Set n
|
|
85 _∋_ : ( A x : ZFSet ) → Set m
|
|
86 _≈_ : ( A B : ZFSet ) → Set m
|
|
87 -- ZF Set constructor
|
|
88 ∅ : ZFSet
|
|
89 _×_ : ( A B : ZFSet ) → ZFSet
|
|
90 Union : ( A : ZFSet ) → ZFSet
|
|
91 Power : ( A : ZFSet ) → ZFSet
|
|
92 Restrict : ( ZFSet → Set m ) → ZFSet
|
|
93 infinite : ZFSet
|
|
94 isZF : IsZF ZFSet _∋_ _≈_ ∅ _×_ Union Power Restrict infinite
|
|
95
|
7
|
96 module reguraliry-m {n m : Level } ( zf : ZF {m} {n} ) where
|
|
97
|
|
98 open import Relation.Nullary
|
|
99 open import Data.Empty
|
|
100 open import Relation.Binary.PropositionalEquality
|
|
101
|
|
102 _≈_ = ZF._≈_ zf
|
|
103 ZFSet = ZF.ZFSet
|
|
104 Restrict = ZF.Restrict zf
|
|
105 ∅ = ZF.∅ zf
|
|
106 _∩_ = ( IsZF._∩_ ) (ZF.isZF zf)
|
|
107 _∋_ = ZF._∋_ zf
|
|
108 replacement = IsZF.replacement ( ZF.isZF zf )
|
|
109
|
|
110 russel : ( x : ZFSet zf ) → x ≈ Restrict ( λ x → ¬ ( x ∋ x ) ) → ⊥
|
|
111 russel x eq with x ∋ x
|
|
112 ... | x∋x with replacement ( λ x → ¬ ( x ∋ x )) x {!!}
|
|
113 ... | r1 = {!!}
|
|
114
|
|
115
|
|
116
|
|
117
|
|
118 data Nat : Set zero where
|
|
119 Zero : Nat
|
|
120 Suc : Nat → Nat
|
|
121
|
|
122 prev : Nat → Nat
|
|
123 prev (Suc n) = n
|
|
124 prev Zero = Zero
|
|
125
|
|
126 open import Relation.Binary.PropositionalEquality
|
|
127
|
|
128
|
8
|
129 data Transtive {n : Level } : ( lv : Nat) → Set n where
|
|
130 Φ : {lv : Nat} → lv ≡ Zero → Transtive lv
|
|
131 T-suc : {lv : Nat} → lv ≡ Zero → Transtive {n} lv → Transtive lv
|
|
132 ℵ : {lv : Nat} → Transtive {n} lv → Transtive (Suc lv)
|
7
|
133
|
|
134
|
|
135
|
|
136
|