MPE Home Metamath Proof Explorer < Previous   Next >
Nearby theorems
Mirrors  >  Home  >  MPE Home  >  Th. List  >  pcprod Structured version   Visualization version   GIF version

Theorem pcprod 16869
Description: The product of the primes taken to their respective powers reconstructs the original number. (Contributed by Mario Carneiro, 12-Mar-2014.)
Hypothesis
Ref Expression
pcprod.1 𝐹 = (𝑛 ∈ β„• ↦ if(𝑛 ∈ β„™, (𝑛↑(𝑛 pCnt 𝑁)), 1))
Assertion
Ref Expression
pcprod (𝑁 ∈ β„• β†’ (seq1( Β· , 𝐹)β€˜π‘) = 𝑁)
Distinct variable group:   𝑛,𝑁
Allowed substitution hint:   𝐹(𝑛)

Proof of Theorem pcprod
Dummy variable 𝑝 is distinct from all other variables.
StepHypRef Expression
1 pcprod.1 . . . . . 6 𝐹 = (𝑛 ∈ β„• ↦ if(𝑛 ∈ β„™, (𝑛↑(𝑛 pCnt 𝑁)), 1))
2 pccl 16823 . . . . . . . . 9 ((𝑛 ∈ β„™ ∧ 𝑁 ∈ β„•) β†’ (𝑛 pCnt 𝑁) ∈ β„•0)
32ancoms 457 . . . . . . . 8 ((𝑁 ∈ β„• ∧ 𝑛 ∈ β„™) β†’ (𝑛 pCnt 𝑁) ∈ β„•0)
43ralrimiva 3142 . . . . . . 7 (𝑁 ∈ β„• β†’ βˆ€π‘› ∈ β„™ (𝑛 pCnt 𝑁) ∈ β„•0)
54adantl 480 . . . . . 6 ((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) β†’ βˆ€π‘› ∈ β„™ (𝑛 pCnt 𝑁) ∈ β„•0)
6 simpr 483 . . . . . 6 ((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) β†’ 𝑁 ∈ β„•)
7 simpl 481 . . . . . 6 ((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) β†’ 𝑝 ∈ β„™)
8 oveq1 7431 . . . . . 6 (𝑛 = 𝑝 β†’ (𝑛 pCnt 𝑁) = (𝑝 pCnt 𝑁))
91, 5, 6, 7, 8pcmpt 16866 . . . . 5 ((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) β†’ (𝑝 pCnt (seq1( Β· , 𝐹)β€˜π‘)) = if(𝑝 ≀ 𝑁, (𝑝 pCnt 𝑁), 0))
10 iftrue 4536 . . . . . . 7 (𝑝 ≀ 𝑁 β†’ if(𝑝 ≀ 𝑁, (𝑝 pCnt 𝑁), 0) = (𝑝 pCnt 𝑁))
1110adantl 480 . . . . . 6 (((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) ∧ 𝑝 ≀ 𝑁) β†’ if(𝑝 ≀ 𝑁, (𝑝 pCnt 𝑁), 0) = (𝑝 pCnt 𝑁))
12 iffalse 4539 . . . . . . . 8 (Β¬ 𝑝 ≀ 𝑁 β†’ if(𝑝 ≀ 𝑁, (𝑝 pCnt 𝑁), 0) = 0)
1312adantl 480 . . . . . . 7 (((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) ∧ Β¬ 𝑝 ≀ 𝑁) β†’ if(𝑝 ≀ 𝑁, (𝑝 pCnt 𝑁), 0) = 0)
14 prmz 16651 . . . . . . . . . 10 (𝑝 ∈ β„™ β†’ 𝑝 ∈ β„€)
15 dvdsle 16292 . . . . . . . . . 10 ((𝑝 ∈ β„€ ∧ 𝑁 ∈ β„•) β†’ (𝑝 βˆ₯ 𝑁 β†’ 𝑝 ≀ 𝑁))
1614, 15sylan 578 . . . . . . . . 9 ((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) β†’ (𝑝 βˆ₯ 𝑁 β†’ 𝑝 ≀ 𝑁))
1716con3dimp 407 . . . . . . . 8 (((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) ∧ Β¬ 𝑝 ≀ 𝑁) β†’ Β¬ 𝑝 βˆ₯ 𝑁)
18 pceq0 16845 . . . . . . . . 9 ((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) β†’ ((𝑝 pCnt 𝑁) = 0 ↔ Β¬ 𝑝 βˆ₯ 𝑁))
1918adantr 479 . . . . . . . 8 (((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) ∧ Β¬ 𝑝 ≀ 𝑁) β†’ ((𝑝 pCnt 𝑁) = 0 ↔ Β¬ 𝑝 βˆ₯ 𝑁))
2017, 19mpbird 256 . . . . . . 7 (((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) ∧ Β¬ 𝑝 ≀ 𝑁) β†’ (𝑝 pCnt 𝑁) = 0)
2113, 20eqtr4d 2770 . . . . . 6 (((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) ∧ Β¬ 𝑝 ≀ 𝑁) β†’ if(𝑝 ≀ 𝑁, (𝑝 pCnt 𝑁), 0) = (𝑝 pCnt 𝑁))
2211, 21pm2.61dan 811 . . . . 5 ((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) β†’ if(𝑝 ≀ 𝑁, (𝑝 pCnt 𝑁), 0) = (𝑝 pCnt 𝑁))
239, 22eqtrd 2767 . . . 4 ((𝑝 ∈ β„™ ∧ 𝑁 ∈ β„•) β†’ (𝑝 pCnt (seq1( Β· , 𝐹)β€˜π‘)) = (𝑝 pCnt 𝑁))
2423ancoms 457 . . 3 ((𝑁 ∈ β„• ∧ 𝑝 ∈ β„™) β†’ (𝑝 pCnt (seq1( Β· , 𝐹)β€˜π‘)) = (𝑝 pCnt 𝑁))
2524ralrimiva 3142 . 2 (𝑁 ∈ β„• β†’ βˆ€π‘ ∈ β„™ (𝑝 pCnt (seq1( Β· , 𝐹)β€˜π‘)) = (𝑝 pCnt 𝑁))
261, 4pcmptcl 16865 . . . . . 6 (𝑁 ∈ β„• β†’ (𝐹:β„•βŸΆβ„• ∧ seq1( Β· , 𝐹):β„•βŸΆβ„•))
2726simprd 494 . . . . 5 (𝑁 ∈ β„• β†’ seq1( Β· , 𝐹):β„•βŸΆβ„•)
28 ffvelcdm 7094 . . . . 5 ((seq1( Β· , 𝐹):β„•βŸΆβ„• ∧ 𝑁 ∈ β„•) β†’ (seq1( Β· , 𝐹)β€˜π‘) ∈ β„•)
2927, 28mpancom 686 . . . 4 (𝑁 ∈ β„• β†’ (seq1( Β· , 𝐹)β€˜π‘) ∈ β„•)
3029nnnn0d 12568 . . 3 (𝑁 ∈ β„• β†’ (seq1( Β· , 𝐹)β€˜π‘) ∈ β„•0)
31 nnnn0 12515 . . 3 (𝑁 ∈ β„• β†’ 𝑁 ∈ β„•0)
32 pc11 16854 . . 3 (((seq1( Β· , 𝐹)β€˜π‘) ∈ β„•0 ∧ 𝑁 ∈ β„•0) β†’ ((seq1( Β· , 𝐹)β€˜π‘) = 𝑁 ↔ βˆ€π‘ ∈ β„™ (𝑝 pCnt (seq1( Β· , 𝐹)β€˜π‘)) = (𝑝 pCnt 𝑁)))
3330, 31, 32syl2anc 582 . 2 (𝑁 ∈ β„• β†’ ((seq1( Β· , 𝐹)β€˜π‘) = 𝑁 ↔ βˆ€π‘ ∈ β„™ (𝑝 pCnt (seq1( Β· , 𝐹)β€˜π‘)) = (𝑝 pCnt 𝑁)))
3425, 33mpbird 256 1 (𝑁 ∈ β„• β†’ (seq1( Β· , 𝐹)β€˜π‘) = 𝑁)
Colors of variables: wff setvar class
Syntax hints:  Β¬ wn 3   β†’ wi 4   ↔ wb 205   ∧ wa 394   = wceq 1533   ∈ wcel 2098  βˆ€wral 3057  ifcif 4530   class class class wbr 5150   ↦ cmpt 5233  βŸΆwf 6547  β€˜cfv 6551  (class class class)co 7424  0cc0 11144  1c1 11145   Β· cmul 11149   ≀ cle 11285  β„•cn 12248  β„•0cn0 12508  β„€cz 12594  seqcseq 14004  β†‘cexp 14064   βˆ₯ cdvds 16236  β„™cprime 16647   pCnt cpc 16810
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1789  ax-4 1803  ax-5 1905  ax-6 1963  ax-7 2003  ax-8 2100  ax-9 2108  ax-10 2129  ax-11 2146  ax-12 2166  ax-ext 2698  ax-sep 5301  ax-nul 5308  ax-pow 5367  ax-pr 5431  ax-un 7744  ax-cnex 11200  ax-resscn 11201  ax-1cn 11202  ax-icn 11203  ax-addcl 11204  ax-addrcl 11205  ax-mulcl 11206  ax-mulrcl 11207  ax-mulcom 11208  ax-addass 11209  ax-mulass 11210  ax-distr 11211  ax-i2m1 11212  ax-1ne0 11213  ax-1rid 11214  ax-rnegex 11215  ax-rrecex 11216  ax-cnre 11217  ax-pre-lttri 11218  ax-pre-lttrn 11219  ax-pre-ltadd 11220  ax-pre-mulgt0 11221  ax-pre-sup 11222
This theorem depends on definitions:  df-bi 206  df-an 395  df-or 846  df-3or 1085  df-3an 1086  df-tru 1536  df-fal 1546  df-ex 1774  df-nf 1778  df-sb 2060  df-mo 2529  df-eu 2558  df-clab 2705  df-cleq 2719  df-clel 2805  df-nfc 2880  df-ne 2937  df-nel 3043  df-ral 3058  df-rex 3067  df-rmo 3372  df-reu 3373  df-rab 3429  df-v 3473  df-sbc 3777  df-csb 3893  df-dif 3950  df-un 3952  df-in 3954  df-ss 3964  df-pss 3966  df-nul 4325  df-if 4531  df-pw 4606  df-sn 4631  df-pr 4633  df-op 4637  df-uni 4911  df-iun 5000  df-br 5151  df-opab 5213  df-mpt 5234  df-tr 5268  df-id 5578  df-eprel 5584  df-po 5592  df-so 5593  df-fr 5635  df-we 5637  df-xp 5686  df-rel 5687  df-cnv 5688  df-co 5689  df-dm 5690  df-rn 5691  df-res 5692  df-ima 5693  df-pred 6308  df-ord 6375  df-on 6376  df-lim 6377  df-suc 6378  df-iota 6503  df-fun 6553  df-fn 6554  df-f 6555  df-f1 6556  df-fo 6557  df-f1o 6558  df-fv 6559  df-riota 7380  df-ov 7427  df-oprab 7428  df-mpo 7429  df-om 7875  df-1st 7997  df-2nd 7998  df-frecs 8291  df-wrecs 8322  df-recs 8396  df-rdg 8435  df-1o 8491  df-2o 8492  df-er 8729  df-en 8969  df-dom 8970  df-sdom 8971  df-fin 8972  df-sup 9471  df-inf 9472  df-pnf 11286  df-mnf 11287  df-xr 11288  df-ltxr 11289  df-le 11290  df-sub 11482  df-neg 11483  df-div 11908  df-nn 12249  df-2 12311  df-3 12312  df-n0 12509  df-z 12595  df-uz 12859  df-q 12969  df-rp 13013  df-fz 13523  df-fl 13795  df-mod 13873  df-seq 14005  df-exp 14065  df-cj 15084  df-re 15085  df-im 15086  df-sqrt 15220  df-abs 15221  df-dvds 16237  df-gcd 16475  df-prm 16648  df-pc 16811
This theorem is referenced by:  pclogsum  27166
  Copyright terms: Public domain W3C validator