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

Theorem ccatcl 13917
Description: The concatenation of two words is a word. (Contributed by FL, 2-Feb-2014.) (Proof shortened by Stefan O'Rear, 15-Aug-2015.) (Proof shortened by AV, 29-Apr-2020.)
Assertion
Ref Expression
ccatcl ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → (𝑆 ++ 𝑇) ∈ Word 𝐵)

Proof of Theorem ccatcl
Dummy variable 𝑥 is distinct from all other variables.
StepHypRef Expression
1 ccatfval 13916 . 2 ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → (𝑆 ++ 𝑇) = (𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ↦ if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆))))))
2 wrdf 13862 . . . . . . 7 (𝑆 ∈ Word 𝐵𝑆:(0..^(♯‘𝑆))⟶𝐵)
32ad2antrr 725 . . . . . 6 (((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) → 𝑆:(0..^(♯‘𝑆))⟶𝐵)
43ffvelrnda 6828 . . . . 5 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ 𝑥 ∈ (0..^(♯‘𝑆))) → (𝑆𝑥) ∈ 𝐵)
5 wrdf 13862 . . . . . . 7 (𝑇 ∈ Word 𝐵𝑇:(0..^(♯‘𝑇))⟶𝐵)
65ad3antlr 730 . . . . . 6 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) → 𝑇:(0..^(♯‘𝑇))⟶𝐵)
7 simpr 488 . . . . . . . 8 (((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) → 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))))
87anim1i 617 . . . . . . 7 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) → (𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))))
9 lencl 13876 . . . . . . . . . 10 (𝑆 ∈ Word 𝐵 → (♯‘𝑆) ∈ ℕ0)
109nn0zd 12073 . . . . . . . . 9 (𝑆 ∈ Word 𝐵 → (♯‘𝑆) ∈ ℤ)
11 lencl 13876 . . . . . . . . . 10 (𝑇 ∈ Word 𝐵 → (♯‘𝑇) ∈ ℕ0)
1211nn0zd 12073 . . . . . . . . 9 (𝑇 ∈ Word 𝐵 → (♯‘𝑇) ∈ ℤ)
1310, 12anim12i 615 . . . . . . . 8 ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → ((♯‘𝑆) ∈ ℤ ∧ (♯‘𝑇) ∈ ℤ))
1413ad2antrr 725 . . . . . . 7 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) → ((♯‘𝑆) ∈ ℤ ∧ (♯‘𝑇) ∈ ℤ))
15 fzocatel 13096 . . . . . . 7 (((𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) ∧ ((♯‘𝑆) ∈ ℤ ∧ (♯‘𝑇) ∈ ℤ)) → (𝑥 − (♯‘𝑆)) ∈ (0..^(♯‘𝑇)))
168, 14, 15syl2anc 587 . . . . . 6 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) → (𝑥 − (♯‘𝑆)) ∈ (0..^(♯‘𝑇)))
176, 16ffvelrnd 6829 . . . . 5 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) → (𝑇‘(𝑥 − (♯‘𝑆))) ∈ 𝐵)
184, 17ifclda 4459 . . . 4 (((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) → if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆)))) ∈ 𝐵)
1918fmpttd 6856 . . 3 ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → (𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ↦ if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆))))):(0..^((♯‘𝑆) + (♯‘𝑇)))⟶𝐵)
20 iswrdi 13861 . . 3 ((𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ↦ if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆))))):(0..^((♯‘𝑆) + (♯‘𝑇)))⟶𝐵 → (𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ↦ if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆))))) ∈ Word 𝐵)
2119, 20syl 17 . 2 ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → (𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ↦ if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆))))) ∈ Word 𝐵)
221, 21eqeltrd 2890 1 ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → (𝑆 ++ 𝑇) ∈ Word 𝐵)
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 399  wcel 2111  ifcif 4425  cmpt 5110  wf 6320  cfv 6324  (class class class)co 7135  0cc0 10526   + caddc 10529  cmin 10859  cz 11969  ..^cfzo 13028  chash 13686  Word cword 13857   ++ cconcat 13913
This theorem was proved from axioms:  ax-mp 5  ax-1 6  ax-2 7  ax-3 8  ax-gen 1797  ax-4 1811  ax-5 1911  ax-6 1970  ax-7 2015  ax-8 2113  ax-9 2121  ax-10 2142  ax-11 2158  ax-12 2175  ax-ext 2770  ax-rep 5154  ax-sep 5167  ax-nul 5174  ax-pow 5231  ax-pr 5295  ax-un 7441  ax-cnex 10582  ax-resscn 10583  ax-1cn 10584  ax-icn 10585  ax-addcl 10586  ax-addrcl 10587  ax-mulcl 10588  ax-mulrcl 10589  ax-mulcom 10590  ax-addass 10591  ax-mulass 10592  ax-distr 10593  ax-i2m1 10594  ax-1ne0 10595  ax-1rid 10596  ax-rnegex 10597  ax-rrecex 10598  ax-cnre 10599  ax-pre-lttri 10600  ax-pre-lttrn 10601  ax-pre-ltadd 10602  ax-pre-mulgt0 10603
This theorem depends on definitions:  df-bi 210  df-an 400  df-or 845  df-3or 1085  df-3an 1086  df-tru 1541  df-ex 1782  df-nf 1786  df-sb 2070  df-mo 2598  df-eu 2629  df-clab 2777  df-cleq 2791  df-clel 2870  df-nfc 2938  df-ne 2988  df-nel 3092  df-ral 3111  df-rex 3112  df-reu 3113  df-rab 3115  df-v 3443  df-sbc 3721  df-csb 3829  df-dif 3884  df-un 3886  df-in 3888  df-ss 3898  df-pss 3900  df-nul 4244  df-if 4426  df-pw 4499  df-sn 4526  df-pr 4528  df-tp 4530  df-op 4532  df-uni 4801  df-int 4839  df-iun 4883  df-br 5031  df-opab 5093  df-mpt 5111  df-tr 5137  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 6116  df-ord 6162  df-on 6163  df-lim 6164  df-suc 6165  df-iota 6283  df-fun 6326  df-fn 6327  df-f 6328  df-f1 6329  df-fo 6330  df-f1o 6331  df-fv 6332  df-riota 7093  df-ov 7138  df-oprab 7139  df-mpo 7140  df-om 7561  df-1st 7671  df-2nd 7672  df-wrecs 7930  df-recs 7991  df-rdg 8029  df-1o 8085  df-oadd 8089  df-er 8272  df-en 8493  df-dom 8494  df-sdom 8495  df-fin 8496  df-card 9352  df-pnf 10666  df-mnf 10667  df-xr 10668  df-ltxr 10669  df-le 10670  df-sub 10861  df-neg 10862  df-nn 11626  df-n0 11886  df-z 11970  df-uz 12232  df-fz 12886  df-fzo 13029  df-hash 13687  df-word 13858  df-concat 13914
This theorem is referenced by:  ccatsymb  13927  ccatass  13933  ccatalpha  13938  ccatws1cl  13961  ccatws1clv  13962  ccatswrd  14021  swrdccat2  14022  ccatpfx  14054  pfxccat1  14055  swrdccatfn  14077  swrdccatin1  14078  swrdccatin2  14082  pfxccatin12lem2c  14083  pfxccatpfx1  14089  pfxccatpfx2  14090  splcl  14105  spllen  14107  splfv1  14108  splfv2a  14109  splval2  14110  revccat  14119  cshwcl  14151  cats1cld  14208  cats1cli  14210  cats2cat  14215  gsumsgrpccat  17996  gsumccatOLD  17997  gsumspl  18001  gsumwspan  18003  frmdplusg  18011  frmdmnd  18016  frmdsssubm  18018  frmdup1  18021  psgnuni  18619  efginvrel2  18845  efgsp1  18855  efgredleme  18861  efgredlemc  18863  efgcpbllemb  18873  efgcpbl2  18875  frgpuplem  18890  frgpup1  18893  psgnghm  20269  wwlksnext  27679  clwwlkccat  27775  clwlkclwwlk2  27788  clwwlkel  27831  wwlksext2clwwlk  27842  numclwwlk1lem2fo  28143  ccatf1  30651  splfv3  30658  cycpmco2f1  30816  cycpmco2rn  30817  cycpmco2lem2  30819  cycpmco2lem3  30820  cycpmco2lem4  30821  cycpmco2lem5  30822  cycpmco2lem6  30823  cycpmco2  30825  cyc3genpm  30844  sseqf  31760  ofcccat  31923  signstfvc  31954  signsvfn  31962  signsvtn  31964  signshf  31968  mrsubccat  32878  mrsubco  32881  frlmfzoccat  39439
  Copyright terms: Public domain W3C validator