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

Theorem lcmn0val 15927
Description: The value of the lcm operator when both operands are nonzero. (Contributed by Steve Rodriguez, 20-Jan-2020.) (Revised by AV, 16-Sep-2020.)
Assertion
Ref Expression
lcmn0val (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∨ 𝑁 = 0)) → (𝑀 lcm 𝑁) = inf({𝑛 ∈ ℕ ∣ (𝑀𝑛𝑁𝑛)}, ℝ, < ))
Distinct variable groups:   𝑛,𝑀   𝑛,𝑁

Proof of Theorem lcmn0val
StepHypRef Expression
1 lcmval 15924 . 2 ((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) → (𝑀 lcm 𝑁) = if((𝑀 = 0 ∨ 𝑁 = 0), 0, inf({𝑛 ∈ ℕ ∣ (𝑀𝑛𝑁𝑛)}, ℝ, < )))
2 iffalse 4472 . 2 (¬ (𝑀 = 0 ∨ 𝑁 = 0) → if((𝑀 = 0 ∨ 𝑁 = 0), 0, inf({𝑛 ∈ ℕ ∣ (𝑀𝑛𝑁𝑛)}, ℝ, < )) = inf({𝑛 ∈ ℕ ∣ (𝑀𝑛𝑁𝑛)}, ℝ, < ))
31, 2sylan9eq 2873 1 (((𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ) ∧ ¬ (𝑀 = 0 ∨ 𝑁 = 0)) → (𝑀 lcm 𝑁) = inf({𝑛 ∈ ℕ ∣ (𝑀𝑛𝑁𝑛)}, ℝ, < ))
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 396  wo 841   = wceq 1528  wcel 2105  {crab 3139  ifcif 4463   class class class wbr 5057  (class class class)co 7145  infcinf 8893  cr 10524  0cc0 10525   < clt 10663  cn 11626  cz 11969  cdvds 15595   lcm clcm 15920
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1787  ax-4 1801  ax-5 1902  ax-6 1961  ax-7 2006  ax-8 2107  ax-9 2115  ax-10 2136  ax-11 2151  ax-12 2167  ax-ext 2790  ax-sep 5194  ax-nul 5201  ax-pow 5257  ax-pr 5320  ax-un 7450  ax-resscn 10582  ax-1cn 10583  ax-icn 10584  ax-addcl 10585  ax-mulcl 10587  ax-i2m1 10593  ax-pre-lttri 10599  ax-pre-lttrn 10600
This theorem depends on definitions:  df-bi 208  df-an 397  df-or 842  df-3or 1080  df-3an 1081  df-tru 1531  df-ex 1772  df-nf 1776  df-sb 2061  df-mo 2615  df-eu 2647  df-clab 2797  df-cleq 2811  df-clel 2890  df-nfc 2960  df-ne 3014  df-nel 3121  df-ral 3140  df-rex 3141  df-rmo 3143  df-rab 3144  df-v 3494  df-sbc 3770  df-csb 3881  df-dif 3936  df-un 3938  df-in 3940  df-ss 3949  df-nul 4289  df-if 4464  df-pw 4537  df-sn 4558  df-pr 4560  df-op 4564  df-uni 4831  df-br 5058  df-opab 5120  df-mpt 5138  df-id 5453  df-po 5467  df-so 5468  df-xp 5554  df-rel 5555  df-cnv 5556  df-co 5557  df-dm 5558  df-rn 5559  df-res 5560  df-ima 5561  df-iota 6307  df-fun 6350  df-fn 6351  df-f 6352  df-f1 6353  df-fo 6354  df-f1o 6355  df-fv 6356  df-ov 7148  df-oprab 7149  df-mpo 7150  df-er 8278  df-en 8498  df-dom 8499  df-sdom 8500  df-sup 8894  df-inf 8895  df-pnf 10665  df-mnf 10666  df-ltxr 10668  df-lcm 15922
This theorem is referenced by:  lcmcllem  15928  lcmledvds  15931  lcmgcdlem  15938
  Copyright terms: Public domain W3C validator