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

Theorem elfzp1b 13075
Description: An integer is a member of a 0-based finite set of sequential integers iff its successor is a member of the corresponding 1-based set. (Contributed by Paul Chapman, 22-Jun-2011.)
Assertion
Ref Expression
elfzp1b ((𝐾 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∈ (0...(𝑁 − 1)) ↔ (𝐾 + 1) ∈ (1...𝑁)))

Proof of Theorem elfzp1b
StepHypRef Expression
1 peano2z 12104 . . . 4 (𝐾 ∈ ℤ → (𝐾 + 1) ∈ ℤ)
2 1z 12093 . . . . 5 1 ∈ ℤ
3 fzsubel 13034 . . . . . 6 (((1 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ((𝐾 + 1) ∈ ℤ ∧ 1 ∈ ℤ)) → ((𝐾 + 1) ∈ (1...𝑁) ↔ ((𝐾 + 1) − 1) ∈ ((1 − 1)...(𝑁 − 1))))
42, 3mpanl1 700 . . . . 5 ((𝑁 ∈ ℤ ∧ ((𝐾 + 1) ∈ ℤ ∧ 1 ∈ ℤ)) → ((𝐾 + 1) ∈ (1...𝑁) ↔ ((𝐾 + 1) − 1) ∈ ((1 − 1)...(𝑁 − 1))))
52, 4mpanr2 704 . . . 4 ((𝑁 ∈ ℤ ∧ (𝐾 + 1) ∈ ℤ) → ((𝐾 + 1) ∈ (1...𝑁) ↔ ((𝐾 + 1) − 1) ∈ ((1 − 1)...(𝑁 − 1))))
61, 5sylan2 596 . . 3 ((𝑁 ∈ ℤ ∧ 𝐾 ∈ ℤ) → ((𝐾 + 1) ∈ (1...𝑁) ↔ ((𝐾 + 1) − 1) ∈ ((1 − 1)...(𝑁 − 1))))
76ancoms 462 . 2 ((𝐾 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 + 1) ∈ (1...𝑁) ↔ ((𝐾 + 1) − 1) ∈ ((1 − 1)...(𝑁 − 1))))
8 zcn 12067 . . . . 5 (𝐾 ∈ ℤ → 𝐾 ∈ ℂ)
9 ax-1cn 10673 . . . . 5 1 ∈ ℂ
10 pncan 10970 . . . . 5 ((𝐾 ∈ ℂ ∧ 1 ∈ ℂ) → ((𝐾 + 1) − 1) = 𝐾)
118, 9, 10sylancl 589 . . . 4 (𝐾 ∈ ℤ → ((𝐾 + 1) − 1) = 𝐾)
12 1m1e0 11788 . . . . . 6 (1 − 1) = 0
1312oveq1i 7180 . . . . 5 ((1 − 1)...(𝑁 − 1)) = (0...(𝑁 − 1))
1413a1i 11 . . . 4 (𝐾 ∈ ℤ → ((1 − 1)...(𝑁 − 1)) = (0...(𝑁 − 1)))
1511, 14eleq12d 2827 . . 3 (𝐾 ∈ ℤ → (((𝐾 + 1) − 1) ∈ ((1 − 1)...(𝑁 − 1)) ↔ 𝐾 ∈ (0...(𝑁 − 1))))
1615adantr 484 . 2 ((𝐾 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (((𝐾 + 1) − 1) ∈ ((1 − 1)...(𝑁 − 1)) ↔ 𝐾 ∈ (0...(𝑁 − 1))))
177, 16bitr2d 283 1 ((𝐾 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∈ (0...(𝑁 − 1)) ↔ (𝐾 + 1) ∈ (1...𝑁)))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 209  wa 399   = wceq 1542  wcel 2114  (class class class)co 7170  cc 10613  0cc0 10615  1c1 10616   + caddc 10618  cmin 10948  cz 12062  ...cfz 12981
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1802  ax-4 1816  ax-5 1917  ax-6 1975  ax-7 2020  ax-8 2116  ax-9 2124  ax-10 2145  ax-11 2162  ax-12 2179  ax-ext 2710  ax-sep 5167  ax-nul 5174  ax-pow 5232  ax-pr 5296  ax-un 7479  ax-cnex 10671  ax-resscn 10672  ax-1cn 10673  ax-icn 10674  ax-addcl 10675  ax-addrcl 10676  ax-mulcl 10677  ax-mulrcl 10678  ax-mulcom 10679  ax-addass 10680  ax-mulass 10681  ax-distr 10682  ax-i2m1 10683  ax-1ne0 10684  ax-1rid 10685  ax-rnegex 10686  ax-rrecex 10687  ax-cnre 10688  ax-pre-lttri 10689  ax-pre-lttrn 10690  ax-pre-ltadd 10691  ax-pre-mulgt0 10692
This theorem depends on definitions:  df-bi 210  df-an 400  df-or 847  df-3or 1089  df-3an 1090  df-tru 1545  df-fal 1555  df-ex 1787  df-nf 1791  df-sb 2075  df-mo 2540  df-eu 2570  df-clab 2717  df-cleq 2730  df-clel 2811  df-nfc 2881  df-ne 2935  df-nel 3039  df-ral 3058  df-rex 3059  df-reu 3060  df-rab 3062  df-v 3400  df-sbc 3681  df-csb 3791  df-dif 3846  df-un 3848  df-in 3850  df-ss 3860  df-pss 3862  df-nul 4212  df-if 4415  df-pw 4490  df-sn 4517  df-pr 4519  df-tp 4521  df-op 4523  df-uni 4797  df-iun 4883  df-br 5031  df-opab 5093  df-mpt 5111  df-tr 5137  df-id 5429  df-eprel 5434  df-po 5442  df-so 5443  df-fr 5483  df-we 5485  df-xp 5531  df-rel 5532  df-cnv 5533  df-co 5534  df-dm 5535  df-rn 5536  df-res 5537  df-ima 5538  df-pred 6129  df-ord 6175  df-on 6176  df-lim 6177  df-suc 6178  df-iota 6297  df-fun 6341  df-fn 6342  df-f 6343  df-f1 6344  df-fo 6345  df-f1o 6346  df-fv 6347  df-riota 7127  df-ov 7173  df-oprab 7174  df-mpo 7175  df-om 7600  df-wrecs 7976  df-recs 8037  df-rdg 8075  df-er 8320  df-en 8556  df-dom 8557  df-sdom 8558  df-pnf 10755  df-mnf 10756  df-xr 10757  df-ltxr 10758  df-le 10759  df-sub 10950  df-neg 10951  df-nn 11717  df-n0 11977  df-z 12063  df-fz 12982
This theorem is referenced by:  numclwlk2lem2f  28314
  Copyright terms: Public domain W3C validator