1039 by (simp add: A_def) |
1039 by (simp add: A_def) |
1040 |
1040 |
1041 let ?P = "\<lambda>y. to_nat_on (A y)" |
1041 let ?P = "\<lambda>y. to_nat_on (A y)" |
1042 def pp \<equiv> "map_pmf (\<lambda>(x, y). (y, ?P y x)) pq" |
1042 def pp \<equiv> "map_pmf (\<lambda>(x, y). (y, ?P y x)) pq" |
1043 let ?pp = "\<lambda>y x. pmf pp (y, x)" |
1043 let ?pp = "\<lambda>y x. pmf pp (y, x)" |
1044 { fix x y have "x \<in> A y \<Longrightarrow> pmf pp (y, ?P y x) = pmf pq (x, y)" |
1044 { fix x y have "x \<in> A y \<Longrightarrow> ?pp y (?P y x) = pmf pq (x, y)" |
1045 unfolding pp_def |
1045 unfolding pp_def |
1046 by (intro pmf_map_inj[of "\<lambda>(x, y). (y, ?P y x)" pq "(x, y)", simplified]) |
1046 by (intro pmf_map_inj[of "\<lambda>(x, y). (y, ?P y x)" pq "(x, y)", simplified]) |
1047 (auto simp: inj_on_def A) } |
1047 (auto simp: inj_on_def A) } |
1048 note pmf_pp = this |
1048 note pmf_pp = this |
|
1049 have pp_0: "\<And>y x. pmf q y = 0 \<Longrightarrow> ?pp y x = 0" |
|
1050 proof(erule contrapos_pp) |
|
1051 fix y x |
|
1052 assume "?pp y x \<noteq> 0" |
|
1053 hence "(y, x) \<in> set_pmf pp" by(simp add: set_pmf_iff) |
|
1054 hence "y \<in> set_pmf q" by(auto simp add: pp_def q set_map_pmf intro: rev_image_eqI) |
|
1055 thus "pmf q y \<noteq> 0" by(simp add: set_pmf_iff) |
|
1056 qed |
1049 |
1057 |
1050 def B \<equiv> "\<lambda>y. {z. (y, z) \<in> set_pmf qr}" |
1058 def B \<equiv> "\<lambda>y. {z. (y, z) \<in> set_pmf qr}" |
1051 then have "\<And>y. B y \<subseteq> set_pmf r" by (auto simp add: r set_map_pmf intro: rev_image_eqI) |
1059 then have "\<And>y. B y \<subseteq> set_pmf r" by (auto simp add: r set_map_pmf intro: rev_image_eqI) |
1052 then have [simp]: "\<And>y. countable (B y)" by (rule countable_subset) simp |
1060 then have [simp]: "\<And>y. countable (B y)" by (rule countable_subset) simp |
1053 have B: "\<And>y z. (y, z) \<in> set_pmf qr \<longleftrightarrow> z \<in> B y" |
1061 have B: "\<And>y z. (y, z) \<in> set_pmf qr \<longleftrightarrow> z \<in> B y" |
1054 by (simp add: B_def) |
1062 by (simp add: B_def) |
1055 |
1063 |
1056 let ?R = "\<lambda>y. to_nat_on (B y)" |
1064 let ?R = "\<lambda>y. to_nat_on (B y)" |
1057 def rr \<equiv> "map_pmf (\<lambda>(y, z). (y, ?R y z)) qr" |
1065 def rr \<equiv> "map_pmf (\<lambda>(y, z). (y, ?R y z)) qr" |
1058 let ?rr = "\<lambda>y z. pmf rr (y, z)" |
1066 let ?rr = "\<lambda>y z. pmf rr (y, z)" |
1059 { fix y z have "z \<in> B y \<Longrightarrow> pmf rr (y, ?R y z) = pmf qr (y, z)" |
1067 { fix y z have "z \<in> B y \<Longrightarrow> ?rr y (?R y z) = pmf qr (y, z)" |
1060 unfolding rr_def |
1068 unfolding rr_def |
1061 by (intro pmf_map_inj[of "\<lambda>(y, z). (y, ?R y z)" qr "(y, z)", simplified]) |
1069 by (intro pmf_map_inj[of "\<lambda>(y, z). (y, ?R y z)" qr "(y, z)", simplified]) |
1062 (auto simp: inj_on_def B) } |
1070 (auto simp: inj_on_def B) } |
1063 note pmf_rr = this |
1071 note pmf_rr = this |
1064 |
1072 have rr_0: "\<And>y z. pmf q y = 0 \<Longrightarrow> ?rr y z = 0" |
|
1073 proof(erule contrapos_pp) |
|
1074 fix y z |
|
1075 assume "?rr y z \<noteq> 0" |
|
1076 hence "(y, z) \<in> set_pmf rr" by(simp add: set_pmf_iff) |
|
1077 hence "y \<in> set_pmf q" by(auto simp add: rr_def q' set_map_pmf intro: rev_image_eqI) |
|
1078 thus "pmf q y \<noteq> 0" by(simp add: set_pmf_iff) |
|
1079 qed |
|
1080 |
|
1081 have nn_integral_pp2: "\<And>y. (\<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV) = pmf q y" |
|
1082 by (simp add: nn_integral_pmf' inj_on_def pp_def q) |
|
1083 (auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure]) |
|
1084 have nn_integral_rr1: "\<And>y. (\<integral>\<^sup>+ x. ?rr y x \<partial>count_space UNIV) = pmf q y" |
|
1085 by (simp add: nn_integral_pmf' inj_on_def rr_def q') |
|
1086 (auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure]) |
1065 have eq: "\<And>y. (\<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV) = (\<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV)" |
1087 have eq: "\<And>y. (\<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV) = (\<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV)" |
1066 proof - |
1088 by(simp add: nn_integral_pp2 nn_integral_rr1) |
1067 fix y |
1089 |
1068 have "(\<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV) = pmf q y" |
1090 def assign \<equiv> "\<lambda>y x z. ?pp y x * ?rr y z / pmf q y" |
1069 by (simp add: nn_integral_pmf' inj_on_def pp_def q) |
1091 have assign_nonneg [simp]: "\<And>y x z. 0 \<le> assign y x z" by(simp add: assign_def) |
1070 (auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure]) |
|
1071 also have "\<dots> = (\<integral>\<^sup>+ x. ?rr y x \<partial>count_space UNIV)" |
|
1072 by (simp add: nn_integral_pmf' inj_on_def rr_def q') |
|
1073 (auto simp add: ereal_pmf_map intro!: arg_cong2[where f=emeasure]) |
|
1074 finally show "?thesis y" . |
|
1075 qed |
|
1076 |
|
1077 def assign_aux \<equiv> "\<lambda>y remainder start weight z. |
|
1078 if z < start then 0 |
|
1079 else if z = start then min weight remainder |
|
1080 else if remainder + setsum (?rr y) {Suc start ..<z} < weight then min (weight - remainder - setsum (?rr y) {Suc start..<z}) (?rr y z) else 0" |
|
1081 hence assign_aux_alt_def: "\<And>y remainder start weight z. assign_aux y remainder start weight z = |
|
1082 (if z < start then 0 |
|
1083 else if z = start then min weight remainder |
|
1084 else if remainder + setsum (?rr y) {Suc start ..<z} < weight then min (weight - remainder - setsum (?rr y) {Suc start..<z}) (?rr y z) else 0)" |
|
1085 by simp |
|
1086 { fix y and remainder :: real and start and weight :: real |
|
1087 assume weight_nonneg: "0 \<le> weight" |
|
1088 let ?assign_aux = "assign_aux y remainder start weight" |
|
1089 { fix z |
|
1090 have "setsum ?assign_aux {..<z} = |
|
1091 (if z \<le> start then 0 else if remainder + setsum (?rr y) {Suc start..<z} < weight then remainder + setsum (?rr y) {Suc start..<z} else weight)" |
|
1092 proof(induction z) |
|
1093 case (Suc z) show ?case |
|
1094 by (auto simp add: Suc.IH assign_aux_alt_def[where z=z] not_less) |
|
1095 (metis add.commute add.left_commute add_increasing pmf_nonneg) |
|
1096 qed(auto simp add: assign_aux_def) } |
|
1097 note setsum_start_assign_aux = this |
|
1098 moreover { |
|
1099 assume remainder_nonneg: "0 \<le> remainder" |
|
1100 have [simp]: "\<And>z. 0 \<le> ?assign_aux z" |
|
1101 by(simp add: assign_aux_def weight_nonneg remainder_nonneg) |
|
1102 moreover have "\<And>z. \<lbrakk> ?rr y z = 0; remainder \<le> ?rr y start \<rbrakk> \<Longrightarrow> ?assign_aux z = 0" |
|
1103 using remainder_nonneg weight_nonneg |
|
1104 by(auto simp add: assign_aux_def min_def) |
|
1105 moreover have "(\<integral>\<^sup>+ z. ?assign_aux z \<partial>count_space UNIV) = |
|
1106 min weight (\<integral>\<^sup>+ z. (if z < start then 0 else if z = start then remainder else ?rr y z) \<partial>count_space UNIV)" |
|
1107 (is "?lhs = ?rhs" is "_ = min _ (\<integral>\<^sup>+ y. ?f y \<partial>_)") |
|
1108 proof - |
|
1109 have "?lhs = (SUP n. \<Sum>z<n. ereal (?assign_aux z))" |
|
1110 by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP) |
|
1111 also have "\<dots> = (SUP n. min weight (\<Sum>z<n. ?f z))" |
|
1112 proof(rule arg_cong2[where f=SUPREMUM] ext refl)+ |
|
1113 fix n |
|
1114 have "(\<Sum>z<n. ereal (?assign_aux z)) = min weight ((if n > start then remainder else 0) + setsum ?f {Suc start..<n})" |
|
1115 using weight_nonneg remainder_nonneg by(simp add: setsum_start_assign_aux min_def) |
|
1116 also have "\<dots> = min weight (setsum ?f {start..<n})" |
|
1117 by(simp add: setsum_head_upt_Suc) |
|
1118 also have "\<dots> = min weight (setsum ?f {..<n})" |
|
1119 by(intro arg_cong2[where f=min] setsum.mono_neutral_left) auto |
|
1120 finally show "(\<Sum>z<n. ereal (?assign_aux z)) = \<dots>" . |
|
1121 qed |
|
1122 also have "\<dots> = min weight (SUP n. setsum ?f {..<n})" |
|
1123 unfolding inf_min[symmetric] by(subst inf_SUP) simp |
|
1124 also have "\<dots> = ?rhs" |
|
1125 by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP remainder_nonneg) |
|
1126 finally show ?thesis . |
|
1127 qed |
|
1128 moreover note calculation } |
|
1129 moreover note calculation } |
|
1130 note setsum_start_assign_aux = this(1) |
|
1131 and assign_aux_nonneg [simp] = this(2) |
|
1132 and assign_aux_eq_0_outside = this(3) |
|
1133 and nn_integral_assign_aux = this(4) |
|
1134 { fix y and remainder :: real and start target |
|
1135 have "setsum (?rr y) {Suc start..<target} \<ge> 0" by (simp add: setsum_nonneg) |
|
1136 moreover assume "0 \<le> remainder" |
|
1137 ultimately have "assign_aux y remainder start 0 target = 0" |
|
1138 by(auto simp add: assign_aux_def min_def) } |
|
1139 note assign_aux_weight_0 [simp] = this |
|
1140 |
|
1141 def find_start \<equiv> "\<lambda>y weight. if \<exists>n. weight \<le> setsum (?rr y) {..n} then Some (LEAST n. weight \<le> setsum (?rr y) {..n}) else None" |
|
1142 have find_start_eq_Some_above: |
|
1143 "\<And>y weight n. find_start y weight = Some n \<Longrightarrow> weight \<le> setsum (?rr y) {..n}" |
|
1144 by(drule sym)(auto simp add: find_start_def split: split_if_asm intro: LeastI) |
|
1145 { fix y weight n |
|
1146 assume find_start: "find_start y weight = Some n" |
|
1147 and weight: "0 \<le> weight" |
|
1148 have "setsum (?rr y) {..n} \<le> ?rr y n + weight" |
|
1149 proof(rule ccontr) |
|
1150 assume "\<not> ?thesis" |
|
1151 hence "?rr y n + weight < setsum (?rr y) {..n}" by simp |
|
1152 moreover with weight obtain n' where "n = Suc n'" by(cases n) auto |
|
1153 ultimately have "weight \<le> setsum (?rr y) {..n'}" by simp |
|
1154 hence "(LEAST n. weight \<le> setsum (?rr y) {..n}) \<le> n'" by(rule Least_le) |
|
1155 moreover from find_start have "n = (LEAST n. weight \<le> setsum (?rr y) {..n})" |
|
1156 by(auto simp add: find_start_def split: split_if_asm) |
|
1157 ultimately show False using \<open>n = Suc n'\<close> by auto |
|
1158 qed } |
|
1159 note find_start_eq_Some_least = this |
|
1160 have find_start_0 [simp]: "\<And>y. find_start y 0 = Some 0" |
|
1161 by(auto simp add: find_start_def intro!: exI[where x=0]) |
|
1162 { fix y and weight :: real |
|
1163 assume "weight < \<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV" |
|
1164 also have "(\<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV) = (SUP n. \<Sum>z<n. ereal (?rr y z))" |
|
1165 by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP) |
|
1166 finally obtain n where "weight < (\<Sum>z<n. ?rr y z)" by(auto simp add: less_SUP_iff) |
|
1167 hence "weight \<in> dom (find_start y)" |
|
1168 by(auto simp add: find_start_def)(meson atMost_iff finite_atMost lessThan_iff less_imp_le order_trans pmf_nonneg setsum_mono3 subsetI) } |
|
1169 note in_dom_find_startI = this |
|
1170 { fix y and w w' :: real and m |
|
1171 let ?m' = "LEAST m. w' \<le> setsum (?rr y) {..m}" |
|
1172 assume "w' \<le> w" |
|
1173 also assume "find_start y w = Some m" |
|
1174 hence "w \<le> setsum (?rr y) {..m}" by(rule find_start_eq_Some_above) |
|
1175 finally have "find_start y w' = Some ?m'" by(auto simp add: find_start_def) |
|
1176 moreover from \<open>w' \<le> setsum (?rr y) {..m}\<close> have "?m' \<le> m" by(rule Least_le) |
|
1177 ultimately have "\<exists>m'. find_start y w' = Some m' \<and> m' \<le> m" by blast } |
|
1178 note find_start_mono = this[rotated] |
|
1179 |
|
1180 def assign \<equiv> "\<lambda>y x z. let used = setsum (?pp y) {..<x} |
|
1181 in case find_start y used of None \<Rightarrow> 0 |
|
1182 | Some start \<Rightarrow> assign_aux y (setsum (?rr y) {..start} - used) start (?pp y x) z" |
|
1183 hence assign_alt_def: "\<And>y x z. assign y x z = |
|
1184 (let used = setsum (?pp y) {..<x} |
|
1185 in case find_start y used of None \<Rightarrow> 0 |
|
1186 | Some start \<Rightarrow> assign_aux y (setsum (?rr y) {..start} - used) start (?pp y x) z)" |
|
1187 by simp |
|
1188 have assign_nonneg [simp]: "\<And>y x z. 0 \<le> assign y x z" |
|
1189 by(simp add: assign_def diff_le_iff find_start_eq_Some_above Let_def split: option.split) |
|
1190 have assign_eq_0_outside: "\<And>y x z. \<lbrakk> ?pp y x = 0 \<or> ?rr y z = 0 \<rbrakk> \<Longrightarrow> assign y x z = 0" |
1092 have assign_eq_0_outside: "\<And>y x z. \<lbrakk> ?pp y x = 0 \<or> ?rr y z = 0 \<rbrakk> \<Longrightarrow> assign y x z = 0" |
1191 by(auto simp add: assign_def assign_aux_eq_0_outside diff_le_iff find_start_eq_Some_above find_start_eq_Some_least setsum_nonneg Let_def split: option.split) |
1093 by(auto simp add: assign_def) |
1192 |
|
1193 { fix y x z |
|
1194 have "(\<Sum>n<Suc x. assign y n z) = |
|
1195 (case find_start y (setsum (?pp y) {..<x}) of None \<Rightarrow> ?rr y z |
|
1196 | Some m \<Rightarrow> if z < m then ?rr y z |
|
1197 else min (?rr y z) (max 0 (setsum (?pp y) {..<x} + ?pp y x - setsum (?rr y) {..<z})))" |
|
1198 (is "?lhs x = ?rhs x") |
|
1199 proof(induction x) |
|
1200 case 0 thus ?case |
|
1201 by(auto simp add: assign_def assign_aux_def setsum_head_upt_Suc atLeast0LessThan[symmetric] not_less field_simps max_def) |
|
1202 next |
|
1203 case (Suc x) |
|
1204 have "?lhs (Suc x) = ?lhs x + assign y (Suc x) z" by simp |
|
1205 also have "?lhs x = ?rhs x" by(rule Suc.IH) |
|
1206 also have "?rhs x + assign y (Suc x) z = ?rhs (Suc x)" |
|
1207 proof(cases "find_start y (setsum (?pp y) {..<Suc x})") |
|
1208 case None |
|
1209 thus ?thesis |
|
1210 by(auto split: option.split simp add: assign_def min_def max_def diff_le_iff setsum_nonneg not_le field_simps) |
|
1211 (metis add.commute add_increasing find_start_def lessThan_Suc_atMost less_imp_le option.distinct(1) setsum_lessThan_Suc)+ |
|
1212 next |
|
1213 case (Some m) |
|
1214 have [simp]: "setsum (?rr y) {..m} = ?rr y m + setsum (?rr y) {..<m}" |
|
1215 by(simp add: ivl_disj_un(2)[symmetric]) |
|
1216 from Some obtain m' where m': "find_start y (setsum (?pp y) {..<x}) = Some m'" "m' \<le> m" |
|
1217 by(auto dest: find_start_mono[where w'2="setsum (?pp y) {..<x}"]) |
|
1218 moreover { |
|
1219 assume "z < m" |
|
1220 then have "setsum (?rr y) {..z} \<le> setsum (?rr y) {..<m}" |
|
1221 by(auto intro: setsum_mono3) |
|
1222 also have "\<dots> \<le> setsum (?pp y) {..<Suc x}" using find_start_eq_Some_least[OF Some] |
|
1223 by(simp add: ivl_disj_un(2)[symmetric] setsum_nonneg) |
|
1224 finally have "?rr y z \<le> max 0 (setsum (?pp y) {..<x} + ?pp y x - setsum (?rr y) {..<z})" |
|
1225 by(auto simp add: ivl_disj_un(2)[symmetric] max_def diff_le_iff simp del: pmf_le_0_iff) |
|
1226 } moreover { |
|
1227 assume "m \<le> z" |
|
1228 have "setsum (?pp y) {..<Suc x} \<le> setsum (?rr y) {..m}" |
|
1229 using Some by(rule find_start_eq_Some_above) |
|
1230 also have "\<dots> \<le> setsum (?rr y) {..<Suc z}" using \<open>m \<le> z\<close> by(intro setsum_mono3) auto |
|
1231 finally have "max 0 (setsum (?pp y) {..<x} + ?pp y x - setsum (?rr y) {..<z}) \<le> ?rr y z" by simp |
|
1232 moreover have "z \<noteq> m \<Longrightarrow> setsum (?rr y) {..m} + setsum (?rr y) {Suc m..<z} = setsum (?rr y) {..<z}" |
|
1233 using \<open>m \<le> z\<close> |
|
1234 by(subst ivl_disj_un(8)[where l="Suc m", symmetric]) |
|
1235 (simp_all add: setsum_Un ivl_disj_un(2)[symmetric] setsum.neutral) |
|
1236 moreover note calculation |
|
1237 } moreover { |
|
1238 assume "m < z" |
|
1239 have "setsum (?pp y) {..<Suc x} \<le> setsum (?rr y) {..m}" |
|
1240 using Some by(rule find_start_eq_Some_above) |
|
1241 also have "\<dots> \<le> setsum (?rr y) {..<z}" using \<open>m < z\<close> by(intro setsum_mono3) auto |
|
1242 finally have "max 0 (setsum (?pp y) {..<Suc x} - setsum (?rr y) {..<z}) = 0" by simp } |
|
1243 moreover have "setsum (?pp y) {..<Suc x} \<ge> setsum (?rr y) {..<m}" |
|
1244 using find_start_eq_Some_least[OF Some] |
|
1245 by(simp add: setsum_nonneg ivl_disj_un(2)[symmetric]) |
|
1246 moreover hence "setsum (?pp y) {..<Suc (Suc x)} \<ge> setsum (?rr y) {..<m}" |
|
1247 by(fastforce intro: order_trans) |
|
1248 ultimately show ?thesis using Some |
|
1249 by(auto simp add: assign_def assign_aux_def Let_def field_simps max_def) |
|
1250 qed |
|
1251 finally show ?case . |
|
1252 qed } |
|
1253 note setsum_assign = this |
|
1254 |
|
1255 have nn_integral_assign1: "\<And>y z. (\<integral>\<^sup>+ x. assign y x z \<partial>count_space UNIV) = ?rr y z" |
1094 have nn_integral_assign1: "\<And>y z. (\<integral>\<^sup>+ x. assign y x z \<partial>count_space UNIV) = ?rr y z" |
1256 proof - |
1095 proof - |
1257 fix y z |
1096 fix y z |
1258 have "(\<integral>\<^sup>+ x. assign y x z \<partial>count_space UNIV) = (SUP n. ereal (\<Sum>x<n. assign y x z))" |
1097 have "(\<integral>\<^sup>+ x. assign y x z \<partial>count_space UNIV) = |
1259 by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP) |
1098 (\<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV) * (?rr y z / pmf q y)" |
1260 also have "\<dots> = ?rr y z" |
1099 by(simp add: assign_def nn_integral_multc times_ereal.simps(1)[symmetric] divide_real_def mult.assoc del: times_ereal.simps(1)) |
1261 proof(rule antisym) |
1100 also have "\<dots> = ?rr y z" by(simp add: rr_0 nn_integral_pp2) |
1262 show "(SUP n. ereal (\<Sum>x<n. assign y x z)) \<le> ?rr y z" |
|
1263 proof(rule SUP_least) |
|
1264 fix n |
|
1265 show "ereal (\<Sum>x<n. (assign y x z)) \<le> ?rr y z" |
|
1266 using setsum_assign[of y z "n - 1"] |
|
1267 by(cases n)(simp_all split: option.split) |
|
1268 qed |
|
1269 show "?rr y z \<le> (SUP n. ereal (\<Sum>x<n. assign y x z))" |
|
1270 proof(cases "setsum (?rr y) {..z} < \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV") |
|
1271 case True |
|
1272 then obtain n where "setsum (?rr y) {..z} < setsum (?pp y) {..<n}" |
|
1273 by(auto simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP less_SUP_iff) |
|
1274 moreover have "\<And>k. k < z \<Longrightarrow> setsum (?rr y) {..k} \<le> setsum (?rr y) {..<z}" |
|
1275 by(auto intro: setsum_mono3) |
|
1276 ultimately have "?rr y z \<le> (\<Sum>x<Suc n. assign y x z)" |
|
1277 by(subst setsum_assign)(auto split: option.split dest!: find_start_eq_Some_above simp add: ivl_disj_un(2)[symmetric] add.commute add_increasing le_diff_eq le_max_iff_disj) |
|
1278 also have "\<dots> \<le> (SUP n. ereal (\<Sum>x<n. assign y x z))" |
|
1279 by(rule SUP_upper) simp |
|
1280 finally show ?thesis by simp |
|
1281 next |
|
1282 case False |
|
1283 have "setsum (?rr y) {..z} = \<integral>\<^sup>+ z. ?rr y z \<partial>count_space {..z}" |
|
1284 by(simp add: nn_integral_count_space_finite max_def) |
|
1285 also have "\<dots> \<le> \<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV" |
|
1286 by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_mono) |
|
1287 also have "\<dots> = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" by(simp add: eq) |
|
1288 finally have *: "setsum (?rr y) {..z} = \<dots>" using False by simp |
|
1289 also have "\<dots> = (SUP n. ereal (\<Sum>x<n. ?pp y x))" |
|
1290 by(simp add: nn_integral_count_space_nat suminf_ereal_eq_SUP) |
|
1291 also have "\<dots> \<le> (SUP n. ereal (\<Sum>x<n. assign y x z)) + setsum (?rr y) {..<z}" |
|
1292 proof(rule SUP_least) |
|
1293 fix n |
|
1294 have "setsum (?pp y) {..<n} = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space {..<n}" |
|
1295 by(simp add: nn_integral_count_space_finite max_def) |
|
1296 also have "\<dots> \<le> \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" |
|
1297 by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_mono) |
|
1298 also have "\<dots> = setsum (?rr y) {..z}" using * by simp |
|
1299 finally obtain k where k: "find_start y (setsum (?pp y) {..<n}) = Some k" |
|
1300 by(fastforce simp add: find_start_def) |
|
1301 with \<open>ereal (setsum (?pp y) {..<n}) \<le> setsum (?rr y) {..z}\<close> |
|
1302 have "k \<le> z" by(auto simp add: find_start_def split: split_if_asm intro: Least_le) |
|
1303 then have "setsum (?pp y) {..<n} - setsum (?rr y) {..<z} \<le> ereal (\<Sum>x<Suc n. assign y x z)" |
|
1304 using \<open>ereal (setsum (?pp y) {..<n}) \<le> setsum (?rr y) {..z}\<close> |
|
1305 apply (subst setsum_assign) |
|
1306 apply (auto simp add: field_simps max_def k ivl_disj_un(2)[symmetric]) |
|
1307 apply (meson add_increasing le_cases pmf_nonneg) |
|
1308 done |
|
1309 also have "\<dots> \<le> (SUP n. ereal (\<Sum>x<n. assign y x z))" |
|
1310 by(rule SUP_upper) simp |
|
1311 finally show "ereal (\<Sum>x<n. ?pp y x) \<le> \<dots> + setsum (?rr y) {..<z}" |
|
1312 by(simp add: ereal_minus(1)[symmetric] ereal_minus_le del: ereal_minus(1)) |
|
1313 qed |
|
1314 finally show ?thesis |
|
1315 by(simp add: ivl_disj_un(2)[symmetric] plus_ereal.simps(1)[symmetric] ereal_add_le_add_iff2 del: plus_ereal.simps(1)) |
|
1316 qed |
|
1317 qed |
|
1318 finally show "?thesis y z" . |
1101 finally show "?thesis y z" . |
1319 qed |
1102 qed |
1320 |
1103 have nn_integral_assign2: "\<And>y x. (\<integral>\<^sup>+ z. assign y x z \<partial>count_space UNIV) = ?pp y x" |
1321 { fix y x |
1104 proof - |
1322 have "(\<integral>\<^sup>+ z. assign y x z \<partial>count_space UNIV) = ?pp y x" |
1105 fix x y |
1323 proof(cases "setsum (?pp y) {..<x} = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV") |
1106 have "(\<integral>\<^sup>+ z. assign y x z \<partial>count_space UNIV) = (\<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV) * (?pp y x / pmf q y)" |
1324 case False |
1107 by(simp add: assign_def divide_real_def mult.commute[where a="?pp y x"] mult.assoc nn_integral_multc times_ereal.simps(1)[symmetric] del: times_ereal.simps(1)) |
1325 let ?used = "setsum (?pp y) {..<x}" |
1108 also have "\<dots> = ?pp y x" by(simp add: nn_integral_rr1 pp_0) |
1326 have "?used = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space {..<x}" |
1109 finally show "?thesis y x" . |
1327 by(simp add: nn_integral_count_space_finite max_def) |
1110 qed |
1328 also have "\<dots> \<le> \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" |
|
1329 by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_mono) |
|
1330 finally have "?used < \<dots>" using False by auto |
|
1331 also note eq finally have "?used \<in> dom (find_start y)" by(rule in_dom_find_startI) |
|
1332 then obtain k where k: "find_start y ?used = Some k" by auto |
|
1333 let ?f = "\<lambda>z. if z < k then 0 else if z = k then setsum (?rr y) {..k} - ?used else ?rr y z" |
|
1334 let ?g = "\<lambda>x'. if x' < x then 0 else ?pp y x'" |
|
1335 have "?pp y x = ?g x" by simp |
|
1336 also have "?g x \<le> \<integral>\<^sup>+ x'. ?g x' \<partial>count_space UNIV" by(rule nn_integral_ge_point) simp |
|
1337 also { |
|
1338 have "?used = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space {..<x}" |
|
1339 by(simp add: nn_integral_count_space_finite max_def) |
|
1340 also have "\<dots> = \<integral>\<^sup>+ x'. (if x' < x then ?pp y x' else 0) \<partial>count_space UNIV" |
|
1341 by(simp add: nn_integral_count_space_indicator indicator_def if_distrib zero_ereal_def cong del: if_cong) |
|
1342 also have "(\<integral>\<^sup>+ x'. ?g x' \<partial>count_space UNIV) + \<dots> = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" |
|
1343 by(subst nn_integral_add[symmetric])(auto intro: nn_integral_cong) |
|
1344 also note calculation } |
|
1345 ultimately have "ereal (?pp y x) + ?used \<le> \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" |
|
1346 by (metis (no_types, lifting) ereal_add_mono order_refl) |
|
1347 also note eq |
|
1348 also have "(\<integral>\<^sup>+ z. ?rr y z \<partial>count_space UNIV) = (\<integral>\<^sup>+ z. ?f z \<partial>count_space UNIV) + (\<integral>\<^sup>+ z. (if z < k then ?rr y z else if z = k then ?used - setsum (?rr y) {..<k} else 0) \<partial>count_space UNIV)" |
|
1349 using k by(subst nn_integral_add[symmetric])(auto intro!: nn_integral_cong simp add: ivl_disj_un(2)[symmetric] setsum_nonneg dest: find_start_eq_Some_least find_start_eq_Some_above) |
|
1350 also have "(\<integral>\<^sup>+ z. (if z < k then ?rr y z else if z = k then ?used - setsum (?rr y) {..<k} else 0) \<partial>count_space UNIV) = |
|
1351 (\<integral>\<^sup>+ z. (if z < k then ?rr y z else if z = k then ?used - setsum (?rr y) {..<k} else 0) \<partial>count_space {..k})" |
|
1352 by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_cong) |
|
1353 also have "\<dots> = ?used" |
|
1354 using k by(auto simp add: nn_integral_count_space_finite max_def ivl_disj_un(2)[symmetric] diff_le_iff setsum_nonneg dest: find_start_eq_Some_least) |
|
1355 finally have "?pp y x \<le> (\<integral>\<^sup>+ z. ?f z \<partial>count_space UNIV)" |
|
1356 by(cases "\<integral>\<^sup>+ z. ?f z \<partial>count_space UNIV") simp_all |
|
1357 then show ?thesis using k |
|
1358 by(simp add: assign_def nn_integral_assign_aux diff_le_iff find_start_eq_Some_above min_def) |
|
1359 next |
|
1360 case True |
|
1361 have "setsum (?pp y) {..x} = \<integral>\<^sup>+ x. ?pp y x \<partial>count_space {..x}" |
|
1362 by(simp add: nn_integral_count_space_finite max_def) |
|
1363 also have "\<dots> \<le> \<integral>\<^sup>+ x. ?pp y x \<partial>count_space UNIV" |
|
1364 by(auto simp add: nn_integral_count_space_indicator indicator_def intro: nn_integral_mono) |
|
1365 also have "\<dots> = setsum (?pp y) {..<x}" by(simp add: True) |
|
1366 finally have "?pp y x = 0" by(simp add: ivl_disj_un(2)[symmetric] eq_iff del: pmf_le_0_iff) |
|
1367 thus ?thesis |
|
1368 by(cases "find_start y (setsum (?pp y) {..<x})")(simp_all add: assign_def diff_le_iff find_start_eq_Some_above) |
|
1369 qed } |
|
1370 note nn_integral_assign2 = this |
|
1371 |
1111 |
1372 def a \<equiv> "embed_pmf (\<lambda>(y, x, z). assign y x z)" |
1112 def a \<equiv> "embed_pmf (\<lambda>(y, x, z). assign y x z)" |
1373 { fix y x z |
1113 { fix y x z |
1374 have "assign y x z = pmf a (y, x, z)" |
1114 have "assign y x z = pmf a (y, x, z)" |
1375 unfolding a_def |
1115 unfolding a_def |