src/HOL/Analysis/Bochner_Integration.thy
changeset 68403 223172b97d0b
parent 68073 fad29d2a17a5
child 68794 63e84bd8e1f6
equal deleted inserted replaced
68387:691c02d1699b 68403:223172b97d0b
   474     unfolding simple_integral_def
   474     unfolding simple_integral_def
   475     by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ennreal (f x)" and v=enn2real])
   475     by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ennreal (f x)" and v=enn2real])
   476        (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases
   476        (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases
   477              intro!: sum.cong ennreal_cong_mult
   477              intro!: sum.cong ennreal_cong_mult
   478              simp: ac_simps ennreal_mult
   478              simp: ac_simps ennreal_mult
   479              reorient: sum_ennreal)
   479              simp flip: sum_ennreal)
   480   also have "\<dots> = (\<integral>\<^sup>+x. f x \<partial>M)"
   480   also have "\<dots> = (\<integral>\<^sup>+x. f x \<partial>M)"
   481     using f
   481     using f
   482     by (intro nn_integral_eq_simple_integral[symmetric])
   482     by (intro nn_integral_eq_simple_integral[symmetric])
   483        (auto simp: simple_function_compose1 simple_bochner_integrable.simps)
   483        (auto simp: simple_function_compose1 simple_bochner_integrable.simps)
   484   finally show ?thesis .
   484   finally show ?thesis .
   502     by (auto intro!: simple_bochner_integral_norm_bound)
   502     by (auto intro!: simple_bochner_integral_norm_bound)
   503   also have "\<dots> = (\<integral>\<^sup>+x. norm (s x - t x) \<partial>M)"
   503   also have "\<dots> = (\<integral>\<^sup>+x. norm (s x - t x) \<partial>M)"
   504     using simple_bochner_integrable_compose2[of "\<lambda>x y. norm (x - y)" M "s" "t"] s t
   504     using simple_bochner_integrable_compose2[of "\<lambda>x y. norm (x - y)" M "s" "t"] s t
   505     by (auto intro!: simple_bochner_integral_eq_nn_integral)
   505     by (auto intro!: simple_bochner_integral_eq_nn_integral)
   506   also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \<partial>M)"
   506   also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \<partial>M)"
   507     by (auto intro!: nn_integral_mono reorient: ennreal_plus)
   507     by (auto intro!: nn_integral_mono simp flip: ennreal_plus)
   508        (metis (erased, hide_lams) add_diff_cancel_left add_diff_eq diff_add_eq order_trans
   508        (metis (erased, hide_lams) add_diff_cancel_left add_diff_eq diff_add_eq order_trans
   509               norm_minus_commute norm_triangle_ineq4 order_refl)
   509               norm_minus_commute norm_triangle_ineq4 order_refl)
   510   also have "\<dots> = ?S + ?T"
   510   also have "\<dots> = ?S + ?T"
   511    by (rule nn_integral_add) auto
   511    by (rule nn_integral_add) auto
   512   finally show ?thesis .
   512   finally show ?thesis .
   592     show "eventually (\<lambda>i. ?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) \<partial>M) + \<integral>\<^sup>+ x. (norm (g x - sg i x)) \<partial>M) sequentially"
   592     show "eventually (\<lambda>i. ?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) \<partial>M) + \<integral>\<^sup>+ x. (norm (g x - sg i x)) \<partial>M) sequentially"
   593       (is "eventually (\<lambda>i. ?f i \<le> ?g i) sequentially")
   593       (is "eventually (\<lambda>i. ?f i \<le> ?g i) sequentially")
   594     proof (intro always_eventually allI)
   594     proof (intro always_eventually allI)
   595       fix i have "?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \<partial>M)"
   595       fix i have "?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \<partial>M)"
   596         by (auto intro!: nn_integral_mono norm_diff_triangle_ineq
   596         by (auto intro!: nn_integral_mono norm_diff_triangle_ineq
   597                  reorient: ennreal_plus)
   597                  simp flip: ennreal_plus)
   598       also have "\<dots> = ?g i"
   598       also have "\<dots> = ?g i"
   599         by (intro nn_integral_add) auto
   599         by (intro nn_integral_add) auto
   600       finally show "?f i \<le> ?g i" .
   600       finally show "?f i \<le> ?g i" .
   601     qed
   601     qed
   602     show "?g \<longlonglongrightarrow> 0"
   602     show "?g \<longlonglongrightarrow> 0"
   745   also have "\<dots> < \<infinity>"
   745   also have "\<dots> < \<infinity>"
   746     using s by (subst nn_integral_cmult_indicator) (auto simp: \<open>0 \<le> m\<close> simple_bochner_integrable.simps ennreal_mult_less_top less_top)
   746     using s by (subst nn_integral_cmult_indicator) (auto simp: \<open>0 \<le> m\<close> simple_bochner_integrable.simps ennreal_mult_less_top less_top)
   747   finally have s_fin: "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>" .
   747   finally have s_fin: "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>" .
   748 
   748 
   749   have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \<partial>M)"
   749   have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \<partial>M)"
   750     by (auto intro!: nn_integral_mono reorient: ennreal_plus)
   750     by (auto intro!: nn_integral_mono simp flip: ennreal_plus)
   751        (metis add.commute norm_triangle_sub)
   751        (metis add.commute norm_triangle_sub)
   752   also have "\<dots> = (\<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) + (\<integral>\<^sup>+x. norm (s i x) \<partial>M)"
   752   also have "\<dots> = (\<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) + (\<integral>\<^sup>+x. norm (s i x) \<partial>M)"
   753     by (rule nn_integral_add) auto
   753     by (rule nn_integral_add) auto
   754   also have "\<dots> < \<infinity>"
   754   also have "\<dots> < \<infinity>"
   755     using s_fin f_s_fin by auto
   755     using s_fin f_s_fin by auto
   781         by (auto intro!: simple_bochner_integral_norm_bound)
   781         by (auto intro!: simple_bochner_integral_norm_bound)
   782       also have "\<dots> = (\<integral>\<^sup>+x. norm (s n x) \<partial>M)"
   782       also have "\<dots> = (\<integral>\<^sup>+x. norm (s n x) \<partial>M)"
   783         by (intro simple_bochner_integral_eq_nn_integral)
   783         by (intro simple_bochner_integral_eq_nn_integral)
   784            (auto intro: s simple_bochner_integrable_compose2)
   784            (auto intro: s simple_bochner_integrable_compose2)
   785       also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \<partial>M)"
   785       also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \<partial>M)"
   786         by (auto intro!: nn_integral_mono reorient: ennreal_plus)
   786         by (auto intro!: nn_integral_mono simp flip: ennreal_plus)
   787            (metis add.commute norm_minus_commute norm_triangle_sub)
   787            (metis add.commute norm_minus_commute norm_triangle_sub)
   788       also have "\<dots> = ?t n"
   788       also have "\<dots> = ?t n"
   789         by (rule nn_integral_add) auto
   789         by (rule nn_integral_add) auto
   790       finally show "norm (?s n) \<le> ?t n" .
   790       finally show "norm (?s n) \<le> ?t n" .
   791     qed
   791     qed
   826       by (intro always_eventually allI simple_bochner_integral_bounded s t f)
   826       by (intro always_eventually allI simple_bochner_integral_bounded s t f)
   827     show "(\<lambda>i. ?S i + ?T i) \<longlonglongrightarrow> ennreal 0"
   827     show "(\<lambda>i. ?S i + ?T i) \<longlonglongrightarrow> ennreal 0"
   828       using tendsto_add[OF \<open>?S \<longlonglongrightarrow> 0\<close> \<open>?T \<longlonglongrightarrow> 0\<close>] by simp
   828       using tendsto_add[OF \<open>?S \<longlonglongrightarrow> 0\<close> \<open>?T \<longlonglongrightarrow> 0\<close>] by simp
   829   qed
   829   qed
   830   then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> 0"
   830   then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> 0"
   831     by (simp reorient: ennreal_0)
   831     by (simp flip: ennreal_0)
   832   ultimately have "norm (x - y) = 0"
   832   ultimately have "norm (x - y) = 0"
   833     by (rule LIMSEQ_unique)
   833     by (rule LIMSEQ_unique)
   834   then show "x = y" by simp
   834   then show "x = y" by simp
   835 qed
   835 qed
   836 
   836 
  1172         using M[OF n] by auto
  1172         using M[OF n] by auto
  1173       have "norm (?s n - ?s m) \<le> ?S n + ?S m"
  1173       have "norm (?s n - ?s m) \<le> ?S n + ?S m"
  1174         by (intro simple_bochner_integral_bounded s f)
  1174         by (intro simple_bochner_integral_bounded s f)
  1175       also have "\<dots> < ennreal (e / 2) + e / 2"
  1175       also have "\<dots> < ennreal (e / 2) + e / 2"
  1176         by (intro add_strict_mono M n m)
  1176         by (intro add_strict_mono M n m)
  1177       also have "\<dots> = e" using \<open>0<e\<close> by (simp reorient: ennreal_plus)
  1177       also have "\<dots> = e" using \<open>0<e\<close> by (simp flip: ennreal_plus)
  1178       finally show "dist (?s n) (?s m) < e"
  1178       finally show "dist (?s n) (?s m) < e"
  1179         using \<open>0<e\<close> by (simp add: dist_norm ennreal_less_iff)
  1179         using \<open>0<e\<close> by (simp add: dist_norm ennreal_less_iff)
  1180     qed
  1180     qed
  1181   qed
  1181   qed
  1182   then obtain x where "?s \<longlonglongrightarrow> x" ..
  1182   then obtain x where "?s \<longlonglongrightarrow> x" ..
  1217       using u'
  1217       using u'
  1218     proof eventually_elim
  1218     proof eventually_elim
  1219       fix x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
  1219       fix x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
  1220       from tendsto_diff[OF tendsto_const[of "u' x"] this]
  1220       from tendsto_diff[OF tendsto_const[of "u' x"] this]
  1221       show "(\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
  1221       show "(\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
  1222         by (simp add: tendsto_norm_zero_iff reorient: ennreal_0)
  1222         by (simp add: tendsto_norm_zero_iff flip: ennreal_0)
  1223     qed
  1223     qed
  1224   qed (insert bnd w_nonneg, auto)
  1224   qed (insert bnd w_nonneg, auto)
  1225   then show ?thesis by simp
  1225   then show ?thesis by simp
  1226 qed
  1226 qed
  1227 
  1227 
  2115     }
  2115     }
  2116     then show "eventually (\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)) F"
  2116     then show "eventually (\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M)) \<le> (\<integral>\<^sup>+x. norm(u n x - f x) \<partial>M)) F"
  2117       by auto
  2117       by auto
  2118   qed
  2118   qed
  2119   then have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F"
  2119   then have "((\<lambda>n. norm((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F"
  2120     by (simp reorient: ennreal_0)
  2120     by (simp flip: ennreal_0)
  2121   then have "((\<lambda>n. ((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F" using tendsto_norm_zero_iff by blast
  2121   then have "((\<lambda>n. ((\<integral>x. u n x \<partial>M) - (\<integral>x. f x \<partial>M))) \<longlongrightarrow> 0) F" using tendsto_norm_zero_iff by blast
  2122   then show ?thesis using Lim_null by auto
  2122   then show ?thesis using Lim_null by auto
  2123 qed
  2123 qed
  2124 
  2124 
  2125 text \<open>The next lemma asserts that, if a sequence of functions converges in $L^1$, then
  2125 text \<open>The next lemma asserts that, if a sequence of functions converges in $L^1$, then
  2213     moreover then have "liminf (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
  2213     moreover then have "liminf (\<lambda>n. ennreal (norm(u (r n) x))) \<le> 0"
  2214       by (rule order_trans[rotated]) (auto intro: Liminf_le_Limsup)
  2214       by (rule order_trans[rotated]) (auto intro: Liminf_le_Limsup)
  2215     ultimately have "(\<lambda>n. ennreal (norm(u (r n) x))) \<longlonglongrightarrow> 0"
  2215     ultimately have "(\<lambda>n. ennreal (norm(u (r n) x))) \<longlonglongrightarrow> 0"
  2216       using tendsto_Limsup[of sequentially "\<lambda>n. ennreal (norm(u (r n) x))"] by auto
  2216       using tendsto_Limsup[of sequentially "\<lambda>n. ennreal (norm(u (r n) x))"] by auto
  2217     then have "(\<lambda>n. norm(u (r n) x)) \<longlonglongrightarrow> 0"
  2217     then have "(\<lambda>n. norm(u (r n) x)) \<longlonglongrightarrow> 0"
  2218       by (simp reorient: ennreal_0)
  2218       by (simp flip: ennreal_0)
  2219     then have "(\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0"
  2219     then have "(\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0"
  2220       by (simp add: tendsto_norm_zero_iff)
  2220       by (simp add: tendsto_norm_zero_iff)
  2221   }
  2221   }
  2222   ultimately have "AE x in M. (\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0" by auto
  2222   ultimately have "AE x in M. (\<lambda>n. u (r n) x) \<longlonglongrightarrow> 0" by auto
  2223   then show ?thesis using \<open>strict_mono r\<close> by auto
  2223   then show ?thesis using \<open>strict_mono r\<close> by auto