Skip to content

Commit 2b47ec0

Browse files
1 parent 06326c9 commit 2b47ec0

File tree

2 files changed

+41
-2
lines changed

2 files changed

+41
-2
lines changed

Mathlib/MeasureTheory/Measure/MeasureSpace.lean

Lines changed: 16 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johannes Hölzl, Mario Carneiro
55
66
! This file was ported from Lean 3 source module measure_theory.measure.measure_space
7-
! leanprover-community/mathlib commit 88fcb83fe7996142dfcfe7368d31304a9adc874a
7+
! leanprover-community/mathlib commit 343e80208d29d2d15f8050b929aa50fe4ce71b55
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -1058,6 +1058,21 @@ instance instCompleteLattice [MeasurableSpace α] : CompleteLattice (Measure α)
10581058

10591059
end sInf
10601060

1061+
@[simp]
1062+
theorem _root_.MeasureTheory.OuterMeasure.toMeasure_top [MeasurableSpace α] :
1063+
(⊤ : OuterMeasure α).toMeasure (by rw [OuterMeasure.top_caratheodory]; exact le_top) =
1064+
(⊤ : Measure α) :=
1065+
top_unique fun s hs => by
1066+
cases' s.eq_empty_or_nonempty with h h <;>
1067+
simp [h, toMeasure_apply ⊤ _ hs, OuterMeasure.top_apply]
1068+
#align measure_theory.outer_measure.to_measure_top MeasureTheory.OuterMeasure.toMeasure_top
1069+
1070+
@[simp]
1071+
theorem toOuterMeasure_top [MeasurableSpace α] :
1072+
(⊤ : Measure α).toOuterMeasure = (⊤ : OuterMeasure α) := by
1073+
rw [← OuterMeasure.toMeasure_top, toMeasure_toOuterMeasure, OuterMeasure.trim_top]
1074+
#align measure_theory.measure.to_outer_measure_top MeasureTheory.Measure.toOuterMeasure_top
1075+
10611076
@[simp]
10621077
theorem top_add : ⊤ + μ = ⊤ :=
10631078
top_unique <| Measure.le_add_right le_rfl

Mathlib/MeasureTheory/Measure/OuterMeasure.lean

Lines changed: 25 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johannes Hölzl, Mario Carneiro
55
66
! 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
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -882,6 +882,20 @@ theorem le_bounded_by' {μ : OuterMeasure α} :
882882
cases' s.eq_empty_or_nonempty with h h <;> simp [h]
883883
#align measure_theory.outer_measure.le_bounded_by' MeasureTheory.OuterMeasure.le_bounded_by'
884884

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+
885899
theorem smul_boundedBy {c : ℝ≥0∞} (hc : c ≠ ∞) : c • boundedBy m = boundedBy (c • m) := by
886900
simp only [boundedBy , smul_ofFunction hc]
887901
congr 1 with s : 1
@@ -1334,6 +1348,11 @@ theorem extend_congr {β : Type _} {Pb : β → Prop} {mb : ∀ s : β, Pb s →
13341348
iInf_congr_Prop hP fun _h => hm _ _
13351349
#align measure_theory.extend_congr MeasureTheory.extend_congr
13361350

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+
13371356
end Extend
13381357

13391358
section ExtendSet
@@ -1648,6 +1667,11 @@ theorem trim_trim (m : OuterMeasure α) : m.trim.trim = m.trim :=
16481667
trim_eq_trim_iff.2 fun _s => m.trim_eq
16491668
#align measure_theory.outer_measure.trim_trim MeasureTheory.OuterMeasure.trim_trim
16501669

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+
16511675
@[simp]
16521676
theorem trim_zero : (0 : OuterMeasure α).trim = 0 :=
16531677
ext fun s =>

0 commit comments

Comments
 (0)