|
| 1 | +/- |
| 2 | +Copyright (c) 2021 Kexing Ying. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Kexing Ying |
| 5 | +-/ |
| 6 | +import measure_theory.decomposition.signed_hahn |
| 7 | + |
| 8 | +/-! |
| 9 | +# Jordan decomposition |
| 10 | +
|
| 11 | +This file proves the existence and uniqueness of the Jordan decomposition for signed measures. |
| 12 | +The Jordan decomposition theorem states that, given a signed measure `s`, there exists a |
| 13 | +unique pair of mutually singular measures `μ` and `ν`, such that `s = μ - ν`. |
| 14 | +
|
| 15 | +The Jordan decomposition theorem for measures is a corollary of the Hahn decomposition theorem and |
| 16 | +is useful for the Lebesgue decomposition theorem. |
| 17 | +
|
| 18 | +## Main definitions |
| 19 | +
|
| 20 | +* `measure_theory.jordan_decomposition`: a Jordan decomposition of a measurable space is a |
| 21 | + pair of mutually singular finite measures. We say `j` is a Jordan decomposition of a signed |
| 22 | + meausre `s` if `s = j.pos_part - j.neg_part`. |
| 23 | +* `measure_theory.signed_measure.to_jordan_decomposition`: the Jordan decomposition of a |
| 24 | + signed measure. |
| 25 | +* `measure_theory.signed_measure.to_jordan_decomposition_equiv`: is the `equiv` between |
| 26 | + `measure_theory.signed_measure` and `measure_theory.jordan_decomposition` formed by |
| 27 | + `measure_theory.signed_measure.to_jordan_decomposition`. |
| 28 | +
|
| 29 | +## Main results |
| 30 | +
|
| 31 | +* `measure_theory.signed_measure.to_signed_measure_to_jordan_decomposition` : the Jordan |
| 32 | + decomposition theorem. |
| 33 | +* `measure_thoery.signed_measure.to_signed_measure_injective` : the Jordan decomposition of a |
| 34 | + signed measure is unique. |
| 35 | +
|
| 36 | +## Tags |
| 37 | +
|
| 38 | +Jordan decomposition theorem |
| 39 | +-/ |
| 40 | + |
| 41 | +noncomputable theory |
| 42 | +open_locale classical measure_theory ennreal |
| 43 | + |
| 44 | +variables {α β : Type*} [measurable_space α] |
| 45 | + |
| 46 | +namespace measure_theory |
| 47 | + |
| 48 | +/-- A Jordan decomposition of a measurable space is a pair of mutually singular, |
| 49 | +finite measures. -/ |
| 50 | +@[ext] structure jordan_decomposition (α : Type*) [measurable_space α] := |
| 51 | +(pos_part neg_part : measure α) |
| 52 | +[pos_part_finite : finite_measure pos_part] |
| 53 | +[neg_part_finite : finite_measure neg_part] |
| 54 | +(mutually_singular : pos_part ⊥ₘ neg_part) |
| 55 | + |
| 56 | +attribute [instance] jordan_decomposition.pos_part_finite |
| 57 | +attribute [instance] jordan_decomposition.neg_part_finite |
| 58 | + |
| 59 | +namespace jordan_decomposition |
| 60 | + |
| 61 | +open measure vector_measure |
| 62 | + |
| 63 | +variable (j : jordan_decomposition α) |
| 64 | + |
| 65 | +instance : has_zero (jordan_decomposition α) := |
| 66 | +{ zero := ⟨0, 0, mutually_singular.zero⟩ } |
| 67 | + |
| 68 | +instance : inhabited (jordan_decomposition α) := |
| 69 | +{ default := 0 } |
| 70 | + |
| 71 | +instance : has_neg (jordan_decomposition α) := |
| 72 | +{ neg := λ j, ⟨j.neg_part, j.pos_part, j.mutually_singular.symm⟩ } |
| 73 | + |
| 74 | +@[simp] lemma zero_pos_part : (0 : jordan_decomposition α).pos_part = 0 := rfl |
| 75 | +@[simp] lemma zero_neg_part : (0 : jordan_decomposition α).neg_part = 0 := rfl |
| 76 | + |
| 77 | +@[simp] lemma neg_pos_part : (-j).pos_part = j.neg_part := rfl |
| 78 | +@[simp] lemma neg_neg_part : (-j).neg_part = j.pos_part := rfl |
| 79 | + |
| 80 | +/-- The signed measure associated with a Jordan decomposition. -/ |
| 81 | +def to_signed_measure : signed_measure α := |
| 82 | +j.pos_part.to_signed_measure - j.neg_part.to_signed_measure |
| 83 | + |
| 84 | +lemma to_signed_measure_zero : (0 : jordan_decomposition α).to_signed_measure = 0 := |
| 85 | +begin |
| 86 | + ext1 i hi, |
| 87 | + erw [to_signed_measure, to_signed_measure_sub_apply hi, sub_self, zero_apply], |
| 88 | +end |
| 89 | + |
| 90 | +lemma to_signed_measure_neg : (-j).to_signed_measure = -j.to_signed_measure := |
| 91 | +begin |
| 92 | + ext1 i hi, |
| 93 | + rw [neg_apply, to_signed_measure, to_signed_measure, |
| 94 | + to_signed_measure_sub_apply hi, to_signed_measure_sub_apply hi, neg_sub], |
| 95 | + refl, |
| 96 | +end |
| 97 | + |
| 98 | +/-- A Jordan decomposition provides a Hahn decomposition. -/ |
| 99 | +lemma exists_compl_positive_negative : |
| 100 | + ∃ S : set α, measurable_set S ∧ |
| 101 | + j.to_signed_measure ≤[S] 0 ∧ 0 ≤[Sᶜ] j.to_signed_measure ∧ |
| 102 | + j.pos_part S = 0 ∧ j.neg_part Sᶜ = 0 := |
| 103 | +begin |
| 104 | + obtain ⟨S, hS₁, hS₂, hS₃⟩ := j.mutually_singular, |
| 105 | + refine ⟨S, hS₁, _, _, hS₂, hS₃⟩, |
| 106 | + { refine restrict_le_restrict_of_subset_le _ _ (λ A hA hA₁, _), |
| 107 | + rw [to_signed_measure, to_signed_measure_sub_apply hA, |
| 108 | + show j.pos_part A = 0, by exact nonpos_iff_eq_zero.1 (hS₂ ▸ measure_mono hA₁), |
| 109 | + ennreal.zero_to_real, zero_sub, neg_le, zero_apply, neg_zero], |
| 110 | + exact ennreal.to_real_nonneg }, |
| 111 | + { refine restrict_le_restrict_of_subset_le _ _ (λ A hA hA₁, _), |
| 112 | + rw [to_signed_measure, to_signed_measure_sub_apply hA, |
| 113 | + show j.neg_part A = 0, by exact nonpos_iff_eq_zero.1 (hS₃ ▸ measure_mono hA₁), |
| 114 | + ennreal.zero_to_real, sub_zero], |
| 115 | + exact ennreal.to_real_nonneg }, |
| 116 | +end |
| 117 | + |
| 118 | +end jordan_decomposition |
| 119 | + |
| 120 | +namespace signed_measure |
| 121 | + |
| 122 | +open measure vector_measure jordan_decomposition classical |
| 123 | + |
| 124 | +variables {s : signed_measure α} {μ ν : measure α} [finite_measure μ] [finite_measure ν] |
| 125 | + |
| 126 | +/-- Given a signed measure `s`, `s.to_jordan_decomposition` is the Jordan decomposition `j`, |
| 127 | +such that `s = j.to_signed_measure`. This property is known as the Jordan decomposition |
| 128 | +theorem, and is shown by |
| 129 | +`measure_theory.signed_measure.to_signed_measure_to_jordan_decomposition`. -/ |
| 130 | +def to_jordan_decomposition (s : signed_measure α) : jordan_decomposition α := |
| 131 | +let i := some s.exists_compl_positive_negative in |
| 132 | +let hi := some_spec s.exists_compl_positive_negative in |
| 133 | +{ pos_part := s.to_measure_of_zero_le i hi.1 hi.2.1, |
| 134 | + neg_part := s.to_measure_of_le_zero iᶜ hi.1.compl hi.2.2, |
| 135 | + pos_part_finite := infer_instance, |
| 136 | + neg_part_finite := infer_instance, |
| 137 | + mutually_singular := |
| 138 | + begin |
| 139 | + refine ⟨iᶜ, hi.1.compl, _, _⟩, |
| 140 | + { rw [to_measure_of_zero_le_apply _ _ hi.1 hi.1.compl], simpa }, |
| 141 | + { rw [to_measure_of_le_zero_apply _ _ hi.1.compl hi.1.compl.compl], simpa } |
| 142 | + end } |
| 143 | + |
| 144 | +lemma to_jordan_decomposition_spec (s : signed_measure α) : |
| 145 | + ∃ (i : set α) (hi₁ : measurable_set i) (hi₂ : 0 ≤[i] s) (hi₃ : s ≤[iᶜ] 0), |
| 146 | + s.to_jordan_decomposition.pos_part = s.to_measure_of_zero_le i hi₁ hi₂ ∧ |
| 147 | + s.to_jordan_decomposition.neg_part = s.to_measure_of_le_zero iᶜ hi₁.compl hi₃ := |
| 148 | +begin |
| 149 | + set i := some s.exists_compl_positive_negative, |
| 150 | + obtain ⟨hi₁, hi₂, hi₃⟩ := some_spec s.exists_compl_positive_negative, |
| 151 | + exact ⟨i, hi₁, hi₂, hi₃, rfl, rfl⟩, |
| 152 | +end |
| 153 | + |
| 154 | +/-- **The Jordan decomposition theorem**: Given a signed measure `s`, there exists a pair of |
| 155 | +mutually singular measures `μ` and `ν` such that `s = μ - ν`. In this case, the measures `μ` |
| 156 | +and `ν` are given by `s.to_jordan_decomposition.pos_part` and |
| 157 | +`s.to_jordan_decomposition.neg_part` respectively. |
| 158 | +
|
| 159 | +Note that we use `measure_theory.jordan_decomposition.to_signed_measure` to represent the |
| 160 | +signed measure corresponding to |
| 161 | +`s.to_jordan_decomposition.pos_part - s.to_jordan_decomposition.neg_part`. -/ |
| 162 | +@[simp] lemma to_signed_measure_to_jordan_decomposition (s : signed_measure α) : |
| 163 | + s.to_jordan_decomposition.to_signed_measure = s := |
| 164 | +begin |
| 165 | + obtain ⟨i, hi₁, hi₂, hi₃, hμ, hν⟩ := s.to_jordan_decomposition_spec, |
| 166 | + simp only [jordan_decomposition.to_signed_measure, hμ, hν], |
| 167 | + ext k hk, |
| 168 | + rw [to_signed_measure_sub_apply hk, to_measure_of_zero_le_apply _ hi₂ hi₁ hk, |
| 169 | + to_measure_of_le_zero_apply _ hi₃ hi₁.compl hk], |
| 170 | + simp only [ennreal.coe_to_real, subtype.coe_mk, ennreal.some_eq_coe, sub_neg_eq_add], |
| 171 | + rw [← of_union _ (measurable_set.inter hi₁ hk) (measurable_set.inter hi₁.compl hk), |
| 172 | + set.inter_comm i, set.inter_comm iᶜ, set.inter_union_compl _ _], |
| 173 | + { apply_instance }, |
| 174 | + { rintro x ⟨⟨hx₁, _⟩, hx₂, _⟩, |
| 175 | + exact false.elim (hx₂ hx₁) } |
| 176 | +end |
| 177 | + |
| 178 | +section |
| 179 | + |
| 180 | +variables {u v w : set α} |
| 181 | + |
| 182 | +/-- A subset `v` of a null-set `w` has zero measure if `w` is a subset of a positive set `u`. -/ |
| 183 | +lemma subset_positive_null_set |
| 184 | + (hu : measurable_set u) (hv : measurable_set v) (hw : measurable_set w) |
| 185 | + (hsu : 0 ≤[u] s) (hw₁ : s w = 0) (hw₂ : w ⊆ u) (hwt : v ⊆ w) : s v = 0 := |
| 186 | +begin |
| 187 | + have : s v + s (w \ v) = 0, |
| 188 | + { rw [← hw₁, ← of_union set.disjoint_diff hv (hw.diff hv), |
| 189 | + set.union_diff_self, set.union_eq_self_of_subset_left hwt], |
| 190 | + apply_instance }, |
| 191 | + have h₁ := nonneg_of_zero_le_restrict _ (restrict_le_restrict_subset _ _ hu hsu (hwt.trans hw₂)), |
| 192 | + have h₂ := nonneg_of_zero_le_restrict _ |
| 193 | + (restrict_le_restrict_subset _ _ hu hsu ((w.diff_subset v).trans hw₂)), |
| 194 | + linarith, |
| 195 | +end |
| 196 | + |
| 197 | +/-- A subset `v` of a null-set `w` has zero measure if `w` is a subset of a negative set `u`. -/ |
| 198 | +lemma subset_negative_null_set |
| 199 | + (hu : measurable_set u) (hv : measurable_set v) (hw : measurable_set w) |
| 200 | + (hsu : s ≤[u] 0) (hw₁ : s w = 0) (hw₂ : w ⊆ u) (hwt : v ⊆ w) : s v = 0 := |
| 201 | +begin |
| 202 | + rw [← s.neg_le_neg_iff _ hu, neg_zero] at hsu, |
| 203 | + have := subset_positive_null_set hu hv hw hsu, |
| 204 | + simp only [pi.neg_apply, neg_eq_zero, coe_neg] at this, |
| 205 | + exact this hw₁ hw₂ hwt, |
| 206 | +end |
| 207 | + |
| 208 | +/-- If the symmetric difference of two positive sets is a null-set, then so are the differences |
| 209 | +between the two sets. -/ |
| 210 | +lemma of_diff_eq_zero_of_symm_diff_eq_zero_positive |
| 211 | + (hu : measurable_set u) (hv : measurable_set v) |
| 212 | + (hsu : 0 ≤[u] s) (hsv : 0 ≤[v] s) (hs : s (u Δ v) = 0) : |
| 213 | + s (u \ v) = 0 ∧ s (v \ u) = 0 := |
| 214 | +begin |
| 215 | + rw restrict_le_restrict_iff at hsu hsv, |
| 216 | + have a := hsu (hu.diff hv) (u.diff_subset v), |
| 217 | + have b := hsv (hv.diff hu) (v.diff_subset u), |
| 218 | + erw [of_union (set.disjoint_of_subset_left (u.diff_subset v) set.disjoint_diff) |
| 219 | + (hu.diff hv) (hv.diff hu)] at hs, |
| 220 | + rw zero_apply at a b, |
| 221 | + split, |
| 222 | + all_goals { linarith <|> apply_instance <|> assumption }, |
| 223 | +end |
| 224 | + |
| 225 | +/-- If the symmetric difference of two negative sets is a null-set, then so are the differences |
| 226 | +between the two sets. -/ |
| 227 | +lemma of_diff_eq_zero_of_symm_diff_eq_zero_negative |
| 228 | + (hu : measurable_set u) (hv : measurable_set v) |
| 229 | + (hsu : s ≤[u] 0) (hsv : s ≤[v] 0) (hs : s (u Δ v) = 0) : |
| 230 | + s (u \ v) = 0 ∧ s (v \ u) = 0 := |
| 231 | +begin |
| 232 | + rw [← s.neg_le_neg_iff _ hu, neg_zero] at hsu, |
| 233 | + rw [← s.neg_le_neg_iff _ hv, neg_zero] at hsv, |
| 234 | + have := of_diff_eq_zero_of_symm_diff_eq_zero_positive hu hv hsu hsv, |
| 235 | + simp only [pi.neg_apply, neg_eq_zero, coe_neg] at this, |
| 236 | + exact this hs, |
| 237 | +end |
| 238 | + |
| 239 | +lemma of_inter_eq_of_symm_diff_eq_zero_positive |
| 240 | + (hu : measurable_set u) (hv : measurable_set v) (hw : measurable_set w) |
| 241 | + (hsu : 0 ≤[u] s) (hsv : 0 ≤[v] s) (hs : s (u Δ v) = 0) : |
| 242 | + s (w ∩ u) = s (w ∩ v) := |
| 243 | +begin |
| 244 | + have hwuv : s ((w ∩ u) Δ (w ∩ v)) = 0, |
| 245 | + { refine subset_positive_null_set (hu.union hv) ((hw.inter hu).symm_diff (hw.inter hv)) |
| 246 | + (hu.symm_diff hv) (restrict_le_restrict_union _ _ hu hsu hv hsv) hs _ _, |
| 247 | + { exact symm_diff_le_sup u v }, |
| 248 | + { rintro x (⟨⟨hxw, hxu⟩, hx⟩ | ⟨⟨hxw, hxv⟩, hx⟩); |
| 249 | + rw [set.mem_inter_eq, not_and] at hx, |
| 250 | + { exact or.inl ⟨hxu, hx hxw⟩ }, |
| 251 | + { exact or.inr ⟨hxv, hx hxw⟩ } } }, |
| 252 | + obtain ⟨huv, hvu⟩ := of_diff_eq_zero_of_symm_diff_eq_zero_positive |
| 253 | + (hw.inter hu) (hw.inter hv) |
| 254 | + (restrict_le_restrict_subset _ _ hu hsu (w.inter_subset_right u)) |
| 255 | + (restrict_le_restrict_subset _ _ hv hsv (w.inter_subset_right v)) hwuv, |
| 256 | + rw [← of_diff_of_diff_eq_zero (hw.inter hu) (hw.inter hv) hvu, huv, zero_add] |
| 257 | +end |
| 258 | + |
| 259 | +lemma of_inter_eq_of_symm_diff_eq_zero_negative |
| 260 | + (hu : measurable_set u) (hv : measurable_set v) (hw : measurable_set w) |
| 261 | + (hsu : s ≤[u] 0) (hsv : s ≤[v] 0) (hs : s (u Δ v) = 0) : |
| 262 | + s (w ∩ u) = s (w ∩ v) := |
| 263 | +begin |
| 264 | + rw [← s.neg_le_neg_iff _ hu, neg_zero] at hsu, |
| 265 | + rw [← s.neg_le_neg_iff _ hv, neg_zero] at hsv, |
| 266 | + have := of_inter_eq_of_symm_diff_eq_zero_positive hu hv hw hsu hsv, |
| 267 | + simp only [pi.neg_apply, neg_inj, neg_eq_zero, coe_neg] at this, |
| 268 | + exact this hs, |
| 269 | +end |
| 270 | + |
| 271 | +end |
| 272 | + |
| 273 | +end signed_measure |
| 274 | + |
| 275 | +namespace jordan_decomposition |
| 276 | + |
| 277 | +open measure vector_measure signed_measure function |
| 278 | + |
| 279 | +private lemma eq_of_pos_part_eq_pos_part {j₁ j₂ : jordan_decomposition α} |
| 280 | + (hj : j₁.pos_part = j₂.pos_part) (hj' : j₁.to_signed_measure = j₂.to_signed_measure) : |
| 281 | + j₁ = j₂ := |
| 282 | +begin |
| 283 | + ext1, |
| 284 | + { exact hj }, |
| 285 | + { rw ← to_signed_measure_eq_to_signed_measure_iff, |
| 286 | + suffices : j₁.pos_part.to_signed_measure - j₁.neg_part.to_signed_measure = |
| 287 | + j₁.pos_part.to_signed_measure - j₂.neg_part.to_signed_measure, |
| 288 | + { exact sub_right_inj.mp this }, |
| 289 | + convert hj' } |
| 290 | +end |
| 291 | + |
| 292 | +/-- The Jordan decomposition of a signed measure is unique. -/ |
| 293 | +theorem to_signed_measure_injective : |
| 294 | + injective $ @jordan_decomposition.to_signed_measure α _ := |
| 295 | +begin |
| 296 | + /- The main idea is that two Jordan decompositions of a signed measure provide two |
| 297 | + Hahn decompositions for that measure. Then, from `of_symm_diff_compl_positive_negative`, |
| 298 | + the symmetric difference of the two Hahn decompositions has measure zero, thus, allowing us to |
| 299 | + show the equality of the underlying measures of the Jordan decompositions. -/ |
| 300 | + intros j₁ j₂ hj, |
| 301 | + -- obtain the two Hahn decompositions from the Jordan decompositions |
| 302 | + obtain ⟨S, hS₁, hS₂, hS₃, hS₄, hS₅⟩ := j₁.exists_compl_positive_negative, |
| 303 | + obtain ⟨T, hT₁, hT₂, hT₃, hT₄, hT₅⟩ := j₂.exists_compl_positive_negative, |
| 304 | + rw ← hj at hT₂ hT₃, |
| 305 | + -- the symmetric differences of the two Hahn decompositions have measure zero |
| 306 | + obtain ⟨hST₁, -⟩ := of_symm_diff_compl_positive_negative hS₁.compl hT₁.compl |
| 307 | + ⟨hS₃, (compl_compl S).symm ▸ hS₂⟩ ⟨hT₃, (compl_compl T).symm ▸ hT₂⟩, |
| 308 | + -- it suffices to show the Jordan decompositions have the same positive parts |
| 309 | + refine eq_of_pos_part_eq_pos_part _ hj, |
| 310 | + ext1 i hi, |
| 311 | + -- we see that the positive parts of the two Jordan decompositions are equal to their |
| 312 | + -- associated signed measures restricted on their associated Hahn decompositions |
| 313 | + have hμ₁ : (j₁.pos_part i).to_real = j₁.to_signed_measure (i ∩ Sᶜ), |
| 314 | + { rw [to_signed_measure, to_signed_measure_sub_apply (hi.inter hS₁.compl), |
| 315 | + show j₁.neg_part (i ∩ Sᶜ) = 0, by exact nonpos_iff_eq_zero.1 |
| 316 | + (hS₅ ▸ measure_mono (set.inter_subset_right _ _)), |
| 317 | + ennreal.zero_to_real, sub_zero], |
| 318 | + conv_lhs { rw ← set.inter_union_compl i S }, |
| 319 | + rw [measure_union, show j₁.pos_part (i ∩ S) = 0, by exact nonpos_iff_eq_zero.1 |
| 320 | + (hS₄ ▸ measure_mono (set.inter_subset_right _ _)), zero_add], |
| 321 | + { refine set.disjoint_of_subset_left (set.inter_subset_right _ _) |
| 322 | + (set.disjoint_of_subset_right (set.inter_subset_right _ _) disjoint_compl_right) }, |
| 323 | + { exact hi.inter hS₁ }, |
| 324 | + { exact hi.inter hS₁.compl } }, |
| 325 | + have hμ₂ : (j₂.pos_part i).to_real = j₂.to_signed_measure (i ∩ Tᶜ), |
| 326 | + { rw [to_signed_measure, to_signed_measure_sub_apply (hi.inter hT₁.compl), |
| 327 | + show j₂.neg_part (i ∩ Tᶜ) = 0, by exact nonpos_iff_eq_zero.1 |
| 328 | + (hT₅ ▸ measure_mono (set.inter_subset_right _ _)), |
| 329 | + ennreal.zero_to_real, sub_zero], |
| 330 | + conv_lhs { rw ← set.inter_union_compl i T }, |
| 331 | + rw [measure_union, show j₂.pos_part (i ∩ T) = 0, by exact nonpos_iff_eq_zero.1 |
| 332 | + (hT₄ ▸ measure_mono (set.inter_subset_right _ _)), zero_add], |
| 333 | + { exact set.disjoint_of_subset_left (set.inter_subset_right _ _) |
| 334 | + (set.disjoint_of_subset_right (set.inter_subset_right _ _) disjoint_compl_right) }, |
| 335 | + { exact hi.inter hT₁ }, |
| 336 | + { exact hi.inter hT₁.compl } }, |
| 337 | + -- since the two signed measures associated with the Jordan decompositions are the same, |
| 338 | + -- and the symmetric difference of the Hahn decompositions have measure zero, the result follows |
| 339 | + rw [← ennreal.to_real_eq_to_real (measure_lt_top _ _) (measure_lt_top _ _), |
| 340 | + hμ₁, hμ₂, ← hj], |
| 341 | + exact of_inter_eq_of_symm_diff_eq_zero_positive hS₁.compl hT₁.compl hi hS₃ hT₃ hST₁, |
| 342 | + all_goals { apply_instance }, |
| 343 | +end |
| 344 | + |
| 345 | +@[simp] |
| 346 | +lemma to_jordan_decomposition_to_signed_measure (j : jordan_decomposition α) : |
| 347 | + (j.to_signed_measure).to_jordan_decomposition = j := |
| 348 | +(@to_signed_measure_injective _ _ j (j.to_signed_measure).to_jordan_decomposition (by simp)).symm |
| 349 | + |
| 350 | +end jordan_decomposition |
| 351 | + |
| 352 | +namespace signed_measure |
| 353 | + |
| 354 | +open jordan_decomposition |
| 355 | + |
| 356 | +/-- `measure_theory.signed_measure.to_jordan_decomposition` and |
| 357 | +`measure_theory.jordan_decomposition.to_signed_measure` form a `equiv`. -/ |
| 358 | +@[simps apply symm_apply] |
| 359 | +def to_jordan_decomposition_equiv (α : Type*) [measurable_space α] : |
| 360 | + signed_measure α ≃ jordan_decomposition α := |
| 361 | +{ to_fun := to_jordan_decomposition, |
| 362 | + inv_fun := to_signed_measure, |
| 363 | + left_inv := to_signed_measure_to_jordan_decomposition, |
| 364 | + right_inv := to_jordan_decomposition_to_signed_measure } |
| 365 | + |
| 366 | +lemma to_jordan_decomposition_zero : (0 : signed_measure α).to_jordan_decomposition = 0 := |
| 367 | +begin |
| 368 | + apply to_signed_measure_injective, |
| 369 | + simp [to_signed_measure_zero], |
| 370 | +end |
| 371 | + |
| 372 | +lemma to_jordan_decomposition_neg (s : signed_measure α) : |
| 373 | + (-s).to_jordan_decomposition = -s.to_jordan_decomposition := |
| 374 | +begin |
| 375 | + apply to_signed_measure_injective, |
| 376 | + simp [to_signed_measure_neg], |
| 377 | +end |
| 378 | + |
| 379 | +end signed_measure |
| 380 | + |
| 381 | +end measure_theory |
0 commit comments