comparison zf.agda @ 23:7293a151d949

Sup
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 18 May 2019 08:29:08 +0900
parents 627a79e61116
children fce60b99dc55
comparison
equal deleted inserted replaced
22:3da2c00bd24d 23:7293a151d949
1 module zf where 1 module zf where
2 2
3 open import Level 3 open import Level
4 4
5 data Bool : Set where
6 true : Bool
7 false : Bool
5 8
6 record _∧_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where 9 record _∧_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where
7 field 10 field
8 proj1 : A 11 proj1 : A
9 proj2 : B 12 proj2 : B
58 -- ∀ X ∃ A∀ t(t ∈ A ⇔ ∃ x ∈ X(t ∈ x)) 61 -- ∀ X ∃ A∀ t(t ∈ A ⇔ ∃ x ∈ X(t ∈ x))
59 union→ : ( X x y : ZFSet ) → X ∋ x → x ∋ y → Union X ∋ y 62 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 63 union← : ( X x y : ZFSet ) → Union X ∋ y → X ∋ x → x ∋ y
61 _∈_ : ( A B : ZFSet ) → Set m 64 _∈_ : ( A B : ZFSet ) → Set m
62 A ∈ B = B ∋ A 65 A ∈ B = B ∋ A
63 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → ∀{ A∋x : Set m } → Set m 66 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → Set m
64 _⊆_ A B {x} {A∋x} = B ∋ x 67 _⊆_ A B {x} = A ∋ x → B ∋ x
65 _∩_ : ( A B : ZFSet ) → ZFSet 68 _∩_ : ( A B : ZFSet ) → ZFSet
66 A ∩ B = Select (A , B) ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) 69 A ∩ B = Select (A , B) ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) )
67 _∪_ : ( A B : ZFSet ) → ZFSet 70 _∪_ : ( A B : ZFSet ) → ZFSet
68 A ∪ B = Select (A , B) ( λ x → ( A ∋ x ) ∨ ( B ∋ x ) ) 71 A ∪ B = Select (A , B) ( λ x → ( A ∋ x ) ∨ ( B ∋ x ) )
69 infixr 200 _∈_ 72 infixr 200 _∈_
70 infixr 230 _∩_ _∪_ 73 infixr 230 _∩_ _∪_
71 infixr 220 _⊆_ 74 infixr 220 _⊆_
72 field 75 field
73 empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x ) 76 empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x )
74 -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ t ⊆ X ) ) 77 -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ t ⊆ X ) )
75 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} {y} → _⊆_ t A {x} {y} 78 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → _⊆_ t A {x}
76 power← : ∀( A t : ZFSet ) → ∀ {x} {y} → _⊆_ t A {x} {y} → Power A ∋ t 79 power← : ∀( A t : ZFSet ) → ∀ {x} → _⊆_ t A {x} → Power A ∋ t
77 -- extentionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w ) 80 -- extentionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w )
78 extentionality : ( A B z : ZFSet ) → (( A ∋ z ) ⇔ (B ∋ z) ) → A ≈ B 81 extentionality : ( A B z : ZFSet ) → (( A ∋ z ) ⇔ (B ∋ z) ) → A ≈ B
79 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) ) 82 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) )
80 minimul : ZFSet → ZFSet 83 minimul : ZFSet → ZFSet
81 regularity : ∀( x : ZFSet ) → ¬ (x ≈ ∅) → ( minimul x ∈ x ∧ ( minimul x ∩ x ≈ ∅ ) ) 84 regularity : ∀( x : ZFSet ) → ¬ (x ≈ ∅) → ( minimul x ∈ x ∧ ( minimul x ∩ x ≈ ∅ ) )