|
877 | 877 | @[simp] lemma restrict_apply (ht : measurable_set t) : μ.restrict s t = μ (t ∩ s) :=
|
878 | 878 | by simp [← restrictₗ_apply, restrictₗ, ht]
|
879 | 879 |
|
| 880 | +lemma restrict_eq_self (h_meas_t : measurable_set t) (h : t ⊆ s) : μ.restrict s t = μ t := |
| 881 | +by rw [restrict_apply h_meas_t, subset_iff_inter_eq_left.1 h] |
| 882 | + |
| 883 | +lemma restrict_apply_self (μ:measure α) (h_meas_s : measurable_set s) : |
| 884 | + (μ.restrict s) s = μ s := (restrict_eq_self h_meas_s (set.subset.refl _)) |
| 885 | + |
880 | 886 | lemma restrict_apply_univ (s : set α) : μ.restrict s univ = μ s :=
|
881 | 887 | by rw [restrict_apply measurable_set.univ, set.univ_inter]
|
882 | 888 |
|
@@ -1929,6 +1935,67 @@ end
|
1929 | 1935 |
|
1930 | 1936 | end measure_sub
|
1931 | 1937 |
|
| 1938 | +lemma restrict_sub_eq_restrict_sub_restrict (h_meas_s : measurable_set s) : |
| 1939 | + (μ - ν).restrict s = (μ.restrict s) - (ν.restrict s) := |
| 1940 | +begin |
| 1941 | + repeat {rw sub_def}, |
| 1942 | + have h_nonempty : {d | μ ≤ d + ν}.nonempty, |
| 1943 | + { apply @set.nonempty_of_mem _ _ μ, rw mem_set_of_eq, intros t h_meas, |
| 1944 | + apply le_add_right (le_refl (μ t)) }, |
| 1945 | + rw restrict_Inf_eq_Inf_restrict h_nonempty h_meas_s, |
| 1946 | + apply le_antisymm, |
| 1947 | + { apply @Inf_le_Inf_of_forall_exists_le (measure α) _, |
| 1948 | + intros ν' h_ν'_in, rw mem_set_of_eq at h_ν'_in, apply exists.intro (ν'.restrict s), |
| 1949 | + split, |
| 1950 | + { rw mem_image, apply exists.intro (ν' + (⊤ : measure_theory.measure α).restrict sᶜ), |
| 1951 | + rw mem_set_of_eq, |
| 1952 | + split, |
| 1953 | + { rw [add_assoc, add_comm _ ν, ← add_assoc, measure_theory.measure.le_iff], |
| 1954 | + intros t h_meas_t, |
| 1955 | + have h_inter_inter_eq_inter : ∀ t' : set α , t ∩ t' ∩ t' = t ∩ t', |
| 1956 | + { intro t', rw set.inter_eq_self_of_subset_left, apply set.inter_subset_right t t' }, |
| 1957 | + have h_meas_t_inter_s : measurable_set (t ∩ s) := |
| 1958 | + h_meas_t.inter h_meas_s, |
| 1959 | + repeat {rw measure_eq_inter_diff h_meas_t h_meas_s, rw set.diff_eq}, |
| 1960 | + apply add_le_add _ _; rw add_apply, |
| 1961 | + { apply le_add_right _, |
| 1962 | + rw add_apply, |
| 1963 | + rw ← @restrict_eq_self _ _ μ s _ h_meas_t_inter_s (set.inter_subset_right _ _), |
| 1964 | + rw ← @restrict_eq_self _ _ ν s _ h_meas_t_inter_s (set.inter_subset_right _ _), |
| 1965 | + apply h_ν'_in _ h_meas_t_inter_s }, |
| 1966 | + cases (@set.eq_empty_or_nonempty _ (t ∩ sᶜ)) with h_inter_empty h_inter_nonempty, |
| 1967 | + { simp [h_inter_empty] }, |
| 1968 | + { have h_meas_inter_compl := |
| 1969 | + h_meas_t.inter (measurable_set.compl h_meas_s), |
| 1970 | + rw [restrict_apply h_meas_inter_compl, h_inter_inter_eq_inter sᶜ], |
| 1971 | + have h_mu_le_add_top : μ ≤ ν' + ν + ⊤, |
| 1972 | + { rw add_comm, |
| 1973 | + have h_le_top : μ ≤ ⊤ := le_top, |
| 1974 | + apply (λ t₂ h_meas, le_add_right (h_le_top t₂ h_meas)) }, |
| 1975 | + apply h_mu_le_add_top _ h_meas_inter_compl } }, |
| 1976 | + { ext1 t h_meas_t, |
| 1977 | + simp [restrict_apply h_meas_t, |
| 1978 | + restrict_apply (h_meas_t.inter h_meas_s), |
| 1979 | + set.inter_assoc] } }, |
| 1980 | + { apply restrict_le_self } }, |
| 1981 | + { apply @Inf_le_Inf_of_forall_exists_le (measure α) _, |
| 1982 | + |
| 1983 | + intros s h_s_in, cases h_s_in with t h_t, cases h_t with h_t_in h_t_eq, subst s, |
| 1984 | + apply exists.intro (t.restrict s), split, |
| 1985 | + { rw [set.mem_set_of_eq, ← restrict_add], |
| 1986 | + apply restrict_mono (set.subset.refl _) h_t_in }, |
| 1987 | + { apply le_refl _ } }, |
| 1988 | +end |
| 1989 | + |
| 1990 | +lemma sub_apply_eq_zero_of_restrict_le_restrict |
| 1991 | + (h_le : μ.restrict s ≤ ν.restrict s) (h_meas_s : measurable_set s) : |
| 1992 | + (μ - ν) s = 0 := |
| 1993 | +begin |
| 1994 | + rw [← restrict_apply_self _ h_meas_s, restrict_sub_eq_restrict_sub_restrict, |
| 1995 | + sub_eq_zero_of_le], |
| 1996 | + repeat {simp [*]}, |
| 1997 | +end |
| 1998 | + |
1932 | 1999 | end measure
|
1933 | 2000 |
|
1934 | 2001 | end measure_theory
|
|
0 commit comments