Database
ZF (ZERMELO-FRAENKEL) SET THEORY
ZF Set Theory - start with the Axiom of Extensionality
The difference, union, and intersection of two classes
The symmetric difference of two classes
elsymdifxor
Metamath Proof Explorer
Description: Membership in a symmetric difference is an exclusive-or relationship.
(Contributed by David A. Wheeler , 26-Apr-2020) (Proof shortened by BJ , 13-Aug-2022)
Ref
Expression
Assertion
elsymdifxor
⊢ ( 𝐴 ∈ ( 𝐵 △ 𝐶 ) ↔ ( 𝐴 ∈ 𝐵 ⊻ 𝐴 ∈ 𝐶 ) )
Proof
Step
Hyp
Ref
Expression
1
elsymdif
⊢ ( 𝐴 ∈ ( 𝐵 △ 𝐶 ) ↔ ¬ ( 𝐴 ∈ 𝐵 ↔ 𝐴 ∈ 𝐶 ) )
2
df-xor
⊢ ( ( 𝐴 ∈ 𝐵 ⊻ 𝐴 ∈ 𝐶 ) ↔ ¬ ( 𝐴 ∈ 𝐵 ↔ 𝐴 ∈ 𝐶 ) )
3
1 2
bitr4i
⊢ ( 𝐴 ∈ ( 𝐵 △ 𝐶 ) ↔ ( 𝐴 ∈ 𝐵 ⊻ 𝐴 ∈ 𝐶 ) )