Metamath Proof Explorer


Theorem lswccats1

Description: The last symbol of a word concatenated with a singleton word is the symbol of the singleton word. (Contributed by AV, 6-Aug-2018) (Proof shortened by AV, 22-Oct-2018)

Ref Expression
Assertion lswccats1 ( ( 𝑊 ∈ Word 𝑉𝑆𝑉 ) → ( lastS ‘ ( 𝑊 ++ ⟨“ 𝑆 ”⟩ ) ) = 𝑆 )

Proof

Step Hyp Ref Expression
1 simpl ( ( 𝑊 ∈ Word 𝑉𝑆𝑉 ) → 𝑊 ∈ Word 𝑉 )
2 s1cl ( 𝑆𝑉 → ⟨“ 𝑆 ”⟩ ∈ Word 𝑉 )
3 2 adantl ( ( 𝑊 ∈ Word 𝑉𝑆𝑉 ) → ⟨“ 𝑆 ”⟩ ∈ Word 𝑉 )
4 s1nz ⟨“ 𝑆 ”⟩ ≠ ∅
5 4 a1i ( ( 𝑊 ∈ Word 𝑉𝑆𝑉 ) → ⟨“ 𝑆 ”⟩ ≠ ∅ )
6 lswccatn0lsw ( ( 𝑊 ∈ Word 𝑉 ∧ ⟨“ 𝑆 ”⟩ ∈ Word 𝑉 ∧ ⟨“ 𝑆 ”⟩ ≠ ∅ ) → ( lastS ‘ ( 𝑊 ++ ⟨“ 𝑆 ”⟩ ) ) = ( lastS ‘ ⟨“ 𝑆 ”⟩ ) )
7 1 3 5 6 syl3anc ( ( 𝑊 ∈ Word 𝑉𝑆𝑉 ) → ( lastS ‘ ( 𝑊 ++ ⟨“ 𝑆 ”⟩ ) ) = ( lastS ‘ ⟨“ 𝑆 ”⟩ ) )
8 lsws1 ( 𝑆𝑉 → ( lastS ‘ ⟨“ 𝑆 ”⟩ ) = 𝑆 )
9 8 adantl ( ( 𝑊 ∈ Word 𝑉𝑆𝑉 ) → ( lastS ‘ ⟨“ 𝑆 ”⟩ ) = 𝑆 )
10 7 9 eqtrd ( ( 𝑊 ∈ Word 𝑉𝑆𝑉 ) → ( lastS ‘ ( 𝑊 ++ ⟨“ 𝑆 ”⟩ ) ) = 𝑆 )