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

Theorem ccatcl 14015
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 14014 . 2 ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → (𝑆 ++ 𝑇) = (𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ↦ if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆))))))
2 wrdf 13960 . . . . . . 7 (𝑆 ∈ Word 𝐵𝑆:(0..^(♯‘𝑆))⟶𝐵)
32ad2antrr 726 . . . . . 6 (((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) → 𝑆:(0..^(♯‘𝑆))⟶𝐵)
43ffvelrnda 6861 . . . . 5 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ 𝑥 ∈ (0..^(♯‘𝑆))) → (𝑆𝑥) ∈ 𝐵)
5 wrdf 13960 . . . . . . 7 (𝑇 ∈ Word 𝐵𝑇:(0..^(♯‘𝑇))⟶𝐵)
65ad3antlr 731 . . . . . 6 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) → 𝑇:(0..^(♯‘𝑇))⟶𝐵)
7 simpr 488 . . . . . . . 8 (((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) → 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))))
87anim1i 618 . . . . . . 7 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) → (𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))))
9 lencl 13974 . . . . . . . . . 10 (𝑆 ∈ Word 𝐵 → (♯‘𝑆) ∈ ℕ0)
109nn0zd 12166 . . . . . . . . 9 (𝑆 ∈ Word 𝐵 → (♯‘𝑆) ∈ ℤ)
11 lencl 13974 . . . . . . . . . 10 (𝑇 ∈ Word 𝐵 → (♯‘𝑇) ∈ ℕ0)
1211nn0zd 12166 . . . . . . . . 9 (𝑇 ∈ Word 𝐵 → (♯‘𝑇) ∈ ℤ)
1310, 12anim12i 616 . . . . . . . 8 ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → ((♯‘𝑆) ∈ ℤ ∧ (♯‘𝑇) ∈ ℤ))
1413ad2antrr 726 . . . . . . 7 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) → ((♯‘𝑆) ∈ ℤ ∧ (♯‘𝑇) ∈ ℤ))
15 fzocatel 13192 . . . . . . 7 (((𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) ∧ ((♯‘𝑆) ∈ ℤ ∧ (♯‘𝑇) ∈ ℤ)) → (𝑥 − (♯‘𝑆)) ∈ (0..^(♯‘𝑇)))
168, 14, 15syl2anc 587 . . . . . 6 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) → (𝑥 − (♯‘𝑆)) ∈ (0..^(♯‘𝑇)))
176, 16ffvelrnd 6862 . . . . 5 ((((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) ∧ ¬ 𝑥 ∈ (0..^(♯‘𝑆))) → (𝑇‘(𝑥 − (♯‘𝑆))) ∈ 𝐵)
184, 17ifclda 4449 . . . 4 (((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) ∧ 𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇)))) → if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆)))) ∈ 𝐵)
1918fmpttd 6889 . . 3 ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → (𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ↦ if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆))))):(0..^((♯‘𝑆) + (♯‘𝑇)))⟶𝐵)
20 iswrdi 13959 . . 3 ((𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ↦ if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆))))):(0..^((♯‘𝑆) + (♯‘𝑇)))⟶𝐵 → (𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ↦ if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆))))) ∈ Word 𝐵)
2119, 20syl 17 . 2 ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → (𝑥 ∈ (0..^((♯‘𝑆) + (♯‘𝑇))) ↦ if(𝑥 ∈ (0..^(♯‘𝑆)), (𝑆𝑥), (𝑇‘(𝑥 − (♯‘𝑆))))) ∈ Word 𝐵)
221, 21eqeltrd 2833 1 ((𝑆 ∈ Word 𝐵𝑇 ∈ Word 𝐵) → (𝑆 ++ 𝑇) ∈ Word 𝐵)
Colors of variables: wff setvar class
Syntax hints:  ¬ wn 3  wi 4  wa 399  wcel 2114  ifcif 4414  cmpt 5110  wf 6335  cfv 6339  (class class class)co 7170  0cc0 10615   + caddc 10618  cmin 10948  cz 12062  ..^cfzo 13124  chash 13782  Word cword 13955   ++ cconcat 14011
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 2020  ax-8 2116  ax-9 2124  ax-10 2145  ax-11 2162  ax-12 2179  ax-ext 2710  ax-rep 5154  ax-sep 5167  ax-nul 5174  ax-pow 5232  ax-pr 5296  ax-un 7479  ax-cnex 10671  ax-resscn 10672  ax-1cn 10673  ax-icn 10674  ax-addcl 10675  ax-addrcl 10676  ax-mulcl 10677  ax-mulrcl 10678  ax-mulcom 10679  ax-addass 10680  ax-mulass 10681  ax-distr 10682  ax-i2m1 10683  ax-1ne0 10684  ax-1rid 10685  ax-rnegex 10686  ax-rrecex 10687  ax-cnre 10688  ax-pre-lttri 10689  ax-pre-lttrn 10690  ax-pre-ltadd 10691  ax-pre-mulgt0 10692
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 2075  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-rab 3062  df-v 3400  df-sbc 3681  df-csb 3791  df-dif 3846  df-un 3848  df-in 3850  df-ss 3860  df-pss 3862  df-nul 4212  df-if 4415  df-pw 4490  df-sn 4517  df-pr 4519  df-tp 4521  df-op 4523  df-uni 4797  df-int 4837  df-iun 4883  df-br 5031  df-opab 5093  df-mpt 5111  df-tr 5137  df-id 5429  df-eprel 5434  df-po 5442  df-so 5443  df-fr 5483  df-we 5485  df-xp 5531  df-rel 5532  df-cnv 5533  df-co 5534  df-dm 5535  df-rn 5536  df-res 5537  df-ima 5538  df-pred 6129  df-ord 6175  df-on 6176  df-lim 6177  df-suc 6178  df-iota 6297  df-fun 6341  df-fn 6342  df-f 6343  df-f1 6344  df-fo 6345  df-f1o 6346  df-fv 6347  df-riota 7127  df-ov 7173  df-oprab 7174  df-mpo 7175  df-om 7600  df-1st 7714  df-2nd 7715  df-wrecs 7976  df-recs 8037  df-rdg 8075  df-1o 8131  df-er 8320  df-en 8556  df-dom 8557  df-sdom 8558  df-fin 8559  df-card 9441  df-pnf 10755  df-mnf 10756  df-xr 10757  df-ltxr 10758  df-le 10759  df-sub 10950  df-neg 10951  df-nn 11717  df-n0 11977  df-z 12063  df-uz 12325  df-fz 12982  df-fzo 13125  df-hash 13783  df-word 13956  df-concat 14012
This theorem is referenced by:  ccatsymb  14025  ccatass  14031  ccatalpha  14036  ccatws1cl  14059  ccatws1clv  14060  ccatswrd  14119  swrdccat2  14120  ccatpfx  14152  pfxccat1  14153  swrdccatfn  14175  swrdccatin1  14176  swrdccatin2  14180  pfxccatin12lem2c  14181  pfxccatpfx1  14187  pfxccatpfx2  14188  splcl  14203  spllen  14205  splfv1  14206  splfv2a  14207  splval2  14208  revccat  14217  cshwcl  14249  cats1cld  14306  cats1cli  14308  cats2cat  14313  gsumsgrpccat  18120  gsumccatOLD  18121  gsumspl  18125  gsumwspan  18127  frmdplusg  18135  frmdmnd  18140  frmdsssubm  18142  frmdup1  18145  psgnuni  18745  efginvrel2  18971  efgsp1  18981  efgredleme  18987  efgredlemc  18989  efgcpbllemb  18999  efgcpbl2  19001  frgpuplem  19016  frgpup1  19019  psgnghm  20396  wwlksnext  27831  clwwlkccat  27927  clwlkclwwlk2  27940  clwwlkel  27983  wwlksext2clwwlk  27994  numclwwlk1lem2fo  28295  ccatf1  30798  splfv3  30805  cycpmco2f1  30968  cycpmco2rn  30969  cycpmco2lem2  30971  cycpmco2lem3  30972  cycpmco2lem4  30973  cycpmco2lem5  30974  cycpmco2lem6  30975  cycpmco2  30977  cyc3genpm  30996  sseqf  31929  ofcccat  32092  signstfvc  32123  signsvfn  32131  signsvtn  32133  signshf  32137  mrsubccat  33051  mrsubco  33054  frlmfzoccat  39818
  Copyright terms: Public domain W3C validator