diff --git a/thys/Ergodic_Theory/Gouezel_Karlsson.thy b/thys/Ergodic_Theory/Gouezel_Karlsson.thy --- a/thys/Ergodic_Theory/Gouezel_Karlsson.thy +++ b/thys/Ergodic_Theory/Gouezel_Karlsson.thy @@ -1,1810 +1,1810 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Gouezel-Karlsson\ theory Gouezel_Karlsson imports Asymptotic_Density Kingman begin text \This section is devoted to the proof of the main ergodic result of the article "Subadditive and multiplicative ergodic theorems" by Gouezel and Karlsson~\cite{gouezel_karlsson}. It is a version of Kingman theorem ensuring that, for subadditive cocycles, there are almost surely many times $n$ where the cocycle is nearly additive at \emph{all} times between $0$ and $n$. This theorem is then used in this article to construct horofunctions characterizing the behavior at infinity of compositions of semi-contractions. This requires too many further notions to be implemented in current Isabelle/HOL, but the main ergodic result is completely proved below, in Theorem~\verb+Gouezel_Karlsson+, following the arguments in the paper (but in a slightly more general setting here as we are not making any ergodicity assumption). To simplify the exposition, the theorem is proved assuming that the limit of the subcocycle vanishes almost everywhere, in the locale \verb+Gouezel_Karlsson_Kingman+. The final result is proved by an easy reduction to this case. The main steps of the proof are as follows: \begin{itemize} \item assume first that the map is invertible, and consider the inverse map and the corresponding inverse subcocycle. With combinatorial arguments that only work for this inverse subcocycle, we control the density of bad times given some allowed error $d>0$, in a precise quantitative way, in Lemmas~\verb+upper_density_all_times+ and~\verb+upper_density_large_k+. We put these estimates together in Lemma~\verb+upper_density_delta+. \item These estimates are then transfered to the original time direction and the original subcocycle in Lemma~\verb+upper_density_good_direction_invertible+. The fact that we have quantitative estimates in terms of asymptotic densities is central here, just having some infiniteness statement would not be enough. \item The invertibility assumption is removed in Lemma~\verb+upper_density_good_direction+ by using the result in the natural extension. \item Finally, the main result is deduced in Lemma~\verb+infinite_AE+ (still assuming that the asymptotic limit vanishes almost everywhere), and in full generality in Theorem~\verb+Gouezel_Karlsson_Kingman+. \end{itemize} \ lemma upper_density_eventually_measure: fixes a::real assumes [measurable]: "\n. {x \ space M. P x n} \ sets M" and "emeasure M {x \ space M. upper_asymptotic_density {n. P x n} < a} > b" shows "\N. emeasure M {x \ space M. \n \ N. card ({n. P x n} \ {.. b" proof - define G where "G = {x \ space M. upper_asymptotic_density {n. P x n} < a}" define H where "H = (\N. {x \ space M. \n \ N. card ({n. P x n} \ {.. sets M" "\N. H N \ sets M" unfolding G_def H_def by auto have "G \ (\N. H N)" proof fix x assume "x \ G" then have "x \ space M" unfolding G_def by simp have "eventually (\n. card({n. P x n} \ {..x \ G\ unfolding G_def using upper_asymptotic_densityD by auto then obtain N where "\n. n \ N \ card({n. P x n} \ {.. H N" unfolding H_def using \x \ space M\ by auto then show "x \ (\N. H N)" by blast qed have "b < emeasure M G" using assms(2) unfolding G_def by simp also have "... \ emeasure M (\N. H N)" apply (rule emeasure_mono) using \G \ (\N. H N)\ by auto finally have "emeasure M (\N. H N) > b" by simp moreover have "(\N. emeasure M (H N)) \ emeasure M (\N. H N)" apply (rule Lim_emeasure_incseq) unfolding H_def incseq_def by auto ultimately have "eventually (\N. emeasure M (H N) > b) sequentially" by (simp add: order_tendsto_iff) then obtain N where "emeasure M (H N) > b" using eventually_False_sequentially eventually_mono by blast then show ?thesis unfolding H_def by blast qed locale Gouezel_Karlsson_Kingman = pmpt + fixes u::"nat \ 'a \ real" assumes subu: "subcocycle u" and subu_fin: "subcocycle_avg_ereal u > -\" and subu_0: "AE x in M. subcocycle_lim u x = 0" begin lemma int_u [measurable]: "integrable M (u n)" using subu unfolding subcocycle_def by auto text \Next lemma is Lemma 2.1 in~\cite{gouezel_karlsson}.\ lemma upper_density_all_times: assumes "d > (0::real)" shows "\c> (0::real). emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) x \ - c * l} < d} > 1 - d" proof - define f where "f = (\x. abs (u 1 x))" have [measurable]: "f \ borel_measurable M" unfolding f_def by auto define G where "G = {x \ space M. (\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x \ (\n. u n x / n) \ 0}" have [measurable]: "G \ sets M" unfolding G_def by auto have "AE x in M. (\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" apply (rule birkhoff_theorem_AE_nonergodic) using subu unfolding f_def subcocycle_def by auto moreover have "AE x in M. (\n. u n x / n) \ 0" using subu_0 kingman_theorem_nonergodic(1)[OF subu subu_fin] by auto ultimately have "AE x in M. x \ G" unfolding G_def by auto then have "emeasure M G = 1" by (simp add: emeasure_eq_1_AE) define V where "V = (\c x. {n. \l \ {1..n}. u n x - u (n-l) x \ - c * l})" define Good where "Good = (\c. {x \ G. upper_asymptotic_density (V c x) < d})" have [measurable]: "Good c \ sets M" for c unfolding Good_def V_def by auto have I: "upper_asymptotic_density (V c x) \ real_cond_exp M Invariants f x / c" if "c>0" "x \ G" for c x proof - have [simp]: "c>0" "c \ 0" "c \ 0" using \c>0\ by auto define U where "U = (\n. abs(u 0 x) + birkhoff_sum f n x - c * card (V c x \ {1..n}))" have main: "u n x \ U n" for n proof (rule nat_less_induct) fix n assume H: "\m U m" consider "n = 0" | "n\1 \ n \ V c x" | "n\1 \ n \ V c x" by linarith then show "u n x \ U n" proof (cases) assume "n = 0" then show ?thesis unfolding U_def by auto next assume A: "n\1 \ n \ V c x" then have "n \ 1" by simp then have "n-1 {n}" using \1 \ n\ atLeastLessThanSuc by auto then have *: "card (V c x \ {1..n}) = card (V c x \ {1..n-1})" using A by auto have "u n x \ u (n-1) x + u 1 ((T^^(n-1)) x)" using \n\1\ subu unfolding subcocycle_def by (metis le_add_diff_inverse2) also have "... \ U (n-1) + f ((T^^(n-1)) x)" unfolding f_def using H \n-1 by auto also have "... = abs(u 0 x) + birkhoff_sum f (n-1) x + f ((T^^(n-1)) x) - c * card (V c x \ {1..n-1})" unfolding U_def by auto also have "... = abs(u 0 x) + birkhoff_sum f n x - c * card (V c x \ {1..n})" using * birkhoff_sum_cocycle[of f "n-1" 1 x] \1 \ n\ by auto also have "... = U n" unfolding U_def by simp finally show ?thesis by auto next assume B: "n\1 \ n \ V c x" then obtain l where l: "l\{1..n}" "u n x - u (n-l) x \ - c * l" unfolding V_def by blast then have "n-l < n" by simp have m: "- (r * ra) - r * rb = - (r * (rb + ra))" for r ra rb::real by (simp add: algebra_simps) have "card(V c x \ {1..n}) \ card ((V c x \ {1..n-l}) \ {n-l+1..n})" by (rule card_mono, auto) also have "... \ card (V c x \ {1..n-l}) + card {n-l+1..n}" by (rule card_Un_le) also have "... \ card (V c x \ {1..n-l}) + l" by auto finally have "card(V c x \ {1..n}) \ card (V c x \ {1..n-l}) + real l" by auto then have *: "-c * card (V c x \ {1..n-l}) - c * l \ -c * card(V c x \ {1..n})" using m by auto have "birkhoff_sum f ((n-l) + l) x = birkhoff_sum f (n-l) x + birkhoff_sum f l ((T^^(n-l))x)" by (rule birkhoff_sum_cocycle) moreover have "birkhoff_sum f l ((T^^(n-l))x) \ 0" unfolding f_def birkhoff_sum_def using sum_nonneg by auto ultimately have **: "birkhoff_sum f (n-l) x \ birkhoff_sum f n x" using l(1) by auto have "u n x \ u (n-l) x - c * l" using l by simp also have "... \ U (n-l) - c* l" using H \n-l < n\ by auto also have "... = abs(u 0 x) + birkhoff_sum f (n-l) x - c * card (V c x \ {1..n-l}) - c*l" unfolding U_def by auto also have "... \ abs(u 0 x) + birkhoff_sum f n x - c * card (V c x \ {1..n})" using * ** by simp finally show ?thesis unfolding U_def by auto qed qed have "(\n. abs(u 0 x) * (1/n) + birkhoff_sum f n x / n - u n x / n) \ abs(u 0 x) * 0 + real_cond_exp M Invariants f x - 0" apply (intro tendsto_intros) using \x \ G\ unfolding G_def by auto moreover have "(abs(u 0 x) + birkhoff_sum f n x - u n x)/n = abs(u 0 x) * (1/n) + birkhoff_sum f n x / n - u n x / n" for n by (auto simp add: add_divide_distrib diff_divide_distrib) ultimately have "(\n. (abs(u 0 x) + birkhoff_sum f n x - u n x)/n) \ real_cond_exp M Invariants f x" by auto then have a: "limsup (\n. (abs(u 0 x) + birkhoff_sum f n x - u n x)/n) = real_cond_exp M Invariants f x" by (simp add: assms lim_imp_Limsup) have "c * card (V c x \ {1..n})/n \ (abs(u 0 x) + birkhoff_sum f n x - u n x)/n" for n using main[of n] unfolding U_def by (simp add: divide_right_mono) then have "limsup (\n. c * card (V c x \ {1..n})/n) \ limsup (\n. (abs(u 0 x) + birkhoff_sum f n x - u n x)/n)" by (simp add: Limsup_mono) then have b: "limsup (\n. c * card (V c x \ {1..n})/n) \ real_cond_exp M Invariants f x" using a by simp have "ereal(upper_asymptotic_density (V c x)) = limsup (\n. card (V c x \ {1..n})/n)" using upper_asymptotic_density_shift[of "V c x" 1 0] by auto also have "... = limsup (\n. ereal(1/c) * ereal(c * card (V c x \ {1..n})/n))" by auto also have "... = (1/c) * limsup (\n. c * card (V c x \ {1..n})/n)" by (rule limsup_ereal_mult_left, auto) also have "... \ ereal (1/c) * real_cond_exp M Invariants f x" by (rule ereal_mult_left_mono[OF b], auto) finally show "upper_asymptotic_density (V c x) \ real_cond_exp M Invariants f x / c" by auto qed { fix r::real obtain c::nat where "r / d < c" using reals_Archimedean2 by auto then have "r/d < real c+1" by auto then have "r / (real c+1) < d" using \d>0\ by (simp add: divide_less_eq mult.commute) then have "\c::nat. r / (real c+1) < d" by auto } then have unG: "(\c::nat. {x \ G. real_cond_exp M Invariants f x / (c+1) < d}) = G" by auto have *: "r < d * (real n + 1)" if "m \ n" "r < d * (real m + 1)" for m n r proof - have "d * (real m + 1) \ d * (real n + 1)" using \d>0\ \m \ n\ by auto then show ?thesis using \r < d * (real m + 1)\ by auto qed have "(\c. emeasure M {x \ G. real_cond_exp M Invariants f x / (real c+1) < d}) \ emeasure M (\c::nat. {x \ G. real_cond_exp M Invariants f x / (c+1) < d})" apply (rule Lim_emeasure_incseq) unfolding incseq_def by (auto simp add: divide_simps *) then have "(\c. emeasure M {x \ G. real_cond_exp M Invariants f x / (real c+1) < d}) \ emeasure M G" using unG by auto then have "(\c. emeasure M {x \ G. real_cond_exp M Invariants f x / (real c+1) < d}) \ 1" using \emeasure M G = 1\ by simp then have "eventually (\c. emeasure M {x \ G. real_cond_exp M Invariants f x / (real c+1) < d} > 1 - d) sequentially" apply (rule order_tendstoD) apply (insert \0, auto simp add: ennreal_1[symmetric] ennreal_lessI simp del: ennreal_1) done then obtain c0 where c0: "emeasure M {x \ G. real_cond_exp M Invariants f x / (real c0+1) < d} > 1 - d" using eventually_sequentially by auto define c where "c = real c0 + 1" then have "c > 0" by auto have *: "emeasure M {x \ G. real_cond_exp M Invariants f x / c < d} > 1 - d" unfolding c_def using c0 by auto have "{x \ G. real_cond_exp M Invariants f x / c < d} \ {x \ space M. upper_asymptotic_density (V c x) < d}" apply auto using G_def apply blast using I[OF \c>0\] by fastforce then have "emeasure M {x \ G. real_cond_exp M Invariants f x / c < d} \ emeasure M {x \ space M. upper_asymptotic_density (V c x) < d}" apply (rule emeasure_mono) unfolding V_def by auto then have "emeasure M {x \ space M. upper_asymptotic_density (V c x) < d} > 1 - d" using * by auto then show ?thesis unfolding V_def using \c>0\ by auto qed text \Next lemma is Lemma 2.2 in~\cite{gouezel_karlsson}.\ lemma upper_density_large_k: assumes "d > (0::real)" "d \ 1" shows "\k::nat. emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} < d} > 1 - d" proof - have [simp]: "d>0" "d \ 0" "d \ 0" using \d>0\ by auto define rho where "rho = d * d * d / 4" have [simp]: "rho > 0" "rho \ 0" "rho \ 0" unfolding rho_def using assms by auto text \First step: choose a time scale $s$ at which all the computations will be done. the integral of $u_s$ should be suitably small -- how small precisely is given by $\rho$.\ have "ennreal(\x. abs(u n x / n) \M) = (\\<^sup>+x. abs(u n x /n - subcocycle_lim u x) \M)" for n proof - have "ennreal(\x. abs(u n x / n) \M) = (\\<^sup>+x. abs(u n x /n) \M)" apply (rule nn_integral_eq_integral[symmetric]) using int_u by auto also have "... = (\\<^sup>+x. abs(u n x /n - subcocycle_lim u x) \M)" apply (rule nn_integral_cong_AE) using subu_0 by auto finally show ?thesis by simp qed moreover have "(\n. \\<^sup>+x. abs(u n x /n - subcocycle_lim u x) \M) \ 0" by (rule kingman_theorem_nonergodic(3)[OF subu subu_fin]) ultimately have "(\n. ennreal(\x. abs(u n x / n) \M)) \ 0" by auto then have "(\n. (\x. abs(u n x / n) \M)) \ 0" by (simp add: ennreal_0[symmetric] del: ennreal_0) then have "eventually (\n. (\x. abs(u n x / n) \M) < rho) sequentially" by (rule order_tendstoD(2), auto) then obtain s::nat where [simp]: "s>0" and s_int: "(\x. abs(u s x / s) \M) < rho" by (metis (mono_tags, lifting) neq0_conv eventually_sequentially gr_implies_not0 linorder_not_le of_nat_0_eq_iff order_refl zero_neq_one) text \Second step: a truncation argument, to decompose $|u_1|$ as a sum of a constant (its contribution will be small if $k$ is large at the end of the argument) and of a function with small integral).\ have "(\n. (\x. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x \M)) \ (\x. 0 \M)" proof (rule integral_dominated_convergence[where ?w = "\x. abs(u 1 x)"]) show "AE x in M. norm (abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x) \ abs(u 1 x)" for n unfolding indicator_def by auto { fix x have "abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x = (0::real)" if "n > abs(u 1 x)" for n::nat unfolding indicator_def using that by auto then have "eventually (\n. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x = 0) sequentially" by (metis (mono_tags, lifting) eventually_at_top_linorder reals_Archimedean2 less_le_trans of_nat_le_iff) then have "(\n. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x) \ 0" by (rule tendsto_eventually) } then show "AE x in M. (\n. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x) \ 0" by simp qed (auto simp add: int_u) then have "eventually (\n. (\x. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ n} x \M) < rho) sequentially" by (rule order_tendstoD(2), auto) then obtain Knat::nat where Knat: "Knat > 0" "(\x. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ Knat} x \M) < rho" by (metis (mono_tags, lifting) eventually_sequentially gr_implies_not0 neq0_conv linorder_not_le of_nat_0_eq_iff order_refl zero_neq_one) define K where "K = real Knat" then have [simp]: "K \ 0" "K>0" and K: "(\x. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ K} x \M) < rho" using Knat by auto define F where "F = (\x. abs(u 1 x) * indicator {x. abs(u 1 x) \ K} x)" have int_F [measurable]: "integrable M F" unfolding F_def apply (rule Bochner_Integration.integrable_bound[where ?f = "\x. abs(u 1 x)"]) unfolding indicator_def by (auto simp add: int_u) have "(\x. F x \M) = (\x. abs(u 1 x) * indicator {x \ space M. abs(u 1 x) \ K} x \M)" apply (rule integral_cong_AE) unfolding F_def by (auto simp add: indicator_def) then have F_int: "(\x. F x \M) < rho" using K by auto have F_pos: "F x \ 0" for x unfolding F_def by auto have u1_bound: "abs(u 1 x) \ K + F x" for x unfolding F_def indicator_def apply (cases "x \ {x \ space M. K \ \u 1 x\}") by auto define F2 where "F2 = (\x. F x + abs(u s x/s))" have int_F2 [measurable]: "integrable M F2" unfolding F2_def using int_F int_u[of s] by auto have F2_pos: "F2 x \ 0" for x unfolding F2_def using F_pos by auto have "(\x. F2 x \M) = (\x. F x \M) + (\x. abs(u s x/s) \M)" unfolding F2_def apply (rule Bochner_Integration.integral_add) using int_F int_u by auto then have F2_int: "(\x. F2 x \M) < 2 * rho" using F_int s_int by auto text \We can now choose $k$, large enough. The reason for our choice will only appear at the end of the proof.\ define k where "k = max (2 * s + 1) (nat(ceiling((2 * d * s + 2 * K * s)/(d/2))))" have "k > 2 * s" unfolding k_def by auto have "k \ (2 * d * s + 2 * K * s)/(d/2)" unfolding k_def by linarith then have "(2 * d * s + 2 * K * s)/k \ d/2" using \k > 2 * s\ by (simp add: divide_simps mult.commute) text \Third step: definition of a good set $G$ where everything goes well.\ define G where "G = {x \ space M. (\n. u n x / n) \ 0 \ (\n. birkhoff_sum (\x. abs(u s x / s)) n x / n) \ real_cond_exp M Invariants (\x. abs(u s x / s)) x \ (\n. birkhoff_sum F n x / n) \ real_cond_exp M Invariants F x \ real_cond_exp M Invariants F x + real_cond_exp M Invariants (\x. abs(u s x / s)) x = real_cond_exp M Invariants F2 x}" have [measurable]: "G \ sets M" unfolding G_def by auto have "AE x in M. (\n. u n x / n) \ 0" using kingman_theorem_nonergodic(1)[OF subu subu_fin] subu_0 by auto moreover have "AE x in M.(\n. birkhoff_sum (\x. abs(u s x / s)) n x / n) \ real_cond_exp M Invariants (\x. abs(u s x / s)) x" apply (rule birkhoff_theorem_AE_nonergodic) using int_u[of s] by auto moreover have "AE x in M. (\n. birkhoff_sum F n x / n) \ real_cond_exp M Invariants F x" by (rule birkhoff_theorem_AE_nonergodic[OF int_F]) moreover have "AE x in M. real_cond_exp M Invariants F x + real_cond_exp M Invariants (\x. abs(u s x / s)) x = real_cond_exp M Invariants F2 x" unfolding F2_def apply (rule AE_symmetric[OF real_cond_exp_add]) using int_u[of s] int_F int_u[of s] by auto ultimately have "AE x in M. x \ G" unfolding G_def by auto then have "emeasure M G = 1" by (simp add: emeasure_eq_1_AE) text \Estimation of asymptotic densities of bad times, for points in $G$. There is a trivial part, named $U$ below, that has to be treated separately because it creates problematic boundary effects.\ define U where "U = (\x. {n. \l \ {n-s<..n}. u n x - u (n-l) x \ - d * l})" define V where "V = (\x. {n. \l \ {k..n-s}. u n x - u (n-l) x \ - d * l})" text \Trivial estimate for $U(x)$: this set is finite for $x\in G$.\ have densU: "upper_asymptotic_density (U x) = 0" if "x \ G" for x proof - define C where "C = Max {abs(u m x) |m. m {n. u n x \ C - d * n}" proof (auto) fix n assume "n \ U x" then obtain l where l: "l\ {n-s <..n}" "u n x - u (n-l) x \ - d * l" unfolding U_def by auto define m where "m = n-l" have "m < s" unfolding m_def using l by auto have "u n x \ u m x - d * l" using l m_def by auto also have "... \ abs(u m x) - d * n + d * m" unfolding m_def using l by (simp add: algebra_simps of_nat_diff) also have "... \ Max {abs(u m x) |m. mm < s\ apply (auto) by (rule Max_ge, auto) also have "... \ Max {abs(u m x) |m. mm < s\ \d>0\ by auto finally show "u n x \ C - d * n" unfolding C_def by auto qed have "eventually (\n. u n x / n > -d/2) sequentially" apply (rule order_tendstoD(1)) using \x \ G\ \d>0\ unfolding G_def by auto then obtain N where N: "\n. n \ N \ u n x / n > - d/2" using eventually_sequentially by auto { fix n assume *: "u n x \ C - d * n" "n > N" then have "n \ N" "n > 0" by auto have "2 * u n x \ 2 * C - 2 * d * n" using * by auto moreover have "2 * u n x \ - d * n" using N[OF \n \ N\] \n > 0\ by (simp add: divide_simps) ultimately have "- d * n \ 2 * C - 2 * d * n" by auto then have "d * n \ 2 * C" by auto then have "n \ 2 * C / d" using \d>0\ by (simp add: mult.commute divide_simps) } then have "{n. u n x \ C - d * n} \ {..max (nat (floor(2*C/d))) N}" by (auto, meson le_max_iff_disj le_nat_floor not_le) then have "finite {n. u n x \ C - d * n}" using finite_subset by blast then have "finite (U x)" using * finite_subset by blast then show ?thesis using upper_asymptotic_density_finite by auto qed text \Main step: control of $u$ along the sequence $ns+t$, with a term $-(d/2) Card(V(x) \cap [1,ns+t])$ on the right. Then, after averaging in $t$, Birkhoff theorem will imply that $Card(V(x) \cap [1,n])$ is suitably small.\ define Z where "Z = (\t n x. Max {u i x|i. i< s} + (\i {1..n * s + t}))" have Main: "u (n * s + t) x \ Z t n x" if "t < s" for n x t proof (rule nat_less_induct[where ?n = n]) fix n assume H: "\m Z t m x" consider "n = 0"|"n>0 \ V x \ {(n-1) * s+t<..n * s+t} = {}"|"n>0 \ V x \ {(n-1) * s+t<..n * s+t} \ {}" by auto then show "u (n * s+t) x \ Z t n x" proof (cases) assume "n = 0" then have "V x \ {1..n * s + t} = {}" unfolding V_def using \t \k>2* s\ by auto then have *: "card(V x \ {1..n * s+t}) = 0" by simp have **: "0 \ (\in = 0\ by auto also have "... \ Max {u i x|i. i< s}" by (rule Max_ge, auto simp add: \t) also have "... \ Z t n x" unfolding Z_def birkhoff_sum_def using \n = 0\ * ** by auto finally show ?thesis by simp next assume A: "n>0 \ V x \ {(n-1) * s+t<..n * s+t} = {}" then have "n\1" by simp have "n * s + t = (n-1) * s + t + s" using \n\1\ by (simp add: add.commute add.left_commute mult_eq_if) have "V x \ {1..n * s + t} = V x \ {1..(n-1) * s + t} \ V x \ {(n-1) * s + t<..n * s + t}" using \n\1\ by (auto, simp add: mult_eq_if) then have *: "card(V x \ {1..n * s+t}) = card(V x \ {1..(n-1) * s+t})" using A by auto have **: "birkhoff_sum F ((n-1) * s + t) x \ birkhoff_sum F (n * s + t) x" unfolding birkhoff_sum_def apply (rule sum_mono2) using \n * s+t = (n-1) * s+t + s\ F_pos by auto have "(\i (\i (\in\1\ lessThan_Suc_atMost sum.lessThan_Suc[of "\i. abs(u s ((T^^(i* s+t))x))" "n-1", symmetric] by auto finally have ***: "(\i (\in\1\ by (simp add: add.commute add.left_commute mult_eq_if) also have "... \ u ((n-1) * s+t) x + u s ((T^^((n-1) * s+t)) x)" using subcocycle_ineq[OF subu, of "(n-1) * s+t" s x] by simp also have "... \ Max {u i x|i. i< s} + (\i {1..(n-1) * s+t}) + u s ((T^^((n-1) * s+t)) x)" using H \n\1\ unfolding Z_def by auto also have "... \ Max {u i x|i. i< s} + (\i {1..n * s+t})" using * ** *** by auto also have "... \ Z t n x" unfolding Z_def by (auto simp add: divide_simps) finally show ?thesis by simp next assume B: "n>0 \ V x \ {(n-1) * s+t<..n * s+t} \ {}" then have [simp]: "n>0" "n\1" "n \ 0" by auto obtain m where m: "m \ V x \ {(n-1) * s + t<..n * s + t}" using B by blast then obtain l where l: "l \ {k..m-s}" "u m x - u (m-l) x \ - d * l" unfolding V_def by auto then have "m-s>0" using \k>2* s\ by auto then have "m-l \ s" using l by auto define p where "p = (m-l-t) div s" have p1: "m-l \ p * s + t" unfolding p_def using \m-l \ s\ \s>t\ minus_mod_eq_div_mult [symmetric, of "m - l - t" s] by simp have p2: "m-l < p* s + t + s" unfolding p_def using \m-l \ s\ \s>t\ div_mult_mod_eq[of "m-l-t" s] mod_less_divisor[OF \s>0\, of "m-l-t"] by linarith then have "l \ m - p * s - t -s" by auto then have "l \ (n-1) * s + t -p * s - t- s" using m by auto then have "l + 2 * s\ (n * s + t) - (p * s+t)" by (simp add: diff_mult_distrib) have "(p+1) * s + t \ (n-1) * s + t" using \k> 2 * s\ m l(1) p1 by (auto simp add: algebra_simps) then have "p+1 \ n-1" using \s>0\ by (meson add_le_cancel_right mult_le_cancel2) then have "p \ n-1" "p (n * s + t)" using m l(1) p1 by (auto simp add: algebra_simps) then have "(1::real) \ ((n * s + t) - (p* s + t)) / k" using \k > 2* s\ by auto have In: "u (n * s + t) x \ u m x + (\i \ {(n-1) * s + t..i \ {(n-1) * s+t.. 0" by (rule sum_nonneg, auto) then show ?thesis using True by auto next case False then have m2: "n * s + t - m >0" "(n-1) * s+t \ m" using m by auto have "birkhoff_sum (u 1) (n * s+t-m) ((T^^m) x) = (\iij \ {m.. (\j \ {m.. (\j \ {(n-1) * s+t..j \ {m..j \ {(n-1) * s+t.. (\j \ {(n-1) * s+t.. u m x + u (n * s+t-m) ((T^^m) x)" using subcocycle_ineq[OF subu, of m "n * s+t-m"] m2 by auto also have "... \ u m x + birkhoff_sum (u 1) (n * s+t-m) ((T^^m) x)" using subcocycle_bounded_by_birkhoff1[OF subu \n * s+t - m >0\, of "(T^^m)x"] by simp finally show ?thesis using * by auto qed have Ip: "u (m-l) x \ u (p* s+t) x + (\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x)))" proof (cases "m-l = p* s+t") case True have "(\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) \ 0" by (rule sum_nonneg, auto) then show ?thesis using True by auto next case False then have "m-l - (p* s+t) > 0" using p1 by auto have I: "p * s + t + (m - l - (p * s + t)) = m - l" using p1 by auto have "birkhoff_sum (u 1) (m-l - (p* s+t)) ((T^^(p* s+t)) x) = (\iij \ {p* s+t.. (\j \ {p* s+t.. (\j \ {p* s+t..j \ {m-l..<(p+1)* s+t}. abs(u 1 ((T^^j) x)))" by auto also have "... = (\j \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^j) x)))" apply (rule sum.atLeastLessThan_concat) using p1 p2 by auto finally have *: "birkhoff_sum (u 1) (m-l - (p* s+t)) ((T^^(p* s+t)) x) \ (\j \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^j) x)))" by auto have "u (m-l) x \ u (p* s+t) x + u (m-l - (p* s+t)) ((T^^(p* s+t)) x)" using subcocycle_ineq[OF subu, of "p* s+t" "m-l - (p* s+t)" x] I by auto also have "... \ u (p* s+t) x + birkhoff_sum (u 1) (m-l - (p* s+t)) ((T^^(p* s+t)) x)" using subcocycle_bounded_by_birkhoff1[OF subu \m-l - (p* s+t) > 0\, of "(T^^(p* s+t)) x"] by simp finally show ?thesis using * by auto qed have "(\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) \ (\i \ {p* s+t..<(p+1)* s+t}. K + F ((T^^i) x))" apply (rule sum_mono) using u1_bound by auto moreover have "(\i \ {(n-1) * s+t.. (\i \ {(n-1) * s+t..i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) + (\i \ {(n-1) * s+t.. (\i \ {p* s+t..<(p+1)* s+t}. K + F ((T^^i) x)) + (\i \ {(n-1) * s+t..i \ {p* s+t..<(p+1)* s+t}. F ((T^^i) x)) + (\i \{(n-1) * s+t.. 2* K * s + (\i \ {p* s+t..<(n-1) * s+t}. F ((T^^i) x)) + (\i \{(n-1) * s+t..(p+1)* s+t\(n-1) * s+t\ F_pos by auto also have "... = 2* K * s + (\i \ {p* s+t..p\n-1\ by auto finally have A0: "(\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) + (\i \ {(n-1) * s+t.. 2* K * s + (\i \ {p* s+t.. {p * s + t<.. n * s+t}) \ card {p * s + t<.. n * s+t}" by (rule card_mono, auto) have "2 * d * s + 2 * K * s > 0" using \K>0\ \s>0\ \d>0\ by (metis add_pos_pos mult_2 mult_zero_left of_nat_0_less_iff pos_divide_less_eq times_divide_eq_right) then have "2 * d * s + 2 * K * s \ ((n * s + t) - (p* s + t)) * ((2 * d * s + 2 * K * s) / k)" using \1 \ ((n * s + t) - (p* s + t)) / k\ by (simp add: le_divide_eq_1 pos_le_divide_eq) also have "... \ ((n * s + t) - (p* s + t)) * (d/2)" apply (rule mult_left_mono) using \(2 * d * s + 2 * K * s)/k \ d/2\ by auto finally have "2 * d * s + 2 * K * s \ ((n * s + t) - (p* s + t)) * (d/2)" by auto then have "-d * ((n * s+t) - (p* s+t)) + 2 * d * s + 2 * K * s \ -d * ((n * s+t) - (p* s+t)) + ((n * s + t) - (p* s + t)) * (d/2)" by linarith also have "... = (-d/2) * card {p * s + t<.. n * s+t}" by auto also have "... \ (-d/2) * card(V x \ {p * s + t<.. n * s+t})" using \card(V x \ {p * s + t<.. n * s+t}) \ card {p * s + t<.. n * s+t}\ by auto finally have A1: "-d * ((n * s+t) - (p* s+t)) + 2 * d * s + 2 * K * s \ (-d/2) * card(V x \ {p * s + t<.. n * s+t})" by simp have "V x \ {1.. n * s+t} = V x \ {1..p * s + t} \ V x \ {p * s + t<.. n * s+t}" using \p * s + t + k \ n * s + t\ by auto then have "card (V x \ {1.. n * s+t}) = card(V x \ {1..p * s + t} \ V x \ {p * s + t<.. n * s+t})" by auto also have "... = card (V x \ {1..p * s + t}) + card (V x \ {p * s + t<.. n * s+t})" by (rule card_Un_disjoint, auto) finally have A2: "card (V x \ {1..p * s + t}) + card (V x \ {p * s + t<.. n * s+t}) = card (V x \ {1.. n * s+t})" by simp have A3: "(\i (\ip\n-1\ by auto have A4: "birkhoff_sum F (p * s + t) x + (\i \ {p* s+t..p\n-1\ by auto have "u (n * s+t) x \ u m x + (\i \ {(n-1) * s+t.. (u m x - u (m-l) x) + u (m-l) x + (\i \ {(n-1) * s+t.. - d * l + u (p* s+t) x + (\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) + (\i \ {(n-1) * s+t.. - d * ((n * s+t) - (p* s+t)) + 2*d* s + u (p* s+t) x + (\i \ {p* s+t..<(p+1)* s+t}. abs(u 1 ((T^^i) x))) + (\i \ {(n-1) * s+t..l + 2* s\ (n * s+t) - (p* s+t)\ apply (auto simp add: algebra_simps) by (metis assms(1) distrib_left mult.commute mult_2 of_nat_add of_nat_le_iff real_mult_le_cancel_iff2) also have "... \ -d * ((n * s+t) - (p* s+t)) + 2*d* s + Z t p x + 2* K * s + (\i \ {p* s+t..p by auto also have "... \ Z t p x - d/2 * card(V x \ {p * s + t<.. n * s+t}) + (\i \ {p* s+t..i {1..p * s + t}) - d/2 * card(V x \ {p * s + t<.. n * s+t}) + (\i \ {p* s+t.. Max {u i x |i. i < s} + (\ii \ {p* s+t.. {1..p * s + t}) - d/2 * card(V x \ {p * s + t<.. n * s+t})" using A3 by auto also have "... = Z t n x" unfolding Z_def using A2 A4 by (auto simp add: algebra_simps, metis distrib_left of_nat_add) finally show ?thesis by simp qed qed have Main2: "(d/2) * card(V x \ {1..n}) \ Max {u i x|i. i< s} + birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x + birkhoff_sum F (n + 2 * s) x + (1/s) * (\i< 2 * s. abs(u (n+i) x))" for n x proof - define N where "N = (n div s) + 1" then have "n \ N * s" using \s > 0\ dividend_less_div_times less_or_eq_imp_le by auto have "N * s \ n + s" by (auto simp add: N_def) have eq_t: "(d/2) * card(V x \ {1..n}) \ abs(u(N* s+t) x) + (Max {u i x|i. i< s} + birkhoff_sum F (n + 2* s) x) + (\i birkhoff_sum F (n+ 2* s) x" unfolding birkhoff_sum_def apply (rule sum_mono2) using F_pos \N * s \ n + s\ \t by auto have "card(V x \ {1..n}) \ card(V x \ {1..N* s+t})" apply (rule card_mono) using \n \ N * s\ by auto then have "(d/2) * card(V x \ {1..n}) \ (d/2) * card(V x \ {1..N* s+t})" by auto also have "... \ - u (N* s+t) x + Max {u i x|i. i< s} + (\it < s\, of N x] unfolding Z_def by auto also have "... \ abs(u(N* s+t) x) + Max {u i x|i. i< s} + birkhoff_sum F (n + 2* s) x + (\iti\{N* s.. (\i\{n..n \ N * s\ \N * s \ n + s\ by auto also have "... = (\i<2* s. abs (u (n+i) x))" by (rule sum.reindex_bij_betw[symmetric], rule bij_betw_byWitness[where ?f' = "\i. i - n"], auto) finally have **: "(\t (\i<2* s. abs (u (n+i) x))" by simp have "(\tii (\iN * s \ n + s\ by auto finally have ***: "(\ti s * birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x" unfolding birkhoff_sum_def using \s>0\ by (auto simp add: sum_divide_distrib[symmetric]) have ****: "s * (\ii {1..n}) = (\t {1..n}))" by auto also have "... \ (\tittti (\i<2* s. abs (u (n+i) x)) + s * (Max {u i x|i. i< s} + birkhoff_sum F (n + 2* s) x) + s * birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x" using ** *** by auto also have "... = s * ((1/s) * (\i<2* s. abs (u (n+i) x)) + Max {u i x|i. i< s} + birkhoff_sum F (n + 2* s) x + birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x)" by (auto simp add: divide_simps mult.commute distrib_left) finally show ?thesis by auto qed have densV: "upper_asymptotic_density (V x) \ (2/d) * real_cond_exp M Invariants F2 x" if "x \ G" for x proof - have *: "(\n. abs(u n x/n)) \ 0" apply (rule tendsto_rabs_zero) using \x\G\ unfolding G_def by auto define Bound where "Bound = (\n. (Max {u i x|i. i< s}*(1/n) + birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x / n + birkhoff_sum F (n + 2* s) x / n + (1/s) * (\i<2* s. abs(u (n+i) x) / n)))" have "Bound \ (Max {u i x|i. i< s} * 0 + real_cond_exp M Invariants (\x. abs(u s x/s)) x + real_cond_exp M Invariants F x + (1/s) * (\i < 2 * s. 0))" unfolding Bound_def apply (intro tendsto_intros) using \x\G\ * unfolding G_def by auto moreover have "real_cond_exp M Invariants (\x. abs(u s x/s)) x + real_cond_exp M Invariants F x = real_cond_exp M Invariants F2 x" using \x \ G\ unfolding G_def by auto ultimately have B_conv: "Bound \ real_cond_exp M Invariants F2 x" by simp have *: "(d/2) * card(V x \ {1..n}) / n \ Bound n" for n proof - have "(d/2) * card(V x \ {1..n}) / n \ (Max {u i x|i. i< s} + birkhoff_sum (\x. abs(u s x/ s)) (n+2* s) x + birkhoff_sum F (n + 2* s) x + (1/s) * (\i<2* s. abs(u (n+i) x)))/n" using Main2[of x n] using divide_right_mono of_nat_0_le_iff by blast also have "... = Bound n" unfolding Bound_def by (auto simp add: add_divide_distrib sum_divide_distrib[symmetric]) finally show ?thesis by simp qed have "ereal(d/2 * upper_asymptotic_density (V x)) = ereal(d/2) * ereal(upper_asymptotic_density (V x))" by auto also have "... = ereal (d/2) * limsup(\n. card(V x \ {1..n}) / n)" using upper_asymptotic_density_shift[of "V x" 1 0] by auto also have "... = limsup(\n. ereal (d/2) * (card(V x \ {1..n}) / n))" by (rule limsup_ereal_mult_left[symmetric], auto) also have "... \ limsup Bound" apply (rule Limsup_mono) using * not_eventuallyD by auto also have "... = ereal(real_cond_exp M Invariants F2 x)" using B_conv convergent_limsup_cl convergent_def convergent_real_imp_convergent_ereal limI by force finally have "d/2 * upper_asymptotic_density (V x) \ real_cond_exp M Invariants F2 x" by auto then show ?thesis by (simp add: divide_simps mult.commute) qed define epsilon where "epsilon = 2 * rho / d" have [simp]: "epsilon > 0" "epsilon \ 0" "epsilon \ 0" unfolding epsilon_def by auto have "emeasure M {x\space M. real_cond_exp M Invariants F2 x \ epsilon} \ (1/epsilon) * (\x. real_cond_exp M Invariants F2 x \M)" apply (intro integral_Markov_inequality real_cond_exp_pos real_cond_exp_int(1)) by (auto simp add: int_F2 F2_pos) also have "... = (1/epsilon) * (\x. F2 x \M)" apply (intro arg_cong[where f = ennreal]) by (simp, rule real_cond_exp_int(2), simp add: int_F2) also have "... < (1/epsilon) * 2 * rho" using F2_int by (intro ennreal_lessI) (auto simp add: divide_simps) also have "... = d" unfolding epsilon_def by auto finally have *: "emeasure M {x\space M. real_cond_exp M Invariants F2 x \ epsilon} < d" by simp define G2 where "G2 = {x \ G. real_cond_exp M Invariants F2 x < epsilon}" have [measurable]: "G2 \ sets M" unfolding G2_def by simp have "1 = emeasure M G" using \emeasure M G = 1\ by simp also have "... \ emeasure M (G2 \ {x\space M. real_cond_exp M Invariants F2 x \ epsilon})" apply (rule emeasure_mono) unfolding G2_def using sets.sets_into_space[OF \G \ sets M\] by auto also have "... \ emeasure M G2 + emeasure M {x\space M. real_cond_exp M Invariants F2 x \ epsilon}" by (rule emeasure_subadditive, auto) also have "... < emeasure M G2 + d" using * by auto finally have "1 - d < emeasure M G2" using emeasure_eq_measure \d \ 1\ by (auto intro!: ennreal_less_iff[THEN iffD2] simp del: ennreal_plus simp add: ennreal_plus[symmetric]) have "upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} < d" if "x \ G2" for x proof - have "x \ G" using \x \ G2\ unfolding G2_def by auto have "{n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} \ U x \ V x" unfolding U_def V_def by fastforce then have "upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} \ upper_asymptotic_density (U x \ V x)" by (rule upper_asymptotic_density_subset) also have "... \ upper_asymptotic_density (U x) + upper_asymptotic_density (V x)" by (rule upper_asymptotic_density_union) also have "... \ (2/d) * real_cond_exp M Invariants F2 x" using densU[OF \x \ G\] densV[OF \x \ G\] by auto also have "... < (2/d) * epsilon" using \x \ G2\ unfolding G2_def by (simp add: divide_simps) text \This is where the choice of $\rho$ at the beginning of the proof is relevant: we choose it so that the above term is at most $d$.\ also have "... = d" unfolding epsilon_def rho_def by auto finally show ?thesis by simp qed then have "G2 \ {x \ space M. upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} < d}" using sets.sets_into_space[OF \G2 \ sets M\] by blast then have "emeasure M G2 \ emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} < d}" by (rule emeasure_mono, auto) then have "emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - d * l} < d} > 1 -d" using \emeasure M G2 > 1 - d\ by auto then show ?thesis by blast qed text \The two previous lemmas are put together in the following lemma, corresponding to Lemma 2.3 in~\cite{gouezel_karlsson}.\ lemma upper_density_delta: fixes d::real assumes "d > 0" "d \ 1" shows "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ emeasure M {x \ space M. \(N::nat). card {n \{..l \ {1..n}. u n x - u (n-l) x \ - delta l * l} \ d * N} > 1 - d" proof - define d2 where "d2 = d/2" have [simp]: "d2 > 0" unfolding d2_def using assms by simp then have "\ d2 < 0" using not_less [of d2 0] by (simp add: less_le) have "d2/2 > 0" by simp obtain c0 where c0: "c0> (0::real)" "emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) x \ - c0 * l} < d2/2} > 1 - (d2/2)" using upper_density_all_times[OF \d2/2 > 0\] by blast have "\N. emeasure M {x \ space M. \n \ N. card ({n. \l \ {1..n}. u n x - u (n-l) x \ - c0 * l} \ {.. 1 - (d2/2)" apply (rule upper_density_eventually_measure) using c0(2) by auto then obtain N1 where N1: "emeasure M {x \ space M. \B \ N1. card ({n. \l \ {1..n}. u n x - u (n-l) x \ - c0 * l} \ {.. 1 - (d2/2)" by blast define O1 where "O1 = {x \ space M. \B \ N1. card ({n. \l \ {1..n}. u n x - u (n-l) x \ - c0 * l} \ {.. sets M" unfolding O1_def by auto have "emeasure M O1 > 1 - (d2/2)" unfolding O1_def using N1 by auto have *: "\N. emeasure M {x \ space M. \B \ N. card({n. \l \ {N..n}. u n x - u (n-l) x \ - e * l} \ {.. 1 - e" if "e>0" "e \ 1" for e::real proof - obtain k where k: "emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {k..n}. u n x - u (n-l) x \ - e * l} < e} > 1 - e" using upper_density_large_k[OF \e>0\ \e \ 1\] by blast then obtain N0 where N0: "emeasure M {x \ space M. \B \ N0. card({n. \l \ {k..n}. u n x - u (n-l) x \ - e * l} \ {.. 1 - e" using upper_density_eventually_measure[OF _ k] by auto define N where "N = max k N0" have "emeasure M {x \ space M. \B \ N0. card({n. \l \ {k..n}. u n x - u (n-l) x \ - e * l} \ {.. emeasure M {x \ space M. \B \ N. card({n. \l \ {N..n}. u n x - u (n-l) x \ - e * l} \ {.. space M" "\B\N0. card ({n. \l\{k..n}. u n x - u (n - l) x \ - (e * real l)} \ {.. B" have "card({n. \l \ {N..n}. u n x - u (n-l) x \ - (e * real l)} \ {.. card({n. \l \ {k..n}. u n x - u (n-l) x \ -(e * real l)} \ {..l \ {N..n}. u n x - u (n-l) x \ - (e * real l)} \ {.. card({n. \l \ {k..n}. u n x - u (n-l) x \ -(e * real l)} \ {..B\N\ unfolding N_def by auto finally show "card ({n. \l\{N..n}. u n x - u (n - l) x \ - (e * real l)} \ {.. space M. \B \ N. card({n. \l \ {N..n}. u n x - u (n-l) x \ - e * l} \ {.. 1 - e" using N0 by simp then show ?thesis by auto qed define Ne where "Ne = (\(e::real). SOME N. emeasure M {x \ space M. \B \ N. card({n. \l \ {N..n}. u n x - u (n-l) x \ - e * l} \ {.. 1 - e)" have Ne: "emeasure M {x \ space M. \B \ Ne e. card({n. \l \ {Ne e..n}. u n x - u (n-l) x \ - e * l} \ {.. 1 - e" if "e>0" "e \ 1" for e::real unfolding Ne_def by (rule someI_ex[OF *[OF that]]) define eps where "eps = (\(n::nat). d2 * (1/2)^n)" have [simp]: "eps n > 0" for n unfolding eps_def by auto then have [simp]: "eps n \ 0" for n by (rule less_imp_le) have "eps n \ (1 / 2) * 1" for n unfolding eps_def d2_def using \d \ 1\ by (intro mult_mono power_le_one) auto also have "\ < 1" by auto finally have [simp]: "eps n < 1" for n by simp then have [simp]: "eps n \ 1" for n by (rule less_imp_le) have "(\n. d2 * (1/2)^n) \ d2 * 0" by (rule tendsto_mult, auto simp add: LIMSEQ_realpow_zero) then have "eps \ 0" unfolding eps_def by auto define Nf where "Nf = (\N. (if (N = 0) then 0 else if (N = 1) then N1 + 1 else max (N1+1) (Max {Ne(eps n)|n. n \ N}) + N))" have "Nf N < Nf (N+1)" for N proof - consider "N = 0" | "N = 1" | "N>1" by fastforce then show ?thesis proof (cases) assume "N>1" have "Max {Ne (eps n) |n. n \ N} \ Max {Ne (eps n) |n. n \ Suc N}" by (rule Max_mono, auto) then show ?thesis unfolding Nf_def by auto qed (auto simp add: Nf_def) qed then have "strict_mono Nf" using strict_mono_Suc_iff by auto define On where "On = (\(N::nat). (if (N = 1) then O1 else {x \ space M. \B \ Nf N. card({n. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N) * l} \ {.. sets M" for N unfolding On_def by auto have "emeasure M (On N) > 1 - eps N" if "N>0" for N proof - consider "N = 1" | "N>1" using \N>0\ by linarith then show ?thesis proof (cases) case 1 then show ?thesis unfolding On_def eps_def using \emeasure M O1 > 1 - (d2/2)\ by auto next case 2 have "Ne (eps N) \ Max {Ne(eps n)|n. n \ N}" by (rule Max.coboundedI, auto) also have "... \ Nf N" unfolding Nf_def using \N>1\ by auto finally have "Ne (eps N) \ Nf N" by simp have "1 - eps N < emeasure M {x \ space M. \B \ Ne(eps N). card({n. \l \ {Ne(eps N)..n}. u n x - u (n-l) x \ - (eps N) * l} \ {.. emeasure M {x \ space M. \B \ Nf N. card({n. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N) * l} \ {.. space M" "\n\Ne (eps N). card ({n. \l\{Ne (eps N)..n}. u n x - u (n - l) x \ - (eps N * l)} \ {.. n" have "card({n. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. card({n. \l \ {Ne(eps N)..n}. u n x - u (n-l) x \ -(eps N) * l} \ {..Ne (eps N) \ Nf N\ by auto then have "real(card({n. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. card({n. \l \ {Ne(eps N)..n}. u n x - u (n-l) x \ -(eps N) * l} \ {..n \ Nf N\ \Ne (eps N) \ Nf N\ by auto finally show "real (card ({n. \l\{Nf N..n}. u n x - u (n - l) x \ - (eps N * l)} \ {..N>1\ by auto finally show ?thesis by simp qed qed then have *: "emeasure M (On (N+1)) > 1 - eps (N+1)" for N by simp define Ogood where "Ogood = (\N. On (N+1))" have [measurable]: "Ogood \ sets M" unfolding Ogood_def by auto have "emeasure M Ogood \ 1 - (\N. eps(N+1))" unfolding Ogood_def apply (intro emeasure_intersection, auto) using * by (auto simp add: eps_def summable_mult summable_divide summable_geometric less_imp_le) moreover have "(\N. eps(N+1)) = d2" unfolding eps_def apply (subst suminf_mult) using sums_unique[OF power_half_series, symmetric] by (auto intro!: summable_divide summable_geometric) finally have "emeasure M Ogood \ 1 - d2" by simp then have "emeasure M Ogood > 1 - d" unfolding d2_def using \d>0\ \d \ 1\ by (simp add: emeasure_eq_measure field_sum_of_halves ennreal_less_iff) have Ogood_union: "Ogood = (\(K::nat). Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)})" apply auto using sets.sets_into_space[OF \Ogood \ sets M\] apply blast proof - fix x define M where "M = Max {abs(u n x - u (n-l) x)/l | n l. n \ {1..Nf 1} \ l \ {1..n}}" obtain N::nat where "N > M" using reals_Archimedean2 by blast have "finite { (n, l) | n l. n \ {1..Nf 1} \ l \ {1..n}}" by (rule finite_subset[where ?B = "{1.. Nf 1} \ {1..Nf 1}"], auto) moreover have "{abs(u n x - u (n-l) x)/l | n l. n \ {1..Nf 1} \ l \ {1..n}} = (\ (n, l). abs(u n x - u (n-l) x)/l)` { (n, l) | n l. n \ {1..Nf 1} \ l \ {1..n}}" by auto ultimately have fin: "finite {abs(u n x - u (n-l) x)/l | n l. n \ {1..Nf 1} \ l \ {1..n}}" by auto { fix n l assume nl: "n \ {1..Nf 1} \ l \ {1..n}" then have "real l>0" by simp have "abs(u n x - u (n-l) x)/l \ M" unfolding M_def apply (rule Max_ge) using fin nl by auto then have "abs(u n x - u (n-l) x)/l < real N" using \N>M\ by simp then have "abs(u n x - u (n-l) x)< real N * l" using \0 < real l\ pos_divide_less_eq by blast then have "u n x - u (n-l) x > - (real N * l)" by simp } then have "\n\{Suc 0..Nf (Suc 0)}. \l\{Suc 0..n}. - (real N * real l) < u n x - u (n - l) x" by auto then show "\N. \n\{Suc 0..Nf (Suc 0)}. \l\{Suc 0..n}. - (real N * real l) < u n x - u (n - l) x" by auto qed have "(\K. emeasure M (Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)})) \ emeasure M (\(K::nat). Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)})" apply (rule Lim_emeasure_incseq, auto) unfolding incseq_def apply auto proof - fix m n x na l assume "m \ (n::nat)" "\n\{Suc 0..Nf (Suc 0)}. \l\{Suc 0..n}. - (real m * real l) < u n x - u (n - l) x" "Suc 0 \ l" "l \ na" "na \ Nf (Suc 0)" then have "- (real m * real l) < u na x - u (na - l) x" by auto moreover have "- (real n * real l) \ - (real m * real l)" using \m \ n\ by (simp add: mult_mono) ultimately show "- (real n * real l) < u na x - u (na - l) x" by auto qed moreover have "emeasure M (\(K::nat). Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)}) > 1 - d" using Ogood_union \emeasure M Ogood > 1 - d\ by auto ultimately have a: "eventually (\K. emeasure M (Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)}) > 1 - d) sequentially" by (rule order_tendstoD(1)) have b: "eventually (\K. K \ max c0 d2) sequentially" using eventually_at_top_linorder nat_ceiling_le_eq by blast have "eventually (\K. K \ max c0 d2 \ emeasure M (Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)}) > 1 - d) sequentially" by (rule eventually_elim2[OF a b], auto) then obtain K where K: "K\max c0 d2" "emeasure M (Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)}) > 1 - d" using eventually_False_sequentially eventually_elim2 by blast define Og where "Og = Ogood \ {x \ space M. \n \ {1..Nf 1}. \l \ {1..n}. u n x - u (n-l) x > - (real K * l)}" have [measurable]: "Og \ sets M" unfolding Og_def by simp have "emeasure M Og > 1 - d" unfolding Og_def using K by simp have fin: "finite {N. Nf N \ n}" for n using pseudo_inverse_finite_set[OF filterlim_subseq[OF \strict_mono Nf\]] by auto define prev_N where "prev_N = (\n. Max {N. Nf N \ n})" define delta where "delta = (\l. if (prev_N l \ 1) then K else eps (prev_N l))" have "\l. delta l > 0" unfolding delta_def using \K\max c0 d2\ \c0>0\ by auto have "LIM n sequentially. prev_N n :> at_top" unfolding prev_N_def apply (rule tendsto_at_top_pseudo_inverse2) using \strict_mono Nf\ by (simp add: filterlim_subseq) then have "eventually (\l. prev_N l > 1) sequentially" by (simp add: filterlim_iff) then have "eventually (\l. delta l = eps(prev_N l)) sequentially" unfolding delta_def by (simp add: eventually_mono) moreover have "(\l. eps(prev_N l)) \ 0" by (rule filterlim_compose[OF \eps \ 0\ \LIM n sequentially. prev_N n :> at_top\]) ultimately have "delta \ 0" by (simp add: filterlim_cong) have "delta n \ K" for n proof - have *: "d2 * (1 / 2) ^ prev_N n \ real K * 1" apply (rule mult_mono') using \K \ max c0 d2\ \d2>0\ by (auto simp add: power_le_one less_imp_le) then show ?thesis unfolding delta_def apply auto unfolding eps_def using * by auto qed define bad_times where "bad_times = (\x. {n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ (\N\{2..}. {n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)}))" have card_bad_times: "card (bad_times x \ {.. d2 * B" if "x \ Og" for x B proof - have "(\N\{.. (\N. (1/2)^N)" by (rule sum_le_suminf, auto simp add: summable_geometric) also have "... = 2" using suminf_geometric[of "1/(2::real)"] by auto finally have *: "(\N\{.. 2" by simp have "(\ N \ {2.. (\ N \ {2..N\{2..N\{..N\{..N\{.. (d2 * (1/4) * B) * 2" apply (rule mult_left_mono) using * \d2 > 0\ apply auto by (metis \0 < d2\ mult_eq_0_iff mult_le_0_iff not_le of_nat_eq_0_iff of_nat_le_0_iff) finally have I: "(\ N \ {2.. d2/2 * B" by auto have "x \ On 1" using \x \ Og\ unfolding Og_def Ogood_def by auto then have "x \ O1" unfolding On_def by auto have B1: "real(card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. (d2/2) * B" for B proof (cases "B \ N1") case True have "card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. card({n. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. (d2/2) * B" using \x \ O1\ unfolding O1_def using True by auto finally show ?thesis by auto next case False then have "B < Nf 1" unfolding Nf_def by auto then have "{n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. (d2/2) * B" using \\ d2 < 0\ by simp finally show ?thesis by simp qed have BN: "real(card ({n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. eps N * B" if "N \ 2" for N B proof - have "x \ On ((N-1) + 1)" using \x \ Og\ unfolding Og_def Ogood_def by auto then have "x \ On N" using \N\2\ by auto show ?thesis proof (cases "B \ Nf N") case True have "card ({n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. card ({n. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. eps N * B" using \x \ On N\ \N\2\ True unfolding On_def by auto finally show ?thesis by simp next case False then have "{n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. eps N * B" by (metis \\n. 0 < eps n\ le_less mult_eq_0_iff mult_pos_pos of_nat_0 of_nat_0_le_iff) finally show ?thesis by simp qed qed { fix N assume "N \ B" have "Nf N \ B" using seq_suble[OF \strict_mono Nf\, of N] \N \ B\ by simp then have "{Nf N..} \ {.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {..N\{B..}. {n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {2.. {B..}" by auto then have "(\N\{2..}. {n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. (\N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. (\N\{B..}. {n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {..N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {.. {n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. (\N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {.. card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. (\N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {..N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {.. real(card({n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)} \ {.. N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. (d2/2 * B) + (\ N\{2.. {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)} \ {.. (d2/2 * B) + (\ N \ {2.. (d2/2 * B) + (d2/2*B)" using I by auto finally show ?thesis by simp qed have ineq_on_Og: "u n x - u (n-l) x > - delta l * l" if "l \ {1..n}" "n \ bad_times x" "x \ Og" for n x l proof - consider "n < Nf 1" | "n \ Nf 1 \ prev_N l \ 1" | "n \ Nf 1 \ prev_N l \ 2" by linarith then show ?thesis proof (cases) assume "n < Nf 1" then have "{N. Nf N \ n} = {0}" apply auto using \strict_mono Nf\ unfolding strict_mono_def apply (metis le_trans less_Suc0 less_imp_le_nat linorder_neqE_nat not_less) unfolding Nf_def by auto then have "prev_N n = 0" unfolding prev_N_def by auto moreover have "prev_N l \ prev_N n" unfolding prev_N_def apply (rule Max_mono) using \l \ {1..n}\ fin apply auto unfolding Nf_def by auto ultimately have "prev_N l = 0" using \prev_N l \ prev_N n\ by auto then have "delta l = K" unfolding delta_def by auto have "1 \ {N. Nf N \ n}" using fin[of n] by (metis (full_types) Max_ge \prev_N n = 0\ fin not_one_le_zero prev_N_def) then have "n < Nf 1" by auto moreover have "n \ 1" using \l \ {1..n}\ by auto ultimately have "n \ {1..Nf 1}" by auto then have "u n x - u (n-l) x > - (real K * l)" using \x \ Og\ unfolding Og_def using \l \ {1..n}\ by auto then show ?thesis using \delta l = K\ by auto next assume H: "n \ Nf 1 \ prev_N l \ 1" then have "delta l = K" unfolding delta_def by auto have "n \ {n \ {Nf 1..}. \l\{1..n}. u n x - u (n-l) x \ - (c0 * l)}" using \n \ bad_times x\ unfolding bad_times_def by auto then have "u n x - u (n-l) x > - (c0 * l)" using H \l \ {1..n}\ by force moreover have "- (c0 * l) \ - (real K * l)" using K(1) by (simp add: mult_mono) ultimately show ?thesis using \delta l = K\ by auto next assume H: "n \ Nf 1 \ prev_N l \ 2" define N where "N = prev_N l" have "N \ 2" unfolding N_def using H by auto have "prev_N l \ {N. Nf N \ l}" unfolding prev_N_def apply (rule Max_in, auto simp add: fin) unfolding Nf_def by auto then have "Nf N \ l" unfolding N_def by auto then have "Nf N \ n" using \l \ {1..n}\ by auto have "n \ {n \ {Nf N..}. \l \ {Nf N..n}. u n x - u (n-l) x \ - (eps N * l)}" using \n \ bad_times x\ \N\2\ unfolding bad_times_def by auto then have "u n x - u (n-l) x > - (eps N * l)" using \Nf N \ n\ \Nf N \ l\ \l \ {1..n}\ by force moreover have "eps N = delta l" unfolding delta_def N_def using H by auto ultimately show ?thesis by auto qed qed have "Og \ {x \ space M. \(B::nat). card {n \{..l \ {1..n}. u n x - u (n-l) x \ - delta l * l} \ d * B}" proof (auto) fix x assume "x \ Og" then show "x \ space M" unfolding Og_def by auto next fix x B assume "x \ Og" have *: "{n. n < B \ (\l\{Suc 0..n}. u n x - u (n - l) x \ - (delta l * real l))} \ bad_times x \ {..x\Og\ by force have "card {n. n < B \ (\l\{Suc 0..n}. u n x - u (n - l) x \ - (delta l * real l))} \ card (bad_times x \ {.. d2 * B" using card_bad_times \x \ Og\ by auto also have "... \ d * B" unfolding d2_def using \d>0\ by auto finally show "card {n. n < B \ (\l\{Suc 0..n}. u n x - u (n - l) x \ - (delta l * real l))} \ d * B" by simp qed then have "emeasure M Og \ emeasure M {x \ space M. \(B::nat). card {n \{..l \ {1..n}. u n x - u (n-l) x \ - delta l * l} \ d * B}" by (rule emeasure_mono, auto) then have "emeasure M {x \ space M. \(B::nat). card {n \{..l \ {1..n}. u n x - u (n-l) x \ - delta l * l} \ d * B} > 1-d" using \emeasure M Og > 1 - d\ by auto then show ?thesis using \delta \ 0\ \\l. delta l > 0\ by auto qed text \We go back to the natural time direction, by using the previous result for the inverse map and the inverse subcocycle, and a change of variables argument. The price to pay is that the estimates we get are weaker: we have a control on a set of upper asymptotic density close to $1$, while having a set of lower asymptotic density close to $1$ as before would be stronger. This will nevertheless be sufficient for our purposes below.\ lemma upper_density_good_direction_invertible: assumes "invertible_qmpt" "d>(0::real)" "d \ 1" shows "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} \ 1-d} \ ennreal(1-d)" proof - interpret I: Gouezel_Karlsson_Kingman M Tinv "(\n x. u n ((Tinv^^n) x))" proof show "Tinv \ quasi_measure_preserving M M" using Tinv_qmpt[OF \invertible_qmpt\] unfolding qmpt_def qmpt_axioms_def by simp show "Tinv \ measure_preserving M M" using Tinv_mpt[OF \invertible_qmpt\] unfolding mpt_def mpt_axioms_def by simp show "mpt.subcocycle M Tinv (\n x. u n ((Tinv ^^ n) x))" using subcocycle_u_Tinv[OF subu \invertible_qmpt\] by simp show "- \ < subcocycle_avg_ereal (\n x. u n ((Tinv ^^ n) x))" using subcocycle_avg_ereal_Tinv[OF subu \invertible_qmpt\] subu_fin by simp show "AE x in M. fmpt.subcocycle_lim M Tinv (\n x. u n ((Tinv ^^ n) x)) x = 0" using subcocycle_lim_Tinv[OF subu \invertible_qmpt\] subu_0 by auto qed have bij: "bij T" using \invertible_qmpt\ unfolding invertible_qmpt_def by simp define e where "e = d * d / 2" have "e>0" "e\1" unfolding e_def using \d>0\ \d \ 1\ by (auto, meson less_imp_le mult_left_le one_le_numeral order_trans) obtain delta::"nat \ real" where d: "\l. delta l > 0" "delta \ 0" "emeasure M {x \ space M. \N. card {n \ {..l\{1..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) \ - delta l * real l} \ e * real N} > 1-e" using I.upper_density_delta[OF \e>0\ \e\1\] by blast define S where "S = {x \ space M. \N. card {n \ {..l\{1..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) \ - delta l * real l} \ e * real N}" have [measurable]: "S \ sets M" unfolding S_def by auto have "emeasure M S > 1 - e" unfolding S_def using d(3) by simp define Og where "Og = (\n. {x \ space M. \l\{1..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) > - delta l * real l})" have [measurable]: "Og n \ sets M" for n unfolding Og_def by auto define Pg where "Pg = (\n. {x \ space M. \l\{1..n}. u n x - u (n - l) ((T^^l) x) > - delta l * real l})" have [measurable]: "Pg n \ sets M" for n unfolding Pg_def by auto define Bad where "Bad = (\i::nat. {x \ space M. \N\i. card {n \ {.. Pg n} \ (1-d) * real N})" have [measurable]: "Bad i \ sets M" for i unfolding Bad_def by auto then have "range Bad \ sets M" by auto have "incseq Bad" unfolding Bad_def incseq_def by auto have inc: "{x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d} \ (\i. Bad i)" proof fix x assume H: "x \ {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d}" then have "x \ space M" by simp define A where "A = {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l}" have "upper_asymptotic_density A < 1-d" using H unfolding A_def by simp then have "\i. \N\i. card (A \ {.. (1-d)* real N" using upper_asymptotic_densityD[of A "1-d"] by (metis (no_types, lifting) eventually_sequentially less_imp_le) then obtain i where "card (A \ {.. (1-d)* real N" if "N\i" for N by blast moreover have "A \ {.. (\l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l)}" for N unfolding A_def by auto ultimately have "x \ Bad i" unfolding Bad_def Pg_def using \x \ space M\ by auto then show "x \ (\i. Bad i)" by blast qed have "emeasure M (Og n) \ emeasure M (Pg n)" for n proof - have *: "(T^^n)-`(Og n) \ space M \ Pg n" for n proof fix x assume x: "x \ (T^^n)-`(Og n) \ space M" define y where "y = (T^^n) x" then have "y \ Og n" using x by auto have "y \ space M" using sets.sets_into_space[OF \Og n \ sets M\] \y \ Og n\ by auto have "x = (Tinv^^n) y" unfolding y_def Tinv_def using inv_fn_o_fn_is_id[OF bij, of n] by (metis comp_apply) { fix l assume "l \ {1..n}" have "(T^^l) x = (T^^l) ((Tinv^^l) ((Tinv^^(n-l))y))" apply (subst \x = (Tinv^^n) y\) using funpow_add[of l "n-l" Tinv] \l \ {1..n}\ by fastforce then have *: "(T^^l) x = (Tinv^^(n-l)) y" unfolding Tinv_def using fn_o_inv_fn_is_id[OF bij] by (metis comp_apply) then have "u n x - u (n-l) ((T^^l) x) = u n ((Tinv^^n) y) - u (n-l) ((Tinv^^(n-l)) y)" using \x = (Tinv^^n) y\ by auto also have "... > - delta l * real l" using \y \ Og n\ \l \ {1..n}\ unfolding Og_def by auto finally have "u n x - u (n-l) ((T^^l) x) > - delta l * real l" by simp } then show "x \ Pg n" unfolding Pg_def using x by auto qed have "emeasure M (Og n) = emeasure M ((T^^n)-`(Og n) \ space M)" using T_vrestr_same_emeasure(2) unfolding vimage_restr_def by auto also have "... \ emeasure M (Pg n)" apply (rule emeasure_mono) using * by auto finally show ?thesis by simp qed { fix N::nat assume "N \ 1" have I: "card {n\{.. Og n} \ (1-e) * real N" if "x \ S" for x proof - have "x \ space M" using \x \ S\ sets.sets_into_space[OF \S \ sets M\] by auto have a: "real (card {n. n < N \ (\l\{Suc 0..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) \ - (delta l * real l))}) \ e * real N" using \x \ S\ unfolding S_def by auto have *: "{n. n < N} = {n. n < N \ (\l\{Suc 0..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) \ - (delta l * real l))} \ {n. n < N \ x \ Og n}" unfolding Og_def using \x \ space M\ by (auto, meson atLeastAtMost_iff linorder_not_le) have "N = card {n. n < N}" by auto also have "... = card {n. n < N \ (\l\{Suc 0..n}. u n ((Tinv ^^ n) x) - u (n - l) ((Tinv ^^ (n - l)) x) \ - (delta l * real l))} + card {n. n < N \ x \ Og n}" apply (subst *, rule card_Un_disjoint) unfolding Og_def by auto ultimately have "real N \ e * real N + card {n. n < N \ x \ Og n}" using a by auto then show ?thesis by (auto simp add: algebra_simps) qed define m where "m = measure M (Bad N)" have "m \ 0" "1-m \ 0" unfolding m_def by auto have *: "1-e \ emeasure M S" using \emeasure M S > 1 - e\ by auto have "ennreal((1-e) * ((1-e) * real N)) = ennreal(1-e) * ennreal((1-e) * real N)" apply (rule ennreal_mult) using \e \ 1\ by auto also have "... \ emeasure M S * ennreal((1-e) * real N)" using mult_right_mono[OF *] by simp also have "... = (\\<^sup>+ x\S. ((1-e) * real N) \M)" by (metis \S \ events\ mult.commute nn_integral_cmult_indicator) also have "... \ (\\<^sup>+x \ S. ennreal(card {n\{.. Og n}) \M)" apply (rule nn_integral_mono) using I unfolding indicator_def by (simp) also have "... \ (\\<^sup>+x \ space M. ennreal(card {n\{.. Og n}) \M)" by (rule nn_set_integral_set_mono, simp only: sets.sets_into_space[OF \S \ sets M\]) also have "... = (\\<^sup>+x. ennreal(card {n\{.. Og n}) \M)" by (rule nn_set_integral_space) also have "... = (\\<^sup>+x. ennreal (\n\{..M)" apply (rule nn_integral_cong) using sum_indicator_eq_card2[of "{..\<^sup>+x. (\n\{..M)" apply (rule nn_integral_cong, auto, simp only: sum_ennreal[symmetric]) by (metis ennreal_0 ennreal_eq_1 indicator_eq_1_iff indicator_simps(2) real_of_nat_indicator) also have "... = (\n \{..\<^sup>+x. (indicator (Og n) x) \M))" by (rule nn_integral_sum, simp) also have "... = (\n \{.. (\n \{..\n. emeasure M (Og n) \ emeasure M (Pg n)\ by simp also have "... = (\n \{..\<^sup>+x. (indicator (Pg n) x) \M))" by simp also have "... = (\\<^sup>+x. (\n\{..M)" by (rule nn_integral_sum[symmetric], simp) also have "... = (\\<^sup>+x. ennreal (\n\{..M)" apply (rule nn_integral_cong, auto, simp only: sum_ennreal[symmetric]) by (metis ennreal_0 ennreal_eq_1 indicator_eq_1_iff indicator_simps(2) real_of_nat_indicator) also have "... = (\\<^sup>+x. ennreal(card {n\{.. Pg n}) \M)" apply (rule nn_integral_cong) using sum_indicator_eq_card2[of "{..\<^sup>+x \ space M. ennreal(card {n\{.. Pg n}) \M)" by (rule nn_set_integral_space[symmetric]) also have "... = (\\<^sup>+x \ Bad N \ (space M - Bad N). ennreal(card {n\{.. Pg n}) \M)" apply (rule nn_integral_cong) unfolding indicator_def by auto also have "... = (\\<^sup>+x \ Bad N. ennreal(card {n\{.. Pg n}) \M) + (\\<^sup>+x \ space M - Bad N. ennreal(card {n\{.. Pg n}) \M)" by (rule nn_integral_disjoint_pair, auto) also have "... \ (\\<^sup>+x \ Bad N. ennreal((1-d) * real N) \M) + (\\<^sup>+x \ space M - Bad N. ennreal(real N) \M)" apply (rule add_mono) apply (rule nn_integral_mono, simp add: Bad_def indicator_def, auto) apply (rule nn_integral_mono, simp add: indicator_def, auto) using card_Collect_less_nat[of N] card_mono[of "{n. n < N}"] by (simp add: Collect_mono_iff) also have "... = ennreal((1-d) * real N) * emeasure M (Bad N) + ennreal(real N) * emeasure M (space M - Bad N)" by (simp add: nn_integral_cmult_indicator) also have "... = ennreal((1-d) * real N) * ennreal(m) + ennreal(real N) * ennreal(1-m)" unfolding m_def by (simp add: emeasure_eq_measure prob_compl) also have "... = ennreal((1-d) * real N * m + real N * (1-m))" using \m \ 0\ \1-m \ 0\ \d \ 1\ ennreal_plus ennreal_mult by auto finally have "ennreal((1-e) * ((1-e) * real N)) \ ennreal((1-d) * real N * m + real N * (1-m))" by simp moreover have "(1-d) * real N * m + real N * (1-m) \ 0" using \m \ 0\ \1-m \ 0\ \d \ 1\ by auto ultimately have "(1-e) * ((1-e) * real N) \ (1-d) * real N * m + real N * (1-m)" using ennreal_le_iff by auto then have "0 \ (e * 2 - d * m - e * e) * real N" by (auto simp add: algebra_simps) then have "0 \ e * 2 - d * m - e * e" using \N \ 1\ by (simp add: zero_le_mult_iff) also have "... \ e * 2 - d * m" using \e > 0\ by auto finally have "m \ e * 2 / d" using \d>0\ by (auto simp add: algebra_simps divide_simps) then have "m \ d" unfolding e_def using \d>0\ by (auto simp add: divide_simps) then have "emeasure M (Bad N) \ d" unfolding m_def by (simp add: emeasure_eq_measure ennreal_leI) } then have "emeasure M (\i. Bad i) \ d" using LIMSEQ_le_const2[OF Lim_emeasure_incseq[OF \range Bad \ sets M\ \incseq Bad\]] by auto then have "emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d} \ d" using emeasure_mono[OF inc, of M] by auto then have *: "measure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d} \ d" using emeasure_eq_measure \d>0\ by fastforce have "{x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} \ 1-d} = space M - {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d}" by auto then have "measure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} \ 1-d} = measure M (space M - {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d})" by simp also have "... = measure M (space M) - measure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} < 1-d}" by (rule measure_Diff, auto) also have "... \ 1 - d" using * prob_space by linarith finally have "emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} \ 1-d} \ 1 - d" using emeasure_eq_measure by auto then show ?thesis using d(1) d(2) by blast qed text \Now, we want to remove the invertibility assumption in the previous lemma. The idea is to go to the natural extension of the system, use the result there and project it back. However, if the system is not defined on a polish space, there is no reason why it should have a natural extension, so we have first to project the original system on a polish space on which the subcocycle is defined. This system is obtained by considering the joint distribution of the subcocycle and all its iterates (this is indeed a polish system, as a space of functions from $\mathbb{N}^2$ to $\mathbb{R}$).\ lemma upper_density_good_direction: assumes "d>(0::real)" "d \ 1" shows "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l} \ 1-d} \ ennreal(1-d)" proof - define U where "U = (\x. (\n. u n x))" define projJ where "projJ = (\x. (\n. U ((T^^n)x)))" define MJ where "MJ = (distr M borel (\x. (\n. U ((T^^n)x))))" define TJ::"(nat \ nat \ real) \ (nat \ nat \ real)" where "TJ = nat_left_shift" have *: "mpt_factor projJ MJ TJ" unfolding projJ_def MJ_def TJ_def apply (rule fmpt_factor_projection) unfolding U_def by (rule measurable_coordinatewise_then_product, simp) interpret J: polish_pmpt MJ TJ unfolding projJ_def polish_pmpt_def apply (auto) apply (rule pmpt_factor) using * apply simp unfolding polish_pmpt_axioms_def MJ_def by auto have [simp]: "projJ \ measure_preserving M MJ" using mpt_factorE(1)[OF *] by simp then have [measurable]: "projJ \ measurable M MJ" by (simp add: measure_preservingE(1)) text \We define a subcocycle $uJ$ in the projection corresponding to the original subcocycle $u$ above. (With the natural definition, it is only a subcocycle almost everywhere.) We check that it shares most properties of $u$.\ define uJ::"nat \ (nat \ nat \ real) \ real" where "uJ = (\n x. x 0 n)" have [measurable]: "uJ n \ borel_measurable borel" for n unfolding uJ_def by (metis measurable_product_coordinates measurable_product_then_coordinatewise) moreover have "measurable borel borel = measurable MJ borel" apply (rule measurable_cong_sets) unfolding MJ_def by auto ultimately have [measurable]: "uJ n \ borel_measurable MJ" for n by fast have uJ_proj: "u n x = uJ n (projJ x)" for n x unfolding uJ_def projJ_def U_def by auto have uJ_int: "integrable MJ (uJ n)" for n apply (rule measure_preserving_preserves_integral'(1)[OF \projJ \ measure_preserving M MJ\]) apply (subst uJ_proj[of n, symmetric]) using int_u[of n] by auto have uJ_int2: "(\x. uJ n x \MJ) = (\x. u n x \M)" for n unfolding uJ_proj apply (rule measure_preserving_preserves_integral'(2)[OF \projJ \ measure_preserving M MJ\]) apply (subst uJ_proj[of n, symmetric]) using int_u[of n] by auto have uJ_AE: "AE x in MJ. uJ (n+m) x \ uJ n x + uJ m ((TJ^^n) x)" for m n proof - have "AE x in M. uJ (n+m) (projJ x) \ uJ n (projJ x) + uJ m (projJ ((T^^n) x))" unfolding uJ_proj[symmetric] using subcocycle_ineq[OF subu] by auto moreover have "AE x in M. projJ ((T^^n) x) = (TJ^^n) (projJ x)" using qmpt_factor_iterates[OF mpt_factor_is_qmpt_factor[OF *]] by auto ultimately have *: "AE x in M. uJ (n+m) (projJ x) \ uJ n (projJ x) + uJ m ((TJ^^n) (projJ x))" by auto show ?thesis apply (rule quasi_measure_preserving_AE'[OF measure_preserving_is_quasi_measure_preserving[OF \projJ \ measure_preserving M MJ\], OF *]) by auto qed have uJ_0: "AE x in MJ. (\n. uJ n x / n) \ 0" proof - have "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by (rule kingman_theorem_nonergodic(1)[OF subu subu_fin]) moreover have "AE x in M. subcocycle_lim u x = 0" using subu_0 by simp ultimately have *: "AE x in M. (\n. uJ n (projJ x) / n) \ 0" unfolding uJ_proj by auto show ?thesis apply (rule quasi_measure_preserving_AE'[OF measure_preserving_is_quasi_measure_preserving[OF \projJ \ measure_preserving M MJ\], OF *]) by auto qed text \Then, we go to the natural extension of $TJ$, to have an invertible system.\ define MI where "MI = J.natural_extension_measure" define TI where "TI = J.natural_extension_map" define projI where "projI = J.natural_extension_proj" interpret I: pmpt MI TI unfolding MI_def TI_def by (rule J.natural_extension(1)) have "I.mpt_factor projI MJ TJ" unfolding projI_def using I.mpt_factorE(1) J.natural_extension(3) MI_def TI_def by auto then have [simp]: "projI \ measure_preserving MI MJ" using I.mpt_factorE(1) by auto then have [measurable]: "projI \ measurable MI MJ" by (simp add: measure_preservingE(1)) have "I.invertible_qmpt" using J.natural_extension(2) MI_def TI_def by auto text \We define a natural subcocycle $uI$ there, and check its properties.\ define uI where uI_proj: "uI = (\n x. uJ n (projI x))" have [measurable]: "uI n \ borel_measurable MI" for n unfolding uI_proj by auto have uI_int: "integrable MI (uI n)" for n unfolding uI_proj by (rule measure_preserving_preserves_integral(1)[OF \projI \ measure_preserving MI MJ\ uJ_int]) have "(\x. uJ n x \MJ) = (\x. uI n x \MI)" for n unfolding uI_proj by (rule measure_preserving_preserves_integral(2)[OF \projI \ measure_preserving MI MJ\ uJ_int]) then have uI_int2: "(\x. uI n x \MI) = (\x. u n x \M)" for n using uJ_int2 by simp have uI_AE: "AE x in MI. uI (n+m) x \ uI n x + uI m (((TI)^^n) x)" for m n proof - have "AE x in MI. uJ (n+m) (projI x) \ uJ n (projI x) + uJ m (((TJ)^^n) (projI x))" apply (rule quasi_measure_preserving_AE[OF measure_preserving_is_quasi_measure_preserving[OF \projI \ measure_preserving MI MJ\]]) using uJ_AE by auto moreover have "AE x in MI. ((TJ)^^n) (projI x) = projI (((TI)^^n) x)" using I.qmpt_factor_iterates[OF I.mpt_factor_is_qmpt_factor[OF \I.mpt_factor projI MJ TJ\]] by auto ultimately show ?thesis unfolding uI_proj by auto qed have uI_0: "AE x in MI. (\n. uI n x / n) \ 0" unfolding uI_proj apply (rule quasi_measure_preserving_AE[OF measure_preserving_is_quasi_measure_preserving[OF \projI \ measure_preserving MI MJ\]]) using uJ_0 by simp text \As $uI$ is only a subcocycle almost everywhere, we correct it to get a genuine subcocycle, to which we will apply Lemma \verb+upper_density_good_direction_invertible+.\ obtain vI where H: "I.subcocycle vI" "AE x in MI. \n. vI n x = uI n x" using I.subcocycle_AE[OF uI_AE uI_int] by blast have [measurable]: "\n. vI n \ borel_measurable MI" "\n. integrable MI (vI n)" using I.subcocycle_integrable[OF H(1)] by auto have "(\x. vI n x \MI) = (\x. uI n x \MI)" for n apply (rule integral_cong_AE) using H(2) by auto then have "(\x. vI n x \MI) = (\x. u n x \M)" for n using uI_int2 by simp then have "I.subcocycle_avg_ereal vI = subcocycle_avg_ereal u" unfolding I.subcocycle_avg_ereal_def subcocycle_avg_ereal_def by auto then have vI_fin: "I.subcocycle_avg_ereal vI > -\" using subu_fin by simp have "AE x in MI. (\n. vI n x / n) \ 0" using uI_0 H(2) by auto moreover have "AE x in MI. (\n. vI n x / n) \ I.subcocycle_lim vI x" by (rule I.kingman_theorem_nonergodic(1)[OF H(1) vI_fin]) ultimately have vI_0: "AE x in MI. I.subcocycle_lim vI x = 0" using LIMSEQ_unique by auto interpret GKK: Gouezel_Karlsson_Kingman MI TI vI apply standard using H(1) vI_fin vI_0 by auto obtain delta where delta: "\l. delta l > 0" "delta \ 0" "emeasure MI {x \ space MI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d } \ 1 - d" using GKK.upper_density_good_direction_invertible[OF \I.invertible_qmpt\ \d>0\ \d\1\] by blast text \Then, we need to go back to the original system, showing that the estimates for $TI$ carry over. First, we go to $TJ$.\ have BJ: "emeasure MJ {x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d } \ 1 - d" proof - have *: "AE x in MI. uJ n (projI x) = vI n x" for n using uI_proj H(2) by auto have **: "AE x in MI. \n. uJ n (projI x) = vI n x" by (subst AE_all_countable, auto intro: *) have "AE x in MI. \m n. uJ n (projI ((TI^^m) x)) = vI n ((TI^^m) x)" by (rule I.T_AE_iterates[OF **]) then have "AE x in MI. (\m n. uJ n (projI ((TI^^m) x)) = vI n ((TI^^m) x)) \ (\n. projI ((TI^^n) x) = (TJ^^n) (projI x))" using I.qmpt_factor_iterates[OF I.mpt_factor_is_qmpt_factor[OF \I.mpt_factor projI MJ TJ\]] by auto then obtain ZI where ZI: "\x. x \ space MI - ZI \ (\m n. uJ n (projI ((TI^^m) x)) = vI n ((TI^^m) x)) \ (\n. projI ((TI^^n) x) = (TJ^^n) (projI x))" "ZI \ null_sets MI" using AE_E3 by blast have *: "uJ n (projI x) - uJ (n - l) ((TJ ^^ l) (projI x)) = vI n x - vI (n - l) ((TI ^^ l) x)" if "x \ space MI - ZI" for x n l proof - have "(TI^^0) x = x" "(TJ^^0) (projI x) = (projI x)" by auto then show ?thesis using ZI(1)[OF that] by metis qed have "projI-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space MI - ZI = {x \ space MI - ZI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n (projI x) - uJ (n - l) ((TJ ^^ l) (projI x))} \ 1 - d}" by (auto simp add: measurable_space[OF \projI \ measurable MI MJ\]) also have "... = {x \ space MI - ZI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d}" using * by auto also have "... = {x \ space MI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d} - ZI" by auto finally have *: "projI-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space MI - ZI = {x \ space MI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d} - ZI" by simp have "emeasure MJ {x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} = emeasure MI (projI-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space MI)" by (rule measure_preservingE(2)[symmetric], auto) also have "... = emeasure MI ((projI-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space MI) - ZI)" by (rule emeasure_Diff_null_set[OF \ZI \ null_sets MI\, symmetric], measurable) also have "... = emeasure MI ({x \ space MI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d} - ZI)" using * by simp also have "... = emeasure MI {x \ space MI. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < vI n x - vI (n - l) ((TI ^^ l) x)} \ 1 - d}" by (rule emeasure_Diff_null_set[OF \ZI \ null_sets MI\], measurable) also have "... \ 1-d" using delta(3) by simp finally show ?thesis by simp qed text \Then, we go back to $T$ with virtually the same argument.\ have "emeasure M {x \ space M. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d } \ 1 - d" proof - obtain Z where Z: "\x. x \ space M - Z \ (\n. projJ ((T^^n) x) = (TJ^^n) (projJ x))" "Z \ null_sets M" using AE_E3[OF qmpt_factor_iterates[OF mpt_factor_is_qmpt_factor[OF \mpt_factor projJ MJ TJ\]]] by blast have *: "uJ n (projJ x) - uJ (n - l) ((TJ ^^ l) (projJ x)) = u n x - u (n - l) ((T^^ l) x)" if "x \ space M - Z" for x n l proof - have "(T^^0) x = x" "(TJ^^0) (projJ x) = (projJ x)" by auto then show ?thesis using Z(1)[OF that] uJ_proj by metis qed have "projJ-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space M - Z = {x \ space M - Z. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n (projJ x) - uJ (n - l) ((TJ ^^ l) (projJ x))} \ 1 - d}" by (auto simp add: measurable_space[OF \projJ \ measurable M MJ\]) also have "... = {x \ space M - Z. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d}" using * by auto also have "... = {x \ space M. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d} - Z" by auto finally have *: "projJ-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space M - Z = {x \ space M. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d} - Z" by simp have "emeasure MJ {x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} = emeasure M (projJ-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space M)" by (rule measure_preservingE(2)[symmetric], auto) also have "... = emeasure M ((projJ-`{x \ space MJ. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < uJ n x - uJ (n - l) ((TJ ^^ l) x)} \ 1 - d} \ space M) - Z)" by (rule emeasure_Diff_null_set[OF \Z \ null_sets M\, symmetric], measurable) also have "... = emeasure M ({x \ space M. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d} - Z)" using * by simp also have "... = emeasure M {x \ space M. upper_asymptotic_density {n. \l\{1..n}. - delta l * real l < u n x - u (n - l) ((T ^^ l) x)} \ 1 - d}" by (rule emeasure_Diff_null_set[OF \Z \ null_sets M\], measurable) finally show ?thesis using BJ by simp qed then show ?thesis using delta(1) delta(2) by auto qed text \From the quantitative lemma above, we deduce the qualitative statement we are after, still in the setting of the locale.\ lemma infinite_AE: shows "AE x in M. \delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l})" proof - have "\deltaf::real \ nat \ real. \d. ((d > 0 \ d \ 1) \ ((\l. deltaf d l > 0) \ (deltaf d \ 0) \ emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - (deltaf d l) * l} \ 1-d} \ ennreal(1-d)))" apply (subst choice_iff'[symmetric]) using upper_density_good_direction by auto then obtain deltaf::"real \ nat \ real" where H: "\d. d > 0 \ d \1 \ (\l. deltaf d l > 0) \ (deltaf d \ 0) \ emeasure M {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - (deltaf d l) * l} \ 1-d} \ ennreal(1-d)" by blast define U where "U = (\d. {x \ space M. upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - (deltaf d l) * l} \ 1-d})" have [measurable]: "U d \ sets M" for d unfolding U_def by auto have *: "emeasure M (U d) \ 1 - d" if "d>0 \ d\ 1" for d unfolding U_def using H that by auto define V where "V = (\n::nat. U (1/(n+2)))" have [measurable]: "V \ sets M" unfolding V_def by auto have a: "emeasure M V \ 1 - 1 / (n + 2)" for n::nat proof - have "1 - 1 / (n + 2) = 1 - 1 / (real n + 2)" by auto also have "... \ emeasure M (U (1/(real n+2)))" using *[of "1 / (real n + 2)"] by auto also have "... \ emeasure M V" apply (rule Measure_Space.emeasure_mono) unfolding V_def by auto finally show ?thesis by simp qed have b: "(\n::nat. 1 - 1 / (n + 2)) \ ennreal(1 - 0)" by (intro tendsto_intros LIMSEQ_ignore_initial_segment) have "emeasure M V \ 1 - 0" apply (rule Lim_bounded) using a b by auto then have "emeasure M V = 1" by (simp add: emeasure_ge_1_iff) then have "AE x in M. x \ V" by (simp add: emeasure_eq_measure prob_eq_1) moreover { fix x assume "x \ V" then obtain n::nat where "x \ U (1/(real n+2))" unfolding V_def by blast define d where "d = 1/(real n + 2)" have "0 < d" "d\1" unfolding d_def by auto have "0 < 1-d" unfolding d_def by auto also have "... \ upper_asymptotic_density {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - (deltaf d l) * l}" using \x \ U (1/(real n+2))\ unfolding U_def d_def by auto finally have "infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - (deltaf d l) * l}" using upper_asymptotic_density_finite by force then have "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) > - delta l * l})" using H \0 < d\ \d\1\ by auto } ultimately show ?thesis by auto qed end text \Finally, we obtain the full statement, by reducing to the previous situation where the asymptotic average vanishes.\ theorem (in pmpt) Gouezel_Karlsson_Kingman: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. \delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - delta l * l})" proof - have [measurable]: "integrable M (u n)" "u n \ borel_measurable M" for n using subcocycle_integrable[OF assms(1)] by auto define v where "v = birkhoff_sum (\x. -subcocycle_lim u x)" have int [measurable]: "integrable M (\x. -subcocycle_lim u x)" using kingman_theorem_nonergodic(2)[OF assms] by auto have "subcocycle v" unfolding v_def apply (rule subcocycle_birkhoff) using assms \integrable M (\x. -subcocycle_lim u x)\ unfolding subcocycle_def by auto have "subcocycle_avg_ereal v > - \" unfolding v_def using subcocycle_avg_ereal_birkhoff[OF int] kingman_theorem_nonergodic(2)[OF assms] by auto have "AE x in M. subcocycle_lim v x = real_cond_exp M Invariants (\x. -subcocycle_lim u x) x" unfolding v_def by (rule subcocycle_lim_birkhoff[OF int]) moreover have "AE x in M. real_cond_exp M Invariants (\x. - subcocycle_lim u x) x = - subcocycle_lim u x" by (rule real_cond_exp_F_meas[OF int], auto) ultimately have AEv: "AE x in M. subcocycle_lim v x = - subcocycle_lim u x" by auto define w where "w = (\n x. u n x + v n x)" have a: "subcocycle w" unfolding w_def by (rule subcocycle_add[OF assms(1) \subcocycle v\]) have b: "subcocycle_avg_ereal w > -\" unfolding w_def by (rule subcocycle_avg_add(1)[OF assms(1) \subcocycle v\ assms(2) \subcocycle_avg_ereal v > - \\]) have "AE x in M. subcocycle_lim w x = subcocycle_lim u x + subcocycle_lim v x" unfolding w_def by (rule subcocycle_lim_add[OF assms(1) \subcocycle v\ assms(2) \subcocycle_avg_ereal v > - \\]) then have c: "AE x in M. subcocycle_lim w x = 0" using AEv by auto interpret Gouezel_Karlsson_Kingman M T w apply standard using a b c by auto have "AE x in M. \delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. w n x - w (n-l) ((T^^l) x) > - delta l * l})" using infinite_AE by auto moreover { fix x assume H: "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. w n x - w (n-l) ((T^^l) x) > - delta l * l})" "x \ space M" have *: "v n x = - n * subcocycle_lim u x" for n unfolding v_def using birkhoff_sum_of_invariants[OF _ \x \ space M\] by auto have **: "v n ((T^^l) x) = - n * subcocycle_lim u x" for n l proof - have "v n ((T^^l) x) = - n * subcocycle_lim u ((T^^l) x)" unfolding v_def using birkhoff_sum_of_invariants[OF _ T_spaceM_stable(2)[OF \x \ space M\]] by auto also have "... = - n * subcocycle_lim u x" using Invariants_func_is_invariant_n[OF subcocycle_lim_meas_Inv(2) \x \ space M\] by auto finally show ?thesis by simp qed have "w n x - w (n-l) ((T^^l) x) = u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x" if "l \ {1..n}" for n l unfolding w_def using *[of n] **[of "n-l" l] that apply (auto simp add: algebra_simps) by (metis comm_semiring_class.distrib diff_add_inverse nat_le_iff_add of_nat_add) then have "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - delta l * l})" using H(1) by auto } ultimately show ?thesis by auto qed text \The previous theorem only contains a lower bound. The corresponding upper bound follows readily from Kingman's theorem. The next statement combines both upper and lower bounds.\ theorem (in pmpt) Gouezel_Karlsson_Kingman': assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. \delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l})" proof - { fix x assume x: "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - delta l * l})" "(\l. u l x/l) \ subcocycle_lim u x" then obtain alpha::"nat \ real" where a: "\l. alpha l > 0" "alpha \ 0" "infinite {n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - alpha l * l}" by auto define delta::"nat \ real" where "delta = (\l. alpha l + norm(u l x / l - subcocycle_lim u x))" { fix n assume *: "\l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - alpha l * l" have H: "x > -a \ x < a \ abs x < a" for a x::real by simp have "abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l" if "l\{1..n}" for l proof (rule H) have "u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x \ u l x - l * subcocycle_lim u x" using assms(1) subcocycle_ineq[OF assms(1), of l "n-l" x] that by auto also have "... \ l * norm(u l x/l - subcocycle_lim u x)" using that by (auto simp add: algebra_simps divide_simps) also have "... < delta l * l" unfolding delta_def using a(1)[of l] that by auto finally show "u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x < delta l * l" by simp have "- (delta l * l) \ -alpha l * l" unfolding delta_def by (auto simp add: algebra_simps) also have "... < u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x" using * that by auto finally show "u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > -(delta l * l)" by simp qed then have "\l \ {1..n}. abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l" by auto } then have "{n. \l \ {1..n}. u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x > - alpha l * l} \ {n. \l \ {1..n}. abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l}" by auto then have "infinite {n. \l \ {1..n}. abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l}" using a(3) finite_subset by blast moreover have "delta \ 0 + 0" unfolding delta_def using x(2) by (intro tendsto_intros a(2) tendsto_norm_zero LIM_zero) moreover have "delta l > 0" for l unfolding delta_def using a(1)[of l] by auto ultimately have "\delta::nat\real. (\l. delta l > 0) \ (delta \ 0) \ (infinite {n. \l \ {1..n}. abs(u n x - u (n-l) ((T^^l) x) - l * subcocycle_lim u x) < delta l * l})" by auto } then show ?thesis using Gouezel_Karlsson_Kingman[OF assms] kingman_theorem_nonergodic(1)[OF assms] by auto qed end (*of Gouezel_Karlsson.thy*) diff --git a/thys/Ergodic_Theory/Invariants.thy b/thys/Ergodic_Theory/Invariants.thy --- a/thys/Ergodic_Theory/Invariants.thy +++ b/thys/Ergodic_Theory/Invariants.thy @@ -1,1796 +1,1796 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \The invariant sigma-algebra, Birkhoff theorem\ theory Invariants imports Recurrence "HOL-Probability.Conditional_Expectation" begin subsection \The sigma-algebra of invariant subsets\ text \The invariant sigma-algebra of a qmpt is made of those sets that are invariant under the dynamics. When the transformation is ergodic, it is made of sets of zero or full measure. In general, the Birkhoff theorem is expressed in terms of the conditional expectation of an integrable function with respect to the invariant sigma-algebra.\ context qmpt begin text \We define the invariant sigma-algebra, as the sigma algebra of sets which are invariant under the dynamics, i.e., they coincide with their preimage under $T$.\ definition Invariants where "Invariants = sigma (space M) {A \ sets M. T-`A \ space M = A}" text \For this definition to make sense, we need to check that it really defines a sigma-algebra: otherwise, the \verb+sigma+ operation would make garbage out of it. This is the content of the next lemma.\ lemma Invariants_sets: "sets Invariants = {A \ sets M. T-`A \ space M = A}" proof - have "sigma_algebra (space M) {A \ sets M. T-`A \ space M = A}" proof - define I where "I = {A. T-`A \ space M = A}" have i: "\A. A \ I \ A \ space M" unfolding I_def by auto have "algebra (space M) I" proof (subst algebra_iff_Un) have a: "I \ Pow (space M)" using i by auto have b: "{} \ I" unfolding I_def by auto { fix A assume *: "A \ I" then have "T-`(space M - A) = T-`(space M) - T-`A" by auto then have "T-`(space M - A) \ space M = T-`(space M) \ (space M) - T-`A \ (space M)" by auto also have "... = space M - A" using * I_def by (simp add: inf_absorb2 subsetI) finally have "space M - A \ I" unfolding I_def by simp } then have c: "(\a\I. space M - a \ I)" by simp have d: "(\a\I. \b\I. a \ b \ I)" unfolding I_def by auto show "I \ Pow (space M) \ {} \ I \ (\a\I. space M - a \ I) \ (\a\I. \b\I. a \ b \ I)" using a b c d by blast qed moreover have "(\F. range F \ I \ (\i::nat. F i) \ I)" unfolding I_def by auto ultimately have "sigma_algebra (space M) I" using sigma_algebra_iff by auto moreover have "sigma_algebra (space M) (sets M)" using measure_space measure_space_def by auto ultimately have "sigma_algebra (space M) (I \ (sets M))" using sigma_algebra_intersection by auto moreover have "I \ sets M = {A \ sets M. T-`A \ space M = A}" unfolding I_def by auto ultimately show ?thesis by simp qed then show ?thesis unfolding Invariants_def using sigma_algebra.sets_measure_of_eq by blast qed text \By definition, the invariant subalgebra is a subalgebra of the original algebra. This is expressed in the following lemmas.\ lemma Invariants_is_subalg: "subalgebra M Invariants" unfolding subalgebra_def using Invariants_sets Invariants_def by (simp add: space_measure_of_conv) lemma Invariants_in_sets: assumes "A \ sets Invariants" shows "A \ sets M" using Invariants_sets assms by blast lemma Invariants_measurable_func: assumes "f \ measurable Invariants N" shows "f \ measurable M N" using Invariants_is_subalg measurable_from_subalg assms by auto text \We give several trivial characterizations of invariant sets or functions.\ lemma Invariants_vrestr: assumes "A \ sets Invariants" shows "T--`A = A" using assms Invariants_sets Invariants_in_sets[OF assms] by auto lemma Invariants_points: assumes "A \ sets Invariants" "x \ A" shows "T x \ A" using assms Invariants_sets by auto lemma Invariants_func_is_invariant: fixes f::"_ \ 'b::t2_space" assumes "f \ borel_measurable Invariants" "x \ space M" shows "f (T x) = f x" proof - have "{f x} \ sets borel" by simp then have "f-`({f x}) \ space M \ Invariants" using assms(1) by (metis (no_types, lifting) Invariants_def measurable_sets space_measure_of_conv) moreover have "x \ f-`({f x}) \ space M" using assms(2) by blast ultimately have "T x \ f-`({f x}) \ space M" by (rule Invariants_points) then show ?thesis by simp qed lemma Invariants_func_is_invariant_n: fixes f::"_ \ 'b::t2_space" assumes "f \ borel_measurable Invariants" "x \ space M" shows "f ((T^^n) x) = f x" by (induction n, auto simp add: assms Invariants_func_is_invariant) lemma Invariants_func_charac: assumes [measurable]: "f \ measurable M N" and "\x. x \ space M \ f(T x) = f x" shows "f \ measurable Invariants N" proof (rule measurableI) fix A assume "A \ sets N" have "space Invariants = space M" using Invariants_is_subalg subalgebra_def by force show "f -` A \ space Invariants \ sets Invariants" apply (subst Invariants_sets) apply (auto simp add: assms \A \ sets N\ \space Invariants = space M\) using \A \ sets N\ assms(1) measurable_sets by blast next fix x assume "x \ space Invariants" have "space Invariants = space M" using Invariants_is_subalg subalgebra_def by force then show "f x \ space N" using assms(1) \x \ space Invariants\ by (metis measurable_space) qed lemma birkhoff_sum_of_invariants: fixes f::" _ \ real" assumes "f \ borel_measurable Invariants" "x \ space M" shows "birkhoff_sum f n x = n * f x" unfolding birkhoff_sum_def using Invariants_func_is_invariant_n[OF assms] by auto text \There are two possible definitions of the invariant sigma-algebra, competing in the literature: one could also use the sets such that $T^{-1}A$ coincides with $A$ up to a measure $0$ set. It turns out that this is equivalent to being invariant (in our sense) up to a measure $0$ set. Therefore, for all interesting purposes, the two definitions would give the same results. For the proof, we start from an almost invariant set, and build a genuinely invariant set that coincides with it by adding or throwing away null parts. \ proposition Invariants_quasi_Invariants_sets: assumes [measurable]: "A \ sets M" shows "(\B \ sets Invariants. A \ B \ null_sets M) \ (T--`A \ A \ null_sets M)" proof assume "\B \ sets Invariants. A \ B \ null_sets M" then obtain B where "B \ sets Invariants" "A \ B \ null_sets M" by auto then have [measurable]: "B \ sets M" using Invariants_in_sets by simp have "B = T--` B" using Invariants_vrestr \B \ sets Invariants\ by simp then have "T--`A \ B = T--`(A \ B)" by simp moreover have "T--`(A \ B) \ null_sets M" by (rule T_quasi_preserves_null2(1)[OF \A \ B \ null_sets M\]) ultimately have "T--`A \ B \ null_sets M" by simp then show "T--`A \ A \ null_sets M" by (rule null_sym_diff_transitive) (auto simp add: \A \ B \ null_sets M\ Un_commute) next assume H: "T --` A \ A \ null_sets M" have [measurable]: "\n. (T^^n)--`A \ sets M" by simp { fix K assume [measurable]: "K \ sets M" and "T--`K \ K \ null_sets M" fix n::nat have "(T^^n)--`K \ K \ null_sets M" proof (induction n) case 0 have "(T^^0)--` K = K" using T_vrestr_0 by simp then show ?case using Diff_cancel sup.idem by (metis null_sets.empty_sets) next case (Suc n) have "T--`((T^^n)--`K \ K) \ null_sets M" using Suc.IH T_quasi_preserves_null(1)[of "((T^^n)--`K \ K)"] by simp then have *: "(T^^(Suc n))--`K \ T--`K \ null_sets M" using T_vrestr_composed(2)[OF \K \ sets M\] by simp then show ?case by (rule null_sym_diff_transitive, simp add: \T--`K \ K \ null_sets M\ \K \ sets M\, measurable) qed } note * = this define C where "C = (\n. (T^^n)--`A)" have [measurable]: "C \ sets M" unfolding C_def by simp have "C \ A \ (\n. (T^^n)--`A \ A)" unfolding C_def by auto moreover have "(\n. (T^^n)--`A \ A) \ null_sets M" using * null_sets_UN assms \T --` A \ A \ null_sets M\ by auto ultimately have CA: "C \ A \ null_sets M" by (meson \C \ sets M\ assms sets.Diff sets.Un null_sets_subset) then have "T--`(C \ A) \ null_sets M" by (rule T_quasi_preserves_null2(1)) then have "T--`C \ T--`A \ null_sets M" by simp then have "T--`C \ A \ null_sets M" by (rule null_sym_diff_transitive, auto simp add: H) then have TCC: "T--`C \ C \ null_sets M" apply (rule null_sym_diff_transitive) using CA by (auto simp add: Un_commute) have "C \ (\n\{1..}. (T^^n)--`A)" unfolding C_def by auto moreover have "T--`C = (\n\{1..}. (T^^n)--`A)" using T_vrestr_composed(2)[OF assms] by (simp add: C_def atLeast_Suc_greaterThan greaterThan_0) ultimately have "C \ T--`C" by blast then have "(T^^0)--`C \ (T^^1)--`C" using T_vrestr_0 by auto moreover have "(T^^1)--`C \ (\n\{1..}. (T^^n)--`C)" by auto ultimately have "(T^^0)--`C \ (\n\{1..}. (T^^n)--`C)" by auto then have "(T^^0)--`C \ (\n\{1..}. (T^^n)--`C) = (\n\{1..}. (T^^n)--`C)" by auto moreover have "(\n. (T^^n)--`C) = (T^^0)--`C \ (\n\{1..}. (T^^n)--`C)" by (rule union_insert_0) ultimately have C2: "(\n. (T^^n)--`C) = (\n\{1..}. (T^^n)--`C)" by simp define B where "B = (\n. (T^^n)--`C)" have [measurable]: "B \ sets M" unfolding B_def by simp have "B \ C \ (\n. (T^^n)--`C \ C)" unfolding B_def by auto moreover have "(\n. (T^^n)--`C \ C) \ null_sets M" using * null_sets_UN assms TCC by auto ultimately have "B \ C \ null_sets M" by (meson \B \ sets M\ \C \ sets M\ assms sets.Diff sets.Un null_sets_subset) then have "B \ A \ null_sets M" by (rule null_sym_diff_transitive, auto simp add: CA) then have a: "A \ B \ null_sets M" by (simp add: Un_commute) have "T--`B = (\n\{1..}. (T^^n)--`C)" using T_vrestr_composed(2)[OF \C \ sets M\] by (simp add: B_def atLeast_Suc_greaterThan greaterThan_0) then have "T--`B = B" unfolding B_def using C2 by simp then have "B \ sets Invariants" using Invariants_sets vimage_restr_def by auto then show "\B \ sets Invariants. A \ B \ null_sets M" using a by blast qed text \In a conservative setting, it is enough to be included in its image or its preimage to be almost invariant: otherwise, since the difference set has disjoint preimages, and is therefore null by conservativity.\ lemma (in conservative) preimage_included_then_almost_invariant: assumes [measurable]: "A \ sets M" and "T--`A \ A" shows "A \ (T--`A) \ null_sets M" proof - define B where "B = A - T--`A" then have [measurable]: "B \ sets M" by simp have "(T^^(Suc n))--`A \ (T^^n)--`A" for n using T_vrestr_composed(3)[OF assms(1)] vrestr_inclusion[OF assms(2)] by auto then have "disjoint_family (\n. (T^^n)--`A - (T^^(Suc n))--`A)" by (rule disjoint_family_Suc2[where ?A = "\n. (T^^n)--`A"]) moreover have "(T^^n)--`A - (T^^(Suc n))--`A = (T^^n)--`B" for n unfolding B_def Suc_eq_plus1 using T_vrestr_composed(3)[OF assms(1)] by auto ultimately have "disjoint_family (\n. (T^^n)--` B)" by simp then have "\n. n \ 0 \ ((T^^n)--`B) \ B = {}" unfolding disjoint_family_on_def by (metis UNIV_I T_vrestr_0(1)[OF \B \ sets M\]) then have "\n. n > 0 \ (T^^n)-`B \ B = {}" unfolding vimage_restr_def by (simp add: Int_assoc) then have "B \ null_sets M" using disjoint_then_null[OF \B \ sets M\] Int_commute by auto then show ?thesis unfolding B_def using assms(2) by (simp add: Diff_mono Un_absorb2) qed lemma (in conservative) preimage_includes_then_almost_invariant: assumes [measurable]: "A \ sets M" and "A \ T--`A" shows "A \ (T--`A) \ null_sets M" proof - define B where "B = T--`A - A" then have [measurable]: "B \ sets M" by simp have "\n. (T^^(Suc n))--`A \ (T^^n)--`A" using T_vrestr_composed(3)[OF assms(1)] vrestr_inclusion[OF assms(2)] by auto then have "disjoint_family (\n. (T^^(Suc n))--`A - (T^^n)--`A)" by (rule disjoint_family_Suc[where ?A = "\n. (T^^n)--`A"]) moreover have "\n. (T^^(Suc n))--`A - (T^^n)--`A = (T^^n)--`B" unfolding B_def Suc_eq_plus1 using T_vrestr_composed(3)[OF assms(1)] by auto ultimately have "disjoint_family (\n. (T^^n)--` B)" by simp then have "\n. n \ 0 \ ((T^^n)--`B) \ B = {}" unfolding disjoint_family_on_def by (metis UNIV_I T_vrestr_0(1)[OF \B \ sets M\]) then have "\n. n > 0 \ (T^^n)-`B \ B = {}" unfolding vimage_restr_def by (simp add: Int_assoc) then have "B \ null_sets M" using disjoint_then_null[OF \B \ sets M\] Int_commute by auto then show ?thesis unfolding B_def using assms(2) by (simp add: Diff_mono Un_absorb1) qed text \The above properties for sets are also true for functions: if $f$ and $f \circ T$ coincide almost everywhere, i.e., $f$ is almost invariant, then $f$ coincides almost everywhere with a true invariant function. The idea of the proof is straightforward: throw away the orbits on which $f$ is not really invariant (say this is the complement of the good set), and replace it by $0$ there. However, this does not work directly: the good set is not invariant, some points may have a non-constant value of $f$ on their orbit but reach the good set eventually. One can however define $g$ to be equal to the eventual value of $f$ along the orbit, if the orbit reaches the good set, and $0$ elsewhere.\ proposition Invariants_quasi_Invariants_functions: fixes f::"_ \ 'b::{second_countable_topology, t2_space}" assumes f_meas [measurable]: "f \ borel_measurable M" shows "(\g \ borel_measurable Invariants. AE x in M. f x = g x) \ (AE x in M. f(T x) = f x)" proof assume "\g\borel_measurable Invariants. AE x in M. f x = g x" then obtain g where g:"g\borel_measurable Invariants" "AE x in M. f x = g x" by blast then have [measurable]: "g \ borel_measurable M" using Invariants_measurable_func by auto define A where "A = {x \ space M. f x = g x}" have [measurable]: "A \ sets M" unfolding A_def by simp define B where "B = space M - A" have [measurable]: "B \ sets M" unfolding B_def by simp moreover have "AE x in M. x \ B" unfolding B_def A_def using g(2) by auto ultimately have "B \ null_sets M" using AE_iff_null_sets by blast then have "T--`B \ null_sets M" by (rule T_quasi_preserves_null2(1)) then have "B \ T--`B \ null_sets M" using \B \ null_sets M\ by auto then have "AE x in M. x \ (B \ T--`B)" using AE_iff_null_sets null_setsD2 by blast then have i: "AE x in M. x \ space M - (B \ T--`B)" by auto { fix x assume *: "x \ space M - (B \ T--`B)" then have "x \ A" unfolding B_def by blast then have "f x = g x" unfolding A_def by blast have "T x \ A" using * B_def by auto then have "f(T x) = g(T x)" unfolding A_def by blast moreover have "g(T x) = g x" apply (rule Invariants_func_is_invariant) using * by (auto simp add: assms \g\borel_measurable Invariants\) ultimately have "f(T x) = f x" using \f x = g x\ by simp } then show "AE x in M. f(T x) = f x" using i by auto next assume *: "AE x in M. f (T x) = f x" text \\verb+good_set+ is the set of points for which $f$ is constant on their orbit. Here, we define $g = f$. If a point ever enters \verb+good_set+, then we take $g$ to be the value of $f$ there. Otherwise, $g$ takes an arbitrary value, say $y_0$.\ define good_set where "good_set = {x \ space M. \n. f((T^^(Suc n)) x) = f((T^^n) x)}" define good_time where "good_time = (\x. Inf {n. (T^^n) x \ good_set})" have "AE x in M. x \ good_set" using T_AE_iterates[OF *] by (simp add: good_set_def) have [measurable]: "good_set \ sets M" unfolding good_set_def by auto obtain y0::'b where True by auto define g where "g = (\x. if (\n. (T^^n) x \ good_set) then f((T^^(good_time x)) x) else y0)" have [measurable]: "good_time \ measurable M (count_space UNIV)" unfolding good_time_def by measurable have [measurable]: "g \ borel_measurable M" unfolding g_def by measurable have "f x = g x" if "x \ good_set" for x proof - have a: "0 \ {n. (T^^n) x \ good_set}" using that by simp have "good_time x = 0" unfolding good_time_def apply (intro cInf_eq_non_empty) using a by blast+ moreover have "{n. (T^^n) x \ good_set} \ {}" using a by blast ultimately show "f x = g x" unfolding g_def by auto qed then have "AE x in M. f x = g x" using \AE x in M. x \ good_set\ by auto have *: "f((T^^(Suc 0)) x) = f((T^^0) x)" if "x \ good_set" for x using that unfolding good_set_def by blast have good_1: "T x \ good_set \ f(T x) = f x" if "x \ good_set" for x using *[OF that] that unfolding good_set_def apply (auto) unfolding T_Tn_T_compose by blast then have good_k: "\x. x \ good_set \ (T^^k) x \ good_set \ f((T^^k) x) = f x" for k by (induction k, auto) have "g(T x) = g x" if "x \ space M" for x proof (cases) assume *: "\n. (T^^n) (T x) \ good_set" define n where "n = Inf {n. (T^^n) (T x) \ good_set}" have "(T^^n)(T x) \ good_set" using * Inf_nat_def1 by (metis empty_iff mem_Collect_eq n_def) then have a: "(T^^(n+1)) x \ good_set" by (metis Suc_eq_plus1 comp_eq_dest_lhs funpow.simps(2) funpow_swap1) then have **: "\m. (T^^m) x \ good_set" by blast define m where "m = Inf {m. (T^^m) x \ good_set}" then have "(T^^m) x \ good_set" using ** Inf_nat_def1 by (metis empty_iff mem_Collect_eq) have "n+1 \ {m. (T^^m) x \ good_set}" using a by simp then have "m \ n+1" using m_def by (simp add: Inf_nat_def Least_le) then obtain k where "n+1 = m + k" using le_iff_add by blast have "g x = f((T^^m) x)" unfolding g_def good_time_def using ** m_def by simp also have "... = f((T^^k) ((T^^m) x))" using \(T^^m) x \ good_set\ good_k by simp also have "... = f((T^^(n+1))x)" using \n+1 = m + k\[symmetric] funpow_add by (metis add.commute comp_apply) also have "... = f((T^^n) (T x))" using funpow_Suc_right by (metis Suc_eq_plus1 comp_apply) also have "... = g(T x)" unfolding g_def good_time_def using * n_def by simp finally show "g(T x) = g x" by simp next assume *: "\(\n. (T^^n) (T x) \ good_set)" then have "g(T x) = y0" unfolding g_def by simp have **: "\(\n. (T^^(Suc n)) x \ good_set)" using funpow_Suc_right * by (metis comp_apply) have "T x \ good_set" using good_k * by blast then have "x \ good_set" using good_1 by auto then have "\(\n. (T^^n) x \ good_set)" using ** using good_1 by fastforce then have "g x = y0" unfolding g_def by simp then show "g(T x) = g x" using \g(T x) = y0\ by simp qed then have "g \ borel_measurable Invariants" by (rule Invariants_func_charac[OF \g \ borel_measurable M\]) then show "\g\borel_measurable Invariants. AE x in M. f x = g x" using \AE x in M. f x = g x\ by blast qed text \In a conservative setting, it suffices to have an almost everywhere inequality to get an almost everywhere equality, as the set where there is strict inequality has $0$ measure as its iterates are disjoint, by conservativity.\ proposition (in conservative) AE_decreasing_then_invariant: fixes f::"_ \ 'b::{linorder_topology, second_countable_topology}" assumes "AE x in M. f(T x) \ f x" and [measurable]: "f \ borel_measurable M" shows "AE x in M. f(T x) = f x" proof - obtain D::"'b set" where D: "countable D" "(\x y. x < y \ (\d \ D. x \ d \ d < y))" using countable_separating_set_linorder2 by blast define A where "A = {x \ space M. f(T x) \ f x}" then have [measurable]: "A \ sets M" by simp define B where "B = {x \ space M. \n. f((T^^(n+1)) x) \ f((T^^n)x)}" then have [measurable]: "B \ sets M" by simp have "space M - A \ null_sets M" unfolding A_def using assms by (simp add: assms(1) AE_iff_null_sets) then have "(\n. (T^^n)--`(space M - A)) \ null_sets M" by (metis null_sets_UN T_quasi_preserves_null2(2)) moreover have "space M - B = (\n. (T^^n)--`(space M - A))" unfolding B_def A_def by auto ultimately have "space M - B \ null_sets M" by simp have *: "B = (\n. (T^^n)--`A)" unfolding B_def A_def by auto then have "T--`B = (\n. T--` (T^^n)--`A)" by auto also have "... = (\n. (T^^(n+1))--`A)" using T_vrestr_composed(2)[OF \A \ sets M\] by simp also have "... \ (\n. (T^^n)--`A)" by blast finally have B1: "B \ T--`B" using * by simp have "B \ A" using * T_vrestr_0[OF \A \ sets M\] by blast then have B2: "\x. x \ B \ f(T x) \ f x" unfolding A_def by auto define C where "C = (\t. {x \ B. f x \ t})" { fix t have "C t = B \ f-`{..t} \ space M" unfolding C_def using sets.sets_into_space[OF \B \ sets M\] by auto then have [measurable]: "C t \ sets M" using assms(2) by simp have "C t \ T--`(C t)" using B1 unfolding C_def vimage_restr_def apply auto using B2 order_trans by blast then have "T--`(C t) - C t \ null_sets M" by (metis Diff_mono Un_absorb1 preimage_includes_then_almost_invariant[OF \C t \ sets M\]) } then have "(\d\D. T--`(C d) - C d) \ null_sets M" using \countable D\ by (simp add: null_sets_UN') then have "(space M - B) \ (\d\D. T--`(C d) - C d) \ null_sets M" using \space M - B \ null_sets M\ by auto then have "AE x in M. x \ (space M - B) \ (\d\D. T--`(C d) - C d)" using AE_not_in by blast moreover { fix x assume x: "x \ space M" "x \ (space M - B) \ (\d\D. T--`(C d) - C d)" then have "x \ B" by simp then have "T x \ B" using B1 by auto have "f(T x) = f x" proof (rule ccontr) assume "f(T x) \ f x" then have "f(T x) < f x" using B2[OF \x \ B\] by simp then obtain d where d: "d \ D" "f(T x) \ d \ d < f x" using D by auto then have "T x \ C d" using \T x \ B\ unfolding C_def by simp then have "x \ T--`(C d)" using \x \ space M\ by simp then have "x \ C d" using x \d \ D\ by simp then have "f x \ d" unfolding C_def by simp then show False using d by auto qed } ultimately show ?thesis by auto qed proposition (in conservative) AE_increasing_then_invariant: fixes f::"_ \ 'b::{linorder_topology, second_countable_topology}" assumes "AE x in M. f(T x) \ f x" and [measurable]: "f \ borel_measurable M" shows "AE x in M. f(T x) = f x" proof - obtain D::"'b set" where D: "countable D" "(\x y. x < y \ (\d \ D. x < d \ d \ y))" using countable_separating_set_linorder1 by blast define A where "A = {x \ space M. f(T x) \ f x}" then have [measurable]: "A \ sets M" by simp define B where "B = {x \ space M. \n. f((T^^(n+1)) x) \ f((T^^n)x)}" then have [measurable]: "B \ sets M" by simp have "space M - A \ null_sets M" unfolding A_def using assms by (simp add: assms(1) AE_iff_null_sets) then have "(\n. (T^^n)--`(space M - A)) \ null_sets M" by (metis null_sets_UN T_quasi_preserves_null2(2)) moreover have "space M - B = (\n. (T^^n)--`(space M - A))" unfolding B_def A_def by auto ultimately have "space M - B \ null_sets M" by simp have *: "B = (\n. (T^^n)--`A)" unfolding B_def A_def by auto then have "T--`B = (\n. T--` (T^^n)--`A)" by auto also have "... = (\n. (T^^(n+1))--`A)" using T_vrestr_composed(2)[OF \A \ sets M\] by simp also have "... \ (\n. (T^^n)--`A)" by blast finally have B1: "B \ T--`B" using * by simp have "B \ A" using * T_vrestr_0[OF \A \ sets M\] by blast then have B2: "\x. x \ B \ f(T x) \ f x" unfolding A_def by auto define C where "C = (\t. {x \ B. f x \ t})" { fix t have "C t = B \ f-`{t..} \ space M" unfolding C_def using sets.sets_into_space[OF \B \ sets M\] by auto then have [measurable]: "C t \ sets M" using assms(2) by simp have "C t \ T--`(C t)" using B1 unfolding C_def vimage_restr_def apply auto using B2 order_trans by blast then have "T--`(C t) - C t \ null_sets M" by (metis Diff_mono Un_absorb1 preimage_includes_then_almost_invariant[OF \C t \ sets M\]) } then have "(\d\D. T--`(C d) - C d) \ null_sets M" using \countable D\ by (simp add: null_sets_UN') then have "(space M - B) \ (\d\D. T--`(C d) - C d) \ null_sets M" using \space M - B \ null_sets M\ by auto then have "AE x in M. x \ (space M - B) \ (\d\D. T--`(C d) - C d)" using AE_not_in by blast moreover { fix x assume x: "x \ space M" "x \ (space M - B) \ (\d\D. T--`(C d) - C d)" then have "x \ B" by simp then have "T x \ B" using B1 by auto have "f(T x) = f x" proof (rule ccontr) assume "f(T x) \ f x" then have "f(T x) > f x" using B2[OF \x \ B\] by simp then obtain d where d: "d \ D" "f(T x) \ d \ d > f x" using D by auto then have "T x \ C d" using \T x \ B\ unfolding C_def by simp then have "x \ T--`(C d)" using \x \ space M\ by simp then have "x \ C d" using x \d \ D\ by simp then have "f x \ d" unfolding C_def by simp then show False using d by auto qed } ultimately show ?thesis by auto qed text \For an invertible map, the invariants of $T$ and $T^{-1}$ are the same.\ lemma Invariants_Tinv: assumes "invertible_qmpt" shows "qmpt.Invariants M Tinv = Invariants" proof - interpret I: qmpt M Tinv using Tinv_qmpt[OF assms] by auto have "(T -` A \ space M = A) \ (Tinv -` A \ space M = A)" if "A \ sets M" for A proof assume "T -` A \ space M = A" then show "Tinv -` A \ space M = A" using assms that unfolding Tinv_def invertible_qmpt_def apply auto apply (metis IntE UNIV_I bij_def imageE inv_f_f vimageE) apply (metis I.T_spaceM_stable(1) Int_iff Tinv_def bij_inv_eq_iff vimageI) done next assume "Tinv -` A \ space M = A" then show "T -` A \ space M = A" using assms that unfolding Tinv_def invertible_qmpt_def apply auto apply (metis IntE bij_def inv_f_f vimageE) apply (metis T_Tinv_of_set T_meas Tinv_def assms qmpt.vrestr_of_set qmpt_axioms vrestr_image(3)) done qed then have "{A \ sets M. Tinv -` A \ space M = A} = {A \ sets M. T -` A \ space M = A}" by blast then show ?thesis unfolding Invariants_def I.Invariants_def by auto qed end sublocale fmpt \ finite_measure_subalgebra M Invariants unfolding finite_measure_subalgebra_def finite_measure_subalgebra_axioms_def using Invariants_is_subalg by (simp add: finite_measureI) context fmpt begin text \The conditional expectation with respect to the invariant sigma-algebra is the same for $f$ or $f \circ T$, essentially by definition.\ lemma Invariants_of_foTn: fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "AE x in M. real_cond_exp M Invariants (f o (T^^n)) x = real_cond_exp M Invariants f x" proof (rule real_cond_exp_charact) fix A assume [measurable]: "A \ sets Invariants" then have [measurable]: "A \ sets M" using Invariants_in_sets by blast then have ind_meas [measurable]: "((indicator A)::('a \ real)) \ borel_measurable Invariants" by auto have "set_lebesgue_integral M A (f \ (T^^n)) = (\x. indicator A x * f((T^^n) x) \M)" by (auto simp: comp_def set_lebesgue_integral_def) also have "... = (\x. indicator A ((T^^n) x) * f ((T^^n) x) \M)" by (rule Bochner_Integration.integral_cong, auto simp add: Invariants_func_is_invariant_n[OF ind_meas]) also have "... = (\x. indicator A x * f x \M)" apply (rule Tn_integral_preserving(2)) using integrable_mult_indicator[OF \A \ sets M\ assms] by auto also have "... = (\x. indicator A x * real_cond_exp M Invariants f x \M)" apply (rule real_cond_exp_intg(2)[symmetric]) using integrable_mult_indicator[OF \A \ sets M\ assms] by auto also have "... = set_lebesgue_integral M A (real_cond_exp M Invariants f)" by (auto simp: set_lebesgue_integral_def) finally show "set_lebesgue_integral M A (f \ (T^^n)) = set_lebesgue_integral M A (real_cond_exp M Invariants f)" by simp qed (auto simp add: assms real_cond_exp_int Tn_integral_preserving(1)[OF assms] comp_def) lemma Invariants_of_foT: fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "AE x in M. real_cond_exp M Invariants f x = real_cond_exp M Invariants (f o T) x" using Invariants_of_foTn[OF assms, where ?n = 1] by auto lemma birkhoff_sum_Invariants: fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "AE x in M. real_cond_exp M Invariants (birkhoff_sum f n) x = n * real_cond_exp M Invariants f x" proof - define F where "F = (\i. f o (T^^i))" have [measurable]: "\i. F i \ borel_measurable M" unfolding F_def by auto have *: "integrable M (F i)" for i unfolding F_def by (subst comp_def, rule Tn_integral_preserving(1)[OF assms, of i]) have "AE x in M. n * real_cond_exp M Invariants f x = (\i\{..i\{..i\{..i\{..x. \i\{..x. \i\{..Birkhoff theorem\ subsubsection \Almost everywhere version of Birkhoff theorem\ text \This paragraph is devoted to the proof of Birkhoff theorem, arguably the most fundamental result of ergodic theory. This theorem asserts that Birkhoff averages of an integrable function $f$ converge almost surely, to the conditional expectation of $f$ with respect to the invariant sigma algebra. This result implies for instance the strong law of large numbers (in probability theory). There are numerous proofs of this statement, but none is really easy. We follow the very efficient argument given in Katok-Hasselblatt. To help the reader, here is the same proof informally. The first part of the proof is formalized in \verb+birkhoff_lemma1+, the second one in \verb+birkhoff_lemma+, and the conclusion in \verb+birkhoff_theorem+. Start with an integrable function $g$. let $G_n(x) = \max_{k\leq n} S_k g(x)$. Then $\limsup S_n g/n \leq 0$ outside of $A$, the set where $G_n$ tends to infinity. Moreover, $G_{n+1} - G_n \circ T$ is bounded by $g$, and tends to $g$ on $A$. It follows from the dominated convergence theorem that $\int_A G_{n+1} - G_n \circ T \to \int_A g$. As $\int_A G_{n+1} - G_n \circ T = \int_A G_{n+1} - G_n \geq 0$, we obtain $\int_A g \geq 0$. Apply now this result to the function $g = f - E(f | I) - \epsilon$, where $\epsilon>0$ is fixed. Then $\int_A g = -\epsilon \mu(A)$, then have $\mu(A) = 0$. Thus, almost surely, $\limsup S_n g/n\leq 0$, i.e., $\limsup S_n f/n \leq E(f|I)+\epsilon$. Letting $\epsilon$ tend to $0$ gives $\limsup S_n f/n \leq E(f|I)$. Applying the same result to $-f$ gives $S_n f/n \to E(f|I)$. \ context fmpt begin lemma birkhoff_aux1: fixes f::"'a \ real" assumes [measurable]: "integrable M f" defines "A \ {x \ space M. limsup (\n. ereal(birkhoff_sum f n x)) = \}" shows "A \ sets Invariants" "(\x. f x * indicator A x \M) \ 0" proof - let ?bsf = "birkhoff_sum f" have [measurable]: "A \ sets M" unfolding A_def by simp have Ainv: "x \ A \ T x \ A" if "x \ space M" for x proof - have "ereal(?bsf (1 + n) x) = ereal(f x) + ereal(?bsf n (T x))" for n unfolding birkhoff_sum_cocycle birkhoff_sum_1 by simp moreover have "limsup (\n. ereal(f x) + ereal(?bsf n (T x))) = ereal(f x) + limsup(\n. ereal(?bsf n (T x)))" by (rule ereal_limsup_lim_add, auto) moreover have "limsup (\n. ereal(?bsf (n+1) x)) = limsup (\n. ereal(?bsf n x))" using limsup_shift by simp ultimately have "limsup (\n. ereal(birkhoff_sum f n x)) = ereal(f x) + limsup (\n. ereal(?bsf n (T x)))" by simp then have "limsup (\n. ereal(?bsf n x)) = \ \ limsup (\n. ereal(?bsf n (T x))) = \" by simp then show "x \ A \ T x \ A" using \x \ space M\ A_def by simp qed then show "A \ sets Invariants" using assms(2) Invariants_sets by auto define F where "F = (\n x. MAX k \{0..n}. ?bsf k x)" have [measurable]: "\n. F n \ borel_measurable M" unfolding F_def by measurable have intFn: "integrable M (F n)" for n unfolding F_def by (rule integrable_MAX, auto simp add: birkhoff_sum_integral(1)[OF assms(1)]) have Frec: "F (n+1) x - F n (T x) = max (-F n (T x)) (f x)" for n x proof - have "{0..n+1} = {0} \ {1..n+1}" by auto then have "(\k. ?bsf k x) ` {0..n+1} = (\k. ?bsf k x) ` {0} \ (\k. ?bsf k x) ` {1..n+1}" by blast then have *: "(\k. ?bsf k x) ` {0..n+1} = {0} \ (\k. ?bsf k x) ` {1..n+1}" using birkhoff_sum_1(1) by simp have b: "F (n+1) x = max (Max {0}) (MAX k \{1..n+1}. ?bsf k x)" by (subst F_def, subst *, rule Max.union, auto) have "(\k. ?bsf k x) ` {1..n+1} = (\k. ?bsf (1+k) x) ` {0..n}" using Suc_le_D by fastforce also have "... = (\k. f x + ?bsf k (T x)) ` {0..n}" by (subst birkhoff_sum_cocycle, subst birkhoff_sum_1(2), auto) finally have c: "F (n+1) x = max 0 (MAX k \{0..n}. ?bsf k (T x) + f x)" using b by (simp add: add_ac) have "{f x + birkhoff_sum f k (T x) |k. k \{0..n}} = (+) (f x) ` {birkhoff_sum f k (T x) |k. k \{0..n}}" by blast have "(MAX k \{0..n}. ?bsf k (T x) + f x) = (MAX k \{0..n}. ?bsf k (T x)) + f x" by (rule Max_add_commute) auto also have "... = F n (T x) + f x" unfolding F_def by simp finally have "(MAX k \{0..n}. ?bsf k (T x) + f x) = f x + F n (T x)" by simp then have "F (n+1) x = max 0 (f x + F n (T x))" using c by simp then show "F (n+1) x - F n (T x) = max (-F n (T x)) (f x)" by auto qed have a: "abs((F (n+1) x - F n (T x)) * indicator A x) \ abs(f x)" for n x proof - have "F (n+1) x -F n (T x) \ f x" using Frec by simp then have *: "F (n+1) x -F n (T x) \ - abs(f x)" by simp have "F n (T x) \ birkhoff_sum f 0 (T x)" unfolding F_def apply (rule Max_ge, simp) using atLeastAtMost_iff by blast then have "F n (T x) \ 0" using birkhoff_sum_1(1) by simp then have "-F n (T x) \ abs (f x)" by simp moreover have "f x \ abs(f x)" by simp ultimately have "F (n+1) x -F n (T x) \ abs(f x)" using Frec by simp then have "abs(F (n+1) x - F n (T x)) \ abs(f x)" using * by simp then show "abs((F (n+1) x - F n (T x)) * indicator A x) \ abs(f x)" unfolding indicator_def by auto qed have b: "(\n. (F (n+1) x - F n (T x)) * indicator A x) \ f x * indicator A x" for x proof (rule tendsto_eventually, cases) assume "x \ A" then have "T x \ A" using Ainv A_def by auto then have "limsup (\n. ereal(birkhoff_sum f n (T x))) > ereal(-f x)" unfolding A_def by simp then obtain N where "ereal(?bsf N (T x)) > ereal(-f x)" using Limsup_obtain by blast then have *: "?bsf N (T x) > -f x" by simp { fix n assume "n\N" then have "?bsf N (T x) \ (\k. ?bsf k (T x)) ` {0..n}" by auto then have "F n (T x) \ ?bsf N (T x)" unfolding F_def by simp then have "F n (T x) \ -f x" using * by simp then have "max (-F n (T x)) (f x) = f x" by simp then have "F (n+1) x - F n (T x) = f x" using Frec by simp then have "(F (n+1) x - F n (T x)) * indicator A x = f x * indicator A x" by simp } then show "eventually (\n. (F (n+1) x - F n (T x)) * indicator A x = f x * indicator A x) sequentially" using eventually_sequentially by blast next assume "\(x \ A)" then have "indicator A x = (0::real)" by simp then show "eventually (\n. (F (n+1) x - F n (T x)) * indicator A x = f x * indicator A x) sequentially" by auto qed have lim: "(\n. (\x. (F (n+1) x - F n (T x)) * indicator A x \M)) \ (\x. f x * indicator A x \M)" proof (rule integral_dominated_convergence[where ?w = "(\x. abs(f x))"]) show "integrable M (\x. \f x\)" using assms(1) by auto show "AE x in M. (\n. (F (n + 1) x - F n (T x)) * indicator A x) \ f x * indicator A x" using b by auto show "\n. AE x in M. norm ((F (n + 1) x - F n (T x)) * indicator A x) \ \f x\" using a by auto qed (simp_all) have "(\x. (F (n+1) x - F n (T x)) * indicator A x \M) \ 0" for n proof - have "(\x. F n (T x) * indicator A x \M) = (\x. (\x. F n x * indicator A x) (T x) \M)" by (rule Bochner_Integration.integral_cong, auto simp add: Ainv indicator_def) also have "... = (\x. F n x * indicator A x \M)" by (rule T_integral_preserving, auto simp add: intFn integrable_real_mult_indicator) finally have i: "(\x. F n (T x) * indicator A x \M) = (\x. F n x * indicator A x \M)" by simp have "(\x. (F (n+1) x - F n (T x)) * indicator A x \M) = (\x. F (n+1) x * indicator A x - F n (T x) * indicator A x \M)" by (simp add: mult.commute right_diff_distrib) also have "... = (\x. F (n+1) x * indicator A x \M) - (\x. F n (T x) * indicator A x \M)" by (rule Bochner_Integration.integral_diff, auto simp add: intFn integrable_real_mult_indicator T_meas T_integral_preserving(1)) also have "... = (\x. F (n+1) x * indicator A x \M) - (\x. F n x * indicator A x \M)" using i by simp also have "... = (\x. F (n+1) x * indicator A x - F n x * indicator A x \M)" by (rule Bochner_Integration.integral_diff[symmetric], auto simp add: intFn integrable_real_mult_indicator) also have "... = (\x. (F (n+1) x - F n x) * indicator A x \M)" by (simp add: mult.commute right_diff_distrib) finally have *: "(\x. (F (n+1) x - F n (T x)) * indicator A x \M) = (\x. (F (n+1) x - F n x) * indicator A x \M)" by simp have "F n x \ F (n+1) x" for x unfolding F_def by (rule Max_mono, auto) then have "(F (n+1) x - F n x) * indicator A x \ 0" for x by simp then have "integral\<^sup>L M (\x. 0) \ integral\<^sup>L M (\x. (F (n+1) x - F n x) * indicator A x)" by (auto simp add: intFn integrable_real_mult_indicator intro: integral_mono) then have "(\x. (F (n+1) x - F n x) * indicator A x \M) \ 0" by simp then show "(\x. (F (n+1) x - F n (T x)) * indicator A x \M) \ 0" using * by simp qed then show "(\x. f x * indicator A x \M) \ 0" using lim by (simp add: LIMSEQ_le_const) qed lemma birkhoff_aux2: fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "AE x in M. limsup (\n. ereal(birkhoff_sum f n x / n)) \ real_cond_exp M Invariants f x" proof - { fix \ assume "\ > (0::real)" define g where "g = (\x. f x - real_cond_exp M Invariants f x - \)" then have intg: "integrable M g" using assms real_cond_exp_int(1) assms by auto define A where "A = {x \ space M. limsup (\n. ereal(birkhoff_sum g n x)) = \}" have Ag: "A \ sets Invariants" "(\x. g x * indicator A x \M) \ 0" unfolding A_def by (rule birkhoff_aux1[where ?f = g, OF intg])+ then have [measurable]: "A \ sets M" by (simp add: Invariants_in_sets) have eq: "(\x. indicator A x * real_cond_exp M Invariants f x \M) = (\x. indicator A x * f x \M)" proof (rule real_cond_exp_intg[where ?f = "\x. (indicator A x)::real" and ?g = f]) have "(\x. indicator A x * f x) = (\x. f x * indicator A x)" by auto then show "integrable M (\x. indicator A x * f x)" using integrable_real_mult_indicator[OF \A \ sets M\ assms] by simp show "indicator A \ borel_measurable Invariants" using \A \ sets Invariants\ by measurable qed (simp) have "0 \ (\x. g x * indicator A x \M)" using Ag by simp also have "... = (\x. f x * indicator A x - real_cond_exp M Invariants f x * indicator A x - \ * indicator A x \M)" unfolding g_def by (simp add: left_diff_distrib) also have "... = (\x. f x * indicator A x \M) - (\x. real_cond_exp M Invariants f x * indicator A x \M) - (\x. \ * indicator A x \M)" using assms real_cond_exp_int(1)[OF assms] integrable_real_mult_indicator[OF \A \ sets M\] by (auto simp: simp del: integrable_mult_left_iff) also have "... = - (\x. \ * indicator A x \M)" by (auto simp add: eq mult.commute) also have "... = - \ * measure M A" by auto finally have "0 \ - \ * measure M A" by simp then have "measure M A = 0" using \\ > 0\ by (simp add: measure_le_0_iff mult_le_0_iff) then have "A \ null_sets M" by (simp add: emeasure_eq_measure null_setsI) then have "AE x in M. x \ space M - A" by (metis (no_types, lifting) AE_cong Diff_iff AE_not_in) moreover { fix x assume "x \ space M - A" then have "limsup (\n. ereal(birkhoff_sum g n x)) < \" unfolding A_def by auto then obtain C where C: "\n. birkhoff_sum g n x \ C" using limsup_finite_then_bounded by presburger { fix n::nat assume "n > 0" have "birkhoff_sum g n x = birkhoff_sum f n x - birkhoff_sum (real_cond_exp M Invariants f) n x - birkhoff_sum (\x. \) n x" unfolding g_def using birkhoff_sum_add birkhoff_sum_diff by auto moreover have "birkhoff_sum (real_cond_exp M Invariants f) n x = n * real_cond_exp M Invariants f x" using birkhoff_sum_of_invariants using \x \ space M - A\ by auto moreover have "birkhoff_sum (\x. \) n x = n * \" unfolding birkhoff_sum_def by auto ultimately have "birkhoff_sum g n x = birkhoff_sum f n x - n * real_cond_exp M Invariants f x - n * \" by simp then have "birkhoff_sum f n x = birkhoff_sum g n x + n * real_cond_exp M Invariants f x + n * \" by simp then have "birkhoff_sum f n x / n = birkhoff_sum g n x / n + real_cond_exp M Invariants f x + \" using \n > 0\ by (simp add: field_simps) then have "birkhoff_sum f n x / n \ C/n + real_cond_exp M Invariants f x + \" using C[of n] \n > 0\ by (simp add: divide_right_mono) then have "ereal(birkhoff_sum f n x / n) \ ereal(C/n + real_cond_exp M Invariants f x + \)" by simp } then have "eventually (\n. ereal(birkhoff_sum f n x / n) \ ereal(C/n + real_cond_exp M Invariants f x + \)) sequentially" by (simp add: eventually_at_top_dense) then have b: "limsup (\n. ereal(birkhoff_sum f n x / n)) \ limsup (\n. ereal(C/n + real_cond_exp M Invariants f x + \))" by (simp add: Limsup_mono) have "(\n. ereal(C*(1/real n) + real_cond_exp M Invariants f x + \)) \ ereal(C * 0 + real_cond_exp M Invariants f x + \)" by (intro tendsto_intros) then have "limsup (\n. ereal(C/real n + real_cond_exp M Invariants f x + \)) = real_cond_exp M Invariants f x + \" using sequentially_bot tendsto_iff_Liminf_eq_Limsup by force then have "limsup (\n. ereal(birkhoff_sum f n x / n)) \ real_cond_exp M Invariants f x + \" using b by simp } ultimately have "AE x in M. limsup (\n. ereal(birkhoff_sum f n x / n)) \ real_cond_exp M Invariants f x + \" by auto then have "AE x in M. limsup (\n. ereal(birkhoff_sum f n x / n)) \ ereal(real_cond_exp M Invariants f x) + \" by auto } then show ?thesis by (rule AE_upper_bound_inf_ereal) qed theorem birkhoff_theorem_AE_nonergodic: fixes f::"'a \ real" assumes "integrable M f" shows "AE x in M. (\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" proof - { fix x assume i: "limsup (\n. ereal(birkhoff_sum f n x /n)) \ real_cond_exp M Invariants f x" and ii: "limsup (\n. ereal(birkhoff_sum (\x. -f x) n x / n)) \ real_cond_exp M Invariants (\x. -f x) x" and iii: "real_cond_exp M Invariants (\x. -f x) x = - real_cond_exp M Invariants f x" have "\n. birkhoff_sum (\x. -f x) n x = - birkhoff_sum f n x" using birkhoff_sum_cmult[where ?c = "-1" and ?f = f] by auto then have "\n. ereal(birkhoff_sum (\x. -f x) n x / n) = - ereal(birkhoff_sum f n x / n)" by auto moreover have "limsup (\n. - ereal(birkhoff_sum f n x / n)) = - liminf (\n. ereal(birkhoff_sum f n x /n))" by (rule ereal_Limsup_uminus) ultimately have "-liminf (\n. ereal(birkhoff_sum f n x /n)) = limsup (\n. ereal(birkhoff_sum (\x. -f x) n x / n))" by simp then have "-liminf (\n. ereal(birkhoff_sum f n x /n)) \ - real_cond_exp M Invariants f x" using ii iii by simp then have "liminf (\n. ereal(birkhoff_sum f n x /n)) \ real_cond_exp M Invariants f x" by (simp add: ereal_uminus_le_reorder) then have "(\n. birkhoff_sum f n x /n) \ real_cond_exp M Invariants f x" using i by (simp add: limsup_le_liminf_real) } note * = this moreover have "AE x in M. limsup (\n. ereal(birkhoff_sum f n x /n)) \ real_cond_exp M Invariants f x" using birkhoff_aux2 assms by simp moreover have "AE x in M. limsup (\n. ereal(birkhoff_sum (\x. -f x) n x / n)) \ real_cond_exp M Invariants (\x. -f x) x" using birkhoff_aux2 assms by simp moreover have "AE x in M. real_cond_exp M Invariants (\x. -f x) x = - real_cond_exp M Invariants f x" using real_cond_exp_cmult[where ?c = "-1"] assms by force ultimately show ?thesis by auto qed text \If a function $f$ is integrable, then $E(f\circ T - f | I) = E(f\circ T | I) - E(f|I) = 0$. Hence, $S_n(f \circ T - f) / n$ converges almost everywhere to $0$, i.e., $f(T^n x)/n \to 0$. It is remarkable (and sometimes useful) that this holds under the weaker condition that $f\circ T - f$ is integrable (but not necessarily $f$), where this naive argument fails. The reason is that the Birkhoff sum of $f \circ T - f$ is $f\circ T^n - f$. If $n$ is such that $x$ and $T^n(x)$ belong to a set where $f$ is bounded, it follows that this Birkhoff sum is also bounded. Along such a sequence of times, $S_n(f\circ T - f)/n$ tends to $0$. By Poincare recurrence theorem, there are such times for almost every points. As it also converges to $E(f \circ T - f | I)$, it follows that this function is almost everywhere $0$. Then $f (T^n x)/n = S_n(f\circ T^n - f)/n - f/n$ tends almost surely to $E(f\circ T-f |I) = 0$. \ lemma limit_foTn_over_n: fixes f::"'a \ real" assumes [measurable]: "f \ borel_measurable M" and "integrable M (\x. f(T x) - f x)" shows "AE x in M. real_cond_exp M Invariants (\x. f(T x) - f x) x = 0" "AE x in M. (\n. f((T^^n) x) / n) \ 0" proof - define E::"nat \ 'a set" where "E k = {x \ space M. \f x\ \ k}" for k have [measurable]: "E k \ sets M" for k unfolding E_def by auto have *: "(\k. E k) = space M" unfolding E_def by (auto simp add: real_arch_simple) define F::"nat \ 'a set" where "F k = recurrent_subset_infty (E k)" for k have [measurable]: "F k \ sets M" for k unfolding F_def by auto have **: "E k - F k \ null_sets M" for k unfolding F_def using Poincare_recurrence_thm by auto have "space M - (\k. F k) \ null_sets M" apply (rule null_sets_subset[of "(\k. E k - F k)"]) unfolding *[symmetric] using ** by auto with AE_not_in[OF this] have "AE x in M. x \ (\k. F k)" by auto moreover have "AE x in M. (\n. birkhoff_sum (\x. f(T x) - f x) n x / n) \ real_cond_exp M Invariants (\x. f(T x) - f x) x" by (rule birkhoff_theorem_AE_nonergodic[OF assms(2)]) moreover have "real_cond_exp M Invariants (\x. f(T x) - f x) x = 0 \ (\n. f((T^^n) x) / n) \ 0" if H: "(\n. birkhoff_sum (\x. f(T x) - f x) n x / n) \ real_cond_exp M Invariants (\x. f(T x) - f x) x" "x \ (\k. F k)" for x proof - have "f((T^^n) x) = birkhoff_sum (\x. f(T x) - f x) n x + f x" for n unfolding birkhoff_sum_def by (induction n, auto) then have "f((T^^n) x) / n = birkhoff_sum (\x. f(T x) - f x) n x / n + f x * (1/n)" for n by (auto simp add: divide_simps) moreover have "(\n. birkhoff_sum (\x. f(T x) - f x) n x / n + f x * (1/n)) \ real_cond_exp M Invariants (\x. f(T x) - f x) x + f x * 0" by (intro tendsto_intros H(1)) ultimately have lim: "(\n. f((T^^n) x) / n) \ real_cond_exp M Invariants (\x. f(T x) - f x) x" by auto obtain k where "x \ F k" using H(2) by auto then have "infinite {n. (T^^n) x \ E k}" unfolding F_def recurrent_subset_infty_inf_returns by auto with infinite_enumerate[OF this] obtain r :: "nat \ nat" where r: "strict_mono r" "\n. r n \ {n. (T^^n) x \ E k}" by auto have A: "(\n. k * (1/r n)) \ real k * 0" apply (intro tendsto_intros) using LIMSEQ_subseq_LIMSEQ[OF lim_1_over_n \strict_mono r\] unfolding comp_def by auto have B: "\f((T^^(r n)) x) / r n\ \ k / (r n)" for n using r(2) unfolding E_def by (auto simp add: divide_simps) have "(\n. f((T^^(r n)) x) / r n) \ 0" apply (rule tendsto_rabs_zero_cancel, rule tendsto_sandwich[of "\n. 0" _ _ "\n. k * (1/r n)"]) using A B by auto moreover have "(\n. f((T^^(r n)) x) / r n) \ real_cond_exp M Invariants (\x. f(T x) - f x) x" using LIMSEQ_subseq_LIMSEQ[OF lim \strict_mono r\] unfolding comp_def by auto ultimately have *: "real_cond_exp M Invariants (\x. f(T x) - f x) x = 0" using LIMSEQ_unique by auto then have "(\n. f((T^^n) x) / n) \ 0" using lim by auto then show ?thesis using * by auto qed ultimately show "AE x in M. real_cond_exp M Invariants (\x. f(T x) - f x) x = 0" "AE x in M. (\n. f((T^^n) x) / n) \ 0" by auto qed text \We specialize the previous statement to the case where $f$ itself is integrable.\ lemma limit_foTn_over_n': fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "AE x in M. (\n. f((T^^n) x) / n) \ 0" by (rule limit_foTn_over_n, simp, rule Bochner_Integration.integrable_diff) (auto intro: assms T_integral_preserving(1)) text \It is often useful to show that a function is cohomologous to a nicer function, i.e., to prove that a given $f$ can be written as $f = g + u - u \circ T$ where $g$ is nicer than $f$. We show below that any integrable function is cohomologous to a function which is arbitrarily close to $E(f|I)$. This is an improved version of Lemma 2.1 in [Benoist-Quint, Annals of maths, 2011]. Note that the function $g$ to which $f$ is cohomologous is very nice (and, in particular, integrable), but the transfer function is only measurable in this argument. The fact that the control on conditional expectation is nevertheless preserved throughout the argument follows from Lemma~\verb+limit_foTn_over_n+ above.\ text \We start with the lemma (and the proof) of [BQ2011]. It shows that, if a function has a conditional expectation with respect to invariants which is positive, then it is cohomologous to a nonnegative function. The argument is the clever remark that $g = \max (0, \inf_n S_n f)$ and $u = \min (0, \inf_n S_n f)$ work (where these expressions are well defined as $S_n f$ tends to infinity thanks to our assumption).\ lemma cohomologous_approx_cond_exp_aux: fixes f::"'a \ real" assumes [measurable]: "integrable M f" and "AE x in M. real_cond_exp M Invariants f x > 0" shows "\u g. u \ borel_measurable M \ (integrable M g) \ (AE x in M. g x \ 0 \ g x \ max 0 (f x)) \ (\x. f x = g x + u x - u (T x))" proof - define h::"'a \ real" where "h = (\x. (INF n\{1..}. birkhoff_sum f n x))" define u where "u = (\x. min (h x) 0)" define g where "g = (\x. f x - u x + u (T x))" have [measurable]: "h \ borel_measurable M" "u \ borel_measurable M" "g \ borel_measurable M" unfolding g_def h_def u_def by auto have "f x = g x + u x - u (T x)" for x unfolding g_def by auto { fix x assume H: "real_cond_exp M Invariants f x > 0" "(\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" have "eventually (\n. ereal(birkhoff_sum f n x / n) * ereal n = ereal(birkhoff_sum f n x)) sequentially" unfolding eventually_sequentially by (rule exI[of _ 1], auto) moreover have "(\n. ereal(birkhoff_sum f n x / n) * ereal n) \ ereal(real_cond_exp M Invariants f x) * \" apply (intro tendsto_intros) using H by auto ultimately have "(\n. ereal(birkhoff_sum f n x)) \ ereal(real_cond_exp M Invariants f x) * \" by (blast intro: Lim_transform_eventually) then have "(\n. ereal(birkhoff_sum f n x)) \ \" using H by auto then have B: "\C. \n. C \ birkhoff_sum f n x" by (intro liminf_finite_then_bounded_below, simp add: liminf_PInfty) have "h x \ f x" unfolding h_def apply (rule cInf_lower) using B by force+ have "{birkhoff_sum f n (T x) |n. n \ {1..}} = {birkhoff_sum f (1+n) (x) - f x |n. n \ {1..}}" unfolding birkhoff_sum_cocycle by auto also have "... = {birkhoff_sum f n x - f x |n. n \ {2..}}" by (metis (no_types, hide_lams) Suc_1 Suc_eq_plus1_left Suc_le_D Suc_le_mono atLeast_iff) finally have *: "{birkhoff_sum f n (T x) |n. n \ {1..}} = (\t. t - (f x))`{birkhoff_sum f n x |n. n \ {2..}}" by auto have "h(T x) = Inf {birkhoff_sum f n (T x) |n. n \ {1..}}" unfolding h_def by (metis Setcompr_eq_image) also have "... = (\t\{birkhoff_sum f n x |n. n \ {2..}}. t - f x)" by (simp only: *) also have "... = (\t. t - (f x)) (Inf {birkhoff_sum f n x |n. n \ {2..}})" using B by (auto intro!: monoI bijI mono_bij_cInf [symmetric]) finally have I: "Inf {birkhoff_sum f n x |n. n \ {2..}} = f x + h (T x)" by auto have "max 0 (h x) + u x = h x" unfolding u_def by auto also have "... = Inf {birkhoff_sum f n x |n. n \ {1..}}" unfolding h_def by (metis Setcompr_eq_image) also have "... = Inf ({birkhoff_sum f n x |n. n \ {1}} \ {birkhoff_sum f n x |n. n \ {2..}})" by (auto intro!: arg_cong[of _ _ Inf], metis One_nat_def Suc_1 antisym birkhoff_sum_1(2) not_less_eq_eq, force) also have "Inf ({birkhoff_sum f n x |n. n \ {1}} \ {birkhoff_sum f n x |n. n \ {2..}}) = min (Inf {birkhoff_sum f n x |n. n \ {1}}) (Inf {birkhoff_sum f n x |n. n \ {2..}})" unfolding inf_min[symmetric] apply (intro cInf_union_distrib) using B by auto also have "... = min (f x) (f x + h (T x))" using I by auto also have "... = f x + u (T x)" unfolding u_def by auto finally have "max 0 (h x) = f x + u (T x) - u x" by auto then have "g x = max 0 (h x)" unfolding g_def by auto then have "g x \ 0 \ g x \ max 0 (f x)" using \h x \ f x\ by auto } then have *: "AE x in M. g x \ 0 \ g x \ max 0 (f x)" using assms(2) birkhoff_theorem_AE_nonergodic[OF assms(1)] by auto moreover have "integrable M g" apply (rule Bochner_Integration.integrable_bound[of _ f]) using * by (auto simp add: assms) ultimately have "u \ borel_measurable M \ integrable M g \ (AE x in M. 0 \ g x \ g x \ max 0 (f x)) \ (\x. f x = g x + u x - u (T x))" using \\x. f x = g x + u x - u (T x)\ \u \ borel_measurable M\ by auto then show ?thesis by blast qed text \To deduce the stronger version that $f$ is cohomologous to an arbitrarily good approximation of $E(f|I)$, we apply the previous lemma twice, to control successively the negative and the positive side. The sign control in the conclusion of the previous lemma implies that the second step does not spoil the first one.\ lemma cohomologous_approx_cond_exp: fixes f::"'a \ real" and B::"'a \ real" assumes [measurable]: "integrable M f" "B \ borel_measurable M" and "AE x in M. B x > 0" shows "\g u. u \ borel_measurable M \ integrable M g \ (\x. f x = g x + u x - u (T x)) \ (AE x in M. abs(g x - real_cond_exp M Invariants f x) \ B x)" proof - define C where "C = (\x. min (B x) 1)" have [measurable]: "integrable M C" apply (rule Bochner_Integration.integrable_bound[of _ "\_. (1::real)"], auto) unfolding C_def using assms(3) by auto have "C x \ B x" for x unfolding C_def by auto have "AE x in M. C x > 0" unfolding C_def using assms(3) by auto have AECI: "AE x in M. real_cond_exp M Invariants C x > 0" by (intro real_cond_exp_gr_c \integrable M C\ \AE x in M. C x > 0\) define f1 where "f1 = (\x. f x - real_cond_exp M Invariants f x)" have "integrable M f1" unfolding f1_def by (intro Bochner_Integration.integrable_diff \integrable M f\ real_cond_exp_int(1)) have "AE x in M. real_cond_exp M Invariants f1 x = real_cond_exp M Invariants f x - real_cond_exp M Invariants (real_cond_exp M Invariants f) x" unfolding f1_def apply (rule real_cond_exp_diff) by (intro Bochner_Integration.integrable_diff \integrable M f\ \integrable M C\ real_cond_exp_int(1))+ moreover have "AE x in M. real_cond_exp M Invariants (real_cond_exp M Invariants f) x = real_cond_exp M Invariants f x" by (intro real_cond_exp_nested_subalg subalg \integrable M f\, auto) ultimately have AEf1: "AE x in M. real_cond_exp M Invariants f1 x = 0" by auto have A [measurable]: "integrable M (\x. f1 x + C x)" by (intro Bochner_Integration.integrable_add \integrable M f1\ \integrable M C\) have "AE x in M. real_cond_exp M Invariants (\x. f1 x + C x) x = real_cond_exp M Invariants f1 x + real_cond_exp M Invariants C x" by (intro real_cond_exp_add \integrable M f1\ \integrable M C\) then have B: "AE x in M. real_cond_exp M Invariants (\x. f1 x + C x) x > 0" using AECI AEf1 by auto obtain u2 g2 where H2: "u2 \ borel_measurable M" "integrable M g2" "AE x in M. g2 x \ 0 \ g2 x \ max 0 (f1 x + C x)" "\x. f1 x + C x = g2 x + u2 x - u2 (T x)" using cohomologous_approx_cond_exp_aux[OF A B] by blast define f2 where "f2 = (\x. (g2 x - C x))" have *: "u2(T x) - u2 x = f2 x -f1 x" for x unfolding f2_def using H2(4)[of x] by auto have "AE x in M. f2 x \ - C x" using H2(3) unfolding f2_def by auto have "integrable M f2" unfolding f2_def by (intro Bochner_Integration.integrable_diff \integrable M g2\ \integrable M C\) have "AE x in M. real_cond_exp M Invariants (\x. u2(T x) - u2 x) x = 0" proof (rule limit_foTn_over_n) show "integrable M (\x. u2(T x) - u2 x)" unfolding * by (intro Bochner_Integration.integrable_diff \integrable M f1\ \integrable M f2\) qed (simp add: \u2 \ borel_measurable M\) then have "AE x in M. real_cond_exp M Invariants (\x. f2 x - f1 x) x = 0" unfolding * by simp moreover have "AE x in M. real_cond_exp M Invariants (\x. f2 x - f1 x) x = real_cond_exp M Invariants f2 x - real_cond_exp M Invariants f1 x" by (intro real_cond_exp_diff \integrable M f2\ \integrable M f1\) ultimately have AEf2: "AE x in M. real_cond_exp M Invariants f2 x = 0" using AEf1 by auto have A [measurable]: "integrable M (\x. C x - f2 x)" by (intro Bochner_Integration.integrable_diff \integrable M f2\ \integrable M C\) have "AE x in M. real_cond_exp M Invariants (\x. C x - f2 x) x = real_cond_exp M Invariants C x - real_cond_exp M Invariants f2 x" by (intro real_cond_exp_diff \integrable M f2\ \integrable M C\) then have B: "AE x in M. real_cond_exp M Invariants (\x. C x - f2 x) x > 0" using AECI AEf2 by auto obtain u3 g3 where H3: "u3 \ borel_measurable M" "integrable M g3" "AE x in M. g3 x \ 0 \ g3 x \ max 0 (C x - f2 x)" "\x. C x - f2 x = g3 x + u3 x - u3 (T x)" using cohomologous_approx_cond_exp_aux[OF A B] by blast define f3 where "f3 = (\x. C x - g3 x)" have "AE x in M. f3 x \ min (C x) (f2 x)" unfolding f3_def using H3(3) by auto then have "AE x in M. f3 x \ -C x" using \AE x in M. f2 x \ - C x\ \AE x in M. C x > 0\ by auto moreover have "AE x in M. f3 x \ C x" unfolding f3_def using H3(3) by auto ultimately have "AE x in M. abs(f3 x) \ C x" by auto then have *: "AE x in M. abs(f3 x) \ B x" using order_trans[OF _ \\x. C x \ B x\] by auto define g where "g = (\x. f3 x + real_cond_exp M Invariants f x)" define u where "u = (\x. u2 x - u3 x)" have "AE x in M. abs (g x - real_cond_exp M Invariants f x) \ B x" unfolding g_def using * by auto moreover have "f x = g x + u x - u(T x)" for x using H3(4)[of x] H2(4)[of x] unfolding u_def g_def f3_def f2_def f1_def by auto moreover have "u \ borel_measurable M" unfolding u_def using \u2 \ borel_measurable M\ \u3 \ borel_measurable M\ by auto moreover have "integrable M g" unfolding g_def f3_def by (intro Bochner_Integration.integrable_add Bochner_Integration.integrable_diff \integrable M C\ \integrable M g3\ \integrable M f\ real_cond_exp_int(1)) ultimately show ?thesis by auto qed subsubsection \$L^1$ version of Birkhoff theorem\ text \The $L^1$ convergence in Birkhoff theorem follows from the almost everywhere convergence and general considerations on $L^1$ convergence (Scheffe's lemma) explained in \verb+AE_and_int_bound_implies_L1_conv2+. This argument works neatly for nonnegative functions, the general case reduces to this one by taking the positive and negative parts of a given function. One could also prove it by truncation: for bounded functions, the $L^1$ convergence follows from the boundedness and almost sure convergence. The general case follows by density, but it is a little bit tedious to write as one need to make sure that the conditional expectation of the truncation converges to the conditional expectation of the original function. This is true in $L^1$ as the conditional expectation is a contraction in $L^1$, it follows almost everywhere after taking a subsequence. All in all, the argument based on Scheffe's lemma seems more economical.\ lemma birkhoff_lemma_L1: fixes f::"'a \ real" assumes "\x. f x \ 0" and [measurable]: "integrable M f" shows "(\n. \\<^sup>+x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ 0" proof (rule Scheffe_lemma2) show i: "integrable M (real_cond_exp M Invariants f)" using assms by (simp add: real_cond_exp_int(1)) show "AE x in M. (\n. birkhoff_sum f n x / real n) \ real_cond_exp M Invariants f x" using birkhoff_theorem_AE_nonergodic assms by simp fix n have [measurable]: "(\x. ennreal \birkhoff_sum f n x\) \ borel_measurable M" by measurable show [measurable]: "(\x. birkhoff_sum f n x / real n) \ borel_measurable M" by measurable have "AE x in M. real_cond_exp M Invariants f x \ 0" using assms(1) real_cond_exp_pos by simp then have *: "AE x in M. norm (real_cond_exp M Invariants f x) = real_cond_exp M Invariants f x" by auto have **: "(\ x. norm (real_cond_exp M Invariants f x) \M) = (\ x. real_cond_exp M Invariants f x \M)" apply (rule integral_cong_AE) using * by auto have "(\\<^sup>+x. ennreal (norm (real_cond_exp M Invariants f x)) \M) = (\ x. norm (real_cond_exp M Invariants f x) \M)" by (rule nn_integral_eq_integral) (auto simp add: i) also have "... = (\ x. real_cond_exp M Invariants f x \M)" using ** by simp also have "... = (\ x. f x \M)" using real_cond_exp_int(2) assms(2) by auto also have "... = (\x. norm(f x) \M)" using assms by auto also have "... = (\\<^sup>+x. norm(f x) \M)" by (rule nn_integral_eq_integral[symmetric], auto simp add: assms(2)) finally have eq: "(\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M) = (\\<^sup>+ x. norm(f x) \M)" by simp { fix x have "norm(birkhoff_sum f n x) \ birkhoff_sum (\x. norm(f x)) n x" using birkhoff_sum_abs by fastforce then have "norm(birkhoff_sum f n x) \ birkhoff_sum (\x. ennreal(norm(f x))) n x" unfolding birkhoff_sum_def by auto } then have "(\\<^sup>+x. norm(birkhoff_sum f n x) \M) \ (\\<^sup>+x. birkhoff_sum (\x. ennreal(norm(f x))) n x \M)" by (simp add: nn_integral_mono) also have "... = n * (\\<^sup>+x. norm(f x) \M)" by (rule birkhoff_sum_nn_integral, auto) also have "... = n * (\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M)" using eq by simp finally have *: "(\\<^sup>+x. norm(birkhoff_sum f n x) \M) \ n * (\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M)" by simp show "(\\<^sup>+ x. ennreal (norm (birkhoff_sum f n x / real n)) \M) \ (\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M)" proof (cases) assume "n = 0" then show ?thesis by auto next assume "\(n = 0)" then have "n > 0" by simp then have "1/ennreal(real n) \ 0" by simp have "(\\<^sup>+ x. ennreal (norm (birkhoff_sum f n x / real n)) \M) = (\\<^sup>+ x. ennreal (norm (birkhoff_sum f n x)) / ennreal(real n) \M)" using \n > 0\ by (auto simp: divide_ennreal) also have "... = (\\<^sup>+ x. (1/ennreal(real n)) * ennreal (norm (birkhoff_sum f n x)) \M)" by (simp add: \0 < n\ divide_ennreal_def mult.commute) also have "... = (1/ennreal(real n) * (\\<^sup>+ x. ennreal (norm (birkhoff_sum f n x)) \M))" by (subst nn_integral_cmult) auto also have "... \ (1/ennreal(real n)) * (ennreal(real n) * (\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M))" using * by (intro mult_mono) (auto simp: ennreal_of_nat_eq_real_of_nat) also have "... = (\\<^sup>+ x. norm (real_cond_exp M Invariants f x) \M)" using \n > 0\ by (auto simp del: ennreal_1 simp add: ennreal_1[symmetric] divide_ennreal ennreal_mult[symmetric] mult.assoc[symmetric]) simp finally show ?thesis by simp qed qed theorem birkhoff_theorem_L1_nonergodic: fixes f::"'a \ real" assumes [measurable]: "integrable M f" shows "(\n. \\<^sup>+x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ 0" proof - define g where "g = (\x. max (f x) 0)" have g_int [measurable]: "integrable M g" unfolding g_def using assms by auto define h where "h = (\x. max (-f x) 0)" have h_int [measurable]: "integrable M h" unfolding h_def using assms by auto have "f = (\x. g x - h x)" unfolding g_def h_def by auto { fix n::nat assume "n > 0" have "\x. birkhoff_sum f n x = birkhoff_sum g n x - birkhoff_sum h n x" using birkhoff_sum_diff \f = (\x. g x - h x)\ by auto then have "\x. birkhoff_sum f n x / n = birkhoff_sum g n x / n - birkhoff_sum h n x / n" using \n > 0\ by (simp add: diff_divide_distrib) moreover have "AE x in M. real_cond_exp M Invariants g x - real_cond_exp M Invariants h x = real_cond_exp M Invariants f x" using AE_symmetric[OF real_cond_exp_diff] g_int h_int \f = (\x. g x - h x)\ by auto ultimately have "AE x in M. birkhoff_sum f n x / n - real_cond_exp M Invariants f x = (birkhoff_sum g n x / n - real_cond_exp M Invariants g x) - (birkhoff_sum h n x / n - real_cond_exp M Invariants h x)" by auto then have *: "AE x in M. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \ norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) + norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x)" by auto have "(\\<^sup>+ x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ (\\<^sup>+ x. ennreal(norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x)) + norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)" apply (rule nn_integral_mono_AE) using * by (simp add: ennreal_plus[symmetric] del: ennreal_plus) also have "... = (\\<^sup>+ x. norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) \M) + (\\<^sup>+ x. norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)" apply (rule nn_integral_add) apply auto using real_cond_exp_F_meas borel_measurable_cond_exp2 by measurable finally have "(\\<^sup>+ x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ (\\<^sup>+ x. norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) \M) + (\\<^sup>+ x. norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)" by simp } then have *: "eventually (\n. (\\<^sup>+ x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ (\\<^sup>+ x. norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) \M) + (\\<^sup>+ x. norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)) sequentially" using eventually_at_top_dense by auto have **: "eventually (\n. (\\<^sup>+ x. norm(birkhoff_sum f n x / n - real_cond_exp M Invariants f x) \M) \ 0) sequentially" by simp have "(\n. (\\<^sup>+ x. norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) \M)) \ 0" apply (rule birkhoff_lemma_L1, auto simp add: g_int) unfolding g_def by auto moreover have "(\n. (\\<^sup>+ x. norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)) \ 0" apply (rule birkhoff_lemma_L1, auto simp add: h_int) unfolding h_def by auto ultimately have "(\n. (\\<^sup>+ x. norm(birkhoff_sum g n x / n - real_cond_exp M Invariants g x) \M) + (\\<^sup>+ x. norm(birkhoff_sum h n x / n - real_cond_exp M Invariants h x) \M)) \ 0" using tendsto_add[of _ 0 _ _ 0] by auto then show ?thesis using tendsto_sandwich[OF ** *] by auto qed subsubsection \Conservativity of skew products\ text \The behaviour of skew-products of the form $(x, y) \mapsto (Tx, y + f x)$ is directly related to Birkhoff theorem, as the iterates involve the Birkhoff sums in the fiber. Birkhoff theorem implies that such a skew product is conservative when the function $f$ has vanishing conditional expectation. To prove the theorem, assume by contradiction that a set $A$ with positive measure does not intersect its preimages. Replacing $A$ with a smaller set $C$, we can assume that $C$ is bounded in the $y$-direction, by a constant $N$, and also that all its nonempty vertical fibers, above the projection $Cx$, have a measure bounded from below. Then, by Birkhoff theorem, for any $r>0$, most of the first $n$ preimages of $C$ are contained in the set $\{|y| \leq r n+N\}$, of measure $O(r n)$. Hence, they can not be disjoint if $r < \mu(C)$. To make this argument rigorous, one should only consider the preimages whose $x$-component belongs to a set $Dx$ where the Birkhoff sums are small. This condition has a positive measure if $\mu(Cx) + \mu(Dx) > \mu(M)$, which one can ensure by taking $Dx$ large enough.\ theorem (in fmpt) skew_product_conservative: fixes f::"'a \ real" assumes [measurable]: "integrable M f" and "AE x in M. real_cond_exp M Invariants f x = 0" shows "conservative_mpt (M \\<^sub>M lborel) (\(x,y). (T x, y + f x))" proof (rule conservative_mptI) let ?TS = "(\(x,y). (T x, y + f x))" let ?MS = "M \\<^sub>M (lborel::real measure)" have f_meas [measurable]: "f \ borel_measurable M" by auto have "mpt M T" by (simp add: mpt_axioms) with mpt_skew_product_real[OF this f_meas] show "mpt ?MS ?TS" by simp then interpret TS: mpt ?MS ?TS by auto fix A::"('a \ real) set" assume A1 [measurable]: "A \ sets ?MS" and A2: "emeasure ?MS A > 0" have "A = (\N::nat. A \ {(x,y). abs(y) \ N})" by (auto simp add: real_arch_simple) then have *: "emeasure ?MS (\N::nat. A \ {(x,y). abs(y) \ N}) > 0" using A2 by simp have "space ?MS = space M \ space (lborel::real measure)" using space_pair_measure by auto then have A_inc: "A \ space M \ space (lborel::real measure)" using sets.sets_into_space[OF A1] by auto { fix N::nat have "{(x, y). abs(y) \ real N \ x \ space M} = space M \ {-(real N)..(real N)}" by auto then have "{(x, y). \y\ \ real N \ x \ space M} \ sets ?MS" by auto then have "A \ {(x, y). \y\ \ real N \ x \ space M} \ sets ?MS" using A1 by auto moreover have "A \ {(x,y). abs(y) \ real N} = A \ {(x, y). \y\ \ real N \ x \ space M}" using A_inc by blast ultimately have "A \ {(x,y). abs(y) \ real N} \ sets ?MS" by auto } then have [measurable]: "\N::nat. A \ {(x, y). \y\ \ real N} \ sets (M \\<^sub>M borel)" by auto have "\N::nat. emeasure ?MS (A \ {(x,y). abs(y) \ N}) > 0" apply (rule emeasure_pos_unionE) using * by auto then obtain N::nat where N: "emeasure ?MS (A \ {(x,y). abs(y) \ N}) > 0" by auto define B where "B = A \ {(x,y). abs(y) \ N}" have B_meas [measurable]: "B \ sets (M \\<^sub>M lborel)" unfolding B_def by auto have "0 < emeasure (M \\<^sub>M lborel) B" unfolding B_def using N by auto also have "... = (\\<^sup>+x. emeasure lborel (Pair x -` B) \M)" apply (rule sigma_finite_measure.emeasure_pair_measure_alt) using B_meas by (auto simp add: lborel.sigma_finite_measure_axioms) finally have *: "(\\<^sup>+x. emeasure lborel (Pair x -` B) \M) > 0" by simp have "\Cx\sets M. \e::real>0. emeasure M Cx > 0 \ (\x \ Cx. emeasure lborel (Pair x -` B) \ e)" by (rule not_AE_zero_int_ennreal_E, auto simp add: *) then obtain Cx e where [measurable]: "Cx \ sets M" and Cxe: "e>(0::real)" "emeasure M Cx > 0" "\x. x \ Cx \ emeasure lborel (Pair x -` B) \ e" by blast define C where "C = B \ (Cx \ (UNIV::real set))" have C_meas [measurable]: "C \ sets (M \\<^sub>M lborel)" unfolding C_def using B_meas by auto have Cx_fibers: "\x. x \ Cx \ emeasure lborel (Pair x -` C) \ e" using Cxe(3) C_def by auto define c where "c = (measure M Cx)/2" have "c > 0" unfolding c_def using Cxe(2) by (simp add: emeasure_eq_measure) text \We will apply Birkhoff theorem to show that most preimages of $C$ at time $n$ are contained in a cylinder of height roughly $r n$, for some suitably small $r$. How small $r$ should be to get a contradiction can be determined at the end of the proof. It turns out that the good condition is the following one -- this is by no means obvious now.\ define r where "r = (if measure M (space M) = 0 then 1 else e * c / (4 * measure M (space M)))" have "r > 0" using \e > 0\ \c > 0\ unfolding r_def apply auto using measure_le_0_iff by fastforce have pos: "e*c-2*r*measure M (space M) > 0" using \e > 0\ \c > 0\ unfolding r_def by auto define Bgood where "Bgood = {x \ space M. (\n. birkhoff_sum f n x / n) \ 0}" have [measurable]: "Bgood \ sets M" unfolding Bgood_def by auto have *: "AE x in M. x \ Bgood" unfolding Bgood_def using birkhoff_theorem_AE_nonergodic[OF assms(1)] assms(2) by auto then have "emeasure M Bgood = emeasure M (space M)" by (intro emeasure_eq_AE) auto { fix x assume "x \ Bgood" then have "x \ space M" unfolding Bgood_def by auto have "(\n. birkhoff_sum f n x / n) \ 0" using \x \ Bgood\ unfolding Bgood_def by auto moreover have "0 \ {-r<..r>0\ by auto ultimately have "eventually (\n. birkhoff_sum f n x / n \ {-r<..0" "\n. n \ n0 \ birkhoff_sum f n x / n \ {-r<.. n0" then have "n>0" using \n0>0\ by auto with n0(2)[OF \n \ n0\] have "abs(birkhoff_sum f n x / n) \ r" by auto then have "abs(birkhoff_sum f n x) \ r * n" using \n>0\ by (simp add: divide_le_eq) } then have "x \ (\n0. {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n})" using \x \ space M\ by blast } then have "AE x in M. x \ space M - (\n0. {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n})" using * by auto then have eqM: "emeasure M (\n0. {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n}) = emeasure M (space M)" by (intro emeasure_eq_AE) auto have "(\n0. emeasure M {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n} + c) \ emeasure M (\n0. {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n}) + c" by (intro tendsto_intros Lim_emeasure_incseq) (auto simp add: incseq_def) moreover have "emeasure M (\n0. {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n}) + c > emeasure M (space M)" using eqM \c > 0\ emeasure_eq_measure by auto ultimately have "eventually (\n0. emeasure M {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n} + c > emeasure M (space M)) sequentially" unfolding order_tendsto_iff by auto then obtain n0 where n0: "emeasure M {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n} + c > emeasure M (space M)" using eventually_sequentially by auto define Dx where "Dx = {x \ space M. \n\{n0..}. abs(birkhoff_sum f n x) \ r * n}" have Dx_meas [measurable]: "Dx \ sets M" unfolding Dx_def by auto have "emeasure M Dx + c \ emeasure M (space M)" using n0 Dx_def by auto obtain n1::nat where n1: "n1 > max n0 ((measure M (space M) * 2 * N + e*c*n0 - e*c) / (e*c-2*r*measure M (space M)))" by (metis mult.commute mult.left_neutral numeral_One reals_Archimedean3 zero_less_numeral) then have "n1 > n0" by auto have n1_ineq: "n1 * (e*c-2*r*measure M (space M)) > (measure M (space M) * 2 * N + e*c*n0 - e*c)" using n1 pos by (simp add: pos_divide_less_eq) define D where "D = (\n. Dx \ {-r*n1-N..r*n1+N} \ (?TS^^n)-`C)" have Dn_meas [measurable]: "D n \ sets (M \\<^sub>M lborel)" for n unfolding D_def apply (rule TS.T_intersec_meas(2)) using C_meas by auto have "emeasure ?MS (D n) \ e * c" if "n \ {n0..n1}" for n proof - have "n \ n0" "n \ n1" using that by auto { fix x assume [simp]: "x \ space M" define F where "F = {y \ {-r*n1-N..r*n1+N}. y + birkhoff_sum f n x \ Pair ((T^^n)x) -`C}" have [measurable]: "F \ sets lborel" unfolding F_def by measurable { fix y::real have "(?TS^^n)(x,y) = ((T^^n)x, y + birkhoff_sum f n x)" using skew_product_real_iterates by simp then have "(indicator C ((?TS^^n) (x,y))::ennreal) = indicator Cx ((T^^n)x) * indicator (Pair ((T^^n)x) -`C) (y + birkhoff_sum f n x)" using C_def by (simp add: indicator_def) moreover have "(indicator (D n) (x, y)::ennreal) = indicator Dx x * indicator {-r*n1-N..r*n1+N} y * indicator C ((?TS^^n) (x,y))" unfolding D_def by (simp add: indicator_def) ultimately have "(indicator (D n) (x, y)::ennreal) = indicator Dx x * indicator {-r*n1-N..r*n1+N} y * indicator Cx ((T^^n)x) * indicator (Pair ((T^^n)x) -`C) (y + birkhoff_sum f n x)" by (simp add: mult.assoc) then have "(indicator (D n) (x, y)::ennreal) = indicator (Dx \ (T^^n)-`Cx) x * indicator F y" unfolding F_def by (simp add: indicator_def) } then have "(\\<^sup>+y. indicator (D n) (x, y) \lborel) = (\\<^sup>+y. indicator (Dx \ (T^^n)-`Cx) x * indicator F y \lborel)" by auto also have "... = indicator (Dx \ (T^^n)-`Cx) x * (\\<^sup>+y. indicator F y \lborel)" by (rule nn_integral_cmult, auto) also have "... = indicator (Dx \ (T^^n)-`Cx) x * emeasure lborel F" using \F \ sets lborel\ by auto finally have A: "(\\<^sup>+y. indicator (D n) (x, y) \lborel) = indicator (Dx \ (T^^n)-`Cx) x * emeasure lborel F" by simp have "(\\<^sup>+y. indicator (D n) (x, y) \lborel) \ ennreal e * indicator (Dx \ (T^^n)-`Cx) x" proof (cases) assume "indicator (Dx \ (T^^n)-`Cx) x = (0::ennreal)" then show ?thesis by auto next assume "\(indicator (Dx \ (T^^n)-`Cx) x = (0::ennreal))" then have "x \ Dx \ (T^^n)-`Cx" by (simp add: indicator_eq_0_iff) then have "x \ Dx" "(T^^n) x \ Cx" by auto then have "abs(birkhoff_sum f n x) \ r * n" using \n \ {n0..n1}\ Dx_def by auto then have *: "abs(birkhoff_sum f n x) \ r * n1" using \n \ n1\ \r>0\ by (meson of_nat_le_iff order_trans real_mult_le_cancel_iff2) have F_expr: "F = {-r*n1-N..r*n1+N} \ (+)(birkhoff_sum f n x) -` (Pair ((T^^n)x) -`C)" unfolding F_def by (auto simp add: add.commute) have "(Pair ((T^^n)x) -`C) \ {real_of_int (- int N)..real N}" unfolding C_def B_def by auto then have "((+)(birkhoff_sum f n x)) -` (Pair ((T^^n)x) -`C) \ {-N-birkhoff_sum f n x..N-birkhoff_sum f n x}" by auto also have "... \ {-r * n1 - N .. r * n1 + N}" using * by auto finally have "F = ((+)(birkhoff_sum f n x)) -` (Pair ((T^^n)x) -`C)" unfolding F_expr by auto then have "emeasure lborel F = emeasure lborel ((+)(birkhoff_sum f n x) -` (Pair ((T^^n)x) -`C))" by auto also have "... = emeasure lborel (((+)(birkhoff_sum f n x) -` (Pair ((T^^n)x) -`C)) \ space lborel)" by simp also have "... = emeasure (distr lborel borel ((+) (birkhoff_sum f n x))) (Pair ((T^^n)x) -`C)" apply (rule emeasure_distr[symmetric]) using C_meas by auto also have "... = emeasure lborel (Pair ((T^^n)x) -`C)" using lborel_distr_plus[of "birkhoff_sum f n x"] by simp also have "... \ e" using Cx_fibers \(T^^n) x \ Cx\ by auto finally have "emeasure lborel F \ e" by auto then show ?thesis using A by (simp add: indicator_def) qed } moreover have "emeasure ?MS (D n) = (\\<^sup>+x. (\\<^sup>+y. indicator (D n) (x, y) \lborel) \M)" using Dn_meas lborel.emeasure_pair_measure by blast ultimately have "emeasure ?MS (D n) \ (\\<^sup>+x. ennreal e * indicator (Dx \ (T ^^ n) -` Cx) x \M)" by (simp add: nn_integral_mono) also have "(\\<^sup>+x. ennreal e * indicator (Dx \ (T ^^ n) -` Cx) x \M) = e * (\\<^sup>+x. indicator (Dx \ (T ^^ n) -` Cx) x \M)" apply (rule nn_integral_cmult) using \e>0\ by auto also have "... = ennreal e * emeasure M (Dx \ (T ^^ n) -` Cx)" by simp finally have *: "emeasure ?MS (D n) \ ennreal e * emeasure M (Dx \ (T ^^ n) -` Cx)" by auto have "c + emeasure M (space M) \ emeasure M Dx + emeasure M Cx" using \emeasure M Dx + c \ emeasure M (space M)\ unfolding c_def by (auto simp: emeasure_eq_measure ennreal_plus[symmetric] simp del: ennreal_plus) also have "... = emeasure M Dx + emeasure M ((T^^n)--`Cx)" by (simp add: T_vrestr_same_emeasure(2)) also have "... = emeasure M (Dx \ ((T^^n)--`Cx)) + emeasure M (Dx \ ((T^^n)--`Cx))" - by (rule emeasure_union_inter, auto) + by (rule emeasure_Un_Int, auto) also have "... \ emeasure M (space M) + emeasure M (Dx \ ((T^^n)-`Cx))" proof - have "emeasure M (Dx \ ((T^^n)--`Cx)) \ emeasure M (space M)" by (rule emeasure_mono, auto simp add: sets.sets_into_space) moreover have "Dx \ ((T^^n)--`Cx) = Dx \ ((T^^n)-`Cx)" by (simp add: vrestr_intersec_in_space) ultimately show ?thesis by (metis add.commute add_left_mono) qed finally have "emeasure M (Dx \ ((T^^n)-`Cx)) \ c" by (simp add: emeasure_eq_measure) then have "ennreal e * emeasure M (Dx \ ((T^^n)-`Cx)) \ ennreal e * c" using \e > 0\ using mult_left_mono by fastforce with * show "emeasure ?MS (D n) \ e * c" using \0 \0 by (auto simp: ennreal_mult[symmetric]) qed have "\(disjoint_family_on D {n0..n1})" proof assume D: "disjoint_family_on D {n0..n1}" have "emeasure lborel {-r*n1-N..r*n1+N} = (r * real n1 + real N) - (- r * real n1 - real N)" apply (rule emeasure_lborel_Icc) using \r>0\ by auto then have *: "emeasure lborel {-r*n1-N..r*n1+N} = ennreal(2 * r * real n1 + 2 * real N)" by (auto simp: ac_simps) have "ennreal(e * c) * (real n1 - real n0 + 1) = ennreal(e*c) * card {n0..n1}" using \n1 > n0\ by (auto simp: ennreal_of_nat_eq_real_of_nat Suc_diff_le ac_simps of_nat_diff) also have "... = (\n\{n0..n1}. ennreal(e*c))" by (simp add: ac_simps) also have "... \ (\n\{n0..n1}. emeasure ?MS (D n))" using \\n. n \ {n0..n1} \ emeasure ?MS (D n) \ e * c\ by (meson sum_mono) also have "... = emeasure ?MS (\n\{n0..n1}. D n)" apply (rule sum_emeasure) using Dn_meas by (auto simp add: D) also have "... \ emeasure ?MS (space M \ {-r*n1-N..r*n1+N})" apply (rule emeasure_mono) unfolding D_def using sets.sets_into_space[OF Dx_meas] by auto also have "... = emeasure M (space M) * emeasure lborel {-r*n1-N..r*n1+N}" by (rule sigma_finite_measure.emeasure_pair_measure_Times, auto simp add: lborel.sigma_finite_measure_axioms) also have "... = emeasure M (space M) * ennreal(2 * r * real n1 + 2 * real N)" using * by auto finally have "ennreal(e * c) * (real n1- real n0+1) \ emeasure M (space M) * ennreal(2 * r * real n1 + 2 * real N)" by simp then have "e*c * (real n1- real n0 + 1) \ measure M (space M) * (2 * r * real n1 + 2 * real N)" using \0 \0 \0 \n0 < n1\ emeasure_eq_measure by (auto simp: ennreal_mult'[symmetric] simp del: ennreal_plus) then have "0 \ measure M (space M) * (2 * r * real n1 + 2 * real N) - e*c * (real n1- real n0 + 1)" by auto also have "... = (measure M (space M) * 2 * N + e*c*n0 - e*c) - n1 * (e*c-2*r*measure M (space M))" by algebra finally have "n1 * (e*c-2*r*measure M (space M)) \ measure M (space M) * 2 * N + e*c*n0 - e*c" by linarith then show False using n1_ineq by auto qed then obtain n m where nm: "n D n \ {}" unfolding disjoint_family_on_def by (metis inf_sup_aci(1) linorder_cases) define k where "k = m-n" then have "k>0" "D (n+k) \ D n \ {}" using nm by auto then have "((?TS^^(n+k))-`A) \ ((?TS^^n)-`A) \ {}" unfolding D_def C_def B_def by auto moreover have "((?TS^^(n+k))-`A) \ ((?TS^^n)-`A) = (?TS^^n)-`(((?TS^^k)-`A) \ A)" using funpow_add by (simp add: add.commute funpow_add set.compositionality) ultimately have "((?TS^^k)-`A) \ A \ {}" by auto then show "\k>0. ((?TS^^k)-`A) \ A \ {}" using \k>0\ by auto qed subsubsection \Oscillations around the limit in Birkhoff theorem\ text \In this paragraph, we prove that, in Birkhoff theorem with vanishing limit, the Birkhoff sums are infinitely many times arbitrarily close to $0$, both on the positive and the negative side. In the ergodic case, this statement implies for instance that if the Birkhoff sums of an integrable function tend to infinity almost everywhere, then the integral of the function can not vanish, it has to be strictly positive (while Birkhoff theorem per se does not exclude the convergence to infinity, at a rate slower than linear). This converts a qualitative information (convergence to infinity at an unknown rate) to a quantitative information (linear convergence to infinity). This result (sometimes known as Atkinson's Lemma) has been reinvented many times, for instance by Kesten and by Guivarch. It plays an important role in the study of random products of matrices. This is essentially a consequence of the conservativity of the corresponding skew-product, proved in \verb+skew_product_conservative+. Indeed, this implies that, starting from a small set $X \times [-e/2, e/2]$, the skew-product comes back infinitely often to itself, which implies that the Birkhoff sums at these times are bounded by $e$. To show that the Birkhoff sums come back to $[0,e]$ is a little bit more tricky. Argue by contradiction, and induce on $A \times [0,e/2]$ where $A$ is the set of points where the Birkhoff sums don't come back to $[0,e]$. Then the second coordinate decreases strictly when one iterates the skew product, which is not compatible with conservativity.\ lemma birkhoff_sum_small_asymp_lemma: assumes [measurable]: "integrable M f" and "AE x in M. real_cond_exp M Invariants f x = 0" "e>(0::real)" shows "AE x in M. infinite {n. birkhoff_sum f n x \ {0..e}}" proof - have [measurable]: "f \ borel_measurable M" by auto have [measurable]: "\N. {x \ space M. \N. \n\{N..}. birkhoff_sum f n x \ {0..e}} \ sets M" by auto { fix N assume "N>(0::nat)" define Ax where "Ax = {x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}}" then have [measurable]: "Ax \ sets M" by auto define A where "A = Ax \ {0..e/2}" then have A_meas [measurable]: "A \ sets (M \\<^sub>M lborel)" by auto define TN where "TN = T^^N" interpret TN: fmpt M TN unfolding TN_def using fmpt_power by auto define fN where "fN = birkhoff_sum f N" have "TN.birkhoff_sum fN n x = birkhoff_sum f (n*N) x" for n x proof (induction n) case 0 then show ?case by auto next case (Suc n) have "TN.birkhoff_sum fN (Suc n) x = TN.birkhoff_sum fN n x + fN ((TN^^n) x)" using TN.birkhoff_sum_cocycle[of fN n 1] by auto also have "... = birkhoff_sum f (n*N) x + birkhoff_sum f N ((TN^^n) x)" using Suc.IH fN_def by auto also have "... = birkhoff_sum f (n*N+N) x" unfolding TN_def by (subst funpow_mult, subst mult.commute[of N n], rule birkhoff_sum_cocycle[of f "n*N" N x, symmetric]) finally show ?case by (simp add: add.commute) qed then have not0e: "\x n. x \ Ax \ n \ 0 \ TN.birkhoff_sum fN n x \ {0..e}" unfolding Ax_def by auto let ?TS = "(\(x,y). (T x, y + f x))" let ?MS = "M \\<^sub>M (lborel::real measure)" interpret TS: conservative_mpt ?MS ?TS by (rule skew_product_conservative, auto simp add: assms) let ?TSN = "(\(x,y). (TN x, y + fN x))" have *:"?TSN = ?TS^^N" unfolding TN_def fN_def using skew_product_real_iterates by auto interpret TSN: conservative_mpt ?MS ?TSN using * TS.conservative_mpt_power by auto define MA TA where "MA = restrict_space ?MS A" and "TA = TSN.induced_map A" interpret TA: conservative_mpt MA TA unfolding MA_def TA_def by (rule TSN.induced_map_conservative_mpt, measurable) have *: "\ x y. snd (TA (x,y)) = snd (x,y) + TN.birkhoff_sum fN (TSN.return_time_function A (x,y)) x" unfolding TA_def TSN.induced_map_def using TN.skew_product_real_iterates Pair_def by auto have [measurable]: "snd \ borel_measurable ?MS" by auto then have [measurable]: "snd \ borel_measurable MA" unfolding MA_def using measurable_restrict_space1 by blast have "AE z in MA. z \ TSN.recurrent_subset A" unfolding MA_def using TSN.induced_map_recurrent_typical(1)[OF A_meas]. moreover { fix z assume z: "z \ TSN.recurrent_subset A" define x y where "x = fst z" and "y = snd z" then have "z = (x,y)" by simp have "z \ A" using z "TSN.recurrent_subset_incl"(1) by auto then have "x \ Ax" "y \ {0..e/2}" unfolding A_def x_def y_def by auto define y2 where "y2 = y + TN.birkhoff_sum fN (TSN.return_time_function A (x,y)) x" have "y2 = snd (TA z)" unfolding y2_def using * \z = (x, y)\ by force moreover have "TA z \ A" unfolding TA_def using \z \ A\ TSN.induced_map_stabilizes_A by auto ultimately have "y2 \ {0..e/2}" unfolding A_def by auto have "TSN.return_time_function A (x,y) \ 0" using \z = (x,y)\ \z \ TSN.recurrent_subset A\ TSN.return_time0[of A] by fast then have "TN.birkhoff_sum fN (TSN.return_time_function A (x,y)) x \ {0..e}" using not0e[OF \x \ Ax\] by auto moreover have "TN.birkhoff_sum fN (TSN.return_time_function A (x,y)) x \ {-e..e}" using \y \ {0..e/2}\ \y2 \ {0..e/2}\ y2_def by auto ultimately have "TN.birkhoff_sum fN (TSN.return_time_function A (x,y)) x \ {-e..<0}" by auto then have "y2 < y" using y2_def by auto then have "snd(TA z) < snd z" unfolding y_def using \y2 = snd (TA z)\ by auto } ultimately have a: "AE z in MA. snd(TA z) < snd z" by auto then have "AE z in MA. snd(TA z) \ snd z" by auto then have "AE z in MA. snd(TA z) = snd z" using TA.AE_decreasing_then_invariant[of snd] by auto then have "AE z in MA. False" using a by auto then have "space MA \ null_sets MA" by (simp add: AE_iff_null_sets) then have "emeasure MA A = 0" by (metis A_meas MA_def null_setsD1 space_restrict_space2) then have "emeasure ?MS A = 0" unfolding MA_def by (metis A_meas emeasure_restrict_space sets.sets_into_space sets.top space_restrict_space space_restrict_space2) moreover have "emeasure ?MS A = emeasure M Ax * emeasure lborel {0..e/2}" unfolding A_def by (intro lborel.emeasure_pair_measure_Times) auto ultimately have "emeasure M {x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}} = 0" using \e>0\ Ax_def by simp then have "{x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}} \ null_sets M" by auto } then have "(\N\{0<..}. {x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}}) \ null_sets M" by (auto simp: greaterThan_0) moreover have "{x \ space M. \(infinite {n. birkhoff_sum f n x \ {0..e}})} \ (\N\{0<..}. {x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}})" proof fix x assume H: "x \ {x \ space M. \(infinite {n. birkhoff_sum f n x \ {0..e}})}" then have "x \ space M" by auto have *: "finite {n. birkhoff_sum f n x \ {0..e}}" using H by auto then obtain N where "\n. n \ N \ n \ {n. birkhoff_sum f n x \ {0..e}}" by (metis finite_nat_set_iff_bounded not_less) then have "x \ {x \ space M. \n\{N+1..}. birkhoff_sum f n x \ {0..e}}" using \x \ space M\ by auto moreover have "N+1>0" by auto ultimately show "x \ (\N\{0<..}. {x \ space M. \n\{N..}. birkhoff_sum f n x \ {0..e}})" by auto qed ultimately show ?thesis unfolding eventually_ae_filter by auto qed theorem birkhoff_sum_small_asymp_pos_nonergodic: assumes [measurable]: "integrable M f" and "e > (0::real)" shows "AE x in M. infinite {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x .. n * real_cond_exp M Invariants f x + e}}" proof - define g where "g = (\x. f x - real_cond_exp M Invariants f x)" have g_meas [measurable]: "integrable M g" unfolding g_def using real_cond_exp_int(1)[OF assms(1)] assms(1) by auto have "AE x in M. real_cond_exp M Invariants (real_cond_exp M Invariants f) x = real_cond_exp M Invariants f x" by (rule real_cond_exp_F_meas, auto simp add: real_cond_exp_int(1)[OF assms(1)]) then have *: "AE x in M. real_cond_exp M Invariants g x = 0" unfolding g_def using real_cond_exp_diff[OF assms(1) real_cond_exp_int(1)[OF assms(1)]] by auto have "AE x in M. infinite {n. birkhoff_sum g n x \ {0..e}}" by (rule birkhoff_sum_small_asymp_lemma, auto simp add: \e>0\ * g_meas) moreover { fix x assume "x \ space M" "infinite {n. birkhoff_sum g n x \ {0..e}}" { fix n assume H: "birkhoff_sum g n x \ {0..e}" have "birkhoff_sum g n x = birkhoff_sum f n x - birkhoff_sum (real_cond_exp M Invariants f) n x" unfolding g_def using birkhoff_sum_diff by auto also have "... = birkhoff_sum f n x - n * real_cond_exp M Invariants f x" using birkhoff_sum_of_invariants \x \ space M\ by auto finally have "birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x .. n * real_cond_exp M Invariants f x + e}" using H by simp } then have "{n. birkhoff_sum g n x \ {0..e}} \ {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x .. n * real_cond_exp M Invariants f x + e}}" by auto then have "infinite {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x .. n * real_cond_exp M Invariants f x + e}}" using \infinite {n. birkhoff_sum g n x \ {0..e}}\ finite_subset by blast } ultimately show ?thesis by auto qed theorem birkhoff_sum_small_asymp_neg_nonergodic: assumes [measurable]: "integrable M f" and "e > (0::real)" shows "AE x in M. infinite {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x - e .. n * real_cond_exp M Invariants f x}}" proof - define g where "g = (\x. real_cond_exp M Invariants f x - f x)" have g_meas [measurable]: "integrable M g" unfolding g_def using real_cond_exp_int(1)[OF assms(1)] assms(1) by auto have "AE x in M. real_cond_exp M Invariants (real_cond_exp M Invariants f) x = real_cond_exp M Invariants f x" by (rule real_cond_exp_F_meas, auto simp add: real_cond_exp_int(1)[OF assms(1)]) then have *: "AE x in M. real_cond_exp M Invariants g x = 0" unfolding g_def using real_cond_exp_diff[OF real_cond_exp_int(1)[OF assms(1)] assms(1)] by auto have "AE x in M. infinite {n. birkhoff_sum g n x \ {0..e}}" by (rule birkhoff_sum_small_asymp_lemma, auto simp add: \e>0\ * g_meas) moreover { fix x assume "x \ space M" "infinite {n. birkhoff_sum g n x \ {0..e}}" { fix n assume H: "birkhoff_sum g n x \ {0..e}" have "birkhoff_sum g n x = birkhoff_sum (real_cond_exp M Invariants f) n x - birkhoff_sum f n x" unfolding g_def using birkhoff_sum_diff by auto also have "... = n * real_cond_exp M Invariants f x - birkhoff_sum f n x" using birkhoff_sum_of_invariants \x \ space M\ by auto finally have "birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x - e .. n * real_cond_exp M Invariants f x}" using H by simp } then have "{n. birkhoff_sum g n x \ {0..e}} \ {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x - e .. n * real_cond_exp M Invariants f x}}" by auto then have "infinite {n. birkhoff_sum f n x \ {n * real_cond_exp M Invariants f x - e .. n * real_cond_exp M Invariants f x}}" using \infinite {n. birkhoff_sum g n x \ {0..e}}\ finite_subset by blast } ultimately show ?thesis by auto qed subsubsection \Conditional expectation for the induced map\ text \Thanks to Birkhoff theorem, one can relate conditional expectations with respect to the invariant sigma algebra, for a map and for a corresponding induced map, as follows.\ proposition Invariants_cond_exp_induced_map: fixes f::"'a \ real" assumes [measurable]: "A \ sets M" "integrable M f" defines "MA \ restrict_space M A" and "TA \ induced_map A" and "fA \ induced_function A f" shows "AE x in MA. real_cond_exp MA (qmpt.Invariants MA TA) fA x = real_cond_exp M Invariants f x * real_cond_exp MA (qmpt.Invariants MA TA) (return_time_function A) x" proof - interpret A: fmpt MA TA unfolding MA_def TA_def by (rule induced_map_fmpt[OF assms(1)]) have "integrable M fA" unfolding fA_def using induced_function_integral_nonergodic(1) assms by auto then have "integrable MA fA" unfolding MA_def by (metis assms(1) integrable_mult_indicator integrable_restrict_space sets.Int_space_eq2) then have a: "AE x in MA. (\n. A.birkhoff_sum fA n x / n) \ real_cond_exp MA A.Invariants fA x" using A.birkhoff_theorem_AE_nonergodic by auto have "AE x in M. (\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" using birkhoff_theorem_AE_nonergodic assms(2) by auto then have b: "AE x in MA. (\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" unfolding MA_def by (metis (mono_tags, lifting) AE_restrict_space_iff assms(1) eventually_mono sets.Int_space_eq2) define phiA where "phiA = (\x. return_time_function A x)" have "integrable M phiA" unfolding phiA_def using return_time_integrable by auto then have "integrable MA phiA" unfolding MA_def by (metis assms(1) integrable_mult_indicator integrable_restrict_space sets.Int_space_eq2) then have c: "AE x in MA. (\n. A.birkhoff_sum (\x. real(phiA x)) n x / n) \ real_cond_exp MA A.Invariants phiA x" using A.birkhoff_theorem_AE_nonergodic by auto have "A-recurrent_subset A \ null_sets M" using Poincare_recurrence_thm(1)[OF assms(1)] by auto then have "A - recurrent_subset A \ null_sets MA" unfolding MA_def by (metis Diff_subset assms(1) emeasure_restrict_space null_setsD1 null_setsD2 null_setsI sets.Int_space_eq2 sets_restrict_space_iff) then have "AE x in MA. x \ recurrent_subset A" by (simp add: AE_iff_null MA_def null_setsD2 set_diff_eq space_restrict_space2) moreover have "\x. x \ recurrent_subset A \ phiA x > 0" unfolding phiA_def using return_time0 by fastforce ultimately have *: "AE x in MA. phiA x > 0" by auto have d: "AE x in MA. real_cond_exp MA A.Invariants phiA x > 0" by (rule A.real_cond_exp_gr_c, auto simp add: * \integrable MA phiA\) { fix x assume A: "(\n. A.birkhoff_sum fA n x / n) \ real_cond_exp MA A.Invariants fA x" and B: "(\n. birkhoff_sum f n x / n) \ real_cond_exp M Invariants f x" and C: "(\n. A.birkhoff_sum (\x. real(phiA x)) n x / n) \ real_cond_exp MA A.Invariants phiA x" and D: "real_cond_exp MA A.Invariants phiA x > 0" define R where "R = (\n. A.birkhoff_sum phiA n x)" have D2: "ereal(real_cond_exp MA A.Invariants phiA x) > 0" using D by simp have "\n. real(R n)/ n = A.birkhoff_sum (\x. real(phiA x)) n x / n" unfolding R_def A.birkhoff_sum_def by auto moreover have "(\n. A.birkhoff_sum (\x. real(phiA x)) n x / n) \ real_cond_exp MA A.Invariants phiA x" using C by auto ultimately have Rnn: "(\n. real(R n)/n) \ real_cond_exp MA A.Invariants phiA x" by presburger have "\n. ereal(real(R n))/ n = ereal(A.birkhoff_sum (\x. real(phiA x)) n x / n)" unfolding R_def A.birkhoff_sum_def by auto moreover have "(\n. ereal(A.birkhoff_sum (\x. real(phiA x)) n x / n)) \ real_cond_exp MA A.Invariants phiA x" using C by auto ultimately have i: "(\n. ereal(real(R n))/n) \ real_cond_exp MA A.Invariants phiA x" by auto have ii: "(\n. real n) \ \" by (rule id_nat_ereal_tendsto_PInf) have iii: "(\n. ereal(real(R n))/n * real n) \ \" using tendsto_mult_ereal_PInf[OF i D2 ii] by simp have "\n. n > 0 \ ereal(real(R n))/n * real n = R n" by auto then have "eventually (\n. ereal(real(R n))/n * real n = R n) sequentially" using eventually_at_top_dense by auto then have "(\n. real(R n)) \ \" using iii by (simp add: filterlim_cong) then have "(\n. birkhoff_sum f (R n) x / (R n)) \ real_cond_exp M Invariants f x" using limit_along_weak_subseq B by auto then have l: "(\n. (birkhoff_sum f (R n) x / (R n)) * ((R n)/n)) \ real_cond_exp M Invariants f x * real_cond_exp MA A.Invariants phiA x" by (rule tendsto_mult, simp add: Rnn) obtain N where N: "\n. n > N \ R n > 0" using \(\n. real(R n)) \ \\ by (metis (full_types) eventually_at_top_dense filterlim_iff filterlim_weak_subseq) then have "\n. n > N \ (birkhoff_sum f (R n) x / (R n)) * ((R n)/n) = birkhoff_sum f (R n) x / n" by auto then have "eventually (\n. (birkhoff_sum f (R n) x / (R n)) * ((R n)/n) = birkhoff_sum f (R n) x / n) sequentially" by simp with tendsto_cong[OF this] have main_limit: "(\n. birkhoff_sum f (R n) x / n) \ real_cond_exp M Invariants f x * real_cond_exp MA A.Invariants phiA x" using l by auto have "\n. birkhoff_sum f (R n) x = A.birkhoff_sum fA n x" unfolding R_def fA_def phiA_def TA_def using induced_function_birkhoff_sum[OF assms(1)] by simp then have "\n. birkhoff_sum f (R n) x /n = A.birkhoff_sum fA n x / n" by simp then have "(\n. A.birkhoff_sum fA n x / n) \ real_cond_exp M Invariants f x * real_cond_exp MA A.Invariants phiA x" using main_limit by presburger then have "real_cond_exp MA A.Invariants fA x = real_cond_exp M Invariants f x * real_cond_exp MA A.Invariants phiA x" using A LIMSEQ_unique by auto } then show ?thesis using a b c d unfolding phiA_def by auto qed corollary Invariants_cond_exp_induced_map_0: fixes f::"'a \ real" assumes [measurable]: "A \ sets M" "integrable M f" and "AE x in M. real_cond_exp M Invariants f x = 0" defines "MA \ restrict_space M A" and "TA \ induced_map A" and "fA \ induced_function A f" shows "AE x in MA. real_cond_exp MA (qmpt.Invariants MA TA) fA x = 0" proof - have "AE x in MA. real_cond_exp M Invariants f x = 0" unfolding MA_def apply (subst AE_restrict_space_iff) using assms(3) by auto then show ?thesis unfolding MA_def TA_def fA_def using Invariants_cond_exp_induced_map[OF assms(1) assms(2)] by auto qed end (*of locale fmpt*) end (*of Invariants.thy*) diff --git a/thys/Ergodic_Theory/Measure_Preserving_Transformations.thy b/thys/Ergodic_Theory/Measure_Preserving_Transformations.thy --- a/thys/Ergodic_Theory/Measure_Preserving_Transformations.thy +++ b/thys/Ergodic_Theory/Measure_Preserving_Transformations.thy @@ -1,1638 +1,1638 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Measure preserving or quasi-preserving maps\ theory Measure_Preserving_Transformations imports SG_Library_Complement begin text \Ergodic theory in general is the study of the properties of measure preserving or quasi-preserving dynamical systems. In this section, we introduce the basic definitions in this respect.\ subsection \The different classes of transformations\ definition quasi_measure_preserving::"'a measure \ 'b measure \ ('a \ 'b) set" where "quasi_measure_preserving M N = {f \ measurable M N. \ A \ sets N. (f -` A \ space M \ null_sets M) = (A \ null_sets N)}" lemma quasi_measure_preservingI [intro]: assumes "f \ measurable M N" "\A. A \ sets N \ (f -` A \ space M \ null_sets M) = (A \ null_sets N)" shows "f \ quasi_measure_preserving M N" using assms unfolding quasi_measure_preserving_def by auto lemma quasi_measure_preservingE: assumes "f \ quasi_measure_preserving M N" shows "f \ measurable M N" "\A. A \ sets N \ (f -` A \ space M \ null_sets M) = (A \ null_sets N)" using assms unfolding quasi_measure_preserving_def by auto lemma id_quasi_measure_preserving: "(\x. x) \ quasi_measure_preserving M M" unfolding quasi_measure_preserving_def by auto lemma quasi_measure_preserving_composition: assumes "f \ quasi_measure_preserving M N" "g \ quasi_measure_preserving N P" shows "(\x. g(f x)) \ quasi_measure_preserving M P" proof (rule quasi_measure_preservingI) have f_meas [measurable]: "f \ measurable M N" by (rule quasi_measure_preservingE(1)[OF assms(1)]) have g_meas [measurable]: "g \ measurable N P" by (rule quasi_measure_preservingE(1)[OF assms(2)]) then show [measurable]: "(\x. g (f x)) \ measurable M P" by auto fix C assume [measurable]: "C \ sets P" define B where "B = g-`C \ space N" have [measurable]: "B \ sets N" unfolding B_def by simp have *: "B \ null_sets N \ C \ null_sets P" unfolding B_def using quasi_measure_preservingE(2)[OF assms(2)] by simp define A where "A = f-`B \ space M" have [measurable]: "A \ sets M" unfolding A_def by simp have "A \ null_sets M \ B \ null_sets N" unfolding A_def using quasi_measure_preservingE(2)[OF assms(1)] by simp then have "A \ null_sets M \ C \ null_sets P" using * by simp moreover have "A = (\x. g (f x)) -` C \ space M" by (auto simp add: A_def B_def) (meson f_meas measurable_space) ultimately show "((\x. g (f x)) -` C \ space M \ null_sets M) \ C \ null_sets P" by simp qed lemma quasi_measure_preserving_comp: assumes "f \ quasi_measure_preserving M N" "g \ quasi_measure_preserving N P" shows "g o f \ quasi_measure_preserving M P" unfolding comp_def using assms quasi_measure_preserving_composition by blast lemma quasi_measure_preserving_AE: assumes "f \ quasi_measure_preserving M N" "AE x in N. P x" shows "AE x in M. P (f x)" proof - obtain A where "\x. x \ space N - A \ P x" "A \ null_sets N" using AE_E3[OF assms(2)] by blast define B where "B = f-`A \ space M" have "B \ null_sets M" unfolding B_def using quasi_measure_preservingE(2)[OF assms(1)] \A \ null_sets N\ by auto moreover have "x \ space M - B \ P (f x)" for x using \\x. x \ space N - A \ P x\ quasi_measure_preservingE(1)[OF assms(1)] unfolding B_def by (metis (no_types, lifting) Diff_iff IntI measurable_space vimage_eq) ultimately show ?thesis using AE_not_in AE_space by force qed lemma quasi_measure_preserving_AE': assumes "f \ quasi_measure_preserving M N" "AE x in M. P (f x)" "{x \ space N. P x} \ sets N" shows "AE x in N. P x" proof - have [measurable]: "f \ measurable M N" using quasi_measure_preservingE(1)[OF assms(1)] by simp define U where "U = {x \ space N. \(P x)}" have [measurable]: "U \ sets N" unfolding U_def using assms(3) by auto have "f-`U \ space M = {x \ space M. \(P (f x))}" unfolding U_def using \f \ measurable M N\ by (auto, meson measurable_space) also have "... \ null_sets M" apply (subst AE_iff_null[symmetric]) using assms by auto finally have "U \ null_sets N" using quasi_measure_preservingE(2)[OF assms(1) \U \ sets N\] by auto then show ?thesis unfolding U_def using AE_iff_null by blast qed text \The push-forward under a quasi-measure preserving map $f$ of a measure absolutely continuous with respect to $M$ is absolutely continuous with respect to $N$.\ lemma quasi_measure_preserving_absolutely_continuous: assumes "f \ quasi_measure_preserving M N" "u \ borel_measurable M" shows "absolutely_continuous N (distr (density M u) N f)" proof - have [measurable]: "f \ measurable M N" using quasi_measure_preservingE[OF assms(1)] by auto have "S \ null_sets (distr (density M u) N f)" if [measurable]: "S \ null_sets N" for S proof - have [measurable]: "S \ sets N" using null_setsD2[OF that] by auto have *: "AE x in N. x \ S" by (metis AE_not_in that) have "AE x in M. f x \ S" by (rule quasi_measure_preserving_AE[OF _ *], simp add: assms) then have *: "AE x in M. indicator S (f x) * u x = 0" by force have "emeasure (distr (density M u) N f) S = (\\<^sup>+x. indicator S x \(distr (density M u) N f))" by auto also have "... = (\\<^sup>+x. indicator S (f x) \(density M u))" by (rule nn_integral_distr, auto) also have "... = (\\<^sup>+x. indicator S (f x) * u x \M)" by (rule nn_integral_densityR[symmetric], auto simp add: assms) also have "... = (\\<^sup>+x. 0 \M)" using * by (rule nn_integral_cong_AE) finally have "emeasure (distr (density M u) N f) S = 0" by auto then show ?thesis by auto qed then show ?thesis unfolding absolutely_continuous_def by auto qed definition measure_preserving::"'a measure \ 'b measure \ ('a \ 'b) set" where "measure_preserving M N = {f \ measurable M N. (\ A \ sets N. emeasure M (f-`A \ space M) = emeasure N A)}" lemma measure_preservingE: assumes "f \ measure_preserving M N" shows "f \ measurable M N" "\A. A \ sets N \ emeasure M (f-`A \ space M) = emeasure N A" using assms unfolding measure_preserving_def by auto lemma measure_preservingI [intro]: assumes "f \ measurable M N" "\A. A \ sets N \ emeasure M (f-`A \ space M) = emeasure N A" shows "f \ measure_preserving M N" using assms unfolding measure_preserving_def by auto lemma measure_preserving_distr: assumes "f \ measure_preserving M N" shows "distr M N f = N" proof - let ?N2 = "distr M N f" have "sets ?N2 = sets N" by simp moreover have "emeasure ?N2 A = emeasure N A" if "A \ sets N" for A proof - have "emeasure ?N2 A = emeasure M (f-`A \ space M)" using \A \ sets N\ assms emeasure_distr measure_preservingE(1)[OF assms] by blast then show "emeasure ?N2 A = emeasure N A" using \A \ sets N\ measure_preservingE(2)[OF assms] by auto qed ultimately show ?thesis by (metis measure_eqI) qed lemma measure_preserving_distr': assumes "f \ measurable M N" shows "f \ measure_preserving M (distr M N f)" proof (rule measure_preservingI) show "f \ measurable M (distr M N f)" using assms(1) by auto show "emeasure M (f-`A \ space M) = emeasure (distr M N f) A" if "A \ sets (distr M N f)" for A using that emeasure_distr[OF assms] by auto qed lemma measure_preserving_preserves_nn_integral: assumes "T \ measure_preserving M N" "f \ borel_measurable N" shows "(\\<^sup>+x. f x \N) = (\\<^sup>+x. f (T x) \M)" proof - have "(\\<^sup>+x. f (T x) \M) = (\\<^sup>+y. f y \distr M N T)" using assms nn_integral_distr[of T M N f, OF measure_preservingE(1)[OF assms(1)]] by simp also have "... = (\\<^sup>+y. f y \N)" using measure_preserving_distr[OF assms(1)] by simp finally show ?thesis by simp qed lemma measure_preserving_preserves_integral: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "T \ measure_preserving M N" and [measurable]: "integrable N f" shows "integrable M (\x. f(T x))" "(\x. f x \N) = (\x. f (T x) \M)" proof - have a [measurable]: "T \ measurable M N" by (rule measure_preservingE(1)[OF assms(1)]) have b [measurable]: "f \ borel_measurable N" by simp have "distr M N T = N" using measure_preserving_distr[OF assms(1)] by simp then have "integrable (distr M N T) f" using assms(2) by simp then show "integrable M (\x. f(T x))" using integrable_distr_eq[OF a b] by simp have "(\x. f (T x) \M) = (\y. f y \distr M N T)" using integral_distr[OF a b] by simp then show "(\x. f x \N) = (\x. f (T x) \M)" using \distr M N T = N\ by simp qed lemma measure_preserving_preserves_integral': fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "T \ measure_preserving M N" and [measurable]: "integrable M (\x. f (T x))" "f \ borel_measurable N" shows "integrable N f" "(\x. f x \N) = (\x. f (T x) \M)" proof - have a [measurable]: "T \ measurable M N" by (rule measure_preservingE(1)[OF assms(1)]) have "integrable M (\x. f(T x))" using assms(2) unfolding comp_def by auto then have "integrable (distr M N T) f" using integrable_distr_eq[OF a assms(3)] by simp then show *: "integrable N f" using measure_preserving_distr[OF assms(1)] by simp then show "(\x. f x \N) = (\x. f (T x) \M)" using measure_preserving_preserves_integral[OF assms(1) *] by simp qed lemma id_measure_preserving: "(\x. x) \ measure_preserving M M" unfolding measure_preserving_def by auto lemma measure_preserving_is_quasi_measure_preserving: assumes "f \ measure_preserving M N" shows "f \ quasi_measure_preserving M N" using assms unfolding measure_preserving_def quasi_measure_preserving_def apply auto by (metis null_setsD1 null_setsI, metis measurable_sets null_setsD1 null_setsI) lemma measure_preserving_composition: assumes "f \ measure_preserving M N" "g \ measure_preserving N P" shows "(\x. g(f x)) \ measure_preserving M P" proof (rule measure_preservingI) have f [measurable]: "f \ measurable M N" by (rule measure_preservingE(1)[OF assms(1)]) have g [measurable]: "g \ measurable N P" by (rule measure_preservingE(1)[OF assms(2)]) show [measurable]: "(\x. g (f x)) \ measurable M P" by auto fix C assume [measurable]: "C \ sets P" define B where "B = g-`C \ space N" have [measurable]: "B \ sets N" unfolding B_def by simp have *: "emeasure N B = emeasure P C" unfolding B_def using measure_preservingE(2)[OF assms(2)] by simp define A where "A = f-`B \ space M" have [measurable]: "A \ sets M" unfolding A_def by simp have "emeasure M A = emeasure N B" unfolding A_def using measure_preservingE(2)[OF assms(1)] by simp then have "emeasure M A = emeasure P C" using * by simp moreover have "A = (\x. g(f x))-`C \ space M" by (auto simp add: A_def B_def) (meson f measurable_space) ultimately show "emeasure M ((\x. g(f x))-`C \ space M) = emeasure P C" by simp qed lemma measure_preserving_comp: assumes "f \ measure_preserving M N" "g \ measure_preserving N P" shows "g o f \ measure_preserving M P" unfolding o_def using measure_preserving_composition assms by blast lemma measure_preserving_total_measure: assumes "f \ measure_preserving M N" shows "emeasure M (space M) = emeasure N (space N)" proof - have "f \ measurable M N" by (rule measure_preservingE(1)[OF assms(1)]) then have "f-`(space N) \ space M = space M" by (meson Int_absorb1 measurable_space subsetI vimageI) then show "emeasure M (space M) = emeasure N (space N)" by (metis (mono_tags, lifting) measure_preservingE(2)[OF assms(1)] sets.top) qed lemma measure_preserving_finite_measure: assumes "f \ measure_preserving M N" shows "finite_measure M \ finite_measure N" using measure_preserving_total_measure[OF assms] by (metis finite_measure.emeasure_finite finite_measureI infinity_ennreal_def) lemma measure_preserving_prob_space: assumes "f \ measure_preserving M N" shows "prob_space M \ prob_space N" using measure_preserving_total_measure[OF assms] by (metis prob_space.emeasure_space_1 prob_spaceI) locale qmpt = sigma_finite_measure + fixes T assumes Tqm: "T \ quasi_measure_preserving M M" locale mpt = qmpt + assumes Tm: "T \ measure_preserving M M" locale fmpt = mpt + finite_measure locale pmpt = fmpt + prob_space lemma qmpt_I: assumes "sigma_finite_measure M" "T \ measurable M M" "\A. A \ sets M \ ((T-`A \ space M) \ null_sets M) \ (A \ null_sets M)" shows "qmpt M T" unfolding qmpt_def qmpt_axioms_def quasi_measure_preserving_def by (auto simp add: assms) lemma mpt_I: assumes "sigma_finite_measure M" "T \ measurable M M" "\A. A \ sets M \ emeasure M (T-`A \ space M) = emeasure M A" shows "mpt M T" proof - have *: "T \ measure_preserving M M" by (rule measure_preservingI[OF assms(2) assms(3)]) then have **: "T \ quasi_measure_preserving M M" using measure_preserving_is_quasi_measure_preserving by auto show "mpt M T" unfolding mpt_def qmpt_def qmpt_axioms_def mpt_axioms_def using * ** assms(1) by auto qed lemma fmpt_I: assumes "finite_measure M" "T \ measurable M M" "\A. A \ sets M \ emeasure M (T-`A \ space M) = emeasure M A" shows "fmpt M T" proof - have *: "T \ measure_preserving M M" by (rule measure_preservingI[OF assms(2) assms(3)]) then have **: "T \ quasi_measure_preserving M M" using measure_preserving_is_quasi_measure_preserving by auto show "fmpt M T" unfolding fmpt_def mpt_def qmpt_def mpt_axioms_def qmpt_axioms_def using * ** assms(1) finite_measure_def by auto qed lemma pmpt_I: assumes "prob_space M" "T \ measurable M M" "\A. A \ sets M \ emeasure M (T-`A \ space M) = emeasure M A" shows "pmpt M T" proof - have *: "T \ measure_preserving M M" by (rule measure_preservingI[OF assms(2) assms(3)]) then have **: "T \ quasi_measure_preserving M M" using measure_preserving_is_quasi_measure_preserving by auto show "pmpt M T" unfolding pmpt_def fmpt_def mpt_def qmpt_def mpt_axioms_def qmpt_axioms_def using * ** assms(1) prob_space_imp_sigma_finite prob_space.finite_measure by auto qed subsection \Examples\ lemma fmpt_null_space: assumes "emeasure M (space M) = 0" "T \ measurable M M" shows "fmpt M T" apply (rule fmpt_I) apply (auto simp add: assms finite_measureI) apply (metis assms emeasure_eq_0 measurable_sets sets.sets_into_space sets.top) done lemma fmpt_empty_space: assumes "space M = {}" shows "fmpt M T" by (rule fmpt_null_space, auto simp add: assms measurable_empty_iff) text \Translations are measure-preserving\ lemma mpt_translation: fixes c :: "'a::euclidean_space" shows "mpt lborel (\x. x + c)" proof (rule mpt_I, auto simp add: lborel.sigma_finite_measure_axioms) fix A::"'a set" assume [measurable]: "A \ sets borel" have "emeasure lborel ((\x. x + c) -` A) = emeasure lborel ((((+))c)-`A)" by (meson add.commute) also have "... = emeasure lborel ((((+))c)-`A \ space lborel)" by simp also have "... = emeasure (distr lborel borel ((+) c)) A" by (rule emeasure_distr[symmetric], auto) - also have "... = emeasure lborel A" using lborel_distr_plus2[of c] by simp + also have "... = emeasure lborel A" using lborel_distr_plus[of c] by simp finally show "emeasure lborel ((\x. x + c) -` A) = emeasure lborel A" by simp qed text \Skew products are fibered maps of the form $(x,y)\mapsto (Tx, U(x,y))$. If the base map and the fiber maps all are measure preserving, so is the skew product.\ lemma pair_measure_null_product: assumes "emeasure M (space M) = 0" shows "emeasure (M \\<^sub>M N) (space (M \\<^sub>M N)) = 0" proof - have "(\\<^sup>+x. (\\<^sup>+y. indicator X (x,y) \N) \M) = 0" for X proof - have "(\\<^sup>+x. (\\<^sup>+y. indicator X (x,y) \N) \M) = (\\<^sup>+x. 0 \M)" by (intro nn_integral_cong_AE emeasure_0_AE[OF assms]) then show ?thesis by auto qed then have "M \\<^sub>M N = measure_of (space M \ space N) {a \ b | a b. a \ sets M \ b \ sets N} (\X. 0)" unfolding pair_measure_def by auto then show ?thesis by (simp add: emeasure_sigma) qed lemma mpt_skew_product: assumes "mpt M T" "AE x in M. mpt N (U x)" and [measurable]: "(\(x,y). U x y) \ measurable (M \\<^sub>M N) N" shows "mpt (M \\<^sub>M N) (\(x,y). (T x, U x y))" proof (cases) assume H: "emeasure M (space M) = 0" then have *: "emeasure (M \\<^sub>M N) (space (M \\<^sub>M N)) = 0" using pair_measure_null_product by auto have [measurable]: "T \ measurable M M" using assms(1) unfolding mpt_def qmpt_def qmpt_axioms_def quasi_measure_preserving_def by auto then have [measurable]: "(\(x, y). (T x, U x y)) \ measurable (M \\<^sub>M N) (M \\<^sub>M N)" by auto with fmpt_null_space[OF *] show ?thesis by (simp add: fmpt.axioms(1)) next assume "\(emeasure M (space M) = 0)" show ?thesis proof (rule mpt_I) have "sigma_finite_measure M" using assms(1) unfolding mpt_def qmpt_def by auto then interpret M: sigma_finite_measure M . have "\p. \ almost_everywhere M p" by (metis (lifting) AE_E \emeasure M (space M) \ 0\ emeasure_eq_AE emeasure_notin_sets) then have "\x. mpt N (U x)" using assms(2) \\(emeasure M (space M) = 0)\ by (metis (full_types) \AE x in M. mpt N (U x)\ eventually_mono) then have "sigma_finite_measure N" unfolding mpt_def qmpt_def by auto then interpret N: sigma_finite_measure N . show "sigma_finite_measure (M \\<^sub>M N)" by (rule sigma_finite_pair_measure) standard+ have [measurable]: "T \ measurable M M" using assms(1) unfolding mpt_def qmpt_def qmpt_axioms_def quasi_measure_preserving_def by auto show [measurable]: "(\(x, y). (T x, U x y)) \ measurable (M \\<^sub>M N) (M \\<^sub>M N)" by auto have "T \ measure_preserving M M" using assms(1) by (simp add: mpt.Tm) fix A assume [measurable]: "A \ sets (M \\<^sub>M N)" then have [measurable]: "(\ (x,y). (indicator A (x,y))::ennreal) \ borel_measurable (M \\<^sub>M N)" by auto then have [measurable]: "(\x. \\<^sup>+ y. indicator A (x, y) \N) \ borel_measurable M" by simp define B where "B = (\(x, y). (T x, U x y)) -` A \ space (M \\<^sub>M N)" then have [measurable]: "B \ sets (M \\<^sub>M N)" by auto have "(\\<^sup>+y. indicator B (x,y) \N) = (\\<^sup>+y. indicator A (T x, y) \N)" if "x \ space M" "mpt N (U x)" for x proof - have "T x \ space M" by (meson \T \ measurable M M\ \x \ space M\ measurable_space) then have 1: "(\y. (indicator A (T x, y))::ennreal) \ borel_measurable N" using \A \ sets (M \\<^sub>M N)\ by auto have 2: "\y. ((indicator B (x, y))::ennreal) = indicator A (T x, U x y) * indicator (space M) x * indicator (space N) y" unfolding B_def by (simp add: indicator_def space_pair_measure) have 3: "U x \ measure_preserving N N" using assms(2) that(2) by (simp add: mpt.Tm) have "(\\<^sup>+y. indicator B (x,y) \N) = (\\<^sup>+y. indicator A (T x, U x y) \N)" using 2 by (intro nn_integral_cong_simp) (auto simp add: indicator_def \x \ space M\) also have "... = (\\<^sup>+y. indicator A (T x, y) \N)" by (rule measure_preserving_preserves_nn_integral[OF 3, symmetric], metis 1) finally show ?thesis by simp qed then have *: "AE x in M. (\\<^sup>+y. indicator B (x,y) \N) = (\\<^sup>+y. indicator A (T x, y) \N)" using assms(2) by auto have "emeasure (M \\<^sub>M N) B = (\\<^sup>+ x. (\\<^sup>+y. indicator B (x,y) \N) \M)" using \B \ sets (M \\<^sub>M N)\ \sigma_finite_measure N\ sigma_finite_measure.emeasure_pair_measure by fastforce also have "... = (\\<^sup>+ x. (\\<^sup>+y. indicator A (T x, y) \N) \M)" by (intro nn_integral_cong_AE *) also have "... = (\\<^sup>+ x. (\\<^sup>+y. indicator A (x, y) \N) \M)" by (rule measure_preserving_preserves_nn_integral[OF \T \ measure_preserving M M\, symmetric]) auto also have "... = emeasure (M \\<^sub>M N) A" by (simp add: \sigma_finite_measure N\ sigma_finite_measure.emeasure_pair_measure) finally show "emeasure (M \\<^sub>M N) ((\(x, y). (T x, U x y)) -` A \ space (M \\<^sub>M N)) = emeasure (M \\<^sub>M N) A" unfolding B_def by simp qed qed lemma mpt_skew_product_real: fixes f::"'a \ 'b::euclidean_space" assumes "mpt M T" and [measurable]: "f \ borel_measurable M" shows "mpt (M \\<^sub>M lborel) (\(x,y). (T x, y + f x))" by (rule mpt_skew_product, auto simp add: mpt_translation assms(1)) subsection \Preimages restricted to $space M$\ context qmpt begin text \One is all the time lead to take the preimages of sets, and restrict them to \verb+space M+ where the dynamics is living. We introduce a shortcut for this notion.\ definition vimage_restr :: "('a \ 'a) \ 'a set \ 'a set" (infixr "--`" 90) where "f --` A \ f-` (A \ space M) \ space M" lemma vrestr_eq [simp]: "a \ f--` A \ a \ space M \ f a \ A \ space M" unfolding vimage_restr_def by auto lemma vrestr_intersec [simp]: "f--` (A \ B) = (f--`A) \ (f--` B)" using vimage_restr_def by auto lemma vrestr_union [simp]: "f--` (A \ B) = f--`A \ f--`B" using vimage_restr_def by auto lemma vrestr_difference [simp]: "f--`(A-B) = f--`A - f--`B" using vimage_restr_def by auto lemma vrestr_inclusion: "A \ B \ f--`A \ f--`B" using vimage_restr_def by auto lemma vrestr_Union [simp]: "f --` (\A) = (\X\A. f --` X)" using vimage_restr_def by auto lemma vrestr_UN [simp]: "f --` (\x\A. B x) = (\x\A. f --` B x)" using vimage_restr_def by auto lemma vrestr_Inter [simp]: assumes "A \ {}" shows "f --` (\A) = (\X\A. f --` X)" using vimage_restr_def assms by auto lemma vrestr_INT [simp]: assumes "A \ {}" shows "f --` (\x\A. B x) = (\x\A. f --` B x)" using vimage_restr_def assms by auto lemma vrestr_empty [simp]: "f--`{} = {}" using vimage_restr_def by auto lemma vrestr_sym_diff [simp]: "f--`(A \ B) = (f--`A) \ (f--`B)" by auto lemma vrestr_image: assumes "x \ f--`A" shows "x \ space M" "f x \ space M" "f x \ A" using assms unfolding vimage_restr_def by auto lemma vrestr_intersec_in_space: assumes "A \ sets M" "B \ sets M" shows "A \ f--`B = A \ f-`B" unfolding vimage_restr_def using assms sets.sets_into_space by auto lemma vrestr_compose: assumes "g \ measurable M M" shows "(\ x. f(g x))--` A = g--` (f--` A)" proof - define B where "B = A \ space M" have "(\ x. f(g x))--` A = (\ x. f(g x)) -` B \ space M" using B_def vimage_restr_def by blast moreover have "(\ x. f(g x)) -` B \ space M = g-` (f-` B \ space M) \ space M" using measurable_space[OF \g \ measurable M M\] by auto moreover have "g-` (f-` B \ space M) \ space M = g--` (f--` A)" using B_def vimage_restr_def by simp ultimately show ?thesis by auto qed lemma vrestr_comp: assumes "g \ measurable M M" shows "(f o g)--` A = g--` (f--` A)" proof - have "f o g = (\ x. f(g x))" by auto then have "(f o g)--` A = (\ x. f(g x))--` A" by auto moreover have "(\ x. f(g x))--` A = g--` (f--` A)" using vrestr_compose assms by auto ultimately show ?thesis by simp qed lemma vrestr_of_set: assumes "g \ measurable M M" shows "A \ sets M \ g--`A = g-`A \ space M" by (simp add: vimage_restr_def) lemma vrestr_meas [measurable (raw)]: assumes "g \ measurable M M" "A \ sets M" shows "g--`A \ sets M" using assms vimage_restr_def by auto lemma vrestr_same_emeasure_f: assumes "f \ measure_preserving M M" "A \ sets M" shows "emeasure M (f--`A) = emeasure M A" by (metis (mono_tags, lifting) assms measure_preserving_def mem_Collect_eq sets.Int_space_eq2 vimage_restr_def) lemma vrestr_same_measure_f: assumes "f \ measure_preserving M M" "A \ sets M" shows "measure M (f--`A) = measure M A" proof - have "measure M (f--`A) = enn2real (emeasure M (f--`A))" by (simp add: Sigma_Algebra.measure_def) also have "... = enn2real (emeasure M A)" using vrestr_same_emeasure_f[OF assms] by simp also have "... = measure M A" by (simp add: Sigma_Algebra.measure_def) finally show "measure M (f--` A) = measure M A" by simp qed subsection \Basic properties of qmpt\ lemma T_meas [measurable (raw)]: "T \ measurable M M" by (rule quasi_measure_preservingE(1)[OF Tqm]) lemma Tn_quasi_measure_preserving: "T^^n \ quasi_measure_preserving M M" proof (induction n) case 0 show ?case using id_quasi_measure_preserving by simp next case (Suc n) then show ?case using Tqm quasi_measure_preserving_comp by (metis funpow_Suc_right) qed lemma Tn_meas [measurable (raw)]: "T^^n \ measurable M M" by (rule quasi_measure_preservingE(1)[OF Tn_quasi_measure_preserving]) lemma T_vrestr_meas [measurable]: assumes "A \ sets M" shows "T--` A \ sets M" "(T^^n)--` A \ sets M" by (auto simp add: vrestr_meas assms) text \We state the next lemma both with $T^0$ and with $id$ as sometimes the simplifier simplifies $T^0$ to $id$ before applying the first instance of the lemma.\ lemma T_vrestr_0 [simp]: assumes "A \ sets M" shows "(T^^0)--`A = A" "id--`A = A" using sets.sets_into_space[OF assms] by auto lemma T_vrestr_composed: assumes "A \ sets M" shows "(T^^n)--` (T^^m)--` A = (T^^(n+m))--` A" "T--` (T^^m)--` A = (T^^(m+1))--` A" "(T^^m)--` T--` A = (T^^(m+1))--` A" proof - show "(T^^n)--` (T^^m)--` A = (T^^(n+m))--` A" by (simp add: Tn_meas funpow_add add.commute vrestr_comp) show "T--` (T^^m)--` A = (T^^(m+1))--` A" by (metis Suc_eq_plus1 T_meas funpow_Suc_right vrestr_comp) show "(T^^m)--` T--` A = (T^^(m+1))--` A" by (simp add: Tn_meas vrestr_comp) qed text \In the next two lemmas, we give measurability statements that show up all the time for the usual preimage.\ lemma T_intersec_meas [measurable]: assumes [measurable]: "A \ sets M" "B \ sets M" shows "A \ T-`B \ sets M" "A \ (T^^n)-`B \ sets M" "T-`A \ B \ sets M" "(T^^n)-`A \ B \ sets M" "A \ (T \ T ^^ n) -` B \ sets M" "(T \ T ^^ n) -` A \ B \ sets M" by (metis T_meas Tn_meas assms(1) assms(2) measurable_comp sets.Int inf_commute vrestr_intersec_in_space vrestr_meas)+ lemma T_diff_meas [measurable]: assumes [measurable]: "A \ sets M" "B \ sets M" shows "A - T-`B \ sets M" "A - (T^^n)-`B \ sets M" proof - have "A - T-`B = A \ space M - (T-`B \ space M)" using sets.sets_into_space[OF assms(1)] by auto then show "A - T-`B \ sets M" by auto have "A - (T^^n)-`B = A \ space M - ((T^^n)-`B \ space M)" using sets.sets_into_space[OF assms(1)] by auto then show "A - (T^^n)-`B \ sets M" by auto qed lemma T_spaceM_stable [simp]: assumes "x \ space M" shows "T x \ space M" "(T^^n) x \ space M" proof - show "T x \ space M" by (meson measurable_space T_meas measurable_def assms) show "(T^^n) x \ space M" by (meson measurable_space Tn_meas measurable_def assms) qed lemma T_quasi_preserves_null: assumes "A \ sets M" shows "A \ null_sets M \ T--` A \ null_sets M" "A \ null_sets M \ (T^^n)--` A \ null_sets M" using Tqm Tn_quasi_measure_preserving unfolding quasi_measure_preserving_def by (auto simp add: assms vimage_restr_def) lemma T_quasi_preserves: assumes "A \ sets M" shows "emeasure M A = 0 \ emeasure M (T--` A) = 0" "emeasure M A = 0 \ emeasure M ((T^^n)--` A) = 0" using T_quasi_preserves_null[OF assms] T_vrestr_meas assms by blast+ lemma T_quasi_preserves_null2: assumes "A \ null_sets M" shows "T--` A \ null_sets M" "(T^^n)--` A \ null_sets M" using T_quasi_preserves_null[OF null_setsD2[OF assms]] assms by auto lemma T_composition_borel [measurable]: assumes "f \ borel_measurable M" shows "(\x. f(T x)) \ borel_measurable M" "(\x. f((T^^k) x)) \ borel_measurable M" using T_meas Tn_meas assms measurable_compose by auto lemma T_AE_iterates: assumes "AE x in M. P x" shows "AE x in M. \n. P ((T^^n) x)" proof - have "AE x in M. P ((T^^n) x)" for n by (rule quasi_measure_preserving_AE[OF Tn_quasi_measure_preserving[of n] assms]) then show ?thesis unfolding AE_all_countable by simp qed lemma qmpt_power: "qmpt M (T^^n)" by (standard, simp add: Tn_quasi_measure_preserving) lemma T_Tn_T_compose: "T ((T^^n) x) = (T^^(Suc n)) x" "(T^^n) (T x) = (T^^(Suc n)) x" by (auto simp add: funpow_swap1) lemma (in qmpt) qmpt_density: assumes [measurable]: "h \ borel_measurable M" and "AE x in M. h x \ 0" "AE x in M. h x \ \" shows "qmpt (density M h) T" proof - interpret A: sigma_finite_measure "density M h" apply (subst sigma_finite_iff_density_finite) using assms by auto show ?thesis apply (standard) apply (rule quasi_measure_preservingI) unfolding null_sets_density[OF \h \ borel_measurable M\ \AE x in M. h x \ 0\] sets_density space_density using quasi_measure_preservingE(2)[OF Tqm] by auto qed end subsection \Basic properties of mpt\ context mpt begin lemma Tn_measure_preserving: "T^^n \ measure_preserving M M" proof (induction n) case (Suc n) then show ?case using Tm measure_preserving_comp by (metis funpow_Suc_right) qed (simp add: id_measure_preserving) lemma T_integral_preserving: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "integrable M f" shows "integrable M (\x. f(T x))" "(\x. f(T x) \M) = (\x. f x \M)" using measure_preserving_preserves_integral[OF Tm assms] by auto lemma Tn_integral_preserving: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "integrable M f" shows "integrable M (\x. f((T^^n) x))" "(\x. f((T^^n) x) \M) = (\x. f x \M)" using measure_preserving_preserves_integral[OF Tn_measure_preserving assms] by auto lemma T_nn_integral_preserving: fixes f :: "'a \ ennreal" assumes "f \ borel_measurable M" shows "(\\<^sup>+x. f(T x) \M) = (\\<^sup>+x. f x \M)" using measure_preserving_preserves_nn_integral[OF Tm assms] by auto lemma Tn_nn_integral_preserving: fixes f :: "'a \ ennreal" assumes "f \ borel_measurable M" shows "(\\<^sup>+x. f((T^^n) x) \M) = (\\<^sup>+x. f x \M)" using measure_preserving_preserves_nn_integral[OF Tn_measure_preserving assms(1)] by auto lemma mpt_power: "mpt M (T^^n)" by (standard, simp_all add: Tn_quasi_measure_preserving Tn_measure_preserving) lemma T_vrestr_same_emeasure: assumes "A \ sets M" shows "emeasure M (T--` A) = emeasure M A" "emeasure M ((T ^^ n)--`A) = emeasure M A" by (auto simp add: vrestr_same_emeasure_f Tm Tn_measure_preserving assms) lemma T_vrestr_same_measure: assumes "A \ sets M" shows "measure M (T--` A) = measure M A" "measure M ((T ^^ n)--`A) = measure M A" by (auto simp add: vrestr_same_measure_f Tm Tn_measure_preserving assms) lemma (in fmpt) fmpt_power: "fmpt M (T^^n)" by (standard, simp_all add: Tn_quasi_measure_preserving Tn_measure_preserving) end subsection \Birkhoff sums\ text \Birkhoff sums, obtained by summing a function along the orbit of a map, are basic objects to be understood in ergodic theory.\ context qmpt begin definition birkhoff_sum::"('a \ 'b::comm_monoid_add) \ nat \ 'a \ 'b" where "birkhoff_sum f n x = (\i\{.. 'b::{second_countable_topology, topological_comm_monoid_add}" assumes "f \ borel_measurable M" shows "birkhoff_sum f n \ borel_measurable M" proof - define F where "F = (\i x. f((T^^i)x))" have "\i. F i \ borel_measurable M" using assms F_def by auto then have "(\x. (\i borel_measurable M" by measurable then have "(\x. birkhoff_sum f n x) \ borel_measurable M" unfolding birkhoff_sum_def F_def by auto then show ?thesis by simp qed lemma birkhoff_sum_1 [simp]: "birkhoff_sum f 0 x = 0" "birkhoff_sum f 1 x = f x" "birkhoff_sum f (Suc 0) x = f x" unfolding birkhoff_sum_def by auto lemma birkhoff_sum_cocycle: "birkhoff_sum f (n+m) x = birkhoff_sum f n x + birkhoff_sum f m ((T^^n)x)" proof - have "(\iij\{n..< m+n}. f ((T ^^j) x))" using atLeast0LessThan sum.shift_bounds_nat_ivl[where ?g = "\j. f((T^^j)x)" and ?k = n and ?m = 0 and ?n = m, symmetric] add.commute add.left_neutral by auto finally have *: "birkhoff_sum f m ((T^^n)x) = (\j\{n..< m+n}. f ((T ^^j) x))" unfolding birkhoff_sum_def by auto have "birkhoff_sum f (n+m) x = (\ii\{n..i\{n.. real" assumes "\x. f x \ g x" shows "birkhoff_sum f n x \ birkhoff_sum g n x" unfolding birkhoff_sum_def by (simp add: assms sum_mono) lemma birkhoff_sum_abs: fixes f::"_ \ 'b::real_normed_vector" shows "norm(birkhoff_sum f n x) \ birkhoff_sum (\x. norm(f x)) n x" unfolding birkhoff_sum_def using norm_sum by auto lemma birkhoff_sum_add: "birkhoff_sum (\x. f x + g x) n x = birkhoff_sum f n x + birkhoff_sum g n x" unfolding birkhoff_sum_def by (simp add: sum.distrib) lemma birkhoff_sum_diff: fixes f g::"_ \ real" shows "birkhoff_sum (\x. f x - g x) n x = birkhoff_sum f n x - birkhoff_sum g n x" unfolding birkhoff_sum_def by (simp add: sum_subtractf) lemma birkhoff_sum_cmult: fixes f::"_ \ real" shows "birkhoff_sum (\x. c * f x) n x = c * birkhoff_sum f n x" unfolding birkhoff_sum_def by (simp add: sum_distrib_left) lemma skew_product_real_iterates: fixes f::"'a \ real" shows "((\(x,y). (T x, y + f x))^^n) (x,y) = ((T^^n) x, y + birkhoff_sum f n x)" apply (induction n) apply (auto) apply (metis (no_types, lifting) Suc_eq_plus1 birkhoff_sum_cocycle qmpt.birkhoff_sum_1(2) qmpt_axioms) done end lemma (in mpt) birkhoff_sum_integral: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "integrable M f" shows "integrable M (birkhoff_sum f n)" "(\x. birkhoff_sum f n x \M) = n *\<^sub>R (\x. f x \M)" proof - have a: "\k. integrable M (\x. f((T^^k) x))" using Tn_integral_preserving(1) assms by blast then have "integrable M (\x. \k\{..x. birkhoff_sum f n x)" unfolding birkhoff_sum_def by auto then show "integrable M (birkhoff_sum f n)" by simp have b: "\k. (\x. f((T^^k)x) \M) = (\x. f x \M)" using Tn_integral_preserving(2) assms by blast have "(\x. birkhoff_sum f n x \M) = (\x. (\k\{..M)" unfolding birkhoff_sum_def by blast also have "... = (\k\{..x. f((T^^k) x) \M))" by (rule Bochner_Integration.integral_sum, simp add: a) also have "... = (\k\{..x. f x \M))" using b by simp also have "... = n *\<^sub>R (\x. f x \M)" by (simp add: sum_constant_scaleR) finally show "(\x. birkhoff_sum f n x \M) = n *\<^sub>R (\x. f x \M)" by simp qed lemma (in mpt) birkhoff_sum_nn_integral: fixes f :: "'a \ ennreal" assumes [measurable]: "f \ borel_measurable M" and pos: "\x. f x \ 0" shows "(\\<^sup>+x. birkhoff_sum f n x \M) = n * (\\<^sup>+x. f x \M)" proof - have [measurable]: "\k. (\x. f((T^^k)x)) \ borel_measurable M" by simp have posk: "\k x. f((T^^k)x) \ 0" using pos by simp have b: "\k. (\\<^sup>+x. f((T^^k)x) \M) = (\\<^sup>+x. f x \M)" using Tn_nn_integral_preserving assms by blast have "(\\<^sup>+x. birkhoff_sum f n x \M) = (\\<^sup>+x. (\k\{..M)" unfolding birkhoff_sum_def by blast also have "... = (\k\{..\<^sup>+x. f((T^^k) x) \M))" by (rule nn_integral_sum, auto simp add: posk) also have "... = (\k\{..\<^sup>+x. f x \M))" using b by simp also have "... = n * (\\<^sup>+x. f x \M)" by simp finally show "(\\<^sup>+x. birkhoff_sum f n x \M) = n * (\\<^sup>+x. f x \M)" by simp qed subsection \Inverse map\ context qmpt begin definition "invertible_qmpt \ (bij T \ inv T \ measurable M M)" definition "Tinv \ inv T" lemma T_Tinv_of_set: assumes "invertible_qmpt" "A \ sets M" shows "T-`(Tinv-`A \ space M) \ space M = A" using assms sets.sets_into_space unfolding Tinv_def invertible_qmpt_def apply (auto simp add: bij_betw_def) using T_spaceM_stable(1) by blast lemma Tinv_quasi_measure_preserving: assumes "invertible_qmpt" shows "Tinv \ quasi_measure_preserving M M" proof (rule quasi_measure_preservingI, auto) fix A assume [measurable]: "A \ sets M" "Tinv -` A \ space M \ null_sets M" then have "T-`(Tinv -` A \ space M) \ space M \ null_sets M" by (metis T_quasi_preserves_null2(1) null_sets.Int_space_eq2 vimage_restr_def) then show "A \ null_sets M" using T_Tinv_of_set[OF assms \A \ sets M\] by auto next show [measurable]: "Tinv \ measurable M M" using assms unfolding Tinv_def invertible_qmpt_def by blast fix A assume [measurable]: "A \ sets M" "A \ null_sets M" then have "T-`(Tinv -` A \ space M) \ space M \ null_sets M" using T_Tinv_of_set[OF assms \A \ sets M\] by auto moreover have [measurable]: "Tinv-`A \ space M \ sets M" by auto ultimately show "Tinv -` A \ space M \ null_sets M" using T_meas T_quasi_preserves_null(1) vrestr_of_set by presburger qed lemma Tinv_qmpt: assumes "invertible_qmpt" shows "qmpt M Tinv" unfolding qmpt_def qmpt_axioms_def using Tinv_quasi_measure_preserving[OF assms] by (simp add: sigma_finite_measure_axioms) end lemma (in mpt) Tinv_measure_preserving: assumes "invertible_qmpt" shows "Tinv \ measure_preserving M M" proof (rule measure_preservingI) show [measurable]: "Tinv \ measurable M M" using assms unfolding Tinv_def invertible_qmpt_def by blast fix A assume [measurable]: "A \ sets M" have "A = T-`(Tinv -` A \ space M) \ space M" using T_Tinv_of_set[OF assms \A \ sets M\] by auto then show "emeasure M (Tinv -` A \ space M) = emeasure M A" by (metis T_vrestr_same_emeasure(1) \A \ sets M\ \Tinv \ M \\<^sub>M M\ measurable_sets sets.Int_space_eq2 vimage_restr_def) qed lemma (in mpt) Tinv_mpt: assumes "invertible_qmpt" shows "mpt M Tinv" unfolding mpt_def mpt_axioms_def using Tinv_qmpt[OF assms] Tinv_measure_preserving[OF assms] by auto lemma (in fmpt) Tinv_fmpt: assumes "invertible_qmpt" shows "fmpt M Tinv" unfolding fmpt_def using Tinv_mpt[OF assms] by (simp add: finite_measure_axioms) lemma (in pmpt) Tinv_fmpt: assumes "invertible_qmpt" shows "pmpt M Tinv" unfolding pmpt_def using Tinv_fmpt[OF assms] by (simp add: prob_space_axioms) subsection \Factors\ text \Factors of a system are quotients of this system, i.e., systems that can be obtained by a projection, forgetting some part of the dynamics. It is sometimes possible to transfer a result from a factor to the original system, making it possible to prove theorems by reduction to a simpler situation. The dual notion, extension, is equally important and useful. We only mention factors below, as the results for extension readily follow by considering the original system as a factor of its extension. In this paragraph, we define factors both in the qmpt and mpt categories, and prove their basic properties. \ definition (in qmpt) qmpt_factor::"('a \ 'b) \ ('b measure) \ ('b \ 'b) \ bool" where "qmpt_factor proj M2 T2 = ((proj \ quasi_measure_preserving M M2) \ (AE x in M. proj (T x) = T2 (proj x)) \ qmpt M2 T2)" lemma (in qmpt) qmpt_factorE: assumes "qmpt_factor proj M2 T2" shows "proj \ quasi_measure_preserving M M2" "AE x in M. proj (T x) = T2 (proj x)" "qmpt M2 T2" using assms unfolding qmpt_factor_def by auto lemma (in qmpt) qmpt_factor_iterates: assumes "qmpt_factor proj M2 T2" shows "AE x in M. \n. proj ((T^^n) x) = (T2^^n) (proj x)" proof - have "AE x in M. \n. proj (T ((T^^n) x)) = T2 (proj ((T^^n) x))" by (rule T_AE_iterates[OF qmpt_factorE(2)[OF assms]]) moreover { fix x assume "\n. proj (T ((T^^n) x)) = T2 (proj ((T^^n) x))" then have H: "proj (T ((T^^n) x)) = T2 (proj ((T^^n) x))" for n by auto have "proj ((T^^n) x) = (T2^^n) (proj x)" for n apply (induction n) using H by auto then have "\n. proj ((T^^n) x) = (T2^^n) (proj x)" by auto } ultimately show ?thesis by fast qed lemma (in qmpt) qmpt_factorI: assumes "proj \ quasi_measure_preserving M M2" "AE x in M. proj (T x) = T2 (proj x)" "qmpt M2 T2" shows "qmpt_factor proj M2 T2" using assms unfolding qmpt_factor_def by auto text \When there is a quasi-measure-preserving projection, then the quotient map automatically is quasi-measure-preserving. The same goes for measure-preservation below.\ lemma (in qmpt) qmpt_factorI': assumes "proj \ quasi_measure_preserving M M2" "AE x in M. proj (T x) = T2 (proj x)" "sigma_finite_measure M2" "T2 \ measurable M2 M2" shows "qmpt_factor proj M2 T2" proof - have [measurable]: "T \ measurable M M" "T2 \ measurable M2 M2" "proj \ measurable M M2" using assms(4) quasi_measure_preservingE(1)[OF assms(1)] by auto have *: "(T2 -` A \ space M2 \ null_sets M2) = (A \ null_sets M2)" if "A \ sets M2" for A proof - obtain U where U: "\x. x \ space M - U \ proj (T x) = T2 (proj x)" "U \ null_sets M" using AE_E3[OF assms(2)] by blast then have [measurable]: "U \ sets M" by auto have [measurable]: "A \ sets M2" using that by simp have e1: "(T-`(proj-`A \ space M)) \ space M = T-`(proj-`A) \ space M" using subset_eq by auto have e2: "T-`(proj-`A) \ space M - U = proj-`(T2-`A) \ space M - U" using U(1) by auto have e3: "proj-`(T2-`A) \ space M = proj-`(T2-`A \ space M2) \ space M" by (auto, meson \proj \ M \\<^sub>M M2\ measurable_space) have "A \ null_sets M2 \ proj-`A \ space M \ null_sets M" using quasi_measure_preservingE(2)[OF assms(1)] by simp also have "... \ (T-`(proj-`A \ space M)) \ space M \ null_sets M" by (rule quasi_measure_preservingE(2)[OF Tqm, symmetric], auto) also have "... \ T-`(proj-`A) \ space M \ null_sets M" using e1 by simp also have "... \ T-`(proj-`A) \ space M - U \ null_sets M" using emeasure_Diff_null_set[OF \U \ null_sets M\] unfolding null_sets_def by auto also have "... \ proj-`(T2-`A) \ space M - U \ null_sets M" using e2 by simp also have "... \ proj-`(T2-`A) \ space M \ null_sets M" using emeasure_Diff_null_set[OF \U \ null_sets M\] unfolding null_sets_def by auto also have "... \ proj-`(T2-`A \ space M2) \ space M \ null_sets M" using e3 by simp also have "... \ T2-`A \ space M2 \ null_sets M2" using quasi_measure_preservingE(2)[OF assms(1), of "T2-`A \ space M2"] by simp finally show "T2-`A \ space M2 \ null_sets M2 \ A \ null_sets M2" by simp qed show ?thesis by (intro qmpt_factorI qmpt_I) (auto simp add: assms *) qed lemma qmpt_factor_compose: assumes "qmpt M1 T1" "qmpt.qmpt_factor M1 T1 proj1 M2 T2" "qmpt.qmpt_factor M2 T2 proj2 M3 T3" shows "qmpt.qmpt_factor M1 T1 (proj2 o proj1) M3 T3" proof - have *: "proj1 \ quasi_measure_preserving M1 M2 \ AE x in M2. proj2 (T2 x) = T3 (proj2 x) \ (AE x in M1. proj1 (T1 x) = T2 (proj1 x) \ proj2 (T2 (proj1 x)) = T3 (proj2 (proj1 x)))" proof - assume "AE y in M2. proj2 (T2 y) = T3 (proj2 y)" "proj1 \ quasi_measure_preserving M1 M2" then have "AE x in M1. proj2 (T2 (proj1 x)) = T3 (proj2 (proj1 x))" using quasi_measure_preserving_AE by auto moreover { fix x assume "proj2 (T2 (proj1 x)) = T3 (proj2 (proj1 x))" then have "proj1 (T1 x) = T2 (proj1 x) \ proj2 (T2 (proj1 x)) = T3 (proj2 (proj1 x))" by auto } ultimately show "AE x in M1. proj1 (T1 x) = T2 (proj1 x) \ proj2 (T2 (proj1 x)) = T3 (proj2 (proj1 x))" by auto qed interpret I: qmpt M1 T1 using assms(1) by simp interpret J: qmpt M2 T2 using I.qmpt_factorE(3)[OF assms(2)] by simp show "I.qmpt_factor (proj2 o proj1) M3 T3" apply (rule I.qmpt_factorI) using I.qmpt_factorE[OF assms(2)] J.qmpt_factorE[OF assms(3)] by (auto simp add: quasi_measure_preserving_comp *) qed text \The left shift on natural integers is a very natural dynamical system, that can be used to model many systems as we see below. For invertible systems, one uses rather all the integers.\ definition nat_left_shift::"(nat \ 'a) \ (nat \ 'a)" where "nat_left_shift x = (\i. x (i+1))" lemma nat_left_shift_continuous [intro, continuous_intros]: "continuous_on UNIV nat_left_shift" by (rule continuous_on_coordinatewise_then_product, auto simp add: nat_left_shift_def) lemma nat_left_shift_measurable [intro, measurable]: "nat_left_shift \ measurable borel borel" by (rule borel_measurable_continuous_onI, auto) definition int_left_shift::"(int \ 'a) \ (int \ 'a)" where "int_left_shift x = (\i. x (i+1))" definition int_right_shift::"(int \ 'a) \ (int \ 'a)" where "int_right_shift x = (\i. x (i-1))" lemma int_shift_continuous [intro, continuous_intros]: "continuous_on UNIV int_left_shift" "continuous_on UNIV int_right_shift" apply (rule continuous_on_coordinatewise_then_product, auto simp add: int_left_shift_def) apply (rule continuous_on_coordinatewise_then_product, auto simp add: int_right_shift_def) done lemma int_shift_measurable [intro, measurable]: "int_left_shift \ measurable borel borel" "int_right_shift \ measurable borel borel" by (rule borel_measurable_continuous_onI, auto)+ lemma int_shift_bij: "bij int_left_shift" "inv int_left_shift = int_right_shift" "bij int_right_shift" "inv int_right_shift = int_left_shift" proof - show "bij int_left_shift" apply (rule bij_betw_byWitness[where ?f' = "\x. (\i. x (i-1))"]) unfolding int_left_shift_def by auto show "inv int_left_shift = int_right_shift" apply (rule inv_equality) unfolding int_left_shift_def int_right_shift_def by auto show "bij int_right_shift" apply (rule bij_betw_byWitness[where ?f' = "\x. (\i. x (i+1))"]) unfolding int_right_shift_def by auto show "inv int_right_shift = int_left_shift" apply (rule inv_equality) unfolding int_left_shift_def int_right_shift_def by auto qed lemma (in qmpt) qmpt_factor_projection: fixes f::"'a \ ('b::second_countable_topology)" assumes [measurable]: "f \ borel_measurable M" and "sigma_finite_measure (distr M borel (\x n. f ((T ^^ n) x)))" shows "qmpt_factor (\x. (\n. f ((T^^n)x))) (distr M borel (\x. (\n. f ((T^^n)x)))) nat_left_shift" proof (rule qmpt_factorI') have * [measurable]: "(\x. (\n. f ((T^^n)x))) \ borel_measurable M" using measurable_coordinatewise_then_product by measurable show "(\x n. f ((T ^^ n) x)) \ quasi_measure_preserving M (distr M borel (\x n. f ((T ^^ n) x)))" by (rule measure_preserving_is_quasi_measure_preserving[OF measure_preserving_distr'[OF *]]) have "(\n. f ((T ^^ n) (T x))) = nat_left_shift (\n. f ((T ^^ n) x))" for x unfolding nat_left_shift_def by (auto simp add: funpow_swap1) then show "AE x in M. (\n. f ((T ^^ n) (T x))) = nat_left_shift (\n. f ((T ^^ n) x))" by simp qed (auto simp add: assms(2)) text \Let us now define factors of measure-preserving transformations, in the same way as above.\ definition (in mpt) mpt_factor::"('a \ 'b) \ ('b measure) \ ('b \ 'b) \ bool" where "mpt_factor proj M2 T2 = ((proj \ measure_preserving M M2) \ (AE x in M. proj (T x) = T2 (proj x)) \ mpt M2 T2)" lemma (in mpt) mpt_factor_is_qmpt_factor: assumes "mpt_factor proj M2 T2" shows "qmpt_factor proj M2 T2" using assms unfolding mpt_factor_def qmpt_factor_def by (simp add: measure_preserving_is_quasi_measure_preserving mpt_def) lemma (in mpt) mpt_factorE: assumes "mpt_factor proj M2 T2" shows "proj \ measure_preserving M M2" "AE x in M. proj (T x) = T2 (proj x)" "mpt M2 T2" using assms unfolding mpt_factor_def by auto lemma (in mpt) mpt_factorI: assumes "proj \ measure_preserving M M2" "AE x in M. proj (T x) = T2 (proj x)" "mpt M2 T2" shows "mpt_factor proj M2 T2" using assms unfolding mpt_factor_def by auto text \When there is a measure-preserving projection commuting with the dynamics, and the dynamics above preserves the measure, then so does the dynamics below.\ lemma (in mpt) mpt_factorI': assumes "proj \ measure_preserving M M2" "AE x in M. proj (T x) = T2 (proj x)" "sigma_finite_measure M2" "T2 \ measurable M2 M2" shows "mpt_factor proj M2 T2" proof - have [measurable]: "T \ measurable M M" "T2 \ measurable M2 M2" "proj \ measurable M M2" using assms(4) measure_preservingE(1)[OF assms(1)] by auto have *: "emeasure M2 (T2 -` A \ space M2) = emeasure M2 A" if "A \ sets M2" for A proof - obtain U where U: "\x. x \ space M - U \ proj (T x) = T2 (proj x)" "U \ null_sets M" using AE_E3[OF assms(2)] by blast then have [measurable]: "U \ sets M" by auto have [measurable]: "A \ sets M2" using that by simp have e1: "(T-`(proj-`A \ space M)) \ space M = T-`(proj-`A) \ space M" using subset_eq by auto have e2: "T-`(proj-`A) \ space M - U = proj-`(T2-`A) \ space M - U" using U(1) by auto have e3: "proj-`(T2-`A) \ space M = proj-`(T2-`A \ space M2) \ space M" by (auto, meson \proj \ M \\<^sub>M M2\ measurable_space) have "emeasure M2 A = emeasure M (proj-`A \ space M)" using measure_preservingE(2)[OF assms(1)] by simp also have "... = emeasure M (T-`(proj-`A \ space M) \ space M)" by (rule measure_preservingE(2)[OF Tm, symmetric], auto) also have "... = emeasure M (T-`(proj-`A) \ space M)" using e1 by simp also have "... = emeasure M (T-`(proj-`A) \ space M - U)" using emeasure_Diff_null_set[OF \U \ null_sets M\] by auto also have "... = emeasure M (proj-`(T2-`A) \ space M - U)" using e2 by simp also have "... = emeasure M (proj-`(T2-`A) \ space M)" using emeasure_Diff_null_set[OF \U \ null_sets M\] by auto also have "... = emeasure M (proj-`(T2-`A \ space M2) \ space M)" using e3 by simp also have "... = emeasure M2 (T2-`A \ space M2)" using measure_preservingE(2)[OF assms(1), of "T2-`A \ space M2"] by simp finally show "emeasure M2 (T2-`A \ space M2) = emeasure M2 A" by simp qed show ?thesis by (intro mpt_factorI mpt_I) (auto simp add: assms *) qed lemma (in fmpt) mpt_factorI'': assumes "proj \ measure_preserving M M2" "AE x in M. proj (T x) = T2 (proj x)" "T2 \ measurable M2 M2" shows "mpt_factor proj M2 T2" apply (rule mpt_factorI', auto simp add: assms) using measure_preserving_finite_measure[OF assms(1)] finite_measure_axioms finite_measure_def by blast lemma (in fmpt) fmpt_factor: assumes "mpt_factor proj M2 T2" shows "fmpt M2 T2" unfolding fmpt_def using mpt_factorE(3)[OF assms] measure_preserving_finite_measure[OF mpt_factorE(1)[OF assms]] finite_measure_axioms by auto lemma (in pmpt) pmpt_factor: assumes "mpt_factor proj M2 T2" shows "pmpt M2 T2" unfolding pmpt_def using fmpt_factor[OF assms] measure_preserving_prob_space[OF mpt_factorE(1)[OF assms]] prob_space_axioms by auto lemma mpt_factor_compose: assumes "mpt M1 T1" "mpt.mpt_factor M1 T1 proj1 M2 T2" "mpt.mpt_factor M2 T2 proj2 M3 T3" shows "mpt.mpt_factor M1 T1 (proj2 o proj1) M3 T3" proof - have *: "proj1 \ measure_preserving M1 M2 \ AE x in M2. proj2 (T2 x) = T3 (proj2 x) \ (AE x in M1. proj1 (T1 x) = T2 (proj1 x) \ proj2 (T2 (proj1 x)) = T3 (proj2 (proj1 x)))" proof - assume "AE y in M2. proj2 (T2 y) = T3 (proj2 y)" "proj1 \ measure_preserving M1 M2" then have "AE x in M1. proj2 (T2 (proj1 x)) = T3 (proj2 (proj1 x))" using quasi_measure_preserving_AE measure_preserving_is_quasi_measure_preserving by blast moreover { fix x assume "proj2 (T2 (proj1 x)) = T3 (proj2 (proj1 x))" then have "proj1 (T1 x) = T2 (proj1 x) \ proj2 (T2 (proj1 x)) = T3 (proj2 (proj1 x))" by auto } ultimately show "AE x in M1. proj1 (T1 x) = T2 (proj1 x) \ proj2 (T2 (proj1 x)) = T3 (proj2 (proj1 x))" by auto qed interpret I: mpt M1 T1 using assms(1) by simp interpret J: mpt M2 T2 using I.mpt_factorE(3)[OF assms(2)] by simp show "I.mpt_factor (proj2 o proj1) M3 T3" apply (rule I.mpt_factorI) using I.mpt_factorE[OF assms(2)] J.mpt_factorE[OF assms(3)] by (auto simp add: measure_preserving_comp *) qed text \Left shifts are naturally factors of finite measure preserving transformations.\ lemma (in mpt) mpt_factor_projection: fixes f::"'a \ ('b::second_countable_topology)" assumes [measurable]: "f \ borel_measurable M" and "sigma_finite_measure (distr M borel (\x n. f ((T ^^ n) x)))" shows "mpt_factor (\x. (\n. f ((T^^n)x))) (distr M borel (\x. (\n. f ((T^^n)x)))) nat_left_shift" proof (rule mpt_factorI') have * [measurable]: "(\x. (\n. f ((T^^n)x))) \ borel_measurable M" using measurable_coordinatewise_then_product by measurable show "(\x n. f ((T ^^ n) x)) \ measure_preserving M (distr M borel (\x n. f ((T ^^ n) x)))" by (rule measure_preserving_distr'[OF *]) have "(\n. f ((T ^^ n) (T x))) = nat_left_shift (\n. f ((T ^^ n) x))" for x unfolding nat_left_shift_def by (auto simp add: funpow_swap1) then show "AE x in M. (\n. f ((T ^^ n) (T x))) = nat_left_shift (\n. f ((T ^^ n) x))" by simp qed (auto simp add: assms(2)) lemma (in fmpt) fmpt_factor_projection: fixes f::"'a \ ('b::second_countable_topology)" assumes [measurable]: "f \ borel_measurable M" shows "mpt_factor (\x. (\n. f ((T^^n)x))) (distr M borel (\x. (\n. f ((T^^n)x)))) nat_left_shift" proof (rule mpt_factor_projection, simp add: assms) have * [measurable]: "(\x. (\n. f ((T^^n)x))) \ borel_measurable M" using measurable_coordinatewise_then_product by measurable have **: "(\x n. f ((T ^^ n) x)) \ measure_preserving M (distr M borel (\x n. f ((T ^^ n) x)))" by (rule measure_preserving_distr'[OF *]) have a: "finite_measure (distr M borel (\x n. f ((T ^^ n) x)))" using measure_preserving_finite_measure[OF **] finite_measure_axioms by blast then show "sigma_finite_measure (distr M borel (\x n. f ((T ^^ n) x)))" by (simp add: finite_measure_def) qed subsection \Natural extension\ text \Many probability preserving dynamical systems are not invertible, while invertibility is often useful in proofs. The notion of natural extension is a solution to this problem: it shows that (essentially) any system has an extension which is invertible. This extension is constructed by considering the space of orbits indexed by integer numbers, with the left shift acting on it. If one considers the orbits starting from time $-N$ (for some fixed $N$), then there is a natural measure on this space: such an orbit is parameterized by its starting point at time $-N$, hence one may use the original measure on this point. The invariance of the measure ensures that these measures are compatible with each other. Their projective limit (when $N$ tends to infinity) is thus an invariant measure on the bilateral shift. The shift with this measure is the desired extension of the original system. There is a difficulty in the above argument: one needs to make sure that the projective limit of a system of compatible measures is well defined. This requires some topological conditions on the measures (they should be inner regular, i.e., the measure of any set should be approximated from below by compact subsets -- this is automatic on polish spaces). The existence of projective limits is proved in \verb+Projective_Limits.thy+ under the (sufficient) polish condition. We use this theory, so we need the underlying space to be a polish space and the measure to be a Borel measure. This is almost completely satisfactory. What is not completely satisfactory is that the completion of a Borel measure on a polish space (i.e., we add all subsets of sets of measure $0$ into the sigma algebra) does not fit into this setting, while this is an important framework in dynamical systems. It would readily follow once \verb+Projective_Limits.thy+ is extended to the more general inner regularity setting (the completion of a Borel measure on a polish space is always inner regular). \ locale polish_pmpt = pmpt "M::('a::polish_space measure)" T for M T + assumes M_eq_borel: "sets M = sets borel" begin definition natural_extension_map where "natural_extension_map = (int_left_shift::((int \ 'a) \ (int \ 'a)))" definition natural_extension_measure::"(int \ 'a) measure" where "natural_extension_measure = projective_family.lim UNIV (\I. distr M (\\<^sub>M i\I. borel) (\x. (\i\I. (T^^(nat(i- Min I))) x))) (\i. borel)" definition natural_extension_proj::"(int \ 'a) \ 'a" where "natural_extension_proj = (\x. x 0)" theorem natural_extension: "pmpt natural_extension_measure natural_extension_map" "qmpt.invertible_qmpt natural_extension_measure natural_extension_map" "mpt.mpt_factor natural_extension_measure natural_extension_map natural_extension_proj M T" proof - define P::"int set \ (int \ 'a) measure" where "P = (\I. distr M (\\<^sub>M i\I. borel) (\x. (\i\I. (T^^(nat(i- Min I))) x)))" have [measurable]: "(T^^n) \ measurable M borel" for n using M_eq_borel by auto interpret polish_projective UNIV P unfolding polish_projective_def projective_family_def proof (auto) show "prob_space (P I)" if "finite I" for I unfolding P_def by (rule prob_space_distr, auto) fix J H::"int set" assume "J \ H" "finite H" then have "H \ J = J" by blast have "((\f. restrict f J) o (\x. (\i\H. (T^^(nat(i- Min H))) x))) x = ((\x. (\i\J. (T^^(nat(i- Min J))) x)) o (T^^(nat(Min J - Min H)))) x" for x proof - have "nat(i- Min H) = nat(i- Min J) + nat(Min J - Min H)" if "i \ J" for i proof - have "finite J" using \J \ H\ \finite H\ finite_subset by auto then have "Min J \ J" using Min_in \i \ J\ by auto then have "Min J \ H" using \J \ H\ by blast then have "Min H \ Min J" using Min.coboundedI[OF \finite H\] by auto moreover have "Min J \ i" using Min.coboundedI[OF \finite J\ \i \ J\] by auto ultimately show ?thesis by auto qed then show ?thesis unfolding comp_def by (auto simp add: \H \ J = J\ funpow_add) qed then have *: "(\f. restrict f J) o (\x. (\i\H. (T^^(nat(i- Min H))) x)) = (\x. (\i\J. (T^^(nat(i- Min J))) x)) o (T^^(nat(Min J - Min H)))" by auto have "distr (P H) (Pi\<^sub>M J (\_. borel)) (\f. restrict f J) = distr M (\\<^sub>M i\J. borel) ((\f. restrict f J) o (\x. (\i\H. (T^^(nat(i- Min H))) x)))" unfolding P_def by (rule distr_distr, auto simp add: \J \ H\ measurable_restrict_subset) also have "... = distr M (\\<^sub>M i\J. borel) ((\x. (\i\J. (T^^(nat(i- Min J))) x)) o (T^^(nat(Min J - Min H))))" using * by auto also have "... = distr (distr M M (T^^(nat(Min J - Min H)))) (\\<^sub>M i\J. borel) (\x. (\i\J. (T^^(nat(i- Min J))) x))" by (rule distr_distr[symmetric], auto) also have "... = distr M (\\<^sub>M i\J. borel) (\x. (\i\J. (T^^(nat(i- Min J))) x))" using measure_preserving_distr[OF Tn_measure_preserving] by auto also have "... = P J" unfolding P_def by auto finally show "P J = distr (P H) (Pi\<^sub>M J (\_. borel)) (\f. restrict f J)" by simp qed have S: "sets (Pi\<^sub>M UNIV (\_. borel)) = sets (borel::(int \ 'a) measure)" by (rule sets_PiM_equal_borel) have "natural_extension_measure = lim" unfolding natural_extension_measure_def P_def by simp have "measurable lim lim = measurable borel borel" by (rule measurable_cong_sets, auto simp add: S) then have [measurable]: "int_left_shift \ measurable lim lim" "int_right_shift \ measurable lim lim" using int_shift_measurable by fast+ have [simp]: "space lim = UNIV" unfolding space_lim space_PiM space_borel by auto show "pmpt natural_extension_measure natural_extension_map" proof (rule pmpt_I) show "prob_space natural_extension_measure" unfolding \natural_extension_measure = lim\ by (simp add: P.prob_space_axioms) show "natural_extension_map \ measurable natural_extension_measure natural_extension_measure" unfolding natural_extension_map_def \natural_extension_measure = lim\ by simp define E where "E = {(\\<^sub>E i\UNIV. X i) |X::(int \ 'a set). (\i. X i \ sets borel) \ finite {i. X i \ UNIV}}" have "lim = distr lim lim int_left_shift" proof (rule measure_eqI_generator_eq[of E UNIV, where ?A = "\_. UNIV"]) show "sets lim = sigma_sets UNIV E" unfolding E_def using sets_PiM_finite[of "UNIV::int set" "\_. (borel::'a measure)"] by (simp add: PiE_def) moreover have "sets (distr lim lim int_left_shift) = sets lim" by auto ultimately show "sets (distr lim lim int_left_shift) = sigma_sets UNIV E" by simp show "emeasure lim UNIV \ \" by (simp add: P.prob_space_axioms) have "UNIV = (\\<^sub>E i\(UNIV::int set). (UNIV::'a set))" by (simp add: PiE_def) moreover have "... \ E" unfolding E_def by auto ultimately show "range (\(i::nat). (UNIV::(int \ 'a) set)) \ E" by auto show "Int_stable E" proof (rule Int_stableI) fix U V assume "U \ E" "V \ E" then obtain X Y where H: "U = (\\<^sub>E i\UNIV. X i)" "\i. X i \ sets borel" "finite {i. X i \ UNIV}" "V = (\\<^sub>E i\UNIV. Y i)" "\i. Y i \ sets borel" "finite {i. Y i \ UNIV}" unfolding E_def by blast define Z where "Z = (\i. X i \ Y i)" have "{i. Z i \ UNIV} \ {i. X i \ UNIV} \ {i. Y i \ UNIV}" unfolding Z_def by auto then have "finite {i. Z i \ UNIV}" using H(3) H(6) finite_subset by auto moreover have "U \ V = (\\<^sub>E i\UNIV. Z i)" unfolding Z_def using H(1) H(4) by auto moreover have "\i. Z i \ sets borel" unfolding Z_def using H(2) H(5) by auto ultimately show "U \ V \ E" unfolding E_def by auto qed fix U assume "U \ E" then obtain X where H [measurable]: "U = (\\<^sub>E i\UNIV. X i)" "\i. X i \ sets borel" "finite {i. X i \ UNIV}" unfolding E_def by blast define I where "I = {i. X i \ UNIV}" have [simp]: "finite I" unfolding I_def using H(3) by auto have [measurable]: "(\\<^sub>E i\I. X i) \ sets (Pi\<^sub>M I (\i. borel))" using H(2) by simp have *: "U = emb UNIV I (\\<^sub>E i\I. X i)" unfolding H(1) I_def prod_emb_def space_borel apply (auto simp add: PiE_def) by (metis (mono_tags, lifting) PiE UNIV_I mem_Collect_eq restrict_Pi_cancel) have "emeasure lim U = emeasure lim (int_left_shift-`U)" proof (cases "I = {}") case True then have "U = UNIV" unfolding H(1) I_def by auto then show ?thesis by auto next case False have "emeasure lim U = emeasure (P I) (\\<^sub>E i\I. X i)" unfolding * by (rule emeasure_lim_emb, auto) also have "... = emeasure M (((\x. (\i\I. (T^^(nat(i- Min I))) x)))-`(\\<^sub>E i\I. X i) \ space M)" unfolding P_def by (rule emeasure_distr, auto) finally have A: "emeasure lim U = emeasure M (((\x. (\i\I. (T^^(nat(i- Min I))) x)))-`(\\<^sub>E i\I. X i) \ space M)" by simp have i: "int_left_shift-`U = (\\<^sub>E i\UNIV. X (i-1))" unfolding H(1) apply (auto simp add: int_left_shift_def PiE_def) by (metis PiE UNIV_I diff_add_cancel, metis Pi_mem add.commute add_diff_cancel_left' iso_tuple_UNIV_I) define Im where "Im = {i. X (i-1) \ UNIV}" have "Im = (\i. i+1)`I" unfolding I_def Im_def using image_iff by (auto, fastforce) then have [simp]: "finite Im" by auto have *: "int_left_shift-`U = emb UNIV Im (\\<^sub>E i\Im. X (i-1))" unfolding i Im_def prod_emb_def space_borel apply (auto simp add: PiE_def) by (metis (mono_tags, lifting) PiE UNIV_I mem_Collect_eq restrict_Pi_cancel) have "emeasure lim (int_left_shift-`U) = emeasure (P Im) (\\<^sub>E i\Im. X (i-1))" unfolding * by (rule emeasure_lim_emb, auto) also have "... = emeasure M (((\x. (\i\Im. (T^^(nat(i- Min Im))) x)))-`(\\<^sub>E i\Im. X (i-1)) \ space M)" unfolding P_def by (rule emeasure_distr, auto) finally have B: "emeasure lim (int_left_shift-`U) = emeasure M (((\x. (\i\Im. (T^^(nat(i- Min Im))) x)))-`(\\<^sub>E i\Im. X (i-1)) \ space M)" by simp have "Min Im = Min I + 1" unfolding \Im = (\i. i+1)`I\ by (rule mono_Min_commute[symmetric], auto simp add: False monoI) have "((\x. (\i\Im. (T^^(nat(i- Min Im))) x)))-`(\\<^sub>E i\Im. X (i-1)) = ((\x. (\i\I. (T^^(nat(i- Min I))) x)))-`(\\<^sub>E i\I. X i)" unfolding \Min Im = Min I + 1\ unfolding \Im = (\i. i+1)`I\ by (auto simp add: Pi_iff) then show "emeasure lim U = emeasure lim (int_left_shift -` U)" using A B by auto qed also have "... = emeasure lim (int_left_shift-`U \ space lim)" unfolding \space lim = UNIV\ by auto also have "... = emeasure (distr lim lim int_left_shift) U" apply (rule emeasure_distr[symmetric], auto) using * by auto finally show "emeasure lim U = emeasure (distr lim lim int_left_shift) U" by simp qed (auto) fix U assume "U \ sets natural_extension_measure" then have [measurable]: "U \ sets lim" using \natural_extension_measure = lim\ by simp have "emeasure natural_extension_measure (natural_extension_map -` U \ space natural_extension_measure) = emeasure lim (int_left_shift-`U \ space lim)" unfolding \natural_extension_measure = lim\ natural_extension_map_def by simp also have "... = emeasure (distr lim lim int_left_shift) U" apply (rule emeasure_distr[symmetric], auto) using \U \ P.events\ by auto also have "... = emeasure lim U" using \lim = distr lim lim int_left_shift\ by simp also have "... = emeasure natural_extension_measure U" using \natural_extension_measure = lim\ by simp finally show "emeasure natural_extension_measure (natural_extension_map -` U \ space natural_extension_measure) = emeasure natural_extension_measure U" by simp qed then interpret I: pmpt natural_extension_measure natural_extension_map by simp show "I.invertible_qmpt" unfolding I.invertible_qmpt_def unfolding natural_extension_map_def \natural_extension_measure = lim\ by (auto simp add: int_shift_bij) show "I.mpt_factor natural_extension_proj M T" unfolding I.mpt_factor_def proof (auto) show "mpt M T" by (simp add: mpt_axioms) show "natural_extension_proj \ measure_preserving natural_extension_measure M" unfolding \natural_extension_measure = lim\ proof have *: "measurable lim M = measurable borel borel" apply (rule measurable_cong_sets) using sets_PiM_equal_borel M_eq_borel by auto show "natural_extension_proj \ measurable lim M" unfolding * natural_extension_proj_def by auto fix U assume [measurable]: "U \ sets M" have *: "(((\x. \i\{0}. (T ^^ nat (i - Min {0})) x))-` ({0} \\<^sub>E U) \ space M) = U" using sets.sets_into_space[OF \U \ sets M\] by auto have "natural_extension_proj-`U \ space lim = emb UNIV {0} (\\<^sub>E i\{0}. U)" unfolding \space lim = UNIV\ natural_extension_proj_def prod_emb_def by (auto simp add: PiE_iff) then have "emeasure lim (natural_extension_proj-`U \ space lim) = emeasure lim (emb UNIV {0} (\\<^sub>E i\{0}. U))" by simp also have "... = emeasure (P {0}) (\\<^sub>E i\{0}. U)" apply (rule emeasure_lim_emb, auto) using \U \ sets M\ M_eq_borel by auto also have "... = emeasure M (((\x. \i\{0}. (T ^^ nat (i - Min {0})) x))-` ({0} \\<^sub>E U) \ space M)" unfolding P_def apply (rule emeasure_distr) using \U \ sets M\ M_eq_borel by auto also have "... = emeasure M U" using * by simp finally show "emeasure lim (natural_extension_proj-`U \ space lim) = emeasure M U" by simp qed define U::"(int \ 'a) set" where "U = {x \ space (Pi\<^sub>M {0, 1} (\i. borel)). x 1 = T (x 0)}" have *: "((\x. \i\{0, 1}. (T ^^ nat (i - Min {0, 1})) x))-` U \ space M = space M" unfolding U_def space_PiM space_borel by auto have [measurable]: "T \ measurable borel borel" using M_eq_borel by auto have [measurable]: "U \ sets (Pi\<^sub>M {0, 1} (\i. borel))" unfolding U_def by (rule measurable_equality_set, auto) have "emeasure natural_extension_measure (emb UNIV {0, 1} U) = emeasure (P {0, 1}) U" unfolding \natural_extension_measure = lim\ by (rule emeasure_lim_emb, auto) also have "... = emeasure M (((\x. \i\{0, 1}. (T ^^ nat (i - Min {0, 1})) x))-` U \ space M)" unfolding P_def by (rule emeasure_distr, auto) also have "... = emeasure M (space M)" using * by simp also have "... = 1" by (simp add: emeasure_space_1) finally have *: "emeasure natural_extension_measure (emb UNIV {0, 1} U) = 1" by simp have "AE x in natural_extension_measure. x \ emb UNIV {0, 1} U" apply (rule I.AE_prob_1) using * by (simp add: I.emeasure_eq_measure) moreover { fix x assume "x \ emb UNIV {0, 1} U" then have "x 1 = T (x 0)" unfolding prod_emb_def U_def by auto then have "natural_extension_proj (natural_extension_map x) = T (natural_extension_proj x)" unfolding natural_extension_proj_def natural_extension_map_def int_left_shift_def by auto } ultimately show "AE x in natural_extension_measure. natural_extension_proj (natural_extension_map x) = T (natural_extension_proj x)" by auto qed qed end end (*of Measure_Preserving_Transformations.thy*) diff --git a/thys/Ergodic_Theory/SG_Library_Complement.thy b/thys/Ergodic_Theory/SG_Library_Complement.thy --- a/thys/Ergodic_Theory/SG_Library_Complement.thy +++ b/thys/Ergodic_Theory/SG_Library_Complement.thy @@ -1,1453 +1,1319 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \SG Libary complements\ theory SG_Library_Complement imports "HOL-Probability.Probability" begin text \In this file are included many statements that were useful to me, but belong rather naturally to existing theories. In a perfect world, some of these statements would get included into these files. I tried to indicate to which of these classical theories the statements could be added. \ subsection \Basic logic\ text \This one is certainly available, but I could not locate it...\ lemma equiv_neg: "\ P \ Q; \P \ \Q \ \ (P\Q)" by blast subsection \Basic set theory\ lemma compl_compl_eq_id [simp]: "UNIV - (UNIV - s) = s" by auto abbreviation sym_diff :: "'a set \ 'a set \ 'a set" (infixl "\" 70) where "sym_diff A B \ ((A - B) \ (B-A))" text \Not sure the next lemmas are useful, as they are proved solely by auto, so they could be reproved automatically whenever necessary.\ lemma sym_diff_inc: "A \ C \ A \ B \ B \ C" by auto lemma sym_diff_vimage [simp]: "f-`(A \ B) = (f-`A) \ (f-`B)" by auto -lemma card_finite_union: - assumes "finite I" - shows "card(\i\I. A i) \ (\i\I. card(A i))" -using assms apply (induct, auto) -using card_Un_le nat_add_left_cancel_le order_trans by blast - subsection \Set-Interval.thy\ text \The next two lemmas belong naturally to \verb+Set_Interval.thy+, next to \verb+UN_le_add_shift+. They are not trivially equivalent to the corresponding lemmas with large inequalities, due to the difference when $n = 0$.\ -lemma UN_le_add_shift_strict: - "(\ii\{k.. ?A" - proof - fix x assume "x \ ?B" - then obtain i where i: "i \ {k.. M(i)" by auto - then have "i - k < n \ x \ M((i-k) + k)" by auto - then show "x \ ?A" using UN_le_add_shift by blast - qed -qed (fastforce) - lemma UN_le_eq_Un0_strict: "(\ii\{1.. M 0" (is "?A = ?B") proof show "?A \ ?B" proof fix x assume "x \ ?A" then obtain i where i: "i M i" by auto show "x \ ?B" proof(cases i) case 0 with i show ?thesis by simp next case (Suc j) with i show ?thesis by auto qed qed qed (auto) text \I use repeatedly this one, but I could not find it directly\ lemma union_insert_0: "(\n::nat. A n) = A 0 \ (\n\{1..}. A n)" by (metis UN_insert Un_insert_left sup_bot.left_neutral One_nat_def atLeast_0 atLeast_Suc_greaterThan ivl_disj_un_singleton(1)) text \Next one could be close to \verb+sum.nat_group+\ lemma sum_arith_progression: "(\r<(N::nat). (\ijr j \ {i*N..riiri j \ {i*N..jMiscellanous basic results\ lemma ind_from_1 [case_names 1 Suc, consumes 1]: assumes "n > 0" assumes "P 1" and "\n. n > 0 \ P n \ P (Suc n)" shows "P n" proof - have "(n = 0) \ P n" proof (induction n) case 0 then show ?case by auto next case (Suc k) consider "Suc k = 1" | "Suc k > 1" by linarith then show ?case apply (cases) using assms Suc.IH by auto qed then show ?thesis using \n > 0\ by auto qed text \This lemma is certainly available somewhere, but I couldn't locate it\ lemma tends_to_real_e: fixes u::"nat \ real" - assumes "u \ l" - "e>0" + assumes "u \ l" "e>0" shows "\N. \n>N. abs(u n -l) < e" -proof- - have "eventually (\n. dist (u n) l < e) sequentially" using assms tendstoD by auto - then have "\\<^sub>F n in sequentially. abs(u n - l) < e" using dist_real_def by auto - then show ?thesis by (simp add: eventually_at_top_dense) -qed + by (metis assms dist_real_def le_less lim_sequentially) lemma nat_mod_cong: assumes "a = b+(c::nat)" "a mod n = b mod n" shows "c mod n = 0" proof - let ?k = "a mod n" obtain a1 where "a = a1*n + ?k" by (metis div_mult_mod_eq) moreover obtain b1 where "b = b1*n + ?k" using assms(2) by (metis div_mult_mod_eq) ultimately have "a1 * n + ?k = b1 * n + ?k + c" using assms(1) by arith then have "c = (a1 - b1) * n" by (simp add: diff_mult_distrib) then show ?thesis by simp qed lemma funpow_add': "(f ^^ (m + n)) x = (f ^^ m) ((f ^^ n) x)" by (simp add: funpow_add) text \The next two lemmas are not directly equivalent, since $f$ might not be injective.\ lemma abs_Max_sum: fixes A::"real set" assumes "finite A" "A \ {}" shows "abs(Max A) \ (\a\A. abs(a))" -using assms by (induct rule: finite_ne_induct, auto) + by (simp add: assms member_le_sum) lemma abs_Max_sum2: fixes f::"_ \ real" assumes "finite A" "A \ {}" shows "abs(Max (f`A)) \ (\a\A. abs(f a))" using assms by (induct rule: finite_ne_induct, auto) subsection \Conditionally-Complete-Lattices.thy\ lemma mono_cInf: fixes f :: "'a::conditionally_complete_lattice \ 'b::conditionally_complete_lattice" assumes "mono f" "A \ {}" "bdd_below A" shows "f(Inf A) \ Inf (f`A)" using assms by (simp add: cINF_greatest cInf_lower monoD) lemma mono_bij_cInf: fixes f :: "'a::conditionally_complete_linorder \ 'b::conditionally_complete_linorder" assumes "mono f" "bij f" "A \ {}" "bdd_below A" shows "f (Inf A) = Inf (f`A)" proof - have "(inv f) (Inf (f`A)) \ Inf ((inv f)`(f`A))" apply (rule cInf_greatest, auto simp add: assms(3)) using mono_inv[OF assms(1) assms(2)] assms by (simp add: mono_def bdd_below_image_mono cInf_lower) then have "Inf (f`A) \ f (Inf ((inv f)`(f`A)))" by (metis (no_types, lifting) assms(1) assms(2) mono_def bij_inv_eq_iff) also have "... = f(Inf A)" using assms by (simp add: bij_is_inj) finally show ?thesis using mono_cInf[OF assms(1) assms(3) assms(4)] by auto qed subsection \Topological-spaces.thy\ lemma open_less_abs [simp]: "open {x. (C::real) < abs x}" proof - have *: "{x. C < abs x} = abs-`{C<..}" by auto show ?thesis unfolding * by (auto intro!: continuous_intros) qed lemma closed_le_abs [simp]: "closed {x. (C::real) \ abs x}" proof - have *: "{x. C \ \x\} = abs-`{C..}" by auto show ?thesis unfolding * by (auto intro!: continuous_intros) qed text \The next statements come from the same statements for true subsequences\ lemma eventually_weak_subseq: fixes u::"nat \ nat" assumes "(\n. real(u n)) \ \" "eventually P sequentially" shows "eventually (\n. P (u n)) sequentially" proof - obtain N where *: "\n\N. P n" using assms(2) unfolding eventually_sequentially by auto obtain M where "\m\M. ereal(u m) \ N" using assms(1) by (meson Lim_PInfty) then have "\m. m \ M \ u m \ N" by auto then have "\m. m \ M \ P(u m)" using \\n\N. P n\ by simp then show ?thesis unfolding eventually_sequentially by auto qed lemma filterlim_weak_subseq: fixes u::"nat \ nat" assumes "(\n. real(u n)) \ \" shows "LIM n sequentially. u n:> at_top" unfolding filterlim_iff by (metis assms eventually_weak_subseq) lemma limit_along_weak_subseq: fixes u::"nat \ nat" and v::"nat \ _" assumes "(\n. real(u n)) \ \" "v \ l" shows "(\ n. v(u n)) \ l" using filterlim_compose[of v, OF _ filterlim_weak_subseq] assms by auto lemma frontier_indist_le: assumes "x \ frontier {y. infdist y S \ r}" shows "infdist x S = r" proof - have "infdist x S = r" if H: "\e>0. (\y. infdist y S \ r \ dist x y < e) \ (\z. \ infdist z S \ r \ dist x z < e)" proof - have "infdist x S < r + e" if "e > 0" for e proof - obtain y where "infdist y S \ r" "dist x y < e" using H \e > 0\ by blast then show ?thesis by (metis add.commute add_mono_thms_linordered_field(3) infdist_triangle le_less_trans) qed then have A: "infdist x S \ r" by (meson field_le_epsilon order.order_iff_strict) have "r < infdist x S + e" if "e > 0" for e proof - obtain y where "\(infdist y S \ r)" "dist x y < e" using H \e > 0\ by blast then have "r < infdist y S" by auto also have "... \ infdist x S + dist y x" by (rule infdist_triangle) finally show ?thesis using \dist x y < e\ by (simp add: dist_commute) qed then have B: "r \ infdist x S" by (meson field_le_epsilon order.order_iff_strict) show ?thesis using A B by auto qed then show ?thesis using assms unfolding frontier_straddle by auto qed -subsection \Connected.thy\ - -text \The next lemma is missing in the library, contrary to its cousin \verb+continuous_infdist+.\ - -lemma continuous_on_infdist [continuous_intros]: - assumes "continuous_on S f" - shows "continuous_on S (\x. infdist (f x) A)" -using assms unfolding continuous_on by (auto intro: tendsto_infdist) - - -subsection \Transcendental.thy\ - -text \In \verb+Transcendental.thy+, the assumptions of the next two lemmas are -$x > 0$ and $y > 0$, while large inequalities are sufficient, with the same proof.\ - -lemma powr_divide: "0 \ x \ 0 \ y \ (x / y) powr a = (x powr a) / (y powr a)" - for a b x :: real - apply (simp add: divide_inverse positive_imp_inverse_positive powr_mult) - apply (simp add: powr_def exp_minus [symmetric] exp_add [symmetric] ln_inverse) - done - -lemma powr_mult_base: "0 \ x \x * x powr y = x powr (1 + y)" - for x :: real - by (auto simp: powr_add) - - - subsection \Limits\ -text \\verb+lim_1_over_n+ is very useful, it could --should?-- be declared as simp and tendsto-intros.\ - -declare lim_1_over_n [tendsto_intros] - text \The next lemmas are not very natural, but I needed them several times\ lemma tendsto_shift_1_over_n [tendsto_intros]: fixes f::"nat \ real" assumes "(\n. f n / n) \ l" shows "(\n. f (n+k) / n) \ l" proof - have "(1+k*(1/n))* (f(n+k)/(n+k)) = f(n+k)/n" if "n>0" for n using that by (auto simp add: divide_simps) with eventually_mono[OF eventually_gt_at_top[of "0::nat"] this] have "eventually (\n.(1+k*(1/n))* (f(n+k)/(n+k)) = f(n+k)/n) sequentially" by auto moreover have "(\n. (1+k*(1/n))* (f(n+k)/(n+k))) \ (1+real k*0) * l" by (intro tendsto_intros LIMSEQ_ignore_initial_segment assms) ultimately show ?thesis using Lim_transform_eventually by auto qed lemma tendsto_shift_1_over_n' [tendsto_intros]: fixes f::"nat \ real" assumes "(\n. f n / n) \ l" shows "(\n. f (n-k) / n) \ l" proof - have "(1-k*(1/(n+k)))* (f n/ n) = f n/(n+k)" if "n>0" for n using that by (auto simp add: divide_simps) with eventually_mono[OF eventually_gt_at_top[of "0::nat"] this] have "eventually (\n. (1-k*(1/(n+k)))* (f n/ n) = f n/(n+k)) sequentially" by auto moreover have "(\n. (1-k*(1/(n+k)))* (f n/ n)) \ (1-real k*0) * l" by (intro tendsto_intros assms LIMSEQ_ignore_initial_segment) ultimately have "(\n. f n / (n+k)) \ l" using Lim_transform_eventually by auto then have a: "(\n. f(n-k)/(n-k+k)) \ l" using seq_offset_neg by auto have "f(n-k)/(n-k+k) = f(n-k)/n" if "n>k" for n using that by auto with eventually_mono[OF eventually_gt_at_top[of k] this] have "eventually (\n. f(n-k)/(n-k+k) = f(n-k)/n) sequentially" by auto with Lim_transform_eventually[OF a this] show ?thesis by auto qed declare LIMSEQ_realpow_zero [tendsto_intros] subsection \Topology-Euclidean-Space\ -lemma Inf_as_limit: - fixes A::"'a::{linorder_topology, first_countable_topology, complete_linorder} set" - assumes "A \ {}" - shows "\u. (\n. u n \ A) \ u \ Inf A" -proof (cases "Inf A \ A") - case True - show ?thesis - by (rule exI[of _ "\n. Inf A"], auto simp add: True) -next - case False - obtain y where "y \ A" using assms by auto - then have "Inf A < y" using False Inf_lower less_le by auto - obtain F :: "nat \ 'a set" where F: "\i. open (F i)" "\i. Inf A \ F i" - "\u. (\n. u n \ F n) \ u \ Inf A" - by (metis first_countable_topology_class.countable_basis) - define u where "u = (\n. SOME z. z \ F n \ z \ A)" - have "\z. z \ U \ z \ A" if "Inf A \ U" "open U" for U - proof - - obtain b where "b > Inf A" "{Inf A .. U" - using open_right[OF \open U\ \Inf A \ U\ \Inf A < y\] by auto - obtain z where "z < b" "z \ A" - using \Inf A < b\ Inf_less_iff by auto - then have "z \ {Inf A ..z \ A\ \{Inf A .. U\ by auto - qed - then have *: "u n \ F n \ u n \ A" for n - using \Inf A \ F n\ \open (F n)\ unfolding u_def by (metis (no_types, lifting) someI_ex) - then have "u \ Inf A" using F(3) by simp - then show ?thesis using * by auto -qed - text \A (more usable) variation around \verb+continuous_on_closure_sequentially+. The assumption that the spaces are metric spaces is definitely too strong, but sufficient for most applications.\ lemma continuous_on_closure_sequentially': fixes f::"'a::metric_space \ 'b::metric_space" assumes "continuous_on (closure C) f" "\(n::nat). u n \ C" "u \ l" shows "(\n. f (u n)) \ f l" proof - have "l \ closure C" unfolding closure_sequential using assms by auto then show ?thesis using \continuous_on (closure C) f\ unfolding comp_def continuous_on_closure_sequentially using assms by auto qed subsection \Convexity\ lemma convex_on_mean_ineq: fixes f::"real \ real" assumes "convex_on A f" "x \ A" "y \ A" shows "f ((x+y)/2) \ (f x + f y) / 2" using convex_onD[OF assms(1), of "1/2" x y] using assms by (auto simp add: divide_simps) lemma convex_on_closure: assumes "convex (C::'a::real_normed_vector set)" "convex_on C f" "continuous_on (closure C) f" shows "convex_on (closure C) f" proof (rule convex_onI) fix x y::'a and t::real assume "x \ closure C" "y \ closure C" "0 < t" "t < 1" obtain u v::"nat \ 'a" where *: "\n. u n \ C" "u \ x" "\n. v n \ C" "v \ y" using \x \ closure C\ \y \ closure C\ unfolding closure_sequential by blast define w where "w = (\n. (1-t) *\<^sub>R (u n) + t *\<^sub>R (v n))" have "w n \ C" for n using \0 < t\ \t< 1\ convexD[OF \convex C\ *(1)[of n] *(3)[of n]] unfolding w_def by auto have "w \ ((1-t) *\<^sub>R x + t *\<^sub>R y)" unfolding w_def using *(2) *(4) by (intro tendsto_intros) have *: "f(w n) \ (1-t) * f(u n) + t * f (v n)" for n using *(1) *(3) \convex_on C f\ \0 \t<1\ less_imp_le unfolding w_def convex_on_alt[OF assms(1)] by (simp add: add.commute) have i: "(\n. f (w n)) \ f ((1-t) *\<^sub>R x + t *\<^sub>R y)" by (rule continuous_on_closure_sequentially'[OF assms(3) \\n. w n \ C\ \w \ ((1-t) *\<^sub>R x + t *\<^sub>R y)\]) have ii: "(\n. (1-t) * f(u n) + t * f (v n)) \ (1-t) * f x + t * f y" apply (intro tendsto_intros) apply (rule continuous_on_closure_sequentially'[OF assms(3) \\n. u n \ C\ \u \ x\]) apply (rule continuous_on_closure_sequentially'[OF assms(3) \\n. v n \ C\ \v \ y\]) done show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" apply (rule LIMSEQ_le[OF i ii]) using * by auto qed lemma convex_on_norm [simp]: "convex_on UNIV (\(x::'a::real_normed_vector). norm x)" using convex_on_dist[of UNIV "0::'a"] by auto lemma continuous_abs_powr [continuous_intros]: assumes "p > 0" shows "continuous_on UNIV (\(x::real). \x\ powr p)" apply (rule continuous_on_powr') using assms by (auto intro: continuous_intros) lemma continuous_mult_sgn [continuous_intros]: fixes f::"real \ real" assumes "continuous_on UNIV f" "f 0 = 0" shows "continuous_on UNIV (\x. sgn x * f x)" proof - have *: "continuous_on {0..} (\x. sgn x * f x)" apply (subst continuous_on_cong[of "{0..}" "{0..}" _ f], auto simp add: sgn_real_def assms(2)) by (rule continuous_on_subset[OF assms(1)], auto) have **: "continuous_on {..0} (\x. sgn x * f x)" apply (subst continuous_on_cong[of "{..0}" "{..0}" _ "\x. -f x"], auto simp add: sgn_real_def assms(2)) by (rule continuous_on_subset[of UNIV], auto simp add: assms intro!: continuous_intros) show ?thesis using continuous_on_closed_Un[OF _ _ * **] apply (auto intro: continuous_intros) using continuous_on_subset by fastforce qed lemma DERIV_abs_powr [derivative_intros]: assumes "p > (1::real)" shows "DERIV (\x. \x\ powr p) x :> p * sgn x * \x\ powr (p - 1)" proof - consider "x = 0" | "x>0" | "x < 0" by linarith then show ?thesis proof (cases) case 1 have "continuous_on UNIV (\x. sgn x * \x\ powr (p - 1))" by (auto simp add: assms intro!:continuous_intros) then have "(\h. sgn h * \h\ powr (p-1)) \0\ (\h. sgn h * \h\ powr (p-1)) 0" using continuous_on_def by blast moreover have "\h\ powr p / h = sgn h * \h\ powr (p-1)" for h proof - have "\h\ powr p / h = sgn h * \h\ powr p / \h\" by (auto simp add: algebra_simps divide_simps sgn_real_def) also have "... = sgn h * \h\ powr (p-1)" using assms apply (cases "h = 0") apply (auto) by (metis abs_ge_zero powr_diff [symmetric] powr_one_gt_zero_iff times_divide_eq_right) finally show ?thesis by simp qed ultimately have "(\h. \h\ powr p / h) \0\ 0" by auto then show ?thesis unfolding DERIV_def by (auto simp add: \x = 0\) next case 2 have *: "\\<^sub>F y in nhds x. \y\ powr p = y powr p" unfolding eventually_nhds apply (rule exI[of _ "{0<..}"]) using \x > 0\ by auto show ?thesis apply (subst DERIV_cong_ev[of _ x _ "(\x. x powr p)" _ "p * x powr (p-1)"]) using \x > 0\ by (auto simp add: * has_real_derivative_powr) next case 3 have *: "\\<^sub>F y in nhds x. \y\ powr p = (-y) powr p" unfolding eventually_nhds apply (rule exI[of _ "{..<0}"]) using \x < 0\ by auto show ?thesis apply (subst DERIV_cong_ev[of _ x _ "(\x. (-x) powr p)" _ "p * (- x) powr (p - real 1) * - 1"]) using \x < 0\ apply (simp, simp add: *, simp) apply (rule DERIV_fun_powr[of "\y. -y" "-1" "x" p]) using \x < 0\ by (auto simp add: derivative_intros) qed qed lemma convex_abs_powr: assumes "p \ 1" shows "convex_on UNIV (\x::real. \x\ powr p)" proof (cases "p = 1") case True have "convex_on UNIV (\x::real. norm x)" by (rule convex_on_norm) moreover have "\x\ powr p = norm x" for x using True by auto ultimately show ?thesis by simp next case False then have "p > 1" using assms by auto define g where "g = (\x::real. p * sgn x * \x\ powr (p - 1))" have *: "DERIV (\x. \x\ powr p) x :> g x" for x unfolding g_def using \p>1\ by (intro derivative_intros) have **: "g x \ g y" if "x \ y" for x y proof - consider "x \ 0 \ y \ 0" | "x \ 0 \ y \ 0" | "x < 0 \ y > 0" using \x \ y\ by linarith then show ?thesis proof (cases) case 1 then show ?thesis unfolding g_def sgn_real_def using \p>1\ \x \ y\ by (auto simp add: powr_mono2) next case 2 then show ?thesis unfolding g_def sgn_real_def using \p>1\ \x \ y\ by (auto simp add: powr_mono2) next case 3 then have "g x \ 0" "0 \ g y" unfolding g_def using \p > 1\ by auto then show ?thesis by simp qed qed show ?thesis apply (rule convex_on_realI[of _ _ g]) using * ** by auto qed lemma convex_powr: assumes "p \ 1" shows "convex_on {0..} (\x::real. x powr p)" proof - have "convex_on {0..} (\x::real. \x\ powr p)" using convex_abs_powr[OF \p \ 1\] convex_on_subset by auto moreover have "\x\ powr p = x powr p" if "x \ {0..}" for x using that by auto ultimately show ?thesis by (simp add: convex_on_def) qed lemma convex_powr': assumes "p > 0" "p \ 1" shows "convex_on {0..} (\x::real. - (x powr p))" proof - have "convex_on {0<..} (\x::real. - (x powr p))" apply (rule convex_on_realI[of _ _ "\x. -p * x powr (p-1)"]) apply (auto intro!:derivative_intros simp add: has_real_derivative_powr) using \p > 0\ \p \ 1\ by (auto simp add: algebra_simps divide_simps powr_mono2') moreover have "continuous_on {0..} (\x::real. - (x powr p))" by (rule continuous_on_minus, rule continuous_on_powr', auto simp add: \p > 0\ intro!: continuous_intros) moreover have "{(0::real)..} = closure {0<..}" "convex {(0::real)<..}" by auto ultimately show ?thesis using convex_on_closure by metis qed lemma convex_fx_plus_fy_ineq: fixes f::"real \ real" assumes "convex_on {0..} f" "x \ 0" "y \ 0" "f 0 = 0" shows "f x + f y \ f (x+y)" proof - have *: "f a + f b \ f (a+b)" if "a \ 0" "b \ a" for a b proof (cases "a = 0") case False then have "a > 0" "b > 0" using \b \ a\ \a \ 0\ by auto have "(f 0 - f a) / (0 - a) \ (f 0 - f (a+b))/ (0 - (a+b))" apply (rule convex_on_diff[OF \convex_on {0..} f\]) using \a > 0\ \b > 0\ by auto also have "... \ (f b - f (a+b)) / (b - (a+b))" apply (rule convex_on_diff[OF \convex_on {0..} f\]) using \a > 0\ \b > 0\ by auto finally show ?thesis using \a > 0\ \b > 0\ \f 0 = 0\ by (auto simp add: divide_simps algebra_simps) qed (simp add: \f 0 = 0\) then show ?thesis using \x \ 0\ \y \ 0\ by (metis add.commute le_less not_le) qed lemma x_plus_y_p_le_xp_plus_yp: fixes p x y::real assumes "p > 0" "p \ 1" "x \ 0" "y \ 0" shows "(x + y) powr p \ x powr p + y powr p" using convex_fx_plus_fy_ineq[OF convex_powr'[OF \p > 0\ \p \ 1\] \x \ 0\ \y \ 0\] by auto subsection \Nonnegative-extended-real.thy\ lemma x_plus_top_ennreal [simp]: "x + \ = (\::ennreal)" by simp lemma ennreal_ge_nat_imp_PInf: fixes x::ennreal assumes "\N. x \ of_nat N" shows "x = \" using assms apply (cases x, auto) by (meson not_less reals_Archimedean2) lemma ennreal_archimedean: assumes "x \ (\::ennreal)" shows "\n::nat. x \ n" -proof (cases x) - case (real r) - then show ?thesis using real_arch_simple[of r] by auto -next - case top - then show ?thesis using assms by auto -qed + using assms ennreal_ge_nat_imp_PInf linear by blast lemma e2ennreal_mult: fixes a b::ereal assumes "a \ 0" shows "e2ennreal(a * b) = e2ennreal a * e2ennreal b" by (metis assms e2ennreal_neg eq_onp_same_args ereal_mult_le_0_iff linear times_ennreal.abs_eq) lemma e2ennreal_mult': fixes a b::ereal assumes "b \ 0" shows "e2ennreal(a * b) = e2ennreal a * e2ennreal b" using e2ennreal_mult[OF assms, of a] by (simp add: mult.commute) lemma SUP_real_ennreal: assumes "A \ {}" "bdd_above (f`A)" shows "(SUP a\A. ennreal (f a)) = ennreal(SUP a\A. f a)" apply (rule antisym, simp add: SUP_least assms(2) cSUP_upper ennreal_leI) by (metis assms(1) ennreal_SUP ennreal_less_top le_less) lemma e2ennreal_Liminf: "F \ bot \ e2ennreal (Liminf F f) = Liminf F (\n. e2ennreal (f n))" by (rule Liminf_compose_continuous_mono[symmetric]) (auto simp: mono_def e2ennreal_mono continuous_on_e2ennreal) lemma e2ennreal_eq_infty[simp]: "0 \ x \ e2ennreal x = top \ x = \" by (cases x) (auto) lemma ennreal_Inf_cmult: assumes "c>(0::real)" shows "Inf {ennreal c * x |x. P x} = ennreal c * Inf {x. P x}" proof - have "(\x::ennreal. c * x) (Inf {x::ennreal. P x}) = Inf ((\x::ennreal. c * x)`{x::ennreal. P x})" apply (rule mono_bij_Inf) apply (simp add: monoI mult_left_mono) apply (rule bij_betw_byWitness[of _ "\x. (x::ennreal) / c"], auto simp add: assms) apply (metis assms ennreal_lessI ennreal_neq_top mult.commute mult_divide_eq_ennreal not_less_zero) apply (metis assms divide_ennreal_def ennreal_less_zero_iff ennreal_neq_top less_irrefl mult.assoc mult.left_commute mult_divide_eq_ennreal) done then show ?thesis by (simp only: setcompr_eq_image[symmetric]) qed lemma continuous_on_const_minus_ennreal: fixes f :: "'a :: topological_space \ ennreal" shows "continuous_on A f \ continuous_on A (\x. a - f x)" including ennreal.lifting proof (transfer fixing: A; clarsimp) fix f :: "'a \ ereal" and a :: "ereal" assume "0 \ a" "\x. 0 \ f x" and f: "continuous_on A f" then show "continuous_on A (\x. max 0 (a - f x))" proof cases assume "\r. a = ereal r" with f show ?thesis by (auto simp: continuous_on_def minus_ereal_def ereal_Lim_uminus[symmetric] intro!: tendsto_add_ereal_general tendsto_max) next assume "\r. a = ereal r" with \0 \ a\ have "a = \" by (cases a) auto then show ?thesis by (simp add: continuous_on_const) qed qed lemma const_minus_Liminf_ennreal: fixes a :: ennreal shows "F \ bot \ a - Liminf F f = Limsup F (\x. a - f x)" by (intro Limsup_compose_continuous_antimono[symmetric]) (auto simp: antimono_def ennreal_mono_minus continuous_on_id continuous_on_const_minus_ennreal) -declare tendsto_ennrealI [tendsto_intros] - -lemma tendsto_mult_ennreal [tendsto_intros]: - fixes f g::"_ \ ennreal" - assumes "(f \ l) F" "(g \ m) F" "\((l = 0 \ m = \) \ (m = 0 \ l = \))" - shows "((\x. f x * g x) \ l * m) F" -proof - - have "((\x. enn2ereal(f x) * enn2ereal(g x)) \ enn2ereal l * enn2ereal m) F" - apply (auto intro!: tendsto_intros simp add: assms) - by (insert assms le_zero_eq less_eq_ennreal.rep_eq, fastforce)+ - then have "((\x. enn2ereal(f x * g x)) \ enn2ereal(l * m)) F" - using times_ennreal.rep_eq by auto - then show ?thesis - by (auto intro!: tendsto_intros) -qed - lemma tendsto_cmult_ennreal [tendsto_intros]: fixes c l::ennreal assumes "\(c = \ \ l = 0)" "(f \ l) F" shows "((\x. c * f x) \ c * l) F" by (cases "c = 0", insert assms, auto intro!: tendsto_intros) subsection \Indicator-Function.thy\ text \There is something weird with \verb+sum_mult_indicator+: it is defined both in Indicator.thy and BochnerIntegration.thy, with a different meaning. I am surprised there is no name collision... Here, I am using the version from BochnerIntegration.\ lemma sum_indicator_eq_card2: assumes "finite I" shows "(\i\I. (indicator (P i) x)::nat) = card {i\I. x \ P i}" using sum_mult_indicator [OF assms, of "\y. 1::nat" P "\y. x"] unfolding card_eq_sum by auto lemma disjoint_family_indicator_le_1: assumes "disjoint_family_on A I" shows "(\ i\ I. indicator (A i) x) \ (1::'a:: {comm_monoid_add,zero_less_one})" proof (cases "finite I") case True then have *: "(\ i\ I. indicator (A i) x) = ((indicator (\i\I. A i) x)::'a)" by (simp add: indicator_UN_disjoint[OF True assms(1), of x]) show ?thesis unfolding * unfolding indicator_def by (simp add: order_less_imp_le) next case False then show ?thesis by (simp add: order_less_imp_le) qed subsection \sigma-algebra.thy\ lemma algebra_intersection: assumes "algebra \ A" "algebra \ B" shows "algebra \ (A \ B)" apply (subst algebra_iff_Un) using assms by (auto simp add: algebra_iff_Un) lemma sigma_algebra_intersection: assumes "sigma_algebra \ A" "sigma_algebra \ B" shows "sigma_algebra \ (A \ B)" apply (subst sigma_algebra_iff) using assms by (auto simp add: sigma_algebra_iff algebra_intersection) lemma subalgebra_M_M [simp]: "subalgebra M M" by (simp add: subalgebra_def) text \The next one is \verb+disjoint_family_Suc+ with inclusions reversed.\ lemma disjoint_family_Suc2: assumes Suc: "\n. A (Suc n) \ A n" shows "disjoint_family (\i. A i - A (Suc i))" proof - have "A (m+n) \ A n" for m n proof (induct m) case 0 show ?case by simp next case (Suc m) then show ?case by (metis Suc_eq_plus1 assms add.commute add.left_commute subset_trans) qed then have "A m \ A n" if "m > n" for m n by (metis that add.commute le_add_diff_inverse nat_less_le) then show ?thesis by (auto simp add: disjoint_family_on_def) (metis insert_absorb insert_subset le_SucE le_antisym not_le_imp_less) qed subsection \Measure-Space.thy\ -lemma emeasure_union_inter: - assumes "A \ sets M" "B \ sets M" - shows "emeasure M A + emeasure M B = emeasure M (A \ B) + emeasure M (A \ B)" -proof - - have "A = (A-B) \ (A \ B)" by auto - then have 1: "emeasure M A = emeasure M (A-B) + emeasure M (A \ B)" - by (metis Diff_Diff_Int Diff_disjoint assms(1) assms(2) plus_emeasure sets.Diff) - - have "A \ B = (A-B) \ B" by auto - then have 2: "emeasure M (A \ B) = emeasure M (A-B) + emeasure M B" - by (metis Diff_disjoint Int_commute assms(1) assms(2) plus_emeasure sets.Diff) - - show ?thesis using 1 2 by (metis add.assoc add.commute) -qed - lemma AE_equal_sum: assumes "\i. AE x in M. f i x = g i x" shows "AE x in M. (\i\I. f i x) = (\i\I. g i x)" proof (cases) assume "finite I" have "\A. A \ null_sets M \ (\x\ (space M - A). f i x = g i x)" for i using assms(1)[of i] by (metis (mono_tags, lifting) AE_E3) then obtain A where A: "\i. A i \ null_sets M \ (\x\ (space M -A i). f i x = g i x)" by metis define B where "B = (\i\I. A i)" have "B \ null_sets M" using \finite I\ A B_def by blast then have "AE x in M. x \ space M - B" by (simp add: AE_not_in) moreover { fix x assume "x \ space M - B" then have "\i. i \ I \ f i x = g i x" unfolding B_def using A by auto then have "(\i\I. f i x) = (\i\I. g i x)" by auto } ultimately show ?thesis by auto qed (simp) lemma emeasure_pos_unionE: assumes "\ (N::nat). A N \ sets M" "emeasure M (\N. A N) > 0" shows "\N. emeasure M (A N) > 0" proof (rule ccontr) assume "\(\N. emeasure M (A N) > 0)" then have "\N. A N \ null_sets M" using assms(1) by auto then have "(\N. A N) \ null_sets M" by auto then show False using assms(2) by auto qed lemma (in prob_space) emeasure_intersection: fixes e::"nat \ real" assumes [measurable]: "\n. U n \ sets M" and [simp]: "\n. 0 \ e n" "summable e" and ge: "\n. emeasure M (U n) \ 1 - (e n)" shows "emeasure M (\n. U n) \ 1 - (\n. e n)" proof - define V where "V = (\n. space M - (U n))" have [measurable]: "V n \ sets M" for n unfolding V_def by auto have *: "emeasure M (V n) \ e n" for n unfolding V_def using ge[of n] by (simp add: emeasure_eq_measure prob_compl ennreal_leI) have "emeasure M (\n. V n) \ (\n. emeasure M (V n))" by (rule emeasure_subadditive_countably, auto) also have "... \ (\n. ennreal (e n))" using * by (intro suminf_le) auto also have "... = ennreal (\n. e n)" by (intro suminf_ennreal_eq) auto finally have "emeasure M (\n. V n) \ suminf e" by simp then have "1 - suminf e \ emeasure M (space M - (\n. V n))" by (simp add: emeasure_eq_measure prob_compl suminf_nonneg) also have "... \ emeasure M (\n. U n)" by (rule emeasure_mono) (auto simp: V_def) finally show ?thesis by simp qed lemma null_sym_diff_transitive: assumes "A \ B \ null_sets M" "B \ C \ null_sets M" and [measurable]: "A \ sets M" "C \ sets M" shows "A \ C \ null_sets M" proof - have "A \ B \ B \ C \ null_sets M" using assms(1) assms(2) by auto moreover have "A \ C \ A \ B \ B \ C" by auto ultimately show ?thesis by (meson null_sets_subset assms(3) assms(4) sets.Diff sets.Un) qed lemma Delta_null_of_null_is_null: assumes "B \ sets M" "A \ B \ null_sets M" "A \ null_sets M" shows "B \ null_sets M" proof - have "B \ A \ (A \ B)" by auto then show ?thesis using assms by (meson null_sets.Un null_sets_subset) qed lemma Delta_null_same_emeasure: assumes "A \ B \ null_sets M" and [measurable]: "A \ sets M" "B \ sets M" shows "emeasure M A = emeasure M B" proof - have "A = (A \ B) \ (A-B)" by blast moreover have "A-B \ null_sets M" using assms null_sets_subset by blast ultimately have a: "emeasure M A = emeasure M (A \ B)" using emeasure_Un_null_set by (metis assms(2) assms(3) sets.Int) have "B = (A \ B) \ (B-A)" by blast moreover have "B-A \ null_sets M" using assms null_sets_subset by blast ultimately have "emeasure M B = emeasure M (A \ B)" using emeasure_Un_null_set by (metis assms(2) assms(3) sets.Int) then show ?thesis using a by auto qed lemma AE_upper_bound_inf_ereal: fixes F G::"'a \ ereal" assumes "\e. (e::real) > 0 \ AE x in M. F x \ G x + e" shows "AE x in M. F x \ G x" proof - have "AE x in M. \n::nat. F x \ G x + ereal (1 / Suc n)" using assms by (auto simp: AE_all_countable) then show ?thesis proof (eventually_elim) fix x assume x: "\n::nat. F x \ G x + ereal (1 / Suc n)" show "F x \ G x" proof (intro ereal_le_epsilon2[of _ "G x"] allI impI) fix e :: real assume "0 < e" then obtain n where n: "1 / Suc n < e" by (blast elim: nat_approx_posE) have "F x \ G x + 1 / Suc n" using x by simp also have "\ \ G x + e" using n by (intro add_mono ennreal_leI) auto finally show "F x \ G x + ereal e" . qed qed qed text \Egorov theorem asserts that, if a sequence of functions converges almost everywhere to a limit, then the convergence is uniform on a subset of close to full measure. The first step in the proof is the following lemma, often useful by itself, asserting the same result for predicates: if a property $P_n x$ is eventually true for almost every $x$, then there exists $N$ such that $P_n x$ is true for all $n\geq N$ and all $x$ in a set of close to full measure. \ lemma (in finite_measure) Egorov_lemma: assumes [measurable]: "\n. (P n) \ measurable M (count_space UNIV)" and "AE x in M. eventually (\n. P n x) sequentially" "epsilon > 0" shows "\U N. U \ sets M \ (\n \ N. \x \ U. P n x) \ emeasure M (space M - U) < epsilon" proof - define K where "K = (\n. {x \ space M. \k\n. \(P k x)})" have [measurable]: "K n \ sets M" for n unfolding K_def by auto have "x \ (\n. K n)" if "eventually (\n. P n x) sequentially" for x unfolding K_def using that unfolding K_def eventually_sequentially by auto then have "AE x in M. x \ (\n. K n)" using assms by auto then have Z: "0 = emeasure M (\n. K n)" using AE_iff_measurable[of "(\n. K n)" M "\x. x \ (\n. K n)"] unfolding K_def by auto have *: "(\n. emeasure M (K n)) \ 0" unfolding Z apply (rule Lim_emeasure_decseq) using order_trans by (auto simp add: K_def decseq_def) have "eventually (\n. emeasure M (K n) < epsilon) sequentially" by (rule order_tendstoD(2)[OF * \epsilon > 0\]) then obtain N where N: "\n. n \ N \ emeasure M (K n) < epsilon" unfolding eventually_sequentially by auto define U where "U = space M - K N" have A [measurable]: "U \ sets M" unfolding U_def by auto have "space M - U = K N" unfolding U_def K_def by auto then have B: "emeasure M (space M - U) < epsilon" using N by auto have "\n \ N. \x \ U. P n x" unfolding U_def K_def by auto then show ?thesis using A B by blast qed text \The next lemma asserts that, in an uncountable family of disjoint sets, then there is one set with zero measure (and in fact uncountably many). It is often applied to the boundaries of $r$-neighborhoods of a given set, to show that one could choose $r$ for which this boundary has zero measure (this shows up often in relation with weak convergence).\ lemma (in finite_measure) uncountable_disjoint_family_then_exists_zero_measure: assumes [measurable]: "\i. i \ I \ A i \ sets M" and "uncountable I" "disjoint_family_on A I" shows "\i\I. measure M (A i) = 0" proof - define f where "f = (\(r::real). {i \ I. measure M (A i) > r})" have *: "finite (f r)" if "r > 0" for r proof - obtain N::nat where N: "measure M (space M)/r \ N" using real_arch_simple by blast have "finite (f r) \ card (f r) \ N" proof (rule finite_if_finite_subsets_card_bdd) fix G assume G: "G \ f r" "finite G" then have "G \ I" unfolding f_def by auto have "card G * r = (\i \ G. r)" by auto also have "... \ (\i \ G. measure M (A i))" apply (rule sum_mono) using G unfolding f_def by auto also have "... = measure M (\i\G. A i)" apply (rule finite_measure_finite_Union[symmetric]) using \finite G\ \G \ I\ \disjoint_family_on A I\ disjoint_family_on_mono by auto also have "... \ measure M (space M)" by (simp add: bounded_measure) finally have "card G \ measure M (space M)/r" using \r > 0\ by (simp add: divide_simps) then show "card G \ N" using N by auto qed then show ?thesis by simp qed have "countable (\n. f (((1::real)/2)^n))" by (rule countable_UN, auto intro!: countable_finite *) then have "I - (\n. f (((1::real)/2)^n)) \ {}" using assms(2) by (metis countable_empty uncountable_minus_countable) then obtain i where "i \ I" "i \ (\n. f ((1/2)^n))" by auto then have "measure M (A i) \ (1 / 2) ^ n" for n unfolding f_def using linorder_not_le by auto moreover have "(\n. ((1::real) / 2) ^ n) \ 0" by (intro tendsto_intros, auto) ultimately have "measure M (A i) \ 0" using LIMSEQ_le_const by force then have "measure M (A i) = 0" by (simp add: measure_le_0_iff) then show ?thesis using \i \ I\ by auto qed text \The next statements are useful measurability statements.\ lemma measurable_Inf [measurable]: assumes [measurable]: "\(n::nat). P n \ measurable M (count_space UNIV)" shows "(\x. Inf {n. P n x}) \ measurable M (count_space UNIV)" (is "?f \ _") proof - define A where "A = (\n. (P n)-`{True} \ space M - (\m space M))" have A_meas [measurable]: "A n \ sets M" for n unfolding A_def by measurable define B where "B = (\n. if n = 0 then (space M - (\n. A n)) else A (n-1))" show ?thesis proof (rule measurable_piecewise_restrict2[of B]) show "B n \ sets M" for n unfolding B_def by simp show "space M = (\n. B n)" unfolding B_def using sets.sets_into_space [OF A_meas] by auto have *: "?f x = n" if "x \ A n" for x n apply (rule cInf_eq_minimum) using that unfolding A_def by auto moreover have **: "?f x = (Inf ({}::nat set))" if "x \ space M - (\n. A n)" for x proof - have "\(P n x)" for n apply (induction n rule: nat_less_induct) using that unfolding A_def by auto then show ?thesis by simp qed ultimately have "\c. \x \ B n. ?f x = c" for n apply (cases "n = 0") unfolding B_def by auto then show "\h \ measurable M (count_space UNIV). \x \ B n. ?f x = h x" for n by fastforce qed qed lemma measurable_T_iter [measurable]: fixes f::"'a \ nat" assumes [measurable]: "T \ measurable M M" "f \ measurable M (count_space UNIV)" shows "(\x. (T^^(f x)) x) \ measurable M M" proof - have [measurable]: "(T^^n) \ measurable M M" for n::nat by (induction n, auto) show ?thesis by (rule measurable_compose_countable, auto) qed lemma measurable_infdist [measurable]: "(\x. infdist x S) \ borel_measurable borel" by (rule borel_measurable_continuous_onI, intro continuous_intros) text \The next lemma shows that, in a sigma finite measure space, sets with large measure can be approximated by sets with large but finite measure.\ lemma (in sigma_finite_measure) approx_with_finite_emeasure: assumes W_meas: "W \ sets M" and W_inf: "emeasure M W > C" obtains Z where "Z \ sets M" "Z \ W" "emeasure M Z < \" "emeasure M Z > C" proof (cases "emeasure M W = \") case True obtain r where r: "C = ennreal r" using W_inf by (cases C, auto) obtain Z where "Z \ sets M" "Z \ W" "emeasure M Z < \" "emeasure M Z > C" unfolding r using approx_PInf_emeasure_with_finite[OF W_meas True, of r] by auto then show ?thesis using that by blast next case False then have "W \ sets M" "W \ W" "emeasure M W < \" "emeasure M W > C" using assms apply auto using top.not_eq_extremum by blast then show ?thesis using that by blast qed subsection \Nonnegative-Lebesgue-Integration.thy\ text \The next lemma is a variant of \verb+nn_integral_density+, with the density on the right instead of the left, as seems more common.\ lemma nn_integral_densityR: assumes [measurable]: "f \ borel_measurable F" "g \ borel_measurable F" shows "(\\<^sup>+ x. f x * g x \F) = (\\<^sup>+ x. f x \(density F g))" proof - have "(\\<^sup>+ x. f x * g x \F) = (\\<^sup>+ x. g x * f x \F)" by (simp add: mult.commute) also have "... = (\\<^sup>+ x. f x \(density F g))" by (rule nn_integral_density[symmetric], simp_all add: assms) finally show ?thesis by simp qed lemma not_AE_zero_int_ennreal_E: fixes f::"'a \ ennreal" assumes "(\\<^sup>+x. f x \M) > 0" and [measurable]: "f \ borel_measurable M" shows "\A\sets M. \e::real>0. emeasure M A > 0 \ (\x \ A. f x \ e)" proof (rule not_AE_zero_ennreal_E, auto simp add: assms) assume *: "AE x in M. f x = 0" have "(\\<^sup>+x. f x \M) = (\\<^sup>+x. 0 \M)" by (rule nn_integral_cong_AE, simp add: *) then have "(\\<^sup>+x. f x \M) = 0" by simp then show False using assms by simp qed lemma (in finite_measure) nn_integral_bounded_eq_bound_then_AE: assumes "AE x in M. f x \ ennreal c" "(\\<^sup>+x. f x \M) = c * emeasure M (space M)" and [measurable]: "f \ borel_measurable M" shows "AE x in M. f x = c" proof (cases) assume "emeasure M (space M) = 0" then show ?thesis by (rule emeasure_0_AE) next assume "emeasure M (space M) \ 0" have fin: "AE x in M. f x \ top" using assms by (auto simp: top_unique) define g where "g = (\x. c - f x)" have [measurable]: "g \ borel_measurable M" unfolding g_def by auto have "(\\<^sup>+x. g x \M) = (\\<^sup>+x. c \M) - (\\<^sup>+x. f x \M)" unfolding g_def by (rule nn_integral_diff, auto simp add: assms ennreal_mult_eq_top_iff) also have "\ = 0" using assms(2) by (auto simp: ennreal_mult_eq_top_iff) finally have "AE x in M. g x = 0" by (subst nn_integral_0_iff_AE[symmetric]) auto then have "AE x in M. c \ f x" unfolding g_def using fin by (auto simp: ennreal_minus_eq_0) then show ?thesis using assms(1) by auto qed lemma null_sets_density: assumes [measurable]: "h \ borel_measurable M" and "AE x in M. h x \ 0" shows "null_sets (density M h) = null_sets M" proof - have *: "A \ sets M \ (AE x\A in M. h x = 0) \ A \ null_sets M" for A proof (auto) assume "A \ sets M" "AE x\A in M. h x = 0" then show "A \ null_sets M" unfolding AE_iff_null_sets[OF \A \ sets M\] using assms(2) by auto next assume "A \ null_sets M" then show "AE x\A in M. h x = 0" by (metis (mono_tags, lifting) AE_not_in eventually_mono) qed show ?thesis apply (rule set_eqI) unfolding null_sets_density_iff[OF \h \ borel_measurable M\] using * by auto qed text \The next proposition asserts that, if a function $h$ is integrable, then its integral on any set with small enough measure is small. The good conceptual proof is by considering the distribution of the function $h$ on $\mathbb{R}$ and looking at its tails. However, there is a less conceptual but more direct proof, based on dominated convergence and a proof by contradiction. This is the proof we give below.\ proposition integrable_small_integral_on_small_sets: fixes h::"'a \ real" assumes [measurable]: "integrable M h" and "delta > 0" shows "\epsilon>(0::real). \U \ sets M. emeasure M U < epsilon \ abs (\x\U. h x \M) < delta" proof (rule ccontr) assume H: "\ (\epsilon>0. \U\sets M. emeasure M U < ennreal epsilon \ abs(set_lebesgue_integral M U h) < delta)" have "\f. \epsilon\{0<..}. f epsilon \sets M \ emeasure M (f epsilon) < ennreal epsilon \ \(abs(set_lebesgue_integral M (f epsilon) h) < delta)" apply (rule bchoice) using H by auto then obtain f::"real \ 'a set" where f: "\epsilon. epsilon > 0 \ f epsilon \sets M" "\epsilon. epsilon > 0 \ emeasure M (f epsilon) < ennreal epsilon" "\epsilon. epsilon > 0 \ \(abs(set_lebesgue_integral M (f epsilon) h) < delta)" by blast define A where "A = (\n::nat. f ((1/2)^n))" have [measurable]: "A n \ sets M" for n unfolding A_def using f(1) by auto have *: "emeasure M (A n) < ennreal ((1/2)^n)" for n unfolding A_def using f(2) by auto have Large: "\(abs(set_lebesgue_integral M (A n) h) < delta)" for n unfolding A_def using f(3) by auto have S: "summable (\n. Sigma_Algebra.measure M (A n))" apply (rule summable_comparison_test'[of "\n. (1/2)^n" 0]) apply (rule summable_geometric, auto) apply (subst ennreal_le_iff[symmetric], simp) using less_imp_le[OF *] by (metis * emeasure_eq_ennreal_measure top.extremum_strict) have "AE x in M. eventually (\n. x \ space M - A n) sequentially" apply (rule borel_cantelli_AE1, auto simp add: S) by (metis * top.extremum_strict top.not_eq_extremum) moreover have "(\n. indicator (A n) x * h x) \ 0" if "eventually (\n. x \ space M - A n) sequentially" for x proof - have "eventually (\n. indicator (A n) x * h x = 0) sequentially" apply (rule eventually_mono[OF that]) unfolding indicator_def by auto then show ?thesis unfolding eventually_sequentially using lim_explicit by force qed ultimately have A: "AE x in M. ((\n. indicator (A n) x * h x) \ 0)" by auto have I: "integrable M (\x. abs(h x))" using \integrable M h\ by auto have L: "(\n. abs (\x. indicator (A n) x * h x \M)) \ abs (\x. 0 \M)" apply (intro tendsto_intros) apply (rule integral_dominated_convergence[OF _ _ I A]) unfolding indicator_def by auto have "eventually (\n. abs (\x. indicator (A n) x * h x \M) < delta) sequentially" apply (rule order_tendstoD[OF L]) using \delta > 0\ by auto then show False using Large by (auto simp: set_lebesgue_integral_def) qed text \We also give the version for nonnegative ennreal valued functions. It follows from the previous one.\ proposition small_nn_integral_on_small_sets: fixes h::"'a \ ennreal" assumes [measurable]: "h \ borel_measurable M" and "delta > (0::real)" "(\\<^sup>+x. h x \M) \ \" shows "\epsilon>(0::real). \U \ sets M. emeasure M U < epsilon \ (\\<^sup>+x\U. h x \M) < delta" proof - define f where "f = (\x. enn2real(h x))" have "AE x in M. h x \ \" using assms by (metis nn_integral_PInf_AE) then have *: "AE x in M. ennreal (f x) = h x" unfolding f_def using ennreal_enn2real_if by auto have **: "(\\<^sup>+x. ennreal (f x) \M) \ \" using nn_integral_cong_AE[OF *] assms by auto have [measurable]: "f \ borel_measurable M" unfolding f_def by auto have "integrable M f" apply (rule integrableI_nonneg) using assms * f_def ** apply auto using top.not_eq_extremum by blast obtain epsilon::real where H: "epsilon > 0" "\U. U \ sets M \ emeasure M U < epsilon \ abs(\x\U. f x \M) < delta" using integrable_small_integral_on_small_sets[OF \integrable M f\ \delta > 0\] by blast have "(\\<^sup>+x\U. h x \M) < delta" if [measurable]: "U \ sets M" "emeasure M U < epsilon" for U proof - have "(\\<^sup>+x. indicator U x * h x \M) = (\\<^sup>+x. ennreal(indicator U x * f x) \M)" apply (rule nn_integral_cong_AE) using * unfolding indicator_def by auto also have "... = ennreal (\x. indicator U x * f x \M)" apply (rule nn_integral_eq_integral) apply (rule Bochner_Integration.integrable_bound[OF \integrable M f\]) unfolding indicator_def f_def by auto also have "... < ennreal delta" apply (rule ennreal_lessI) using H(2)[OF that] by (auto simp: set_lebesgue_integral_def) finally show ?thesis by (auto simp add: mult.commute) qed then show ?thesis using \epsilon > 0\ by auto qed -subsection \Lebesgue-Measure.thy\ - -text \The next lemma is the same as \verb+lborel_distr_plus+ which is only formulated -on the real line, but on any euclidean space\ - -lemma lborel_distr_plus2: - fixes c :: "'a::euclidean_space" - shows "distr lborel borel ((+) c) = lborel" -by (subst lborel_affine[of 1 c], auto simp: density_1) - - subsection \Probability-measure.thy\ text \The next lemmas ensure that, if sets have a probability close to $1$, then their intersection also does.\ lemma (in prob_space) sum_measure_le_measure_inter: assumes "A \ sets M" "B \ sets M" shows "prob A + prob B \ 1 + prob (A \ B)" proof - have "prob A + prob B = prob (A \ B) + prob (A \ B)" - by (simp add: assms(1) assms(2) fmeasurable_eq_sets measure_Un3) + by (simp add: assms fmeasurable_eq_sets measure_Un3) also have "... \ 1 + prob (A \ B)" by auto finally show ?thesis by simp qed lemma (in prob_space) sum_measure_le_measure_inter3: assumes [measurable]: "A \ sets M" "B \ sets M" "C \ sets M" shows "prob A + prob B + prob C \ 2 + prob (A \ B \ C)" using sum_measure_le_measure_inter[of B C] sum_measure_le_measure_inter[of A "B \ C"] by (auto simp add: inf_assoc) lemma (in prob_space) sum_measure_le_measure_Inter: assumes [measurable]: "finite I" "I \ {}" "\i. i \ I \ A i \ sets M" shows "(\i\I. prob (A i)) \ real(card I) - 1 + prob (\i\I. A i)" using assms proof (induct I rule: finite_ne_induct) fix x F assume H: "finite F" "F \ {}" "x \ F" "((\i. i \ F \ A i \ events) \ (\i\F. prob (A i)) \ real (card F) - 1 + prob (\(A ` F)))" and [measurable]: "(\i. i \ insert x F \ A i \ events)" have "(\x\F. A x) \ events" using \finite F\ \F \ {}\ by auto have "(\i\insert x F. prob (A i)) = (\i\F. prob (A i)) + prob (A x)" using H(1) H(3) by auto also have "... \ real (card F)-1 + prob (\(A ` F)) + prob (A x)" using H(4) by auto also have "... \ real (card F) + prob ((\(A ` F)) \ A x)" using sum_measure_le_measure_inter[OF \(\x\F. A x) \ events\, of "A x"] by auto also have "... = real (card (insert x F)) - 1 + prob (\(A ` (insert x F)))" using H(1) H(2) unfolding card_insert_disjoint[OF \finite F\ \x \ F\] by (simp add: inf_commute) finally show "(\i\insert x F. prob (A i)) \ real (card (insert x F)) - 1 + prob (\(A ` (insert x F)))" by simp qed (auto) text \A random variable gives a small mass to small neighborhoods of infinity.\ lemma (in prob_space) random_variable_small_tails: assumes "alpha > 0" and [measurable]: "f \ borel_measurable M" shows "\(C::real). prob {x \ space M. abs(f x) \ C} < alpha \ C \ K" proof - have *: "(\(n::nat). {x\space M. abs(f x) \ n}) = {}" apply auto by (metis real_arch_simple add.right_neutral add_mono_thms_linordered_field(4) not_less zero_less_one) have **: "(\n. prob {x \ space M. abs(f x) \ n}) \ prob (\(n::nat). {x \ space M. abs(f x) \ n})" by (rule finite_Lim_measure_decseq, auto simp add: decseq_def) have "eventually (\n. prob {x \ space M. abs(f x) \ n} < alpha) sequentially" apply (rule order_tendstoD[OF _ \alpha > 0\]) using ** unfolding * by auto then obtain N::nat where N: "\n::nat. n \ N \ prob {x \ space M. abs(f x) \ n} < alpha" unfolding eventually_sequentially by blast have "\n::nat. n \ N \ n \ K" by (meson le_cases of_nat_le_iff order.trans real_arch_simple) then obtain n::nat where n: "n \ N" "n \ K" by blast show ?thesis apply (rule exI[of _ "of_nat n"]) using N n by auto qed subsection \Distribution-functions.thy\ text \There is a locale called \verb+finite_borel_measure+ in \verb+distribution-functions.thy+. However, it only deals with real measures, and real weak convergence. I will not need the weak convergence in more general settings, but still it seems more natural to me to do the proofs in the natural settings. Let me introduce the locale \verb+finite_borel_measure'+ for this, although it would be better to rename the locale in the library file.\ locale finite_borel_measure' = finite_measure M for M :: "('a::metric_space) measure" + assumes M_is_borel [simp, measurable_cong]: "sets M = sets borel" begin lemma space_eq_univ [simp]: "space M = UNIV" using M_is_borel[THEN sets_eq_imp_space_eq] by simp lemma measurable_finite_borel [simp]: "f \ borel_measurable borel \ f \ borel_measurable M" by (rule borel_measurable_subalgebra[where N = borel]) auto text \Any closed set can be slightly enlarged to obtain a set whose boundary has $0$ measure.\ lemma approx_closed_set_with_set_zero_measure_boundary: assumes "closed S" "epsilon > 0" "S \ {}" shows "\r. r < epsilon \ r > 0 \ measure M {x. infdist x S = r} = 0 \ measure M {x. infdist x S \ r} < measure M S + epsilon" proof - have [measurable]: "S \ sets M" using \closed S\ by auto define T where "T = (\r. {x. infdist x S \ r})" have [measurable]: "T r \ sets borel" for r unfolding T_def by measurable have *: "(\n. T ((1/2)^n)) = S" unfolding T_def proof (auto) fix x assume *: "\n. infdist x S \ (1 / 2) ^n" have "infdist x S \ 0" apply (rule LIMSEQ_le_const[of "\n. (1/2)^n"], intro tendsto_intros) using * by auto then show "x \ S" using assms infdist_pos_not_in_closed by fastforce qed have A: "((1::real)/2)^n \ (1/2)^m" if "m \ n" for m n::nat using that by (simp add: power_decreasing) have "(\n. measure M (T ((1/2)^n))) \ measure M S" unfolding *[symmetric] apply (rule finite_Lim_measure_decseq, auto simp add: T_def decseq_def) using A order.trans by blast then have B: "eventually (\n. measure M (T ((1/2)^n)) < measure M S + epsilon) sequentially" apply (rule order_tendstoD) using \epsilon > 0\ by simp have C: "eventually (\n. (1/2)^n < epsilon) sequentially" by (rule order_tendstoD[OF _ \epsilon > 0\], intro tendsto_intros, auto) obtain n where n: "(1/2)^n < epsilon" "measure M (T ((1/2)^n)) < measure M S + epsilon" using eventually_conj[OF B C] unfolding eventually_sequentially by auto have "\r\{0<..<(1/2)^n}. measure M {x. infdist x S = r} = 0" apply (rule uncountable_disjoint_family_then_exists_zero_measure, auto simp add: disjoint_family_on_def) using uncountable_open_interval by fastforce then obtain r where r: "r\{0<..<(1/2)^n}" "measure M {x. infdist x S = r} = 0" by blast then have r2: "r > 0" "r < epsilon" using n by auto have "measure M {x. infdist x S \ r} \ measure M {x. infdist x S \ (1/2)^n}" apply (rule finite_measure_mono) using r by auto then have "measure M {x. infdist x S \ r} < measure M S + epsilon" using n(2) unfolding T_def by auto then show ?thesis using r(2) r2 by auto qed end (* of locale finite_borel_measure'*) sublocale finite_borel_measure \ finite_borel_measure' by (standard, simp add: M_is_borel) subsection \Weak-convergence.thy\ text \Since weak convergence is not implemented as a topology, the fact that the convergence of a sequence implies the convergence of a subsequence is not automatic. We prove it in the lemma below..\ lemma weak_conv_m_subseq: assumes "weak_conv_m M_seq M" "strict_mono r" shows "weak_conv_m (\n. M_seq (r n)) M" using assms LIMSEQ_subseq_LIMSEQ unfolding weak_conv_m_def weak_conv_def comp_def by auto context fixes \ :: "nat \ real measure" and M :: "real measure" assumes \: "\n. real_distribution (\ n)" assumes M: "real_distribution M" assumes \_to_M: "weak_conv_m \ M" begin text \The measure of a closed set behaves upper semicontinuously with respect to weak convergence: if $\mu_n \to \mu$, then $\limsup \mu_n(F) \leq \mu(F)$ (and the inequality can be strict, think of the situation where $\mu$ is a Dirac mass at $0$ and $F = \{0\}$, but $\mu_n$ has a density so that $\mu_n(\{0\}) = 0$).\ lemma closed_set_weak_conv_usc: assumes "closed S" "measure M S < l" shows "eventually (\n. measure (\ n) S < l) sequentially" proof (cases "S = {}") case True then show ?thesis using \measure M S < l\ by auto next case False interpret real_distribution M using M by simp define epsilon where "epsilon = l - measure M S" have "epsilon > 0" unfolding epsilon_def using assms(2) by auto obtain r where r: "r > 0" "r < epsilon" "measure M {x. infdist x S = r} = 0" "measure M {x. infdist x S \ r} < measure M S + epsilon" using approx_closed_set_with_set_zero_measure_boundary[OF \closed S\ \epsilon > 0\ \S \ {}\] by blast define T where "T = {x. infdist x S \ r}" have [measurable]: "T \ sets borel" unfolding T_def by auto have "S \ T" unfolding T_def using \closed S\ \r > 0\ by auto have "measure M T < l" using r(4) unfolding T_def epsilon_def by auto have "measure M (frontier T) \ measure M {x. infdist x S = r}" apply (rule finite_measure_mono) unfolding T_def using frontier_indist_le by auto then have "measure M (frontier T) = 0" using \measure M {x. infdist x S = r} = 0\ by (auto simp add: measure_le_0_iff) then have "(\n. measure (\ n) T) \ measure M T" using \_to_M by (simp add: \ emeasure_eq_measure real_distribution_axioms weak_conv_imp_continuity_set_conv) then have *: "eventually (\n. measure (\ n) T < l) sequentially" apply (rule order_tendstoD) using \measure M T < l\ by simp have **: "measure (\ n) S \ measure (\ n) T" for n apply (rule finite_measure.finite_measure_mono) using \ apply (simp add: finite_borel_measure.axioms(1) real_distribution.finite_borel_measure_M) using \S \ T\ apply simp by (simp add: \ real_distribution.events_eq_borel) show ?thesis apply (rule eventually_mono[OF *]) using ** le_less_trans by auto qed text \In the same way, the measure of an open set behaves lower semicontinuously with respect to weak convergence: if $\mu_n \to \mu$, then $\liminf \mu_n(U) \geq \mu(U)$ (and the inequality can be strict). This follows from the same statement for closed sets by passing to the complement.\ lemma open_set_weak_conv_lsc: assumes "open S" "measure M S > l" shows "eventually (\n. measure (\ n) S > l) sequentially" proof - interpret real_distribution M using M by auto have [measurable]: "S \ events" using assms(1) by auto have "eventually (\n. measure (\ n) (UNIV - S) < 1 - l) sequentially" apply (rule closed_set_weak_conv_usc) using assms prob_compl[of S] by auto moreover have "measure (\ n) (UNIV - S) = 1 - measure (\ n) S" for n proof - interpret mu: real_distribution "\ n" using \ by auto have "S \ mu.events" using assms(1) by auto then show ?thesis using mu.prob_compl[of S] by auto qed ultimately show ?thesis by auto qed end (*of context weak_conv_m*) end (*of SG_Library_Complement.thy*) diff --git a/thys/Gromov_Hyperbolicity/Library_Complements.thy b/thys/Gromov_Hyperbolicity/Library_Complements.thy --- a/thys/Gromov_Hyperbolicity/Library_Complements.thy +++ b/thys/Gromov_Hyperbolicity/Library_Complements.thy @@ -1,1765 +1,1760 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Additions to the library\ theory Library_Complements imports "HOL-Analysis.Analysis" "HOL-Cardinals.Cardinal_Order_Relation" begin subsection \Mono intros\ text \We have a lot of (large) inequalities to prove. It is very convenient to have a set of introduction rules for this purpose (a lot should be added to it, I have put here all the ones I needed). The typical use case is when one wants to prove some inequality, say $ \exp (x*x) \leq y + \exp(1 + z * z + y)$, assuming $y \geq 0$ and $0 \leq x \leq z$. One would write it has \begin{verbatim} have "0 + \exp(0 + x * x + 0) < = y + \exp(1 + z * z + y)" using `y > = 0` `x < = z` by (intro mono_intros) \end{verbatim} When the left and right hand terms are written in completely analogous ways as above, then the introduction rules (that contain monotonicity of addition, of the exponential, and so on) reduce this to comparison of elementary terms in the formula. This is a very naive strategy, that fails in many situations, but that is very efficient when used correctly. \ named_theorems mono_intros "structural introduction rules to prove inequalities" declare le_imp_neg_le [mono_intros] declare add_left_mono [mono_intros] declare add_right_mono [mono_intros] declare add_strict_left_mono [mono_intros] declare add_strict_right_mono [mono_intros] declare add_mono [mono_intros] declare add_less_le_mono [mono_intros] declare diff_right_mono [mono_intros] declare diff_left_mono [mono_intros] declare diff_mono [mono_intros] declare mult_left_mono [mono_intros] declare mult_right_mono [mono_intros] declare mult_mono [mono_intros] declare max.mono [mono_intros] declare min.mono [mono_intros] declare power_mono [mono_intros] declare ln_ge_zero [mono_intros] declare ln_le_minus_one [mono_intros] declare ennreal_minus_mono [mono_intros] declare ennreal_leI [mono_intros] declare e2ennreal_mono [mono_intros] declare enn2ereal_nonneg [mono_intros] declare zero_le [mono_intros] declare top_greatest [mono_intros] declare bot_least [mono_intros] declare dist_triangle [mono_intros] declare dist_triangle2 [mono_intros] declare dist_triangle3 [mono_intros] declare exp_ge_add_one_self [mono_intros] declare exp_gt_one [mono_intros] declare exp_less_mono [mono_intros] declare dist_triangle [mono_intros] declare abs_triangle_ineq [mono_intros] declare abs_triangle_ineq2 [mono_intros] declare abs_triangle_ineq2_sym [mono_intros] declare abs_triangle_ineq3 [mono_intros] declare abs_triangle_ineq4 [mono_intros] declare Liminf_le_Limsup [mono_intros] declare ereal_liminf_add_mono [mono_intros] declare le_of_int_ceiling [mono_intros] declare ereal_minus_mono [mono_intros] declare infdist_triangle [mono_intros] declare divide_right_mono [mono_intros] declare self_le_power [mono_intros] lemma ln_le_cancelI [mono_intros]: assumes "(0::real) < x" "x \ y" shows "ln x \ ln y" using assms by auto lemma exp_le_cancelI [mono_intros]: assumes "x \ (y::real)" shows "exp x \ exp y" using assms by simp lemma mult_ge1_mono [mono_intros]: assumes "a \ (0::'a::linordered_idom)" "b \ 1" shows "a \ a * b" "a \ b * a" using assms mult_le_cancel_left1 mult_le_cancel_right1 by force+ text \A few convexity inequalities we will need later on.\ lemma xy_le_uxx_vyy [mono_intros]: assumes "u > 0" "u * v = (1::real)" shows "x * y \ u * x^2/2 + v * y^2/2" proof - have "v > 0" using assms by (metis (full_types) dual_order.strict_implies_order le_less_linear mult_nonneg_nonpos not_one_le_zero) then have *: "sqrt u * sqrt v = 1" using assms by (metis real_sqrt_mult real_sqrt_one) have "(sqrt u * x - sqrt v * y)^2 \ 0" by auto then have "u * x^2 + v * y^2 - 2 * 1 * x * y \ 0" unfolding power2_eq_square *[symmetric] using \u > 0\ \v > 0\ by (auto simp add: algebra_simps) then show ?thesis by (auto simp add: algebra_simps divide_simps) qed lemma xy_le_xx_yy [mono_intros]: "x * y \ x^2/2 + y^2/2" for x y::real using xy_le_uxx_vyy[of 1 1] by auto lemma ln_squared_bound [mono_intros]: "(ln x)^2 \ 2 * x - 2" if "x \ 1" for x::real proof - define f where "f = (\x::real. 2 * x - 2 - ln x * ln x)" have *: "DERIV f x :> 2 - 2 * ln x / x" if "x > 0" for x::real unfolding f_def using that by (auto intro!: derivative_eq_intros) have "f 1 \ f x" if "x \ 1" for x proof (rule DERIV_nonneg_imp_nondecreasing[OF that]) fix t::real assume "t \ 1" show "\y. (f has_real_derivative y) (at t) \ 0 \ y" apply (rule exI[of _ "2 - 2 * ln t / t"]) using *[of t] \t \ 1\ by (auto simp add: divide_simps ln_bound) qed then show ?thesis unfolding f_def power2_eq_square using that by auto qed text \In the next lemma, the assumptions are too strong (negative numbers less than $-1$ also work well to have a square larger than $1$), but in practice one proves inequalities with nonnegative numbers, so this version is really the useful one for \verb+mono_intros+.\ lemma mult_ge1_powers [mono_intros]: assumes "a \ (1::'a::linordered_idom)" shows "1 \ a * a" "1 \ a * a * a" "1 \ a * a * a * a" using assms by (meson assms dual_order.trans mult_ge1_mono(1) zero_le_one)+ lemmas [mono_intros] = ln_bound lemma mono_cSup: fixes f :: "'a::conditionally_complete_lattice \ 'b::conditionally_complete_lattice" assumes "bdd_above A" "A \ {}" "mono f" shows "Sup (f`A) \ f (Sup A)" by (metis assms(1) assms(2) assms(3) cSUP_least cSup_upper mono_def) lemma mono_cSup_bij: fixes f :: "'a::conditionally_complete_linorder \ 'b::conditionally_complete_linorder" assumes "bdd_above A" "A \ {}" "mono f" "bij f" shows "Sup (f`A) = f(Sup A)" proof - have "Sup ((inv f)`(f`A)) \ (inv f) (Sup (f`A))" apply (rule mono_cSup) using mono_inv[OF assms(3) assms(4)] assms(2) bdd_above_image_mono[OF assms(3) assms(1)] by auto then have "f (Sup ((inv f)`(f`A))) \ Sup (f`A)" using assms mono_def by (metis (no_types, hide_lams) bij_betw_imp_surj_on surj_f_inv_f) moreover have "f (Sup ((inv f)`(f`A))) = f(Sup A)" using assms by (simp add: bij_is_inj) ultimately show ?thesis using mono_cSup[OF assms(1) assms(2) assms(3)] by auto qed subsection \More topology\ text \In situations of interest to us later on, convergence is well controlled only for sequences living in some dense subset of the space (but the limit can be anywhere). This is enough to establish continuity of the function, if the target space is well enough separated. The statement we give below is very general, as we do not assume that the function is continuous inside the original set $S$, it will typically only be continuous at a set $T$ contained in the closure of $S$. In many applications, $T$ will be the closure of $S$, but we are also thinking of the case where one constructs an extension of a function inside a space, to its boundary, and the behaviour at the boundary is better than inside the space. The example we have in mind is the extension of a quasi-isometry to the boundary of a Gromov hyperbolic space. In the following criterion, we assume that if $u_n$ inside $S$ converges to a point at the boundary $T$, then $f(u_n)$ converges (where $f$ is some function inside). Then, we can extend the function $f$ at the boundary, by picking the limit value of $f(u_n)$ for some sequence converging to $u_n$. Then the lemma asserts that $f$ is continuous at every point $b$ on the boundary. The proof is done in two steps: \begin{enumerate} \item First, if $v_n$ is another inside sequence tending to the same point $b$ on the boundary, then $f(v_n)$ converges to the same value as $f(u_n)$: this is proved by considering the sequence $w$ equal to $u$ at even times and to $v$ at odd times, and saying that $f(w_n)$ converges. Its limit is equal to the limit of $f(u_n)$ and of $f(v_n)$, so they have to coincide. \item Now, consider a general sequence $v$ (in the space or the boundary) converging to $b$. We want to show that $f(v_n)$ tends to $f(b)$. If $v_n$ is inside $S$, we have already done it in the first step. If it is on the boundary, on the other hand, we can approximate it by an inside point $w_n$ for which $f(w_n)$ is very close to $f(v_n)$. Then $w_n$ is an inside sequence converging to $b$, hence $f(w_n)$ converges to $f(b)$ by the first step, and then $f(v_n)$ also converges to $f(b)$. The precise argument is more conveniently written by contradiction. It requires good separation properties of the target space. \end{enumerate}\ text \First, we introduce the material to interpolate between two sequences, one at even times and the other one at odd times.\ definition even_odd_interpolate::"(nat \ 'a) \ (nat \ 'a) \ (nat \ 'a)" where "even_odd_interpolate u v n = (if even n then u (n div 2) else v (n div 2))" lemma even_odd_interpolate_compose: "even_odd_interpolate (f o u) (f o v) = f o (even_odd_interpolate u v)" unfolding even_odd_interpolate_def comp_def by auto lemma even_odd_interpolate_filterlim: "filterlim u F sequentially \ filterlim v F sequentially \ filterlim (even_odd_interpolate u v) F sequentially" proof (auto) assume H: "filterlim (even_odd_interpolate u v) F sequentially" define r::"nat \ nat" where "r = (\n. 2 * n)" have "strict_mono r" unfolding r_def strict_mono_def by auto then have "filterlim r sequentially sequentially" by (simp add: filterlim_subseq) have "filterlim (\n. (even_odd_interpolate u v) (r n)) F sequentially" by (rule filterlim_compose[OF H filterlim_subseq[OF \strict_mono r\]]) moreover have "even_odd_interpolate u v (r n) = u n" for n unfolding r_def even_odd_interpolate_def by auto ultimately show "filterlim u F sequentially" by auto define r::"nat \ nat" where "r = (\n. 2 * n + 1)" have "strict_mono r" unfolding r_def strict_mono_def by auto then have "filterlim r sequentially sequentially" by (simp add: filterlim_subseq) have "filterlim (\n. (even_odd_interpolate u v) (r n)) F sequentially" by (rule filterlim_compose[OF H filterlim_subseq[OF \strict_mono r\]]) moreover have "even_odd_interpolate u v (r n) = v n" for n unfolding r_def even_odd_interpolate_def by auto ultimately show "filterlim v F sequentially" by auto next assume H: "filterlim u F sequentially" "filterlim v F sequentially" show "filterlim (even_odd_interpolate u v) F sequentially" unfolding filterlim_iff eventually_sequentially proof (auto) fix P assume *: "eventually P F" obtain N1 where N1: "\n. n \ N1 \ P (u n)" using H(1) unfolding filterlim_iff eventually_sequentially using * by auto obtain N2 where N2: "\n. n \ N2 \ P (v n)" using H(2) unfolding filterlim_iff eventually_sequentially using * by auto have "P (even_odd_interpolate u v n)" if "n \ 2 * N1 + 2 * N2" for n proof (cases "even n") case True have "n div 2 \ N1" using that by auto then show ?thesis unfolding even_odd_interpolate_def using True N1 by auto next case False have "n div 2 \ N2" using that by auto then show ?thesis unfolding even_odd_interpolate_def using False N2 by auto qed then show "\N. \n \ N. P (even_odd_interpolate u v n)" by auto qed qed text \Then, we prove the continuity criterion for extensions of functions to the boundary $T$ of a set $S$. The first assumption is that $f(u_n)$ converges when $f$ converges to the boundary, and the second one that the extension of $f$ to the boundary has been defined using the limit along some sequence tending to the point under consideration. The following criterion is the most general one, but this is not the version that is most commonly applied so we use a prime in its name.\ lemma continuous_at_extension_sequentially': fixes f :: "'a::{first_countable_topology, t2_space} \ 'b::t3_space" assumes "b \ T" "\u b. (\n. u n \ S) \ b \ T \ u \ b \ convergent (\n. f (u n))" "\b. b \ T \ \u. (\n. u n \ S) \ u \ b \ ((\n. f (u n)) \ f b)" shows "continuous (at b within (S \ T)) f" proof - have first_step: "(\n. f (u n)) \ f c" if "\n. u n \ S" "u \ c" "c \ T" for u c proof - obtain v where v: "\n. v n \ S" "v \ c" "(\n. f (v n)) \ f c" using assms(3)[OF \c \ T\] by blast then have A: "even_odd_interpolate u v \ c" unfolding even_odd_interpolate_filterlim[symmetric] using \u \ c\ by auto moreover have B: "\n. even_odd_interpolate u v n \ S" using \\n. u n \ S\ \\n. v n \ S\ unfolding even_odd_interpolate_def by auto have "convergent (\n. f (even_odd_interpolate u v n))" by (rule assms(2)[OF B \c \ T\ A]) then obtain m where "(\n. f (even_odd_interpolate u v n)) \ m" unfolding convergent_def by auto then have "even_odd_interpolate (f o u) (f o v) \ m" unfolding even_odd_interpolate_compose unfolding comp_def by auto then have "(f o u) \ m" "(f o v) \ m" unfolding even_odd_interpolate_filterlim[symmetric] by auto then have "m = f c" using v(3) unfolding comp_def using LIMSEQ_unique by auto then show ?thesis using \(f o u) \ m\ unfolding comp_def by auto qed show "continuous (at b within (S \ T)) f" proof (rule ccontr) assume "\ ?thesis" then obtain U where U: "open U" "f b \ U" "\(\\<^sub>F x in at b within S \ T. f x \ U)" unfolding continuous_within tendsto_def[where l = "f b"] using sequentially_imp_eventually_nhds_within by auto have "\V W. open V \ open W \ f b \ V \ (UNIV - U) \ W \ V \ W = {}" apply (rule t3_space) using U by auto then obtain V W where VW: "open V" "open W" "f b \ V" "UNIV - U \ W" "V \ W = {}" by auto obtain A :: "nat \ 'a set" where *: "\i. open (A i)" "\i. b \ A i" "\F. \n. F n \ A n \ F \ b" by (rule first_countable_topology_class.countable_basis) blast with * U(3) have "\F. \n. F n \ S \ T \ F n \ A n \ \ (f(F n) \ U)" unfolding at_within_def eventually_inf_principal eventually_nhds by (intro choice) (meson DiffE) then obtain F where F: "\n. F n \ S \ T" "\n. F n \ A n" "\n. f(F n) \ U" by auto have "\y. y \ S \ y \ A n \ f y \ W" for n proof (cases "F n \ S") case True show ?thesis apply (rule exI[of _ "F n"]) using F VW True by auto next case False then have "F n \ T" using \F n \ S \ T\ by auto obtain u where u: "\p. u p \ S" "u \ F n" "(\p. f (u p)) \ f(F n)" using assms(3)[OF \F n \ T\] by auto moreover have "f(F n) \ W" using F VW by auto ultimately have "eventually (\p. f (u p) \ W) sequentially" using \open W\ by (simp add: tendsto_def) moreover have "eventually (\p. u p \ A n) sequentially" using \F n \ A n\ u \open (A n)\ by (simp add: tendsto_def) ultimately have "\p. f(u p) \ W \ u p \ A n" using eventually_False_sequentially eventually_elim2 by blast then show ?thesis using u(1) by auto qed then have "\u. \n. u n \ S \ u n \ A n \ f (u n) \ W" by (auto intro: choice) then obtain u where u: "\n. u n \ S" "\n. u n \ A n" "\n. f (u n) \ W" by blast then have "u \ b" using *(3) by auto then have "(\n. f (u n)) \ f b" using first_step assms u by auto then have "eventually (\n. f (u n) \ V) sequentially" using VW by (simp add: tendsto_def) then have "\n. f (u n) \ V" using eventually_False_sequentially eventually_elim2 by blast then show False using u(3) \V \ W = {}\ by auto qed qed text \We can specialize the previous statement to the common case where one already knows the sequential continuity of $f$ along sequences in $S$ converging to a point in $T$. This will be the case in most --but not all-- applications. This is a straightforward application of the above criterion.\ proposition continuous_at_extension_sequentially: fixes f :: "'a::{first_countable_topology, t2_space} \ 'b::t3_space" assumes "a \ T" "T \ closure S" "\u b. (\n. u n \ S) \ b \ T \ u \ b \ (\n. f (u n)) \ f b" shows "continuous (at a within (S \ T)) f" apply (rule continuous_at_extension_sequentially'[OF \a \ T\]) using assms(3) convergent_def apply blast by (metis assms(2) assms(3) closure_sequential subset_iff) text \We also give global versions. We can only express the continuity on $T$, so this is slightly weaker than the previous statements since we are not saying anything on inside sequences tending to $T$ -- but in cases where $T$ contains $S$ these statements contain all the information.\ lemma continuous_on_extension_sequentially': fixes f :: "'a::{first_countable_topology, t2_space} \ 'b::t3_space" assumes "\u b. (\n. u n \ S) \ b \ T \ u \ b \ convergent (\n. f (u n))" "\b. b \ T \ \u. (\n. u n \ S) \ u \ b \ ((\n. f (u n)) \ f b)" shows "continuous_on T f" unfolding continuous_on_eq_continuous_within apply (auto intro!: continuous_within_subset[of _ "S \ T" f T]) by (intro continuous_at_extension_sequentially'[OF _ assms], auto) lemma continuous_on_extension_sequentially: fixes f :: "'a::{first_countable_topology, t2_space} \ 'b::t3_space" assumes "T \ closure S" "\u b. (\n. u n \ S) \ b \ T \ u \ b \ (\n. f (u n)) \ f b" shows "continuous_on T f" unfolding continuous_on_eq_continuous_within apply (auto intro!: continuous_within_subset[of _ "S \ T" f T]) by (intro continuous_at_extension_sequentially[OF _ assms], auto) subsubsection \Homeomorphisms\ lemma homeomorphism_cong: "homeomorphism X' Y' f' g'" if "homeomorphism X Y f g" "X' = X" "Y' = Y" "\x. x \ X \ f' x = f x" "\y. y \ Y \ g' y = g y" using that by (auto simp add: homeomorphism_def) lemma homeomorphism_empty [simp]: "homeomorphism {} {} f g" unfolding homeomorphism_def by auto text \A variant around the notion of homeomorphism, which is only expressed in terms of the function and not of its inverse.\ definition homeomorphism_on::"'a set \ ('a::topological_space \ 'b::topological_space) \ bool" where "homeomorphism_on S f = (\g. homeomorphism S (f`S) f g)" lemma homeomorphism_on_continuous: assumes "homeomorphism_on S f" shows "continuous_on S f" using assms unfolding homeomorphism_on_def homeomorphism_def by auto lemma homeomorphism_on_bij: assumes "homeomorphism_on S f" shows "bij_betw f S (f`S)" using assms unfolding homeomorphism_on_def homeomorphism_def by auto (metis inj_on_def inj_on_imp_bij_betw) lemma homeomorphism_on_homeomorphic: assumes "homeomorphism_on S f" shows "S homeomorphic (f`S)" using assms unfolding homeomorphism_on_def homeomorphic_def by auto lemma homeomorphism_on_compact: fixes f::"'a::topological_space \ 'b::t2_space" assumes "continuous_on S f" "compact S" "inj_on f S" shows "homeomorphism_on S f" unfolding homeomorphism_on_def using homeomorphism_compact[OF assms(2) assms(1) _ assms(3)] by auto lemma homeomorphism_on_subset: assumes "homeomorphism_on S f" "T \ S" shows "homeomorphism_on T f" using assms homeomorphism_of_subsets unfolding homeomorphism_on_def by blast lemma homeomorphism_on_empty [simp]: "homeomorphism_on {} f" unfolding homeomorphism_on_def using homeomorphism_empty[of f] by auto lemma homeomorphism_on_cong: assumes "homeomorphism_on X f" "X' = X" "\x. x \ X \ f' x = f x" shows "homeomorphism_on X' f'" proof - obtain g where g:"homeomorphism X (f`X) f g" using assms unfolding homeomorphism_on_def by auto have "homeomorphism X' (f'`X') f' g" apply (rule homeomorphism_cong[OF g]) using assms by (auto simp add: rev_image_eqI) then show ?thesis unfolding homeomorphism_on_def by auto qed lemma homeomorphism_on_inverse: fixes f::"'a::topological_space \ 'b::topological_space" assumes "homeomorphism_on X f" shows "homeomorphism_on (f`X) (inv_into X f)" proof - obtain g where g: "homeomorphism X (f`X) f g" using assms unfolding homeomorphism_on_def by auto then have "g`f`X = X" by (simp add: homeomorphism_def) then have "homeomorphism_on (f`X) g" unfolding homeomorphism_on_def using homeomorphism_symD[OF g] by auto moreover have "g x = inv_into X f x" if "x \ f`X" for x using g that unfolding homeomorphism_def by (auto, metis f_inv_into_f inv_into_into that) ultimately show ?thesis using homeomorphism_on_cong by force qed text \Characterization of homeomorphisms in terms of sequences: a map is a homeomorphism if and only if it respects convergent sequences.\ lemma homeomorphism_on_compose: assumes "homeomorphism_on S f" "x \ S" "eventually (\n. u n \ S) F" shows "(u \ x) F \ ((\n. f (u n)) \ f x) F" proof assume "(u \ x) F" then show "((\n. f (u n)) \ f x) F" using continuous_on_tendsto_compose[OF homeomorphism_on_continuous[OF assms(1)] _ assms(2) assms(3)] by simp next assume *: "((\n. f (u n)) \ f x) F" have I: "inv_into S f (f y) = y" if "y \ S" for y using homeomorphism_on_bij[OF assms(1)] by (meson bij_betw_inv_into_left that) then have A: "eventually (\n. u n = inv_into S f (f (u n))) F" using assms eventually_mono by force have "((\n. (inv_into S f) (f (u n))) \ (inv_into S f) (f x)) F" apply (rule continuous_on_tendsto_compose[OF homeomorphism_on_continuous[OF homeomorphism_on_inverse[OF assms(1)]] *]) using assms eventually_mono by (auto) fastforce then show "(u \ x) F" unfolding tendsto_cong[OF A] I[OF \x \ S\] by simp qed lemma homeomorphism_on_sequentially: fixes f::"'a::{first_countable_topology, t2_space} \ 'b::{first_countable_topology, t2_space}" assumes "\x u. x \ S \ (\n. u n \ S) \ u \ x \ (\n. f (u n)) \ f x" shows "homeomorphism_on S f" proof - have "x = y" if "f x = f y" "x \ S" "y \ S" for x y proof - have "(\n. f x) \ f y" using that by auto then have "(\n. x) \ y" using assms(1) that by auto then show "x = y" using LIMSEQ_unique by auto qed then have "inj_on f S" by (simp add: inj_on_def) have Cf: "continuous_on S f" apply (rule continuous_on_sequentiallyI) using assms by auto define g where "g = inv_into S f" have Cg: "continuous_on (f`S) g" proof (rule continuous_on_sequentiallyI) fix v b assume H: "\n. v n \ f ` S" "b \ f ` S" "v \ b" define u where "u = (\n. g (v n))" define a where "a = g b" have "u n \ S" "f (u n) = v n" for n unfolding u_def g_def using H(1) by (auto simp add: inv_into_into f_inv_into_f) have "a \ S" "f a = b" unfolding a_def g_def using H(2) by (auto simp add: inv_into_into f_inv_into_f) show "(\n. g(v n)) \ g b" unfolding u_def[symmetric] a_def[symmetric] apply (rule iffD2[OF assms]) using \\n. u n \ S\ \a \ S\ \v \ b\ unfolding \\n. f (u n) = v n\ \f a = b\ by auto qed have "homeomorphism S (f`S) f g" apply (rule homeomorphismI[OF Cf Cg]) unfolding g_def using \inj_on f S\ by auto then show ?thesis unfolding homeomorphism_on_def by auto qed lemma homeomorphism_on_UNIV_sequentially: fixes f::"'a::{first_countable_topology, t2_space} \ 'b::{first_countable_topology, t2_space}" assumes "\x u. u \ x \ (\n. f (u n)) \ f x" shows "homeomorphism_on UNIV f" using assms by (auto intro!: homeomorphism_on_sequentially) text \Now, we give similar characterizations in terms of sequences living in a dense subset. As in the sequential continuity criteria above, we first give a very general criterion, where the map does not have to be continuous on the approximating set $S$, only on the limit set $T$, without any a priori identification of the limit. Then, we specialize this statement to a less general but often more usable version.\ lemma homeomorphism_on_extension_sequentially_precise: fixes f::"'a::{first_countable_topology, t3_space} \ 'b::{first_countable_topology, t3_space}" assumes "\u b. (\n. u n \ S) \ b \ T \ u \ b \ convergent (\n. f (u n))" "\u c. (\n. u n \ S) \ c \ f`T \ (\n. f (u n)) \ c \ convergent u" "\b. b \ T \ \u. (\n. u n \ S) \ u \ b \ ((\n. f (u n)) \ f b)" "\n. u n \ S \ T" "l \ T" shows "u \ l \ (\n. f (u n)) \ f l" proof assume H: "u \ l" have "continuous (at l within (S \ T)) f" apply (rule continuous_at_extension_sequentially'[OF \l \ T\]) using assms(1) assms(3) by auto then show "(\n. f (u n)) \ f l" apply (rule continuous_within_tendsto_compose) using H assms(4) by auto next text \For the reverse implication, we would like to use the continuity criterion \verb+ continuous_at_extension_sequentially'+ applied to the inverse of $f$. Unfortunately, this inverse is only well defined on $T$, while our sequence takes values in $S \cup T$. So, instead, we redo by hand the proof of the continuity criterion, but in the opposite direction.\ assume H: "(\n. f (u n)) \ f l" show "u \ l" proof (rule ccontr) assume "\ ?thesis" then obtain U where U: "open U" "l \ U" "\(\\<^sub>F n in sequentially. u n \ U)" unfolding continuous_within tendsto_def[where l = l] using sequentially_imp_eventually_nhds_within by auto obtain A :: "nat \ 'b set" where *: "\i. open (A i)" "\i. f l \ A i" "\F. \n. F n \ A n \ F \ f l" by (rule first_countable_topology_class.countable_basis) blast have B: "eventually (\n. f (u n) \ A i) sequentially" for i using \open (A i)\ \f l \ A i\ H topological_tendstoD by fastforce have M: "\r. r \ N \ (u r \ U) \ f (u r) \ A i" for N i using U(3) B[of i] unfolding eventually_sequentially by (meson dual_order.trans le_cases) have "\r. \n. (u (r n) \ U \ f (u (r n)) \ A n) \ r (Suc n) \ r n + 1" apply (rule dependent_nat_choice) using M by auto then obtain r where r: "\n. u (r n) \ U" "\n. f (u (r n)) \ A n" "\n. r (Suc n) \ r n + 1" by auto then have "strict_mono r" by (metis Suc_eq_plus1 Suc_le_lessD strict_monoI_Suc) have "\V W. open V \ open W \ l \ V \ (UNIV - U) \ W \ V \ W = {}" apply (rule t3_space) using U by auto then obtain V W where VW: "open V" "open W" "l \ V" "UNIV - U \ W" "V \ W = {}" by auto have "\z. z \ S \ f z \ A n \ z \ W" for n proof - define z where "z = u (r n)" have "f z \ A n" unfolding z_def using r(2) by auto have "z \ S \ T" "z \ U" unfolding z_def using r(1) assms(4) by auto then have "z \ W" using VW by auto show ?thesis proof (cases "z \ T") case True obtain u::"nat \ 'a" where u: "\p. u p \ S" "u \ z" "(\p. f (u p)) \ f z" using assms(3)[OF \z \ T\] by auto then have "eventually (\p. f (u p) \ A n) sequentially" using \open (A n)\ \f z \ A n\ unfolding tendsto_def by simp moreover have "eventually (\p. u p \ W) sequentially" using \open W\ \z \ W\ u unfolding tendsto_def by simp ultimately have "\p. u p \ W \ f (u p) \ A n" using eventually_False_sequentially eventually_elim2 by blast then show ?thesis using u(1) by auto next case False then have "z \ S" using \z \ S \ T\ by auto then show ?thesis using \f z \ A n\ \z \ W\ by auto qed qed then have "\v. \n. v n \ S \ f (v n) \ A n \ v n \ W" by (auto intro: choice) then obtain v where v: "\n. v n \ S" "\n. f (v n) \ A n" "\n. v n \ W" by blast then have I: "(\n. f (v n)) \ f l" using *(3) by auto obtain w where w: "\n. w n \ S" "w \ l" "((\n. f (w n)) \ f l)" using assms(3)[OF \l \ T\] by auto have "even_odd_interpolate (f o v) (f o w) \ f l" unfolding even_odd_interpolate_filterlim[symmetric] comp_def using v w I by auto then have *: "(\n. f (even_odd_interpolate v w n)) \ f l" unfolding even_odd_interpolate_compose unfolding comp_def by auto have "convergent (even_odd_interpolate v w)" apply (rule assms(2)[OF _ _ *]) unfolding even_odd_interpolate_def using v(1) w(1) \l \ T\ by auto then obtain z where "even_odd_interpolate v w \ z" unfolding convergent_def by auto then have *: "v \ z" "w \ z" unfolding even_odd_interpolate_filterlim[symmetric] by auto then have "z = l" using v(2) w(2) LIMSEQ_unique by auto then have "v \ l" using * by simp then have "eventually (\n. v n \ V) sequentially" using VW by (simp add: tendsto_def) then have "\n. v n \ V" using eventually_False_sequentially eventually_elim2 by blast then show False using v(3) \V \ W = {}\ by auto qed qed lemma homeomorphism_on_extension_sequentially': fixes f::"'a::{first_countable_topology, t3_space} \ 'b::{first_countable_topology, t3_space}" assumes "\u b. (\n. u n \ S) \ b \ T \ u \ b \ convergent (\n. f (u n))" "\u c. (\n. u n \ S) \ c \ f`T \ (\n. f (u n)) \ c \ convergent u" "\b. b \ T \ \u. (\n. u n \ S) \ u \ b \ ((\n. f (u n)) \ f b)" shows "homeomorphism_on T f" apply (rule homeomorphism_on_sequentially, rule homeomorphism_on_extension_sequentially_precise[of S T]) using assms by auto proposition homeomorphism_on_extension_sequentially: fixes f::"'a::{first_countable_topology, t3_space} \ 'b::{first_countable_topology, t3_space}" assumes "\u b. (\n. u n \ S) \ u \ b \ (\n. f (u n)) \ f b" "T \ closure S" shows "homeomorphism_on T f" apply (rule homeomorphism_on_extension_sequentially'[of S]) using assms(1) convergent_def apply fastforce using assms(1) convergent_def apply blast by (metis assms(1) assms(2) closure_sequential subsetCE) lemma homeomorphism_on_UNIV_extension_sequentially: fixes f::"'a::{first_countable_topology, t3_space} \ 'b::{first_countable_topology, t3_space}" assumes "\u b. (\n. u n \ S) \ u \ b \ (\n. f (u n)) \ f b" "closure S = UNIV" shows "homeomorphism_on UNIV f" apply (rule homeomorphism_on_extension_sequentially[of S]) using assms by auto subsubsection \Proper spaces\ text \Proper spaces, i.e., spaces in which every closed ball is compact -- or, equivalently, any closed bounded set is compact.\ definition proper::"('a::metric_space) set \ bool" where "proper S \ (\ x r. compact (cball x r \ S))" lemma properI: assumes "\x r. compact (cball x r \ S)" shows "proper S" using assms unfolding proper_def by auto lemma proper_compact_cball: assumes "proper (UNIV::'a::metric_space set)" shows "compact (cball (x::'a) r)" using assms unfolding proper_def by auto lemma proper_compact_bounded_closed: assumes "proper (UNIV::'a::metric_space set)" "closed (S::'a set)" "bounded S" shows "compact S" proof - obtain x r where "S \ cball x r" using \bounded S\ bounded_subset_cball by blast then have *: "S = S \ cball x r" by auto show ?thesis apply (subst *, rule closed_Int_compact) using assms unfolding proper_def by auto qed lemma proper_real [simp]: "proper (UNIV::real set)" unfolding proper_def by auto lemma complete_of_proper: assumes "proper S" shows "complete S" proof - have "\l\S. u \ l" if "Cauchy u" "\n. u n \ S" for u proof - have "bounded (range u)" using \Cauchy u\ cauchy_imp_bounded by auto then obtain x r where *: "\n. dist x (u n) \ r" unfolding bounded_def by auto then have "u n \ (cball x r) \ S" for n using \u n \ S\ by auto moreover have "complete ((cball x r) \ S)" apply (rule compact_imp_complete) using assms unfolding proper_def by auto ultimately show ?thesis unfolding complete_def using \Cauchy u\ by auto qed then show ?thesis unfolding complete_def by auto qed lemma proper_of_compact: assumes "compact S" shows "proper S" using assms by (auto intro: properI) lemma proper_Un: assumes "proper A" "proper B" shows "proper (A \ B)" using assms unfolding proper_def by (auto simp add: compact_Un inf_sup_distrib1) subsubsection \Miscellaneous topology\ text \When manipulating the triangle inequality, it is very frequent to deal with 4 points (and automation has trouble doing it automatically). Even sometimes with 5 points...\ lemma dist_triangle4 [mono_intros]: "dist x t \ dist x y + dist y z + dist z t" using dist_triangle[of x z y] dist_triangle[of x t z] by auto lemma dist_triangle5 [mono_intros]: "dist x u \ dist x y + dist y z + dist z t + dist t u" using dist_triangle4[of x u y z] dist_triangle[of z u t] by auto text \A thickening of a compact set is closed.\ lemma compact_has_closed_thickening: assumes "compact C" "continuous_on C f" shows "closed (\x\C. cball x (f x))" proof (auto simp add: closed_sequential_limits) fix u l assume *: "\n::nat. \x\C. dist x (u n) \ f x" "u \ l" have "\x::nat\'a. \n. x n \ C \ dist (x n) (u n) \ f (x n)" apply (rule choice) using * by auto then obtain x::"nat \ 'a" where x: "\n. x n \ C" "\n. dist (x n) (u n) \ f (x n)" by blast obtain r c where "strict_mono r" "c \ C" "(x o r) \ c" using x(1) \compact C\ by (meson compact_eq_seq_compact_metric seq_compact_def) then have "c \ C" using x(1) \compact C\ by auto have lim: "(\n. f (x (r n)) - dist (x (r n)) (u (r n))) \ f c - dist c l" apply (intro tendsto_intros, rule continuous_on_tendsto_compose[of C f]) using *(2) x(1) \(x o r) \ c\ \continuous_on C f\ \c \ C\ \strict_mono r\ LIMSEQ_subseq_LIMSEQ unfolding comp_def by auto have "f c - dist c l \ 0" apply (rule LIMSEQ_le_const[OF lim]) using x(2) by auto then show "\x\C. dist x l \ f x" using \c \ C\ by auto qed text \congruence rule for continuity. The assumption that $f y = g y$ is necessary since \verb+at x+ is the pointed neighborhood at $x$.\ lemma continuous_within_cong: assumes "continuous (at y within S) f" "eventually (\x. f x = g x) (at y within S)" "f y = g y" shows "continuous (at y within S) g" using assms(1) assms(2) Lim_transform_eventually unfolding continuous_within assms(3) by auto text \A function which tends to infinity at infinity, on a proper set, realizes its infimum\ lemma continuous_attains_inf_proper: fixes f :: "'a::metric_space \ 'b::linorder_topology" assumes "proper s" "a \ s" "continuous_on s f" "\z. z \ s - cball a r \ f z \ f a" shows "\x\s. \y\s. f x \ f y" proof (cases "r \ 0") case True have "\x\cball a r \ s. \y \ cball a r \ s. f x \ f y" apply (rule continuous_attains_inf) using assms True unfolding proper_def apply (auto simp add: continuous_on_subset) using centre_in_cball by blast then obtain x where x: "x \ cball a r \ s" "\y. y \ cball a r \ s \ f x \ f y" by auto have "f x \ f y" if "y \ s" for y proof (cases "y \ cball a r") case True then show ?thesis using x(2) that by auto next case False have "f x \ f a" apply (rule x(2)) using assms True by auto then show ?thesis using assms(4)[of y] that False by auto qed then show ?thesis using x(1) by auto next case False show ?thesis apply (rule bexI[of _ a]) using assms False by auto qed subsubsection \Measure of balls\ text \The image of a ball by an affine map is still a ball, with explicit center and radius.\ lemma affine_image_ball [simp]: "(\y. R *\<^sub>R y + x) ` cball 0 1 = cball (x::('a::real_normed_vector)) \R\" proof have "dist x (R *\<^sub>R y + x) \ \R\" if "dist 0 y \ 1" for y proof - have "dist x (R *\<^sub>R y + x) = norm ((R *\<^sub>R y + x) - x)" by (simp add: dist_norm) also have "... = \R\ * norm y" by auto finally show ?thesis using that by (simp add: mult_left_le) qed then show "(\y. R *\<^sub>R y + x) ` cball 0 1 \ cball x \R\" by auto show "cball x \R\ \ (\y. R *\<^sub>R y + x) ` cball 0 1" proof (cases "\R\ = 0") case True then have "cball x \R\ = {x}" by auto moreover have "x = R *\<^sub>R 0 + x \ 0 \ cball 0 1" by auto ultimately show ?thesis by auto next case False have "z \ (\y. R *\<^sub>R y + x) ` cball 0 1" if "z \ cball x \R\" for z proof - define y where "y = (z - x) /\<^sub>R R" have "R *\<^sub>R y + x = z" unfolding y_def using False by auto moreover have "y \ cball 0 1" using \z \ cball x \R\\ False unfolding y_def by (auto simp add: dist_norm[symmetric] divide_simps dist_commute) ultimately show ?thesis by auto qed then show ?thesis by auto qed qed text \From the rescaling properties of Lebesgue measure in a euclidean space, it follows that the measure of any ball can be expressed in terms of the measure of the unit ball.\ lemma lebesgue_measure_ball: assumes "R \ 0" shows "measure lborel (cball (x::('a::euclidean_space)) R) = R^(DIM('a)) * measure lborel (cball (0::'a) 1)" "emeasure lborel (cball (x::('a::euclidean_space)) R) = R^(DIM('a)) * emeasure lborel (cball (0::'a) 1)" using measure_lebesgue_affine[of R x "cball 0 1"] emeasure_lebesgue_affine[of R x "cball 0 1"] assms unfolding affine_image_ball by (auto simp add: ennreal_power) text \We show that the unit ball has positive measure -- this is obvious, but useful. We could show it by arguing that it contains a box, whose measure can be computed, but instead we say that if the measure vanished then the measure of any ball would also vanish, contradicting the fact that the space has infinite measure. This avoids all computations.\ lemma lebesgue_measure_ball_pos: "emeasure lborel (cball (0::'a::euclidean_space) 1) > 0" "measure lborel (cball (0::'a::euclidean_space) 1) > 0" proof - show "emeasure lborel (cball (0::'a::euclidean_space) 1) > 0" proof (rule ccontr) assume "\(emeasure lborel (cball (0::'a::euclidean_space) 1) > 0)" then have "emeasure lborel (cball (0::'a) 1) = 0" by auto then have "emeasure lborel (cball (0::'a) n) = 0" for n::nat using lebesgue_measure_ball(2)[of "real n" 0] by (metis mult_zero_right of_nat_0_le_iff) then have "emeasure lborel (\n. cball (0::'a) (real n)) = 0" by (metis (mono_tags, lifting) borel_closed closed_cball emeasure_UN_eq_0 imageE sets_lborel subsetI) moreover have "(\n. cball (0::'a) (real n)) = UNIV" by (auto simp add: real_arch_simple) ultimately show False by simp qed moreover have "emeasure lborel (cball (0::'a::euclidean_space) 1) < \" by (rule emeasure_bounded_finite, auto) ultimately show "measure lborel (cball (0::'a::euclidean_space) 1) > 0" by (metis borel_closed closed_cball ennreal_0 has_integral_iff_emeasure_lborel has_integral_measure_lborel less_irrefl order_refl zero_less_measure_iff) qed subsubsection \infdist and closest point projection\ text \The distance to a union of two sets is the minimum of the distance to the two sets.\ lemma infdist_union_min [mono_intros]: assumes "A \ {}" "B \ {}" shows "infdist x (A \ B) = min (infdist x A) (infdist x B)" using assms by (simp add: infdist_def cINF_union inf_real_def) text \The distance to a set is non-increasing with the set.\ lemma infdist_mono [mono_intros]: assumes "A \ B" "A \ {}" shows "infdist x B \ infdist x A" proof - have "(INF a\B. dist x a) \ (INF a\A. dist x a)" by (rule cINF_superset_mono[OF assms(2) _ assms(1)], auto) then show ?thesis unfolding infdist_def using assms by auto qed text \If a set is proper, then the infimum of the distances to this set is attained.\ lemma infdist_proper_attained: assumes "proper C" "C \ {}" shows "\c\C. infdist x C = dist x c" proof - obtain a where "a \ C" using assms by auto have *: "dist x a \ dist x z" if "dist a z \ 2 * dist a x" for z proof - have "2 * dist a x \ dist a z" using that by simp also have "... \ dist a x + dist x z" by (intro mono_intros) finally show ?thesis by (simp add: dist_commute) qed have "\c\C. \d\C. dist x c \ dist x d" apply (rule continuous_attains_inf_proper[OF assms(1) \a \ C\, of _ "2 * dist a x"]) using * by (auto intro: continuous_intros) then show ?thesis unfolding infdist_def using \C \ {}\ by (metis antisym bdd_below_image_dist cINF_lower le_cINF_iff) qed lemma infdist_almost_attained: assumes "infdist x X < a" "X \ {}" shows "\y\X. dist x y < a" using assms using cInf_less_iff[of "(dist x)`X"] unfolding infdist_def by auto lemma infdist_triangle_abs [mono_intros]: "\infdist x A - infdist y A\ \ dist x y" by (metis (full_types) abs_diff_le_iff diff_le_eq dist_commute infdist_triangle) text \The next lemma is missing in the library, contrary to its cousin \verb+continuous_infdist+.\ -lemma continuous_on_infdist [continuous_intros]: - assumes "continuous_on S f" - shows "continuous_on S (\x. infdist (f x) A)" -using assms unfolding continuous_on by (auto intro: tendsto_infdist) - text \The infimum of the distance to a singleton set is simply the distance to the unique member of the set.\ lemma infdist_point [simp]: "infdist x {y} = dist x y" unfolding infdist_def by (metis cInf_singleton image_empty image_insert insert_not_empty) text \The closest point projection of $x$ on $A$. It is not unique, so we choose one point realizing the minimal distance. And if there is no such point, then we use $x$, to make some statements true without any assumption.\ definition proj_set::"'a::metric_space \ 'a set \ 'a set" where "proj_set x A = {y \ A. dist x y = infdist x A}" definition distproj::"'a::metric_space \ 'a set \ 'a" where "distproj x A = (if proj_set x A \ {} then SOME y. y \ proj_set x A else x)" lemma proj_setD: assumes "y \ proj_set x A" shows "y \ A" "dist x y = infdist x A" using assms unfolding proj_set_def by auto lemma proj_setI: assumes "y \ A" "dist x y \ infdist x A" shows "y \ proj_set x A" using assms infdist_le[OF \y \ A\, of x] unfolding proj_set_def by auto lemma proj_setI': assumes "y \ A" "\z. z \ A \ dist x y \ dist x z" shows "y \ proj_set x A" proof (rule proj_setI[OF \y \ A\]) show "dist x y \ infdist x A" apply (subst infdist_notempty) using assms by (auto intro!: cInf_greatest) qed lemma distproj_in_proj_set: assumes "proj_set x A \ {}" shows "distproj x A \ proj_set x A" "distproj x A \ A" "dist x (distproj x A) = infdist x A" proof - show "distproj x A \ proj_set x A" using assms unfolding distproj_def using some_in_eq by auto then show "distproj x A \ A" "dist x (distproj x A) = infdist x A" unfolding proj_set_def by auto qed lemma proj_set_nonempty_of_proper: assumes "proper A" "A \ {}" shows "proj_set x A \ {}" proof - have "\y. y \ A \ dist x y = infdist x A" using infdist_proper_attained[OF assms, of x] by auto then show "proj_set x A \ {}" unfolding proj_set_def by auto qed lemma distproj_self [simp]: assumes "x \ A" shows "proj_set x A = {x}" "distproj x A = x" proof - show "proj_set x A = {x}" unfolding proj_set_def using assms by auto then show "distproj x A = x" unfolding distproj_def by auto qed lemma distproj_closure [simp]: assumes "x \ closure A" shows "distproj x A = x" proof (cases "proj_set x A \ {}") case True show ?thesis using distproj_in_proj_set(3)[OF True] assms by (metis closure_empty dist_eq_0_iff distproj_self(2) in_closure_iff_infdist_zero) next case False then show ?thesis unfolding distproj_def by auto qed lemma distproj_le: assumes "y \ A" shows "dist x (distproj x A) \ dist x y" proof (cases "proj_set x A \ {}") case True show ?thesis using distproj_in_proj_set(3)[OF True] infdist_le[OF assms] by auto next case False then show ?thesis unfolding distproj_def by auto qed lemma proj_set_dist_le: assumes "y \ A" "p \ proj_set x A" shows "dist x p \ dist x y" using assms infdist_le unfolding proj_set_def by auto subsection \Material on ereal and ennreal\ text \We add the simp rules that we needed to make all computations become more or less automatic.\ lemma ereal_of_real_of_ereal_iff [simp]: "ereal(real_of_ereal x) = x \ x \ \ \ x \ - \" "x = ereal(real_of_ereal x) \ x \ \ \ x \ - \" by (metis MInfty_neq_ereal(1) PInfty_neq_ereal(2) real_of_ereal.elims)+ lemma inverse_eq_infinity_iff_eq_zero [simp]: "1/(x::ereal) = \ \ x = 0" by (simp add: divide_ereal_def) declare ereal_inverse_eq_0 [simp] declare ereal_0_gt_inverse [simp] declare ereal_inverse_le_0_iff [simp] declare ereal_divide_eq_0_iff [simp] declare ereal_mult_le_0_iff [simp] declare ereal_zero_le_0_iff [simp] declare ereal_mult_less_0_iff [simp] declare ereal_zero_less_0_iff [simp] declare ereal_uminus_eq_reorder [simp] declare ereal_minus_le_iff [simp] lemma ereal_inverse_noteq_minus_infinity [simp]: "1/(x::ereal) \ -\" by (simp add: divide_ereal_def) lemma ereal_inverse_positive_iff_nonneg_not_infinity [simp]: "0 < 1/(x::ereal) \ (x \ 0 \ x \ \)" by (cases x, auto simp add: one_ereal_def) lemma ereal_inverse_negative_iff_nonpos_not_infinity' [simp]: "0 > inverse (x::ereal) \ (x < 0 \ x \ -\)" by (cases x, auto simp add: one_ereal_def) lemma ereal_divide_pos_iff [simp]: "0 < x/(y::ereal) \ (y \ \ \ y \ -\) \ ((x > 0 \ y > 0) \ (x < 0 \ y < 0) \ (y = 0 \ x > 0))" unfolding divide_ereal_def by auto lemma ereal_divide_neg_iff [simp]: "0 > x/(y::ereal) \ (y \ \ \ y \ -\) \ ((x > 0 \ y < 0) \ (x < 0 \ y > 0) \ (y = 0 \ x < 0))" unfolding divide_ereal_def by auto lemma ereal_distrib_left: fixes a b c :: ereal assumes "a \ \ \ b \ -\" and "a \ -\ \ b \ \" and "\c\ \ \" shows "c * (a + b) = c * a + c * b" using assms by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) lemma ereal_distrib_minus_left: fixes a b c :: ereal assumes "a \ \ \ b \ \" and "a \ -\ \ b \ -\" and "\c\ \ \" shows "c * (a - b) = c * a - c * b" using assms by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) lemma ereal_distrib_minus_right: fixes a b c :: ereal assumes "a \ \ \ b \ \" and "a \ -\ \ b \ -\" and "\c\ \ \" shows "(a - b) * c = a * c - b * c" using assms by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) text \The next one is missing close to its friend \verb+Liminf_ereal_mult_right+.\ lemma Liminf_ereal_mult_left: assumes "F \ bot" "(c::real) \ 0" shows "Liminf F (\n. ereal c * f n) = ereal c * Liminf F f" using Liminf_ereal_mult_right[OF assms] by (subst (1 2) mult.commute) text \More additions to \verb+mono_intros+.\ lemma ereal_leq_imp_neg_leq [mono_intros]: fixes x y::ereal assumes "x \ y" shows "-y \ -x" using assms by auto lemma ereal_le_imp_neg_le [mono_intros]: fixes x y::ereal assumes "x < y" shows "-y < -x" using assms by auto declare ereal_mult_left_mono [mono_intros] declare ereal_mult_right_mono [mono_intros] declare ereal_mult_strict_right_mono [mono_intros] declare ereal_mult_strict_left_mono [mono_intros] text \Monotonicity of basic inclusions.\ lemma ennreal_mono': "mono ennreal" by (simp add: ennreal_leI monoI) lemma enn2ereal_mono': "mono enn2ereal" by (simp add: less_eq_ennreal.rep_eq mono_def) lemma e2ennreal_mono': "mono e2ennreal" by (simp add: e2ennreal_mono mono_def) lemma enn2ereal_mono [mono_intros]: assumes "x \ y" shows "enn2ereal x \ enn2ereal y" using assms less_eq_ennreal.rep_eq by auto lemma ereal_mono: "mono ereal" unfolding mono_def by auto lemma ereal_strict_mono: "strict_mono ereal" unfolding strict_mono_def by auto lemma ereal_mono2 [mono_intros]: assumes "x \ y" shows "ereal x \ ereal y" by (simp add: assms) lemma ereal_strict_mono2 [mono_intros]: assumes "x < y" shows "ereal x < ereal y" using assms by auto lemma enn2ereal_a_minus_b_plus_b [mono_intros]: "enn2ereal a \ enn2ereal (a - b) + enn2ereal b" by (metis diff_add_self_ennreal less_eq_ennreal.rep_eq linear plus_ennreal.rep_eq) text \The next lemma follows from the same assertion in ereals.\ lemma enn2ereal_strict_mono [mono_intros]: assumes "x < y" shows "enn2ereal x < enn2ereal y" using assms less_ennreal.rep_eq by auto declare ennreal_mult_strict_left_mono [mono_intros] declare ennreal_mult_strict_right_mono [mono_intros] lemma ennreal_ge_0 [mono_intros]: assumes "0 < x" shows "0 < ennreal x" by (simp add: assms) text \The next lemma is true and useful in ereal. Note that variants such as $a + b \leq c + d$ implies $a-d \leq c -b$ are not true -- take $a = c = \infty$ and $b = d = 0$...\ lemma ereal_minus_le_minus_plus [mono_intros]: fixes a b c::ereal assumes "a \ b + c" shows "-b \ -a + c" using assms apply (cases a, cases b, cases c, auto) using ereal_infty_less_eq2(2) ereal_plus_1(4) by fastforce lemma tendsto_ennreal_0 [tendsto_intros]: assumes "(u \ 0) F" shows "((\n. ennreal(u n)) \ 0) F" unfolding ennreal_0[symmetric] by (intro tendsto_intros assms) lemma tendsto_ennreal_1 [tendsto_intros]: assumes "(u \ 1) F" shows "((\n. ennreal(u n)) \ 1) F" unfolding ennreal_1[symmetric] by (intro tendsto_intros assms) subsection \Miscellaneous\ declare lim_1_over_n [tendsto_intros] declare lim_ln_over_n [tendsto_intros] lemma lim_log_over_n [tendsto_intros]: "(\n. log k n/n) \ 0" proof - have *: "log k n/n = (1/ln k) * (ln n / n)" for n unfolding log_def by auto have "(\n. (1/ln k) * (ln n / n)) \ (1/ln k) * 0" by (intro tendsto_intros) then show ?thesis unfolding * by auto qed lemma lim_ceiling_over_n [tendsto_intros]: assumes "(\n. u n/n) \ l" shows "(\n. ceiling(u n)/n) \ l" proof (rule tendsto_sandwich[of "\n. u n/n" _ _ "\n. u n/n + 1/n"]) show "\\<^sub>F n in sequentially. u n / real n \ real_of_int \u n\ / real n" unfolding eventually_sequentially by (rule exI[of _ 1], auto simp add: divide_simps) show "\\<^sub>F n in sequentially. real_of_int \u n\ / real n \ u n / real n + 1 / real n" unfolding eventually_sequentially by (rule exI[of _ 1], auto simp add: divide_simps) have "(\n. u n / real n + 1 / real n) \ l + 0" by (intro tendsto_intros assms) then show "(\n. u n / real n + 1 / real n) \ l" by auto qed (simp add: assms) lemma power4_eq_xxxx: fixes x::"'a::monoid_mult" shows "x^4 = x * x * x * x" by (simp add: mult.assoc power_numeral_even) subsubsection \Liminfs and Limsups\ text \More facts on liminfs and limsups\ lemma Limsup_obtain': fixes u::"'a \ 'b::complete_linorder" assumes "Limsup F u > c" "eventually P F" shows "\n. P n \ u n > c" proof - have *: "(INF P\{P. eventually P F}. SUP x\{x. P x}. u x) > c" using assms by (simp add: Limsup_def) have **: "c < (SUP x\{x. P x}. u x)" using less_INF_D[OF *, of P] assms by auto then show ?thesis by (simp add: less_SUP_iff) qed lemma limsup_obtain: fixes u::"nat \ 'a :: complete_linorder" assumes "limsup u > c" shows "\n \ N. u n > c" using Limsup_obtain'[OF assms, of "\n. n \ N"] unfolding eventually_sequentially by auto lemma Liminf_obtain': fixes u::"'a \ 'b::complete_linorder" assumes "Liminf F u < c" "eventually P F" shows "\n. P n \ u n < c" proof - have *: "(SUP P\{P. eventually P F}. INF x\{x. P x}. u x) < c" using assms by (simp add: Liminf_def) have **: "(INF x\{x. P x}. u x) < c" using SUP_lessD[OF *, of P] assms by auto then show ?thesis by (simp add: INF_less_iff) qed lemma liminf_obtain: fixes u::"nat \ 'a :: complete_linorder" assumes "liminf u < c" shows "\n \ N. u n < c" using Liminf_obtain'[OF assms, of "\n. n \ N"] unfolding eventually_sequentially by auto text \The Liminf of a minimum is the minimum of the Liminfs.\ lemma Liminf_min_eq_min_Liminf: fixes u v::"nat \ 'a::complete_linorder" shows "Liminf F (\n. min (u n) (v n)) = min (Liminf F u) (Liminf F v)" proof (rule order_antisym) show "Liminf F (\n. min (u n) (v n)) \ min (Liminf F u) (Liminf F v)" by (auto simp add: Liminf_mono) have "Liminf F (\n. min (u n) (v n)) > w" if H: "min (Liminf F u) (Liminf F v) > w" for w proof (cases "{w<..n. u n > w) F" "eventually (\n. v n > w) F" using H le_Liminf_iff by fastforce+ then have "eventually (\n. min (u n) (v n) > w) F" apply auto using eventually_elim2 by fastforce moreover have "z > w \ z \ min (Liminf F u) (Liminf F v)" for z using H True not_le_imp_less by fastforce ultimately have "eventually (\n. min (u n) (v n) \ min (Liminf F u) (Liminf F v)) F" by (simp add: eventually_mono) then have "min (Liminf F u) (Liminf F v) \ Liminf F (\n. min (u n) (v n))" by (metis Liminf_bounded) then show ?thesis using H less_le_trans by blast next case False then obtain z where "z \ {w<..n. u n > z) F" "eventually (\n. v n > z) F" using le_Liminf_iff by fastforce+ then have "eventually (\n. min (u n) (v n) > z) F" apply auto using eventually_elim2 by fastforce then have "Liminf F (\n. min (u n) (v n)) \ z" by (simp add: Liminf_bounded eventually_mono less_imp_le) then show ?thesis using H(1) by auto qed then show "min (Liminf F u) (Liminf F v) \ Liminf F (\n. min (u n) (v n))" using not_le_imp_less by blast qed text \The Limsup of a maximum is the maximum of the Limsups.\ lemma Limsup_max_eq_max_Limsup: fixes u::"'a \ 'b::complete_linorder" shows "Limsup F (\n. max (u n) (v n)) = max (Limsup F u) (Limsup F v)" proof (rule order_antisym) show "max (Limsup F u) (Limsup F v) \ Limsup F (\n. max (u n) (v n))" by (auto intro: Limsup_mono) have "Limsup F (\n. max (u n) (v n)) < e" if "max (Limsup F u) (Limsup F v) < e" for e proof (cases "\t. max (Limsup F u) (Limsup F v) < t \ t < e") case True then obtain t where t: "t < e" "max (Limsup F u) (Limsup F v) < t" by auto then have "Limsup F u < t" "Limsup F v < t" using that max_def by auto then have *: "eventually (\n. u n < t) F" "eventually (\n. v n < t) F" by (auto simp: Limsup_lessD) have "eventually (\n. max (u n) (v n) < t) F" using eventually_mono[OF eventually_conj[OF *]] by auto then have "Limsup F (\n. max (u n) (v n)) \ t" by (meson Limsup_obtain' not_le_imp_less order.asym) then show ?thesis using t by auto next case False have "Limsup F u < e" "Limsup F v < e" using that max_def by auto then have *: "eventually (\n. u n < e) F" "eventually (\n. v n < e) F" by (auto simp: Limsup_lessD) have "eventually (\n. max (u n) (v n) \ max (Limsup F u) (Limsup F v)) F" apply (rule eventually_mono[OF eventually_conj[OF *]]) using False not_le_imp_less by force then have "Limsup F (\n. max (u n) (v n)) \ max (Limsup F u) (Limsup F v)" by (meson Limsup_obtain' leD leI) then show ?thesis using that le_less_trans by blast qed then show "Limsup F (\n. max (u n) (v n)) \ max (Limsup F u) (Limsup F v)" using not_le_imp_less by blast qed subsubsection \Bounding the cardinality of a finite set\ text \A variation with real bounds.\ lemma finite_finite_subset_caract': fixes C::real assumes "\G. G \ F \ finite G \ card G \ C" shows "finite F \ card F \ C" by (meson assms finite_if_finite_subsets_card_bdd le_nat_floor order_refl) text \To show that a set has cardinality at most one, it suffices to show that any two of its elements coincide.\ lemma finite_at_most_singleton: assumes "\x y. x \ F \ y \ F \ x = y" shows "finite F \ card F \ 1" proof (cases "F = {}") case True then show ?thesis by auto next case False then obtain x where "x \ F" by auto then have "F = {x}" using assms by auto then show ?thesis by auto qed text \Bounded sets of integers are finite.\ lemma finite_real_int_interval [simp]: "finite (range real_of_int \ {a..b})" proof - have "range real_of_int \ {a..b} \ real_of_int`{floor a..ceiling b}" by (auto, metis atLeastAtMost_iff ceiling_mono ceiling_of_int floor_mono floor_of_int image_eqI) then show ?thesis using finite_subset by blast qed text \Well separated sets of real numbers are finite, with controlled cardinality.\ lemma separated_in_real_card_bound: assumes "T \ {a..(b::real)}" "d > 0" "\x y. x \ T \ y \ T \ y > x \ y \ x + d" shows "finite T" "card T \ nat (floor ((b-a)/d) + 1)" proof - define f where "f = (\x. floor ((x-a) / d))" have "f`{a..b} \ {0..floor ((b-a)/d)}" unfolding f_def using \d > 0\ by (auto simp add: floor_mono frac_le) then have *: "f`T \ {0..floor ((b-a)/d)}" using \T \ {a..b}\ by auto then have "finite (f`T)" by (rule finite_subset, auto) have "card (f`T) \ card {0..floor ((b-a)/d)}" apply (rule card_mono) using * by auto then have card_le: "card (f`T) \ nat (floor ((b-a)/d) + 1)" using card_atLeastAtMost_int by auto have *: "f x \ f y" if "y \ x + d" for x y proof - have "(y-a)/d \ (x-a)/d + 1" using \d > 0\ that by (auto simp add: divide_simps) then show ?thesis unfolding f_def by linarith qed have "inj_on f T" unfolding inj_on_def using * assms(3) by (auto, metis not_less_iff_gr_or_eq) show "finite T" using \finite (f`T)\ \inj_on f T\ finite_image_iff by blast have "card T = card (f`T)" using \inj_on f T\ by (simp add: card_image) then show "card T \ nat (floor ((b-a)/d) + 1)" using card_le by auto qed subsection \Manipulating finite ordered sets\ text \We will need below to contruct finite sets of real numbers with good properties expressed in terms of consecutive elements of the set. We introduce tools to manipulate such sets, expressing in particular the next and the previous element of the set and controlling how they evolve when one inserts a new element in the set. It works in fact in any linorder, and could also prove useful to construct sets of integer numbers. Manipulating the next and previous elements work well, except at the top (respectively bottom). In our constructions, these will be fixed and called $b$ and $a$.\ text \Notations for the next and the previous elements.\ definition next_in::"'a set \ 'a \ ('a::linorder)" where "next_in A u = Min (A \ {u<..})" definition prev_in::"'a set \ 'a \ ('a::linorder)" where "prev_in A u = Max (A \ {.. {a..b}" "a \ A" "b \ A" "a < b" begin text \Basic properties of the next element, when one starts from an element different from top.\ lemma next_in_basics: assumes "u \ {a.. A" "next_in A u > u" "A \ {u<.. A \ {u<..}" unfolding next_in_def apply (rule Min_in) using assms \finite A\ \b \ A\ by auto then show "next_in A u \ A" "next_in A u > u" by auto show "A \ {u<.. {a.. A" "v > u" "{u<.. A = {}" shows "next_in A u = v" using assms next_in_basics[OF \u \ {a..] by fastforce text \Basic properties of the previous element, when one starts from an element different from bottom.\ lemma prev_in_basics: assumes "u \ {a<..b}" shows "prev_in A u \ A" "prev_in A u < u" "A \ {prev_in A u<.. A \ {..finite A\ \a \ A\ by auto then show "prev_in A u \ A" "prev_in A u < u" by auto show "A \ {prev_in A u<.. {a<..b}" "v \ A" "v < u" "{v<.. A = {}" shows "prev_in A u = v" using assms prev_in_basics[OF \u \ {a<..b}\] by (meson disjoint_iff_not_equal greaterThanLessThan_iff less_linear) text \The interval $[a,b]$ is covered by the intervals between the consecutive elements of $A$.\ lemma intervals_decomposition: "(\ U \ {{u..next_in A u} | u. u \ A - {b}}. U) = {a..b}" proof show "(\U\{{u..next_in A u} |u. u \ A - {b}}. U) \ {a..b}" using \A \ {a..b}\ next_in_basics(1) apply auto apply fastforce by (metis \A \ {a..b}\ atLeastAtMost_iff eq_iff le_less_trans less_eq_real_def not_less subset_eq subset_iff_psubset_eq) have "x \ (\U\{{u..next_in A u} |u. u \ A - {b}}. U)" if "x \ {a..b}" for x proof - consider "x = b" | "x \ A - {b}" | "x \ A" by blast then show ?thesis proof(cases) case 1 define u where "u = prev_in A b" have "b \ {a<..b}" using \a < b\ by simp have "u \ A - {b}" unfolding u_def using prev_in_basics[OF \b \ {a<..b}\] by simp then have "u \ {a..A \ {a..b}\ \a < b\ by fastforce have "next_in A u = b" using prev_in_basics[OF \b \ {a<..b}\] next_in_basics[OF \u \ {a..] \A \ {a..b}\ unfolding u_def by force then have "x \ {u..next_in A u}" unfolding 1 using prev_in_basics[OF \b \ {a<..b}\] u_def by auto then show ?thesis using \u \ A - {b}\ by auto next case 2 then have "x \ {a..A \ {a..b}\ \a < b\ by fastforce have "x \ {x.. next_in A x}" using next_in_basics[OF \x \ {a..] by auto then show ?thesis using 2 by auto next case 3 then have "x \ {a<..b}" using that \a \ A\ leI by fastforce define u where "u = prev_in A x" have "u \ A - {b}" unfolding u_def using prev_in_basics[OF \x \ {a<..b}\] that by auto then have "u \ {a..A \ {a..b}\ \a < b\ by fastforce have "x \ {u..next_in A u}" using prev_in_basics[OF \x \ {a<..b}\] next_in_basics[OF \u \ {a..] unfolding u_def by auto then show ?thesis using \u \ A - {b}\ by auto qed qed then show "{a..b} \ (\U\{{u..next_in A u} |u. u \ A - {b}}. U)" by auto qed end text \If one inserts an additional element, then next and previous elements are not modified, except at the location of the insertion.\ lemma next_in_insert: assumes A: "finite A" "A \ {a..b}" "a \ A" "b \ A" "a < b" and "x \ {a..b} - A" shows "\u. u \ A - {b, prev_in A x} \ next_in (insert x A) u = next_in A u" "next_in (insert x A) x = next_in A x" "next_in (insert x A) (prev_in A x) = x" proof - define B where "B = insert x A" have B: "finite B" "B \ {a..b}" "a \ B" "b \ B" "a < b" using assms unfolding B_def by auto have x: "x \ {a.. {a<..b}" using assms leI by fastforce+ show "next_in B x = next_in A x" unfolding B_def by (auto simp add: next_in_def) show "next_in B (prev_in A x) = x" apply (rule next_inI[OF B]) unfolding B_def using prev_in_basics[OF A \x \ {a<..b}\] \A \ {a..b}\ x by auto fix u assume "u \ A - {b, prev_in A x}" then have "u \ {a.. {u<..(x \ {u<.. {u<..x \ {a<..b}\]) using \u \ A - {b, prev_in A x}\ * next_in_basics[OF A \u \ {a..] apply auto by (meson disjoint_iff_not_equal greaterThanLessThan_iff less_trans) then show False using \u \ A - {b, prev_in A x}\ by auto qed show "next_in B u = next_in A u" apply (rule next_inI[OF B \u \ {a..]) unfolding B_def using next_in_basics[OF A \u \ {a..] \x \ {u<.. by auto qed text \If consecutive elements are enough separated, this implies a simple bound on the cardinality of the set.\ lemma separated_in_real_card_bound2: fixes A::"real set" assumes A: "finite A" "A \ {a..b}" "a \ A" "b \ A" "a < b" and B: "\u. u \ A - {b} \ next_in A u \ u + d" "d > 0" shows "card A \ nat (floor ((b-a)/d) + 1)" proof (rule separated_in_real_card_bound[OF \A \ {a..b}\ \d > 0\]) fix x y assume "x \ A" "y \ A" "y > x" then have "x \ A - {b}" "x \ {a..A \ {a..b}\ by auto have "y \ next_in A x" using next_in_basics[OF A \x \ {a..] \y \ A\ \y > x\ by auto then show "y \ x + d" using B(1)[OF \x \ A - {b}\] by auto qed subsection \Well-orders\ text \In this subsection, we give additional lemmas on well-orders or cardinals or whatever, that would well belong to the library, and will be needed below.\ lemma (in wo_rel) max2_underS [simp]: assumes "x \ underS z" "y \ underS z" shows "max2 x y \ underS z" using assms max2_def by auto lemma (in wo_rel) max2_underS' [simp]: assumes "x \ underS y" shows "max2 x y = y" "max2 y x = y" apply (simp add: underS_E assms max2_def) using assms max2_def ANTISYM antisym_def underS_E by fastforce lemma (in wo_rel) max2_xx [simp]: "max2 x x = x" using max2_def by auto declare underS_notIn [simp] text \The abbrevation $=o$ is used both in \verb+Set_Algebras+ and Cardinals. We disable the one from \verb+Set_Algebras+.\ no_notation elt_set_eq (infix "=o" 50) lemma regularCard_ordIso: assumes "Card_order r" "regularCard r" "s =o r" shows "regularCard s" unfolding regularCard_def proof (auto) fix K assume K: "K \ Field s" "cofinal K s" obtain f where f: "bij_betw f (Field s) (Field r)" "embed s r f" using \s =o r\ unfolding ordIso_def iso_def by auto have "f`K \ Field r" using K(1) f(1) bij_betw_imp_surj_on by blast have "cofinal (f`K) r" unfolding cofinal_def proof fix a assume "a \ Field r" then obtain a' where a: "a' \ Field s" "f a' = a" using f by (metis bij_betw_imp_surj_on imageE) then obtain b' where b: "b' \ K" "a' \ b' \ (a', b') \ s" using \cofinal K s\ unfolding cofinal_def by auto have P1: "f b' \ f`K" using b(1) by auto have "a' \ b'" "a' \ Field s" "b' \ Field s" using a(1) b K(1) by auto then have P2: "a \ f b'" unfolding a(2)[symmetric] using f(1) unfolding bij_betw_def inj_on_def by auto have "(a', b') \ s" using b by auto then have P3: "(a, f b') \ r" unfolding a(2)[symmetric] using f by (meson FieldI1 FieldI2 Card_order_ordIso[OF assms(1) assms(3)] card_order_on_def iso_defs(1) iso_iff2) show "\b\f ` K. a \ b \ (a, b) \ r" using P1 P2 P3 by blast qed then have "|f`K| =o r" using \regularCard r\ \f`K \ Field r\ unfolding regularCard_def by auto moreover have "|f`K| =o |K|" using f(1) K(1) by (meson bij_betw_subset card_of_ordIsoI ordIso_symmetric) ultimately show "|K| =o s" using \s =o r\ by (meson ordIso_symmetric ordIso_transitive) qed lemma AboveS_not_empty_in_regularCard: assumes "|S| Field r" assumes r: "Card_order r" "regularCard r" "\finite (Field r)" shows "AboveS r S \ {}" proof - have "\(cofinal S r)" using assms not_ordLess_ordIso unfolding regularCard_def by auto then obtain a where a: "a \ Field r" "\b\S. \(a \ b \ (a,b) \ r)" unfolding cofinal_def by auto have *: "a = b \ (b, a) \ r" if "b \ S" for b proof - have "a = b \ (a,b) \ r" using a that by auto then show ?thesis using \Card_order r\ \a \ Field r\ \b \ S\ \S \ Field r\ unfolding card_order_on_def well_order_on_def linear_order_on_def total_on_def by auto qed obtain c where "c \ Field r" "c \ a" "(a, c) \ r" using a(1) r infinite_Card_order_limit by fastforce then have "c \ AboveS r S" unfolding AboveS_def apply simp using Card_order_trans[OF r(1)] by (metis *) then show ?thesis by auto qed lemma AboveS_not_empty_in_regularCard': assumes "|S| Field r" "T \ S" assumes r: "Card_order r" "regularCard r" "\finite (Field r)" shows "AboveS r (f`T) \ {}" proof - have "|f`T| \o |T|" by simp moreover have "|T| \o |S|" using \T \ S\ by simp ultimately have *: "|f`T| |S| by (meson ordLeq_ordLess_trans) show ?thesis using AboveS_not_empty_in_regularCard[OF * _ r] \T \ S\ \f`S \ Field r\ by auto qed lemma Well_order_extend: assumes WELL: "well_order_on A r" and SUB: "A \ B" shows "\r'. well_order_on B r' \ r \ r'" proof- have r: "Well_order r \ Field r = A" using WELL well_order_on_Well_order by blast let ?C = "B - A" obtain r'' where "well_order_on ?C r''" using well_order_on by blast then have r'': "Well_order r'' \ Field r'' = ?C" using well_order_on_Well_order by blast let ?r' = "r Osum r''" have "Field r Int Field r'' = {}" using r r'' by auto then have "r \o ?r'" using Osum_ordLeq[of r r''] r r'' by blast then have "Well_order ?r'" unfolding ordLeq_def by auto moreover have "Field ?r' = B" using r r'' SUB by (auto simp add: Field_Osum) ultimately have "well_order_on B ?r'" by auto moreover have "r \ ?r'" by (simp add: Osum_def subrelI) ultimately show ?thesis by blast qed text \The next lemma shows that, if the range of a function is endowed with a wellorder, then one can pull back this wellorder by the function, and then extend it in the fibers of the function in order to keep the wellorder property. The proof is done by taking an arbitrary family of wellorders on each of the fibers, and using the lexicographic order: one has $x < y$ if $f x < f y$, or if $f x = f y$ and, in the corresponding fiber of $f$, one has $x < y$. To formalize it, it is however more efficient to use one single wellorder, and restrict it to each fiber.\ lemma Well_order_pullback: assumes "Well_order r" shows "\s. Well_order s \ Field s = UNIV \ (\x y. (f x, f y) \ (r-Id) \ (x, y) \ s)" proof - obtain r2 where r2: "Well_order r2" "Field r2 = UNIV" "r \ r2" using Well_order_extend[OF assms, of UNIV] well_order_on_Well_order by auto obtain s2 where s2: "Well_order s2" "Field s2 = (UNIV::'b set)" by (meson well_ordering) have r2s2: "\x y z. (x, y) \ s2 \ (y, z) \ s2 \ (x, z) \ s2" "\x. (x, x) \ s2" "\x y. (x, y) \ s2 \ (y, x) \ s2" "\x y. (x, y) \ s2 \ (y, x) \ s2 \ x = y" "\x y z. (x, y) \ r2 \ (y, z) \ r2 \ (x, z) \ r2" "\x. (x, x) \ r2" "\x y. (x, y) \ r2 \ (y, x) \ r2" "\x y. (x, y) \ r2 \ (y, x) \ r2 \ x = y" using r2 s2 unfolding well_order_on_def linear_order_on_def partial_order_on_def total_on_def preorder_on_def antisym_def refl_on_def trans_def by (metis UNIV_I)+ define s where "s = {(x,y). (f x, f y) \ r2 \ (f x = f y \ (x, y) \ s2)}" have "linear_order s" unfolding linear_order_on_def partial_order_on_def preorder_on_def proof (auto) show "total_on UNIV s" unfolding s_def apply (rule total_onI, auto) using r2s2 by metis+ show "refl_on UNIV s" unfolding s_def apply (rule refl_onI, auto) using r2s2 by blast+ show "trans s" unfolding s_def apply (rule transI, auto) using r2s2 by metis+ show "antisym s" unfolding s_def apply (rule antisymI, auto) using r2s2 by metis+ qed moreover have "wf (s - Id)" proof (rule wfI_min) fix x::'b and Q assume "x \ Q" obtain z' where z': "z' \ f`Q" "\y. (y, z') \ r2 - Id \ y \ f`Q" proof (rule wfE_min[of "r2-Id" "f x" "f`Q"], auto) show "wf(r2-Id)" using \Well_order r2\ unfolding well_order_on_def by auto show "f x \ f`Q" using \x \ Q\ by auto qed define Q2 where "Q2 = Q \ f-`{z'}" obtain z where z: "z \ Q2" "\y. (y, z) \ s2 - Id \ y \ Q2" proof (rule wfE_min'[of "s2-Id" "Q2"], auto) show "wf(s2-Id)" using \Well_order s2\ unfolding well_order_on_def by auto assume "Q2 = {}" then show False unfolding Q2_def using \z' \ f`Q\ by blast qed have "(y, z) \ (s-Id) \ y \ Q" for y unfolding s_def using z' z Q2_def by auto then show "\z\Q. \y. (y, z) \ s - Id \ y \ Q" using \z \ Q2\ Q2_def by auto qed ultimately have "well_order_on UNIV s" unfolding well_order_on_def by simp moreover have "(f x, f y) \ (r-Id) \ (x, y) \ s" for x y unfolding s_def using \r \ r2\ by auto ultimately show ?thesis using well_order_on_Well_order by metis qed end (*of theory Library_Complements*)