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

Theorem divides 14823
Description: Define the divides relation. 𝑀𝑁 means 𝑀 divides into 𝑁 with no remainder. For example, 3 ∥ 6 (ex-dvds 26705). As proven in dvdsval3 14825, 𝑀𝑁 ↔ (𝑁 mod 𝑀) = 0. See divides 14823 and dvdsval2 14824 for other equivalent expressions. (Contributed by Paul Chapman, 21-Mar-2011.)
Assertion
Ref Expression
divides ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀𝑁 ↔ ∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑁))
Distinct variable groups:   𝑛,𝑀   𝑛,𝑁

Proof of Theorem divides
Dummy variables 𝑥 𝑦 are mutually distinct and distinct from all other variables.
StepHypRef Expression
1 df-br 4584 . . 3 (𝑀𝑁 ↔ ⟨𝑀, 𝑁⟩ ∈ ∥ )
2 df-dvds 14822 . . . 4 ∥ = {⟨𝑥, 𝑦⟩ ∣ ((𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ) ∧ ∃𝑛 ∈ ℤ (𝑛 · 𝑥) = 𝑦)}
32eleq2i 2680 . . 3 (⟨𝑀, 𝑁⟩ ∈ ∥ ↔ ⟨𝑀, 𝑁⟩ ∈ {⟨𝑥, 𝑦⟩ ∣ ((𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ) ∧ ∃𝑛 ∈ ℤ (𝑛 · 𝑥) = 𝑦)})
41, 3bitri 263 . 2 (𝑀𝑁 ↔ ⟨𝑀, 𝑁⟩ ∈ {⟨𝑥, 𝑦⟩ ∣ ((𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ) ∧ ∃𝑛 ∈ ℤ (𝑛 · 𝑥) = 𝑦)})
5 oveq2 6557 . . . . 5 (𝑥 = 𝑀 → (𝑛 · 𝑥) = (𝑛 · 𝑀))
65eqeq1d 2612 . . . 4 (𝑥 = 𝑀 → ((𝑛 · 𝑥) = 𝑦 ↔ (𝑛 · 𝑀) = 𝑦))
76rexbidv 3034 . . 3 (𝑥 = 𝑀 → (∃𝑛 ∈ ℤ (𝑛 · 𝑥) = 𝑦 ↔ ∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑦))
8 eqeq2 2621 . . . 4 (𝑦 = 𝑁 → ((𝑛 · 𝑀) = 𝑦 ↔ (𝑛 · 𝑀) = 𝑁))
98rexbidv 3034 . . 3 (𝑦 = 𝑁 → (∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑦 ↔ ∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑁))
107, 9opelopab2 4921 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (⟨𝑀, 𝑁⟩ ∈ {⟨𝑥, 𝑦⟩ ∣ ((𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ) ∧ ∃𝑛 ∈ ℤ (𝑛 · 𝑥) = 𝑦)} ↔ ∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑁))
114, 10syl5bb 271 1 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀𝑁 ↔ ∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑁))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 195  wa 383   = wceq 1475  wcel 1977  wrex 2897  cop 4131   class class class wbr 4583  {copab 4642  (class class class)co 6549   · cmul 9820  cz 11254  cdvds 14821
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-rex 2902  df-rab 2905  df-v 3175  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-iota 5768  df-fv 5812  df-ov 6552  df-dvds 14822
This theorem is referenced by:  dvdsval2  14824  dvds0lem  14830  dvds1lem  14831  dvds2lem  14832  0dvds  14840  dvdsle  14870  divconjdvds  14875  odd2np1  14903  even2n  14904  oddm1even  14905  opeo  14927  omeo  14928  m1exp1  14931  divalglem4  14957  divalglem9  14962  divalgb  14965  modremain  14970  zeqzmulgcd  15070  bezoutlem4  15097  gcddiv  15106  dvdssqim  15111  coprmdvds2  15206  congr  15216  divgcdcoprm0  15217  cncongr2  15220  dvdsnprmd  15241  prmpwdvds  15446  odmulg  17796  gexdvdsi  17821  lgsquadlem2  24906  dvdspw  30889  dvdsrabdioph  36392  jm2.26a  36585  coskpi2  38749  cosknegpi  38752  fourierswlem  39123  dfeven2  40100
  Copyright terms: Public domain W3C validator