| Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
| Mirrors > Home > MPE Home > Th. List > zmodcl | Structured version Visualization version GIF version | ||
| Description: Closure law for the modulo operation restricted to integers. (Contributed by NM, 27-Nov-2008.) |
| Ref | Expression |
|---|---|
| zmodcl | ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → (𝐴 mod 𝐵) ∈ ℕ0) |
| Step | Hyp | Ref | Expression |
|---|---|---|---|
| 1 | zre 12592 | . . . 4 ⊢ (𝐴 ∈ ℤ → 𝐴 ∈ ℝ) | |
| 2 | nnrp 13020 | . . . 4 ⊢ (𝐵 ∈ ℕ → 𝐵 ∈ ℝ+) | |
| 3 | modval 13888 | . . . 4 ⊢ ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ+) → (𝐴 mod 𝐵) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵))))) | |
| 4 | 1, 2, 3 | syl2an 596 | . . 3 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → (𝐴 mod 𝐵) = (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵))))) |
| 5 | nnz 12609 | . . . . . 6 ⊢ (𝐵 ∈ ℕ → 𝐵 ∈ ℤ) | |
| 6 | 5 | adantl 481 | . . . . 5 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → 𝐵 ∈ ℤ) |
| 7 | nndivre 12281 | . . . . . . 7 ⊢ ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℕ) → (𝐴 / 𝐵) ∈ ℝ) | |
| 8 | 1, 7 | sylan 580 | . . . . . 6 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → (𝐴 / 𝐵) ∈ ℝ) |
| 9 | 8 | flcld 13815 | . . . . 5 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → (⌊‘(𝐴 / 𝐵)) ∈ ℤ) |
| 10 | 6, 9 | zmulcld 12703 | . . . 4 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → (𝐵 · (⌊‘(𝐴 / 𝐵))) ∈ ℤ) |
| 11 | zsubcl 12634 | . . . 4 ⊢ ((𝐴 ∈ ℤ ∧ (𝐵 · (⌊‘(𝐴 / 𝐵))) ∈ ℤ) → (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))) ∈ ℤ) | |
| 12 | 10, 11 | syldan 591 | . . 3 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → (𝐴 − (𝐵 · (⌊‘(𝐴 / 𝐵)))) ∈ ℤ) |
| 13 | 4, 12 | eqeltrd 2834 | . 2 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → (𝐴 mod 𝐵) ∈ ℤ) |
| 14 | modge0 13896 | . . 3 ⊢ ((𝐴 ∈ ℝ ∧ 𝐵 ∈ ℝ+) → 0 ≤ (𝐴 mod 𝐵)) | |
| 15 | 1, 2, 14 | syl2an 596 | . 2 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → 0 ≤ (𝐴 mod 𝐵)) |
| 16 | elnn0z 12601 | . 2 ⊢ ((𝐴 mod 𝐵) ∈ ℕ0 ↔ ((𝐴 mod 𝐵) ∈ ℤ ∧ 0 ≤ (𝐴 mod 𝐵))) | |
| 17 | 13, 15, 16 | sylanbrc 583 | 1 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → (𝐴 mod 𝐵) ∈ ℕ0) |
| Colors of variables: wff setvar class |
| Syntax hints: → wi 4 ∧ wa 395 = wceq 1540 ∈ wcel 2108 class class class wbr 5119 ‘cfv 6531 (class class class)co 7405 ℝcr 11128 0cc0 11129 · cmul 11134 ≤ cle 11270 − cmin 11466 / cdiv 11894 ℕcn 12240 ℕ0cn0 12501 ℤcz 12588 ℝ+crp 13008 ⌊cfl 13807 mod cmo 13886 |
| 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-sep 5266 ax-nul 5276 ax-pow 5335 ax-pr 5402 ax-un 7729 ax-cnex 11185 ax-resscn 11186 ax-1cn 11187 ax-icn 11188 ax-addcl 11189 ax-addrcl 11190 ax-mulcl 11191 ax-mulrcl 11192 ax-mulcom 11193 ax-addass 11194 ax-mulass 11195 ax-distr 11196 ax-i2m1 11197 ax-1ne0 11198 ax-1rid 11199 ax-rnegex 11200 ax-rrecex 11201 ax-cnre 11202 ax-pre-lttri 11203 ax-pre-lttrn 11204 ax-pre-ltadd 11205 ax-pre-mulgt0 11206 ax-pre-sup 11207 |
| 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-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-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 6484 df-fun 6533 df-fn 6534 df-f 6535 df-f1 6536 df-fo 6537 df-f1o 6538 df-fv 6539 df-riota 7362 df-ov 7408 df-oprab 7409 df-mpo 7410 df-om 7862 df-2nd 7989 df-frecs 8280 df-wrecs 8311 df-recs 8385 df-rdg 8424 df-er 8719 df-en 8960 df-dom 8961 df-sdom 8962 df-sup 9454 df-inf 9455 df-pnf 11271 df-mnf 11272 df-xr 11273 df-ltxr 11274 df-le 11275 df-sub 11468 df-neg 11469 df-div 11895 df-nn 12241 df-n0 12502 df-z 12589 df-uz 12853 df-rp 13009 df-fl 13809 df-mod 13887 |
| This theorem is referenced by: zmodcld 13909 zmodfz 13910 modaddmodup 13952 modaddmodlo 13953 cshwlen 14817 cshwidxmod 14821 repswcshw 14830 modfsummods 15809 divalgmod 16425 modgcd 16551 eucalgf 16602 eucalginv 16603 modprmn0modprm0 16827 fldivp1 16917 smndex1iidm 18879 odmodnn0 19521 gexdvds 19565 elqaalem2 26280 lgsmod 27286 dchrisumlem1 27452 congrep 42997 |
| Copyright terms: Public domain | W3C validator |