diff 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
line wrap: on
line diff
--- a/zf.agda	Thu May 16 17:20:45 2019 +0900
+++ b/zf.agda	Sat May 18 08:29:08 2019 +0900
@@ -2,6 +2,9 @@
 
 open import Level
 
+data Bool : Set where
+   true : Bool
+   false : Bool
 
 record  _∧_  {n m : Level} (A  : Set n) ( B : Set m ) : Set (n ⊔ m) where
    field 
@@ -60,8 +63,8 @@
      union← : ( X x y : ZFSet  ) → Union X  ∋ y → X ∋ x  → x ∋ y 
   _∈_ : ( A B : ZFSet  ) → Set m
   A ∈ B = B ∋ A
-  _⊆_ : ( A B : ZFSet  ) → ∀{ x : ZFSet } → ∀{ A∋x : Set m } → Set m
-  _⊆_ A B {x} {A∋x} = B ∋ x
+  _⊆_ : ( A B : ZFSet  ) → ∀{ x : ZFSet } →  Set m
+  _⊆_ A B {x} = A ∋ x →  B ∋ x
   _∩_ : ( A B : ZFSet  ) → ZFSet
   A ∩ B = Select (A , B) (  λ x → ( A ∋ x ) ∧ ( B ∋ x )  )
   _∪_ : ( A B : ZFSet  ) → ZFSet
@@ -72,8 +75,8 @@
   field
      empty :  ∀( x : ZFSet  ) → ¬ ( ∅ ∋ x )
      -- power : ∀ X ∃ A ∀ t ( t ∈ A ↔ t ⊆ X ) )
-     power→ : ∀( A t : ZFSet  ) → Power A ∋ t → ∀ {x} {y} →  _⊆_ t A {x} {y}
-     power← : ∀( A t : ZFSet  ) → ∀ {x} {y} →  _⊆_ t A {x} {y} → Power A ∋ t 
+     power→ : ∀( A t : ZFSet  ) → Power A ∋ t → ∀ {x}  →  _⊆_ t A {x} 
+     power← : ∀( A t : ZFSet  ) → ∀ {x}  →  _⊆_ t A {x}  → Power A ∋ t 
      -- extentionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w )
      extentionality :  ( A B z  : ZFSet  ) → (( A ∋ z ) ⇔ (B ∋ z) ) → A ≈ B
      -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) )