| Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > MPE Home > Th. List > lcmfval | Structured version Visualization version GIF version | ||
| Description: Value of the lcm function. (lcm‘𝑍) is the least common multiple of the integers contained in the finite subset of integers 𝑍. If at least one of the elements of 𝑍 is 0, the result is defined conventionally as 0. (Contributed by AV, 21-Apr-2020.) (Revised by AV, 16-Sep-2020.) |
| Ref | Expression |
|---|---|
| lcmfval | ⊢ ((𝑍 ⊆ ℤ ∧ 𝑍 ∈ Fin) → (lcm‘𝑍) = if(0 ∈ 𝑍, 0, inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}, ℝ, < ))) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | df-lcmf 16608 | . 2 ⊢ lcm = (𝑧 ∈ 𝒫 ℤ ↦ if(0 ∈ 𝑧, 0, inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑧 𝑚 ∥ 𝑛}, ℝ, < ))) | |
| 2 | eleq2 2823 | . . 3 ⊢ (𝑧 = 𝑍 → (0 ∈ 𝑧 ↔ 0 ∈ 𝑍)) | |
| 3 | raleq 3302 | . . . . 5 ⊢ (𝑧 = 𝑍 → (∀𝑚 ∈ 𝑧 𝑚 ∥ 𝑛 ↔ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛)) | |
| 4 | 3 | rabbidv 3423 | . . . 4 ⊢ (𝑧 = 𝑍 → {𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑧 𝑚 ∥ 𝑛} = {𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}) |
| 5 | 4 | infeq1d 9488 | . . 3 ⊢ (𝑧 = 𝑍 → inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑧 𝑚 ∥ 𝑛}, ℝ, < ) = inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}, ℝ, < )) |
| 6 | 2, 5 | ifbieq2d 4527 | . 2 ⊢ (𝑧 = 𝑍 → if(0 ∈ 𝑧, 0, inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑧 𝑚 ∥ 𝑛}, ℝ, < )) = if(0 ∈ 𝑍, 0, inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}, ℝ, < ))) |
| 7 | zex 12595 | . . . . . 6 ⊢ ℤ ∈ V | |
| 8 | 7 | ssex 5291 | . . . . 5 ⊢ (𝑍 ⊆ ℤ → 𝑍 ∈ V) |
| 9 | elpwg 4578 | . . . . 5 ⊢ (𝑍 ∈ V → (𝑍 ∈ 𝒫 ℤ ↔ 𝑍 ⊆ ℤ)) | |
| 10 | 8, 9 | syl 17 | . . . 4 ⊢ (𝑍 ⊆ ℤ → (𝑍 ∈ 𝒫 ℤ ↔ 𝑍 ⊆ ℤ)) |
| 11 | 10 | ibir 268 | . . 3 ⊢ (𝑍 ⊆ ℤ → 𝑍 ∈ 𝒫 ℤ) |
| 12 | 11 | adantr 480 | . 2 ⊢ ((𝑍 ⊆ ℤ ∧ 𝑍 ∈ Fin) → 𝑍 ∈ 𝒫 ℤ) |
| 13 | 0nn0 12514 | . . . 4 ⊢ 0 ∈ ℕ0 | |
| 14 | 13 | a1i 11 | . . 3 ⊢ (((𝑍 ⊆ ℤ ∧ 𝑍 ∈ Fin) ∧ 0 ∈ 𝑍) → 0 ∈ ℕ0) |
| 15 | df-nel 3037 | . . . 4 ⊢ (0 ∉ 𝑍 ↔ ¬ 0 ∈ 𝑍) | |
| 16 | ssrab2 4055 | . . . . . 6 ⊢ {𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛} ⊆ ℕ | |
| 17 | nnssnn0 12502 | . . . . . 6 ⊢ ℕ ⊆ ℕ0 | |
| 18 | 16, 17 | sstri 3968 | . . . . 5 ⊢ {𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛} ⊆ ℕ0 |
| 19 | nnuz 12893 | . . . . . . 7 ⊢ ℕ = (ℤ≥‘1) | |
| 20 | 16, 19 | sseqtri 4007 | . . . . . 6 ⊢ {𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛} ⊆ (ℤ≥‘1) |
| 21 | fissn0dvdsn0 16637 | . . . . . . 7 ⊢ ((𝑍 ⊆ ℤ ∧ 𝑍 ∈ Fin ∧ 0 ∉ 𝑍) → {𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛} ≠ ∅) | |
| 22 | 21 | 3expa 1118 | . . . . . 6 ⊢ (((𝑍 ⊆ ℤ ∧ 𝑍 ∈ Fin) ∧ 0 ∉ 𝑍) → {𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛} ≠ ∅) |
| 23 | infssuzcl 12946 | . . . . . 6 ⊢ (({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛} ⊆ (ℤ≥‘1) ∧ {𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛} ≠ ∅) → inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}, ℝ, < ) ∈ {𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}) | |
| 24 | 20, 22, 23 | sylancr 587 | . . . . 5 ⊢ (((𝑍 ⊆ ℤ ∧ 𝑍 ∈ Fin) ∧ 0 ∉ 𝑍) → inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}, ℝ, < ) ∈ {𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}) |
| 25 | 18, 24 | sselid 3956 | . . . 4 ⊢ (((𝑍 ⊆ ℤ ∧ 𝑍 ∈ Fin) ∧ 0 ∉ 𝑍) → inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}, ℝ, < ) ∈ ℕ0) |
| 26 | 15, 25 | sylan2br 595 | . . 3 ⊢ (((𝑍 ⊆ ℤ ∧ 𝑍 ∈ Fin) ∧ ¬ 0 ∈ 𝑍) → inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}, ℝ, < ) ∈ ℕ0) |
| 27 | 14, 26 | ifclda 4536 | . 2 ⊢ ((𝑍 ⊆ ℤ ∧ 𝑍 ∈ Fin) → if(0 ∈ 𝑍, 0, inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}, ℝ, < )) ∈ ℕ0) |
| 28 | 1, 6, 12, 27 | fvmptd3 7008 | 1 ⊢ ((𝑍 ⊆ ℤ ∧ 𝑍 ∈ Fin) → (lcm‘𝑍) = if(0 ∈ 𝑍, 0, inf({𝑛 ∈ ℕ ∣ ∀𝑚 ∈ 𝑍 𝑚 ∥ 𝑛}, ℝ, < ))) |
| Colors of variables: wff setvar class |
| Syntax hints: ¬ wn 3 → wi 4 ↔ wb 206 ∧ wa 395 = wceq 1540 ∈ wcel 2108 ≠ wne 2932 ∉ wnel 3036 ∀wral 3051 {crab 3415 Vcvv 3459 ⊆ wss 3926 ∅c0 4308 ifcif 4500 𝒫 cpw 4575 class class class wbr 5119 ‘cfv 6530 Fincfn 8957 infcinf 9451 ℝcr 11126 0cc0 11127 1c1 11128 < clt 11267 ℕcn 12238 ℕ0cn0 12499 ℤcz 12586 ℤ≥cuz 12850 ∥ cdvds 16270 lcmclcmf 16606 |
| This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1795 ax-4 1809 ax-5 1910 ax-6 1967 ax-7 2007 ax-8 2110 ax-9 2118 ax-10 2141 ax-11 2157 ax-12 2177 ax-ext 2707 ax-rep 5249 ax-sep 5266 ax-nul 5276 ax-pow 5335 ax-pr 5402 ax-un 7727 ax-inf2 9653 ax-cnex 11183 ax-resscn 11184 ax-1cn 11185 ax-icn 11186 ax-addcl 11187 ax-addrcl 11188 ax-mulcl 11189 ax-mulrcl 11190 ax-mulcom 11191 ax-addass 11192 ax-mulass 11193 ax-distr 11194 ax-i2m1 11195 ax-1ne0 11196 ax-1rid 11197 ax-rnegex 11198 ax-rrecex 11199 ax-cnre 11200 ax-pre-lttri 11201 ax-pre-lttrn 11202 ax-pre-ltadd 11203 ax-pre-mulgt0 11204 ax-pre-sup 11205 |
| This theorem depends on definitions: df-bi 207 df-an 396 df-or 848 df-3or 1087 df-3an 1088 df-tru 1543 df-fal 1553 df-ex 1780 df-nf 1784 df-sb 2065 df-mo 2539 df-eu 2568 df-clab 2714 df-cleq 2727 df-clel 2809 df-nfc 2885 df-ne 2933 df-nel 3037 df-ral 3052 df-rex 3061 df-rmo 3359 df-reu 3360 df-rab 3416 df-v 3461 df-sbc 3766 df-csb 3875 df-dif 3929 df-un 3931 df-in 3933 df-ss 3943 df-pss 3946 df-nul 4309 df-if 4501 df-pw 4577 df-sn 4602 df-pr 4604 df-op 4608 df-uni 4884 df-int 4923 df-iun 4969 df-br 5120 df-opab 5182 df-mpt 5202 df-tr 5230 df-id 5548 df-eprel 5553 df-po 5561 df-so 5562 df-fr 5606 df-se 5607 df-we 5608 df-xp 5660 df-rel 5661 df-cnv 5662 df-co 5663 df-dm 5664 df-rn 5665 df-res 5666 df-ima 5667 df-pred 6290 df-ord 6355 df-on 6356 df-lim 6357 df-suc 6358 df-iota 6483 df-fun 6532 df-fn 6533 df-f 6534 df-f1 6535 df-fo 6536 df-f1o 6537 df-fv 6538 df-isom 6539 df-riota 7360 df-ov 7406 df-oprab 7407 df-mpo 7408 df-om 7860 df-1st 7986 df-2nd 7987 df-frecs 8278 df-wrecs 8309 df-recs 8383 df-rdg 8422 df-1o 8478 df-er 8717 df-en 8958 df-dom 8959 df-sdom 8960 df-fin 8961 df-sup 9452 df-inf 9453 df-oi 9522 df-card 9951 df-pnf 11269 df-mnf 11270 df-xr 11271 df-ltxr 11272 df-le 11273 df-sub 11466 df-neg 11467 df-div 11893 df-nn 12239 df-2 12301 df-3 12302 df-n0 12500 df-z 12587 df-uz 12851 df-rp 13007 df-fz 13523 df-fzo 13670 df-seq 14018 df-exp 14078 df-hash 14347 df-cj 15116 df-re 15117 df-im 15118 df-sqrt 15252 df-abs 15253 df-clim 15502 df-prod 15918 df-dvds 16271 df-lcmf 16608 |
| This theorem is referenced by: lcmfn0val 16640 lcmfpr 16644 |
| Copyright terms: Public domain | W3C validator |