Metamath Proof Explorer
		
		
		
		Description:  The absolute value of a real number is either that number or its
       negative.  (Contributed by NM, 30-Sep-1999)
		
			
				
					|  |  | Ref | Expression | 
					
						|  | Hypothesis | sqrtthi.1 | ⊢ 𝐴  ∈  ℝ | 
				
					|  | Assertion | absori | ⊢  ( ( abs ‘ 𝐴 )  =  𝐴  ∨  ( abs ‘ 𝐴 )  =  - 𝐴 ) | 
			
		
		
			
				Proof
				
					
						| Step | Hyp | Ref | Expression | 
						
							| 1 |  | sqrtthi.1 | ⊢ 𝐴  ∈  ℝ | 
						
							| 2 |  | absor | ⊢ ( 𝐴  ∈  ℝ  →  ( ( abs ‘ 𝐴 )  =  𝐴  ∨  ( abs ‘ 𝐴 )  =  - 𝐴 ) ) | 
						
							| 3 | 1 2 | ax-mp | ⊢ ( ( abs ‘ 𝐴 )  =  𝐴  ∨  ( abs ‘ 𝐴 )  =  - 𝐴 ) |