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 ) infixr 130 _∧_ infixr 140 _∨_ infixr 150 _⇔_ open import Data.Empty open import Relation.Nullary record ZF (n m : Level ) : Set (suc (n ⊔ m)) where coinductive 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 infixl 200 _∋_ infixr 210 _×_ infixr 220 _≈_ field -- ∀ 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 X t : ZFSet ) → A ∋ t → ∀ {x} {y} → _⊆_ t X {x} {y} power← : ( A X t : ZFSet ) → ∀ {x} {y} → _⊆_ t X {x} {y} → 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 ) ) infinite : ZFSet 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