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

Theorem ltoddhalfle 14923
Description: An integer is less than half of an odd number iff it is less than or equal to the half of the predecessor of the odd number (which is an even number). (Contributed by AV, 29-Jun-2021.)
Assertion
Ref Expression
ltoddhalfle ((𝑁 ∈ ℤ ∧ ¬ 2 ∥ 𝑁𝑀 ∈ ℤ) → (𝑀 < (𝑁 / 2) ↔ 𝑀 ≤ ((𝑁 − 1) / 2)))

Proof of Theorem ltoddhalfle
Dummy variable 𝑛 is distinct from all other variables.
StepHypRef Expression
1 odd2np1 14903 . . 3 (𝑁 ∈ ℤ → (¬ 2 ∥ 𝑁 ↔ ∃𝑛 ∈ ℤ ((2 · 𝑛) + 1) = 𝑁))
2 halfre 11123 . . . . . . . . . . . . . . . 16 (1 / 2) ∈ ℝ
32a1i 11 . . . . . . . . . . . . . . 15 (𝑛 ∈ ℤ → (1 / 2) ∈ ℝ)
4 1red 9934 . . . . . . . . . . . . . . 15 (𝑛 ∈ ℤ → 1 ∈ ℝ)
5 zre 11258 . . . . . . . . . . . . . . 15 (𝑛 ∈ ℤ → 𝑛 ∈ ℝ)
63, 4, 53jca 1235 . . . . . . . . . . . . . 14 (𝑛 ∈ ℤ → ((1 / 2) ∈ ℝ ∧ 1 ∈ ℝ ∧ 𝑛 ∈ ℝ))
76adantr 480 . . . . . . . . . . . . 13 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → ((1 / 2) ∈ ℝ ∧ 1 ∈ ℝ ∧ 𝑛 ∈ ℝ))
8 halflt1 11127 . . . . . . . . . . . . 13 (1 / 2) < 1
9 axltadd 9990 . . . . . . . . . . . . 13 (((1 / 2) ∈ ℝ ∧ 1 ∈ ℝ ∧ 𝑛 ∈ ℝ) → ((1 / 2) < 1 → (𝑛 + (1 / 2)) < (𝑛 + 1)))
107, 8, 9mpisyl 21 . . . . . . . . . . . 12 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑛 + (1 / 2)) < (𝑛 + 1))
11 zre 11258 . . . . . . . . . . . . . 14 (𝑀 ∈ ℤ → 𝑀 ∈ ℝ)
1211adantl 481 . . . . . . . . . . . . 13 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → 𝑀 ∈ ℝ)
135, 3readdcld 9948 . . . . . . . . . . . . . 14 (𝑛 ∈ ℤ → (𝑛 + (1 / 2)) ∈ ℝ)
1413adantr 480 . . . . . . . . . . . . 13 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑛 + (1 / 2)) ∈ ℝ)
15 peano2z 11295 . . . . . . . . . . . . . . 15 (𝑛 ∈ ℤ → (𝑛 + 1) ∈ ℤ)
1615zred 11358 . . . . . . . . . . . . . 14 (𝑛 ∈ ℤ → (𝑛 + 1) ∈ ℝ)
1716adantr 480 . . . . . . . . . . . . 13 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑛 + 1) ∈ ℝ)
18 lttr 9993 . . . . . . . . . . . . 13 ((𝑀 ∈ ℝ ∧ (𝑛 + (1 / 2)) ∈ ℝ ∧ (𝑛 + 1) ∈ ℝ) → ((𝑀 < (𝑛 + (1 / 2)) ∧ (𝑛 + (1 / 2)) < (𝑛 + 1)) → 𝑀 < (𝑛 + 1)))
1912, 14, 17, 18syl3anc 1318 . . . . . . . . . . . 12 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → ((𝑀 < (𝑛 + (1 / 2)) ∧ (𝑛 + (1 / 2)) < (𝑛 + 1)) → 𝑀 < (𝑛 + 1)))
2010, 19mpan2d 706 . . . . . . . . . . 11 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑀 < (𝑛 + (1 / 2)) → 𝑀 < (𝑛 + 1)))
21 zleltp1 11305 . . . . . . . . . . . 12 ((𝑀 ∈ ℤ ∧ 𝑛 ∈ ℤ) → (𝑀𝑛𝑀 < (𝑛 + 1)))
2221ancoms 468 . . . . . . . . . . 11 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑀𝑛𝑀 < (𝑛 + 1)))
2320, 22sylibrd 248 . . . . . . . . . 10 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑀 < (𝑛 + (1 / 2)) → 𝑀𝑛))
24 halfgt0 11125 . . . . . . . . . . . 12 0 < (1 / 2)
253, 5jca 553 . . . . . . . . . . . . . 14 (𝑛 ∈ ℤ → ((1 / 2) ∈ ℝ ∧ 𝑛 ∈ ℝ))
2625adantr 480 . . . . . . . . . . . . 13 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → ((1 / 2) ∈ ℝ ∧ 𝑛 ∈ ℝ))
27 ltaddpos 10397 . . . . . . . . . . . . 13 (((1 / 2) ∈ ℝ ∧ 𝑛 ∈ ℝ) → (0 < (1 / 2) ↔ 𝑛 < (𝑛 + (1 / 2))))
2826, 27syl 17 . . . . . . . . . . . 12 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (0 < (1 / 2) ↔ 𝑛 < (𝑛 + (1 / 2))))
2924, 28mpbii 222 . . . . . . . . . . 11 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → 𝑛 < (𝑛 + (1 / 2)))
305adantr 480 . . . . . . . . . . . 12 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → 𝑛 ∈ ℝ)
31 lelttr 10007 . . . . . . . . . . . 12 ((𝑀 ∈ ℝ ∧ 𝑛 ∈ ℝ ∧ (𝑛 + (1 / 2)) ∈ ℝ) → ((𝑀𝑛𝑛 < (𝑛 + (1 / 2))) → 𝑀 < (𝑛 + (1 / 2))))
3212, 30, 14, 31syl3anc 1318 . . . . . . . . . . 11 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → ((𝑀𝑛𝑛 < (𝑛 + (1 / 2))) → 𝑀 < (𝑛 + (1 / 2))))
3329, 32mpan2d 706 . . . . . . . . . 10 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑀𝑛𝑀 < (𝑛 + (1 / 2))))
3423, 33impbid 201 . . . . . . . . 9 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑀 < (𝑛 + (1 / 2)) ↔ 𝑀𝑛))
35 zcn 11259 . . . . . . . . . . . 12 (𝑛 ∈ ℤ → 𝑛 ∈ ℂ)
36 1cnd 9935 . . . . . . . . . . . 12 (𝑛 ∈ ℤ → 1 ∈ ℂ)
37 2cnne0 11119 . . . . . . . . . . . . 13 (2 ∈ ℂ ∧ 2 ≠ 0)
3837a1i 11 . . . . . . . . . . . 12 (𝑛 ∈ ℤ → (2 ∈ ℂ ∧ 2 ≠ 0))
39 muldivdir 10599 . . . . . . . . . . . 12 ((𝑛 ∈ ℂ ∧ 1 ∈ ℂ ∧ (2 ∈ ℂ ∧ 2 ≠ 0)) → (((2 · 𝑛) + 1) / 2) = (𝑛 + (1 / 2)))
4035, 36, 38, 39syl3anc 1318 . . . . . . . . . . 11 (𝑛 ∈ ℤ → (((2 · 𝑛) + 1) / 2) = (𝑛 + (1 / 2)))
4140breq2d 4595 . . . . . . . . . 10 (𝑛 ∈ ℤ → (𝑀 < (((2 · 𝑛) + 1) / 2) ↔ 𝑀 < (𝑛 + (1 / 2))))
4241adantr 480 . . . . . . . . 9 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑀 < (((2 · 𝑛) + 1) / 2) ↔ 𝑀 < (𝑛 + (1 / 2))))
43 2z 11286 . . . . . . . . . . . . . . . . 17 2 ∈ ℤ
4443a1i 11 . . . . . . . . . . . . . . . 16 (𝑛 ∈ ℤ → 2 ∈ ℤ)
45 id 22 . . . . . . . . . . . . . . . 16 (𝑛 ∈ ℤ → 𝑛 ∈ ℤ)
4644, 45zmulcld 11364 . . . . . . . . . . . . . . 15 (𝑛 ∈ ℤ → (2 · 𝑛) ∈ ℤ)
4746zcnd 11359 . . . . . . . . . . . . . 14 (𝑛 ∈ ℤ → (2 · 𝑛) ∈ ℂ)
4847adantr 480 . . . . . . . . . . . . 13 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (2 · 𝑛) ∈ ℂ)
49 pncan1 10333 . . . . . . . . . . . . 13 ((2 · 𝑛) ∈ ℂ → (((2 · 𝑛) + 1) − 1) = (2 · 𝑛))
5048, 49syl 17 . . . . . . . . . . . 12 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (((2 · 𝑛) + 1) − 1) = (2 · 𝑛))
5150oveq1d 6564 . . . . . . . . . . 11 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → ((((2 · 𝑛) + 1) − 1) / 2) = ((2 · 𝑛) / 2))
52 2cnd 10970 . . . . . . . . . . . . 13 (𝑛 ∈ ℤ → 2 ∈ ℂ)
53 2ne0 10990 . . . . . . . . . . . . . 14 2 ≠ 0
5453a1i 11 . . . . . . . . . . . . 13 (𝑛 ∈ ℤ → 2 ≠ 0)
5535, 52, 54divcan3d 10685 . . . . . . . . . . . 12 (𝑛 ∈ ℤ → ((2 · 𝑛) / 2) = 𝑛)
5655adantr 480 . . . . . . . . . . 11 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → ((2 · 𝑛) / 2) = 𝑛)
5751, 56eqtrd 2644 . . . . . . . . . 10 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → ((((2 · 𝑛) + 1) − 1) / 2) = 𝑛)
5857breq2d 4595 . . . . . . . . 9 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑀 ≤ ((((2 · 𝑛) + 1) − 1) / 2) ↔ 𝑀𝑛))
5934, 42, 583bitr4d 299 . . . . . . . 8 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝑀 < (((2 · 𝑛) + 1) / 2) ↔ 𝑀 ≤ ((((2 · 𝑛) + 1) − 1) / 2)))
60 oveq1 6556 . . . . . . . . . 10 (((2 · 𝑛) + 1) = 𝑁 → (((2 · 𝑛) + 1) / 2) = (𝑁 / 2))
6160breq2d 4595 . . . . . . . . 9 (((2 · 𝑛) + 1) = 𝑁 → (𝑀 < (((2 · 𝑛) + 1) / 2) ↔ 𝑀 < (𝑁 / 2)))
62 oveq1 6556 . . . . . . . . . . 11 (((2 · 𝑛) + 1) = 𝑁 → (((2 · 𝑛) + 1) − 1) = (𝑁 − 1))
6362oveq1d 6564 . . . . . . . . . 10 (((2 · 𝑛) + 1) = 𝑁 → ((((2 · 𝑛) + 1) − 1) / 2) = ((𝑁 − 1) / 2))
6463breq2d 4595 . . . . . . . . 9 (((2 · 𝑛) + 1) = 𝑁 → (𝑀 ≤ ((((2 · 𝑛) + 1) − 1) / 2) ↔ 𝑀 ≤ ((𝑁 − 1) / 2)))
6561, 64bibi12d 334 . . . . . . . 8 (((2 · 𝑛) + 1) = 𝑁 → ((𝑀 < (((2 · 𝑛) + 1) / 2) ↔ 𝑀 ≤ ((((2 · 𝑛) + 1) − 1) / 2)) ↔ (𝑀 < (𝑁 / 2) ↔ 𝑀 ≤ ((𝑁 − 1) / 2))))
6659, 65syl5ibcom 234 . . . . . . 7 ((𝑛 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (((2 · 𝑛) + 1) = 𝑁 → (𝑀 < (𝑁 / 2) ↔ 𝑀 ≤ ((𝑁 − 1) / 2))))
6766ex 449 . . . . . 6 (𝑛 ∈ ℤ → (𝑀 ∈ ℤ → (((2 · 𝑛) + 1) = 𝑁 → (𝑀 < (𝑁 / 2) ↔ 𝑀 ≤ ((𝑁 − 1) / 2)))))
6867adantl 481 . . . . 5 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → (𝑀 ∈ ℤ → (((2 · 𝑛) + 1) = 𝑁 → (𝑀 < (𝑁 / 2) ↔ 𝑀 ≤ ((𝑁 − 1) / 2)))))
6968com23 84 . . . 4 ((𝑁 ∈ ℤ ∧ 𝑛 ∈ ℤ) → (((2 · 𝑛) + 1) = 𝑁 → (𝑀 ∈ ℤ → (𝑀 < (𝑁 / 2) ↔ 𝑀 ≤ ((𝑁 − 1) / 2)))))
7069rexlimdva 3013 . . 3 (𝑁 ∈ ℤ → (∃𝑛 ∈ ℤ ((2 · 𝑛) + 1) = 𝑁 → (𝑀 ∈ ℤ → (𝑀 < (𝑁 / 2) ↔ 𝑀 ≤ ((𝑁 − 1) / 2)))))
711, 70sylbid 229 . 2 (𝑁 ∈ ℤ → (¬ 2 ∥ 𝑁 → (𝑀 ∈ ℤ → (𝑀 < (𝑁 / 2) ↔ 𝑀 ≤ ((𝑁 − 1) / 2)))))
72713imp 1249 1 ((𝑁 ∈ ℤ ∧ ¬ 2 ∥ 𝑁𝑀 ∈ ℤ) → (𝑀 < (𝑁 / 2) ↔ 𝑀 ≤ ((𝑁 − 1) / 2)))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wb 195  wa 383  w3a 1031   = wceq 1475  wcel 1977  wne 2780  wrex 2897   class class class wbr 4583  (class class class)co 6549  cc 9813  cr 9814  0cc0 9815  1c1 9816   + caddc 9818   · cmul 9820   < clt 9953  cle 9954  cmin 10145   / cdiv 10563  2c2 10947  cz 11254  cdvds 14821
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1713  ax-4 1728  ax-5 1827  ax-6 1875  ax-7 1922  ax-8 1979  ax-9 1986  ax-10 2006  ax-11 2021  ax-12 2034  ax-13 2234  ax-ext 2590  ax-sep 4709  ax-nul 4717  ax-pow 4769  ax-pr 4833  ax-un 6847  ax-resscn 9872  ax-1cn 9873  ax-icn 9874  ax-addcl 9875  ax-addrcl 9876  ax-mulcl 9877  ax-mulrcl 9878  ax-mulcom 9879  ax-addass 9880  ax-mulass 9881  ax-distr 9882  ax-i2m1 9883  ax-1ne0 9884  ax-1rid 9885  ax-rnegex 9886  ax-rrecex 9887  ax-cnre 9888  ax-pre-lttri 9889  ax-pre-lttrn 9890  ax-pre-ltadd 9891  ax-pre-mulgt0 9892
This theorem depends on definitions:  df-bi 196  df-or 384  df-an 385  df-3or 1032  df-3an 1033  df-tru 1478  df-ex 1696  df-nf 1701  df-sb 1868  df-eu 2462  df-mo 2463  df-clab 2597  df-cleq 2603  df-clel 2606  df-nfc 2740  df-ne 2782  df-nel 2783  df-ral 2901  df-rex 2902  df-reu 2903  df-rmo 2904  df-rab 2905  df-v 3175  df-sbc 3403  df-csb 3500  df-dif 3543  df-un 3545  df-in 3547  df-ss 3554  df-pss 3556  df-nul 3875  df-if 4037  df-pw 4110  df-sn 4126  df-pr 4128  df-tp 4130  df-op 4132  df-uni 4373  df-iun 4457  df-br 4584  df-opab 4644  df-mpt 4645  df-tr 4681  df-eprel 4949  df-id 4953  df-po 4959  df-so 4960  df-fr 4997  df-we 4999  df-xp 5044  df-rel 5045  df-cnv 5046  df-co 5047  df-dm 5048  df-rn 5049  df-res 5050  df-ima 5051  df-pred 5597  df-ord 5643  df-on 5644  df-lim 5645  df-suc 5646  df-iota 5768  df-fun 5806  df-fn 5807  df-f 5808  df-f1 5809  df-fo 5810  df-f1o 5811  df-fv 5812  df-riota 6511  df-ov 6552  df-oprab 6553  df-mpt2 6554  df-om 6958  df-wrecs 7294  df-recs 7355  df-rdg 7393  df-er 7629  df-en 7842  df-dom 7843  df-sdom 7844  df-pnf 9955  df-mnf 9956  df-xr 9957  df-ltxr 9958  df-le 9959  df-sub 10147  df-neg 10148  df-div 10564  df-nn 10898  df-2 10956  df-n0 11170  df-z 11255  df-dvds 14822
This theorem is referenced by:  gausslemma2dlem1a  24890
  Copyright terms: Public domain W3C validator