view zf.agda @ 11:2df90eb0896c

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Mon, 13 May 2019 20:51:45 +0900
parents 8022e14fce74
children b531d2b417ad e11e95d5ddee
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 Func {n m : Level } (ZFSet : Set n) (_≈_ : Rel ZFSet m) : Set (n ⊔ suc m) where
  field
     rel : Rel ZFSet m
     dom : ( y : ZFSet ) → ∀ { x : ZFSet } → rel x y
     fun-eq : { x y z : ZFSet } →  ( rel  x  y  ∧ rel  x  z  ) → y ≈ z 

open Func


record IsZF {n m : Level }
     (ZFSet : Set n)
     (_∋_ : ( A x : ZFSet  ) → Set m)
     (_≈_ : Rel ZFSet m)
     (∅  : ZFSet)
     (_×_ : ( A B : ZFSet  ) → ZFSet)
     (Union : ( A : ZFSet  ) → ZFSet)
     (Power : ( A : ZFSet  ) → ZFSet)
     (Select : ( ZFSet → Set m ) → ZFSet )
     (Replace : ( 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  )
     -- ∀ 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 = Select (  λ x → ( A ∋ x ) ∧ ( B ∋ x )  )
  _∪_ : ( A B : ZFSet  ) → ZFSet
  A ∪ B = Select (  λ 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 = ∅ ) )
     minimul : ZFSet → ZFSet
     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
     -- 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 ψ )  

record ZF {n m : Level } : Set (suc (n ⊔ m)) where
  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
     Select : ( ZFSet → Set m ) → ZFSet
     Replace : ( ZFSet → ZFSet ) → ZFSet
     infinite : ZFSet
     isZF : IsZF ZFSet _∋_ _≈_ ∅ _×_ Union Power Select Replace infinite 

module zf-exapmle {n m : Level } ( zf : ZF {m} {n} ) where

  _≈_ =  ZF._≈_ zf
  ZFSet =  ZF.ZFSet  zf
  Select =  ZF.Select  zf
  ∅ =  ZF.∅ zf
  _∩_ =  ( IsZF._∩_  ) (ZF.isZF zf)
  _∋_ =   ZF._∋_  zf
  replacement =   IsZF.replacement  ( ZF.isZF zf )
  selection =   IsZF.selection  ( ZF.isZF zf )
  minimul =   IsZF.minimul  ( ZF.isZF zf )
  regularity =   IsZF.regularity  ( ZF.isZF zf )

--  russel : Select ( λ x →  x ∋ x  ) ≈ ∅
--  russel with Select ( λ x →  x ∋ x  ) 
--  ... | s = {!!}

