Metamath Proof Explorer |
< Previous
Next >
Nearby theorems |
||
Mirrors > Home > MPE Home > Th. List > rpnnen2lem1 | Structured version Visualization version GIF version |
Description: Lemma for rpnnen2 15933. (Contributed by Mario Carneiro, 13-May-2013.) |
Ref | Expression |
---|---|
rpnnen2.1 | ⊢ 𝐹 = (𝑥 ∈ 𝒫 ℕ ↦ (𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝑥, ((1 / 3)↑𝑛), 0))) |
Ref | Expression |
---|---|
rpnnen2lem1 | ⊢ ((𝐴 ⊆ ℕ ∧ 𝑁 ∈ ℕ) → ((𝐹‘𝐴)‘𝑁) = if(𝑁 ∈ 𝐴, ((1 / 3)↑𝑁), 0)) |
Step | Hyp | Ref | Expression |
---|---|---|---|
1 | nnex 11979 | . . . . 5 ⊢ ℕ ∈ V | |
2 | 1 | elpw2 5273 | . . . 4 ⊢ (𝐴 ∈ 𝒫 ℕ ↔ 𝐴 ⊆ ℕ) |
3 | eleq2 2829 | . . . . . . 7 ⊢ (𝑥 = 𝐴 → (𝑛 ∈ 𝑥 ↔ 𝑛 ∈ 𝐴)) | |
4 | 3 | ifbid 4488 | . . . . . 6 ⊢ (𝑥 = 𝐴 → if(𝑛 ∈ 𝑥, ((1 / 3)↑𝑛), 0) = if(𝑛 ∈ 𝐴, ((1 / 3)↑𝑛), 0)) |
5 | 4 | mpteq2dv 5181 | . . . . 5 ⊢ (𝑥 = 𝐴 → (𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝑥, ((1 / 3)↑𝑛), 0)) = (𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝐴, ((1 / 3)↑𝑛), 0))) |
6 | rpnnen2.1 | . . . . 5 ⊢ 𝐹 = (𝑥 ∈ 𝒫 ℕ ↦ (𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝑥, ((1 / 3)↑𝑛), 0))) | |
7 | 1 | mptex 7096 | . . . . 5 ⊢ (𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝐴, ((1 / 3)↑𝑛), 0)) ∈ V |
8 | 5, 6, 7 | fvmpt 6872 | . . . 4 ⊢ (𝐴 ∈ 𝒫 ℕ → (𝐹‘𝐴) = (𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝐴, ((1 / 3)↑𝑛), 0))) |
9 | 2, 8 | sylbir 234 | . . 3 ⊢ (𝐴 ⊆ ℕ → (𝐹‘𝐴) = (𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝐴, ((1 / 3)↑𝑛), 0))) |
10 | 9 | fveq1d 6773 | . 2 ⊢ (𝐴 ⊆ ℕ → ((𝐹‘𝐴)‘𝑁) = ((𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝐴, ((1 / 3)↑𝑛), 0))‘𝑁)) |
11 | eleq1 2828 | . . . 4 ⊢ (𝑛 = 𝑁 → (𝑛 ∈ 𝐴 ↔ 𝑁 ∈ 𝐴)) | |
12 | oveq2 7279 | . . . 4 ⊢ (𝑛 = 𝑁 → ((1 / 3)↑𝑛) = ((1 / 3)↑𝑁)) | |
13 | 11, 12 | ifbieq1d 4489 | . . 3 ⊢ (𝑛 = 𝑁 → if(𝑛 ∈ 𝐴, ((1 / 3)↑𝑛), 0) = if(𝑁 ∈ 𝐴, ((1 / 3)↑𝑁), 0)) |
14 | eqid 2740 | . . 3 ⊢ (𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝐴, ((1 / 3)↑𝑛), 0)) = (𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝐴, ((1 / 3)↑𝑛), 0)) | |
15 | ovex 7304 | . . . 4 ⊢ ((1 / 3)↑𝑁) ∈ V | |
16 | c0ex 10970 | . . . 4 ⊢ 0 ∈ V | |
17 | 15, 16 | ifex 4515 | . . 3 ⊢ if(𝑁 ∈ 𝐴, ((1 / 3)↑𝑁), 0) ∈ V |
18 | 13, 14, 17 | fvmpt 6872 | . 2 ⊢ (𝑁 ∈ ℕ → ((𝑛 ∈ ℕ ↦ if(𝑛 ∈ 𝐴, ((1 / 3)↑𝑛), 0))‘𝑁) = if(𝑁 ∈ 𝐴, ((1 / 3)↑𝑁), 0)) |
19 | 10, 18 | sylan9eq 2800 | 1 ⊢ ((𝐴 ⊆ ℕ ∧ 𝑁 ∈ ℕ) → ((𝐹‘𝐴)‘𝑁) = if(𝑁 ∈ 𝐴, ((1 / 3)↑𝑁), 0)) |
Colors of variables: wff setvar class |
Syntax hints: → wi 4 ∧ wa 396 = wceq 1542 ∈ wcel 2110 ⊆ wss 3892 ifcif 4465 𝒫 cpw 4539 ↦ cmpt 5162 ‘cfv 6432 (class class class)co 7271 0cc0 10872 1c1 10873 / cdiv 11632 ℕcn 11973 3c3 12029 ↑cexp 13780 |
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 1917 ax-6 1975 ax-7 2015 ax-8 2112 ax-9 2120 ax-10 2141 ax-11 2158 ax-12 2175 ax-ext 2711 ax-rep 5214 ax-sep 5227 ax-nul 5234 ax-pr 5356 ax-un 7582 ax-cnex 10928 ax-1cn 10930 ax-icn 10931 ax-addcl 10932 ax-mulcl 10934 ax-i2m1 10940 |
This theorem depends on definitions: df-bi 206 df-an 397 df-or 845 df-3or 1087 df-3an 1088 df-tru 1545 df-fal 1555 df-ex 1787 df-nf 1791 df-sb 2072 df-mo 2542 df-eu 2571 df-clab 2718 df-cleq 2732 df-clel 2818 df-nfc 2891 df-ne 2946 df-ral 3071 df-rex 3072 df-reu 3073 df-rab 3075 df-v 3433 df-sbc 3721 df-csb 3838 df-dif 3895 df-un 3897 df-in 3899 df-ss 3909 df-pss 3911 df-nul 4263 df-if 4466 df-pw 4541 df-sn 4568 df-pr 4570 df-op 4574 df-uni 4846 df-iun 4932 df-br 5080 df-opab 5142 df-mpt 5163 df-tr 5197 df-id 5490 df-eprel 5496 df-po 5504 df-so 5505 df-fr 5545 df-we 5547 df-xp 5596 df-rel 5597 df-cnv 5598 df-co 5599 df-dm 5600 df-rn 5601 df-res 5602 df-ima 5603 df-pred 6201 df-ord 6268 df-on 6269 df-lim 6270 df-suc 6271 df-iota 6390 df-fun 6434 df-fn 6435 df-f 6436 df-f1 6437 df-fo 6438 df-f1o 6439 df-fv 6440 df-ov 7274 df-om 7707 df-2nd 7825 df-frecs 8088 df-wrecs 8119 df-recs 8193 df-rdg 8232 df-nn 11974 |
This theorem is referenced by: rpnnen2lem3 15923 rpnnen2lem4 15924 rpnnen2lem9 15929 rpnnen2lem10 15930 rpnnen2lem11 15931 |
Copyright terms: Public domain | W3C validator |