ILE Home Intuitionistic Logic Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  ILE Home  >  Th. List  >  modqval GIF version

Theorem modqval 10355
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 numbers 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.) As with flqcl 10304 we only prove this for rationals although other particular kinds of real numbers may be possible. (Contributed by Jim Kingdon, 16-Oct-2021.)
Assertion
Ref Expression
modqval ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐴 mod 𝐵) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))))

Proof of Theorem modqval
Dummy variables 𝑥 𝑦 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 qre 9655 . . 3 (𝐴 ∈ ℚ → 𝐴 ∈ ℝ)
213ad2ant1 1020 . 2 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐴 ∈ ℝ)
3 qre 9655 . . . 4 (𝐵 ∈ ℚ → 𝐵 ∈ ℝ)
433ad2ant2 1021 . . 3 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐵 ∈ ℝ)
5 simp3 1001 . . 3 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 0 < 𝐵)
64, 5elrpd 9723 . 2 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐵 ∈ ℝ+)
75gt0ne0d 8499 . . . . . . 7 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐵 ≠ 0)
8 qdivcl 9673 . . . . . . 7 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 𝐵 ≠ 0) → (𝐴 / 𝐵) ∈ ℚ)
97, 8syld3an3 1294 . . . . . 6 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐴 / 𝐵) ∈ ℚ)
109flqcld 10308 . . . . 5 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (⌊‘(𝐴 / 𝐵)) ∈ ℤ)
1110zred 9405 . . . 4 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (⌊‘(𝐴 / 𝐵)) ∈ ℝ)
124, 11remulcld 8018 . . 3 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐵 · (⌊‘(𝐴 / 𝐵))) ∈ ℝ)
132, 12resubcld 8368 . 2 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))) ∈ ℝ)
14 oveq1 5903 . . . . . 6 (𝑥 = 𝐴 → (𝑥 / 𝑦) = (𝐴 / 𝑦))
1514fveq2d 5538 . . . . 5 (𝑥 = 𝐴 → (⌊‘(𝑥 / 𝑦)) = (⌊‘(𝐴 / 𝑦)))
1615oveq2d 5912 . . . 4 (𝑥 = 𝐴 → (𝑦 · (⌊‘(𝑥 / 𝑦))) = (𝑦 · (⌊‘(𝐴 / 𝑦))))
17 oveq12 5905 . . . 4 ((𝑥 = 𝐴 ∧ (𝑦 · (⌊‘(𝑥 / 𝑦))) = (𝑦 · (⌊‘(𝐴 / 𝑦)))) → (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))) = (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦)))))
1816, 17mpdan 421 . . 3 (𝑥 = 𝐴 → (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))) = (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦)))))
19 oveq2 5904 . . . . . 6 (𝑦 = 𝐵 → (𝐴 / 𝑦) = (𝐴 / 𝐵))
2019fveq2d 5538 . . . . 5 (𝑦 = 𝐵 → (⌊‘(𝐴 / 𝑦)) = (⌊‘(𝐴 / 𝐵)))
21 oveq12 5905 . . . . 5 ((𝑦 = 𝐵 ∧ (⌊‘(𝐴 / 𝑦)) = (⌊‘(𝐴 / 𝐵))) → (𝑦 · (⌊‘(𝐴 / 𝑦))) = (𝐵 · (⌊‘(𝐴 / 𝐵))))
2220, 21mpdan 421 . . . 4 (𝑦 = 𝐵 → (𝑦 · (⌊‘(𝐴 / 𝑦))) = (𝐵 · (⌊‘(𝐴 / 𝐵))))
2322oveq2d 5912 . . 3 (𝑦 = 𝐵 → (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦)))) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))))
24 df-mod 10354 . . 3 mod = (𝑥 ∈ ℝ, 𝑦 ∈ ℝ+ ↦ (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))))
2518, 23, 24ovmpog 6031 . 2 ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ+ ∧ (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))) ∈ ℝ) → (𝐴 mod 𝐵) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))))
262, 6, 13, 25syl3anc 1249 1 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐴 mod 𝐵) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))))
Colors of variables: wff set class
Syntax hints:  wi 4  w3a 980   = wceq 1364  wcel 2160  wne 2360   class class class wbr 4018  cfv 5235  (class class class)co 5896  cr 7840  0cc0 7841   · cmul 7846   < clt 8022  cmin 8158   / cdiv 8659  cq 9649  +crp 9683  cfl 10299   mod cmo 10353
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-ia1 106  ax-ia2 107  ax-ia3 108  ax-in1 615  ax-in2 616  ax-io 710  ax-5 1458  ax-7 1459  ax-gen 1460  ax-ie1 1504  ax-ie2 1505  ax-8 1515  ax-10 1516  ax-11 1517  ax-i12 1518  ax-bndl 1520  ax-4 1521  ax-17 1537  ax-i9 1541  ax-ial 1545  ax-i5r 1546  ax-13 2162  ax-14 2163  ax-ext 2171  ax-sep 4136  ax-pow 4192  ax-pr 4227  ax-un 4451  ax-setind 4554  ax-cnex 7932  ax-resscn 7933  ax-1cn 7934  ax-1re 7935  ax-icn 7936  ax-addcl 7937  ax-addrcl 7938  ax-mulcl 7939  ax-mulrcl 7940  ax-addcom 7941  ax-mulcom 7942  ax-addass 7943  ax-mulass 7944  ax-distr 7945  ax-i2m1 7946  ax-0lt1 7947  ax-1rid 7948  ax-0id 7949  ax-rnegex 7950  ax-precex 7951  ax-cnre 7952  ax-pre-ltirr 7953  ax-pre-ltwlin 7954  ax-pre-lttrn 7955  ax-pre-apti 7956  ax-pre-ltadd 7957  ax-pre-mulgt0 7958  ax-pre-mulext 7959  ax-arch 7960
This theorem depends on definitions:  df-bi 117  df-3or 981  df-3an 982  df-tru 1367  df-fal 1370  df-nf 1472  df-sb 1774  df-eu 2041  df-mo 2042  df-clab 2176  df-cleq 2182  df-clel 2185  df-nfc 2321  df-ne 2361  df-nel 2456  df-ral 2473  df-rex 2474  df-reu 2475  df-rmo 2476  df-rab 2477  df-v 2754  df-sbc 2978  df-csb 3073  df-dif 3146  df-un 3148  df-in 3150  df-ss 3157  df-pw 3592  df-sn 3613  df-pr 3614  df-op 3616  df-uni 3825  df-int 3860  df-iun 3903  df-br 4019  df-opab 4080  df-mpt 4081  df-id 4311  df-po 4314  df-iso 4315  df-xp 4650  df-rel 4651  df-cnv 4652  df-co 4653  df-dm 4654  df-rn 4655  df-res 4656  df-ima 4657  df-iota 5196  df-fun 5237  df-fn 5238  df-f 5239  df-fv 5243  df-riota 5852  df-ov 5899  df-oprab 5900  df-mpo 5901  df-1st 6165  df-2nd 6166  df-pnf 8024  df-mnf 8025  df-xr 8026  df-ltxr 8027  df-le 8028  df-sub 8160  df-neg 8161  df-reap 8562  df-ap 8569  df-div 8660  df-inn 8950  df-n0 9207  df-z 9284  df-q 9650  df-rp 9684  df-fl 10301  df-mod 10354
This theorem is referenced by:  modqvalr  10356  modqcl  10357  modq0  10360  modqge0  10363  modqlt  10364  modqdiffl  10366  modqfrac  10368  modqmulnn  10373  zmodcl  10375  modqid  10380  modqcyc  10390  modqadd1  10392  modqmul1  10408  modqdi  10423  modqsubdir  10424  iexpcyc  10656  dvdsmod  11900  divalgmod  11964  modgcd  12024  prmdiv  12267  odzdvds  12277  fldivp1  12380  mulgmodid  13101
  Copyright terms: Public domain W3C validator