@@ -288,6 +288,20 @@ protected lemma tendsto.mul_const {f : filter α} {m : α → ℝ≥0∞} {a b :
288
288
(hm : tendsto m f (𝓝 a)) (ha : a ≠ 0 ∨ b ≠ ⊤) : tendsto (λx, m x * b) f (𝓝 (a * b)) :=
289
289
by simpa only [mul_comm] using ennreal.tendsto.const_mul hm ha
290
290
291
+ lemma tendsto_finset_prod_of_ne_top {ι : Type *} {f : ι → α → ℝ≥0 ∞} {x : filter α} {a : ι → ℝ≥0 ∞}
292
+ (s : finset ι) (h : ∀ i ∈ s, tendsto (f i) x (𝓝 (a i))) (h' : ∀ i ∈ s, a i ≠ ∞):
293
+ tendsto (λ b, ∏ c in s, f c b) x (𝓝 (∏ c in s, a c)) :=
294
+ begin
295
+ induction s using finset.induction with a s has IH, { simp [tendsto_const_nhds] },
296
+ simp only [finset.prod_insert has],
297
+ apply tendsto.mul (h _ (finset.mem_insert_self _ _)),
298
+ { right,
299
+ exact (prod_lt_top (λ i hi, lt_top_iff_ne_top.2 (h' _ (finset.mem_insert_of_mem hi)))).ne },
300
+ { exact IH (λ i hi, h _ (finset.mem_insert_of_mem hi))
301
+ (λ i hi, h' _ (finset.mem_insert_of_mem hi)) },
302
+ { exact or.inr (h' _ (finset.mem_insert_self _ _)) }
303
+ end
304
+
291
305
protected lemma continuous_at_const_mul {a b : ℝ≥0 ∞} (h : a ≠ ⊤ ∨ b ≠ 0 ) :
292
306
continuous_at ((*) a) b :=
293
307
tendsto.const_mul tendsto_id h.symm
@@ -1101,9 +1115,11 @@ end
1101
1115
metric.diam (closure s) = diam s :=
1102
1116
by simp only [metric.diam, emetric.diam_closure]
1103
1117
1118
+ namespace real
1119
+
1104
1120
/-- For a bounded set `s : set ℝ`, its `emetric.diam` is equal to `Sup s - Inf s` reinterpreted as
1105
1121
`ℝ≥0∞`. -/
1106
- lemma real. ediam_eq {s : set ℝ} (h : bounded s) :
1122
+ lemma ediam_eq {s : set ℝ} (h : bounded s) :
1107
1123
emetric.diam s = ennreal.of_real (Sup s - Inf s) :=
1108
1124
begin
1109
1125
rcases eq_empty_or_nonempty s with rfl|hne, { simp },
@@ -1119,13 +1135,41 @@ begin
1119
1135
end
1120
1136
1121
1137
/-- For a bounded set `s : set ℝ`, its `metric.diam` is equal to `Sup s - Inf s`. -/
1122
- lemma real. diam_eq {s : set ℝ} (h : bounded s) : metric.diam s = Sup s - Inf s :=
1138
+ lemma diam_eq {s : set ℝ} (h : bounded s) : metric.diam s = Sup s - Inf s :=
1123
1139
begin
1124
1140
rw [metric.diam, real.ediam_eq h, ennreal.to_real_of_real],
1125
1141
rw real.bounded_iff_bdd_below_bdd_above at h,
1126
1142
exact sub_nonneg.2 (real.Inf_le_Sup s h.1 h.2 )
1127
1143
end
1128
1144
1145
+ @[simp] lemma ediam_Ioo (a b : ℝ) :
1146
+ emetric.diam (Ioo a b) = ennreal.of_real (b - a) :=
1147
+ begin
1148
+ rcases le_or_lt b a with h|h,
1149
+ { simp [h] },
1150
+ { rw [real.ediam_eq (bounded_Ioo _ _), cSup_Ioo h, cInf_Ioo h] },
1151
+ end
1152
+
1153
+ @[simp] lemma ediam_Icc (a b : ℝ) :
1154
+ emetric.diam (Icc a b) = ennreal.of_real (b - a) :=
1155
+ begin
1156
+ rcases le_or_lt a b with h|h,
1157
+ { rw [real.ediam_eq (bounded_Icc _ _), cSup_Icc h, cInf_Icc h] },
1158
+ { simp [h, h.le] }
1159
+ end
1160
+
1161
+ @[simp] lemma ediam_Ico (a b : ℝ) :
1162
+ emetric.diam (Ico a b) = ennreal.of_real (b - a) :=
1163
+ le_antisymm (ediam_Icc a b ▸ diam_mono Ico_subset_Icc_self)
1164
+ (ediam_Ioo a b ▸ diam_mono Ioo_subset_Ico_self)
1165
+
1166
+ @[simp] lemma ediam_Ioc (a b : ℝ) :
1167
+ emetric.diam (Ioc a b) = ennreal.of_real (b - a) :=
1168
+ le_antisymm (ediam_Icc a b ▸ diam_mono Ioc_subset_Icc_self)
1169
+ (ediam_Ioo a b ▸ diam_mono Ioo_subset_Ioc_self)
1170
+
1171
+ end real
1172
+
1129
1173
/-- If `edist (f n) (f (n+1))` is bounded above by a function `d : ℕ → ℝ≥0∞`,
1130
1174
then the distance from `f n` to the limit is bounded by `∑'_{k=n}^∞ d k`. -/
1131
1175
lemma edist_le_tsum_of_edist_le_of_tendsto {f : ℕ → α} (d : ℕ → ℝ≥0 ∞)
0 commit comments