852 then show ?thesis |
855 then show ?thesis |
853 by (intro pmf_eqI) simp |
856 by (intro pmf_eqI) simp |
854 qed |
857 qed |
855 |
858 |
856 lemma bind_cond_pmf_cancel: |
859 lemma bind_cond_pmf_cancel: |
857 assumes in_S: "\<And>x. x \<in> set_pmf p \<Longrightarrow> x \<in> S x" "\<And>x. x \<in> set_pmf q \<Longrightarrow> x \<in> S x" |
860 assumes [simp]: "\<And>x. x \<in> set_pmf p \<Longrightarrow> set_pmf q \<inter> {y. R x y} \<noteq> {}" |
858 assumes S_eq: "\<And>x y. x \<in> S y \<Longrightarrow> S x = S y" |
861 assumes [simp]: "\<And>y. y \<in> set_pmf q \<Longrightarrow> set_pmf p \<inter> {x. R x y} \<noteq> {}" |
859 and same: "\<And>x. measure (measure_pmf p) (S x) = measure (measure_pmf q) (S x)" |
862 assumes [simp]: "\<And>x y. x \<in> set_pmf p \<Longrightarrow> y \<in> set_pmf q \<Longrightarrow> R x y \<Longrightarrow> measure q {y. R x y} = measure p {x. R x y}" |
860 shows "bind_pmf p (\<lambda>x. cond_pmf q (S x)) = q" (is "?lhs = _") |
863 shows "bind_pmf p (\<lambda>x. cond_pmf q {y. R x y}) = q" |
861 proof (rule pmf_eqI) |
864 proof (rule pmf_eqI) |
862 { fix x |
865 fix i |
863 assume "x \<in> set_pmf p" |
866 have "ereal (pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i) = |
864 hence "set_pmf p \<inter> (S x) \<noteq> {}" using in_S by auto |
867 (\<integral>\<^sup>+x. ereal (pmf q i / measure p {x. R x i}) * ereal (indicator {x. R x i} x) \<partial>p)" |
865 hence "measure (measure_pmf p) (S x) \<noteq> 0" |
868 by (auto simp add: ereal_pmf_bind AE_measure_pmf_iff pmf_cond pmf_eq_0_set_pmf intro!: nn_integral_cong_AE) |
866 by(auto simp add: measure_pmf.prob_eq_0 AE_measure_pmf_iff) |
869 also have "\<dots> = (pmf q i * measure p {x. R x i}) / measure p {x. R x i}" |
867 with same have "measure (measure_pmf q) (S x) \<noteq> 0" by simp |
870 by (simp add: pmf_nonneg measure_nonneg zero_ereal_def[symmetric] ereal_indicator |
868 hence "set_pmf q \<inter> S x \<noteq> {}" |
871 nn_integral_cmult measure_pmf.emeasure_eq_measure) |
869 by(auto simp add: measure_pmf.prob_eq_0 AE_measure_pmf_iff) } |
872 also have "\<dots> = pmf q i" |
870 note [simp] = this |
873 by (cases "pmf q i = 0") (simp_all add: pmf_eq_0_set_pmf measure_measure_pmf_not_zero) |
871 |
874 finally show "pmf (bind_pmf p (\<lambda>x. cond_pmf q {y. R x y})) i = pmf q i" |
872 fix z |
875 by simp |
873 have pmf_q_z: "z \<notin> S z \<Longrightarrow> pmf q z = 0" |
|
874 by(erule contrapos_np)(simp add: pmf_eq_0_set_pmf in_S) |
|
875 |
|
876 have "ereal (pmf ?lhs z) = \<integral>\<^sup>+ x. ereal (pmf (cond_pmf q (S x)) z) \<partial>measure_pmf p" |
|
877 by(simp add: ereal_pmf_bind) |
|
878 also have "\<dots> = \<integral>\<^sup>+ x. ereal (pmf q z / measure p (S z)) * indicator (S z) x \<partial>measure_pmf p" |
|
879 by(rule nn_integral_cong_AE)(auto simp add: AE_measure_pmf_iff pmf_cond same pmf_q_z in_S dest!: S_eq split: split_indicator) |
|
880 also have "\<dots> = pmf q z" using pmf_nonneg[of q z] |
|
881 by (subst nn_integral_cmult)(auto simp add: measure_nonneg measure_pmf.emeasure_eq_measure same measure_pmf.prob_eq_0 AE_measure_pmf_iff pmf_eq_0_set_pmf in_S) |
|
882 finally show "pmf ?lhs z = pmf q z" by simp |
|
883 qed |
876 qed |
884 |
877 |
885 subsection \<open> Relator \<close> |
878 subsection \<open> Relator \<close> |
886 |
879 |
887 inductive rel_pmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf \<Rightarrow> bool" |
880 inductive rel_pmf :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a pmf \<Rightarrow> 'b pmf \<Rightarrow> bool" |
926 from pq obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y" |
919 from pq obtain pq where pq: "\<And>x y. (x, y) \<in> set_pmf pq \<Longrightarrow> R x y" |
927 and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto |
920 and p: "p = map_pmf fst pq" and q: "q = map_pmf snd pq" by cases auto |
928 from qr obtain qr where qr: "\<And>y z. (y, z) \<in> set_pmf qr \<Longrightarrow> S y z" |
921 from qr obtain qr where qr: "\<And>y z. (y, z) \<in> set_pmf qr \<Longrightarrow> S y z" |
929 and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto |
922 and q': "q = map_pmf fst qr" and r: "r = map_pmf snd qr" by cases auto |
930 |
923 |
931 def pr \<equiv> "bind_pmf pq (\<lambda>(x, y). bind_pmf (cond_pmf qr {(y', z). y' = y}) (\<lambda>(y', z). return_pmf (x, z)))" |
924 def pr \<equiv> "bind_pmf pq (\<lambda>xy. bind_pmf (cond_pmf qr {yz. fst yz = snd xy}) (\<lambda>yz. return_pmf (fst xy, snd yz)))" |
932 have pr_welldefined: "\<And>y. y \<in> q \<Longrightarrow> qr \<inter> {(y', z). y' = y} \<noteq> {}" |
925 have pr_welldefined: "\<And>y. y \<in> q \<Longrightarrow> qr \<inter> {yz. fst yz = y} \<noteq> {}" |
933 by (force simp: q') |
926 by (force simp: q') |
934 |
927 |
935 have "rel_pmf (R OO S) p r" |
928 have "rel_pmf (R OO S) p r" |
936 proof (rule rel_pmf.intros) |
929 proof (rule rel_pmf.intros) |
937 fix x z assume "(x, z) \<in> pr" |
930 fix x z assume "(x, z) \<in> pr" |
938 then have "\<exists>y. (x, y) \<in> pq \<and> (y, z) \<in> qr" |
931 then have "\<exists>y. (x, y) \<in> pq \<and> (y, z) \<in> qr" |
939 by (auto simp: q pr_welldefined pr_def split_beta) |
932 by (auto simp: q pr_welldefined pr_def split_beta) |
940 with pq qr show "(R OO S) x z" |
933 with pq qr show "(R OO S) x z" |
941 by blast |
934 by blast |
942 next |
935 next |
943 have "map_pmf snd pr = map_pmf snd (bind_pmf q (\<lambda>y. cond_pmf qr {(y', z). y' = y}))" |
936 have "map_pmf snd pr = map_pmf snd (bind_pmf q (\<lambda>y. cond_pmf qr {yz. fst yz = y}))" |
944 by (simp add: pr_def q split_beta bind_map_pmf map_pmf_def[symmetric] map_bind_pmf map_return_pmf) |
937 by (simp add: pr_def q split_beta bind_map_pmf map_pmf_def[symmetric] map_bind_pmf map_return_pmf map_pmf_comp) |
945 then show "map_pmf snd pr = r" |
938 then show "map_pmf snd pr = r" |
946 unfolding r q' bind_map_pmf by (subst (asm) bind_cond_pmf_cancel) auto |
939 unfolding r q' bind_map_pmf by (subst (asm) bind_cond_pmf_cancel) (auto simp: eq_commute) |
947 qed (simp add: pr_def map_bind_pmf split_beta map_return_pmf map_pmf_def[symmetric] p) } |
940 qed (simp add: pr_def map_bind_pmf split_beta map_return_pmf map_pmf_def[symmetric] p map_pmf_comp) } |
948 then show "rel_pmf R OO rel_pmf S \<le> rel_pmf (R OO S)" |
941 then show "rel_pmf R OO rel_pmf S \<le> rel_pmf (R OO S)" |
949 by(auto simp add: le_fun_def) |
942 by(auto simp add: le_fun_def) |
950 qed (fact natLeq_card_order natLeq_cinfinite)+ |
943 qed (fact natLeq_card_order natLeq_cinfinite)+ |
951 |
944 |
952 lemma rel_pmf_conj[simp]: |
945 lemma rel_pmf_conj[simp]: |
1136 assumes 1: "rel_pmf R p q" |
1129 assumes 1: "rel_pmf R p q" |
1137 assumes 2: "rel_pmf R q p" |
1130 assumes 2: "rel_pmf R q p" |
1138 and refl: "reflp R" and trans: "transp R" |
1131 and refl: "reflp R" and trans: "transp R" |
1139 shows "rel_pmf (inf R R\<inverse>\<inverse>) p q" |
1132 shows "rel_pmf (inf R R\<inverse>\<inverse>) p q" |
1140 proof |
1133 proof |
1141 let ?E = "\<lambda>x. {y. R x y \<and> R y x}" |
1134 let ?R = "\<lambda>x y. R x y \<and> R y x" |
1142 let ?\<mu>E = "\<lambda>x. measure q (?E x)" |
1135 let ?\<mu>R = "\<lambda>y. measure q {x. ?R x y}" |
1143 { fix x |
1136 { fix x |
1144 have "measure p (?E x) = measure p ({y. R x y} - {y. R x y \<and> \<not> R y x})" |
1137 have "measure p {y. ?R x y} = measure p ({y. R x y} - {y. R x y \<and> \<not> R y x})" |
1145 by(auto intro!: arg_cong[where f="measure p"]) |
1138 by(auto intro!: arg_cong[where f="measure p"]) |
1146 also have "\<dots> = measure p {y. R x y} - measure p {y. R x y \<and> \<not> R y x}" |
1139 also have "\<dots> = measure p {y. R x y} - measure p {y. R x y \<and> \<not> R y x}" |
1147 by (rule measure_pmf.finite_measure_Diff) auto |
1140 by (rule measure_pmf.finite_measure_Diff) auto |
1148 also have "measure p {y. R x y \<and> \<not> R y x} = measure q {y. R x y \<and> \<not> R y x}" |
1141 also have "measure p {y. R x y \<and> \<not> R y x} = measure q {y. R x y \<and> \<not> R y x}" |
1149 using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ioi) |
1142 using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ioi) |
1150 also have "measure p {y. R x y} = measure q {y. R x y}" |
1143 also have "measure p {y. R x y} = measure q {y. R x y}" |
1151 using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ici) |
1144 using 1 2 refl trans by(auto intro!: Orderings.antisym measure_Ici) |
1152 also have "measure q {y. R x y} - measure q {y. R x y \<and> ~ R y x} = |
1145 also have "measure q {y. R x y} - measure q {y. R x y \<and> ~ R y x} = |
1153 measure q ({y. R x y} - {y. R x y \<and> \<not> R y x})" |
1146 measure q ({y. R x y} - {y. R x y \<and> \<not> R y x})" |
1154 by(rule measure_pmf.finite_measure_Diff[symmetric]) auto |
1147 by(rule measure_pmf.finite_measure_Diff[symmetric]) auto |
1155 also have "\<dots> = ?\<mu>E x" |
1148 also have "\<dots> = ?\<mu>R x" |
1156 by(auto intro!: arg_cong[where f="measure q"]) |
1149 by(auto intro!: arg_cong[where f="measure q"]) |
1157 also note calculation } |
1150 also note calculation } |
1158 note eq = this |
1151 note eq = this |
1159 |
1152 |
1160 def pq \<equiv> "bind_pmf p (\<lambda>x. bind_pmf (cond_pmf q (?E x)) (\<lambda>y. return_pmf (x, y)))" |
1153 def pq \<equiv> "bind_pmf p (\<lambda>x. bind_pmf (cond_pmf q {y. ?R y x}) (\<lambda>y. return_pmf (x, y)))" |
1161 |
1154 |
1162 show "map_pmf fst pq = p" |
1155 show "map_pmf fst pq = p" |
1163 by(simp add: pq_def map_bind_pmf map_return_pmf bind_return_pmf') |
1156 by(simp add: pq_def map_bind_pmf map_return_pmf bind_return_pmf') |
1164 |
1157 |
|
1158 { fix y assume "y \<in> set_pmf p" then have "set_pmf q \<inter> {x. ?R x y} \<noteq> {}" |
|
1159 unfolding measure_pmf_zero_iff[symmetric] eq[symmetric] by (auto simp: measure_pmf_zero_iff intro: reflpD[OF refl]) } |
|
1160 note set_p = this |
|
1161 moreover |
|
1162 { fix x assume "x \<in> set_pmf q" then have "set_pmf p \<inter> {y. R x y \<and> R y x} \<noteq> {}" |
|
1163 unfolding measure_pmf_zero_iff[symmetric] eq by (auto simp: measure_pmf_zero_iff intro: reflpD[OF refl]) } |
|
1164 ultimately |
1165 show "map_pmf snd pq = q" |
1165 show "map_pmf snd pq = q" |
1166 unfolding pq_def map_bind_pmf map_return_pmf bind_return_pmf' snd_conv |
1166 unfolding pq_def map_bind_pmf map_return_pmf bind_return_pmf' snd_conv |
1167 by(subst bind_cond_pmf_cancel)(auto simp add: reflpD[OF \<open>reflp R\<close>] eq intro: transpD[OF \<open>transp R\<close>]) |
1167 by (subst bind_cond_pmf_cancel) |
1168 |
1168 (auto simp add: eq AE_measure_pmf_iff dest: transpD[OF trans] |
1169 fix x y |
1169 intro!: measure_pmf.finite_measure_eq_AE) |
1170 assume "(x, y) \<in> set_pmf pq" |
1170 |
1171 moreover |
1171 fix x y assume "(x, y) \<in> set_pmf pq" with set_p show "inf R R\<inverse>\<inverse> x y" |
1172 { assume "x \<in> set_pmf p" |
|
1173 hence "measure (measure_pmf p) (?E x) \<noteq> 0" |
|
1174 by (auto simp add: measure_pmf.prob_eq_0 AE_measure_pmf_iff intro: reflpD[OF \<open>reflp R\<close>]) |
|
1175 hence "measure (measure_pmf q) (?E x) \<noteq> 0" using eq by simp |
|
1176 hence "set_pmf q \<inter> {y. R x y \<and> R y x} \<noteq> {}" |
|
1177 by (auto simp add: measure_pmf.prob_eq_0 AE_measure_pmf_iff) } |
|
1178 ultimately show "inf R R\<inverse>\<inverse> x y" |
|
1179 by (auto simp add: pq_def) |
1172 by (auto simp add: pq_def) |
1180 qed |
1173 qed |
1181 |
1174 |
1182 lemma rel_pmf_antisym: |
1175 lemma rel_pmf_antisym: |
1183 fixes p q :: "'a pmf" |
1176 fixes p q :: "'a pmf" |