Metamath Proof Explorer


Theorem fsumrecl

Description: Closure of a finite sum of reals. (Contributed by NM, 9-Nov-2005) (Revised by Mario Carneiro, 22-Apr-2014)

Ref Expression
Hypotheses fsumcl.1 ( 𝜑𝐴 ∈ Fin )
fsumrecl.2 ( ( 𝜑𝑘𝐴 ) → 𝐵 ∈ ℝ )
Assertion fsumrecl ( 𝜑 → Σ 𝑘𝐴 𝐵 ∈ ℝ )

Proof

Step Hyp Ref Expression
1 fsumcl.1 ( 𝜑𝐴 ∈ Fin )
2 fsumrecl.2 ( ( 𝜑𝑘𝐴 ) → 𝐵 ∈ ℝ )
3 ax-resscn ℝ ⊆ ℂ
4 3 a1i ( 𝜑 → ℝ ⊆ ℂ )
5 readdcl ( ( 𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ ) → ( 𝑥 + 𝑦 ) ∈ ℝ )
6 5 adantl ( ( 𝜑 ∧ ( 𝑥 ∈ ℝ ∧ 𝑦 ∈ ℝ ) ) → ( 𝑥 + 𝑦 ) ∈ ℝ )
7 0red ( 𝜑 → 0 ∈ ℝ )
8 4 6 1 2 7 fsumcllem ( 𝜑 → Σ 𝑘𝐴 𝐵 ∈ ℝ )