Metamath Proof Explorer
		
		
		
		Description:  Distribution of multiplication over subtraction.  Theorem I.5 of
       Apostol p. 18.  (Contributed by Mario Carneiro, 27-May-2016)
		
			
				
					 | 
					 | 
					Ref | 
					Expression | 
				
					
						 | 
						Hypotheses | 
						mulm1d.1 | 
						⊢ ( 𝜑  →  𝐴  ∈  ℂ )  | 
					
					
						 | 
						 | 
						mulnegd.2 | 
						⊢ ( 𝜑  →  𝐵  ∈  ℂ )  | 
					
					
						 | 
						 | 
						subdid.3 | 
						⊢ ( 𝜑  →  𝐶  ∈  ℂ )  | 
					
				
					 | 
					Assertion | 
					subdid | 
					⊢  ( 𝜑  →  ( 𝐴  ·  ( 𝐵  −  𝐶 ) )  =  ( ( 𝐴  ·  𝐵 )  −  ( 𝐴  ·  𝐶 ) ) )  | 
				
			
		
		
			
				Proof
				
					
						| Step | 
						Hyp | 
						Ref | 
						Expression | 
					
						
							| 1 | 
							
								
							 | 
							mulm1d.1 | 
							⊢ ( 𝜑  →  𝐴  ∈  ℂ )  | 
						
						
							| 2 | 
							
								
							 | 
							mulnegd.2 | 
							⊢ ( 𝜑  →  𝐵  ∈  ℂ )  | 
						
						
							| 3 | 
							
								
							 | 
							subdid.3 | 
							⊢ ( 𝜑  →  𝐶  ∈  ℂ )  | 
						
						
							| 4 | 
							
								
							 | 
							subdi | 
							⊢ ( ( 𝐴  ∈  ℂ  ∧  𝐵  ∈  ℂ  ∧  𝐶  ∈  ℂ )  →  ( 𝐴  ·  ( 𝐵  −  𝐶 ) )  =  ( ( 𝐴  ·  𝐵 )  −  ( 𝐴  ·  𝐶 ) ) )  | 
						
						
							| 5 | 
							
								1 2 3 4
							 | 
							syl3anc | 
							⊢ ( 𝜑  →  ( 𝐴  ·  ( 𝐵  −  𝐶 ) )  =  ( ( 𝐴  ·  𝐵 )  −  ( 𝐴  ·  𝐶 ) ) )  |