Metamath Proof Explorer


Theorem eqvreldisj2

Description: The elements of the quotient set of an equivalence relation are disjoint (cf. eqvreldisj3 ). (Contributed by Mario Carneiro, 10-Dec-2016) (Revised by Peter Mazsa, 19-Sep-2021)

Ref Expression
Assertion eqvreldisj2 ( EqvRel 𝑅 → ElDisj ( 𝐴 / 𝑅 ) )

Proof

Step Hyp Ref Expression
1 eqvreldisj1 ( EqvRel 𝑅 → ∀ 𝑥 ∈ ( 𝐴 / 𝑅 ) ∀ 𝑦 ∈ ( 𝐴 / 𝑅 ) ( 𝑥 = 𝑦 ∨ ( 𝑥𝑦 ) = ∅ ) )
2 dfeldisj5 ( ElDisj ( 𝐴 / 𝑅 ) ↔ ∀ 𝑥 ∈ ( 𝐴 / 𝑅 ) ∀ 𝑦 ∈ ( 𝐴 / 𝑅 ) ( 𝑥 = 𝑦 ∨ ( 𝑥𝑦 ) = ∅ ) )
3 1 2 sylibr ( EqvRel 𝑅 → ElDisj ( 𝐴 / 𝑅 ) )