@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Johannes Hölzl, Mario Carneiro
5
5
6
6
! This file was ported from Lean 3 source module measure_theory.measure.outer_measure
7
- ! leanprover-community/mathlib commit ec4b2eeb50364487f80421c0b4c41328a611f30d
7
+ ! leanprover-community/mathlib commit 343e80208d29d2d15f8050b929aa50fe4ce71b55
8
8
! Please do not edit these lines, except to modify the commit id
9
9
! if you have ported upstream changes.
10
10
-/
@@ -882,6 +882,20 @@ theorem le_bounded_by' {μ : OuterMeasure α} :
882
882
cases' s.eq_empty_or_nonempty with h h <;> simp [h]
883
883
#align measure_theory.outer_measure.le_bounded_by' MeasureTheory.OuterMeasure.le_bounded_by'
884
884
885
+ @[simp]
886
+ theorem boundedBy_top : boundedBy (⊤ : Set α → ℝ≥0 ∞) = ⊤ := by
887
+ rw [eq_top_iff, le_bounded_by']
888
+ intro s hs
889
+ rw [top_apply hs]
890
+ exact le_rfl
891
+ #align measure_theory.outer_measure.bounded_by_top MeasureTheory.OuterMeasure.boundedBy_top
892
+
893
+ @[simp]
894
+ theorem boundedBy_zero : boundedBy (0 : Set α → ℝ≥0 ∞) = 0 := by
895
+ rw [← coe_bot, eq_bot_iff]
896
+ apply boundedBy_le
897
+ #align measure_theory.outer_measure.bounded_by_zero MeasureTheory.OuterMeasure.boundedBy_zero
898
+
885
899
theorem smul_boundedBy {c : ℝ≥0 ∞} (hc : c ≠ ∞) : c • boundedBy m = boundedBy (c • m) := by
886
900
simp only [boundedBy , smul_ofFunction hc]
887
901
congr 1 with s : 1
@@ -1334,6 +1348,11 @@ theorem extend_congr {β : Type _} {Pb : β → Prop} {mb : ∀ s : β, Pb s →
1334
1348
iInf_congr_Prop hP fun _h => hm _ _
1335
1349
#align measure_theory.extend_congr MeasureTheory.extend_congr
1336
1350
1351
+ @[simp]
1352
+ theorem extend_top {α : Type _} {P : α → Prop } : extend (fun _ _ => ∞ : ∀ s : α, P s → ℝ≥0 ∞) = ⊤ :=
1353
+ funext fun _ => iInf_eq_top.mpr fun _ => rfl
1354
+ #align measure_theory.extend_top MeasureTheory.extend_top
1355
+
1337
1356
end Extend
1338
1357
1339
1358
section ExtendSet
@@ -1648,6 +1667,11 @@ theorem trim_trim (m : OuterMeasure α) : m.trim.trim = m.trim :=
1648
1667
trim_eq_trim_iff.2 fun _s => m.trim_eq
1649
1668
#align measure_theory.outer_measure.trim_trim MeasureTheory.OuterMeasure.trim_trim
1650
1669
1670
+ @[simp]
1671
+ theorem trim_top : (⊤ : OuterMeasure α).trim = ⊤ :=
1672
+ eq_top_iff.2 <| le_trim _
1673
+ #align measure_theory.outer_measure.trim_top MeasureTheory.OuterMeasure.trim_top
1674
+
1651
1675
@[simp]
1652
1676
theorem trim_zero : (0 : OuterMeasure α).trim = 0 :=
1653
1677
ext fun s =>
0 commit comments