Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
Mirrors > Home > MPE Home > Th. List > quoremnn0 | Structured version Visualization version GIF version |
Description: Quotient and remainder of a nonnegative integer divided by a positive integer. (Contributed by NM, 14-Aug-2008.) |
Ref | Expression |
---|---|
quorem.1 | ⊢ 𝑄 = (⌊‘(𝐴 / 𝐵)) |
quorem.2 | ⊢ 𝑅 = (𝐴 − (𝐵 · 𝑄)) |
Ref | Expression |
---|---|
quoremnn0 | ⊢ ((𝐴 ∈ ℕ0 ∧ 𝐵 ∈ ℕ) → ((𝑄 ∈ ℕ0 ∧ 𝑅 ∈ ℕ0) ∧ (𝑅 < 𝐵 ∧ 𝐴 = ((𝐵 · 𝑄) + 𝑅)))) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | quorem.1 | . . 3 ⊢ 𝑄 = (⌊‘(𝐴 / 𝐵)) | |
2 | fldivnn0 13635 | . . 3 ⊢ ((𝐴 ∈ ℕ0 ∧ 𝐵 ∈ ℕ) → (⌊‘(𝐴 / 𝐵)) ∈ ℕ0) | |
3 | 1, 2 | eqeltrid 2841 | . 2 ⊢ ((𝐴 ∈ ℕ0 ∧ 𝐵 ∈ ℕ) → 𝑄 ∈ ℕ0) |
4 | nn0z 12436 | . . 3 ⊢ (𝐴 ∈ ℕ0 → 𝐴 ∈ ℤ) | |
5 | quorem.2 | . . . 4 ⊢ 𝑅 = (𝐴 − (𝐵 · 𝑄)) | |
6 | 1, 5 | quoremz 13668 | . . 3 ⊢ ((𝐴 ∈ ℤ ∧ 𝐵 ∈ ℕ) → ((𝑄 ∈ ℤ ∧ 𝑅 ∈ ℕ0) ∧ (𝑅 < 𝐵 ∧ 𝐴 = ((𝐵 · 𝑄) + 𝑅)))) |
7 | 4, 6 | sylan 580 | . 2 ⊢ ((𝐴 ∈ ℕ0 ∧ 𝐵 ∈ ℕ) → ((𝑄 ∈ ℤ ∧ 𝑅 ∈ ℕ0) ∧ (𝑅 < 𝐵 ∧ 𝐴 = ((𝐵 · 𝑄) + 𝑅)))) |
8 | simpl 483 | . . . . . 6 ⊢ ((𝑄 ∈ ℕ0 ∧ 𝑄 ∈ ℤ) → 𝑄 ∈ ℕ0) | |
9 | 8 | anim1i 615 | . . . . 5 ⊢ (((𝑄 ∈ ℕ0 ∧ 𝑄 ∈ ℤ) ∧ 𝑅 ∈ ℕ0) → (𝑄 ∈ ℕ0 ∧ 𝑅 ∈ ℕ0)) |
10 | 9 | anasss 467 | . . . 4 ⊢ ((𝑄 ∈ ℕ0 ∧ (𝑄 ∈ ℤ ∧ 𝑅 ∈ ℕ0)) → (𝑄 ∈ ℕ0 ∧ 𝑅 ∈ ℕ0)) |
11 | 10 | anim1i 615 | . . 3 ⊢ (((𝑄 ∈ ℕ0 ∧ (𝑄 ∈ ℤ ∧ 𝑅 ∈ ℕ0)) ∧ (𝑅 < 𝐵 ∧ 𝐴 = ((𝐵 · 𝑄) + 𝑅))) → ((𝑄 ∈ ℕ0 ∧ 𝑅 ∈ ℕ0) ∧ (𝑅 < 𝐵 ∧ 𝐴 = ((𝐵 · 𝑄) + 𝑅)))) |
12 | 11 | anasss 467 | . 2 ⊢ ((𝑄 ∈ ℕ0 ∧ ((𝑄 ∈ ℤ ∧ 𝑅 ∈ ℕ0) ∧ (𝑅 < 𝐵 ∧ 𝐴 = ((𝐵 · 𝑄) + 𝑅)))) → ((𝑄 ∈ ℕ0 ∧ 𝑅 ∈ ℕ0) ∧ (𝑅 < 𝐵 ∧ 𝐴 = ((𝐵 · 𝑄) + 𝑅)))) |
13 | 3, 7, 12 | syl2anc 584 | 1 ⊢ ((𝐴 ∈ ℕ0 ∧ 𝐵 ∈ ℕ) → ((𝑄 ∈ ℕ0 ∧ 𝑅 ∈ ℕ0) ∧ (𝑅 < 𝐵 ∧ 𝐴 = ((𝐵 · 𝑄) + 𝑅)))) |
Colors of variables: wff setvar class |
Syntax hints: → wi 4 ∧ wa 396 = wceq 1540 ∈ wcel 2105 class class class wbr 5089 ‘cfv 6473 (class class class)co 7329 + caddc 10967 · cmul 10969 < clt 11102 − cmin 11298 / cdiv 11725 ℕcn 12066 ℕ0cn0 12326 ℤcz 12412 ⌊cfl 13603 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1796 ax-4 1810 ax-5 1912 ax-6 1970 ax-7 2010 ax-8 2107 ax-9 2115 ax-10 2136 ax-11 2153 ax-12 2170 ax-ext 2707 ax-sep 5240 ax-nul 5247 ax-pow 5305 ax-pr 5369 ax-un 7642 ax-cnex 11020 ax-resscn 11021 ax-1cn 11022 ax-icn 11023 ax-addcl 11024 ax-addrcl 11025 ax-mulcl 11026 ax-mulrcl 11027 ax-mulcom 11028 ax-addass 11029 ax-mulass 11030 ax-distr 11031 ax-i2m1 11032 ax-1ne0 11033 ax-1rid 11034 ax-rnegex 11035 ax-rrecex 11036 ax-cnre 11037 ax-pre-lttri 11038 ax-pre-lttrn 11039 ax-pre-ltadd 11040 ax-pre-mulgt0 11041 ax-pre-sup 11042 |
This theorem depends on definitions: df-bi 206 df-an 397 df-or 845 df-3or 1087 df-3an 1088 df-tru 1543 df-fal 1553 df-ex 1781 df-nf 1785 df-sb 2067 df-mo 2538 df-eu 2567 df-clab 2714 df-cleq 2728 df-clel 2814 df-nfc 2886 df-ne 2941 df-nel 3047 df-ral 3062 df-rex 3071 df-rmo 3349 df-reu 3350 df-rab 3404 df-v 3443 df-sbc 3727 df-csb 3843 df-dif 3900 df-un 3902 df-in 3904 df-ss 3914 df-pss 3916 df-nul 4269 df-if 4473 df-pw 4548 df-sn 4573 df-pr 4575 df-op 4579 df-uni 4852 df-iun 4940 df-br 5090 df-opab 5152 df-mpt 5173 df-tr 5207 df-id 5512 df-eprel 5518 df-po 5526 df-so 5527 df-fr 5569 df-we 5571 df-xp 5620 df-rel 5621 df-cnv 5622 df-co 5623 df-dm 5624 df-rn 5625 df-res 5626 df-ima 5627 df-pred 6232 df-ord 6299 df-on 6300 df-lim 6301 df-suc 6302 df-iota 6425 df-fun 6475 df-fn 6476 df-f 6477 df-f1 6478 df-fo 6479 df-f1o 6480 df-fv 6481 df-riota 7286 df-ov 7332 df-oprab 7333 df-mpo 7334 df-om 7773 df-2nd 7892 df-frecs 8159 df-wrecs 8190 df-recs 8264 df-rdg 8303 df-er 8561 df-en 8797 df-dom 8798 df-sdom 8799 df-sup 9291 df-inf 9292 df-pnf 11104 df-mnf 11105 df-xr 11106 df-ltxr 11107 df-le 11108 df-sub 11300 df-neg 11301 df-div 11726 df-nn 12067 df-n0 12327 df-z 12413 df-uz 12676 df-fl 13605 |
This theorem is referenced by: (None) |
Copyright terms: Public domain | W3C validator |