Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
Mirrors > Home > MPE Home > Th. List > flbi2 | Structured version Visualization version GIF version |
Description: A condition equivalent to floor. (Contributed by NM, 15-Aug-2008.) |
Ref | Expression |
---|---|
flbi2 | ⊢ ((𝑁 ∈ ℤ ∧ 𝐹 ∈ ℝ) → ((⌊‘(𝑁 + 𝐹)) = 𝑁 ↔ (0 ≤ 𝐹 ∧ 𝐹 < 1))) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | zre 12059 | . . . 4 ⊢ (𝑁 ∈ ℤ → 𝑁 ∈ ℝ) | |
2 | readdcl 10691 | . . . 4 ⊢ ((𝑁 ∈ ℝ ∧ 𝐹 ∈ ℝ) → (𝑁 + 𝐹) ∈ ℝ) | |
3 | 1, 2 | sylan 583 | . . 3 ⊢ ((𝑁 ∈ ℤ ∧ 𝐹 ∈ ℝ) → (𝑁 + 𝐹) ∈ ℝ) |
4 | simpl 486 | . . 3 ⊢ ((𝑁 ∈ ℤ ∧ 𝐹 ∈ ℝ) → 𝑁 ∈ ℤ) | |
5 | flbi 13270 | . . 3 ⊢ (((𝑁 + 𝐹) ∈ ℝ ∧ 𝑁 ∈ ℤ) → ((⌊‘(𝑁 + 𝐹)) = 𝑁 ↔ (𝑁 ≤ (𝑁 + 𝐹) ∧ (𝑁 + 𝐹) < (𝑁 + 1)))) | |
6 | 3, 4, 5 | syl2anc 587 | . 2 ⊢ ((𝑁 ∈ ℤ ∧ 𝐹 ∈ ℝ) → ((⌊‘(𝑁 + 𝐹)) = 𝑁 ↔ (𝑁 ≤ (𝑁 + 𝐹) ∧ (𝑁 + 𝐹) < (𝑁 + 1)))) |
7 | addge01 11221 | . . . 4 ⊢ ((𝑁 ∈ ℝ ∧ 𝐹 ∈ ℝ) → (0 ≤ 𝐹 ↔ 𝑁 ≤ (𝑁 + 𝐹))) | |
8 | 1re 10712 | . . . . . 6 ⊢ 1 ∈ ℝ | |
9 | ltadd2 10815 | . . . . . 6 ⊢ ((𝐹 ∈ ℝ ∧ 1 ∈ ℝ ∧ 𝑁 ∈ ℝ) → (𝐹 < 1 ↔ (𝑁 + 𝐹) < (𝑁 + 1))) | |
10 | 8, 9 | mp3an2 1450 | . . . . 5 ⊢ ((𝐹 ∈ ℝ ∧ 𝑁 ∈ ℝ) → (𝐹 < 1 ↔ (𝑁 + 𝐹) < (𝑁 + 1))) |
11 | 10 | ancoms 462 | . . . 4 ⊢ ((𝑁 ∈ ℝ ∧ 𝐹 ∈ ℝ) → (𝐹 < 1 ↔ (𝑁 + 𝐹) < (𝑁 + 1))) |
12 | 7, 11 | anbi12d 634 | . . 3 ⊢ ((𝑁 ∈ ℝ ∧ 𝐹 ∈ ℝ) → ((0 ≤ 𝐹 ∧ 𝐹 < 1) ↔ (𝑁 ≤ (𝑁 + 𝐹) ∧ (𝑁 + 𝐹) < (𝑁 + 1)))) |
13 | 1, 12 | sylan 583 | . 2 ⊢ ((𝑁 ∈ ℤ ∧ 𝐹 ∈ ℝ) → ((0 ≤ 𝐹 ∧ 𝐹 < 1) ↔ (𝑁 ≤ (𝑁 + 𝐹) ∧ (𝑁 + 𝐹) < (𝑁 + 1)))) |
14 | 6, 13 | bitr4d 285 | 1 ⊢ ((𝑁 ∈ ℤ ∧ 𝐹 ∈ ℝ) → ((⌊‘(𝑁 + 𝐹)) = 𝑁 ↔ (0 ≤ 𝐹 ∧ 𝐹 < 1))) |
Colors of variables: wff setvar class |
Syntax hints: → wi 4 ↔ wb 209 ∧ wa 399 = wceq 1542 ∈ wcel 2113 class class class wbr 5027 ‘cfv 6333 (class class class)co 7164 ℝcr 10607 0cc0 10608 1c1 10609 + caddc 10611 < clt 10746 ≤ cle 10747 ℤcz 12055 ⌊cfl 13244 |
This theorem was proved from axioms: ax-mp 5 ax-1 6 ax-2 7 ax-3 8 ax-gen 1802 ax-4 1816 ax-5 1916 ax-6 1974 ax-7 2019 ax-8 2115 ax-9 2123 ax-10 2144 ax-11 2161 ax-12 2178 ax-ext 2710 ax-sep 5164 ax-nul 5171 ax-pow 5229 ax-pr 5293 ax-un 7473 ax-cnex 10664 ax-resscn 10665 ax-1cn 10666 ax-icn 10667 ax-addcl 10668 ax-addrcl 10669 ax-mulcl 10670 ax-mulrcl 10671 ax-mulcom 10672 ax-addass 10673 ax-mulass 10674 ax-distr 10675 ax-i2m1 10676 ax-1ne0 10677 ax-1rid 10678 ax-rnegex 10679 ax-rrecex 10680 ax-cnre 10681 ax-pre-lttri 10682 ax-pre-lttrn 10683 ax-pre-ltadd 10684 ax-pre-mulgt0 10685 ax-pre-sup 10686 |
This theorem depends on definitions: df-bi 210 df-an 400 df-or 847 df-3or 1089 df-3an 1090 df-tru 1545 df-fal 1555 df-ex 1787 df-nf 1791 df-sb 2074 df-mo 2540 df-eu 2570 df-clab 2717 df-cleq 2730 df-clel 2811 df-nfc 2881 df-ne 2935 df-nel 3039 df-ral 3058 df-rex 3059 df-reu 3060 df-rmo 3061 df-rab 3062 df-v 3399 df-sbc 3680 df-csb 3789 df-dif 3844 df-un 3846 df-in 3848 df-ss 3858 df-pss 3860 df-nul 4210 df-if 4412 df-pw 4487 df-sn 4514 df-pr 4516 df-tp 4518 df-op 4520 df-uni 4794 df-iun 4880 df-br 5028 df-opab 5090 df-mpt 5108 df-tr 5134 df-id 5425 df-eprel 5430 df-po 5438 df-so 5439 df-fr 5478 df-we 5480 df-xp 5525 df-rel 5526 df-cnv 5527 df-co 5528 df-dm 5529 df-rn 5530 df-res 5531 df-ima 5532 df-pred 6123 df-ord 6169 df-on 6170 df-lim 6171 df-suc 6172 df-iota 6291 df-fun 6335 df-fn 6336 df-f 6337 df-f1 6338 df-fo 6339 df-f1o 6340 df-fv 6341 df-riota 7121 df-ov 7167 df-oprab 7168 df-mpo 7169 df-om 7594 df-wrecs 7969 df-recs 8030 df-rdg 8068 df-er 8313 df-en 8549 df-dom 8550 df-sdom 8551 df-sup 8972 df-inf 8973 df-pnf 10748 df-mnf 10749 df-xr 10750 df-ltxr 10751 df-le 10752 df-sub 10943 df-neg 10944 df-nn 11710 df-n0 11970 df-z 12056 df-uz 12318 df-fl 13246 |
This theorem is referenced by: adddivflid 13272 ico01fl0 13273 divfl0 13278 fldiv4p1lem1div2 13289 fldiv 13312 modid 13348 flodddiv4 15851 bitsp1o 15869 fldivp1 16326 fourierdlem26 43200 zofldiv2ALTV 44632 zofldiv2 45395 |
Copyright terms: Public domain | W3C validator |