Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
Mirrors > Home > MPE Home > Th. List > bitsp1o | Structured version Visualization version GIF version |
Description: The 𝑀 + 1-th bit of 2𝑁 + 1 is the 𝑀-th bit of 𝑁. (Contributed by Mario Carneiro, 5-Sep-2016.) |
Ref | Expression |
---|---|
bitsp1o | ⊢ ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℕ0) → ((𝑀 + 1) ∈ (bits‘((2 · 𝑁) + 1)) ↔ 𝑀 ∈ (bits‘𝑁))) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | 2z 11286 | . . . . . 6 ⊢ 2 ∈ ℤ | |
2 | 1 | a1i 11 | . . . . 5 ⊢ (𝑁 ∈ ℤ → 2 ∈ ℤ) |
3 | id 22 | . . . . 5 ⊢ (𝑁 ∈ ℤ → 𝑁 ∈ ℤ) | |
4 | 2, 3 | zmulcld 11364 | . . . 4 ⊢ (𝑁 ∈ ℤ → (2 · 𝑁) ∈ ℤ) |
5 | 4 | peano2zd 11361 | . . 3 ⊢ (𝑁 ∈ ℤ → ((2 · 𝑁) + 1) ∈ ℤ) |
6 | bitsp1 14991 | . . 3 ⊢ ((((2 · 𝑁) + 1) ∈ ℤ ∧ 𝑀 ∈ ℕ0) → ((𝑀 + 1) ∈ (bits‘((2 · 𝑁) + 1)) ↔ 𝑀 ∈ (bits‘(⌊‘(((2 · 𝑁) + 1) / 2))))) | |
7 | 5, 6 | sylan 487 | . 2 ⊢ ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℕ0) → ((𝑀 + 1) ∈ (bits‘((2 · 𝑁) + 1)) ↔ 𝑀 ∈ (bits‘(⌊‘(((2 · 𝑁) + 1) / 2))))) |
8 | 2re 10967 | . . . . . . . . . . . 12 ⊢ 2 ∈ ℝ | |
9 | 8 | a1i 11 | . . . . . . . . . . 11 ⊢ (𝑁 ∈ ℤ → 2 ∈ ℝ) |
10 | zre 11258 | . . . . . . . . . . 11 ⊢ (𝑁 ∈ ℤ → 𝑁 ∈ ℝ) | |
11 | 9, 10 | remulcld 9949 | . . . . . . . . . 10 ⊢ (𝑁 ∈ ℤ → (2 · 𝑁) ∈ ℝ) |
12 | 11 | recnd 9947 | . . . . . . . . 9 ⊢ (𝑁 ∈ ℤ → (2 · 𝑁) ∈ ℂ) |
13 | 1cnd 9935 | . . . . . . . . 9 ⊢ (𝑁 ∈ ℤ → 1 ∈ ℂ) | |
14 | 2cnd 10970 | . . . . . . . . 9 ⊢ (𝑁 ∈ ℤ → 2 ∈ ℂ) | |
15 | 2ne0 10990 | . . . . . . . . . 10 ⊢ 2 ≠ 0 | |
16 | 15 | a1i 11 | . . . . . . . . 9 ⊢ (𝑁 ∈ ℤ → 2 ≠ 0) |
17 | 12, 13, 14, 16 | divdird 10718 | . . . . . . . 8 ⊢ (𝑁 ∈ ℤ → (((2 · 𝑁) + 1) / 2) = (((2 · 𝑁) / 2) + (1 / 2))) |
18 | zcn 11259 | . . . . . . . . . 10 ⊢ (𝑁 ∈ ℤ → 𝑁 ∈ ℂ) | |
19 | 18, 14, 16 | divcan3d 10685 | . . . . . . . . 9 ⊢ (𝑁 ∈ ℤ → ((2 · 𝑁) / 2) = 𝑁) |
20 | 19 | oveq1d 6564 | . . . . . . . 8 ⊢ (𝑁 ∈ ℤ → (((2 · 𝑁) / 2) + (1 / 2)) = (𝑁 + (1 / 2))) |
21 | 17, 20 | eqtrd 2644 | . . . . . . 7 ⊢ (𝑁 ∈ ℤ → (((2 · 𝑁) + 1) / 2) = (𝑁 + (1 / 2))) |
22 | 21 | fveq2d 6107 | . . . . . 6 ⊢ (𝑁 ∈ ℤ → (⌊‘(((2 · 𝑁) + 1) / 2)) = (⌊‘(𝑁 + (1 / 2)))) |
23 | 0re 9919 | . . . . . . . . 9 ⊢ 0 ∈ ℝ | |
24 | halfre 11123 | . . . . . . . . 9 ⊢ (1 / 2) ∈ ℝ | |
25 | halfgt0 11125 | . . . . . . . . 9 ⊢ 0 < (1 / 2) | |
26 | 23, 24, 25 | ltleii 10039 | . . . . . . . 8 ⊢ 0 ≤ (1 / 2) |
27 | halflt1 11127 | . . . . . . . 8 ⊢ (1 / 2) < 1 | |
28 | 26, 27 | pm3.2i 470 | . . . . . . 7 ⊢ (0 ≤ (1 / 2) ∧ (1 / 2) < 1) |
29 | flbi2 12480 | . . . . . . . 8 ⊢ ((𝑁 ∈ ℤ ∧ (1 / 2) ∈ ℝ) → ((⌊‘(𝑁 + (1 / 2))) = 𝑁 ↔ (0 ≤ (1 / 2) ∧ (1 / 2) < 1))) | |
30 | 24, 29 | mpan2 703 | . . . . . . 7 ⊢ (𝑁 ∈ ℤ → ((⌊‘(𝑁 + (1 / 2))) = 𝑁 ↔ (0 ≤ (1 / 2) ∧ (1 / 2) < 1))) |
31 | 28, 30 | mpbiri 247 | . . . . . 6 ⊢ (𝑁 ∈ ℤ → (⌊‘(𝑁 + (1 / 2))) = 𝑁) |
32 | 22, 31 | eqtrd 2644 | . . . . 5 ⊢ (𝑁 ∈ ℤ → (⌊‘(((2 · 𝑁) + 1) / 2)) = 𝑁) |
33 | 32 | adantr 480 | . . . 4 ⊢ ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℕ0) → (⌊‘(((2 · 𝑁) + 1) / 2)) = 𝑁) |
34 | 33 | fveq2d 6107 | . . 3 ⊢ ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℕ0) → (bits‘(⌊‘(((2 · 𝑁) + 1) / 2))) = (bits‘𝑁)) |
35 | 34 | eleq2d 2673 | . 2 ⊢ ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℕ0) → (𝑀 ∈ (bits‘(⌊‘(((2 · 𝑁) + 1) / 2))) ↔ 𝑀 ∈ (bits‘𝑁))) |
36 | 7, 35 | bitrd 267 | 1 ⊢ ((𝑁 ∈ ℤ ∧ 𝑀 ∈ ℕ0) → ((𝑀 + 1) ∈ (bits‘((2 · 𝑁) + 1)) ↔ 𝑀 ∈ (bits‘𝑁))) |
Colors of variables: wff setvar class |
Syntax hints: → wi 4 ↔ wb 195 ∧ wa 383 = wceq 1475 ∈ wcel 1977 ≠ wne 2780 class class class wbr 4583 ‘cfv 5804 (class class class)co 6549 ℝcr 9814 0cc0 9815 1c1 9816 + caddc 9818 · cmul 9820 < clt 9953 ≤ cle 9954 / cdiv 10563 2c2 10947 ℕ0cn0 11169 ℤcz 11254 ⌊cfl 12453 bitscbits 14979 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1713 ax-4 1728 ax-5 1827 ax-6 1875 ax-7 1922 ax-8 1979 ax-9 1986 ax-10 2006 ax-11 2021 ax-12 2034 ax-13 2234 ax-ext 2590 ax-sep 4709 ax-nul 4717 ax-pow 4769 ax-pr 4833 ax-un 6847 ax-cnex 9871 ax-resscn 9872 ax-1cn 9873 ax-icn 9874 ax-addcl 9875 ax-addrcl 9876 ax-mulcl 9877 ax-mulrcl 9878 ax-mulcom 9879 ax-addass 9880 ax-mulass 9881 ax-distr 9882 ax-i2m1 9883 ax-1ne0 9884 ax-1rid 9885 ax-rnegex 9886 ax-rrecex 9887 ax-cnre 9888 ax-pre-lttri 9889 ax-pre-lttrn 9890 ax-pre-ltadd 9891 ax-pre-mulgt0 9892 ax-pre-sup 9893 |
This theorem depends on definitions: df-bi 196 df-or 384 df-an 385 df-3or 1032 df-3an 1033 df-tru 1478 df-ex 1696 df-nf 1701 df-sb 1868 df-eu 2462 df-mo 2463 df-clab 2597 df-cleq 2603 df-clel 2606 df-nfc 2740 df-ne 2782 df-nel 2783 df-ral 2901 df-rex 2902 df-reu 2903 df-rmo 2904 df-rab 2905 df-v 3175 df-sbc 3403 df-csb 3500 df-dif 3543 df-un 3545 df-in 3547 df-ss 3554 df-pss 3556 df-nul 3875 df-if 4037 df-pw 4110 df-sn 4126 df-pr 4128 df-tp 4130 df-op 4132 df-uni 4373 df-iun 4457 df-br 4584 df-opab 4644 df-mpt 4645 df-tr 4681 df-eprel 4949 df-id 4953 df-po 4959 df-so 4960 df-fr 4997 df-we 4999 df-xp 5044 df-rel 5045 df-cnv 5046 df-co 5047 df-dm 5048 df-rn 5049 df-res 5050 df-ima 5051 df-pred 5597 df-ord 5643 df-on 5644 df-lim 5645 df-suc 5646 df-iota 5768 df-fun 5806 df-fn 5807 df-f 5808 df-f1 5809 df-fo 5810 df-f1o 5811 df-fv 5812 df-riota 6511 df-ov 6552 df-oprab 6553 df-mpt2 6554 df-om 6958 df-2nd 7060 df-wrecs 7294 df-recs 7355 df-rdg 7393 df-er 7629 df-en 7842 df-dom 7843 df-sdom 7844 df-sup 8231 df-inf 8232 df-pnf 9955 df-mnf 9956 df-xr 9957 df-ltxr 9958 df-le 9959 df-sub 10147 df-neg 10148 df-div 10564 df-nn 10898 df-2 10956 df-n0 11170 df-z 11255 df-uz 11564 df-fl 12455 df-seq 12664 df-exp 12723 df-bits 14982 |
This theorem is referenced by: (None) |
Copyright terms: Public domain | W3C validator |