![]() |
Intuitionistic Logic Explorer |
< Previous
Next >
Nearby theorems |
|
Mirrors > Home > ILE Home > Th. List > modqval | GIF version |
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 10345 we only prove this for rationals although other particular kinds of real numbers may be possible. (Contributed by Jim Kingdon, 16-Oct-2021.) |
Ref | Expression |
---|---|
modqval | ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐴 mod 𝐵) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵))))) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | qre 9693 | . . 3 ⊢ (𝐴 ∈ ℚ → 𝐴 ∈ ℝ) | |
2 | 1 | 3ad2ant1 1020 | . 2 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐴 ∈ ℝ) |
3 | qre 9693 | . . . 4 ⊢ (𝐵 ∈ ℚ → 𝐵 ∈ ℝ) | |
4 | 3 | 3ad2ant2 1021 | . . 3 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐵 ∈ ℝ) |
5 | simp3 1001 | . . 3 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 0 < 𝐵) | |
6 | 4, 5 | elrpd 9762 | . 2 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐵 ∈ ℝ+) |
7 | 5 | gt0ne0d 8533 | . . . . . . 7 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → 𝐵 ≠ 0) |
8 | qdivcl 9711 | . . . . . . 7 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 𝐵 ≠ 0) → (𝐴 / 𝐵) ∈ ℚ) | |
9 | 7, 8 | syld3an3 1294 | . . . . . 6 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐴 / 𝐵) ∈ ℚ) |
10 | 9 | flqcld 10349 | . . . . 5 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (⌊‘(𝐴 / 𝐵)) ∈ ℤ) |
11 | 10 | zred 9442 | . . . 4 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (⌊‘(𝐴 / 𝐵)) ∈ ℝ) |
12 | 4, 11 | remulcld 8052 | . . 3 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐵 · (⌊‘(𝐴 / 𝐵))) ∈ ℝ) |
13 | 2, 12 | resubcld 8402 | . 2 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))) ∈ ℝ) |
14 | oveq1 5926 | . . . . . 6 ⊢ (𝑥 = 𝐴 → (𝑥 / 𝑦) = (𝐴 / 𝑦)) | |
15 | 14 | fveq2d 5559 | . . . . 5 ⊢ (𝑥 = 𝐴 → (⌊‘(𝑥 / 𝑦)) = (⌊‘(𝐴 / 𝑦))) |
16 | 15 | oveq2d 5935 | . . . 4 ⊢ (𝑥 = 𝐴 → (𝑦 · (⌊‘(𝑥 / 𝑦))) = (𝑦 · (⌊‘(𝐴 / 𝑦)))) |
17 | oveq12 5928 | . . . 4 ⊢ ((𝑥 = 𝐴 ∧ (𝑦 · (⌊‘(𝑥 / 𝑦))) = (𝑦 · (⌊‘(𝐴 / 𝑦)))) → (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))) = (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦))))) | |
18 | 16, 17 | mpdan 421 | . . 3 ⊢ (𝑥 = 𝐴 → (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦)))) = (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦))))) |
19 | oveq2 5927 | . . . . . 6 ⊢ (𝑦 = 𝐵 → (𝐴 / 𝑦) = (𝐴 / 𝐵)) | |
20 | 19 | fveq2d 5559 | . . . . 5 ⊢ (𝑦 = 𝐵 → (⌊‘(𝐴 / 𝑦)) = (⌊‘(𝐴 / 𝐵))) |
21 | oveq12 5928 | . . . . 5 ⊢ ((𝑦 = 𝐵 ∧ (⌊‘(𝐴 / 𝑦)) = (⌊‘(𝐴 / 𝐵))) → (𝑦 · (⌊‘(𝐴 / 𝑦))) = (𝐵 · (⌊‘(𝐴 / 𝐵)))) | |
22 | 20, 21 | mpdan 421 | . . . 4 ⊢ (𝑦 = 𝐵 → (𝑦 · (⌊‘(𝐴 / 𝑦))) = (𝐵 · (⌊‘(𝐴 / 𝐵)))) |
23 | 22 | oveq2d 5935 | . . 3 ⊢ (𝑦 = 𝐵 → (𝐴 − (𝑦 · (⌊‘(𝐴 / 𝑦)))) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵))))) |
24 | df-mod 10397 | . . 3 ⊢ mod = (𝑥 ∈ ℝ, 𝑦 ∈ ℝ+ ↦ (𝑥 − (𝑦 · (⌊‘(𝑥 / 𝑦))))) | |
25 | 18, 23, 24 | ovmpog 6054 | . 2 ⊢ ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ+ ∧ (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))) ∈ ℝ) → (𝐴 mod 𝐵) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵))))) |
26 | 2, 6, 13, 25 | syl3anc 1249 | 1 ⊢ ((𝐴 ∈ ℚ ∧ 𝐵 ∈ ℚ ∧ 0 < 𝐵) → (𝐴 mod 𝐵) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵))))) |
Colors of variables: wff set class |
Syntax hints: → wi 4 ∧ w3a 980 = wceq 1364 ∈ wcel 2164 ≠ wne 2364 class class class wbr 4030 ‘cfv 5255 (class class class)co 5919 ℝcr 7873 0cc0 7874 · cmul 7879 < clt 8056 − cmin 8192 / cdiv 8693 ℚcq 9687 ℝ+crp 9722 ⌊cfl 10340 mod cmo 10396 |
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 2166 ax-14 2167 ax-ext 2175 ax-sep 4148 ax-pow 4204 ax-pr 4239 ax-un 4465 ax-setind 4570 ax-cnex 7965 ax-resscn 7966 ax-1cn 7967 ax-1re 7968 ax-icn 7969 ax-addcl 7970 ax-addrcl 7971 ax-mulcl 7972 ax-mulrcl 7973 ax-addcom 7974 ax-mulcom 7975 ax-addass 7976 ax-mulass 7977 ax-distr 7978 ax-i2m1 7979 ax-0lt1 7980 ax-1rid 7981 ax-0id 7982 ax-rnegex 7983 ax-precex 7984 ax-cnre 7985 ax-pre-ltirr 7986 ax-pre-ltwlin 7987 ax-pre-lttrn 7988 ax-pre-apti 7989 ax-pre-ltadd 7990 ax-pre-mulgt0 7991 ax-pre-mulext 7992 ax-arch 7993 |
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 2045 df-mo 2046 df-clab 2180 df-cleq 2186 df-clel 2189 df-nfc 2325 df-ne 2365 df-nel 2460 df-ral 2477 df-rex 2478 df-reu 2479 df-rmo 2480 df-rab 2481 df-v 2762 df-sbc 2987 df-csb 3082 df-dif 3156 df-un 3158 df-in 3160 df-ss 3167 df-pw 3604 df-sn 3625 df-pr 3626 df-op 3628 df-uni 3837 df-int 3872 df-iun 3915 df-br 4031 df-opab 4092 df-mpt 4093 df-id 4325 df-po 4328 df-iso 4329 df-xp 4666 df-rel 4667 df-cnv 4668 df-co 4669 df-dm 4670 df-rn 4671 df-res 4672 df-ima 4673 df-iota 5216 df-fun 5257 df-fn 5258 df-f 5259 df-fv 5263 df-riota 5874 df-ov 5922 df-oprab 5923 df-mpo 5924 df-1st 6195 df-2nd 6196 df-pnf 8058 df-mnf 8059 df-xr 8060 df-ltxr 8061 df-le 8062 df-sub 8194 df-neg 8195 df-reap 8596 df-ap 8603 df-div 8694 df-inn 8985 df-n0 9244 df-z 9321 df-q 9688 df-rp 9723 df-fl 10342 df-mod 10397 |
This theorem is referenced by: modqvalr 10399 modqcl 10400 modq0 10403 modqge0 10406 modqlt 10407 modqdiffl 10409 modqfrac 10411 modqmulnn 10416 zmodcl 10418 modqid 10423 modqcyc 10433 modqadd1 10435 modqmul1 10451 modqdi 10466 modqsubdir 10467 iexpcyc 10718 dvdsmod 12007 divalgmod 12071 modgcd 12131 prmdiv 12376 odzdvds 12386 fldivp1 12489 mulgmodid 13234 lgseisenlem4 15230 |
Copyright terms: Public domain | W3C validator |