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

Theorem modval 12532
Description: The value of the modulo operation. The modulo congruence notation of number theory, 𝐽𝐾 (modulo 𝑁), can be expressed in our notation as (𝐽 mod 𝑁) = (𝐾 mod 𝑁). Definition 1 in Knuth, The Art of Computer Programming, Vol. I (1972), p. 38. Knuth uses "mod" for the operation and "modulo" for the congruence. Unlike Knuth, we restrict the second argument to positive reals to simplify certain theorems. (This also gives us future flexibility to extend it to any one of several different conventions for a zero or negative second argument, should there be an advantage in doing so.) (Contributed by NM, 10-Nov-2008.) (Revised by Mario Carneiro, 3-Nov-2013.)
Assertion
Ref Expression
modval ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ+) → (𝐴 mod 𝐵) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))))

Proof of Theorem modval
Dummy variables 𝑥 𝑦 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 oveq1 6556 . . . . 5 (𝑥 = 𝐴 → (𝑥 / 𝑦) = (𝐴 / 𝑦))
21fveq2d 6107 . . . 4 (𝑥 = 𝐴 → (⌊‘(𝑥 / 𝑦)) = (⌊‘(𝐴 / 𝑦)))
32oveq2d 6565 . . 3 (𝑥 = 𝐴 → (𝑦 · (⌊‘(𝑥 / 𝑦))) = (𝑦 · (⌊‘(𝐴 / 𝑦))))
4 oveq12 6558 . . 3 ((𝑥 = 𝐴 ∧ (𝑦 · (⌊‘(𝑥 / 𝑦))) = (𝑦 · (⌊‘(𝐴 / 𝑦)))) → (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))) = (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦)))))
53, 4mpdan 699 . 2 (𝑥 = 𝐴 → (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))) = (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦)))))
6 oveq2 6557 . . . . 5 (𝑦 = 𝐵 → (𝐴 / 𝑦) = (𝐴 / 𝐵))
76fveq2d 6107 . . . 4 (𝑦 = 𝐵 → (⌊‘(𝐴 / 𝑦)) = (⌊‘(𝐴 / 𝐵)))
8 oveq12 6558 . . . 4 ((𝑦 = 𝐵 ∧ (⌊‘(𝐴 / 𝑦)) = (⌊‘(𝐴 / 𝐵))) → (𝑦 · (⌊‘(𝐴 / 𝑦))) = (𝐵 · (⌊‘(𝐴 / 𝐵))))
97, 8mpdan 699 . . 3 (𝑦 = 𝐵 → (𝑦 · (⌊‘(𝐴 / 𝑦))) = (𝐵 · (⌊‘(𝐴 / 𝐵))))
109oveq2d 6565 . 2 (𝑦 = 𝐵 → (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦)))) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))))
11 df-mod 12531 . 2 mod = (𝑥 ∈ ℝ, 𝑦 ∈ ℝ+ ↦ (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))))
12 ovex 6577 . 2 (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))) ∈ V
135, 10, 11, 12ovmpt2 6694 1 ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ+) → (𝐴 mod 𝐵) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wa 383   = wceq 1475  wcel 1977  cfv 5804  (class class class)co 6549  cr 9814   · cmul 9820  cmin 10145   / cdiv 10563  +crp 11708  cfl 12453   mod cmo 12530
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-9 1986  ax-10 2006  ax-11 2021  ax-12 2034  ax-13 2234  ax-ext 2590  ax-sep 4709  ax-nul 4717  ax-pr 4833
This theorem depends on definitions:  df-bi 196  df-or 384  df-an 385  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-ral 2901  df-rex 2902  df-rab 2905  df-v 3175  df-sbc 3403  df-dif 3543  df-un 3545  df-in 3547  df-ss 3554  df-nul 3875  df-if 4037  df-sn 4126  df-pr 4128  df-op 4132  df-uni 4373  df-br 4584  df-opab 4644  df-id 4953  df-xp 5044  df-rel 5045  df-cnv 5046  df-co 5047  df-dm 5048  df-iota 5768  df-fun 5806  df-fv 5812  df-ov 6552  df-oprab 6553  df-mpt2 6554  df-mod 12531
This theorem is referenced by:  modvalr  12533  modcl  12534  mod0  12537  modge0  12540  modlt  12541  moddiffl  12543  modfrac  12545  modmulnn  12550  zmodcl  12552  modid  12557  modcyc  12567  modadd1  12569  modmul1  12585  moddi  12600  modsubdir  12601  modirr  12603  iexpcyc  12831  digit2  12859  dvdsmod  14888  divalgmod  14967  divalgmodOLD  14968  modgcd  15091  bezoutlem3  15096  prmdiv  15328  odzdvds  15338  fldivp1  15439  mulgmodid  17404  odmodnn0  17782  odmod  17788  gexdvds  17822  zringlpirlem3  19653  sineq0  24077  efif1olem2  24093  lgseisenlem4  24903  dchrisumlem1  24978  ostth2lem2  25123  sineq0ALT  38195  ltmod  38705  fourierswlem  39123
  Copyright terms: Public domain W3C validator