@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Floris van Doorn
5
5
-/
6
6
import measure_theory.measure.giry_monad
7
+ import dynamics.ergodic.measure_preserving
7
8
import measure_theory.integral.set_integral
8
9
9
10
/-!
@@ -335,34 +336,14 @@ by simp_rw [prod_apply (hs.prod ht), mk_preimage_prod_right_eq_if, measure_if,
335
336
lintegral_indicator _ hs, lintegral_const, restrict_apply measurable_set.univ,
336
337
univ_inter, mul_comm]
337
338
338
- local attribute [instance] nonempty_measurable_superset
339
339
/-- If we don't assume measurability of `s` and `t`, we can bound the measure of their product. -/
340
340
lemma prod_prod_le (s : set α) (t : set β) : μ.prod ν (s.prod t) ≤ μ s * ν t :=
341
- begin
342
- by_cases hs0 : μ s = 0 ,
343
- { rcases (exists_measurable_superset_of_null hs0) with ⟨s', hs', h2s', h3s'⟩,
344
- convert measure_mono (prod_mono hs' (subset_univ _)),
345
- simp_rw [hs0, prod_prod h2s' measurable_set.univ, h3s', zero_mul] },
346
- by_cases hti : ν t = ∞,
347
- { convert le_top, simp_rw [hti, ennreal.mul_top, hs0, if_false] },
348
- rw [measure_eq_infi' μ],
349
- simp_rw [ennreal.infi_mul hti],
350
- refine le_infi _,
351
- rintro ⟨s', h1s', h2s'⟩,
352
- rw [subtype.coe_mk],
353
- by_cases ht0 : ν t = 0 ,
354
- { rcases (exists_measurable_superset_of_null ht0) with ⟨t', ht', h2t', h3t'⟩,
355
- convert measure_mono (prod_mono (subset_univ _) ht'),
356
- simp_rw [ht0, prod_prod measurable_set.univ h2t', h3t', mul_zero] },
357
- by_cases hsi : μ s' = ∞,
358
- { convert le_top, simp_rw [hsi, ennreal.top_mul, ht0, if_false] },
359
- rw [measure_eq_infi' ν],
360
- simp_rw [ennreal.mul_infi hsi],
361
- refine le_infi _,
362
- rintro ⟨t', h1t', h2t'⟩,
363
- convert measure_mono (prod_mono h1s' h1t'),
364
- simp [prod_prod h2s' h2t'],
365
- end
341
+ calc μ.prod ν (s.prod t) ≤ μ.prod ν ((to_measurable μ s).prod (to_measurable ν t)) :
342
+ measure_mono $ set.prod_mono (subset_to_measurable _ _) (subset_to_measurable _ _)
343
+ ... = μ (to_measurable μ s) * ν (to_measurable ν t) :
344
+ prod_prod (measurable_set_to_measurable _ _) (measurable_set_to_measurable _ _)
345
+ ... = μ s * ν t :
346
+ by rw [measure_to_measurable, measure_to_measurable]
366
347
367
348
lemma ae_measure_lt_top {s : set (α × β)} (hs : measurable_set s)
368
349
(h2s : (μ.prod ν) s ≠ ∞) : ∀ᵐ x ∂μ, ν (prod.mk x ⁻¹' s) < ∞ :=
@@ -561,6 +542,53 @@ begin
561
542
end
562
543
563
544
end measure
545
+
546
+ namespace measure_preserving
547
+
548
+ open measure
549
+
550
+ variables {δ : Type *} [measurable_space δ] {μa : measure α} {μb : measure β}
551
+ {μc : measure γ} {μd : measure δ}
552
+
553
+ lemma skew_product [sigma_finite μb] [sigma_finite μd]
554
+ {f : α → β} (hf : measure_preserving f μa μb) {g : α → γ → δ}
555
+ (hgm : measurable (uncurry g)) (hg : ∀ᵐ x ∂μa, map (g x) μc = μd) :
556
+ measure_preserving (λ p : α × γ, (f p.1 , g p.1 p.2 )) (μa.prod μc) (μb.prod μd) :=
557
+ begin
558
+ classical,
559
+ have : measurable (λ p : α × γ, (f p.1 , g p.1 p.2 )) := (hf.1 .comp measurable_fst).prod_mk hgm,
560
+ /- if `μa = 0`, then the lemma is trivial, otherwise we can use `hg`
561
+ to deduce `sigma_finite μc`. -/
562
+ rcases eq_or_ne μa 0 with (rfl|ha),
563
+ { rw [← hf.map_eq, zero_prod, (map f).map_zero, zero_prod],
564
+ exact ⟨this , (map _).map_zero⟩ },
565
+ haveI : sigma_finite μc,
566
+ { rcases (ae_ne_bot.2 ha).nonempty_of_mem hg with ⟨x, hx : map (g x) μc = μd⟩,
567
+ exact sigma_finite.of_map _ hgm.of_uncurry_left (by rwa hx) },
568
+ -- Thus we can apply `measure.prod_eq` to prove equality of measures.
569
+ refine ⟨this , (prod_eq $ λ s t hs ht, _).symm⟩,
570
+ rw [map_apply this (hs.prod ht)],
571
+ refine (prod_apply (this $ hs.prod ht)).trans _,
572
+ have : ∀ᵐ x ∂μa, μc ((λ y, (f x, g x y)) ⁻¹' s.prod t) = indicator (f ⁻¹' s) (λ y, μd t) x,
573
+ { refine hg.mono (λ x hx, _), unfreezingI { subst hx },
574
+ simp only [mk_preimage_prod_right_fn_eq_if, indicator_apply, mem_preimage],
575
+ split_ifs,
576
+ exacts [(map_apply hgm.of_uncurry_left ht).symm, measure_empty] },
577
+ simp only [preimage_preimage],
578
+ rw [lintegral_congr_ae this , lintegral_indicator _ (hf.1 hs),
579
+ set_lintegral_const, hf.measure_preimage hs, mul_comm]
580
+ end
581
+
582
+ /-- If `f : α → β` sends the measure `μa` to `μb` and `g : γ → δ` sends the measure `μc` to `μd`,
583
+ then `prod.map f g` sends `μa.prod μc` to `μb.prod μd`. -/
584
+ protected lemma prod [sigma_finite μb] [sigma_finite μd] {f : α → β} {g : γ → δ}
585
+ (hf : measure_preserving f μa μb) (hg : measure_preserving g μc μd) :
586
+ measure_preserving (prod.map f g) (μa.prod μc) (μb.prod μd) :=
587
+ have measurable (uncurry $ λ _ : α, g), from (hg.1 .comp measurable_snd),
588
+ hf.skew_product this $ filter.eventually_of_forall $ λ _, hg.map_eq
589
+
590
+ end measure_preserving
591
+
564
592
end measure_theory
565
593
566
594
open measure_theory.measure
0 commit comments