comparison zf.agda @ 30:3b0fdb95618e

problem on Ordinal ( OSuc ℵ )
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Tue, 21 May 2019 00:30:01 +0900
parents fce60b99dc55
children c9ad0d97ce41
comparison
equal deleted inserted replaced
29:fce60b99dc55 30:3b0fdb95618e
78 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → _⊆_ t A {x} 78 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → _⊆_ t A {x}
79 power← : ∀( A t : ZFSet ) → ∀ {x} → _⊆_ t A {x} → Power A ∋ t 79 power← : ∀( A t : ZFSet ) → ∀ {x} → _⊆_ t A {x} → Power A ∋ t
80 -- extentionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w ) 80 -- extentionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w )
81 extentionality : ( A B z : ZFSet ) → (( A ∋ z ) ⇔ (B ∋ z) ) → A ≈ B 81 extentionality : ( A B z : ZFSet ) → (( A ∋ z ) ⇔ (B ∋ z) ) → A ≈ B
82 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) ) 82 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) )
83 minimul : ZFSet → ZFSet 83 minimul : ZFSet → ( ZFSet ∧ ZFSet )
84 regularity : ∀( x : ZFSet ) → ¬ (x ≈ ∅) → ( minimul x ∈ x ∧ ( minimul x ∩ x ≈ ∅ ) ) 84 regularity : ∀( x : ZFSet ) → ¬ (x ≈ ∅) → ( proj1 ( minimul x ) ∈ x ∧ ( proj1 ( minimul x ) ∩ x ≈ ∅ ) )
85 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) ) 85 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) )
86 infinity∅ : ∅ ∈ infinite 86 infinity∅ : ∅ ∈ infinite
87 infinity : ∀( X x : ZFSet ) → x ∈ infinite → ( x ∪ Select X ( λ y → x ≈ y )) ∈ infinite 87 infinity : ∀( X x : ZFSet ) → x ∈ infinite → ( x ∪ Select X ( λ y → x ≈ y )) ∈ infinite
88 selection : { ψ : ZFSet → Set m } → ∀ ( X y : ZFSet ) → ( y ∈ Select X ψ ) → ψ y 88 selection : { ψ : ZFSet → Set m } → ∀ ( X y : ZFSet ) → ( y ∈ Select X ψ ) → ψ y
89 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) ) 89 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) )