Metamath Proof Explorer
		
		
		
		Description:  Degree of a nonzero univariate polynomial.  (Contributed by Stefan
       O'Rear, 23-Mar-2015)  (Revised by Mario Carneiro, 7-Oct-2015)
		
			
				
					|  |  | Ref | Expression | 
					
						|  | Hypotheses | deg1z.d | ⊢ 𝐷  =  ( deg1 ‘ 𝑅 ) | 
					
						|  |  | deg1z.p | ⊢ 𝑃  =  ( Poly1 ‘ 𝑅 ) | 
					
						|  |  | deg1z.z | ⊢  0   =  ( 0g ‘ 𝑃 ) | 
					
						|  |  | deg1nn0cl.b | ⊢ 𝐵  =  ( Base ‘ 𝑃 ) | 
				
					|  | Assertion | deg1nn0cl | ⊢  ( ( 𝑅  ∈  Ring  ∧  𝐹  ∈  𝐵  ∧  𝐹  ≠   0  )  →  ( 𝐷 ‘ 𝐹 )  ∈  ℕ0 ) | 
			
		
		
			
				Proof
				
					
						| Step | Hyp | Ref | Expression | 
						
							| 1 |  | deg1z.d | ⊢ 𝐷  =  ( deg1 ‘ 𝑅 ) | 
						
							| 2 |  | deg1z.p | ⊢ 𝑃  =  ( Poly1 ‘ 𝑅 ) | 
						
							| 3 |  | deg1z.z | ⊢  0   =  ( 0g ‘ 𝑃 ) | 
						
							| 4 |  | deg1nn0cl.b | ⊢ 𝐵  =  ( Base ‘ 𝑃 ) | 
						
							| 5 | 1 | deg1fval | ⊢ 𝐷  =  ( 1o  mDeg  𝑅 ) | 
						
							| 6 |  | eqid | ⊢ ( 1o  mPoly  𝑅 )  =  ( 1o  mPoly  𝑅 ) | 
						
							| 7 | 6 2 3 | ply1mpl0 | ⊢  0   =  ( 0g ‘ ( 1o  mPoly  𝑅 ) ) | 
						
							| 8 | 2 4 | ply1bas | ⊢ 𝐵  =  ( Base ‘ ( 1o  mPoly  𝑅 ) ) | 
						
							| 9 | 5 6 7 8 | mdegnn0cl | ⊢ ( ( 𝑅  ∈  Ring  ∧  𝐹  ∈  𝐵  ∧  𝐹  ≠   0  )  →  ( 𝐷 ‘ 𝐹 )  ∈  ℕ0 ) |