diff zf.agda @ 18:627a79e61116

fix
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Thu, 16 May 2019 10:55:34 +0900
parents e11e95d5ddee
children 7293a151d949
line wrap: on
line diff
--- a/zf.agda	Tue May 14 13:52:19 2019 +0900
+++ b/zf.agda	Thu May 16 10:55:34 2019 +0900
@@ -44,17 +44,17 @@
      (_∋_ : ( A x : ZFSet  ) → Set m)
      (_≈_ : Rel ZFSet m)
      (∅  : ZFSet)
-     (_×_ : ( A B : ZFSet  ) → ZFSet)
+     (_,_ : ( A B : ZFSet  ) → ZFSet)
      (Union : ( A : ZFSet  ) → ZFSet)
      (Power : ( A : ZFSet  ) → ZFSet)
-     (Select : ( ZFSet → Set m ) → ZFSet )
-     (Replace : ( ZFSet → ZFSet ) → ZFSet )
+     (Select : ZFSet → ( ZFSet → Set m ) → ZFSet )
+     (Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet )
      (infinite : ZFSet)
        : Set (suc (n ⊔ m)) where
   field
      isEquivalence : {A B : ZFSet} → IsEquivalence {n} {m} {ZFSet} _≈_ 
      -- ∀ x ∀ y ∃ z(x ∈ z ∧ y ∈ z)
-     pair : ( A B : ZFSet  ) →  ( (A × B)  ∋ A ) ∧ ( (A × B)  ∋ B  )
+     pair : ( A B : ZFSet  ) →  ( (A , B)  ∋ A ) ∧ ( (A , B)  ∋ B  )
      -- ∀ X ∃ A∀ t(t ∈ A ⇔ ∃ x ∈ X(t  ∈ x))
      union→ : ( X x y : ZFSet  ) → X ∋ x  → x ∋ y → Union X  ∋ y
      union← : ( X x y : ZFSet  ) → Union X  ∋ y → X ∋ x  → x ∋ y 
@@ -63,9 +63,9 @@
   _⊆_ : ( A B : ZFSet  ) → ∀{ x : ZFSet } → ∀{ A∋x : Set m } → Set m
   _⊆_ A B {x} {A∋x} = B ∋ x
   _∩_ : ( A B : ZFSet  ) → ZFSet
-  A ∩ B = Select (  λ x → ( A ∋ x ) ∧ ( B ∋ x )  )
+  A ∩ B = Select (A , B) (  λ x → ( A ∋ x ) ∧ ( B ∋ x )  )
   _∪_ : ( A B : ZFSet  ) → ZFSet
-  A ∪ B = Select (  λ x → ( A ∋ x ) ∨ ( B ∋ x )  )
+  A ∪ B = Select (A , B) (  λ x → ( A ∋ x ) ∨ ( B ∋ x )  )
   infixr  200 _∈_
   infixr  230 _∩_ _∪_
   infixr  220 _⊆_ 
@@ -81,13 +81,13 @@
      regularity : ∀( x : ZFSet  ) → ¬ (x ≈ ∅) → (  minimul x ∈ x ∧  ( minimul x  ∩ x  ≈ ∅ ) )
      -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) )
      infinity∅ :  ∅ ∈ infinite
-     infinity :  ∀( x : ZFSet  ) → x ∈ infinite →  ( x ∪ Select (  λ y → x ≈ y )) ∈ infinite 
-     selection : { ψ : ZFSet → Set m } → ∀ ( y : ZFSet  ) →  ( y  ∈  Select ψ )  → ψ y
+     infinity :  ∀( X x : ZFSet  ) → x ∈ infinite →  ( x ∪ Select X (  λ y → x ≈ y )) ∈ infinite 
+     selection : { ψ : ZFSet → Set m } → ∀ ( X y : ZFSet  ) →  ( y  ∈  Select X ψ )  → ψ y
      -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) )
-     replacement : {ψ : ZFSet → ZFSet} → ∀ ( x : ZFSet  ) →  ( ψ x ∈  Replace ψ )  
+     replacement : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet  ) →  ( ψ x ∈  Replace X ψ )  
 
 record ZF {n m : Level } : Set (suc (n ⊔ m)) where
-  infixr  210 _×_
+  infixr  210 _,_
   infixl  200 _∋_ 
   infixr  220 _≈_
   field
@@ -96,13 +96,13 @@
      _≈_ : ( A B : ZFSet  ) → Set m
   -- ZF Set constructor
      ∅  : ZFSet
-     _×_ : ( A B : ZFSet  ) → ZFSet
+     _,_ : ( A B : ZFSet  ) → ZFSet
      Union : ( A : ZFSet  ) → ZFSet
      Power : ( A : ZFSet  ) → ZFSet
-     Select : ( ZFSet → Set m ) → ZFSet
-     Replace : ( ZFSet → ZFSet ) → ZFSet
+     Select : ZFSet → ( ZFSet → Set m ) → ZFSet
+     Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet
      infinite : ZFSet
-     isZF : IsZF ZFSet _∋_ _≈_ ∅ _×_ Union Power Select Replace infinite 
+     isZF : IsZF ZFSet _∋_ _≈_ ∅ _,_ Union Power Select Replace infinite 
 
 module zf-exapmle {n m : Level } ( zf : ZF {m} {n} ) where