comparison zf.agda @ 37:f10ceee99d00

¬ ( y c< x ) → x ≡ od∅
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Thu, 23 May 2019 13:48:27 +0900
parents c9ad0d97ce41
children 83b13f1f4f42
comparison
equal deleted inserted replaced
36:4d64509067d0 37:f10ceee99d00
69 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → _⊆_ t A {x} 69 power→ : ∀( A t : ZFSet ) → Power A ∋ t → ∀ {x} → _⊆_ t A {x}
70 power← : ∀( A t : ZFSet ) → ∀ {x} → _⊆_ t A {x} → Power A ∋ t 70 power← : ∀( A t : ZFSet ) → ∀ {x} → _⊆_ t A {x} → Power A ∋ t
71 -- extentionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w ) 71 -- extentionality : ∀ z ( z ∈ x ⇔ z ∈ y ) ⇒ ∀ w ( x ∈ w ⇔ y ∈ w )
72 extentionality : ( A B z : ZFSet ) → (( A ∋ z ) ⇔ (B ∋ z) ) → A ≈ B 72 extentionality : ( A B z : ZFSet ) → (( A ∋ z ) ⇔ (B ∋ z) ) → A ≈ B
73 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) ) 73 -- regularity : ∀ x ( x ≠ ∅ → ∃ y ∈ x ( y ∩ x = ∅ ) )
74 minimul : ZFSet → ( ZFSet ∧ ZFSet ) 74 minimul : (x : ZFSet ) → ¬ (x ≈ ∅) → ZFSet
75 regularity : ∀( x : ZFSet ) → ¬ (x ≈ ∅) → ( proj1 ( minimul x ) ∈ x ∧ ( proj1 ( minimul x ) ∩ x ≈ ∅ ) ) 75 regularity : ∀( x : ZFSet ) → (not : ¬ (x ≈ ∅)) → ( minimul x not ∈ x ∧ ( minimul x not ∩ x ≈ ∅ ) )
76 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) ) 76 -- infinity : ∃ A ( ∅ ∈ A ∧ ∀ x ∈ A ( x ∪ { x } ∈ A ) )
77 infinity∅ : ∅ ∈ infinite 77 infinity∅ : ∅ ∈ infinite
78 infinity : ∀( X x : ZFSet ) → x ∈ infinite → ( x ∪ Select X ( λ y → x ≈ y )) ∈ infinite 78 infinity : ∀( X x : ZFSet ) → x ∈ infinite → ( x ∪ Select X ( λ y → x ≈ y )) ∈ infinite
79 selection : { ψ : ZFSet → Set m } → ∀ ( X y : ZFSet ) → ( y ∈ Select X ψ ) → ψ y 79 selection : { ψ : ZFSet → Set m } → ∀ ( X y : ZFSet ) → ( y ∈ Select X ψ ) → ψ y
80 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) ) 80 -- replacement : ∀ x ∀ y ∀ z ( ( ψ ( x , y ) ∧ ψ ( x , z ) ) → y = z ) → ∀ X ∃ A ∀ y ( y ∈ A ↔ ∃ x ∈ X ψ ( x , y ) )