comparison zf.agda @ 9:5ed16e2d8eb7

try to fix axiom of replacement
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Sun, 12 May 2019 21:18:38 +0900
parents cb014a103467
children 8022e14fce74
comparison
equal deleted inserted replaced
8:cb014a103467 9:5ed16e2d8eb7
28 28
29 infixr 130 _∧_ 29 infixr 130 _∧_
30 infixr 140 _∨_ 30 infixr 140 _∨_
31 infixr 150 _⇔_ 31 infixr 150 _⇔_
32 32
33 record Func {n m : Level } (ZFSet : Set n) (_≈_ : Rel ZFSet m) : Set (n ⊔ suc m) where
34 field
35 Restrict : ZFSet
36 rel : Rel ZFSet m
37 dom : ( y : ZFSet ) → ∀ { x : ZFSet } → rel x y
38 fun-eq : { x y z : ZFSet } → ( rel x y ∧ rel x z ) → y ≈ z
39
40 open Func
41
42
33 record IsZF {n m : Level } 43 record IsZF {n m : Level }
34 (ZFSet : Set n) 44 (ZFSet : Set n)
35 (_∋_ : ( A x : ZFSet ) → Set m) 45 (_∋_ : ( A x : ZFSet ) → Set m)
36 (_≈_ : ( A B : ZFSet ) → Set m) 46 (_≈_ : Rel ZFSet m)
37 (∅ : ZFSet) 47 (∅ : ZFSet)
38 (_×_ : ( A B : ZFSet ) → ZFSet) 48 (_×_ : ( A B : ZFSet ) → ZFSet)
39 (Union : ( A : ZFSet ) → ZFSet) 49 (Union : ( A : ZFSet ) → ZFSet)
40 (Power : ( A : ZFSet ) → ZFSet) 50 (Power : ( A : ZFSet ) → ZFSet)
41 (Restrict : ( ZFSet → Set m ) → ZFSet)
42 (infinite : ZFSet) 51 (infinite : ZFSet)
43 : Set (suc (n ⊔ m)) where 52 : Set (suc (n ⊔ m)) where
44 field 53 field
45 isEquivalence : {A B : ZFSet} → IsEquivalence {n} {m} {ZFSet} _≈_ 54 isEquivalence : {A B : ZFSet} → IsEquivalence {n} {m} {ZFSet} _≈_
46 -- ∀ x ∀ y ∃ z(x ∈ z ∧ y ∈ z) 55 -- ∀ x ∀ y ∃ z(x ∈ z ∧ y ∈ z)
50 union← : ( X x y : ZFSet ) → Union X ∋ y → X ∋ x → x ∋ y 59 union← : ( X x y : ZFSet ) → Union X ∋ y → X ∋ x → x ∋ y
51 _∈_ : ( A B : ZFSet ) → Set m 60 _∈_ : ( A B : ZFSet ) → Set m
52 A ∈ B = B ∋ A 61 A ∈ B = B ∋ A
53 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → ∀{ A∋x : Set m } → Set m 62 _⊆_ : ( A B : ZFSet ) → ∀{ x : ZFSet } → ∀{ A∋x : Set m } → Set m
54 _⊆_ A B {x} {A∋x} = B ∋ x 63 _⊆_ A B {x} {A∋x} = B ∋ x
64 Repl : ( ψ : ZFSet → Set m ) → Func ZFSet _≈_
65 Repl ψ = record { Restrict = {!!} ; rel = {!!} ; dom = {!!} ; fun-eq = {!!} }
55 _∩_ : ( A B : ZFSet ) → ZFSet 66 _∩_ : ( A B : ZFSet ) → ZFSet
56 A ∩ B = Restrict ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) 67 A ∩ B = Restrict ( Repl ( λ x → ( A ∋ x ) ∧ ( B ∋ x ) ) )
57 _∪_ : ( A B : ZFSet ) → ZFSet 68 _∪_ : ( A B : ZFSet ) → ZFSet
58 A ∪ B = Restrict ( λ x → ( A ∋ x ) ∨ ( B ∋ x ) ) 69 A ∪ B = Restrict ( Repl ( λ x → ( A ∋ x ) ∨ ( B ∋ x ) ) )
59 infixr 200 _∈_ 70 infixr 200 _∈_
60 infixr 230 _∩_ _∪_ 71 infixr 230 _∩_ _∪_
61 infixr 220 _⊆_ 72 infixr 220 _⊆_
62 field 73 field
63 empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x ) 74 empty : ∀( x : ZFSet ) → ¬ ( ∅ ∋ x )
69 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) ) 80 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) )
70 smaller : ZFSet → ZFSet 81 smaller : ZFSet → ZFSet
71 regularity : ∀( x : ZFSet ) → ¬ (x ≈ ∅) → ( smaller x ∈ x ∧ ( smaller x ∩ x ≈ ∅ ) ) 82 regularity : ∀( x : ZFSet ) → ¬ (x ≈ ∅) → ( smaller x ∈ x ∧ ( smaller x ∩ x ≈ ∅ ) )
72 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) ) 83 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) )
73 infinity∅ : ∅ ∈ infinite 84 infinity∅ : ∅ ∈ infinite
74 infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ Restrict ( λ y → x ≈ y )) ∈ infinite 85 infinity : ∀( x : ZFSet ) → x ∈ infinite → ( x ∪ Restrict ( Repl ( λ y → x ≈ y ))) ∈ infinite
75 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) ) 86 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) )
76 -- this form looks like specification 87 -- this form looks like specification
77 replacement : ( ψ : ZFSet → Set m ) → ( x : ZFSet ) → x ∈ Restrict ψ → ψ x 88 replacement : ( ψ : Func ZFSet _≈_ ) → ∀ ( y : ZFSet ) → ( y ∈ Restrict ψ ) → {!!} -- dom ψ y
78 89
79 record ZF {n m : Level } : Set (suc (n ⊔ m)) where 90 record ZF {n m : Level } : Set (suc (n ⊔ m)) where
80 infixr 210 _×_ 91 infixr 210 _×_
81 infixl 200 _∋_ 92 infixl 200 _∋_
82 infixr 220 _≈_ 93 infixr 220 _≈_
87 -- ZF Set constructor 98 -- ZF Set constructor
88 ∅ : ZFSet 99 ∅ : ZFSet
89 _×_ : ( A B : ZFSet ) → ZFSet 100 _×_ : ( A B : ZFSet ) → ZFSet
90 Union : ( A : ZFSet ) → ZFSet 101 Union : ( A : ZFSet ) → ZFSet
91 Power : ( A : ZFSet ) → ZFSet 102 Power : ( A : ZFSet ) → ZFSet
92 Restrict : ( ZFSet → Set m ) → ZFSet
93 infinite : ZFSet 103 infinite : ZFSet
94 isZF : IsZF ZFSet _∋_ _≈_ ∅ _×_ Union Power Restrict infinite 104 isZF : IsZF ZFSet _∋_ _≈_ ∅ _×_ Union Power infinite
95 105
96 module reguraliry-m {n m : Level } ( zf : ZF {m} {n} ) where 106 module reguraliry-m {n m : Level } ( zf : ZF {m} {n} ) where
97 107
98 open import Relation.Nullary 108 open import Relation.Nullary
99 open import Data.Empty 109 open import Data.Empty
100 open import Relation.Binary.PropositionalEquality 110 open import Relation.Binary.PropositionalEquality
101 111
102 _≈_ = ZF._≈_ zf 112 _≈_ = ZF._≈_ zf
103 ZFSet = ZF.ZFSet 113 ZFSet = ZF.ZFSet
104 Restrict = ZF.Restrict zf
105 ∅ = ZF.∅ zf 114 ∅ = ZF.∅ zf
106 _∩_ = ( IsZF._∩_ ) (ZF.isZF zf) 115 _∩_ = ( IsZF._∩_ ) (ZF.isZF zf)
107 _∋_ = ZF._∋_ zf 116 _∋_ = ZF._∋_ zf
108 replacement = IsZF.replacement ( ZF.isZF zf ) 117 replacement = IsZF.replacement ( ZF.isZF zf )
109 118
110 russel : ( x : ZFSet zf ) → x ≈ Restrict ( λ x → ¬ ( x ∋ x ) ) → ⊥ 119 -- russel : ( x : ZFSet zf ) → x ≈ Restrict ( λ x → ¬ ( x ∋ x ) ) → ⊥
111 russel x eq with x ∋ x 120 -- russel x eq with x ∋ x
112 ... | x∋x with replacement ( λ x → ¬ ( x ∋ x )) x {!!} 121 -- ... | x∋x with replacement ( λ x → ¬ ( x ∋ x )) x {!!}
113 ... | r1 = {!!} 122 -- ... | r1 = {!!}
114 123
115 124
116 125
117 126
118 data Nat : Set zero where 127 data Nat : Set zero where