diff zf.agda @ 376:6c72bee25653

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Mon, 20 Jul 2020 16:28:12 +0900
parents 8cade5f660bd
children cc7909f86841
line wrap: on
line diff
--- a/zf.agda	Mon Jul 20 16:22:44 2020 +0900
+++ b/zf.agda	Mon Jul 20 16:28:12 2020 +0900
@@ -16,8 +16,8 @@
      (_,_ : ( A B : ZFSet  ) → ZFSet)
      (Union : ( A : ZFSet  ) → ZFSet)
      (Power : ( A : ZFSet  ) → ZFSet)
-     (Select :  (X : ZFSet  ) → ( ψ : (x : ZFSet ) → (X ∋ x) → Set m ) → ZFSet ) 
-     (Replace : (X : ZFSet) → ( (x : ZFSet ) → ( X ∋ x ) → ZFSet ) → ZFSet )
+     (Select :  (X : ZFSet  ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet ) 
+     (Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet )
      (infinite : ZFSet)
        : Set (suc (n ⊔ suc m)) where
   field
@@ -33,7 +33,7 @@
   _⊆_ : ( A B : ZFSet  ) → ∀{ x : ZFSet } →  Set m
   _⊆_ A B {x} = A ∋ x →  B ∋ x
   _∩_ : ( A B : ZFSet  ) → ZFSet
-  A ∩ B = Select A (  λ x _ → ( A ∋ x ) ∧ ( B ∋ x )  )
+  A ∩ B = Select A (  λ x → ( A ∋ x ) ∧ ( B ∋ x )  )
   _∪_ : ( A B : ZFSet  ) → ZFSet
   A ∪ B = Union (A , B)    
   {_} : ZFSet → ZFSet
@@ -55,10 +55,10 @@
      -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) )
      infinity∅ :  ∅ ∈ infinite
      infinity :  ∀( x : ZFSet  ) → x ∈ infinite →  ( x ∪ { x }) ∈ infinite 
-     selection : ∀ { X y : ZFSet  } →  { ψ : (x : ZFSet ) → X ∋ x → Set m } → (y ∈ X ∧ (( y∈X : y ∈ X ) →  ψ y y∈X)) ⇔ (y ∈ Select X ψ ) 
+     selection : { ψ : ZFSet → Set m } → ∀ { X y : ZFSet  } →  ( ( y ∈ X ) ∧ ψ y ) ⇔ (y ∈  Select X ψ ) 
      -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) )
-     replacement← : ∀ ( X x : ZFSet  ) → (x∈X : x ∈ X ) → {ψ : (x : ZFSet ) → x ∈ X  → ZFSet} → ψ x x∈X ∈  Replace X ψ 
-     replacement→ : ∀ ( X x : ZFSet  ) →  {ψ : (x : ZFSet ) → X ∋ x  → ZFSet} → ( lt : x ∈  Replace X ψ ) → ¬ ( ∀ (y : ZFSet)  → (X∋y : X ∋ y ) →  ¬ ( x ≈ ψ y X∋y  ) )
+     replacement← : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet  ) → x ∈ X → ψ x ∈  Replace X ψ 
+     replacement→ : {ψ : ZFSet → ZFSet} → ∀ ( X x : ZFSet  ) →  ( lt : x ∈  Replace X ψ ) → ¬ ( ∀ (y : ZFSet)  →  ¬ ( x ≈ ψ y ) )
 
 record ZF {n m : Level } : Set (suc (n ⊔ suc m)) where
   infixr  210 _,_
@@ -73,8 +73,8 @@
      _,_ : ( A B : ZFSet  ) → ZFSet
      Union : ( A : ZFSet  ) → ZFSet
      Power : ( A : ZFSet  ) → ZFSet
-     Select :  (X : ZFSet  ) → ( ψ : (x : ZFSet ) → ( X ∋ x ) → Set m ) → ZFSet 
-     Replace : (X : ZFSet ) → ( (x : ZFSet ) → ( X ∋ x ) → ZFSet ) → ZFSet
+     Select :  (X : ZFSet  ) → ( ψ : (x : ZFSet ) → Set m ) → ZFSet 
+     Replace : ZFSet → ( ZFSet → ZFSet ) → ZFSet
      infinite : ZFSet
      isZF : IsZF ZFSet _∋_ _≈_ ∅ _,_ Union Power Select Replace infinite