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

Theorem oddm1even 15529
Description: An integer is odd iff its predecessor is even. (Contributed by Mario Carneiro, 5-Sep-2016.)
Assertion
Ref Expression
oddm1even (𝑁 ∈ ℤ → (¬ 2 ∥ 𝑁 ↔ 2 ∥ (𝑁 − 1)))

Proof of Theorem oddm1even
Dummy variable 𝑛 is distinct from all other variables.
StepHypRef Expression
1 simpl 483 . . . . . 6 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → 𝑁 ∈ ℤ)
21zcnd 11942 . . . . 5 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → 𝑁 ∈ ℂ)
3 1cnd 10489 . . . . 5 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → 1 ∈ ℂ)
4 2cnd 11569 . . . . . 6 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → 2 ∈ ℂ)
5 simpr 485 . . . . . . 7 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → 𝑛 ∈ ℤ)
65zcnd 11942 . . . . . 6 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → 𝑛 ∈ ℂ)
74, 6mulcld 10514 . . . . 5 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → (2 · 𝑛) ∈ ℂ)
82, 3, 7subadd2d 10870 . . . 4 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → ((𝑁 − 1) = (2 · 𝑛) ↔ ((2 · 𝑛) + 1) = 𝑁))
9 eqcom 2804 . . . . 5 ((𝑁 − 1) = (2 · 𝑛) ↔ (2 · 𝑛) = (𝑁 − 1))
104, 6mulcomd 10515 . . . . . 6 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → (2 · 𝑛) = (𝑛 · 2))
1110eqeq1d 2799 . . . . 5 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → ((2 · 𝑛) = (𝑁 − 1) ↔ (𝑛 · 2) = (𝑁 − 1)))
129, 11syl5bb 284 . . . 4 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → ((𝑁 − 1) = (2 · 𝑛) ↔ (𝑛 · 2) = (𝑁 − 1)))
138, 12bitr3d 282 . . 3 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → (((2 · 𝑛) + 1) = 𝑁 ↔ (𝑛 · 2) = (𝑁 − 1)))
1413rexbidva 3261 . 2 (𝑁 ∈ ℤ → (∃𝑛 ∈ ℤ ((2 · 𝑛) + 1) = 𝑁 ↔ ∃𝑛 ∈ ℤ (𝑛 · 2) = (𝑁 − 1)))
15 odd2np1 15527 . 2 (𝑁 ∈ ℤ → (¬ 2 ∥ 𝑁 ↔ ∃𝑛 ∈ ℤ ((2 · 𝑛) + 1) = 𝑁))
16 2z 11868 . . 3 2 ∈ ℤ
17 peano2zm 11879 . . 3 (𝑁 ∈ ℤ → (𝑁 − 1) ∈ ℤ)
18 divides 15446 . . 3 ((2 ∈ ℤ ∧ (𝑁 − 1) ∈ ℤ) → (2 ∥ (𝑁 − 1) ↔ ∃𝑛 ∈ ℤ (𝑛 · 2) = (𝑁 − 1)))
1916, 17, 18sylancr 587 . 2 (𝑁 ∈ ℤ → (2 ∥ (𝑁 − 1) ↔ ∃𝑛 ∈ ℤ (𝑛 · 2) = (𝑁 − 1)))
2014, 15, 193bitr4d 312 1 (𝑁 ∈ ℤ → (¬ 2 ∥ 𝑁 ↔ 2 ∥ (𝑁 − 1)))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wb 207  wa 396   = wceq 1525  wcel 2083  wrex 3108   class class class wbr 4968  (class class class)co 7023  1c1 10391   + caddc 10393   · cmul 10395  cmin 10723  2c2 11546  cz 11835  cdvds 15444
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1781  ax-4 1795  ax-5 1892  ax-6 1951  ax-7 1996  ax-8 2085  ax-9 2093  ax-10 2114  ax-11 2128  ax-12 2143  ax-13 2346  ax-ext 2771  ax-sep 5101  ax-nul 5108  ax-pow 5164  ax-pr 5228  ax-un 7326  ax-resscn 10447  ax-1cn 10448  ax-icn 10449  ax-addcl 10450  ax-addrcl 10451  ax-mulcl 10452  ax-mulrcl 10453  ax-mulcom 10454  ax-addass 10455  ax-mulass 10456  ax-distr 10457  ax-i2m1 10458  ax-1ne0 10459  ax-1rid 10460  ax-rnegex 10461  ax-rrecex 10462  ax-cnre 10463  ax-pre-lttri 10464  ax-pre-lttrn 10465  ax-pre-ltadd 10466  ax-pre-mulgt0 10467
This theorem depends on definitions:  df-bi 208  df-an 397  df-or 843  df-3or 1081  df-3an 1082  df-tru 1528  df-ex 1766  df-nf 1770  df-sb 2045  df-mo 2578  df-eu 2614  df-clab 2778  df-cleq 2790  df-clel 2865  df-nfc 2937  df-ne 2987  df-nel 3093  df-ral 3112  df-rex 3113  df-reu 3114  df-rmo 3115  df-rab 3116  df-v 3442  df-sbc 3712  df-csb 3818  df-dif 3868  df-un 3870  df-in 3872  df-ss 3880  df-pss 3882  df-nul 4218  df-if 4388  df-pw 4461  df-sn 4479  df-pr 4481  df-tp 4483  df-op 4485  df-uni 4752  df-iun 4833  df-br 4969  df-opab 5031  df-mpt 5048  df-tr 5071  df-id 5355  df-eprel 5360  df-po 5369  df-so 5370  df-fr 5409  df-we 5411  df-xp 5456  df-rel 5457  df-cnv 5458  df-co 5459  df-dm 5460  df-rn 5461  df-res 5462  df-ima 5463  df-pred 6030  df-ord 6076  df-on 6077  df-lim 6078  df-suc 6079  df-iota 6196  df-fun 6234  df-fn 6235  df-f 6236  df-f1 6237  df-fo 6238  df-f1o 6239  df-fv 6240  df-riota 6984  df-ov 7026  df-oprab 7027  df-mpo 7028  df-om 7444  df-wrecs 7805  df-recs 7867  df-rdg 7905  df-er 8146  df-en 8365  df-dom 8366  df-sdom 8367  df-pnf 10530  df-mnf 10531  df-xr 10532  df-ltxr 10533  df-le 10534  df-sub 10725  df-neg 10726  df-div 11152  df-nn 11493  df-2 11554  df-n0 11752  df-z 11836  df-dvds 15445
This theorem is referenced by:  oddp1even  15530  n2dvds3OLD  15559  oddpwp1fsum  15580  bitscmp  15624  lighneallem1  43274  lighneallem3  43276  2dvdsoddm1  43326
  Copyright terms: Public domain W3C validator