Metamath Proof Explorer
		
		
		
		Description:  A modus tollens deduction involving disjunction.  (Contributed by Jeff
       Hankins, 15-Jul-2009)
		
			
				
					|  |  | Ref | Expression | 
					
						|  | Hypotheses | mtord.1 | ⊢ ( 𝜑  →  ¬  𝜒 ) | 
					
						|  |  | mtord.2 | ⊢ ( 𝜑  →  ¬  𝜃 ) | 
					
						|  |  | mtord.3 | ⊢ ( 𝜑  →  ( 𝜓  →  ( 𝜒  ∨  𝜃 ) ) ) | 
				
					|  | Assertion | mtord | ⊢  ( 𝜑  →  ¬  𝜓 ) | 
			
		
		
			
				Proof
				
					
						| Step | Hyp | Ref | Expression | 
						
							| 1 |  | mtord.1 | ⊢ ( 𝜑  →  ¬  𝜒 ) | 
						
							| 2 |  | mtord.2 | ⊢ ( 𝜑  →  ¬  𝜃 ) | 
						
							| 3 |  | mtord.3 | ⊢ ( 𝜑  →  ( 𝜓  →  ( 𝜒  ∨  𝜃 ) ) ) | 
						
							| 4 |  | pm2.53 | ⊢ ( ( 𝜒  ∨  𝜃 )  →  ( ¬  𝜒  →  𝜃 ) ) | 
						
							| 5 | 3 1 4 | syl6ci | ⊢ ( 𝜑  →  ( 𝜓  →  𝜃 ) ) | 
						
							| 6 | 2 5 | mtod | ⊢ ( 𝜑  →  ¬  𝜓 ) |