Metamath Proof Explorer
		
		
		
		Description:  If a product is zero, one of its factors must be zero.  Theorem I.11 of
       Apostol p. 18.  (Contributed by Mario Carneiro, 27-May-2016)
		
			
				
					|  |  | Ref | Expression | 
					
						|  | Hypotheses | msq0d.1 | ⊢ ( 𝜑  →  𝐴  ∈  ℂ ) | 
					
						|  |  | mul0ord.2 | ⊢ ( 𝜑  →  𝐵  ∈  ℂ ) | 
				
					|  | Assertion | mul0ord | ⊢  ( 𝜑  →  ( ( 𝐴  ·  𝐵 )  =  0  ↔  ( 𝐴  =  0  ∨  𝐵  =  0 ) ) ) | 
			
		
		
			
				Proof
				
					
						| Step | Hyp | Ref | Expression | 
						
							| 1 |  | msq0d.1 | ⊢ ( 𝜑  →  𝐴  ∈  ℂ ) | 
						
							| 2 |  | mul0ord.2 | ⊢ ( 𝜑  →  𝐵  ∈  ℂ ) | 
						
							| 3 |  | mul0or | ⊢ ( ( 𝐴  ∈  ℂ  ∧  𝐵  ∈  ℂ )  →  ( ( 𝐴  ·  𝐵 )  =  0  ↔  ( 𝐴  =  0  ∨  𝐵  =  0 ) ) ) | 
						
							| 4 | 1 2 3 | syl2anc | ⊢ ( 𝜑  →  ( ( 𝐴  ·  𝐵 )  =  0  ↔  ( 𝐴  =  0  ∨  𝐵  =  0 ) ) ) |