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

Theorem modqval 10433
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 10380 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 9716 . . 3 (𝐴 ∈ ℚ → 𝐴 ∈ ℝ)
213ad2ant1 1020 . 2 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐴 ∈ ℝ)
3 qre 9716 . . . 4 (𝐵 ∈ ℚ → 𝐵 ∈ ℝ)
433ad2ant2 1021 . . 3 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐵 ∈ ℝ)
5 simp3 1001 . . 3 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 0 < 𝐵)
64, 5elrpd 9785 . 2 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐵 ∈ ℝ+)
75gt0ne0d 8556 . . . . . . 7 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐵 ≠ 0)
8 qdivcl 9734 . . . . . . 7 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 𝐵 ≠ 0) → (𝐴 / 𝐵) ∈ ℚ)
97, 8syld3an3 1294 . . . . . 6 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐴 / 𝐵) ∈ ℚ)
109flqcld 10384 . . . . 5 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (⌊‘(𝐴 / 𝐵)) ∈ ℤ)
1110zred 9465 . . . 4 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (⌊‘(𝐴 / 𝐵)) ∈ ℝ)
124, 11remulcld 8074 . . 3 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐵 · (⌊‘(𝐴 / 𝐵))) ∈ ℝ)
132, 12resubcld 8424 . 2 ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))) ∈ ℝ)
14 oveq1 5932 . . . . . 6 (𝑥 = 𝐴 → (𝑥 / 𝑦) = (𝐴 / 𝑦))
1514fveq2d 5565 . . . . 5 (𝑥 = 𝐴 → (⌊‘(𝑥 / 𝑦)) = (⌊‘(𝐴 / 𝑦)))
1615oveq2d 5941 . . . 4 (𝑥 = 𝐴 → (𝑦 · (⌊‘(𝑥 / 𝑦))) = (𝑦 · (⌊‘(𝐴 / 𝑦))))
17 oveq12 5934 . . . 4 ((𝑥 = 𝐴 ∧ (𝑦 · (⌊‘(𝑥 / 𝑦))) = (𝑦 · (⌊‘(𝐴 / 𝑦)))) → (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))) = (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦)))))
1816, 17mpdan 421 . . 3 (𝑥 = 𝐴 → (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))) = (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦)))))
19 oveq2 5933 . . . . . 6 (𝑦 = 𝐵 → (𝐴 / 𝑦) = (𝐴 / 𝐵))
2019fveq2d 5565 . . . . 5 (𝑦 = 𝐵 → (⌊‘(𝐴 / 𝑦)) = (⌊‘(𝐴 / 𝐵)))
21 oveq12 5934 . . . . 5 ((𝑦 = 𝐵 ∧ (⌊‘(𝐴 / 𝑦)) = (⌊‘(𝐴 / 𝐵))) → (𝑦 · (⌊‘(𝐴 / 𝑦))) = (𝐵 · (⌊‘(𝐴 / 𝐵))))
2220, 21mpdan 421 . . . 4 (𝑦 = 𝐵 → (𝑦 · (⌊‘(𝐴 / 𝑦))) = (𝐵 · (⌊‘(𝐴 / 𝐵))))
2322oveq2d 5941 . . 3 (𝑦 = 𝐵 → (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦)))) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))))
24 df-mod 10432 . . 3 mod = (𝑥 ∈ ℝ, 𝑦 ∈ ℝ+ ↦ (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))))
2518, 23, 24ovmpog 6061 . 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 2167  wne 2367   class class class wbr 4034  cfv 5259  (class class class)co 5925  cr 7895  0cc0 7896   · cmul 7901   < clt 8078  cmin 8214   / cdiv 8716  cq 9710  +crp 9745  cfl 10375   mod cmo 10431
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 1461  ax-7 1462  ax-gen 1463  ax-ie1 1507  ax-ie2 1508  ax-8 1518  ax-10 1519  ax-11 1520  ax-i12 1521  ax-bndl 1523  ax-4 1524  ax-17 1540  ax-i9 1544  ax-ial 1548  ax-i5r 1549  ax-13 2169  ax-14 2170  ax-ext 2178  ax-sep 4152  ax-pow 4208  ax-pr 4243  ax-un 4469  ax-setind 4574  ax-cnex 7987  ax-resscn 7988  ax-1cn 7989  ax-1re 7990  ax-icn 7991  ax-addcl 7992  ax-addrcl 7993  ax-mulcl 7994  ax-mulrcl 7995  ax-addcom 7996  ax-mulcom 7997  ax-addass 7998  ax-mulass 7999  ax-distr 8000  ax-i2m1 8001  ax-0lt1 8002  ax-1rid 8003  ax-0id 8004  ax-rnegex 8005  ax-precex 8006  ax-cnre 8007  ax-pre-ltirr 8008  ax-pre-ltwlin 8009  ax-pre-lttrn 8010  ax-pre-apti 8011  ax-pre-ltadd 8012  ax-pre-mulgt0 8013  ax-pre-mulext 8014  ax-arch 8015
This theorem depends on definitions:  df-bi 117  df-3or 981  df-3an 982  df-tru 1367  df-fal 1370  df-nf 1475  df-sb 1777  df-eu 2048  df-mo 2049  df-clab 2183  df-cleq 2189  df-clel 2192  df-nfc 2328  df-ne 2368  df-nel 2463  df-ral 2480  df-rex 2481  df-reu 2482  df-rmo 2483  df-rab 2484  df-v 2765  df-sbc 2990  df-csb 3085  df-dif 3159  df-un 3161  df-in 3163  df-ss 3170  df-pw 3608  df-sn 3629  df-pr 3630  df-op 3632  df-uni 3841  df-int 3876  df-iun 3919  df-br 4035  df-opab 4096  df-mpt 4097  df-id 4329  df-po 4332  df-iso 4333  df-xp 4670  df-rel 4671  df-cnv 4672  df-co 4673  df-dm 4674  df-rn 4675  df-res 4676  df-ima 4677  df-iota 5220  df-fun 5261  df-fn 5262  df-f 5263  df-fv 5267  df-riota 5880  df-ov 5928  df-oprab 5929  df-mpo 5930  df-1st 6207  df-2nd 6208  df-pnf 8080  df-mnf 8081  df-xr 8082  df-ltxr 8083  df-le 8084  df-sub 8216  df-neg 8217  df-reap 8619  df-ap 8626  df-div 8717  df-inn 9008  df-n0 9267  df-z 9344  df-q 9711  df-rp 9746  df-fl 10377  df-mod 10432
This theorem is referenced by:  modqvalr  10434  modqcl  10435  modq0  10438  modqge0  10441  modqlt  10442  modqdiffl  10444  modqfrac  10446  modqmulnn  10451  zmodcl  10453  modqid  10458  modqcyc  10468  modqadd1  10470  modqmul1  10486  modqdi  10501  modqsubdir  10502  iexpcyc  10753  dvdsmod  12044  divalgmod  12109  modgcd  12183  prmdiv  12428  odzdvds  12439  fldivp1  12542  mulgmodid  13367  lgseisenlem4  15398
  Copyright terms: Public domain W3C validator