| Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > MPE Home > Th. List > df-sgn | Structured version Visualization version GIF version | ||
| Description: Signum function. We do not call it "sign", which is homophonic with "sine" (df-sin 16011). Defined as "sgn" in ISO 80000-2:2009(E) operation 2-9.13. It is named "sign" (with the same definition) in the "NIST Digital Library of Mathematical Functions" , front introduction, "Common Notations and Definitions" section at http://dlmf.nist.gov/front/introduction#Sx4 16011. We define this over ℝ* (df-xr 11188) instead of ℝ so that it can accept +∞ and -∞. Note that df-psgn 19397 defines the sign of a permutation, which is different. Value shown in sgnval 15030. (Contributed by David A. Wheeler, 15-May-2015.) |
| Ref | Expression |
|---|---|
| df-sgn | ⊢ sgn = (𝑥 ∈ ℝ* ↦ if(𝑥 = 0, 0, if(𝑥 < 0, -1, 1))) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | csgn 15028 | . 2 class sgn | |
| 2 | vx | . . 3 setvar 𝑥 | |
| 3 | cxr 11183 | . . 3 class ℝ* | |
| 4 | 2 | cv 1539 | . . . . 5 class 𝑥 |
| 5 | cc0 11044 | . . . . 5 class 0 | |
| 6 | 4, 5 | wceq 1540 | . . . 4 wff 𝑥 = 0 |
| 7 | clt 11184 | . . . . . 6 class < | |
| 8 | 4, 5, 7 | wbr 5102 | . . . . 5 wff 𝑥 < 0 |
| 9 | c1 11045 | . . . . . 6 class 1 | |
| 10 | 9 | cneg 11382 | . . . . 5 class -1 |
| 11 | 8, 10, 9 | cif 4484 | . . . 4 class if(𝑥 < 0, -1, 1) |
| 12 | 6, 5, 11 | cif 4484 | . . 3 class if(𝑥 = 0, 0, if(𝑥 < 0, -1, 1)) |
| 13 | 2, 3, 12 | cmpt 5183 | . 2 class (𝑥 ∈ ℝ* ↦ if(𝑥 = 0, 0, if(𝑥 < 0, -1, 1))) |
| 14 | 1, 13 | wceq 1540 | 1 wff sgn = (𝑥 ∈ ℝ* ↦ if(𝑥 = 0, 0, if(𝑥 < 0, -1, 1))) |
| Colors of variables: wff setvar class |
| This definition is referenced by: sgnval 15030 |
| Copyright terms: Public domain | W3C validator |