Metamath Proof Explorer
		
		
		
		Description:  A lemma for Conjunctive Normal Form unit propagation, in double
       deduction form.  (Contributed by Giovanni Mascellani, 19-Mar-2018)
		
			
				
					|  |  | Ref | Expression | 
					
						|  | Hypotheses | cnfn2dd.1 | ⊢ ( 𝜑  →  ( 𝜓  →  𝜃 ) ) | 
					
						|  |  | cnfn2dd.2 | ⊢ ( 𝜑  →  ( 𝜓  →  ( 𝜒  ∨  ¬  𝜃 ) ) ) | 
				
					|  | Assertion | cnfn2dd | ⊢  ( 𝜑  →  ( 𝜓  →  𝜒 ) ) | 
			
		
		
			
				Proof
				
					
						| Step | Hyp | Ref | Expression | 
						
							| 1 |  | cnfn2dd.1 | ⊢ ( 𝜑  →  ( 𝜓  →  𝜃 ) ) | 
						
							| 2 |  | cnfn2dd.2 | ⊢ ( 𝜑  →  ( 𝜓  →  ( 𝜒  ∨  ¬  𝜃 ) ) ) | 
						
							| 3 |  | notnot | ⊢ ( 𝜃  →  ¬  ¬  𝜃 ) | 
						
							| 4 | 1 3 | syl6 | ⊢ ( 𝜑  →  ( 𝜓  →  ¬  ¬  𝜃 ) ) | 
						
							| 5 | 4 2 | cnf2dd | ⊢ ( 𝜑  →  ( 𝜓  →  𝜒 ) ) |