Description: A theorem used to prove the base case of the Eliminability Theorem (see section comment): variable equals abstraction. (Contributed by BJ, 30-Apr-2024) (Proof modification is discouraged.) (New usage is discouraged.)
Ref | Expression | ||
---|---|---|---|
Assertion | eliminable-veqab | ⊢ ( 𝑥 = { 𝑦 ∣ 𝜑 } ↔ ∀ 𝑧 ( 𝑧 ∈ 𝑥 ↔ [ 𝑧 / 𝑦 ] 𝜑 ) ) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | dfcleq | ⊢ ( 𝑥 = { 𝑦 ∣ 𝜑 } ↔ ∀ 𝑧 ( 𝑧 ∈ 𝑥 ↔ 𝑧 ∈ { 𝑦 ∣ 𝜑 } ) ) | |
2 | eliminable-velab | ⊢ ( 𝑧 ∈ { 𝑦 ∣ 𝜑 } ↔ [ 𝑧 / 𝑦 ] 𝜑 ) | |
3 | 2 | bibi2i | ⊢ ( ( 𝑧 ∈ 𝑥 ↔ 𝑧 ∈ { 𝑦 ∣ 𝜑 } ) ↔ ( 𝑧 ∈ 𝑥 ↔ [ 𝑧 / 𝑦 ] 𝜑 ) ) |
4 | 3 | albii | ⊢ ( ∀ 𝑧 ( 𝑧 ∈ 𝑥 ↔ 𝑧 ∈ { 𝑦 ∣ 𝜑 } ) ↔ ∀ 𝑧 ( 𝑧 ∈ 𝑥 ↔ [ 𝑧 / 𝑦 ] 𝜑 ) ) |
5 | 1 4 | bitri | ⊢ ( 𝑥 = { 𝑦 ∣ 𝜑 } ↔ ∀ 𝑧 ( 𝑧 ∈ 𝑥 ↔ [ 𝑧 / 𝑦 ] 𝜑 ) ) |