Theory Rel_PMF_Characterisation

theory Rel_PMF_Characterisation
imports Max_Flow_Min_Cut_Countable
(* Author: Andreas Lochbihler, ETH Zurich *)

theory Rel_PMF_Characterisation imports
  Max_Flow_Min_Cut_Countable
begin

section ‹Characterisation of @{const rel_pmf}›

context begin

private datatype ('a, 'b) vertex = Source | Sink | Left 'a | Right 'b

private lemma inj_Left [simp]: "inj_on Left X"
by(simp add: inj_on_def)

private lemma inj_Right [simp]: "inj_on Right X"
by(simp add: inj_on_def)

proposition rel_pmf_measureI:
  fixes p :: "'a pmf" and q :: "'b pmf"
  assumes le: "⋀A. measure (measure_pmf p) A ≤ measure (measure_pmf q) {y. ∃x∈A. R x y}"
  shows "rel_pmf R p q"
proof -
  have DomR: "∃y∈set_pmf q. R x y" if "x ∈ set_pmf p" for x
  proof(rule ccontr)
    assume *: "¬ ?thesis"
    from le[of "{x}"] have "pmf p x ≤ measure (measure_pmf q) {y. R x y}"
      by(auto simp add: pmf.rep_eq)
    also have "… = 0" using * by(auto simp add: measure_pmf_zero_iff)
    finally show False using that by(auto dest: pmf_positive)
  qed
  have RanR: "∃x∈set_pmf p. R x y" if "y ∈ set_pmf q" for y
  proof(rule ccontr)
    assume *: "¬ ?thesis"
    then have "measure (measure_pmf q) {y. ∃x∈set_pmf p. R x y} + measure (measure_pmf q) {y} = measure (measure_pmf q) ({y. ∃x∈set_pmf p. R x y} ∪ {y})"
      by(intro measure_pmf.finite_measure_Union[symmetric]) auto
    also have "… ≤ 1" by simp
    also have "measure (measure_pmf p) (set_pmf p) = 1" by(simp add: measure_pmf.prob_eq_1 AE_measure_pmf)
    then have "1 ≤ measure (measure_pmf q) {y. ∃x∈set_pmf p. R x y}" using le[of "set_pmf p"] by auto
    ultimately have "measure (measure_pmf q) {y} ≤ 0" by auto
    with that show False by(auto dest: pmf_positive simp add: pmf.rep_eq)
  qed

  def edge'  "λxv yv. case (xv, yv) of
      (Source, Left x) ⇒ x ∈ set_pmf p
    | (Left x, Right y) ⇒ R x y ∧ x ∈ set_pmf p ∧ y ∈ set_pmf q
    | (Right y, Sink) ⇒ y ∈ set_pmf q
    | _ ⇒ False"
  have edge'_simps [simp]:
    "edge' xv (Left x) ⟷ xv = Source ∧ x ∈ set_pmf p"
    "edge' (Left x) (Right y) ⟷ R x y ∧ x ∈ set_pmf p ∧ y ∈ set_pmf q"
    "edge' (Right y) yv ⟷ yv = Sink ∧ y ∈ set_pmf q"
    "edge' Source (Right y) ⟷ False"
    "edge' Source Sink ⟷ False"
    "edge' xv Source ⟷ False"
    "edge' Sink yv ⟷ False"
    "edge' (Left x) Sink ⟷ False"
    for xv yv x y by(simp_all add: edge'_def split: vertex.split)
  have edge'_cases[cases pred]: thesis if "edge' xv yv"
    "⋀x. ⟦ xv = Source; yv = Left x; x ∈ set_pmf p ⟧ ⟹ thesis"
    "⋀x y. ⟦ xv = Left x; yv = Right y; R x y; x ∈ set_pmf p; y ∈ set_pmf q ⟧ ⟹ thesis"
    "⋀y. ⟦ xv = Right y; yv = Sink; y ∈ set_pmf q ⟧ ⟹ thesis"
    for thesis xv yv using that by(simp add: edge'_def split: prod.split_asm vertex.split_asm)
  have edge'_SourceE [elim!]: thesis if "edge' Source yv" "⋀x. ⟦ yv = Left x; x ∈ set_pmf p ⟧ ⟹ thesis"
    for yv thesis using that by(auto elim: edge'_cases)
  have edge'_LeftE [elim!]: thesis if "edge' (Left x) yv" "⋀y. ⟦ yv = Right y; R x y; x ∈ set_pmf p; y ∈ set_pmf q ⟧ ⟹ thesis"
    for x yv thesis using that by(auto elim: edge'_cases)
  have edge'_RightE [elim!]: thesis if "edge' xv (Right y)" "⋀x. ⟦ xv = Left x; R x y; x ∈ set_pmf p; y ∈ set_pmf q ⟧ ⟹ thesis"
    for xv y thesis using that by(auto elim: edge'_cases)
  have edge'_SinkE [elim!]: thesis if "edge' xv Sink" "⋀y. ⟦ xv = Right y; y ∈ set_pmf q ⟧ ⟹ thesis"
    for xv thesis using that by(auto elim: edge'_cases)
  have edge'1: "x ∈ set_pmf p ⟹ edge' Source (Left x)"
    and edge'2: "⟦ R x y; x ∈ set_pmf p; y ∈ set_pmf q ⟧ ⟹ edge' (Left x) (Right y)"
    and edge'3: "y ∈ set_pmf q ⟹ edge' (Right y) Sink"
    for x y by simp_all
  note edge'I[intro] = this

  def cap  "(λ(xv, yv). case (xv, yv) of
      (Source, Left x) ⇒ pmf p x
    | (Left x, Right y) ⇒ if R x y ∧ x ∈ set_pmf p ∧ y ∈ set_pmf q then 2 else 0
    | (Right y, Sink) ⇒ pmf q y
    | _ ⇒ 0) :: ('a, 'b) vertex flow"
  have cap_simps [simp]:
    "cap (xv, Left x) = (if xv = Source then ennreal (pmf p x) else 0)"
    "cap (Left x, Right y) = (if R x y ∧ x ∈ set_pmf p ∧ y ∈ set_pmf q then 2 else 0)"
    "cap (Right y, yv) = (if yv = Sink then ennreal (pmf q y) else 0)"
    "cap (Source, Right y) = 0"
    "cap (Source, Sink) = 0"
    "cap (xv, Source) = 0"
    "cap (Sink, yv) = 0"
    "cap (Left x, Sink) = 0"
    for xv yv x y by(simp_all add: cap_def split: vertex.split)

  def Δ  "⦇edge = edge', capacity = cap, source = Source, sink = Sink⦈"
  write Δ (structure)
  have Δ_sel [simp]:
    "edge Δ = edge'"
    "capacity Δ = cap"
    "source Δ = Source"
    "sink Δ = Sink"
    by(simp_all add: Δ_def)

  have IN_Left [simp]: "IN (Left x) = (if x ∈ set_pmf p then {Source} else {})" for x
    by(auto simp add: incoming_def)
  have OUT_Right [simp]: "OUT (Right y) = (if y ∈ set_pmf q then {Sink} else {})" for y
    by(auto simp add: outgoing_def)

  interpret network: countable_network Δ
  proof
    show "source Δ ≠ sink Δ" by simp
    show "capacity Δ e = 0" if "e ∉ E" for e using that
      by(cases e)(auto simp add: edge'_def pmf_eq_0_set_pmf split: vertex.split_asm)
    show "capacity Δ e ≠ top" for e by(auto simp add: cap_def split: prod.split vertex.split)
    have "E ⊆ ((Pair Source ∘ Left) ` set_pmf p) ∪ (map_prod Left Right ` (set_pmf p × set_pmf q)) ∪ ((λy. (Right y, Sink)) ` set_pmf q)"
      by(auto elim: edge'_cases)
    thus "countable E" by(rule countable_subset) auto
  qed
  from network.max_flow_min_cut obtain f S
    where f: "flow Δ f" and cut: "cut Δ S" and ortho: "orthogonal Δ f S" by blast
  from cut obtain Source: "Source ∈ S" and Sink: "Sink ∉ S" by cases simp

  have f_finite [simp]: "f e < top" for e
    using network.flowD_finite[OF f, of e] by (simp_all add: less_top)

  have OUT_cap_Source: "d_OUT cap Source = 1"
  proof -
    have "d_OUT cap Source = (∑+ y∈range Left. cap (Source, y))"
      by(auto 4 4 simp add: d_OUT_def nn_integral_count_space_indicator intro!: nn_integral_cong network.capacity_outside[simplified] split: split_indicator)
    also have "… = (∑+ y. pmf p y)" by(simp add: nn_integral_count_space_reindex)
    also have "… = 1" by(simp add: nn_integral_pmf)
    finally show ?thesis .
  qed
  have IN_cap_Left: "d_IN cap (Left x) = pmf p x" for x
    by(subst d_IN_alt_def[of _ Δ])(simp_all add: pmf_eq_0_set_pmf pmf_nonneg nn_integral_count_space_indicator max_def)
  have OUT_cap_Right: "d_OUT cap (Right y) = pmf q y" for y
    by(subst d_OUT_alt_def[of _ Δ])(simp_all add: pmf_eq_0_set_pmf pmf_nonneg nn_integral_count_space_indicator max_def)
  have IN_f_Left: "d_IN f (Left x) = f (Source, Left x)" for x
    by(subst d_IN_alt_def[of _ Δ])(simp_all add: nn_integral_count_space_indicator max_def network.flowD_outside[OF f])
  have OUT_f_Right: "d_OUT f (Right y) = f (Right y, Sink)" for y
    by(subst d_OUT_alt_def[of _ Δ])(simp_all add: nn_integral_count_space_indicator max_def network.flowD_outside[OF f])

  have S_LR: "Right y ∈ S" if Left: "Left x ∈ S" and edge: "R x y" "x ∈ set_pmf p" "y ∈ set_pmf q" for x y
  proof(rule ccontr)
    assume Right: "Right y ∉ S"
    from edge have "edge Δ (Left x) (Right y)" by simp
    from orthogonalD_out[OF ortho this Left Right] edge have "2 = f (Left x, Right y)" by simp
    also have "… ≤ d_OUT f (Left x)" unfolding d_OUT_def by(rule nn_integral_ge_point) simp
    also have "… = d_IN f (Left x)" using f by(rule flowD_KIR) simp_all
    also have "… ≤ d_IN cap (Left x)" using flowD_capacity_IN[OF f, of "Left x"] by simp
    also have "… = pmf p x" by(rule IN_cap_Left)
    also have "… ≤ 1" by(simp add: pmf_le_1)
    finally show False
      unfolding ennreal_numeral[symmetric] ennreal_1[symmetric] by (subst (asm) ennreal_le_iff) auto
  qed

  have "value_flow Δ f ≤ 1" using flowD_capacity_OUT[OF f, of Source] by(simp add: OUT_cap_Source)
  moreover have "1 ≤ value_flow Δ f"
  proof -
    let ?L = "Left -` S ∩ set_pmf p"
    let ?R = "{y|y x. x ∈ set_pmf p ∧ Left x ∈ S ∧ R x y ∧ y ∈ set_pmf q}"
    have "value_flow Δ f = (∑+ x∈range Left. f (Source, x))" unfolding d_OUT_def
      by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
    also have "… = (∑+ x. f (Source, Left x) * indicator ?L x) + (∑+ x. f (Source, Left x) * indicator (- ?L) x)"
      by(subst nn_integral_add[symmetric])(auto simp add: nn_integral_count_space_reindex intro!: nn_integral_cong split: split_indicator)
    also have "(∑+ x. f (Source, Left x) * indicator (- ?L) x) = (∑+ x∈- ?L. cap (Source, Left x))"
      using orthogonalD_out[OF ortho _ Source]
      apply(auto simp add: set_pmf_iff network.flowD_outside[OF f] nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
      subgoal for x by(cases "x ∈ set_pmf p")(auto simp add: set_pmf_iff zero_ennreal_def[symmetric] network.flowD_outside[OF f])
      done
    also have "… = (∑+ x∈- ?L. pmf p x)" by simp
    also have "… = emeasure (measure_pmf p) (- ?L)" by(simp add: nn_integral_pmf)
    also have "(∑+ x. f (Source, Left x) * indicator ?L x) = (∑+ x∈?L. d_IN f (Left x))"
      by(subst d_IN_alt_def[of _ Δ])(auto simp add: network.flowD_outside[OF f] nn_integral_count_space_indicator intro!: nn_integral_cong)
    also have "… = (∑+ x∈?L. d_OUT f (Left x))"
      by(rule nn_integral_cong flowD_KIR[OF f, symmetric])+ simp_all
    also have "… = (∑+ x. ∑+ y. f (Left x, y) * indicator (range Right) y * indicator ?L x)"
      by(auto simp add: d_OUT_def nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
    also have "… = (∑+ y∈range Right. ∑+ x. f (Left x, y) * indicator ?L x)"
      by(subst nn_integral_fst_count_space[where f="case_prod _", simplified])
        (simp add: nn_integral_snd_count_space[where f="case_prod _", simplified] nn_integral_count_space_indicator nn_integral_cmult[symmetric] mult_ac)
    also have "… = (∑+ y. ∑+ x. f (Left x, Right y) * indicator ?L x)"
      by(simp add: nn_integral_count_space_reindex)
    also have "… = (∑+ y. ∑+ x. f (Left x, Right y) * indicator ?L x * indicator {y|x. Right y ∈ S ∧ x ∈ set_pmf p ∧ Left x ∈ S ∧ R x y ∧ y ∈ set_pmf q} y)"
      apply(rule nn_integral_cong)+
      subgoal for y x
        by(cases "R x y" "y ∈ set_pmf q" rule: bool.exhaust[case_product bool.exhaust])
          (auto split: split_indicator dest: S_LR simp add: network.flowD_outside[OF f])
      done
    also have "… = (∑+ y. ∑+ x. f (Left x, Right y) * indicator ?R y)"
      apply(rule nn_integral_cong)+
      subgoal for y x
        by(cases "R x y" "x ∈ set_pmf p" rule: bool.exhaust[case_product bool.exhaust])
          (auto split: split_indicator simp add: orthogonalD_in[OF ortho] network.flowD_outside[OF f] dest: S_LR)
      done
    also have "… = (∑+ y∈?R. ∑+ x∈range Left. f (x, Right y))"
      by(simp add: nn_integral_count_space_reindex)(auto simp add:  nn_integral_count_space_indicator nn_integral_multc intro!: nn_integral_cong arg_cong2[where f="op *"] split: split_indicator)
    also have "… = (∑+ y∈?R. d_IN f (Right y))"
      by(auto simp add: d_IN_def nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
    also have "… = (∑+ y∈?R. d_OUT f (Right y))" using flowD_KIR[OF f] by simp
    also have "… = (∑+ y∈?R. d_OUT cap (Right y))"
      by(auto 4 3 intro!: nn_integral_cong simp add: d_OUT_def network.flowD_outside[OF f] Sink dest: S_LR intro: orthogonalD_out[OF ortho, of "Right _" Sink, simplified])
    also have "… = (∑+ y∈?R. pmf q y)" by(simp add: OUT_cap_Right)
    also have "… = emeasure (measure_pmf q) ?R"  by(simp add: nn_integral_pmf)
    also have "… ≥ emeasure (measure_pmf p) ?L" using le[of ?L]
      by(auto elim!: order_trans simp add: measure_pmf.emeasure_eq_measure AE_measure_pmf_iff intro!: measure_pmf.finite_measure_mono_AE)
    ultimately have "value_flow Δ f ≥ emeasure (measure_pmf p) ?L + emeasure (measure_pmf p) (- ?L)"
      by(simp add: add_right_mono)
    also have "emeasure (measure_pmf p) ?L + emeasure (measure_pmf p) (- ?L) = emeasure (measure_pmf p) (?L ∪ - ?L)"
      by(subst plus_emeasure) auto
    also have "?L ∪ -?L = UNIV" by blast
    finally show ?thesis by simp
  qed
  ultimately have val: "value_flow Δ f = 1" by simp

  have SAT_p: "f (Source, Left x) = pmf p x" for x
  proof(rule antisym)
    show "f (Source, Left x) ≤ pmf p x" using flowD_capacity[OF f, of "(Source, Left x)"] by simp
    show "pmf p x ≤ f (Source, Left x)"
    proof(rule ccontr)
      assume *: "¬ ?thesis"
      have finite: "(∑+ y. f (Source, Left y) * indicator (- {x}) y) ≠ ∞"
      proof -
        have "(∑+ y. f (Source, Left y) * indicator (- {x}) y) ≤ (∑+ y∈range Left. f (Source, y))"
          by(auto simp add: nn_integral_count_space_reindex intro!: nn_integral_mono split: split_indicator)
        also have "… = value_flow Δ f"
          by(auto simp add: d_OUT_def nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
        finally show ?thesis using val by (auto simp: top_unique)
      qed
      have "value_flow Δ f = (∑+ y∈range Left. f (Source, y))"
        by(auto simp add: d_OUT_def nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
      also have "… = (∑+ y. f (Source, Left y) * indicator (- {x}) y) + (∑+ y. f (Source, Left y) * indicator {x} y)"
        by(subst nn_integral_add[symmetric])(auto simp add: nn_integral_count_space_reindex intro!: nn_integral_cong split: split_indicator)
      also have "… < (∑+ y. f (Source, Left y) * indicator (- {x}) y) + (∑+ y. pmf p y * indicator {x} y)" using * finite
        by(auto simp add:)
      also have "… ≤ (∑+ y. pmf p y * indicator (- {x}) y) + (∑+ y. pmf p y * indicator {x} y)"
        using flowD_capacity[OF f, of "(Source, Left _)"]
        by(auto intro!: nn_integral_mono split: split_indicator)
      also have "… = (∑+ y. pmf p y)"
        by(subst nn_integral_add[symmetric])(auto simp add: pmf_nonneg intro!: nn_integral_cong split: split_indicator)
      also have "… = 1" unfolding nn_integral_pmf by simp
      finally show False using val by simp
    qed
  qed

  have IN_Sink: "d_IN f Sink = 1"
  proof -
    have "d_IN f Sink = (∑+ x∈range Right. f (x, Sink))" unfolding d_IN_def
      by(auto intro!: nn_integral_cong network.flowD_outside[OF f] simp add: nn_integral_count_space_indicator split: split_indicator)
    also have "… = (∑+ y. d_OUT f (Right y))" by(simp add: nn_integral_count_space_reindex OUT_f_Right)
    also have "… = (∑+ y. d_IN f (Right y))" by(simp add: flowD_KIR[OF f])
    also have "… = (∑+ y. (∑+ x∈range Left. f (x, Right y)))"
      by(auto simp add: d_IN_def nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
    also have "… = (∑+ y. ∑+ x. f (Left x, Right y))" by(simp add: nn_integral_count_space_reindex)
    also have "… = (∑+ x. ∑+ y. f (Left x, Right y))"
      by(subst nn_integral_fst_count_space[where f="case_prod _", simplified])(simp add: nn_integral_snd_count_space[where f="case_prod _", simplified])
    also have "… = (∑+ x. (∑+ y∈range Right. f (Left x, y)))"
      by(simp add: nn_integral_count_space_reindex)
    also have "… = (∑+ x. d_OUT f (Left x))" unfolding d_OUT_def
      by(auto intro!: nn_integral_cong network.flowD_outside[OF f] simp add: nn_integral_count_space_indicator split: split_indicator)
    also have "… = (∑+ x. d_IN f (Left x))" by(simp add: flowD_KIR[OF f])
    also have "… = (∑+ x. pmf p x)" by(simp add: IN_f_Left SAT_p)
    also have "… = 1" unfolding nn_integral_pmf by simp
    finally show ?thesis .
  qed

  have SAT_q: "f (Right y, Sink) = pmf q y" for y
  proof(rule antisym)
    show "f (Right y, Sink) ≤ pmf q y" using flowD_capacity[OF f, of "(Right y, Sink)"] by simp
    show "pmf q y ≤ f (Right y, Sink)"
    proof(rule ccontr)
      assume *: "¬ ?thesis"
      have finite: "(∑+ x. f (Right x, Sink) * indicator (- {y}) x) ≠ ∞"
      proof -
        have "(∑+ x. f (Right x, Sink) * indicator (- {y}) x) ≤ (∑+ x∈range Right. f (x, Sink))"
          by(auto simp add: nn_integral_count_space_reindex intro!: nn_integral_mono split: split_indicator)
        also have "… = d_IN f Sink"
          by(auto simp add: d_IN_def nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
        finally show ?thesis using IN_Sink by (auto simp: top_unique)
      qed
      have "d_IN f Sink = (∑+ x∈range Right. f (x, Sink))"
        by(auto simp add: d_IN_def nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
      also have "… = (∑+ x. f (Right x, Sink) * indicator (- {y}) x) + (∑+ x. f (Right x, Sink) * indicator {y} x)"
        by(subst nn_integral_add[symmetric])(auto simp add: nn_integral_count_space_reindex intro!: nn_integral_cong split: split_indicator)
      also have "… < (∑+ x. f (Right x, Sink) * indicator (- {y}) x) + (∑+ x. pmf q x * indicator {y} x)" using * finite
        by auto
      also have "… ≤ (∑+ x. pmf q x * indicator (- {y}) x) + (∑+ x. pmf q x * indicator {y} x)"
        using flowD_capacity[OF f, of "(Right _, Sink)"]
        by(auto intro!: nn_integral_mono split: split_indicator)
      also have "… = (∑+ x. pmf q x)"
        by(subst nn_integral_add[symmetric])(auto simp add: pmf_nonneg intro!: nn_integral_cong split: split_indicator)
      also have "… = 1" unfolding nn_integral_pmf by simp
      finally show False using IN_Sink by simp
    qed
  qed

  let ?z = "λ(x, y). enn2real (f (Left x, Right y))"
  have nonneg: "⋀xy. 0 ≤ ?z xy" by clarsimp
  have prob: "(∑+ xy. ?z xy) = 1"
  proof -
    have "(∑+ xy. ?z xy) = (∑+ x. ∑+ y. (ennreal ∘ ?z) (x, y))"
      unfolding nn_integral_fst_count_space by(simp add: split_def o_def)
    also have "… = (∑+ x. (∑+y∈range Right. f (Left x, y)))"
      by(auto simp add: nn_integral_count_space_reindex intro!: nn_integral_cong)
    also have "… = (∑+ x. d_OUT f (Left x))"
      by(auto simp add: d_OUT_def nn_integral_count_space_indicator split: split_indicator intro!: nn_integral_cong network.flowD_outside[OF f])
    also have "… = (∑+ x. d_IN f (Left x))" using flowD_KIR[OF f] by simp
    also have "… = (∑+ x∈range Left. f (Source, x))" by(simp add: nn_integral_count_space_reindex IN_f_Left)
    also have "… = value_flow Δ f"
      by(auto simp add: d_OUT_def nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
    finally show ?thesis using val by(simp)
  qed
  note z = nonneg prob
  def z  "embed_pmf ?z"
  have z_sel [simp]: "pmf z (x, y) = enn2real (f (Left x, Right y))" for x y
    by(simp add: z_def pmf_embed_pmf[OF z])

  show ?thesis
  proof
    show "R x y" if "(x, y) ∈ set_pmf z" for x y
      using that network.flowD_outside[OF f, of "(Left x, Right y)"] unfolding set_pmf_iff
      by(auto simp add: enn2real_eq_0_iff)

    show "map_pmf fst z = p"
    proof(rule pmf_eqI)
      fix x
      have "pmf (map_pmf fst z) x = (∑+ e∈range (Pair x). pmf z e)"
        by(auto simp add: ennreal_pmf_map nn_integral_measure_pmf nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
      also have "… = (∑+ y∈range Right. f (Left x, y))" by(simp add: nn_integral_count_space_reindex)
      also have "… = d_OUT f (Left x)"
        by(auto simp add: d_OUT_def nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
      also have "… = d_IN f (Left x)" by(rule flowD_KIR[OF f]) simp_all
      also have "… = f (Source, Left x)" by(simp add: IN_f_Left)
      also have "… = pmf p x" by(simp add: SAT_p)
      finally show "pmf (map_pmf fst z) x = pmf p x" by simp
    qed

    show "map_pmf snd z = q"
    proof(rule pmf_eqI)
      fix y
      have "pmf (map_pmf snd z) y = (∑+ e∈range (λx. (x, y)). pmf z e)"
        by(auto simp add: ennreal_pmf_map nn_integral_measure_pmf nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
      also have "… = (∑+ x∈range Left. f (x, Right y))" by(simp add: nn_integral_count_space_reindex)
      also have "… = d_IN f (Right y)"
        by(auto simp add: d_IN_def nn_integral_count_space_indicator intro!: nn_integral_cong network.flowD_outside[OF f] split: split_indicator)
      also have "… = d_OUT f (Right y)" by(simp add: flowD_KIR[OF f])
      also have "… = f (Right y, Sink)" by(simp add: OUT_f_Right)
      also have "… = pmf q y" by(simp add: SAT_q)
      finally show "pmf (map_pmf snd z) y = pmf q y" by simp
    qed
  qed
qed

end

corollary rel_pmf_distr_mono: "rel_pmf R OO rel_pmf S ≤ rel_pmf (R OO S)"
-- ‹This fact has already been proven for the registration of @{typ "'a pmf"} as a BNF,
  but this proof is much shorter and more elegant. See @{cite HoelzlLochbihlerTraytel2015ITP} for a
  comparison of formalisations.›
proof(intro le_funI le_boolI rel_pmf_measureI, elim relcomppE)
  fix p q r A
  assume pq: "rel_pmf R p q" and qr: "rel_pmf S q r"
  have "measure (measure_pmf p) A ≤ measure (measure_pmf q) {y. ∃x∈A. R x y}"
    (is "_ ≤ measure _ ?B") using pq by(rule rel_pmf_measureD)
  also have "… ≤ measure (measure_pmf r) {z. ∃y∈?B. S y z}"
    using qr by(rule rel_pmf_measureD)
  also have "{z. ∃y∈?B. S y z} = {z. ∃x∈A. (R OO S) x z}" by auto
  finally show "measure (measure_pmf p) A ≤ measure (measure_pmf r) …" .
qed

end