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

Theorem vdwapfval 17020
Description: Define the arithmetic progression function, which takes as input a length 𝑘, a start point 𝑎, and a step 𝑑 and outputs the set of points in this progression. (Contributed by Mario Carneiro, 18-Aug-2014.)
Assertion
Ref Expression
vdwapfval (𝐾 ∈ ℕ0 → (AP‘𝐾) = (𝑎 ∈ ℕ, 𝑑 ∈ ℕ ↦ ran (𝑚 ∈ (0...(𝐾 − 1)) ↦ (𝑎 + (𝑚 · 𝑑)))))
Distinct variable group:   𝑎,𝑑,𝑚,𝐾

Proof of Theorem vdwapfval
Dummy variable 𝑘 is distinct from all other variables.
StepHypRef Expression
1 simp1 1136 . . . . . . 7 ((𝑘 = 𝐾𝑎 ∈ ℕ ∧ 𝑑 ∈ ℕ) → 𝑘 = 𝐾)
21oveq1d 7465 . . . . . 6 ((𝑘 = 𝐾𝑎 ∈ ℕ ∧ 𝑑 ∈ ℕ) → (𝑘 − 1) = (𝐾 − 1))
32oveq2d 7466 . . . . 5 ((𝑘 = 𝐾𝑎 ∈ ℕ ∧ 𝑑 ∈ ℕ) → (0...(𝑘 − 1)) = (0...(𝐾 − 1)))
43mpteq1d 5261 . . . 4 ((𝑘 = 𝐾𝑎 ∈ ℕ ∧ 𝑑 ∈ ℕ) → (𝑚 ∈ (0...(𝑘 − 1)) ↦ (𝑎 + (𝑚 · 𝑑))) = (𝑚 ∈ (0...(𝐾 − 1)) ↦ (𝑎 + (𝑚 · 𝑑))))
54rneqd 5963 . . 3 ((𝑘 = 𝐾𝑎 ∈ ℕ ∧ 𝑑 ∈ ℕ) → ran (𝑚 ∈ (0...(𝑘 − 1)) ↦ (𝑎 + (𝑚 · 𝑑))) = ran (𝑚 ∈ (0...(𝐾 − 1)) ↦ (𝑎 + (𝑚 · 𝑑))))
65mpoeq3dva 7529 . 2 (𝑘 = 𝐾 → (𝑎 ∈ ℕ, 𝑑 ∈ ℕ ↦ ran (𝑚 ∈ (0...(𝑘 − 1)) ↦ (𝑎 + (𝑚 · 𝑑)))) = (𝑎 ∈ ℕ, 𝑑 ∈ ℕ ↦ ran (𝑚 ∈ (0...(𝐾 − 1)) ↦ (𝑎 + (𝑚 · 𝑑)))))
7 df-vdwap 17017 . 2 AP = (𝑘 ∈ ℕ0 ↦ (𝑎 ∈ ℕ, 𝑑 ∈ ℕ ↦ ran (𝑚 ∈ (0...(𝑘 − 1)) ↦ (𝑎 + (𝑚 · 𝑑)))))
8 nnex 12301 . . 3 ℕ ∈ V
98, 8mpoex 8122 . 2 (𝑎 ∈ ℕ, 𝑑 ∈ ℕ ↦ ran (𝑚 ∈ (0...(𝐾 − 1)) ↦ (𝑎 + (𝑚 · 𝑑)))) ∈ V
106, 7, 9fvmpt 7031 1 (𝐾 ∈ ℕ0 → (AP‘𝐾) = (𝑎 ∈ ℕ, 𝑑 ∈ ℕ ↦ ran (𝑚 ∈ (0...(𝐾 − 1)) ↦ (𝑎 + (𝑚 · 𝑑)))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  w3a 1087   = wceq 1537  wcel 2108  cmpt 5249  ran crn 5701  cfv 6575  (class class class)co 7450  cmpo 7452  0cc0 11186  1c1 11187   + caddc 11189   · cmul 11191  cmin 11522  cn 12295  0cn0 12555  ...cfz 13569  APcvdwa 17014
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1793  ax-4 1807  ax-5 1909  ax-6 1967  ax-7 2007  ax-8 2110  ax-9 2118  ax-10 2141  ax-11 2158  ax-12 2178  ax-ext 2711  ax-rep 5303  ax-sep 5317  ax-nul 5324  ax-pow 5383  ax-pr 5447  ax-un 7772  ax-cnex 11242  ax-1cn 11244  ax-addcl 11246
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 847  df-3or 1088  df-3an 1089  df-tru 1540  df-fal 1550  df-ex 1778  df-nf 1782  df-sb 2065  df-mo 2543  df-eu 2572  df-clab 2718  df-cleq 2732  df-clel 2819  df-nfc 2895  df-ne 2947  df-ral 3068  df-rex 3077  df-reu 3389  df-rab 3444  df-v 3490  df-sbc 3805  df-csb 3922  df-dif 3979  df-un 3981  df-in 3983  df-ss 3993  df-pss 3996  df-nul 4353  df-if 4549  df-pw 4624  df-sn 4649  df-pr 4651  df-op 4655  df-uni 4932  df-iun 5017  df-br 5167  df-opab 5229  df-mpt 5250  df-tr 5284  df-id 5593  df-eprel 5599  df-po 5607  df-so 5608  df-fr 5652  df-we 5654  df-xp 5706  df-rel 5707  df-cnv 5708  df-co 5709  df-dm 5710  df-rn 5711  df-res 5712  df-ima 5713  df-pred 6334  df-ord 6400  df-on 6401  df-lim 6402  df-suc 6403  df-iota 6527  df-fun 6577  df-fn 6578  df-f 6579  df-f1 6580  df-fo 6581  df-f1o 6582  df-fv 6583  df-ov 7453  df-oprab 7454  df-mpo 7455  df-om 7906  df-1st 8032  df-2nd 8033  df-frecs 8324  df-wrecs 8355  df-recs 8429  df-rdg 8468  df-nn 12296  df-vdwap 17017
This theorem is referenced by:  vdwapf  17021  vdwapval  17022
  Copyright terms: Public domain W3C validator