Database
REAL AND COMPLEX NUMBERS
Words over a set
Repeated symbol words
repsconst
Metamath Proof Explorer
Description: Construct a function mapping a half-open range of nonnegative integers
to a constant, see also fconstmpt . (Contributed by AV , 4-Nov-2018)
Ref
Expression
Assertion
repsconst
⊢ ( ( 𝑆 ∈ 𝑉 ∧ 𝑁 ∈ ℕ0 ) → ( 𝑆 repeatS 𝑁 ) = ( ( 0 ..^ 𝑁 ) × { 𝑆 } ) )
Proof
Step
Hyp
Ref
Expression
1
reps
⊢ ( ( 𝑆 ∈ 𝑉 ∧ 𝑁 ∈ ℕ0 ) → ( 𝑆 repeatS 𝑁 ) = ( 𝑥 ∈ ( 0 ..^ 𝑁 ) ↦ 𝑆 ) )
2
fconstmpt
⊢ ( ( 0 ..^ 𝑁 ) × { 𝑆 } ) = ( 𝑥 ∈ ( 0 ..^ 𝑁 ) ↦ 𝑆 )
3
1 2
eqtr4di
⊢ ( ( 𝑆 ∈ 𝑉 ∧ 𝑁 ∈ ℕ0 ) → ( 𝑆 repeatS 𝑁 ) = ( ( 0 ..^ 𝑁 ) × { 𝑆 } ) )