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

Theorem mulmoddvds 15787
Description: If an integer is divisible by a positive integer, the product of this integer with another integer modulo the positive integer is 0. (Contributed by Alexander van der Vekens, 30-Aug-2018.) (Proof shortened by AV, 18-Mar-2022.)
Assertion
Ref Expression
mulmoddvds ((𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝑁𝐴 → ((𝐴 · 𝐵) mod 𝑁) = 0))

Proof of Theorem mulmoddvds
StepHypRef Expression
1 simp1 1137 . 2 ((𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → 𝑁 ∈ ℕ)
2 nnz 12097 . . 3 (𝑁 ∈ ℕ → 𝑁 ∈ ℤ)
3 dvdsmultr1 15753 . . 3 ((𝑁 ∈ ℤ ∧ 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝑁𝐴𝑁 ∥ (𝐴 · 𝐵)))
42, 3syl3an1 1164 . 2 ((𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝑁𝐴𝑁 ∥ (𝐴 · 𝐵)))
5 dvdsmod0 15717 . 2 ((𝑁 ∈ ℕ ∧ 𝑁 ∥ (𝐴 · 𝐵)) → ((𝐴 · 𝐵) mod 𝑁) = 0)
61, 4, 5syl6an 684 1 ((𝑁 ∈ ℕ ∧ 𝐴 ∈ ℤ ∧ 𝐵 ∈ ℤ) → (𝑁𝐴 → ((𝐴 · 𝐵) mod 𝑁) = 0))
Colors of variables: wff setvar class
Syntax hints:  wi 4  w3a 1088   = wceq 1542  wcel 2114   class class class wbr 5040  (class class class)co 7182  0cc0 10627   · cmul 10632  cn 11728  cz 12074   mod cmo 13340  cdvds 15711
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1802  ax-4 1816  ax-5 1917  ax-6 1975  ax-7 2020  ax-8 2116  ax-9 2124  ax-10 2145  ax-11 2162  ax-12 2179  ax-ext 2711  ax-sep 5177  ax-nul 5184  ax-pow 5242  ax-pr 5306  ax-un 7491  ax-cnex 10683  ax-resscn 10684  ax-1cn 10685  ax-icn 10686  ax-addcl 10687  ax-addrcl 10688  ax-mulcl 10689  ax-mulrcl 10690  ax-mulcom 10691  ax-addass 10692  ax-mulass 10693  ax-distr 10694  ax-i2m1 10695  ax-1ne0 10696  ax-1rid 10697  ax-rnegex 10698  ax-rrecex 10699  ax-cnre 10700  ax-pre-lttri 10701  ax-pre-lttrn 10702  ax-pre-ltadd 10703  ax-pre-mulgt0 10704  ax-pre-sup 10705
This theorem depends on definitions:  df-bi 210  df-an 400  df-or 847  df-3or 1089  df-3an 1090  df-tru 1545  df-fal 1555  df-ex 1787  df-nf 1791  df-sb 2075  df-mo 2541  df-eu 2571  df-clab 2718  df-cleq 2731  df-clel 2812  df-nfc 2882  df-ne 2936  df-nel 3040  df-ral 3059  df-rex 3060  df-reu 3061  df-rmo 3062  df-rab 3063  df-v 3402  df-sbc 3686  df-csb 3801  df-dif 3856  df-un 3858  df-in 3860  df-ss 3870  df-pss 3872  df-nul 4222  df-if 4425  df-pw 4500  df-sn 4527  df-pr 4529  df-tp 4531  df-op 4533  df-uni 4807  df-iun 4893  df-br 5041  df-opab 5103  df-mpt 5121  df-tr 5147  df-id 5439  df-eprel 5444  df-po 5452  df-so 5453  df-fr 5493  df-we 5495  df-xp 5541  df-rel 5542  df-cnv 5543  df-co 5544  df-dm 5545  df-rn 5546  df-res 5547  df-ima 5548  df-pred 6139  df-ord 6185  df-on 6186  df-lim 6187  df-suc 6188  df-iota 6307  df-fun 6351  df-fn 6352  df-f 6353  df-f1 6354  df-fo 6355  df-f1o 6356  df-fv 6357  df-riota 7139  df-ov 7185  df-oprab 7186  df-mpo 7187  df-om 7612  df-wrecs 7988  df-recs 8049  df-rdg 8087  df-er 8332  df-en 8568  df-dom 8569  df-sdom 8570  df-sup 8991  df-inf 8992  df-pnf 10767  df-mnf 10768  df-xr 10769  df-ltxr 10770  df-le 10771  df-sub 10962  df-neg 10963  df-div 11388  df-nn 11729  df-n0 11989  df-z 12075  df-uz 12337  df-rp 12485  df-fl 13265  df-mod 13341  df-dvds 15712
This theorem is referenced by:  numclwwlk5  28337  numclwwlk7  28340
  Copyright terms: Public domain W3C validator