Metamath Proof Explorer
		
		
		
		Description:  A lemma for not-or-not elimination, in deduction form.  (Contributed by Giovanni Mascellani, 19-Mar-2018)
		
			
				
					|  |  | Ref | Expression | 
					
						|  | Hypothesis | notornotel1.1 | ⊢ ( 𝜑  →  ¬  ( ¬  𝜓  ∨  𝜒 ) ) | 
				
					|  | Assertion | notornotel1 | ⊢  ( 𝜑  →  𝜓 ) | 
			
		
		
			
				Proof
				
					
						| Step | Hyp | Ref | Expression | 
						
							| 1 |  | notornotel1.1 | ⊢ ( 𝜑  →  ¬  ( ¬  𝜓  ∨  𝜒 ) ) | 
						
							| 2 |  | ioran | ⊢ ( ¬  ( ¬  𝜓  ∨  𝜒 )  ↔  ( ¬  ¬  𝜓  ∧  ¬  𝜒 ) ) | 
						
							| 3 | 2 | biimpi | ⊢ ( ¬  ( ¬  𝜓  ∨  𝜒 )  →  ( ¬  ¬  𝜓  ∧  ¬  𝜒 ) ) | 
						
							| 4 |  | simpl | ⊢ ( ( ¬  ¬  𝜓  ∧  ¬  𝜒 )  →  ¬  ¬  𝜓 ) | 
						
							| 5 |  | notnotr | ⊢ ( ¬  ¬  𝜓  →  𝜓 ) | 
						
							| 6 | 1 3 4 5 | 4syl | ⊢ ( 𝜑  →  𝜓 ) |