view zf.agda @ 6:d9b704508281

isEquiv and isZF
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sat, 11 May 2019 11:40:31 +0900
parents c12d964a04c0
children 813f1b3b000b
line wrap: on
line source

module zf where

open import Level


record  _∧_  {n m : Level} (A  : Set n) ( B : Set m ) : Set (n ⊔ m) where
   field 
      proj1 : A
      proj2 : B

open _∧_


data  _∨_  {n m : Level} (A  : Set n) ( B : Set m ) : Set (n ⊔ m) where
   case1 : A → A ∨ B
   case2 : B → A ∨ B

-- open import Relation.Binary.PropositionalEquality 

_⇔_ : {n : Level } → ( A B : Set n )  → Set  n
_⇔_ A B =  ( A → B ) ∧ ( B  → A )

open import Data.Empty
open import Relation.Nullary

open import Relation.Binary
open import Relation.Binary.Core

infixr  130 _∧_
infixr  140 _∨_
infixr  150 _⇔_

record IsZF {n m : Level }
     (ZFSet : Set n)
     (_∋_ : ( A x : ZFSet  ) → Set m)
     (_≈_ : ( A B : ZFSet  ) → Set m)
     (∅  : ZFSet)
     (_×_ : ( A B : ZFSet  ) → ZFSet)
     (Union : ( A : ZFSet  ) → ZFSet)
     (Power : ( A : ZFSet  ) → ZFSet)
     (Restrict : ( ZFSet → Set m ) → 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  )
     -- ∀ 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 
  _∈_ : ( 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  ) → ZFSet
  A ∩ B = Restrict ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) )
  _∪_ : ( A B : ZFSet  ) → ZFSet
  A ∪ B = Restrict ( λ x → ( A ∋ x ) ∨ ( B ∋ x ) )
  infixr  200 _∈_
  infixr  230 _∩_ _∪_
  infixr  220 _⊆_ 
  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 
     -- 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 = ∅ ) )
     -- smaller : ZFSet → ZFSet
     -- regularity : ( x : ZFSet  ) → ¬ (x ≈ ∅) → smaller x  ∩ x  ≈ ∅ 
     -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) )
     infinity∅ :  ∅ ∈ infinite
     infinity :  ( x : ZFSet  ) → x ∈ infinite →  ( x ∪ Restrict ( λ y → x ≈ y )) ∈ infinite 
     -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) )
     replacement : ( ψ :  ZFSet → Set m ) → ( y : ZFSet  ) →  y  ∈  Restrict ψ  → ψ y

record ZF {n m : Level } : Set (suc (n ⊔ m)) where
  coinductive
  infixr  210 _×_
  infixl  200 _∋_ 
  infixr  220 _≈_
  field
     ZFSet : Set n
     _∋_ : ( A x : ZFSet  ) → Set m 
     _≈_ : ( A B : ZFSet  ) → Set m
  -- ZF Set constructor
     ∅  : ZFSet
     _×_ : ( A B : ZFSet  ) → ZFSet
     Union : ( A : ZFSet  ) → ZFSet
     Power : ( A : ZFSet  ) → ZFSet
     Restrict : ( ZFSet → Set m ) → ZFSet
     infinite : ZFSet
  field
     isZF : IsZF ZFSet _∋_ _≈_ ∅ _×_ Union Power Restrict infinite