MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  sgnp Structured version   Visualization version   GIF version

Theorem sgnp 14987
Description: The signum of a positive extended real is 1. (Contributed by David A. Wheeler, 15-May-2015.)
Assertion
Ref Expression
sgnp ((𝐴 ∈ ℝ* ∧ 0 < 𝐴) → (sgn‘𝐴) = 1)

Proof of Theorem sgnp
StepHypRef Expression
1 sgnval 14985 . . 3 (𝐴 ∈ ℝ* → (sgn‘𝐴) = if(𝐴 = 0, 0, if(𝐴 < 0, -1, 1)))
21adantr 481 . 2 ((𝐴 ∈ ℝ* ∧ 0 < 𝐴) → (sgn‘𝐴) = if(𝐴 = 0, 0, if(𝐴 < 0, -1, 1)))
3 0xr 11211 . . . . 5 0 ∈ ℝ*
4 xrltne 13092 . . . . 5 ((0 ∈ ℝ*𝐴 ∈ ℝ* ∧ 0 < 𝐴) → 𝐴 ≠ 0)
53, 4mp3an1 1448 . . . 4 ((𝐴 ∈ ℝ* ∧ 0 < 𝐴) → 𝐴 ≠ 0)
65neneqd 2944 . . 3 ((𝐴 ∈ ℝ* ∧ 0 < 𝐴) → ¬ 𝐴 = 0)
76iffalsed 4502 . 2 ((𝐴 ∈ ℝ* ∧ 0 < 𝐴) → if(𝐴 = 0, 0, if(𝐴 < 0, -1, 1)) = if(𝐴 < 0, -1, 1))
8 xrltnsym 13066 . . . . 5 ((0 ∈ ℝ*𝐴 ∈ ℝ*) → (0 < 𝐴 → ¬ 𝐴 < 0))
93, 8mpan 688 . . . 4 (𝐴 ∈ ℝ* → (0 < 𝐴 → ¬ 𝐴 < 0))
109imp 407 . . 3 ((𝐴 ∈ ℝ* ∧ 0 < 𝐴) → ¬ 𝐴 < 0)
1110iffalsed 4502 . 2 ((𝐴 ∈ ℝ* ∧ 0 < 𝐴) → if(𝐴 < 0, -1, 1) = 1)
122, 7, 113eqtrd 2775 1 ((𝐴 ∈ ℝ* ∧ 0 < 𝐴) → (sgn‘𝐴) = 1)
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 396   = wceq 1541  wcel 2106  wne 2939  ifcif 4491   class class class wbr 5110  cfv 6501  0cc0 11060  1c1 11061  *cxr 11197   < clt 11198  -cneg 11395  sgncsgn 14983
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1797  ax-4 1811  ax-5 1913  ax-6 1971  ax-7 2011  ax-8 2108  ax-9 2116  ax-10 2137  ax-11 2154  ax-12 2171  ax-ext 2702  ax-sep 5261  ax-nul 5268  ax-pow 5325  ax-pr 5389  ax-un 7677  ax-cnex 11116  ax-resscn 11117  ax-1cn 11118  ax-icn 11119  ax-addcl 11120  ax-addrcl 11121  ax-mulcl 11122  ax-i2m1 11128  ax-rnegex 11131  ax-cnre 11133  ax-pre-lttri 11134  ax-pre-lttrn 11135
This theorem depends on definitions:  df-bi 206  df-an 397  df-or 846  df-3or 1088  df-3an 1089  df-tru 1544  df-fal 1554  df-ex 1782  df-nf 1786  df-sb 2068  df-mo 2533  df-eu 2562  df-clab 2709  df-cleq 2723  df-clel 2809  df-nfc 2884  df-ne 2940  df-nel 3046  df-ral 3061  df-rex 3070  df-rab 3406  df-v 3448  df-sbc 3743  df-csb 3859  df-dif 3916  df-un 3918  df-in 3920  df-ss 3930  df-nul 4288  df-if 4492  df-pw 4567  df-sn 4592  df-pr 4594  df-op 4598  df-uni 4871  df-br 5111  df-opab 5173  df-mpt 5194  df-id 5536  df-po 5550  df-so 5551  df-xp 5644  df-rel 5645  df-cnv 5646  df-co 5647  df-dm 5648  df-rn 5649  df-res 5650  df-ima 5651  df-iota 6453  df-fun 6503  df-fn 6504  df-f 6505  df-f1 6506  df-fo 6507  df-f1o 6508  df-fv 6509  df-ov 7365  df-er 8655  df-en 8891  df-dom 8892  df-sdom 8893  df-pnf 11200  df-mnf 11201  df-xr 11202  df-ltxr 11203  df-neg 11397  df-sgn 14984
This theorem is referenced by:  sgnrrp  14988  sgn1  14989  sgnpnf  14990  sgncl  33227  sgnmul  33231  sgnmulrp2  33232  sgnsub  33233  sgnpbi  33235
  Copyright terms: Public domain W3C validator