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

Theorem coprmdvds 15991
Description: Euclid's Lemma (see ProofWiki "Euclid's Lemma", 10-Jul-2021, https://proofwiki.org/wiki/Euclid's_Lemma): If an integer divides the product of two integers and is coprime to one of them, then it divides the other. See also theorem 1.5 in [ApostolNT] p. 16. Generalization of euclemma 16051. (Contributed by Paul Chapman, 22-Jun-2011.) (Proof shortened by AV, 10-Jul-2021.)
Assertion
Ref Expression
coprmdvds ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 ∥ (𝑀 · 𝑁) ∧ (𝐾 gcd 𝑀) = 1) → 𝐾𝑁))

Proof of Theorem coprmdvds
StepHypRef Expression
1 zcn 11980 . . . . . . . . . 10 (𝑀 ∈ ℤ → 𝑀 ∈ ℂ)
2 zcn 11980 . . . . . . . . . 10 (𝑁 ∈ ℤ → 𝑁 ∈ ℂ)
3 mulcom 10617 . . . . . . . . . 10 ((𝑀 ∈ ℂ ∧ 𝑁 ∈ ℂ) → (𝑀 · 𝑁) = (𝑁 · 𝑀))
41, 2, 3syl2an 597 . . . . . . . . 9 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 · 𝑁) = (𝑁 · 𝑀))
54breq2d 5070 . . . . . . . 8 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · 𝑀)))
6 dvdsmulgcd 15899 . . . . . . . . 9 ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝐾 ∥ (𝑁 · 𝑀) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
76ancoms 461 . . . . . . . 8 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑁 · 𝑀) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
85, 7bitrd 281 . . . . . . 7 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
983adant1 1126 . . . . . 6 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
109adantr 483 . . . . 5 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾))))
11 gcdcom 15856 . . . . . . . . . . 11 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ) → (𝐾 gcd 𝑀) = (𝑀 gcd 𝐾))
12113adant3 1128 . . . . . . . . . 10 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝐾 gcd 𝑀) = (𝑀 gcd 𝐾))
1312eqeq1d 2823 . . . . . . . . 9 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 ↔ (𝑀 gcd 𝐾) = 1))
14 oveq2 7158 . . . . . . . . 9 ((𝑀 gcd 𝐾) = 1 → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1))
1513, 14syl6bi 255 . . . . . . . 8 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1)))
1615imp 409 . . . . . . 7 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · (𝑀 gcd 𝐾)) = (𝑁 · 1))
172mulid1d 10652 . . . . . . . . 9 (𝑁 ∈ ℤ → (𝑁 · 1) = 𝑁)
18173ad2ant3 1131 . . . . . . . 8 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑁 · 1) = 𝑁)
1918adantr 483 . . . . . . 7 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · 1) = 𝑁)
2016, 19eqtrd 2856 . . . . . 6 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝑁 · (𝑀 gcd 𝐾)) = 𝑁)
2120breq2d 5070 . . . . 5 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑁 · (𝑀 gcd 𝐾)) ↔ 𝐾𝑁))
2210, 21bitrd 281 . . . 4 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) ↔ 𝐾𝑁))
2322biimpd 231 . . 3 (((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝐾 gcd 𝑀) = 1) → (𝐾 ∥ (𝑀 · 𝑁) → 𝐾𝑁))
2423ex 415 . 2 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 gcd 𝑀) = 1 → (𝐾 ∥ (𝑀 · 𝑁) → 𝐾𝑁)))
2524impcomd 414 1 ((𝐾 ∈ ℤ ∧ 𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝐾 ∥ (𝑀 · 𝑁) ∧ (𝐾 gcd 𝑀) = 1) → 𝐾𝑁))
Colors of variables: wff setvar class
Syntax hints:  wi 4  wb 208  wa 398  w3a 1083   = wceq 1533  wcel 2110   class class class wbr 5058  (class class class)co 7150  cc 10529  1c1 10532   · cmul 10536  cz 11975  cdvds 15601   gcd cgcd 15837
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1792  ax-4 1806  ax-5 1907  ax-6 1966  ax-7 2011  ax-8 2112  ax-9 2120  ax-10 2141  ax-11 2157  ax-12 2173  ax-ext 2793  ax-sep 5195  ax-nul 5202  ax-pow 5258  ax-pr 5321  ax-un 7455  ax-cnex 10587  ax-resscn 10588  ax-1cn 10589  ax-icn 10590  ax-addcl 10591  ax-addrcl 10592  ax-mulcl 10593  ax-mulrcl 10594  ax-mulcom 10595  ax-addass 10596  ax-mulass 10597  ax-distr 10598  ax-i2m1 10599  ax-1ne0 10600  ax-1rid 10601  ax-rnegex 10602  ax-rrecex 10603  ax-cnre 10604  ax-pre-lttri 10605  ax-pre-lttrn 10606  ax-pre-ltadd 10607  ax-pre-mulgt0 10608  ax-pre-sup 10609
This theorem depends on definitions:  df-bi 209  df-an 399  df-or 844  df-3or 1084  df-3an 1085  df-tru 1536  df-ex 1777  df-nf 1781  df-sb 2066  df-mo 2618  df-eu 2650  df-clab 2800  df-cleq 2814  df-clel 2893  df-nfc 2963  df-ne 3017  df-nel 3124  df-ral 3143  df-rex 3144  df-reu 3145  df-rmo 3146  df-rab 3147  df-v 3496  df-sbc 3772  df-csb 3883  df-dif 3938  df-un 3940  df-in 3942  df-ss 3951  df-pss 3953  df-nul 4291  df-if 4467  df-pw 4540  df-sn 4561  df-pr 4563  df-tp 4565  df-op 4567  df-uni 4832  df-iun 4913  df-br 5059  df-opab 5121  df-mpt 5139  df-tr 5165  df-id 5454  df-eprel 5459  df-po 5468  df-so 5469  df-fr 5508  df-we 5510  df-xp 5555  df-rel 5556  df-cnv 5557  df-co 5558  df-dm 5559  df-rn 5560  df-res 5561  df-ima 5562  df-pred 6142  df-ord 6188  df-on 6189  df-lim 6190  df-suc 6191  df-iota 6308  df-fun 6351  df-fn 6352  df-f 6353  df-f1 6354  df-fo 6355  df-f1o 6356  df-fv 6357  df-riota 7108  df-ov 7153  df-oprab 7154  df-mpo 7155  df-om 7575  df-2nd 7684  df-wrecs 7941  df-recs 8002  df-rdg 8040  df-er 8283  df-en 8504  df-dom 8505  df-sdom 8506  df-sup 8900  df-inf 8901  df-pnf 10671  df-mnf 10672  df-xr 10673  df-ltxr 10674  df-le 10675  df-sub 10866  df-neg 10867  df-div 11292  df-nn 11633  df-2 11694  df-3 11695  df-n0 11892  df-z 11976  df-uz 12238  df-rp 12384  df-fl 13156  df-mod 13232  df-seq 13364  df-exp 13424  df-cj 14452  df-re 14453  df-im 14454  df-sqrt 14588  df-abs 14589  df-dvds 15602  df-gcd 15838
This theorem is referenced by:  coprmdvds2  15992  qredeq  15995  cncongr1  16005  euclemma  16051  eulerthlem2  16113  prmdiveq  16117  prmpwdvds  16234  ablfacrp2  19183  dvdsmulf1o  25765  perfectlem1  25799  lgseisenlem1  25945  lgseisenlem2  25946  lgsquadlem2  25951  lgsquadlem3  25952  2sqlem8  25996  2sqmod  26006  nn0prpwlem  33665  coprmdvdsb  39575  jm2.20nn  39587  perfectALTVlem1  43880
  Copyright terms: Public domain W3C validator