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

Theorem lcmeq0 16640
Description: The lcm of two integers is zero iff either is zero. (Contributed by Steve Rodriguez, 20-Jan-2020.)
Assertion
Ref Expression
lcmeq0 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝑀 lcm 𝑁) = 0 ↔ (𝑀 = 0 ∨ 𝑁 = 0)))

Proof of Theorem lcmeq0
StepHypRef Expression
1 lcmn0cl 16637 . . . . 5 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∨ 𝑁 = 0)) → (𝑀 lcm 𝑁) ∈ ℕ)
21nnne0d 12320 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∨ 𝑁 = 0)) → (𝑀 lcm 𝑁) ≠ 0)
32ex 412 . . 3 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (¬ (𝑀 = 0 ∨ 𝑁 = 0) → (𝑀 lcm 𝑁) ≠ 0))
43necon4bd 2959 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝑀 lcm 𝑁) = 0 → (𝑀 = 0 ∨ 𝑁 = 0)))
5 oveq1 7442 . . . . . 6 (𝑀 = 0 → (𝑀 lcm 𝑁) = (0 lcm 𝑁))
6 0z 12628 . . . . . . . 8 0 ∈ ℤ
7 lcmcom 16633 . . . . . . . 8 ((𝑁 ∈ ℤ ∧ 0 ∈ ℤ) → (𝑁 lcm 0) = (0 lcm 𝑁))
86, 7mpan2 691 . . . . . . 7 (𝑁 ∈ ℤ → (𝑁 lcm 0) = (0 lcm 𝑁))
9 lcm0val 16634 . . . . . . 7 (𝑁 ∈ ℤ → (𝑁 lcm 0) = 0)
108, 9eqtr3d 2778 . . . . . 6 (𝑁 ∈ ℤ → (0 lcm 𝑁) = 0)
115, 10sylan9eqr 2798 . . . . 5 ((𝑁 ∈ ℤ ∧ 𝑀 = 0) → (𝑀 lcm 𝑁) = 0)
1211adantll 714 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ 𝑀 = 0) → (𝑀 lcm 𝑁) = 0)
13 oveq2 7443 . . . . . 6 (𝑁 = 0 → (𝑀 lcm 𝑁) = (𝑀 lcm 0))
14 lcm0val 16634 . . . . . 6 (𝑀 ∈ ℤ → (𝑀 lcm 0) = 0)
1513, 14sylan9eqr 2798 . . . . 5 ((𝑀 ∈ ℤ ∧ 𝑁 = 0) → (𝑀 lcm 𝑁) = 0)
1615adantlr 715 . . . 4 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ 𝑁 = 0) → (𝑀 lcm 𝑁) = 0)
1712, 16jaodan 959 . . 3 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ (𝑀 = 0 ∨ 𝑁 = 0)) → (𝑀 lcm 𝑁) = 0)
1817ex 412 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝑀 = 0 ∨ 𝑁 = 0) → (𝑀 lcm 𝑁) = 0))
194, 18impbid 212 1 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → ((𝑀 lcm 𝑁) = 0 ↔ (𝑀 = 0 ∨ 𝑁 = 0)))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wb 206  wa 395  wo 847   = wceq 1538  wcel 2107  wne 2939  (class class class)co 7435  0cc0 11159  cz 12617   lcm clcm 16628
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1793  ax-4 1807  ax-5 1909  ax-6 1966  ax-7 2006  ax-8 2109  ax-9 2117  ax-10 2140  ax-11 2156  ax-12 2176  ax-ext 2707  ax-sep 5303  ax-nul 5313  ax-pow 5372  ax-pr 5439  ax-un 7758  ax-cnex 11215  ax-resscn 11216  ax-1cn 11217  ax-icn 11218  ax-addcl 11219  ax-addrcl 11220  ax-mulcl 11221  ax-mulrcl 11222  ax-mulcom 11223  ax-addass 11224  ax-mulass 11225  ax-distr 11226  ax-i2m1 11227  ax-1ne0 11228  ax-1rid 11229  ax-rnegex 11230  ax-rrecex 11231  ax-cnre 11232  ax-pre-lttri 11233  ax-pre-lttrn 11234  ax-pre-ltadd 11235  ax-pre-mulgt0 11236  ax-pre-sup 11237
This theorem depends on definitions:  df-bi 207  df-an 396  df-or 848  df-3or 1087  df-3an 1088  df-tru 1541  df-fal 1551  df-ex 1778  df-nf 1782  df-sb 2064  df-mo 2539  df-eu 2568  df-clab 2714  df-cleq 2728  df-clel 2815  df-nfc 2891  df-ne 2940  df-nel 3046  df-ral 3061  df-rex 3070  df-rmo 3379  df-reu 3380  df-rab 3435  df-v 3481  df-sbc 3793  df-csb 3910  df-dif 3967  df-un 3969  df-in 3971  df-ss 3981  df-pss 3984  df-nul 4341  df-if 4533  df-pw 4608  df-sn 4633  df-pr 4635  df-op 4639  df-uni 4914  df-iun 4999  df-br 5150  df-opab 5212  df-mpt 5233  df-tr 5267  df-id 5584  df-eprel 5590  df-po 5598  df-so 5599  df-fr 5642  df-we 5644  df-xp 5696  df-rel 5697  df-cnv 5698  df-co 5699  df-dm 5700  df-rn 5701  df-res 5702  df-ima 5703  df-pred 6326  df-ord 6392  df-on 6393  df-lim 6394  df-suc 6395  df-iota 6519  df-fun 6568  df-fn 6569  df-f 6570  df-f1 6571  df-fo 6572  df-f1o 6573  df-fv 6574  df-riota 7392  df-ov 7438  df-oprab 7439  df-mpo 7440  df-om 7892  df-2nd 8020  df-frecs 8311  df-wrecs 8342  df-recs 8416  df-rdg 8455  df-er 8750  df-en 8991  df-dom 8992  df-sdom 8993  df-sup 9486  df-inf 9487  df-pnf 11301  df-mnf 11302  df-xr 11303  df-ltxr 11304  df-le 11305  df-sub 11498  df-neg 11499  df-div 11925  df-nn 12271  df-2 12333  df-3 12334  df-n0 12531  df-z 12618  df-uz 12883  df-rp 13039  df-seq 14046  df-exp 14106  df-cj 15141  df-re 15142  df-im 15143  df-sqrt 15277  df-abs 15278  df-dvds 16294  df-lcm 16630
This theorem is referenced by:  lcmass  16654
  Copyright terms: Public domain W3C validator