module constructible-set  where

  open import Data.Nat renaming ( zero to Zero ; suc to Suc ;  ℕ to Nat ) 
  
  open import  Relation.Binary.PropositionalEquality
  
  data Ordinal {n : Level }  :  Set n where
     Φ : {lv : Nat} → Ordinal {n} lv
     T-suc : {lv : Nat} → Ordinal {n} lv → Ordinal lv
     ℵ_ :  (lv : Nat) → Ordinal (Suc lv)

  data _o<_ {n : Level } :  Ordinal {n}  →  Ordinal {n}  → Set n where
     l< : {lx ly : Nat }  → {x : Ordinal {n} lx } →  {y : Ordinal {n} ly } → lx < ly → x o< y
     Φ<  : {lx : Nat} → {x : Ordinal {n} lx}  →  Φ {n} {lx} o< T-suc {n} {lx} x
     s<  : {lx : Nat} → {x : Ordinal {n} lx}  →  x o< T-suc {n} {lx} x
     ℵΦ< : {lx : Nat} → {x : Ordinal {n} lx } →  Φ {n} {lx} o< (ℵ lx) 
     ℵ<  : {lx : Nat} → {x : Ordinal {n} lx } →  T-suc {n} {lx} x o< (ℵ lx) 

  _o≈_ : {n : Level } {lv : Nat } → Rel ( Ordinal {n} lv ) n
  _o≈_  = {!!}

  triO : {n : Level } → {lx ly : Nat}   → Trichotomous  _o≈_ ( _o<_ {n} {lx} {ly} )
  triO {n} {lv} Φ y = {!!}
  triO {n} {lv} (T-suc x) y = {!!}
  triO {n} {.(Suc lv)} (ℵ lv) y = {!!}


  max = Data.Nat._⊔_
  
  maxα : {n : Level } → { lx ly : Nat } → Ordinal {n} lx  →  Ordinal {n} ly  → Ordinal {n} (max lx ly)
  maxα x y with x o< y
  ... | t = {!!}

  -- X' = { x ∈ X |  ψ  x } ∪ X , Mα = ( ∪ (β < α) Mβ ) '

  data Constructible {n : Level } {lv : Nat} ( α : Ordinal {n} lv )  :  Set (suc n) where
     fsub : ( ψ : Ordinal {n} lv → Set n ) → Constructible  α
     xself : Ordinal {n} lv → Constructible  α
  
  record ConstructibleSet {n : Level } : Set (suc n) where
   field
      level : Nat
      α : Ordinal {n} level 
      constructible : Constructible α
  
  open ConstructibleSet
  
  data _c∋_ {n : Level } : {lv lv' : Nat} {α : Ordinal {n} lv } {α' : Ordinal {n} lv' } →
        Constructible {n} {lv} α → Constructible {n} {lv'} α' → Set n where
     c> : {lv lv' : Nat} {α : Ordinal {n} lv } {α' : Ordinal {n} lv' }
        (ta : Constructible {n} {lv} α ) ( tx : Constructible {n} {lv'} α' ) → α' o< α →  ta c∋ tx
     xself-fsub  : {lv : Nat} {α : Ordinal {n} lv } 
         (ta : Ordinal {n} lv ) ( ψ : Ordinal {n} lv → Set n ) → _c∋_ {n} {_} {_} {α} {α} (xself ta ) ( fsub ψ)  
     fsub-fsub : {lv lv' : Nat} {α : Ordinal {n} lv } 
          ( ψ : Ordinal {n} lv → Set n ) ( ψ₁ : Ordinal {n} lv → Set n ) →
         ( ∀ ( x :  Ordinal {n} lv ) → ψ x →  ψ₁ x ) →  _c∋_ {n} {_} {_} {α} {α} ( fsub ψ ) ( fsub ψ₁) 

  _∋_  : {n : Level} → (ConstructibleSet {n}) → (ConstructibleSet {n} ) → Set n 
  a ∋ x  = constructible a c∋ constructible x

  data _c≈_ {n : Level } : {lv lv' : Nat} {α : Ordinal {n} lv } {α' : Ordinal {n} lv' } →
        Constructible {n} {lv} α → Constructible {n} {lv'} α' → Set n where
      crefl :  {lv : Nat} {α : Ordinal {n} lv } → _c≈_ {n} {_} {_} {α} {α} (xself α ) (xself α )
      feq :  {lv : Nat} {α : Ordinal {n} lv }
          → ( ψ : Ordinal {n} lv → Set n ) ( ψ₁ : Ordinal {n} lv → Set n ) 
          → (∀ ( x :  Ordinal {n} lv ) → ψ x  ⇔ ψ₁ x ) → _c≈_ {n} {_} {_} {α} {α} ( fsub ψ ) ( fsub ψ₁)

  _≈_  : {n : Level} → (ConstructibleSet {n}) → (ConstructibleSet {n} ) → Set n 
  a ≈ x  = constructible a c≈ constructible x
      
  
  ConstructibleSet→ZF : {n : Level } → ZF {suc n} {n}
  ConstructibleSet→ZF {n}  = record {
       ZFSet = ConstructibleSet 
     ; _∋_ = _∋_
     ; _≈_ = _≈_ 
     ; ∅  = record { level = Zero ; α = Φ ; constructible = xself Φ }
     ; _×_ = {!!}
     ; Union = {!!}
     ; Power = {!!}
     ; Select = {!!}
     ; Replace = {!!}
     ; infinite = {!!}
     ; isZF = {!!}
    } where