Metamath Proof Explorer


Theorem mul0ori

Description: If a product is zero, one of its factors must be zero. Theorem I.11 of Apostol p. 18. (Contributed by NM, 7-Oct-1999)

Ref Expression
Hypotheses mul0or.1 𝐴 ∈ ℂ
mul0or.2 𝐵 ∈ ℂ
Assertion mul0ori ( ( 𝐴 · 𝐵 ) = 0 ↔ ( 𝐴 = 0 ∨ 𝐵 = 0 ) )

Proof

Step Hyp Ref Expression
1 mul0or.1 𝐴 ∈ ℂ
2 mul0or.2 𝐵 ∈ ℂ
3 mul0or ( ( 𝐴 ∈ ℂ ∧ 𝐵 ∈ ℂ ) → ( ( 𝐴 · 𝐵 ) = 0 ↔ ( 𝐴 = 0 ∨ 𝐵 = 0 ) ) )
4 1 2 3 mp2an ( ( 𝐴 · 𝐵 ) = 0 ↔ ( 𝐴 = 0 ∨ 𝐵 = 0 ) )