![]() |
Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
|
Mirrors > Home > MPE Home > Th. List > zneo | Structured version Visualization version GIF version |
Description: No even integer equals an odd integer (i.e. no integer can be both even and odd). Exercise 10(a) of [Apostol] p. 28. (Contributed by NM, 31-Jul-2004.) (Proof shortened by Mario Carneiro, 18-May-2014.) |
Ref | Expression |
---|---|
zneo | ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (2 · 𝐴) ≠ ((2 · 𝐵) + 1)) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | halfnz 11668 | . . 3 ⊢ ¬ (1 / 2) ∈ ℤ | |
2 | 2cn 11304 | . . . . . . 7 ⊢ 2 ∈ ℂ | |
3 | zcn 11595 | . . . . . . . 8 ⊢ (𝐴 ∈ ℤ → 𝐴 ∈ ℂ) | |
4 | 3 | adantr 472 | . . . . . . 7 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → 𝐴 ∈ ℂ) |
5 | mulcl 10233 | . . . . . . 7 ⊢ ((2 ∈ ℂ ∧ 𝐴 ∈ ℂ) → (2 · 𝐴) ∈ ℂ) | |
6 | 2, 4, 5 | sylancr 698 | . . . . . 6 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (2 · 𝐴) ∈ ℂ) |
7 | zcn 11595 | . . . . . . . 8 ⊢ (𝐵 ∈ ℤ → 𝐵 ∈ ℂ) | |
8 | 7 | adantl 473 | . . . . . . 7 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → 𝐵 ∈ ℂ) |
9 | mulcl 10233 | . . . . . . 7 ⊢ ((2 ∈ ℂ ∧ 𝐵 ∈ ℂ) → (2 · 𝐵) ∈ ℂ) | |
10 | 2, 8, 9 | sylancr 698 | . . . . . 6 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (2 · 𝐵) ∈ ℂ) |
11 | 1cnd 10269 | . . . . . 6 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → 1 ∈ ℂ) | |
12 | 6, 10, 11 | subaddd 10623 | . . . . 5 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (((2 · 𝐴) − (2 · 𝐵)) = 1 ↔ ((2 · 𝐵) + 1) = (2 · 𝐴))) |
13 | 2 | a1i 11 | . . . . . . . . . 10 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → 2 ∈ ℂ) |
14 | 13, 4, 8 | subdid 10699 | . . . . . . . . 9 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (2 · (𝐴 − 𝐵)) = ((2 · 𝐴) − (2 · 𝐵))) |
15 | 14 | oveq1d 6830 | . . . . . . . 8 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((2 · (𝐴 − 𝐵)) / 2) = (((2 · 𝐴) − (2 · 𝐵)) / 2)) |
16 | zsubcl 11632 | . . . . . . . . . 10 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝐴 − 𝐵) ∈ ℤ) | |
17 | zcn 11595 | . . . . . . . . . 10 ⊢ ((𝐴 − 𝐵) ∈ ℤ → (𝐴 − 𝐵) ∈ ℂ) | |
18 | 16, 17 | syl 17 | . . . . . . . . 9 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝐴 − 𝐵) ∈ ℂ) |
19 | 2ne0 11326 | . . . . . . . . . 10 ⊢ 2 ≠ 0 | |
20 | 19 | a1i 11 | . . . . . . . . 9 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → 2 ≠ 0) |
21 | 18, 13, 20 | divcan3d 11019 | . . . . . . . 8 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((2 · (𝐴 − 𝐵)) / 2) = (𝐴 − 𝐵)) |
22 | 15, 21 | eqtr3d 2797 | . . . . . . 7 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (((2 · 𝐴) − (2 · 𝐵)) / 2) = (𝐴 − 𝐵)) |
23 | 22, 16 | eqeltrd 2840 | . . . . . 6 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (((2 · 𝐴) − (2 · 𝐵)) / 2) ∈ ℤ) |
24 | oveq1 6822 | . . . . . . 7 ⊢ (((2 · 𝐴) − (2 · 𝐵)) = 1 → (((2 · 𝐴) − (2 · 𝐵)) / 2) = (1 / 2)) | |
25 | 24 | eleq1d 2825 | . . . . . 6 ⊢ (((2 · 𝐴) − (2 · 𝐵)) = 1 → ((((2 · 𝐴) − (2 · 𝐵)) / 2) ∈ ℤ ↔ (1 / 2) ∈ ℤ)) |
26 | 23, 25 | syl5ibcom 235 | . . . . 5 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (((2 · 𝐴) − (2 · 𝐵)) = 1 → (1 / 2) ∈ ℤ)) |
27 | 12, 26 | sylbird 250 | . . . 4 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (((2 · 𝐵) + 1) = (2 · 𝐴) → (1 / 2) ∈ ℤ)) |
28 | 27 | necon3bd 2947 | . . 3 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (¬ (1 / 2) ∈ ℤ → ((2 · 𝐵) + 1) ≠ (2 · 𝐴))) |
29 | 1, 28 | mpi 20 | . 2 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → ((2 · 𝐵) + 1) ≠ (2 · 𝐴)) |
30 | 29 | necomd 2988 | 1 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (2 · 𝐴) ≠ ((2 · 𝐵) + 1)) |
Colors of variables: wff setvar class |
Syntax hints: ¬ wn 3 → wi 4 ∧ wa 383 = wceq 1632 ∈ wcel 2140 ≠ wne 2933 (class class class)co 6815 ℂcc 10147 0cc0 10149 1c1 10150 + caddc 10152 · cmul 10154 − cmin 10479 / cdiv 10897 2c2 11283 ℤcz 11590 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1871 ax-4 1886 ax-5 1989 ax-6 2055 ax-7 2091 ax-8 2142 ax-9 2149 ax-10 2169 ax-11 2184 ax-12 2197 ax-13 2392 ax-ext 2741 ax-sep 4934 ax-nul 4942 ax-pow 4993 ax-pr 5056 ax-un 7116 ax-resscn 10206 ax-1cn 10207 ax-icn 10208 ax-addcl 10209 ax-addrcl 10210 ax-mulcl 10211 ax-mulrcl 10212 ax-mulcom 10213 ax-addass 10214 ax-mulass 10215 ax-distr 10216 ax-i2m1 10217 ax-1ne0 10218 ax-1rid 10219 ax-rnegex 10220 ax-rrecex 10221 ax-cnre 10222 ax-pre-lttri 10223 ax-pre-lttrn 10224 ax-pre-ltadd 10225 ax-pre-mulgt0 10226 |
This theorem depends on definitions: df-bi 197 df-or 384 df-an 385 df-3or 1073 df-3an 1074 df-tru 1635 df-ex 1854 df-nf 1859 df-sb 2048 df-eu 2612 df-mo 2613 df-clab 2748 df-cleq 2754 df-clel 2757 df-nfc 2892 df-ne 2934 df-nel 3037 df-ral 3056 df-rex 3057 df-reu 3058 df-rmo 3059 df-rab 3060 df-v 3343 df-sbc 3578 df-csb 3676 df-dif 3719 df-un 3721 df-in 3723 df-ss 3730 df-pss 3732 df-nul 4060 df-if 4232 df-pw 4305 df-sn 4323 df-pr 4325 df-tp 4327 df-op 4329 df-uni 4590 df-iun 4675 df-br 4806 df-opab 4866 df-mpt 4883 df-tr 4906 df-id 5175 df-eprel 5180 df-po 5188 df-so 5189 df-fr 5226 df-we 5228 df-xp 5273 df-rel 5274 df-cnv 5275 df-co 5276 df-dm 5277 df-rn 5278 df-res 5279 df-ima 5280 df-pred 5842 df-ord 5888 df-on 5889 df-lim 5890 df-suc 5891 df-iota 6013 df-fun 6052 df-fn 6053 df-f 6054 df-f1 6055 df-fo 6056 df-f1o 6057 df-fv 6058 df-riota 6776 df-ov 6818 df-oprab 6819 df-mpt2 6820 df-om 7233 df-wrecs 7578 df-recs 7639 df-rdg 7677 df-er 7914 df-en 8125 df-dom 8126 df-sdom 8127 df-pnf 10289 df-mnf 10290 df-xr 10291 df-ltxr 10292 df-le 10293 df-sub 10481 df-neg 10482 df-div 10898 df-nn 11234 df-2 11292 df-n0 11506 df-z 11591 |
This theorem is referenced by: nneo 11674 zeo2 11677 znnenlem 15160 |
Copyright terms: Public domain | W3C validator |