Database
ZF (ZERMELO-FRAENKEL) SET THEORY
ZF Set Theory - start with the Axiom of Extensionality
Unordered and ordered pairs
ralunsn
Metamath Proof Explorer
Description: Restricted quantification over the union of a set and a singleton, using
implicit substitution. (Contributed by Paul Chapman , 17-Nov-2012)
(Revised by Mario Carneiro , 23-Apr-2015)
Ref
Expression
Hypothesis
ralunsn.1
⊢ ( 𝑥 = 𝐵 → ( 𝜑 ↔ 𝜓 ) )
Assertion
ralunsn
⊢ ( 𝐵 ∈ 𝐶 → ( ∀ 𝑥 ∈ ( 𝐴 ∪ { 𝐵 } ) 𝜑 ↔ ( ∀ 𝑥 ∈ 𝐴 𝜑 ∧ 𝜓 ) ) )
Proof
Step
Hyp
Ref
Expression
1
ralunsn.1
⊢ ( 𝑥 = 𝐵 → ( 𝜑 ↔ 𝜓 ) )
2
ralunb
⊢ ( ∀ 𝑥 ∈ ( 𝐴 ∪ { 𝐵 } ) 𝜑 ↔ ( ∀ 𝑥 ∈ 𝐴 𝜑 ∧ ∀ 𝑥 ∈ { 𝐵 } 𝜑 ) )
3
1
ralsng
⊢ ( 𝐵 ∈ 𝐶 → ( ∀ 𝑥 ∈ { 𝐵 } 𝜑 ↔ 𝜓 ) )
4
3
anbi2d
⊢ ( 𝐵 ∈ 𝐶 → ( ( ∀ 𝑥 ∈ 𝐴 𝜑 ∧ ∀ 𝑥 ∈ { 𝐵 } 𝜑 ) ↔ ( ∀ 𝑥 ∈ 𝐴 𝜑 ∧ 𝜓 ) ) )
5
2 4
syl5bb
⊢ ( 𝐵 ∈ 𝐶 → ( ∀ 𝑥 ∈ ( 𝐴 ∪ { 𝐵 } ) 𝜑 ↔ ( ∀ 𝑥 ∈ 𝐴 𝜑 ∧ 𝜓 ) ) )