src/HOL/Probability/Radon_Nikodym.thy
changeset 44918 6a80fbc4e72c
parent 44890 22f665a2e91c
child 44928 7ef6505bde7f
equal deleted inserted replaced
44917:8df4c332cb1c 44918:6a80fbc4e72c
   409     using g_in_G `incseq ?g`
   409     using g_in_G `incseq ?g`
   410     by (auto intro!: positive_integral_monotone_convergence_SUP simp: G_def)
   410     by (auto intro!: positive_integral_monotone_convergence_SUP simp: G_def)
   411   also have "\<dots> = ?y"
   411   also have "\<dots> = ?y"
   412   proof (rule antisym)
   412   proof (rule antisym)
   413     show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y"
   413     show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y"
   414       using g_in_G by (auto intro!: exI Sup_mono simp: SUPR_def)
   414       using g_in_G
       
   415       using [[simp_trace]]
       
   416       by (auto intro!: exI Sup_mono simp: SUPR_def)
   415     show "?y \<le> (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq
   417     show "?y \<le> (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq
   416       by (auto intro!: SUP_mono positive_integral_mono Max_ge)
   418       by (auto intro!: SUP_mono positive_integral_mono Max_ge)
   417   qed
   419   qed
   418   finally have int_f_eq_y: "integral\<^isup>P M f = ?y" .
   420   finally have int_f_eq_y: "integral\<^isup>P M f = ?y" .
   419   have "\<And>x. 0 \<le> f x"
   421   have "\<And>x. 0 \<le> f x"