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

Theorem divides 14985
Description: Define the divides relation. 𝑀𝑁 means 𝑀 divides into 𝑁 with no remainder. For example, 3 ∥ 6 (ex-dvds 27313). As proven in dvdsval3 14987, 𝑀𝑁 ↔ (𝑁 mod 𝑀) = 0. See divides 14985 and dvdsval2 14986 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 4654 . . 3 (𝑀𝑁 ↔ ⟨𝑀, 𝑁⟩ ∈ ∥ )
2 df-dvds 14984 . . . 4 ∥ = {⟨𝑥, 𝑦⟩ ∣ ((𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ) ∧ ∃𝑛 ∈ ℤ (𝑛 · 𝑥) = 𝑦)}
32eleq2i 2693 . . 3 (⟨𝑀, 𝑁⟩ ∈ ∥ ↔ ⟨𝑀, 𝑁⟩ ∈ {⟨𝑥, 𝑦⟩ ∣ ((𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ) ∧ ∃𝑛 ∈ ℤ (𝑛 · 𝑥) = 𝑦)})
41, 3bitri 264 . 2 (𝑀𝑁 ↔ ⟨𝑀, 𝑁⟩ ∈ {⟨𝑥, 𝑦⟩ ∣ ((𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ) ∧ ∃𝑛 ∈ ℤ (𝑛 · 𝑥) = 𝑦)})
5 oveq2 6658 . . . . 5 (𝑥 = 𝑀 → (𝑛 · 𝑥) = (𝑛 · 𝑀))
65eqeq1d 2624 . . . 4 (𝑥 = 𝑀 → ((𝑛 · 𝑥) = 𝑦 ↔ (𝑛 · 𝑀) = 𝑦))
76rexbidv 3052 . . 3 (𝑥 = 𝑀 → (∃𝑛 ∈ ℤ (𝑛 · 𝑥) = 𝑦 ↔ ∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑦))
8 eqeq2 2633 . . . 4 (𝑦 = 𝑁 → ((𝑛 · 𝑀) = 𝑦 ↔ (𝑛 · 𝑀) = 𝑁))
98rexbidv 3052 . . 3 (𝑦 = 𝑁 → (∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑦 ↔ ∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑁))
107, 9opelopab2 4996 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (⟨𝑀, 𝑁⟩ ∈ {⟨𝑥, 𝑦⟩ ∣ ((𝑥 ∈ ℤ ∧ 𝑦 ∈ ℤ) ∧ ∃𝑛 ∈ ℤ (𝑛 · 𝑥) = 𝑦)} ↔ ∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑁))
114, 10syl5bb 272 1 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀𝑁 ↔ ∃𝑛 ∈ ℤ (𝑛 · 𝑀) = 𝑁))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 196  wa 384   = wceq 1483  wcel 1990  wrex 2913  cop 4183   class class class wbr 4653  {copab 4712  (class class class)co 6650   · cmul 9941  cz 11377  cdvds 14983
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1722  ax-4 1737  ax-5 1839  ax-6 1888  ax-7 1935  ax-9 1999  ax-10 2019  ax-11 2034  ax-12 2047  ax-13 2246  ax-ext 2602  ax-sep 4781  ax-nul 4789  ax-pr 4906
This theorem depends on definitions:  df-bi 197  df-or 385  df-an 386  df-3an 1039  df-tru 1486  df-ex 1705  df-nf 1710  df-sb 1881  df-eu 2474  df-mo 2475  df-clab 2609  df-cleq 2615  df-clel 2618  df-nfc 2753  df-rex 2918  df-rab 2921  df-v 3202  df-dif 3577  df-un 3579  df-in 3581  df-ss 3588  df-nul 3916  df-if 4087  df-sn 4178  df-pr 4180  df-op 4184  df-uni 4437  df-br 4654  df-opab 4713  df-iota 5851  df-fv 5896  df-ov 6653  df-dvds 14984
This theorem is referenced by:  dvdsval2  14986  dvds0lem  14992  dvds1lem  14993  dvds2lem  14994  0dvds  15002  dvdsle  15032  divconjdvds  15037  odd2np1  15065  even2n  15066  oddm1even  15067  opeo  15089  omeo  15090  m1exp1  15093  divalglem4  15119  divalglem9  15124  divalgb  15127  modremain  15132  zeqzmulgcd  15232  bezoutlem4  15259  gcddiv  15268  dvdssqim  15273  coprmdvds2  15368  congr  15378  divgcdcoprm0  15379  cncongr2  15382  dvdsnprmd  15403  prmpwdvds  15608  odmulg  17973  gexdvdsi  17998  lgsquadlem2  25106  dvdspw  31636  dvdsrabdioph  37374  jm2.26a  37567  coskpi2  40077  cosknegpi  40080  fourierswlem  40447  dfeven2  41562
  Copyright terms: Public domain W3C validator