Theory Max_Flow_Min_Cut_Countable

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

section ‹Graphs›

type_synonym 'v edge = "'v × 'v"

record 'v graph =
  edge :: "'v ⇒ 'v ⇒ bool"

abbreviation edges :: "('v, 'more) graph_scheme ⇒ 'v edge set" ("Eı")
where "EG ≡ {(x, y). edge G x y}"

definition outgoing :: "('v, 'more) graph_scheme ⇒ 'v ⇒ 'v set" ("OUTı")
where "OUTG x = {y. (x, y) ∈ EG}"

definition incoming :: "('v, 'more) graph_scheme ⇒ 'v ⇒ 'v set" ("INı")
where "ING y = {x. (x, y) ∈ EG}"

text ‹
  Vertices are implicitly defined as the endpoints of edges, so we do not allow isolated vertices.
  For the purpose of flows, this does not matter as isolated vertices cannot contribute to a flow.
  The advantage is that we do not need any invariant on graphs that the endpoints of edges are a
  subset of the vertices. Conversely, this design choice makes a few proofs about reductions on webs
  harder, because we have to adjust other sets which are supposed to be part of the vertices.
›

definition vertex :: "('v, 'more) graph_scheme ⇒ 'v ⇒ bool"
where "vertex G x ⟷ DomainP (edge G) x ∨ RangeP (edge G) x"

lemma vertexI:
  shows vertexI1: "edge Γ x y ⟹ vertex Γ x"
  and vertexI2: "edge Γ x y ⟹ vertex Γ y"
by(auto simp add: vertex_def)

abbreviation vertices :: "('v, 'more) graph_scheme ⇒ 'v set" ("Vı")
where "VG ≡ Collect (vertex G)"

lemma "V_def": "VG = fst ` EG ∪ snd ` EG"
by(auto 4 3 simp add: vertex_def intro: rev_image_eqI prod.expand)

type_synonym 'v path = "'v list"

abbreviation path :: "('v, 'more) graph_scheme ⇒ 'v ⇒ 'v path ⇒ 'v ⇒ bool"
where "path G ≡ rtrancl_path (edge G)"

inductive cycle :: "('v, 'more) graph_scheme ⇒ 'v path ⇒ bool"
  for G
where -- ‹Cycles must not pass through the same node multiple times. Otherwise, the cycle might
  enter a node via two different edges and leave it via just one edge. Thus, the clean-up lemma
  would not hold any more.
›
  cycle: "⟦ path G v p v; p ≠ []; distinct p ⟧ ⟹ cycle G p"

inductive_simps cycle_Nil [simp]: "cycle G Nil"

abbreviation cycles :: "('v, 'more) graph_scheme ⇒ 'v path set"
where "cycles G ≡ Collect (cycle G)"

lemma countable_cycles [simp]:
  assumes "countable (VG)"
  shows "countable (cycles G)"
proof -
  have "cycles G ⊆ lists VG"
    by(auto elim!: cycle.cases dest: rtrancl_path_Range_end rtrancl_path_Range simp add: vertex_def)
  thus ?thesis by(rule countable_subset)(simp add: assms)
qed

definition cycle_edges :: "'v path ⇒ 'v edge list"
where "cycle_edges p = zip p (rotate1 p)"

lemma cycle_edges_not_Nil: "cycle G p ⟹ cycle_edges p ≠ []"
by(auto simp add: cycle_edges_def cycle.simps neq_Nil_conv zip_Cons1 split: list.split)

lemma distinct_cycle_edges:
  "cycle G p ⟹ distinct (cycle_edges p)"
by(erule cycle.cases)(simp add: cycle_edges_def distinct_zipI2)

lemma cycle_enter_leave_same:
  assumes "cycle G p"
  shows "card (set [(x', y) ← cycle_edges p. x' = x]) = card (set [(x', y) ← cycle_edges p. y = x])"
  (is "?lhs = ?rhs")
using assms
proof cases
  case (cycle v)
  from distinct_cycle_edges[OF assms]
  have "?lhs = length [x' ← map fst (cycle_edges p). x' = x]"
    by(subst distinct_card; simp add: filter_map o_def split_def)
  also have "… = (if x ∈ set p then 1 else 0)" using cycle
    by(auto simp add: cycle_edges_def filter_empty_conv length_filter_conv_card card_eq_1_iff in_set_conv_nth dest: nth_eq_iff_index_eq)
  also have "… = length [y ← map snd (cycle_edges p). y = x]" using cycle
    apply(auto simp add: cycle_edges_def filter_empty_conv Suc_length_conv intro!: exI[where x=x])
    apply(drule split_list_first)
    apply(auto dest: split_list_first simp add: append_eq_Cons_conv rotate1_append filter_empty_conv split: if_split_asm dest: in_set_tlD)
    done
  also have "… = ?rhs" using distinct_cycle_edges[OF assms]
    by(subst distinct_card; simp add: filter_map o_def split_def)
  finally show ?thesis .
qed

lemma cycle_leave_ex_enter:
  assumes "cycle G p" and "(x, y) ∈ set (cycle_edges p)"
  shows "∃z. (z, x) ∈ set (cycle_edges p)"
using assms
by(cases)(auto 4 3 simp add: cycle_edges_def cong: conj_cong split: if_split_asm intro: set_zip_rightI dest: set_zip_leftD)

lemma cycle_edges_edges:
  assumes "cycle G p"
  shows "set (cycle_edges p) ⊆ EG"
proof
  fix x
  assume "x ∈ set (cycle_edges p)"
  then obtain i where x: "x = (p ! i, rotate1 p ! i)" and i: "i < length p"
    by(auto simp add: cycle_edges_def set_zip)
  from assms obtain v where p: "path G v p v" and "p ≠ []" and "distinct p" by cases
  let ?i = "Suc i mod length p"
  have "?i < length p" by (simp add: `p ≠ []`)
  note rtrancl_path_nth[OF p this]
  also have "(v # p) ! ?i = p ! i"
  proof(cases "Suc i < length p")
    case True thus ?thesis by simp
  next
    case False
    with i have "Suc i = length p" by simp
    moreover from p ‹p ≠ []› have "last p = v" by(rule rtrancl_path_last)
    ultimately show ?thesis using ‹p ≠ []› by(simp add: last_conv_nth)(metis diff_Suc_Suc diff_zero)
  qed
  also have "p ! ?i = rotate1 p ! i" using i by(simp add: nth_rotate1)
  finally show "x ∈ EG" by(simp add: x)
qed


section ‹Network and Flow›

record 'v network = "'v graph" +
  capacity :: "'v edge ⇒ ennreal"
  source :: "'v"
  sink :: "'v"

type_synonym 'v flow = "'v edge ⇒ ennreal"

inductive_set support_flow :: "'v flow ⇒ 'v edge set"
  for f
where "f e > 0 ⟹ e ∈ support_flow f"

lemma support_flow_conv: "support_flow f = {e. f e > 0}"
by(auto simp add: support_flow.simps)

lemma not_in_support_flowD: "x ∉ support_flow f ⟹ f x = 0"
by(simp add: support_flow_conv)

definition d_OUT :: "'v flow ⇒ 'v ⇒ ennreal"
where "d_OUT g x = (∑+ y. g (x, y))"

definition d_IN :: "'v flow ⇒ 'v ⇒ ennreal"
where "d_IN g y = (∑+ x. g (x, y))"

lemma d_OUT_mono: "(⋀y. f (x, y) ≤ g (x, y)) ⟹ d_OUT f x ≤ d_OUT g x"
by(auto simp add: d_OUT_def le_fun_def intro: nn_integral_mono)

lemma d_IN_mono: "(⋀x. f (x, y) ≤ g (x, y)) ⟹ d_IN f y ≤ d_IN g y"
by(auto simp add: d_IN_def le_fun_def intro: nn_integral_mono)

lemma d_OUT_0 [simp]: "d_OUT (λ_. 0) x = 0"
by(simp add: d_OUT_def)

lemma d_IN_0 [simp]: "d_IN (λ_. 0) x = 0"
by(simp add: d_IN_def)

lemma d_OUT_add: "d_OUT (λe. f e + g e) x = d_OUT f x + d_OUT g x"
unfolding d_OUT_def by(simp add: nn_integral_add)

lemma d_IN_add: "d_IN (λe. f e + g e) x = d_IN f x + d_IN g x"
unfolding d_IN_def by(simp add: nn_integral_add)

lemma d_OUT_cmult: "d_OUT (λe. c * f e) x = c * d_OUT f x"
by(simp add: d_OUT_def nn_integral_cmult)

lemma d_IN_cmult: "d_IN (λe. c * f e) x = c * d_IN f x"
by(simp add: d_IN_def nn_integral_cmult)

lemma d_OUT_ge_point: "f (x, y) ≤ d_OUT f x"
by(auto simp add: d_OUT_def intro!: nn_integral_ge_point)

lemma d_IN_ge_point: "f (y, x) ≤ d_IN f x"
by(auto simp add: d_IN_def intro!: nn_integral_ge_point)

lemma d_OUT_monotone_convergence_SUP:
  assumes "incseq (λn y. f n (x, y))"
  shows "d_OUT (λe. SUP n :: nat. f n e) x = (SUP n. d_OUT (f n) x)"
unfolding d_OUT_def by(rule nn_integral_monotone_convergence_SUP[OF assms]) simp

lemma d_IN_monotone_convergence_SUP:
  assumes "incseq (λn x. f n (x, y))"
  shows "d_IN (λe. SUP n :: nat. f n e) y = (SUP n. d_IN (f n) y)"
unfolding d_IN_def by(rule nn_integral_monotone_convergence_SUP[OF assms]) simp

lemma d_OUT_diff:
  assumes "⋀y. g (x, y) ≤ f (x, y)" "d_OUT g x ≠ ⊤"
  shows "d_OUT (λe. f e - g e) x = d_OUT f x - d_OUT g x"
using assms by(simp add: nn_integral_diff d_OUT_def)

lemma d_IN_diff:
  assumes "⋀x. g (x, y) ≤ f (x, y)" "d_IN g y ≠ ⊤"
  shows "d_IN (λe. f e - g e) y = d_IN f y - d_IN g y"
using assms by(simp add: nn_integral_diff d_IN_def)

lemma fixes G (structure)
  shows d_OUT_alt_def: "(⋀y. (x, y) ∉ E ⟹ g (x, y) = 0) ⟹ d_OUT g x = (∑+  y∈OUT x. g (x, y))"
  and d_IN_alt_def: "(⋀x. (x, y) ∉ E ⟹ g (x, y) = 0) ⟹ d_IN g y = (∑+ x∈IN y. g (x, y))"
unfolding d_OUT_def d_IN_def
by(fastforce simp add: max_def d_OUT_def d_IN_def nn_integral_count_space_indicator outgoing_def incoming_def intro!: nn_integral_cong split: split_indicator)+

lemma d_OUT_alt_def2: "d_OUT g x = (∑+ y∈{y. (x, y) ∈ support_flow g}. g (x, y))"
  and d_IN_alt_def2: "d_IN g y = (∑+ x∈{x. (x, y) ∈ support_flow g}. g (x, y))"
unfolding d_OUT_def d_IN_def
by(auto simp add: max_def d_OUT_def d_IN_def nn_integral_count_space_indicator outgoing_def incoming_def support_flow.simps intro!: nn_integral_cong split: split_indicator)+

definition d_diff :: "('v edge ⇒ ennreal) ⇒ 'v ⇒ ennreal"
where "d_diff g x = d_OUT g x - d_IN g x"

abbreviation KIR :: "('v edge ⇒ ennreal) ⇒ 'v ⇒ bool"
where "KIR f x ≡ d_OUT f x = d_IN f x"

inductive_set SINK :: "('v edge ⇒ ennreal) ⇒ 'v set"
  for f
where SINK: "d_OUT f x = 0 ⟹ x ∈ SINK f"

lemma SINK_mono:
  assumes "⋀e. f e ≤ g e"
  shows "SINK g ⊆ SINK f"
proof(rule subsetI; erule SINK.cases; hypsubst)
  fix x
  assume "d_OUT g x = 0"
  moreover have "d_OUT f x ≤ d_OUT g x" using assms by(rule d_OUT_mono)
  ultimately have "d_OUT f x = 0" by simp
  thus "x ∈ SINK f" ..
qed

lemma SINK_mono': "f ≤ g ⟹ SINK g ⊆ SINK f"
by(rule SINK_mono)(rule le_funD)

lemma support_flow_Sup: "support_flow (Sup Y) = (⋃f∈Y. support_flow f)"
by(auto simp add: support_flow_conv less_SUP_iff)

lemma
  assumes chain: "Complete_Partial_Order.chain op ≤ Y"
  and Y: "Y ≠ {}"
  and countable: "countable (support_flow (Sup Y))"
  shows d_OUT_Sup: "d_OUT (Sup Y) x = (SUP f:Y. d_OUT f x)" (is "?OUT x" is "?lhs1 x = ?rhs1 x")
  and d_IN_Sup: "d_IN (Sup Y) y = (SUP f:Y. d_IN f y)" (is "?IN" is "?lhs2 = ?rhs2")
  and SINK_Sup: "SINK (Sup Y) = (⋂f∈Y. SINK f)" (is "?SINK")
proof -
  have chain': "Complete_Partial_Order.chain op ≤ ((λf y. f (x, y)) ` Y)" for x using chain
    by(rule chain_imageI)(simp add: le_fun_def)
  have countable': "countable {y. (x, y) ∈ support_flow (Sup Y)}" for x
    using _ countable[THEN countable_image[where f=snd]]
    by(rule countable_subset)(auto intro: prod.expand rev_image_eqI)
  { fix x
    have "?lhs1 x = (∑+ y∈{y. (x, y) ∈ support_flow (Sup Y)}. SUP f:Y. f (x, y))"
      by(subst d_OUT_alt_def2; simp)
    also have "… = (SUP f:Y. ∑+ y∈{y. (x, y) ∈ support_flow (Sup Y)}. f (x, y))" using Y
      by(rule nn_integral_monotone_convergence_SUP_countable)(auto simp add: chain' intro: countable')
    also have "… = ?rhs1 x" unfolding d_OUT_alt_def2
      by(auto 4 3 simp add: support_flow_Sup max_def nn_integral_count_space_indicator intro!: nn_integral_cong SUP_cong split: split_indicator dest: not_in_support_flowD)
    finally show "?OUT x" . }
  note out = this

  have chain'': "Complete_Partial_Order.chain op ≤ ((λf x. f (x, y)) ` Y)" for y using chain
    by(rule chain_imageI)(simp add: le_fun_def)
  have countable'': "countable {x. (x, y) ∈ support_flow (Sup Y)}" for y
    using _ countable[THEN countable_image[where f=fst]]
    by(rule countable_subset)(auto intro: prod.expand rev_image_eqI)
  have "?lhs2 = (∑+ x∈{x. (x, y) ∈ support_flow (Sup Y)}. SUP f:Y. f (x, y))"
    by(subst d_IN_alt_def2; simp)
  also have "… = (SUP f:Y. ∑+ x∈{x. (x, y) ∈ support_flow (Sup Y)}. f (x, y))" using Y
    by(rule nn_integral_monotone_convergence_SUP_countable)(simp_all add: chain'' countable'')
  also have "… = ?rhs2" unfolding d_IN_alt_def2
    by(auto 4 3 simp add: support_flow_Sup max_def nn_integral_count_space_indicator intro!: nn_integral_cong SUP_cong split: split_indicator dest: not_in_support_flowD)
  finally show ?IN .

  show ?SINK by(rule set_eqI)(simp add: SINK.simps out Y bot_ennreal[symmetric])
qed

lemma
  assumes chain: "Complete_Partial_Order.chain op ≤ Y"
  and Y: "Y ≠ {}"
  and countable: "countable (support_flow f)"
  and bounded: "⋀g e. g ∈ Y ⟹ g e ≤ f e"
  shows d_OUT_Inf: "d_OUT f x ≠ top ⟹ d_OUT (Inf Y) x = (INF g:Y. d_OUT g x)" (is "_ ⟹ ?OUT" is "_ ⟹ ?lhs1 = ?rhs1")
  and d_IN_Inf: "d_IN f x ≠ top ⟹ d_IN (Inf Y) x = (INF g:Y. d_IN g x)" (is "_ ⟹ ?IN" is "_ ⟹ ?lhs2 = ?rhs2")
proof -
  text ‹We take a detour here via suprema because we have more theorems about @{const nn_integral}
    with suprema than with infinma.›

  from Y obtain g0 where g0: "g0 ∈ Y" by auto
  have g0_le_f: "g0 e ≤ f e" for e by(rule bounded[OF g0])

  have "support_flow (SUP g:Y. (λe. f e - g e)) ⊆ support_flow f"
    by(clarsimp simp add: support_flow.simps less_SUP_iff elim!: less_le_trans intro!: diff_le_self_ennreal)
  then have countable': "countable (support_flow (SUP g:Y. (λe. f e - g e)))" by(rule countable_subset)(rule countable)

  have "Complete_Partial_Order.chain op ≥ Y" using chain by(simp add: chain_dual)
  hence chain': "Complete_Partial_Order.chain op ≤ ((λg e. f e - g e) ` Y)"
    by(rule chain_imageI)(auto simp add: le_fun_def intro: ennreal_minus_mono)

  { assume finite: "d_OUT f x ≠ top"
    have finite' [simp]: "f (x, y) ≠ ⊤" for y using finite
      by(rule neq_top_trans) (rule d_OUT_ge_point)

    have finite'_g: "g (x, y) ≠ ⊤" if "g ∈ Y" for g y using finite'[of y]
      by(rule neq_top_trans)(rule bounded[OF that])

    have finite1: "(∑+ y. f (x, y) - (INF g:Y. g (x, y))) ≠ top"
      using finite by(rule neq_top_trans)(auto simp add: d_OUT_def intro!: nn_integral_mono)
    have finite2: "d_OUT g x ≠ top" if "g ∈ Y" for g using finite
      by(rule neq_top_trans)(auto intro: d_OUT_mono bounded[OF that])

    have bounded1: "(⨅g∈Y. d_OUT g x) ≤ d_OUT f x"
      using Y by (blast intro: INF_lower2 d_OUT_mono bounded)

    have "?lhs1 = (∑+ y. INF g:Y. g (x, y))" by(simp add: d_OUT_def)
    also have "… = d_OUT f x - (∑+ y. f (x, y) - (INF g:Y. g (x, y)))" unfolding d_OUT_def
      using finite1 g0_le_f
      apply(subst nn_integral_diff[symmetric])
      apply(auto simp add: AE_count_space intro!: diff_le_self_ennreal INF_lower2[OF g0] nn_integral_cong diff_diff_ennreal[symmetric])
      done
    also have "(∑+ y. f (x, y) - (INF g:Y. g (x, y))) = d_OUT (λe. SUP g:Y. f e - g e) x"
      unfolding d_OUT_def by(subst SUP_const_minus_ennreal)(simp_all add: Y)
    also have "… = (SUP h:(λg e. f e - g e) ` Y. d_OUT h x)" using countable' chain' Y
      by(subst d_OUT_Sup[symmetric])(simp_all add: SUP_apply[abs_def])
    also have "… = (SUP g:Y. d_OUT (λe. f e - g e) x)" unfolding image_image ..
    also have "… = (SUP g:Y. d_OUT f x - d_OUT g x)"
      by(rule SUP_cong[OF refl] d_OUT_diff)+(auto intro: bounded simp add: finite2)
    also have "… = d_OUT f x - ?rhs1" by(subst SUP_const_minus_ennreal)(simp_all add: Y)
    also have "d_OUT f x - … = ?rhs1"
      using Y by(subst diff_diff_ennreal)(simp_all add: bounded1 finite)
    finally show ?OUT .
  next
    assume finite: "d_IN f x ≠ top"
    have finite' [simp]: "f (y, x) ≠ ⊤" for y using finite
      by(rule neq_top_trans) (rule d_IN_ge_point)

    have finite'_g: "g (y, x) ≠ ⊤" if "g ∈ Y" for g y using finite'[of y]
      by(rule neq_top_trans)(rule bounded[OF that])

    have finite1: "(∑+ y. f (y, x) - (INF g:Y. g (y, x))) ≠ top"
      using finite by(rule neq_top_trans)(auto simp add: d_IN_def diff_le_self_ennreal intro!: nn_integral_mono)
    have finite2: "d_IN g x ≠ top" if "g ∈ Y" for g using finite
      by(rule neq_top_trans)(auto intro: d_IN_mono bounded[OF that])

    have bounded1: "(⨅g∈Y. d_IN g x) ≤ d_IN f x"
      using Y by (blast intro: INF_lower2 d_IN_mono bounded)

    have "?lhs2 = (∑+ y. INF g:Y. g (y, x))" by(simp add: d_IN_def)
    also have "… = d_IN f x - (∑+ y. f (y, x) - (INF g:Y. g (y, x)))" unfolding d_IN_def
      using finite1 g0_le_f
      apply(subst nn_integral_diff[symmetric])
      apply(auto simp add: AE_count_space intro!: diff_le_self_ennreal INF_lower2[OF g0] nn_integral_cong diff_diff_ennreal[symmetric])
      done
    also have "(∑+ y. f (y, x) - (INF g:Y. g (y, x))) = d_IN (λe. SUP g:Y. f e - g e) x"
      unfolding d_IN_def by(subst SUP_const_minus_ennreal)(simp_all add: Y)
    also have "… = (SUP h:(λg e. f e - g e) ` Y. d_IN h x)" using countable' chain' Y
      by(subst d_IN_Sup[symmetric])(simp_all add: SUP_apply[abs_def])
    also have "… = (SUP g:Y. d_IN (λe. f e - g e) x)" unfolding image_image ..
    also have "… = (SUP g:Y. d_IN f x - d_IN g x)"
      by(rule SUP_cong[OF refl] d_IN_diff)+(auto intro: bounded simp add: finite2)
    also have "… = d_IN f x - ?rhs2" by(subst SUP_const_minus_ennreal)(simp_all add: Y)
    also have "d_IN f x - … = ?rhs2"
      by(subst diff_diff_ennreal)(simp_all add: finite bounded1)
    finally show ?IN . }
qed

inductive flow :: "('v, 'more) network_scheme ⇒ 'v flow ⇒ bool"
  for Δ (structure) and f
where
  flow: "⟦ ⋀e. f e ≤ capacity Δ e;
     ⋀x. ⟦ x ≠ source Δ; x ≠ sink Δ ⟧ ⟹ KIR f x ⟧
  ⟹ flow Δ f"

lemma flowD_capacity: "flow Δ f ⟹ f e ≤ capacity Δ e"
by(cases e)(simp add: flow.simps)

lemma flowD_KIR: "⟦ flow Δ f; x ≠ source Δ; x ≠ sink Δ ⟧ ⟹ KIR f x"
by(simp add: flow.simps)

lemma flowD_capacity_OUT: "flow Δ f ⟹ d_OUT f x ≤ d_OUT (capacity Δ) x"
by(rule d_OUT_mono)(erule flowD_capacity)

lemma flowD_capacity_IN: "flow Δ f ⟹ d_IN f x ≤ d_IN (capacity Δ) x"
by(rule d_IN_mono)(erule flowD_capacity)

abbreviation value_flow :: "('v, 'more) network_scheme ⇒ ('v edge ⇒ ennreal) ⇒ ennreal"
where "value_flow Δ f ≡ d_OUT f (source Δ)"

subsection ‹Cut›

type_synonym 'v cut = "'v set"

inductive cut :: "('v, 'more) network_scheme ⇒ 'v cut ⇒ bool"
  for Δ and S
where cut: "⟦ source Δ ∈ S; sink Δ ∉ S ⟧ ⟹ cut Δ S"

inductive orthogonal :: "('v, 'more) network_scheme ⇒ 'v flow ⇒ 'v cut ⇒ bool"
  for Δ f S
where
  "⟦ ⋀x y. ⟦ edge Δ x y; x ∈ S; y ∉ S ⟧ ⟹ f (x, y) = capacity Δ (x, y);
     ⋀x y. ⟦ edge Δ x y; x ∉ S; y ∈ S ⟧ ⟹ f (x, y) = 0 ⟧
  ⟹ orthogonal Δ f S"

lemma orthogonalD_out:
  "⟦ orthogonal Δ f S; edge Δ x y; x ∈ S; y ∉ S ⟧ ⟹ f (x, y) = capacity Δ (x, y)"
by(simp add: orthogonal.simps)

lemma orthogonalD_in:
  "⟦ orthogonal Δ f S; edge Δ x y; x ∉ S; y ∈ S ⟧ ⟹ f (x, y) = 0"
by(simp add: orthogonal.simps)



section ‹Countable network›

locale countable_network =
  fixes Δ :: "('v, 'more) network_scheme" (structure)
  assumes countable_E [simp]: "countable E"
  and source_neq_sink [simp]: "source Δ ≠ sink Δ"
  and capacity_outside: "e ∉ E ⟹ capacity Δ e = 0"
  and capacity_finite [simp]: "capacity Δ e ≠ ⊤"
begin

lemma sink_neq_source [simp]: "sink Δ ≠ source Δ"
using source_neq_sink[symmetric] .

lemma countable_V [simp]: "countable V"
unfolding "V_def" using countable_E by auto

lemma flowD_outside:
  assumes g: "flow Δ g"
  shows "e ∉ E ⟹ g e = 0"
using flowD_capacity[OF g, of e] capacity_outside[of e] by simp

lemma flowD_finite:
  assumes "flow Δ g"
  shows "g e ≠ ⊤"
using flowD_capacity[OF assms, of e] by (auto simp: top_unique)

lemma zero_flow [simp]: "flow Δ (λ_. 0)"
by(rule flow.intros) simp_all

end

section ‹Attainability of flows in networks›

subsection ‹Cleaning up flows›

text ‹If there is a flow along antiparallel edges, it suffices to consider the difference.›

definition cleanup :: "'a flow ⇒ 'a flow"
where "cleanup f = (λ(a, b). if f (a, b) > f (b, a) then f (a, b) - f (b, a) else 0)"

lemma cleanup_simps [simp]:
  "cleanup f (a, b) = (if f (a, b) > f (b, a) then f (a, b) - f (b, a) else 0)"
by(simp add: cleanup_def)

lemma value_flow_cleanup:
  assumes [simp]: "⋀x. f (x, source Δ) = 0"
  shows "value_flow Δ (cleanup f) = value_flow Δ f"
unfolding d_OUT_def
by (auto simp add: not_less intro!: nn_integral_cong intro: antisym)

lemma KIR_cleanup:
  assumes KIR: "KIR f x"
  and finite_IN: "d_IN f x ≠ ⊤"
  shows "KIR (cleanup f) x"
proof -
  from finite_IN KIR have finite_OUT: "d_OUT f x ≠ ⊤" by simp

  have finite_IN: "(∑+ y∈A. f (y, x)) ≠ ⊤" for A
    using finite_IN by(rule neq_top_trans)(auto simp add: d_IN_def nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator)
  have finite_OUT: "(∑+ y∈A. f (x, y)) ≠ ⊤" for A
    using finite_OUT by(rule neq_top_trans)(auto simp add: d_OUT_def nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator)
  have finite_in: "f (x, y) ≠ ⊤" for y using ‹d_OUT f x ≠ ⊤›
    by(rule neq_top_trans) (rule d_OUT_ge_point)

  let ?M = "{y. f (x, y) > f (y, x)}"

  have "d_OUT (cleanup f) x = (∑+ y∈?M. f (x, y) - f (y, x))"
    by(auto simp add: d_OUT_def nn_integral_count_space_indicator intro!: nn_integral_cong)
  also have "… = (∑+ y∈?M. f (x, y)) - (∑+ y∈?M. f (y, x))" using finite_IN
    by(subst nn_integral_diff)(auto simp add: AE_count_space)
  also have "… = (d_OUT f x - (∑+ y∈{y. f (x, y) ≤ f (y, x)}. f (x, y))) - (∑+ y∈?M. f (y, x))"
    unfolding d_OUT_def d_IN_def using finite_IN finite_OUT
    apply(simp add: nn_integral_count_space_indicator)
    apply(subst (2) nn_integral_diff[symmetric])
    apply(auto simp add: AE_count_space finite_in split: split_indicator intro!: arg_cong2[where f="op -"] intro!: nn_integral_cong)
    done
  also have "… = (d_IN f x - (∑+ y∈?M. f (y, x))) - (∑+ y∈{y. f (x, y) ≤ f (y, x)}. f (x, y))"
    using KIR by(simp add: diff_diff_commute_ennreal)
  also have "… = (∑+ y∈{y. f (x, y) ≤ f (y, x)}. f (y, x)) - (∑+ y∈{y. f (x, y) ≤ f (y, x)}. f (x, y))"
    using finite_IN finite_IN[of "{ _ }"]
    apply(simp add: d_IN_def nn_integral_count_space_indicator)
    apply(subst nn_integral_diff[symmetric])
    apply(auto simp add: d_IN_def AE_count_space split: split_indicator intro!: arg_cong2[where f="op -"] intro!: nn_integral_cong)
    done
  also have "… = (∑+ y∈{y. f (x, y) ≤ f (y, x)}. f (y, x) - f (x, y))" using finite_OUT
    by(subst nn_integral_diff)(auto simp add: AE_count_space)
  also have "… = d_IN (cleanup f) x" using finite_in
    by(auto simp add: d_IN_def nn_integral_count_space_indicator intro!: ennreal_diff_self nn_integral_cong split: split_indicator)
  finally show "KIR (cleanup f) x" .
qed

locale flow_attainability = countable_network Δ
  for Δ :: "('v, 'more) network_scheme" (structure)
  +
  assumes finite_capacity: "⋀x. x ≠ sink Δ ⟹ d_IN (capacity Δ) x ≠ ⊤ ∨ d_OUT (capacity Δ) x ≠ ⊤"
  and no_loop: "⋀x. ¬ edge Δ x x"
  and source_in: "⋀x. ¬ edge Δ x (source Δ)"
begin

lemma source_in_not_cycle:
  assumes "cycle Δ p"
  shows "(x, source Δ) ∉ set (cycle_edges p)"
using cycle_edges_edges[OF assms] source_in[of x] by(auto)

lemma source_out_not_cycle:
  "cycle Δ p ⟹ (source Δ, x) ∉ set (cycle_edges p)"
by(auto dest: cycle_leave_ex_enter source_in_not_cycle)

lemma flowD_source_IN:
  assumes "flow Δ f"
  shows "d_IN f (source Δ) = 0"
proof -
  have "d_IN f (source Δ) = (∑+ y∈IN (source Δ). f (y, source Δ))"
    by(rule d_IN_alt_def)(simp add: flowD_outside[OF assms])
  also have "… = (∑+ y∈IN (source Δ). 0)"
    by(rule nn_integral_cong)(simp add: source_in incoming_def)
  finally show ?thesis by simp
qed

lemma flowD_finite_IN:
  assumes f: "flow Δ f" and x: "x ≠ sink Δ"
  shows "d_IN f x ≠ top"
proof(cases "x = source Δ")
  case True thus ?thesis by(simp add: flowD_source_IN[OF f])
next
  case False
  from finite_capacity[OF x] show ?thesis
  proof
    assume *: "d_IN (capacity Δ) x ≠ ⊤"
    from flowD_capacity[OF f] have "d_IN f x ≤ d_IN (capacity Δ) x" by(rule d_IN_mono)
    also have "… < ⊤" using * by (simp add: less_top)
    finally show ?thesis by simp
  next
    assume *: "d_OUT (capacity Δ) x ≠ ⊤"
    have "d_IN f x = d_OUT f x" using flowD_KIR[OF f False x] by simp
    also have "… ≤ d_OUT (capacity Δ) x" using flowD_capacity[OF f] by(rule d_OUT_mono)
    also have "… < ⊤" using * by (simp add: less_top)
    finally show ?thesis by simp
  qed
qed

lemma flowD_finite_OUT:
  assumes "flow Δ f" "x ≠ source Δ" "x ≠ sink Δ"
  shows "d_OUT f x ≠ ⊤"
using flowD_KIR[OF assms] assms by(simp add: flowD_finite_IN)

end

locale flow_network = flow_attainability
  +
  fixes g :: "'v flow"
  assumes g: "flow Δ g"
  and g_finite: "value_flow Δ g ≠ ⊤"
  and nontrivial: "V - {source Δ, sink Δ} ≠ {}"
begin

lemma g_outside: "e ∉ E ⟹ g e = 0"
by(rule flowD_outside)(rule g)

lemma g_loop [simp]: "g (x, x) = 0"
by(rule g_outside)(simp add: no_loop)

lemma finite_IN_g: "x ≠ sink Δ ⟹ d_IN g x ≠ top"
by(rule flowD_finite_IN[OF g])

lemma finite_OUT_g:
  assumes "x ≠ sink Δ"
  shows "d_OUT g x ≠ top"
proof(cases "x = source Δ")
  case True
  with g_finite show ?thesis by simp
next
  case False
  with g have "KIR g x" using assms by(auto dest: flowD_KIR)
  with finite_IN_g[of x] False assms show ?thesis by(simp)
qed

lemma g_source_in [simp]: "g (x, source Δ) = 0"
by(rule g_outside)(simp add: source_in)

lemma finite_g [simp]: "g e ≠ top"
  by(rule flowD_finite[OF g])


definition enum_v :: "nat ⇒ 'v"
where "enum_v n = from_nat_into (V - {source Δ, sink Δ}) (fst (prod_decode n))"

lemma range_enum_v: "range enum_v ⊆ V - {source Δ, sink Δ}"
using from_nat_into[OF nontrivial] by(auto simp add: enum_v_def)

lemma enum_v_repeat:
  assumes x: "x ∈ V" "x ≠ source Δ" "x ≠ sink Δ"
  shows "∃i'>i. enum_v i' = x"
proof -
  let ?V = "V - {source Δ, sink Δ}"
  let ?n = "to_nat_on ?V x"
  let ?A = "{?n} × (UNIV :: nat set)"
  from x have x': "x ∈ V - {source Δ, sink Δ}" by simp
  have "infinite ?A" by(auto dest: finite_cartesian_productD2)
  hence "infinite (prod_encode ` ?A)" by(auto dest: finite_imageD simp add: inj_prod_encode)
  then obtain i' where "i' > i" "i' ∈ prod_encode ` ?A"
    unfolding infinite_nat_iff_unbounded by blast
  from this(2) have "enum_v i' = x" using x by(clarsimp simp add: enum_v_def)
  with ‹i' > i› show ?thesis by blast
qed

fun h_plus :: "nat ⇒ 'v edge ⇒ ennreal"
where
  "h_plus 0 (x, y) = (if x = source Δ then g (x, y) else 0)"
| "h_plus (Suc i) (x, y) =
  (if enum_v (Suc i) = x ∧ d_OUT (h_plus i) x < d_IN (h_plus i) x then
   let total = d_IN (h_plus i) x - d_OUT (h_plus i) x;
       share = g (x, y) - h_plus i (x, y);
       shares = d_OUT g x - d_OUT (h_plus i) x
   in h_plus i (x, y) + share * total / shares
   else h_plus i (x, y))"


lemma h_plus_le_g: "h_plus i e ≤ g e"
proof(induction i arbitrary: e and e)
  case 0 thus ?case by(cases e) simp
next
  case (Suc i)
  { fix x y
    assume enum: "x = enum_v (Suc i)"
    assume less: "d_OUT (h_plus i) x < d_IN (h_plus i) x"
    from enum have x: "x ≠ source Δ" "x ≠ sink Δ" using range_enum_v
      by(auto dest: sym intro: rev_image_eqI)

    def share  "g (x, y) - h_plus i (x, y)"
    def shares  "d_OUT g x - d_OUT (h_plus i) x"
    def total  "d_IN (h_plus i) x - d_OUT (h_plus i) x"
    let ?h = "h_plus i (x, y) + share * total / shares"

    have "d_OUT (h_plus i) x ≤ d_OUT g x" by(rule d_OUT_mono)(rule Suc.IH)
    also have "… < top" using finite_OUT_g[of x] x by (simp add: less_top)
    finally have "d_OUT (h_plus i) x ≠ ⊤" by simp
    then have shares_eq: "shares = (∑+ y. g (x, y) - h_plus i (x, y))" unfolding shares_def d_OUT_def
      by(subst nn_integral_diff)(simp_all add: AE_count_space Suc.IH)

    have *: "share / shares ≤ 1"
    proof (cases "share = 0")
      case True thus ?thesis by(simp)
    next
      case False
      hence "share > 0" using `h_plus i (x, y) ≤ g _`
        by(simp add: share_def dual_order.strict_iff_order)
      moreover have "share ≤ shares" unfolding share_def shares_eq by(rule nn_integral_ge_point)simp
      ultimately show ?thesis by(simp add: divide_le_posI_ennreal)
    qed

    note shares_def[THEN meta_eq_to_obj_eq]
    also have "d_OUT g x = d_IN g x" by(rule flowD_KIR[OF g x])
    also have "d_IN (h_plus i) x ≤ d_IN g x" by(rule d_IN_mono)(rule Suc.IH)
    ultimately have *: "total ≤ shares" unfolding total_def by(simp add: ennreal_minus_mono)
    moreover have "total > 0" unfolding total_def using less by (clarsimp simp add: diff_gr0_ennreal)
    ultimately have "total / shares ≤ 1" by(intro divide_le_posI_ennreal)(simp_all)
    hence "share * (total / shares) ≤ share * 1"
      by(rule mult_left_mono) simp
    hence "?h ≤ h_plus i (x, y) + share" by(simp add: ennreal_times_divide add_mono)
    also have "… = g (x, y)" unfolding share_def using ‹h_plus i (x, y) ≤ g _› finite_g[of "(x, y)"]
      by simp
    moreover
    note calculation }
  note * = this
  show ?case using Suc.IH * by(cases e) clarsimp
qed

lemma h_plus_outside: "e ∉ E ⟹ h_plus i e = 0"
by (metis g_outside h_plus_le_g le_zero_eq)

lemma h_plus_not_infty [simp]: "h_plus i e ≠ top"
using h_plus_le_g[of i e] by (auto simp: top_unique)

lemma h_plus_mono: "h_plus i e ≤ h_plus (Suc i) e"
proof(cases e)
  case [simp]: (Pair x y)
  { assume "d_OUT (h_plus i) x < d_IN (h_plus i) x"
    hence "h_plus i (x, y) + 0 ≤ h_plus i (x, y) + (g (x, y) - h_plus i (x, y)) * (d_IN (h_plus i) x - d_OUT (h_plus i) x) / (d_OUT g x - d_OUT (h_plus i) x)"
      by(intro add_left_mono d_OUT_mono le_funI) (simp_all add: h_plus_le_g) }
  then show ?thesis by clarsimp
qed

lemma h_plus_mono': "i ≤ j ⟹ h_plus i e ≤ h_plus j e"
by(induction rule: dec_induct)(auto intro: h_plus_mono order_trans)

lemma d_OUT_h_plus_not_infty': "x ≠ sink Δ ⟹ d_OUT (h_plus i) x ≠ top"
using d_OUT_mono[of "h_plus i" x g, OF h_plus_le_g] finite_OUT_g[of x] by (auto simp: top_unique)

lemma h_plus_OUT_le_IN:
  assumes "x ≠ source Δ"
  shows "d_OUT (h_plus i) x ≤ d_IN (h_plus i) x"
proof(induction i)
  case 0
  thus ?case using assms by(simp add: d_OUT_def)
next
  case (Suc i)
  have "d_OUT (h_plus (Suc i)) x ≤ d_IN (h_plus i) x"
  proof(cases "enum_v (Suc i) = x ∧ d_OUT (h_plus i) x < d_IN (h_plus i) x")
    case False
    thus ?thesis using Suc.IH by(simp add: d_OUT_def cong: conj_cong)
  next
    case True
    hence x: "x ≠ sink Δ" and le: "d_OUT (h_plus i) x < d_IN (h_plus i) x" using range_enum_v by auto
    let ?r = "λy. (g (x, y) - h_plus i (x, y)) * (d_IN (h_plus i) x - d_OUT (h_plus i) x) / (d_OUT g x - d_OUT (h_plus i) x)"
    have "d_OUT (h_plus (Suc i)) x = d_OUT (h_plus i) x + (∑+ y. ?r y)"
      using True unfolding d_OUT_def h_plus.simps by(simp add: AE_count_space nn_integral_add)
    also from True have "x ≠ source Δ" "x ≠ sink Δ" using range_enum_v by auto
    from flowD_KIR[OF g this] le d_IN_mono[of "h_plus i" x g, OF h_plus_le_g]
    have le': "d_OUT (h_plus i) x < d_OUT g x" by(simp)
    then have "(∑+ y. ?r y) =
      (d_IN (h_plus i) x - d_OUT (h_plus i) x) * ((∑+ y. g (x, y) - h_plus i (x, y)) / (d_OUT g x - d_OUT (h_plus i) x))"
      by(subst mult.commute, subst ennreal_times_divide[symmetric])
        (simp add: nn_integral_cmult nn_integral_divide Suc.IH diff_gr0_ennreal)
    also have "(∑+ y. g (x, y) - h_plus i (x, y)) = d_OUT g x - d_OUT (h_plus i) x" using x
      by(subst nn_integral_diff)(simp_all add: d_OUT_def[symmetric] h_plus_le_g d_OUT_h_plus_not_infty')
    also have "… / … = 1" using le' finite_OUT_g[of x] x
      by(auto intro!: ennreal_divide_self dest: diff_gr0_ennreal simp: less_top[symmetric])
    also have "d_OUT (h_plus i) x + (d_IN (h_plus i) x - d_OUT (h_plus i) x) * 1 = d_IN (h_plus i) x" using x
      by (simp add: Suc)
    finally show ?thesis by simp
  qed
  also have "… ≤ d_IN (h_plus (Suc i)) x" by(rule d_IN_mono)(rule h_plus_mono)
  finally show ?case .
qed

lemma h_plus_OUT_eq_IN:
  assumes enum: "enum_v (Suc i) = x"
  shows "d_OUT (h_plus (Suc i)) x = d_IN (h_plus i) x"
proof(cases "d_OUT (h_plus i) x < d_IN (h_plus i) x")
  case False
  from enum have "x ≠ source Δ" using range_enum_v by auto
  from h_plus_OUT_le_IN[OF this, of i] False have "d_OUT (h_plus i) x = d_IN (h_plus i) x" by auto
  with False enum show ?thesis by(simp add: d_OUT_def)
next
  case True
  from enum have x: "x ≠ source Δ" and sink: "x ≠ sink Δ" using range_enum_v by auto
  let ?r = "λy. (g (x, y) - h_plus i (x, y)) * (d_IN (h_plus i) x - d_OUT (h_plus i) x) / (d_OUT g x - d_OUT (h_plus i) x)"
  have "d_OUT (h_plus (Suc i)) x = d_OUT (h_plus i) x + (∑+ y. ?r y)"
    using True enum unfolding d_OUT_def h_plus.simps by(simp add: AE_count_space nn_integral_add)
  also from True enum have "x ≠ source Δ" "x ≠ sink Δ" using range_enum_v by auto
  from flowD_KIR[OF g this] True d_IN_mono[of "h_plus i" x g, OF h_plus_le_g]
  have le': "d_OUT (h_plus i) x < d_OUT g x" by(simp)
  then have "(∑+ y. ?r y ) =
    (d_IN (h_plus i) x - d_OUT (h_plus i) x) * ((∑+ y. g (x, y) - h_plus i (x, y)) / (d_OUT g x - d_OUT (h_plus i) x))"
    by(subst mult.commute, subst ennreal_times_divide[symmetric])
      (simp add: nn_integral_cmult nn_integral_divide h_plus_OUT_le_IN[OF x] diff_gr0_ennreal)
  also have "(∑+ y. g (x, y) - h_plus i (x, y)) = d_OUT g x - d_OUT (h_plus i) x" using sink
    by(subst nn_integral_diff)(simp_all add: d_OUT_def[symmetric] h_plus_le_g d_OUT_h_plus_not_infty')
  also have "… / … = 1" using le' finite_OUT_g[of x] sink
    by(auto intro!: ennreal_divide_self dest: diff_gr0_ennreal simp: less_top[symmetric])
  also have "d_OUT (h_plus i) x + (d_IN (h_plus i) x - d_OUT (h_plus i) x) * 1 = d_IN (h_plus i) x" using sink
    by (simp add: h_plus_OUT_le_IN x)
  finally show ?thesis .
qed

lemma h_plus_source_in [simp]: "h_plus i (x, source Δ) = 0"
by(induction i)simp_all

lemma h_plus_sum_finite: "(∑+ e. h_plus i e) ≠ top"
proof(induction i)
  case 0
  have "(∑+ e∈UNIV. h_plus 0 e) = (∑+ (x, y). h_plus 0 (x, y))"
    by(simp del: h_plus.simps)
  also have "… = (∑+ (x, y)∈range (Pair (source Δ)). h_plus 0 (x, y))"
    by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_cong)
  also have "… = value_flow Δ g" by(simp add: d_OUT_def nn_integral_count_space_reindex)
  also have "… < ⊤" using g_finite by (simp add: less_top)
  finally show ?case by simp
next
  case (Suc i)
  def xi  "enum_v (Suc i)"
  then have xi: "xi ≠ source Δ" "xi ≠ sink Δ" using range_enum_v by auto
  show ?case
  proof(cases "d_OUT (h_plus i) xi < d_IN (h_plus i) xi")
    case False
    hence "(∑+ e∈UNIV. h_plus (Suc i) e) = (∑+ e. h_plus i e)"
      by(auto intro!: nn_integral_cong simp add: xi_def)
    with Suc.IH show ?thesis by simp
  next
    case True
    have less: "d_OUT (h_plus i) xi < d_OUT g xi"
      using True flowD_KIR[OF g xi] d_IN_mono[of "h_plus i" xi, OF h_plus_le_g]
      by simp

    have "(∑+ e. h_plus (Suc i) e) =
      (∑+ e∈UNIV. h_plus i e) + (∑+ (x, y). ((g (x, y) - h_plus i (x, y)) * (d_IN (h_plus i) x - d_OUT (h_plus i) x) / (d_OUT g x - d_OUT (h_plus i) x)) * indicator (range (Pair xi)) (x, y))"
      (is "_ = ?IH + ?rest" is "_ = _ + ∫+ (x, y). ?f x y * _ ∂_") using xi True
      by(subst nn_integral_add[symmetric])(auto simp add: xi_def split_beta AE_count_space intro!: nn_integral_cong split: split_indicator intro!: h_plus_le_g h_plus_OUT_le_IN d_OUT_mono le_funI)
    also have "?rest = (∑+ (x, y)∈range (Pair xi). ?f x y)"
      by(simp add: nn_integral_count_space_indicator split_def)
    also have "… = (∑+ y. ?f xi y)" by(simp add: nn_integral_count_space_reindex)
    also have "… = (∑+ y. g (xi, y) - h_plus i (xi, y)) * ((d_IN (h_plus i) xi - d_OUT (h_plus i) xi) / (d_OUT g xi - d_OUT (h_plus i) xi))"
      (is "_ = ?integral * ?factor")  using True less
      by(simp add: nn_integral_multc nn_integral_divide diff_gr0_ennreal ennreal_times_divide)
    also have "?integral = d_OUT g xi - d_OUT (h_plus i) xi" unfolding d_OUT_def using xi
      by(subst nn_integral_diff)(simp_all add: h_plus_le_g d_OUT_def[symmetric] d_OUT_h_plus_not_infty')
    also have "… * ?factor = (d_IN (h_plus i) xi - d_OUT (h_plus i) xi)" using xi
      apply (subst ennreal_times_divide)
      apply (subst mult.commute)
      apply (subst ennreal_mult_divide_eq)
      apply (simp_all add: diff_gr0_ennreal finite_OUT_g less zero_less_iff_neq_zero[symmetric])
      done
    also have "… ≠ ⊤" using h_plus_OUT_eq_IN[OF refl, of i, folded xi_def, symmetric] xi
      by(simp add: d_OUT_h_plus_not_infty')
    ultimately show ?thesis using Suc.IH by simp
  qed
qed

lemma d_OUT_h_plus_not_infty [simp]: "d_OUT (h_plus i) x ≠ top"
proof -
  have "d_OUT (h_plus i) x ≤ (∑+ y∈UNIV. ∑+ x. h_plus i (x, y))"
    unfolding d_OUT_def by(rule nn_integral_mono nn_integral_ge_point)+ simp
  also have "… < ⊤" using h_plus_sum_finite by(simp add: nn_integral_snd_count_space less_top)
  finally show ?thesis by simp
qed

definition enum_cycle :: "nat ⇒ 'v path"
where "enum_cycle = from_nat_into (cycles Δ)"

lemma cycle_enum_cycle [simp]: "cycles Δ ≠ {} ⟹ cycle Δ (enum_cycle n)"
unfolding enum_cycle_def using from_nat_into[of "cycles Δ" n] by simp

context
  fixes h' :: "'v flow"
  assumes finite_h': "h' e ≠ top"
begin

fun h_minus_aux :: "nat ⇒ 'v edge ⇒ ennreal"
where
  "h_minus_aux 0 e = 0"
| "h_minus_aux (Suc j) e =
  (if e ∈ set (cycle_edges (enum_cycle j)) then
     h_minus_aux j e + Min {h' e' - h_minus_aux j e'|e'. e'∈set (cycle_edges (enum_cycle j))}
   else h_minus_aux j e)"

lemma h_minus_aux_le_h': "h_minus_aux j e ≤ h' e"
proof(induction j e rule: h_minus_aux.induct)
  case 0: (1 e) show ?case by simp
next
  case Suc: (2 j e)
  { assume e: "e ∈ set (cycle_edges (enum_cycle j))"
    then have "h_minus_aux j e + Min {h' e' - h_minus_aux j e' |e'. e' ∈ set (cycle_edges (enum_cycle j))} ≤
      h_minus_aux j e + (h' e - h_minus_aux j e)"
      using [[simproc add: finite_Collect]] by(cases e rule: prod.exhaust)(auto intro!: add_mono Min_le)
    also have "… = h' e" using e finite_h'[of e] Suc.IH(2)[of e]
      by(cases e rule: prod.exhaust)
        (auto simp add: add_diff_eq_ennreal top_unique intro!: ennreal_add_diff_cancel_left)
    also note calculation }
  then show ?case using Suc by clarsimp
qed

lemma h_minus_aux_finite [simp]: "h_minus_aux j e ≠ top"
using h_minus_aux_le_h'[of j e] finite_h'[of e] by (auto simp: top_unique)

lemma h_minus_aux_mono: "h_minus_aux j e ≤ h_minus_aux (Suc j) e"
proof(cases "e ∈ set (cycle_edges (enum_cycle j)) = True")
  case True
  have "h_minus_aux j e + 0 ≤ h_minus_aux (Suc j) e" unfolding h_minus_aux.simps True if_True
    using True [[simproc add: finite_Collect]]
    by(cases e)(rule add_mono, auto intro!: Min.boundedI simp add: h_minus_aux_le_h')
  thus ?thesis by simp
qed simp

lemma d_OUT_h_minus_aux:
  assumes "cycles Δ ≠ {}"
  shows "d_OUT (h_minus_aux j) x = d_IN (h_minus_aux j) x"
proof(induction j)
  case 0 show ?case by simp
next
  case (Suc j)
  def C  "enum_cycle j"
  def δ  "Min {h' e' - h_minus_aux j e' |e'. e' ∈ set (cycle_edges C)}"

  have "d_OUT (h_minus_aux (Suc j)) x =
    (∑+ y. h_minus_aux j (x, y) + (if (x, y) ∈ set (cycle_edges C) then δ else 0))"
    unfolding d_OUT_def by(simp add: if_distrib C_def δ_def cong del: if_weak_cong)
  also have "… = d_OUT (h_minus_aux j) x + (∑+ y. δ * indicator (set (cycle_edges C)) (x, y))"
    (is "_ = _ + ?add")
    by(subst nn_integral_add)(auto simp add: AE_count_space d_OUT_def intro!: arg_cong2[where f="op +"] nn_integral_cong)
  also have "?add = (∑+ e∈range (Pair x). δ * indicator {(x', y). (x', y) ∈ set (cycle_edges C) ∧ x' = x} e)"
    by(auto simp add: nn_integral_count_space_reindex intro!: nn_integral_cong split: split_indicator)
  also have "… = δ * card (set (filter (λ(x', y). x' = x) (cycle_edges C)))"
    using [[simproc add: finite_Collect]]
    apply(subst nn_integral_cmult_indicator; auto)
    apply(subst emeasure_count_space; auto simp add: split_def)
    done
  also have "card (set (filter (λ(x', y). x' = x) (cycle_edges C))) = card (set (filter (λ(x', y). y = x) (cycle_edges C)))"
    unfolding C_def by(rule cycle_enter_leave_same)(rule cycle_enum_cycle[OF assms])
  also have "δ * … =  (∑+ e∈range (λx'. (x', x)). δ * indicator {(x', y). (x', y) ∈ set (cycle_edges C) ∧ y = x} e)"
    using [[simproc add: finite_Collect]]
    apply(subst nn_integral_cmult_indicator; auto)
    apply(subst emeasure_count_space; auto simp add: split_def)
    done
  also have "… = (∑+ x'. δ * indicator (set (cycle_edges C)) (x', x))"
    by(auto simp add: nn_integral_count_space_reindex intro!: nn_integral_cong split: split_indicator)
  also have "d_OUT (h_minus_aux j) x + … = (∑+ x'. h_minus_aux j (x', x) + δ * indicator (set (cycle_edges C)) (x', x))"
    unfolding Suc.IH d_IN_def by(simp add: nn_integral_add[symmetric])
  also have "… = d_IN (h_minus_aux (Suc j)) x" unfolding d_IN_def
    by(auto intro!: nn_integral_cong simp add: δ_def C_def split: split_indicator)
  finally show ?case .
qed

lemma h_minus_aux_source:
  assumes "cycles Δ ≠ {}"
  shows "h_minus_aux j (source Δ, y) = 0"
proof(induction j)
  case 0 thus ?case by simp
next
  case (Suc j)
  have "(source Δ, y) ∉ set (cycle_edges (enum_cycle j))"
  proof
    assume *: "(source Δ, y) ∈ set (cycle_edges (enum_cycle j))"
    have cycle: "cycle Δ (enum_cycle j)" using assms by(rule cycle_enum_cycle)
    from cycle_leave_ex_enter[OF this *]
    obtain z where "(z, source Δ) ∈ set (cycle_edges (enum_cycle j))" ..
    with cycle_edges_edges[OF cycle] have "(z, source Δ) ∈ E" ..
    thus False using source_in[of z] by simp
  qed
  then show ?case using Suc.IH by simp
qed

lemma h_minus_aux_cycle:
  fixes j defines "C ≡ enum_cycle j"
  assumes "cycles Δ ≠ {}"
  shows "∃e∈set (cycle_edges C). h_minus_aux (Suc j) e = h' e"
proof -
  let ?A = "{h' e' - h_minus_aux j e'|e'. e' ∈ set (cycle_edges C)}"
  from assms have "cycle Δ C" by auto
  from cycle_edges_not_Nil[OF this] have "Min ?A ∈ ?A" using [[simproc add: finite_Collect]]
    by(intro Min_in)(fastforce simp add: neq_Nil_conv)+
  then obtain e' where e: "e' ∈ set (cycle_edges C)"
    and "Min ?A = h' e' - h_minus_aux j e'" by auto
  hence "h_minus_aux (Suc j) e' = h' e'"
    by(simp add: C_def h_minus_aux_le_h')
  with e show ?thesis by blast
qed

end

fun h_minus :: "nat ⇒ 'v edge ⇒ ennreal"
where
  "h_minus 0 e = 0"
| "h_minus (Suc i) e = h_minus i e + (SUP j. h_minus_aux (λe'. h_plus (Suc i) e' - h_minus i e') j e)"

lemma h_minus_le_h_plus: "h_minus i e ≤ h_plus i e"
proof(induction i e rule: h_minus.induct)
  case 0: (1 e) show ?case by simp
next
  case Suc: (2 i e)
  note IH = Suc.IH(2)[OF UNIV_I]
  let ?h' = "λe'. h_plus (Suc i) e' - h_minus i e'"
  have h': "?h' e' ≠ top" for e' using IH(1)[of e'] by(simp add: )

  have "(⨆j. h_minus_aux ?h' j e) ≤ ?h' e" by(rule SUP_least)(rule h_minus_aux_le_h'[OF h'])
  hence "h_minus (Suc i) e ≤ h_minus i e + …" by(simp add: add_mono)
  also have "… = h_plus (Suc i) e" using IH[of e] h_plus_mono[of i e]
    by auto
  finally show ?case .
qed

lemma finite_h': "h_plus (Suc i) e - h_minus i e ≠ top"
  by simp

lemma h_minus_mono: "h_minus i e ≤ h_minus (Suc i) e"
proof -
  have "h_minus i e + 0 ≤ h_minus (Suc i) e" unfolding h_minus.simps
    by(rule add_mono; simp add: SUP_upper2)
  thus ?thesis by simp
qed

lemma h_minus_finite [simp]: "h_minus i e ≠ ⊤"
proof -
  have "h_minus i e ≤ h_plus i e" by(rule h_minus_le_h_plus)
  also have "… < ⊤" by (simp add: less_top[symmetric])
  finally show ?thesis by simp
qed

lemma d_OUT_h_minus:
  assumes cycles: "cycles Δ ≠ {}"
  shows "d_OUT (h_minus i) x = d_IN (h_minus i) x"
proof(induction i)
  case (Suc i)
  let ?h' = "λe. h_plus (Suc i) e - h_minus i e"
  have "d_OUT (λe. h_minus (Suc i) e) x = d_OUT (h_minus i) x + d_OUT (λe. SUP j. h_minus_aux ?h' j e) x"
    by(simp add: d_OUT_add SUP_upper2)
  also have "d_OUT (λe. SUP j. h_minus_aux ?h' j e) x = (SUP j. d_OUT (h_minus_aux ?h' j) x)"
    by(rule d_OUT_monotone_convergence_SUP incseq_SucI le_funI h_minus_aux_mono finite_h')+
  also have "… = (SUP j. d_IN (h_minus_aux ?h' j) x)"
    by(rule SUP_cong[OF refl])(rule d_OUT_h_minus_aux[OF finite_h' cycles])
  also have "… = d_IN (λe. SUP j. h_minus_aux ?h' j e) x"
    by(rule d_IN_monotone_convergence_SUP[symmetric] incseq_SucI le_funI h_minus_aux_mono finite_h')+
  also have "d_OUT (h_minus i) x + … = d_IN (λe. h_minus (Suc i) e) x" using Suc.IH
    by(simp add: d_IN_add SUP_upper2)
  finally show ?case .
qed simp

lemma h_minus_source:
  assumes "cycles Δ ≠ {}"
  shows "h_minus n (source Δ, y) = 0"
by(induction n)(simp_all add: h_minus_aux_source[OF finite_h' assms])

lemma h_minus_source_in [simp]: "h_minus i (x, source Δ) = 0"
using h_minus_le_h_plus[of i "(x, source Δ)"] by simp

lemma h_minus_OUT_finite [simp]: "d_OUT (h_minus i) x ≠ top"
proof -
  have "d_OUT (h_minus i) x ≤ d_OUT (h_plus i) x" by(rule d_OUT_mono)(rule h_minus_le_h_plus)
  also have "… < ⊤" by (simp add: less_top[symmetric])
  finally show ?thesis by simp
qed

lemma h_minus_cycle:
  assumes "cycle Δ C"
  shows "∃e∈set (cycle_edges C). h_minus i e = h_plus i e"
proof(cases i)
  case (Suc i)
  let ?h' = "λe. h_plus (Suc i) e - h_minus i e"
  from assms have cycles: "cycles Δ ≠ {}" by auto
  with assms from_nat_into_surj[of "cycles Δ" C] obtain j where j: "C = enum_cycle j"
    by(auto simp add: enum_cycle_def)
  from h_minus_aux_cycle[of "?h'" j, OF finite_h' cycles] j
  obtain e where e: "e ∈ set (cycle_edges C)" and "h_minus_aux ?h' (Suc j) e = ?h' e" by(auto)
  then have "h_plus (Suc i) e = h_minus i e + h_minus_aux ?h' (Suc j) e"
    using order_trans[OF h_minus_le_h_plus h_plus_mono]
    by (subst eq_commute) simp
  also have "… ≤ h_minus (Suc i) e" unfolding h_minus.simps
    by(intro add_mono SUP_upper; simp)
  finally show ?thesis using e h_minus_le_h_plus[of "Suc i" e] Suc by auto
next
  case 0
  from cycle_edges_not_Nil[OF assms] obtain x y where e: "(x, y) ∈ set (cycle_edges C)"
    by(fastforce simp add: neq_Nil_conv)
  then have "x ≠ source Δ" using assms by(auto dest: source_out_not_cycle)
  hence "h_plus 0 (x, y) = 0" by simp
  with e 0 show ?thesis by(auto simp del: h_plus.simps)
qed

abbreviation lim_h_plus :: "'v edge ⇒ ennreal"
where "lim_h_plus e ≡ SUP n. h_plus n e"

abbreviation lim_h_minus :: "'v edge ⇒ ennreal"
where "lim_h_minus e ≡ SUP n. h_minus n e"

lemma lim_h_plus_le_g: "lim_h_plus e ≤ g e"
by(rule SUP_least)(rule h_plus_le_g)

lemma lim_h_plus_finite [simp]: "lim_h_plus e ≠ top"
proof -
  have "lim_h_plus e ≤ g e" by(rule lim_h_plus_le_g)
  also have "… < top" by (simp add: less_top[symmetric])
  finally show ?thesis unfolding less_top .
qed

lemma lim_h_minus_le_lim_h_plus: "lim_h_minus e ≤ lim_h_plus e"
by(rule SUP_mono)(blast intro: h_minus_le_h_plus)

lemma lim_h_minus_finite [simp]: "lim_h_minus e ≠ top"
proof -
  have "lim_h_minus e ≤ lim_h_plus e" by(rule lim_h_minus_le_lim_h_plus)
  also have "… < top" unfolding less_top[symmetric] by (rule lim_h_plus_finite)
  finally show ?thesis unfolding less_top[symmetric] by simp
qed

lemma lim_h_minus_IN_finite [simp]:
  assumes "x ≠ sink Δ"
  shows "d_IN lim_h_minus x ≠ top"
proof -
  have "d_IN lim_h_minus x ≤ d_IN lim_h_plus x"
    by(intro d_IN_mono le_funI lim_h_minus_le_lim_h_plus)
  also have "… ≤ d_IN g x" by(intro d_IN_mono le_funI lim_h_plus_le_g)
  also have "… < ⊤" using assms by(simp add: finite_IN_g less_top[symmetric])
  finally show ?thesis by simp
qed

lemma lim_h_plus_OUT_IN:
  assumes "x ≠ source Δ" "x ≠ sink Δ"
  shows "d_OUT lim_h_plus x = d_IN lim_h_plus x"
proof(cases "x ∈ V")
  case True
  have "d_OUT lim_h_plus x = (SUP n. d_OUT (h_plus n) x)"
    by(rule d_OUT_monotone_convergence_SUP incseq_SucI le_funI h_plus_mono)+
  also have "… = (SUP n. d_IN (h_plus n) x)" (is "?lhs = ?rhs")
  proof(rule antisym)
    show "?lhs ≤ ?rhs" by(rule SUP_mono)(auto intro: h_plus_OUT_le_IN[OF assms(1)])
    show "?rhs ≤ ?lhs"
    proof(rule SUP_mono)
      fix i
      from enum_v_repeat[OF True assms, of i]
      obtain i' where "i' > i" "enum_v i' = x" by auto
      moreover then obtain i'' where i': "i' = Suc i''" by(cases i') auto
      ultimately have "d_OUT (h_plus i') x = d_IN (h_plus i'') x" using  ‹x ≠ source Δ›
        by(simp add: h_plus_OUT_eq_IN)
      moreover have "i ≤ i''" using `i < i'` i' by simp
      then have "d_IN (h_plus i) x ≤ d_IN (h_plus i'') x" by(intro d_IN_mono h_plus_mono')
      ultimately have "d_IN (h_plus i) x ≤ d_OUT (h_plus i') x" by simp
      thus "∃i'∈UNIV. d_IN (h_plus i) x ≤ d_OUT (h_plus i') x" by blast
    qed
  qed
  also have "… = d_IN lim_h_plus x"
    by(rule d_IN_monotone_convergence_SUP[symmetric] incseq_SucI le_funI h_plus_mono)+
  finally show ?thesis .
next
  case False
  have "(x, y) ∉ support_flow lim_h_plus" for y using False h_plus_outside[of "(x, y)"]
    by(fastforce elim!: support_flow.cases simp add: less_SUP_iff vertex_def)
  moreover have "(y, x) ∉ support_flow lim_h_plus" for y using False h_plus_outside[of "(y, x)"]
    by(fastforce elim!: support_flow.cases simp add: less_SUP_iff vertex_def)
  ultimately show ?thesis
    by(auto simp add: d_OUT_alt_def2 d_IN_alt_def2 AE_count_space intro!: nn_integral_cong_AE)
qed

lemma lim_h_minus_OUT_IN:
  assumes cycles: "cycles Δ ≠ {}"
  shows "d_OUT lim_h_minus x = d_IN lim_h_minus x"
proof -
  have "d_OUT lim_h_minus x = (SUP n. d_OUT (h_minus n) x)"
    by(rule d_OUT_monotone_convergence_SUP incseq_SucI le_funI h_minus_mono)+
  also have "… = (SUP n. d_IN (h_minus n) x)" using cycles by(simp add: d_OUT_h_minus)
  also have "… = d_IN lim_h_minus x"
    by(rule d_IN_monotone_convergence_SUP[symmetric] incseq_SucI le_funI h_minus_mono)+
  finally show ?thesis .
qed

definition h :: "'v edge ⇒ ennreal"
where "h e = lim_h_plus e - (if cycles Δ ≠ {} then lim_h_minus e else 0)"

lemma h_le_lim_h_plus: "h e ≤ lim_h_plus e"
by (simp add: h_def)

lemma h_le_g: "h e ≤ g e"
using h_le_lim_h_plus[of e] lim_h_plus_le_g[of e] by simp

lemma flow_h: "flow Δ h"
proof
  fix e
  have "h e ≤ lim_h_plus e" by(rule h_le_lim_h_plus)
  also have "… ≤ g e" by(rule lim_h_plus_le_g)
  also have "… ≤ capacity Δ e" using g by(rule flowD_capacity)
  finally show "h e ≤ …" .
next
  fix x
  assume "x ≠ source Δ" "x ≠ sink Δ"
  then show "KIR h x"
    by (cases "cycles Δ = {}")
       (auto simp add: h_def[abs_def] lim_h_plus_OUT_IN d_OUT_diff d_IN_diff lim_h_minus_le_lim_h_plus lim_h_minus_OUT_IN)
qed

lemma value_h_plus: "value_flow Δ (h_plus i) = value_flow Δ g" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs ≤ ?rhs" by(rule d_OUT_mono)(rule h_plus_le_g)

  have "?rhs ≤ value_flow Δ (h_plus 0)"
    by(auto simp add: d_OUT_def cong: if_cong intro!: nn_integral_mono)
  also have "… ≤ value_flow Δ (h_plus i)"
    by(rule d_OUT_mono)(rule h_plus_mono'; simp)
  finally show "?rhs ≤ ?lhs" .
qed

lemma value_h: "value_flow Δ h = value_flow Δ g" (is "?lhs = ?rhs")
proof(rule antisym)
  have "?lhs ≤ value_flow Δ lim_h_plus" using ennreal_minus_mono
    by(fastforce simp add: h_def intro!: d_OUT_mono)
  also have "… ≤ ?rhs" by(rule d_OUT_mono)(rule lim_h_plus_le_g)
  finally show "?lhs ≤ ?rhs" .

  show "?rhs ≤ ?lhs"
    by(auto simp add: d_OUT_def h_def h_minus_source cong: if_cong intro!: nn_integral_mono SUP_upper2[where i=0])
qed


definition h_diff :: "nat ⇒ 'v edge ⇒ ennreal"
where "h_diff i e = h_plus i e - (if cycles Δ ≠ {} then h_minus i e else 0)"

lemma d_IN_h_source [simp]: "d_IN (h_diff i) (source Δ) = 0"
by(simp add: d_IN_def h_diff_def cong del: if_weak_cong)

lemma h_diff_le_h_plus: "h_diff i e ≤ h_plus i e"
by(simp add: h_diff_def)

lemma h_diff_le_g: "h_diff i e ≤ g e"
using h_diff_le_h_plus[of i e] h_plus_le_g[of i e] by simp

lemma h_diff_loop [simp]: "h_diff i (x, x) = 0"
using h_diff_le_g[of i "(x, x)"] by simp

lemma supp_h_diff_edges: "support_flow (h_diff i) ⊆ E"
proof
  fix e
  assume "e ∈ support_flow (h_diff i)"
  then have "0 < h_diff i e" by(auto elim: support_flow.cases)
  also have "h_diff i e ≤ h_plus i e" by(rule h_diff_le_h_plus)
  finally show "e ∈ E" using h_plus_outside[of e i] by(cases "e ∈ E") auto
qed

lemma h_diff_OUT_le_IN:
  assumes "x ≠ source Δ"
  shows "d_OUT (h_diff i) x ≤ d_IN (h_diff i) x"
proof(cases "cycles Δ ≠ {}")
  case False
  thus ?thesis using assms by(simp add: h_diff_def[abs_def] h_plus_OUT_le_IN)
next
  case cycles: True
  then have "d_OUT (h_diff i) x = d_OUT (h_plus i) x - d_OUT (h_minus i) x"
    unfolding h_diff_def[abs_def] using assms
    by (simp add: h_minus_le_h_plus d_OUT_diff)
  also have "… ≤ d_IN (h_plus i) x - d_IN (h_minus i) x" using cycles assms
    by(intro ennreal_minus_mono h_plus_OUT_le_IN)(simp_all add: d_OUT_h_minus)
  also have "… = d_IN (h_diff i) x" using cycles
    unfolding h_diff_def[abs_def] by(subst d_IN_diff)(simp_all add: h_minus_le_h_plus d_OUT_h_minus[symmetric])
  finally show ?thesis .
qed

lemma h_diff_cycle:
  assumes "cycle Δ p"
  shows "∃e∈set (cycle_edges p). h_diff i e = 0"
proof -
  from h_minus_cycle[OF assms, of i] obtain e
    where e: "e ∈ set (cycle_edges p)" and "h_minus i e = h_plus i e" by auto
  hence "h_diff i e = 0" using assms by(auto simp add: h_diff_def)
  with e show ?thesis by blast
qed

lemma d_IN_h_le_value': "d_IN (h_diff i) x ≤ value_flow Δ (h_plus i)"
proof -
  let ?supp = "support_flow (h_diff i)"
  def X  "{y. (y, x) ∈ ?supp^*} - {x}"

  { fix x y
    assume x: "x ∉ X" and y: "y ∈ X"
    { assume yx: "(y, x) ∈ ?supp*" and neq: "y ≠ x" and xy: "(x, y) ∈ ?supp"
      from yx obtain p' where "rtrancl_path (λx y. (x, y) ∈ ?supp) y p' x"
        unfolding rtrancl_def rtranclp_eq_rtrancl_path by auto
      then obtain p where p: "rtrancl_path (λx y. (x, y) ∈ ?supp) y p x"
        and distinct: "distinct (y # p)" by(rule rtrancl_path_distinct)
      with neq have "p ≠ []" by(auto elim: rtrancl_path.cases)

      from xy have "(x, y) ∈ E" using supp_h_diff_edges[of i] by(auto)
      moreover from p have "path Δ y p x"
        by(rule rtrancl_path_mono)(auto dest: supp_h_diff_edges[THEN subsetD])
      ultimately have "path Δ x (y # p) x" by(auto intro: rtrancl_path.intros)
      hence cycle: "cycle Δ (y # p)" using _ distinct by(rule cycle) simp
      from h_diff_cycle[OF this, of i] obtain e
        where e: "e ∈ set (cycle_edges (y # p))" and 0: "h_diff i e = 0" by blast
      from e obtain n where e': "e = ((y # p) ! n, (p @ [y]) ! n)" and n: "n < Suc (length p)"
        by(auto simp add: cycle_edges_def set_zip)
      have "e ∈ ?supp"
      proof(cases "n = length p")
        case True
        with rtrancl_path_last[OF p] ‹p ≠ []› have "(y # p) ! n = x"
          by(cases p)(simp_all add: last_conv_nth del: last.simps)
        with e' True have "e = (x, y)" by simp
        with xy show ?thesis by simp
      next
        case False
        with n have "n < length p" by simp
        with rtrancl_path_nth[OF p this] e' show ?thesis by(simp add: nth_append)
      qed
      with 0 have False by(simp add: support_flow.simps) }
    hence "(x, y) ∉ ?supp" using x y
      by(auto simp add: X_def intro: converse_rtrancl_into_rtrancl)
    then have "h_diff i (x, y) = 0"
      by(simp add: support_flow.simps) }
  note acyclic = this

  { fix y
    assume "y ∉ X"
    hence "(y, x) ∉ ?supp" by(auto simp add: X_def support_flow.simps intro: not_in_support_flowD)
    hence "h_diff i (y, x) = 0"  by(simp add: support_flow.simps) }
  note in_X = this

  let ?diff = "λx. (∑+ y. h_diff i (x, y) * indicator X x * indicator X y)"
  have finite2: "(∑+ x. ?diff x) ≠ top" (is "?lhs ≠ _")
  proof -
    have "?lhs ≤ (∑+ x∈UNIV. ∑+ y. h_plus i (x, y))"
      by(intro nn_integral_mono)(auto simp add: h_diff_def split: split_indicator)
    also have "… = (∑+ e. h_plus i e)" by(rule nn_integral_fst_count_space)
    also have "… < ⊤" by(simp add: h_plus_sum_finite less_top[symmetric])
    finally show ?thesis by simp
  qed
  have finite1: "?diff x ≠ top" for x
    using finite2 by(rule neq_top_trans)(rule nn_integral_ge_point, simp)
  have finite3: "(∑+ x. d_OUT (h_diff i) x * indicator (X - {source Δ}) x) ≠ ⊤" (is "?lhs ≠ _")
  proof -
    have "?lhs ≤ (∑+ x∈UNIV. ∑+ y. h_plus i (x, y))" unfolding d_OUT_def
      apply(simp add: nn_integral_multc[symmetric])
      apply(intro nn_integral_mono)
      apply(auto simp add: h_diff_def split: split_indicator)
      done
    also have "… = (∑+ e. h_plus i e)" by(rule nn_integral_fst_count_space)
    also have "… < ⊤" by(simp add: h_plus_sum_finite less_top[symmetric])
    finally show ?thesis by simp
  qed

  have "d_IN (h_diff i) x = (∑+ y. h_diff i (y, x) * indicator X y)" unfolding d_IN_def
    by(rule nn_integral_cong)(simp add: in_X split: split_indicator)
  also have "… ≤ (∑+ x∈- X. ∑+ y. h_diff i (y, x) * indicator X y)"
    by(rule nn_integral_ge_point)(simp add: X_def)
  also have "… = (∑+ x∈UNIV. ∑+ y. h_diff i (y, x) * indicator X y * indicator (- X) x)"
    by(simp add: nn_integral_multc nn_integral_count_space_indicator)
  also have "… = (∑+ x∈UNIV. ∑+ y. h_diff i (x, y) * indicator X x * indicator (- X) y)"
    by(subst nn_integral_snd_count_space[where f="case_prod _", simplified])(simp add: nn_integral_fst_count_space[where f="case_prod _", simplified])
  also have "… = (∑+ x∈UNIV. (∑+ y. h_diff i (x, y) * indicator X x * indicator (- X) y) + (?diff x - ?diff x))"
    by(simp add: finite1)
  also have "… = (∑+ x∈UNIV. (∑+ y. h_diff i (x, y) * indicator X x * indicator (- X) y + h_diff i (x, y) * indicator X x * indicator X y) - ?diff x)"
    apply (subst add_diff_eq_ennreal)
    apply simp
    by(subst nn_integral_add[symmetric])(simp_all add:)
  also have "… = (∑+ x∈UNIV. (∑+ y. h_diff i (x, y) * indicator X x) - ?diff x)"
    by(auto intro!: nn_integral_cong arg_cong2[where f="op -"] split: split_indicator)
  also have "… = (∑+ x∈UNIV. ∑+ y∈UNIV. h_diff i (x, y) * indicator X x) - (∑+ x. ?diff x)"
    by(subst nn_integral_diff)(auto simp add: AE_count_space finite2 intro!: nn_integral_mono split: split_indicator)
  also have "(∑+ x∈UNIV. ∑+ y∈UNIV. h_diff i (x, y) * indicator X x) = (∑+ x. d_OUT (h_diff i) x * indicator X x)"
    unfolding d_OUT_def by(simp add: nn_integral_multc)
  also have "… = (∑+ x. d_OUT (h_diff i) x * indicator (X - {source Δ}) x + value_flow Δ (h_diff i) * indicator X (source Δ) * indicator {source Δ} x)"
    by(rule nn_integral_cong)(simp split: split_indicator)
  also have "… = (∑+ x. d_OUT (h_diff i) x * indicator (X - {source Δ}) x) + value_flow Δ (h_diff i) * indicator X (source Δ)"
    (is "_ = ?out" is "_ = _ + ?value")
    by(subst nn_integral_add) simp_all
  also have "(∑+ x∈UNIV. ∑+ y. h_diff i (x, y) * indicator X x * indicator X y) =
             (∑+ x∈UNIV. ∑+ y. h_diff i (x, y) * indicator X y)"
    using acyclic by(intro nn_integral_cong)(simp split: split_indicator)
  also have "… = (∑+ y∈UNIV. ∑+ x. h_diff i (x, y) * indicator X y)"
    by(subst nn_integral_snd_count_space[where f="case_prod _", simplified])(simp add: nn_integral_fst_count_space[where f="case_prod _", simplified])
  also have "… = (∑+ y. d_IN (h_diff i) y * indicator X y)" unfolding d_IN_def
    by(simp add: nn_integral_multc)
  also have "… = (∑+ y. d_IN (h_diff i) y * indicator (X - {source Δ}) y)"
    by(rule nn_integral_cong)(simp split: split_indicator)
  also have "?out - … ≤ (∑+ x. d_OUT (h_diff i) x * indicator (X - {source Δ}) x) - … + ?value"
    by (auto simp add: add_ac intro!: add_diff_le_ennreal)
  also have "… ≤ 0 + ?value" using h_diff_OUT_le_IN finite3
    by(intro nn_integral_mono add_right_mono)(auto split: split_indicator intro!: diff_eq_0_ennreal nn_integral_mono simp add: less_top)
  also have "… ≤ value_flow Δ (h_diff i)" by(simp split: split_indicator)
  also have "… ≤ value_flow Δ (h_plus i)" by(rule d_OUT_mono le_funI h_diff_le_h_plus)+
  finally show ?thesis .
qed

lemma d_IN_h_le_value: "d_IN h x ≤ value_flow Δ h" (is "?lhs ≤ ?rhs")
proof -
  have [tendsto_intros]: "(λi. h_plus i e) ⇢ lim_h_plus e" for e
    by(rule LIMSEQ_SUP incseq_SucI h_plus_mono)+
  have [tendsto_intros]: "(λi. h_minus i e) ⇢ lim_h_minus e" for e
    by(rule LIMSEQ_SUP incseq_SucI h_minus_mono)+
  have "(λi. h_diff i e) ⇢ lim_h_plus e - (if cycles Δ ≠ {} then lim_h_minus e else 0)" for e
    by(auto intro!: tendsto_intros tendsto_diff_ennreal simp add: h_diff_def simp del: Sup_eq_top_iff SUP_eq_top_iff)
  then have "d_IN h x = (∑+ y. liminf (λi. h_diff i (y, x)))"
    by(simp add: d_IN_def h_def tendsto_iff_Liminf_eq_Limsup)
  also have "… ≤ liminf (λi. d_IN (h_diff i) x)" unfolding d_IN_def
    by(rule nn_integral_liminf) simp_all
  also have "… ≤ liminf (λi. value_flow Δ h)" using d_IN_h_le_value'[of _ x]
    by(intro Liminf_mono eventually_sequentiallyI)(auto simp add: value_h_plus value_h)
  also have "… = value_flow Δ h" by(simp add: Liminf_const)
  finally show ?thesis .
qed

lemma flow_cleanup: -- ‹Lemma 5.4›
  "∃h ≤ g. flow Δ h ∧ value_flow Δ h = value_flow Δ g ∧ (∀x. d_IN h x ≤ value_flow Δ h)"
by(intro exI[where x=h] conjI strip le_funI d_IN_h_le_value flow_h value_h h_le_g)

end

subsection ‹Residual network›

context countable_network begin

definition residual_network :: "'v flow ⇒ ('v, 'more) network_scheme"
where "residual_network f =
  ⦇edge = λx y. edge Δ x y ∨ edge Δ y x ∧ y ≠ source Δ,
   capacity = λ(x, y). if edge Δ x y then capacity Δ (x, y) - f (x, y) else if y = source Δ then 0 else f (y, x),
   source = source Δ, sink = sink Δ, … = network.more Δ ⦈"

lemma residual_network_sel [simp]:
  "edge (residual_network f) x y ⟷ edge Δ x y ∨ edge Δ y x ∧ y ≠ source Δ"
  "capacity (residual_network f) (x, y) = (if edge Δ x y then capacity Δ (x, y) - f (x, y) else if y = source Δ then 0 else f (y, x))"
  "source (residual_network f) = source Δ"
  "sink (residual_network f) = sink Δ"
  "network.more (residual_network f) = network.more Δ"
by(simp_all add: residual_network_def)

lemma "E_residual_network": "Eresidual_network f = E ∪ {(x, y). (y, x) ∈ E ∧ y ≠ source Δ}"
by auto

lemma vertices_residual_network [simp]: "vertex (residual_network f) = vertex Δ"
by(auto simp add: vertex_def fun_eq_iff)

inductive wf_residual_network :: "bool"
where "⟦ ⋀x y. (x, y) ∈ E ⟹ (y, x) ∉ E; (source Δ, sink Δ) ∉ E ⟧ ⟹ wf_residual_network"

lemma wf_residual_networkD:
  "⟦ wf_residual_network; edge Δ x y ⟧ ⟹ ¬ edge Δ y x"
  "⟦ wf_residual_network; e ∈ E ⟧ ⟹ prod.swap e ∉ E"
  "⟦ wf_residual_network; edge Δ (source Δ) (sink Δ) ⟧ ⟹ False"
by(auto simp add: wf_residual_network.simps)

lemma residual_countable_network:
  assumes wf: "wf_residual_network"
  and f: "flow Δ f"
  shows "countable_network (residual_network f)" (is "countable_network ?Δ")
proof
  have "countable (converse E)" by simp
  then have "countable {(x, y). (y, x) ∈ E ∧ y ≠ source Δ}"
    by(rule countable_subset[rotated]) auto
  then show "countable E" unfolding "E_residual_network" by simp

  show "source ?Δ ≠ sink ?Δ" by simp
  show "capacity ?Δ e = 0" if "e ∉ E" for e using that by(cases e)(auto intro: flowD_outside[OF f])
  show "capacity ?Δ e ≠ top" for e
    using flowD_finite[OF f] by(cases e) auto
qed

end

locale antiparallel_edges = countable_network Δ
  for Δ :: "('v, 'more) network_scheme" (structure)
begin

text ‹We eliminate the assumption of antiparallel edges by adding a vertex for every edge.
  Thus, antiparallel edges are split up into a cycle of 4 edges. This idea already appears in
  @{cite Aharoni1983EJC}.›

datatype (plugins del: transfer size) 'v' vertex = Vertex 'v' | Edge 'v' 'v'

inductive edg :: "'v vertex ⇒ 'v vertex ⇒ bool"
where
  OUT: "edge Δ x y ⟹ edg (Vertex x) (Edge x y)"
| IN: "edge Δ x y ⟹ edg (Edge x y) (Vertex y)"

inductive_simps edg_simps [simp]:
  "edg (Vertex x) v"
  "edg (Edge x y) v"
  "edg v (Vertex x)"
  "edg v (Edge x y)"

fun split :: "'v flow ⇒ 'v vertex flow"
where
  "split f (Vertex x, Edge x' y) = (if x' = x then f (x, y) else 0)"
| "split f (Edge x y', Vertex y) = (if y' = y then f (x, y) else 0)"
| "split f _ = 0"

lemma split_Vertex1_eq_0I: "(⋀z. y ≠ Edge x z) ⟹ split f (Vertex x, y) = 0"
by(cases y) auto

lemma split_Vertex2_eq_0I: "(⋀z. y ≠ Edge z x) ⟹ split f (y, Vertex x) = 0"
by(cases y) simp_all

lemma split_Edge1_eq_0I: "(⋀z. y ≠ Vertex x) ⟹ split f (Edge z x, y) = 0"
by(cases y) simp_all

lemma split_Edge2_eq_0I: "(⋀z. y ≠ Vertex x) ⟹ split f (y, Edge x z) = 0"
by(cases y) simp_all

definition Δ'' :: "'v vertex network"
where "Δ'' = ⦇edge = edg, capacity = split (capacity Δ), source = Vertex (source Δ), sink = Vertex (sink Δ)⦈"

lemma Δ''_sel [simp]:
  "edge Δ'' = edg"
  "capacity Δ'' = split (capacity Δ)"
  "source Δ'' = Vertex (source Δ)"
  "sink Δ'' = Vertex (sink Δ)"
by(simp_all add: Δ''_def)

lemma "E_Δ''": "EΔ'' = (λ(x, y). (Vertex x, Edge x y)) ` E ∪ (λ(x, y). (Edge x y, Vertex y)) ` E"
by(auto elim: edg.cases)

lemma "V_Δ''": "VΔ'' = Vertex ` V ∪ case_prod Edge ` E"
by(auto 4 4 simp add: vertex_def elim!: edg.cases)

lemma inj_on_Edge1 [simp]: "inj_on (λx. Edge x y) A"
by(simp add: inj_on_def)

lemma inj_on_Edge2 [simp]: "inj_on (Edge x) A"
by(simp add: inj_on_def)

lemma d_IN_split_Vertex [simp]: "d_IN (split f) (Vertex x) = d_IN f x" (is "?lhs = ?rhs")
proof(rule trans)
  show "?lhs = (∑+ v'∈range (λy. Edge y x). split f (v', Vertex x))"
    by(auto intro!: nn_integral_cong split_Vertex2_eq_0I simp add: d_IN_def nn_integral_count_space_indicator split: split_indicator)
  show "… = ?rhs" by(simp add: nn_integral_count_space_reindex d_IN_def)
qed

lemma d_OUT_split_Vertex [simp]: "d_OUT (split f) (Vertex x) = d_OUT f x" (is "?lhs = ?rhs")
proof(rule trans)
  show "?lhs = (∑+ v'∈range (Edge x). split f (Vertex x, v'))"
    by(auto intro!: nn_integral_cong split_Vertex1_eq_0I simp add: d_OUT_def nn_integral_count_space_indicator split: split_indicator)
  show "… = ?rhs" by(simp add: nn_integral_count_space_reindex d_OUT_def)
qed

lemma d_IN_split_Edge [simp]: "d_IN (split f) (Edge x y) = max 0 (f (x, y))" (is "?lhs = ?rhs")
proof(rule trans)
  show "?lhs = (∑+ v'. split f (v', Edge x y) * indicator {Vertex x} v')"
    unfolding d_IN_def by(rule nn_integral_cong)(simp add: split_Edge2_eq_0I split: split_indicator)
  show "… = ?rhs" by(simp add: max_def)
qed

lemma d_OUT_split_Edge [simp]: "d_OUT (split f) (Edge x y) = max 0 (f (x, y))" (is "?lhs = ?rhs")
proof(rule trans)
  show "?lhs = (∑+ v'. split f (Edge x y, v') * indicator {Vertex y} v')"
    unfolding d_OUT_def by(rule nn_integral_cong)(simp add: split_Edge1_eq_0I split: split_indicator)
  show "… = ?rhs" by(simp add: max_def)
qed

lemma Δ''_countable_network: "countable_network Δ''"
proof
  show "countable EΔ''" unfolding "E_Δ''" by(simp)
  show "source Δ'' ≠ sink Δ''" by auto
  show "capacity Δ'' e = 0" if "e ∉ EΔ''" for e using that
    by(cases "(capacity Δ, e)" rule: split.cases)(auto simp add: capacity_outside)
  show "capacity Δ'' e ≠ top" for e by(cases "(capacity Δ, e)" rule: split.cases)(auto)
qed

interpretation Δ'': countable_network Δ'' by(rule Δ''_countable_network)

lemma Δ''_flow_attainability:
  assumes "flow_attainability_axioms Δ"
  shows "flow_attainability Δ''"
proof -
  interpret flow_attainability Δ using _ assms by(rule flow_attainability.intro) unfold_locales
  show ?thesis
  proof
    show "d_IN (capacity Δ'') v ≠ ⊤ ∨ d_OUT (capacity Δ'') v ≠ ⊤" if "v ≠ sink Δ''" for v
      using that finite_capacity by(cases v)(simp_all add: max_def)
    show "¬ edge Δ'' v v" for v by(auto elim: edg.cases)
    show "¬ edge Δ'' v (source Δ'')" for v by(simp add: source_in)
  qed
qed

lemma flow_split [simp]:
  assumes "flow Δ f"
  shows "flow Δ'' (split f)"
proof
  show "split f e ≤ capacity Δ'' e" for e
    by(cases "(f, e)" rule: split.cases)(auto intro: flowD_capacity[OF assms] intro: SUP_upper2 assms)
  show "KIR (split f) x" if "x ≠ source Δ''" "x ≠ sink Δ''" for x
    using that by(cases "x")(auto dest: flowD_KIR[OF assms])
qed

abbreviation (input) collect :: "'v vertex flow ⇒ 'v flow"
where "collect f ≡ (λ(x, y). f (Edge x y, Vertex y))"

lemma d_OUT_collect:
  assumes f: "flow Δ'' f"
  shows "d_OUT (collect f) x = d_OUT f (Vertex x)"
proof -
  have "d_OUT (collect f) x = (∑+ y. f (Edge x y, Vertex y))"
    by(simp add: nn_integral_count_space_reindex d_OUT_def)
  also have "… = (∑+ y∈range (Edge x). f (Vertex x, y))"
  proof(clarsimp simp add: nn_integral_count_space_reindex intro!: nn_integral_cong)
    fix y
    have "(∑+ z. f (Edge x y, z) * indicator {Vertex y} z) = d_OUT f (Edge x y)"
      unfolding d_OUT_def by(rule nn_integral_cong)(simp split: split_indicator add: Δ''.flowD_outside[OF f])
    also have "… = d_IN f (Edge x y)" using f by(rule flowD_KIR) simp_all
    also have "… = (∑+ z. f (z, Edge x y) * indicator {Vertex x} z)"
      unfolding d_IN_def by(rule nn_integral_cong)(simp split: split_indicator add: Δ''.flowD_outside[OF f])
    finally show "f (Edge x y, Vertex y) = f (Vertex x, Edge x y)"
      by(simp add: max_def)
  qed
  also have "… = d_OUT f (Vertex x)"
    by(auto intro!: nn_integral_cong Δ''.flowD_outside[OF f] simp add: nn_integral_count_space_indicator d_OUT_def split: split_indicator)
  finally show ?thesis .
qed

lemma flow_collect [simp]:
  assumes f: "flow Δ'' f"
  shows "flow Δ (collect f)"
proof
  show "collect f e ≤ capacity Δ e" for e using flowD_capacity[OF f, of "(case_prod Edge e, Vertex (snd e))"]
    by(cases e)(simp)

  fix x
  assume x: "x ≠ source Δ" "x ≠ sink Δ"
  have "d_OUT (collect f) x = d_OUT f (Vertex x)" using f by(rule d_OUT_collect)
  also have "… = d_IN f (Vertex x)" using x flowD_KIR[OF f, of "Vertex x"] by(simp)
  also have "… = (∑+ y∈range (λz. Edge z x). f (y, Vertex x))"
    by(auto intro!: nn_integral_cong Δ''.flowD_outside[OF f] simp add: nn_integral_count_space_indicator d_IN_def split: split_indicator)
  also have "… = d_IN (collect f) x" by(simp add: nn_integral_count_space_reindex d_IN_def)
  finally show "KIR (collect f) x" .
qed

lemma value_collect: "flow Δ'' f ⟹ value_flow Δ (collect f) = value_flow Δ'' f"
by(simp add: d_OUT_collect)

lemma Δ''_wf_residual_network:
  assumes no_loop: "⋀x. ¬ edge Δ x x"
  shows "Δ''.wf_residual_network"
by(auto simp add: Δ''.wf_residual_network.simps assms elim!: edg.cases)

end

subsection ‹The attainability theorem›

context flow_attainability begin

lemma residual_flow_attainability:
  assumes wf: "wf_residual_network"
  and f: "flow Δ f"
  shows "flow_attainability (residual_network f)" (is "flow_attainability ?Δ")
proof -
  interpret res: countable_network "residual_network f" by(rule residual_countable_network[OF assms])
  show ?thesis
  proof
    fix x
    assume sink: "x ≠ sink ?Δ"
    then consider (source) "x = source Δ" | (IN) "d_IN (capacity Δ) x ≠ ⊤" | (OUT) "x ≠ source Δ" "d_OUT (capacity Δ) x ≠ ⊤"
      using finite_capacity[of x] by auto
    then show "d_IN (capacity ?Δ) x ≠ ⊤ ∨ d_OUT (capacity ?Δ) x ≠ ⊤"
    proof(cases)
      case source
      hence "d_IN (capacity ?Δ) x = 0" by(simp add: d_IN_def source_in)
      thus ?thesis by simp
    next
      case IN
      have "d_IN (capacity ?Δ) x =
        (∑+ y. (capacity Δ (y, x) - f (y, x)) * indicator E (y, x) +
              (if x = source Δ then 0 else f (x, y) * indicator E (x, y)))"
        using flowD_outside[OF f] unfolding d_IN_def
        by(auto intro!: nn_integral_cong split: split_indicator dest: wf_residual_networkD[OF wf])
      also have "… = (∑+ y. (capacity Δ (y, x) - f (y, x)) * indicator E (y, x)) +
        (∑+ y. (if x = source Δ then 0 else f (x, y) * indicator E (x, y)))"
        (is "_ = ?in + ?out")
        by(subst nn_integral_add)(auto simp add: AE_count_space split: split_indicator intro!: flowD_capacity[OF f])
      also have "… ≤ d_IN (capacity Δ) x + (if x = source Δ then 0 else d_OUT f x)" (is "_ ≤ ?in + ?rest")
        unfolding d_IN_def d_OUT_def
        by(rule add_mono)(auto intro!: nn_integral_mono split: split_indicator simp add: nn_integral_0_iff_AE AE_count_space intro!: diff_le_self_ennreal)
      also consider (source) "x = source Δ" | (inner) "x ≠ source Δ" "x ≠ sink Δ" using sink by auto
      then have "?rest < ⊤"
      proof cases
        case inner
        show ?thesis using inner flowD_finite_OUT[OF f inner] by (simp add: less_top)
      qed simp
      ultimately show ?thesis using IN sink by (auto simp: less_top[symmetric] top_unique)
    next
      case OUT
      have "d_OUT (capacity ?Δ) x =
        (∑+ y. (capacity Δ (x, y) - f (x, y)) * indicator E (x, y) +
              (if y = source Δ then 0 else f (y, x) * indicator E (y, x)))"
        using flowD_outside[OF f] unfolding d_OUT_def
        by(auto intro!: nn_integral_cong split: split_indicator dest: wf_residual_networkD[OF wf] simp add: source_in)
      also have "… = (∑+ y. (capacity Δ (x, y) - f (x, y)) * indicator E (x, y)) +
        (∑+ y. (if y = source Δ then 0 else f (y, x) * indicator E (y, x)))"
        (is "_ = ?in + ?out")
        by(subst nn_integral_add)(auto simp add: AE_count_space split: split_indicator intro!: flowD_capacity[OF f])
      also have "… ≤ d_OUT (capacity Δ) x + d_IN f x" (is "_ ≤ ?out + ?rest")
        unfolding d_IN_def d_OUT_def
        by(rule add_mono)(auto intro!: nn_integral_mono split: split_indicator simp add: nn_integral_0_iff_AE AE_count_space intro!: diff_le_self_ennreal)
      also have "?rest = d_OUT f x" using flowD_KIR[OF f OUT(1)] sink by simp
      also have "?out + … ≤ ?out + ?out" by(intro add_left_mono d_OUT_mono flowD_capacity[OF f])
      finally show ?thesis using OUT by (auto simp: top_unique)
    qed
  next
    show "¬ edge ?Δ x x" for x by(simp add: no_loop)
    show "¬ edge ?Δ x (source ?Δ)" for x by(simp add: source_in)
  qed
qed

end

definition plus_flow :: "('v, 'more) graph_scheme ⇒ 'v flow ⇒ 'v flow ⇒ 'v flow" (infixr "⊕ı" 65)
where "plus_flow G f g = (λ(x, y). if edge G x y then f (x, y) + g (x, y) - g (y, x) else 0)"

lemma plus_flow_simps [simp]: fixes G (structure) shows
  "(f ⊕ g) (x, y) = (if edge G x y then f (x, y) + g (x, y) - g (y, x) else 0)"
by(simp add: plus_flow_def)

lemma plus_flow_outside: fixes G (structure) shows "e ∉ E ⟹ (f ⊕ g) e = 0"
by(cases e) simp

lemma
  fixes Δ (structure)
  assumes f_outside: "⋀e. e ∉ E ⟹ f e = 0"
  and g_le_f: "⋀x y. edge Δ x y ⟹ g (y, x) ≤ f (x, y)"
  shows OUT_plus_flow: "d_IN g x ≠ top ⟹ d_OUT (f ⊕ g) x = d_OUT f x + (∑+ y∈UNIV. g (x, y) * indicator E (x, y)) - (∑+ y. g (y, x) * indicator E (x, y))"
    (is "_ ⟹ ?OUT" is "_ ⟹ _ = _ + ?g_out - ?g_out'")
  and IN_plus_flow: "d_OUT g x ≠ top ⟹ d_IN (f ⊕ g) x = d_IN f x + (∑+ y∈UNIV. g (y, x) * indicator E (y, x)) - (∑+ y. g (x, y) * indicator E (y, x))"
    (is "_ ⟹ ?IN" is "_ ⟹ _ = _ + ?g_in - ?g_in'")
proof -
  assume "d_IN g x ≠ top"
  then have finite1: "(∑+ y. g (y, x) * indicator A (f y)) ≠ top" for A f
    by(rule neq_top_trans)(auto split: split_indicator simp add: d_IN_def intro!: nn_integral_mono)

  have "d_OUT (f ⊕ g) x = (∑+ y. (g (x, y) + (f (x, y) - g (y, x))) * indicator E (x, y))"
    unfolding d_OUT_def by(rule nn_integral_cong)(simp split: split_indicator add: add_diff_eq_ennreal add.commute ennreal_diff_add_assoc g_le_f)
  also have "… = ?g_out + (∑+ y. (f (x, y) - g (y, x)) * indicator E (x, y))"
    (is "_ = _ + ?rest")
    by(subst nn_integral_add[symmetric])(auto simp add: AE_count_space g_le_f split: split_indicator intro!: nn_integral_cong)
  also have "?rest = (∑+ y. f (x, y) * indicator E (x, y)) - ?g_out'" (is "_ = ?f - _")
    apply(subst nn_integral_diff[symmetric])
    apply(auto intro!: nn_integral_cong split: split_indicator simp add: AE_count_space g_le_f finite1)
    done
  also have "?f = d_OUT f x" unfolding d_OUT_def using f_outside
    by(auto intro!: nn_integral_cong split: split_indicator)
  also have "(∑+ y. g (x, y) * indicator E (x, y)) + (d_OUT f x - (∑+ y. g (y, x) * indicator E (x, y))) =
     d_OUT f x + ?g_out - ?g_out'"
     by (subst ennreal_diff_add_assoc[symmetric])
        (auto simp: ac_simps d_OUT_def intro!: nn_integral_mono g_le_f split: split_indicator)
  finally show ?OUT .
next
  assume "d_OUT g x ≠ top"
  then have finite2: "(∑+ y. g (x, y) * indicator A (f y)) ≠ top" for A f
    by(rule neq_top_trans)(auto split: split_indicator simp add: d_OUT_def intro!: nn_integral_mono)

  have "d_IN (f ⊕ g) x = (∑+ y. (g (y, x) + (f (y, x) - g (x, y))) * indicator E (y, x))"
    unfolding d_IN_def by(rule nn_integral_cong)(simp split: split_indicator add: add_diff_eq_ennreal add.commute ennreal_diff_add_assoc g_le_f)
  also have "… = (∑+ y∈UNIV. g (y, x) * indicator E (y, x)) + (∑+ y. (f (y, x) - g (x, y)) * indicator E (y, x))"
    (is "_ = _ + ?rest")
    by(subst nn_integral_add[symmetric])(auto simp add: AE_count_space g_le_f split: split_indicator intro!: nn_integral_cong)
  also have "?rest = (∑+ y. f (y, x) * indicator E (y, x))- ?g_in'"
    by(subst nn_integral_diff[symmetric])(auto intro!: nn_integral_cong split: split_indicator simp add: add_ac add_diff_eq_ennreal AE_count_space g_le_f finite2)
  also have "(∑+ y. f (y, x) * indicator E (y, x)) = d_IN f x"
    unfolding d_IN_def using f_outside by(auto intro!: nn_integral_cong split: split_indicator)
  also have "(∑+ y. g (y, x) * indicator E (y, x)) + (d_IN f x - (∑+ y. g (x, y) * indicator E (y, x))) =
     d_IN f x + ?g_in - ?g_in'"
     by (subst ennreal_diff_add_assoc[symmetric])
        (auto simp: ac_simps d_IN_def intro!: nn_integral_mono g_le_f split: split_indicator)
  finally show ?IN .
qed

context countable_network begin

lemma d_IN_plus_flow:
  assumes wf: "wf_residual_network"
  and f: "flow Δ f"
  and g: "flow (residual_network f) g"
  shows "d_IN (f ⊕ g) x ≤ d_IN f x + d_IN g x"
proof -
  have "d_IN (f ⊕ g) x ≤ (∑+ y. f (y, x) + g (y, x))" unfolding d_IN_def
    by(rule nn_integral_mono)(auto intro: diff_le_self_ennreal)
  also have "… = d_IN f x + d_IN g x"
    by(subst nn_integral_add)(simp_all add: d_IN_def)
  finally show ?thesis .
qed

lemma scale_flow:
  assumes f: "flow Δ f"
  and c: "c ≤ 1"
  shows "flow Δ (λe. c * f e)"
proof(intro flow.intros)
  fix e
  from c have "c * f e ≤ 1 * f e" by(rule mult_right_mono) simp
  also have "… ≤ capacity Δ e" using flowD_capacity[OF f, of e] by simp
  finally show "c * f e ≤ …" .
next
  fix x
  assume x: "x ≠ source Δ" "x ≠ sink Δ"
  have "d_OUT (λe. c * f e) x = c * d_OUT f x" by(simp add: d_OUT_cmult)
  also have "d_OUT f x = d_IN f x" using f x by(rule flowD_KIR)
  also have "c * … = d_IN (λe. c * f e) x" by(simp add: d_IN_cmult)
  finally show "KIR (λe. c * f e) x" .
qed

lemma value_scale_flow:
  "value_flow Δ (λe. c * f e) = c * value_flow Δ f"
by(rule d_OUT_cmult)

lemma value_flow:
  assumes f: "flow Δ f"
  and source_out: "⋀y. edge Δ (source Δ) y ⟷ y = x"
  shows "value_flow Δ f = f (source Δ, x)"
proof -
  have "value_flow Δ f = (∑+ y∈OUT (source Δ). f (source Δ, y))"
    by(rule d_OUT_alt_def)(simp add: flowD_outside[OF f])
  also have "… = (∑+ y. f (source Δ, y) * indicator {x} y)"
    by(subst nn_integral_count_space_indicator)(auto intro!: nn_integral_cong split: split_indicator simp add: outgoing_def source_out)
  also have "… = f (source Δ, x)" by(simp add: one_ennreal_def[symmetric] max_def)
  finally show ?thesis .
qed

end

context flow_attainability begin

lemma value_plus_flow:
  assumes wf: "wf_residual_network"
  and f: "flow Δ f"
  and g: "flow (residual_network f) g"
  shows "value_flow Δ (f ⊕ g) = value_flow Δ f + value_flow Δ g"
proof -
  interpret RES: countable_network "residual_network f" using wf f by(rule residual_countable_network)
  have "value_flow Δ (f ⊕ g) = (∑+ y. f (source Δ, y) + g (source Δ, y))"
    unfolding d_OUT_def by(rule nn_integral_cong)(simp add: flowD_outside[OF f] RES.flowD_outside[OF g] source_in)
  also have "… = value_flow Δ f + value_flow Δ g" unfolding d_OUT_def
    by(rule nn_integral_add) simp_all
  finally show ?thesis .
qed

lemma flow_residual_add: -- ‹Lemma 5.3›
  assumes wf: "wf_residual_network"
  and f: "flow Δ f"
  and g: "flow (residual_network f) g"
  shows "flow Δ (f ⊕ g)"
proof
  fix e
  { assume e: "e ∈ E"
    hence "(f ⊕ g) e = f e + g e - g (prod.swap e)" by(cases e) simp
    also have "… ≤ f e + g e - 0" by(rule ennreal_minus_mono) simp_all
    also have "… ≤ f e + (capacity Δ e - f e)"
      using e flowD_capacity[OF g, of e] by(simp split: prod.split_asm add: add_mono)
    also have "… = capacity Δ e" using flowD_capacity[OF f, of e]
      by simp
    also note calculation }
  thus "(f ⊕ g) e ≤ capacity Δ e" by(cases e) auto
next
  fix x
  assume x: "x ≠ source Δ" "x ≠ sink Δ"
  have g_le_f: "g (y, x) ≤ f (x, y)" if "edge Δ x y" for x y
    using that flowD_capacity[OF g, of "(y, x)"]
    by(auto split: if_split_asm dest: wf_residual_networkD[OF wf] elim: order_trans)

  interpret RES: flow_attainability "residual_network f" using wf f by(rule residual_flow_attainability)

  have finite1: "(∑+ y. g (y, x) * indicator A (f y)) ≠ ⊤" for A f
    using RES.flowD_finite_IN[OF g, of x]
    by(rule neq_top_trans)(auto simp add: x d_IN_def split: split_indicator intro: nn_integral_mono)
  have finite2: "(∑+ y. g (x, y) * indicator A (f y)) ≠ ⊤" for A f
    using RES.flowD_finite_OUT[OF g, of x]
    by(rule neq_top_trans)(auto simp add: x d_OUT_def split: split_indicator intro: nn_integral_mono)

  have "d_OUT (f ⊕ g) x = d_OUT f x + (∑+ y. g (x, y) * indicator E (x, y)) - (∑+ y. g (y, x) * indicator E (x, y))"
    (is "_ = ?f + ?g_out - ?g_in")
    using flowD_outside[OF f] g_le_f RES.flowD_finite_IN[OF g, of x]
    by(rule OUT_plus_flow)(simp_all add: x)
  also have "?f = d_IN f x" using f x by(auto dest: flowD_KIR)
  also have "?g_out = (∑+ y. g (x, y) * indicator (- E) (y, x))"
  proof -
    have "(∑+ y. g (x, y) * indicator (- E) (y, x)) =
          (∑+ y. g (x, y) * indicator E (x, y)) + (∑+ y. g (x, y) * indicator (- E) (x, y) * indicator (- E) (y, x))"
      by(subst nn_integral_add[symmetric])(auto simp add: AE_count_space dest: wf_residual_networkD[OF wf] split: split_indicator intro!: nn_integral_cong)
    also have "(∑+ y. g (x, y) * indicator (- E) (x, y) * indicator (- E) (y, x)) = 0"
      using RES.flowD_outside[OF g]
      by(auto simp add: nn_integral_0_iff_AE AE_count_space split: split_indicator)
    finally show ?thesis by simp
  qed
  also have "… = (∑+ y. g (x, y) - g (x, y) * indicator E (y, x))"
    by(rule nn_integral_cong)(simp split: split_indicator add: RES.flowD_finite[OF g])
  also have "… = d_OUT g x - (∑+ y. g (x, y) * indicator E (y, x))"
    (is "_ = _ - ?g_in_E") unfolding d_OUT_def
    by(subst nn_integral_diff)(simp_all add: AE_count_space finite2 split: split_indicator)
  also have "d_IN f x + (d_OUT g x - (∑+ y. g (x, y) * indicator E (y, x))) - ?g_in =
    ((d_IN f x + d_OUT g x) - (∑+ y. g (x, y) * indicator E (y, x))) - ?g_in"
    by (subst add_diff_eq_ennreal) (auto simp: d_OUT_def intro!: nn_integral_mono split: split_indicator)
  also have "d_OUT g x = d_IN g x" using x g by(auto dest: flowD_KIR)
  also have "… = (∑+ y∈UNIV. g (y, x) * indicator (- E) (y, x)) + (∑+ y. g (y, x) * indicator E (y, x))"
    (is "_ = ?x + ?g_in_E'")
    by(subst nn_integral_add[symmetric])(auto intro!: nn_integral_cong simp add: d_IN_def AE_count_space split: split_indicator)
  also have "?x = ?g_in"
  proof -
    have "?x = (∑+ y. g (y, x) * indicator (- E) (x, y) * indicator (- E) (y, x)) + ?g_in"
      by(subst nn_integral_add[symmetric])(auto simp add: AE_count_space  dest: wf_residual_networkD[OF wf] split: split_indicator intro!: nn_integral_cong)
    also have "(∑+ y. g (y, x) * indicator (- E) (x, y) * indicator (- E) (y, x)) = 0"
      using RES.flowD_outside[OF g]
      by(auto simp add: nn_integral_0_iff_AE AE_count_space split: split_indicator)
    finally show ?thesis by simp
  qed
  also have "(d_IN f x + (?g_in + ?g_in_E') - ?g_in_E) - ?g_in =
    d_IN f x + ?g_in_E' + ?g_in - ?g_in - ?g_in_E"
    by (subst diff_diff_commute_ennreal) (simp add: ac_simps)
  also have "… = d_IN f x + ?g_in_E' - ?g_in_E"
    by (subst ennreal_add_diff_cancel_right) (simp_all add: finite1)
  also have "… = d_IN (f ⊕ g) x"
    using flowD_outside[OF f]  g_le_f RES.flowD_finite_OUT[OF g, of x]
    by(rule IN_plus_flow[symmetric])(simp_all add: x)
  finally show "KIR (f ⊕ g) x" by simp
qed

definition minus_flow :: "'v flow ⇒ 'v flow ⇒ 'v flow" (infixl "⊖" 65)
where
  "f ⊖ g = (λ(x, y). if edge Δ x y then f (x, y) - g (x, y) else if edge Δ y x then g (y, x) - f (y, x) else 0)"

lemma minus_flow_simps [simp]:
  "(f ⊖ g) (x, y) = (if edge Δ x y then f (x, y) - g (x, y) else if edge Δ y x then g (y, x) - f (y, x) else 0)"
by(simp add: minus_flow_def)

lemma minus_flow:
  assumes wf: "wf_residual_network"
  and f: "flow Δ f"
  and g: "flow Δ g"
  and value_le: "value_flow Δ g ≤ value_flow Δ f"
  and f_finite: "f (source Δ, x) ≠ ⊤"
  and source_out: "⋀y. edge Δ (source Δ) y ⟷ y = x"
  shows "flow (residual_network g) (f ⊖ g)" (is "flow ?Δ ?f")
proof
  show "?f e ≤ capacity ?Δ e" for e
    using value_le f_finite flowD_capacity[OF g] flowD_capacity[OF f]
    by(cases e)(auto simp add: source_in source_out value_flow[OF f source_out] value_flow[OF g source_out] less_top
                     intro!: diff_le_self_ennreal diff_eq_0_ennreal ennreal_minus_mono)

  fix x
  assume "x ≠ source ?Δ" "x ≠ sink ?Δ"
  hence x: "x ≠ source Δ" "x ≠ sink Δ" by simp_all

  have finite_f_in: "(∑+ y. f (y, x) * indicator A y) ≠ top" for A
    using flowD_finite_IN[OF f, of x]
    by(rule neq_top_trans)(auto simp add: x d_IN_def split: split_indicator intro!: nn_integral_mono)
  have finite_f_out: "(∑+ y. f (x, y) * indicator A y) ≠ top" for A
    using flowD_finite_OUT[OF f, of x]
    by(rule neq_top_trans)(auto simp add: x d_OUT_def split: split_indicator intro!: nn_integral_mono)
  have finite_f[simp]: "f (x, y) ≠ top" "f (y, x) ≠ top" for y
    using finite_f_in[of "{y}"] finite_f_out[of "{y}"] by auto

  have finite_g_in: "(∑+ y. g (y, x) * indicator A y) ≠ top" for A
    using flowD_finite_IN[OF g, of x]
    by(rule neq_top_trans)(auto simp add: x d_IN_def split: split_indicator intro!: nn_integral_mono)
  have finite_g_out: "(∑+ y. g (x, y) * indicator A y) ≠ top" for A
    using flowD_finite_OUT[OF g x]
    by(rule neq_top_trans)(auto simp add: x d_OUT_def split: split_indicator intro!: nn_integral_mono)
  have finite_g[simp]: "g (x, y) ≠ top" "g (y, x) ≠ top" for y
    using finite_g_in[of "{y}"] finite_g_out[of "{y}"] by auto

  have "d_OUT (f ⊖ g) x = (∑+ y. (f (x, y) - g (x, y)) * indicator E (x, y) * indicator {y. g (x, y) ≤ f (x, y)} y) +
    (∑+ y. (g (y, x) - f (y, x)) * indicator E (y, x) * indicator {y. f (y, x) < g (y, x)} y)"
    (is "_ = ?out + ?in" is "_ = (∑+ y∈_. _ * ?f_ge_g y) + (∑+ y∈_. _ * ?g_gt_f y)")
    using flowD_finite[OF g]
    apply(subst nn_integral_add[symmetric])
    apply(auto 4 4 simp add: d_OUT_def not_le less_top[symmetric] intro!: nn_integral_cong
           dest!: wf_residual_networkD[OF wf] split: split_indicator intro!: diff_eq_0_ennreal)
    done
  also have "?in = (∑+ y. (g (y, x) - f (y, x)) * ?g_gt_f y)"
    using flowD_outside[OF f] flowD_outside[OF g] by(auto intro!: nn_integral_cong split: split_indicator)
  also have "… = (∑+ y∈UNIV. g (y, x) * ?g_gt_f y) - (∑+ y. f (y, x) * ?g_gt_f y)" (is "_ = ?g_in - ?f_in")
    using finite_f_in
    by(subst nn_integral_diff[symmetric])(auto simp add: AE_count_space split: split_indicator intro!: nn_integral_cong)
  also have "?out = (∑+ y. (f (x, y) - g (x, y)) * ?f_ge_g y)"
    using flowD_outside[OF f] flowD_outside[OF g] by(auto intro!: nn_integral_cong split: split_indicator)
  also have "… = (∑+ y. f (x, y) * ?f_ge_g y) - (∑+ y. g (x, y) * ?f_ge_g y)" (is "_ = ?f_out - ?g_out")
    using finite_g_out
    by(subst nn_integral_diff[symmetric])(auto simp add: AE_count_space split: split_indicator intro!: nn_integral_cong)
  also have "?f_out = d_OUT f x - (∑+ y. f (x, y) * indicator {y. f (x, y) < g (x, y)} y)" (is "_ = _ - ?f_out_less")
    unfolding d_OUT_def using flowD_finite[OF f] using finite_f_out
    by(subst nn_integral_diff[symmetric])(auto split: split_indicator intro!: nn_integral_cong)
  also have "?g_out = d_OUT g x - (∑+ y. g (x, y) * indicator {y. f (x, y) < g (x, y)} y)" (is "_ = _ - ?g_less_f")
    unfolding d_OUT_def using flowD_finite[OF g] finite_g_out
    by(subst nn_integral_diff[symmetric])(auto split: split_indicator intro!: nn_integral_cong)
  also have "d_OUT f x - ?f_out_less - (d_OUT g x - ?g_less_f) + (?g_in - ?f_in) =
    (?g_less_f + (d_OUT f x - ?f_out_less)) - d_OUT g x + (?g_in - ?f_in)"
    by (subst diff_diff_ennreal')
       (auto simp: ac_simps d_OUT_def nn_integral_diff[symmetric] finite_g_out finite_f_out intro!: nn_integral_mono split: split_indicator )
  also have "… = ?g_less_f + d_OUT f x - ?f_out_less - d_OUT g x + (?g_in - ?f_in)"
    by (subst add_diff_eq_ennreal)
       (auto simp: d_OUT_def intro!: nn_integral_mono split: split_indicator)
  also have "… = d_OUT f x + ?g_less_f - ?f_out_less - d_OUT g x + (?g_in - ?f_in)"
    by (simp add: ac_simps)
  also have "… = d_OUT f x + (?g_less_f - ?f_out_less) - d_OUT g x + (?g_in - ?f_in)"
    by (subst add_diff_eq_ennreal[symmetric])
       (auto intro!: nn_integral_mono split: split_indicator)
  also have "… = (?g_in - ?f_in) + ((?g_less_f - ?f_out_less) + d_OUT f x - d_OUT g x)"
    by (simp add: ac_simps)
  also have "… = ((?g_in - ?f_in) + ((?g_less_f - ?f_out_less) + d_OUT f x)) - d_OUT g x"
    apply (subst add_diff_eq_ennreal)
    apply (simp_all add: d_OUT_def)
    apply (subst nn_integral_diff[symmetric])
    apply (auto simp: AE_count_space finite_f_out nn_integral_add[symmetric] not_less diff_add_cancel_ennreal intro!: nn_integral_mono split: split_indicator)
    done
  also have "… = ((?g_less_f - ?f_out_less) + (d_OUT f x + (?g_in - ?f_in))) - d_OUT g x"
    by (simp add: ac_simps)
  also have "… = ((?g_less_f - ?f_out_less) + (d_IN f x + (?g_in - ?f_in))) - d_IN g x"
    unfolding flowD_KIR[OF f x] flowD_KIR[OF g x] ..
  also have "… = (?g_less_f - ?f_out_less) + ((d_IN f x + (?g_in - ?f_in)) - d_IN g x)"
    apply (subst (2) add_diff_eq_ennreal)
    apply (simp_all add: d_IN_def)
    apply (subst nn_integral_diff[symmetric])
    apply (auto simp: AE_count_space finite_f_in finite_f_out nn_integral_add[symmetric] not_less  ennreal_ineq_diff_add[symmetric]
                intro!: nn_integral_mono split: split_indicator)
    done
  also have "… = (?g_less_f - ?f_out_less) + (d_IN f x + ?g_in - d_IN g x - ?f_in)"
    by (subst (2) add_diff_eq_ennreal) (auto intro!: nn_integral_mono split: split_indicator simp: diff_diff_commute_ennreal)
  also have "… = (?g_less_f - ?f_out_less) + (d_IN f x - (d_IN g x - ?g_in) - ?f_in)"
    apply (subst diff_diff_ennreal')
    apply (auto simp: d_IN_def intro!: nn_integral_mono split: split_indicator)
    apply (subst nn_integral_diff[symmetric])
    apply (auto simp: AE_count_space finite_g_in intro!: nn_integral_mono split: split_indicator)
    done
  also have "… =(d_IN f x - ?f_in) - (d_IN g x - ?g_in) + (?g_less_f - ?f_out_less)"
    by (simp add: ac_simps diff_diff_commute_ennreal)
  also have "?g_less_f - ?f_out_less = (∑+ y. (g (x, y) - f (x, y)) * indicator {y. f (x, y) < g (x, y)} y)" using finite_f_out
    by(subst nn_integral_diff[symmetric])(auto simp add: AE_count_space split: split_indicator intro!: nn_integral_cong)
  also have "… = (∑+ y. (g (x, y) - f (x, y)) * indicator E (x, y) * indicator {y. f (x, y) < g (x, y)} y)" (is "_ = ?diff_out")
    using flowD_outside[OF f] flowD_outside[OF g] by(auto intro!: nn_integral_cong split: split_indicator)
  also have "d_IN f x - ?f_in = (∑+ y. f (y, x) * indicator {y. g (y, x) ≤ f (y, x)} y)"
    unfolding d_IN_def using finite_f_in
    apply(subst nn_integral_diff[symmetric])
    apply(auto simp add: AE_count_space split: split_indicator intro!: nn_integral_cong)
    done
  also have "d_IN g x - ?g_in = (∑+ y. g (y, x) * indicator {y. g (y, x) ≤ f (y, x)} y)"
    unfolding d_IN_def using finite_g_in
    by(subst nn_integral_diff[symmetric])(auto simp add: flowD_finite[OF g] AE_count_space split: split_indicator intro!: nn_integral_cong)
  also have "(∑+ y∈UNIV. f (y, x) * indicator {y. g (y, x) ≤ f (y, x)} y) - … = (∑+ y. (f (y, x) - g (y, x)) * indicator {y. g (y, x) ≤ f (y, x)} y)"
    using finite_g_in
    by(subst nn_integral_diff[symmetric])(auto simp add: flowD_finite[OF g] AE_count_space split: split_indicator intro!: nn_integral_cong)
  also have "… = (∑+ y. (f (y, x) - g (y, x)) * indicator E (y, x) * indicator {y. g (y, x) ≤ f (y, x)} y)"
    using flowD_outside[OF f] flowD_outside[OF g] by(auto intro!: nn_integral_cong split: split_indicator)
  also have "… + ?diff_out = d_IN ?f x"
    using flowD_finite[OF g]
    apply(subst nn_integral_add[symmetric])
    apply(auto 4 4 simp add: d_IN_def not_le less_top[symmetric] intro!: nn_integral_cong
               dest!: wf_residual_networkD[OF wf] split: split_indicator intro: diff_eq_0_ennreal)
    done
  finally show "KIR ?f x" .
qed

lemma value_minus_flow:
  assumes f: "flow Δ f"
  and g: "flow Δ g"
  and value_le: "value_flow Δ g ≤ value_flow Δ f"
  and source_out: "⋀y. edge Δ (source Δ) y ⟷ y = x"
  shows "value_flow Δ (f ⊖ g) = value_flow Δ f - value_flow Δ g" (is "?value")
proof -
  have "value_flow Δ (f ⊖ g) = (∑+ y∈OUT (source Δ). (f ⊖ g) (source Δ, y))"
    by(subst d_OUT_alt_def)(auto simp add: flowD_outside[OF f] flowD_outside[OF g] source_in)
  also have "… = (∑+ y. (f (source Δ, y) - g (source Δ, y)) * indicator {x} y)"
    by(subst nn_integral_count_space_indicator)(auto intro!: nn_integral_cong split: split_indicator simp add: outgoing_def source_out)
  also have "… = f (source Δ, x) - g (source Δ, x)"
    using value_le value_flow[OF f source_out] value_flow[OF g source_out]
    by(auto simp add: one_ennreal_def[symmetric] max_def not_le intro: antisym)
  also have "… = value_flow Δ f - value_flow Δ g" using f g source_out by(simp add: value_flow)
  finally show ?value .
qed

context
  fixes α
  defines "α ≡ (SUP g:{g. flow Δ g}. value_flow Δ g)"
begin

lemma flow_by_value:
  assumes "v < α"
  and real[rule_format]: "∀f. α = ⊤ ⟶ flow Δ f ⟶ value_flow Δ f < α"
  obtains f where "flow Δ f" "value_flow Δ f = v"
proof -
  have α_pos: "α > 0" using assms by (auto simp add: zero_less_iff_neq_zero)
  from ‹v < α› obtain f where f: "flow Δ f" and v: "v < value_flow Δ f"
    unfolding α_def less_SUP_iff by blast
  have [simp]: "value_flow Δ f ≠ ⊤"
  proof
    assume val: "value_flow Δ f = ⊤"
    from f have "value_flow Δ f ≤ α" unfolding α_def by(blast intro: SUP_upper2)
    with val have "α = ⊤" by (simp add: top_unique)
    from real[OF this f] val show False by simp
  qed
  let ?f = "λe. (v / value_flow Δ f) * f e"
  note f
  moreover
  have *: "0 < value_flow Δ f"
    using ‹v < value_flow Δ f› by (auto simp add: zero_less_iff_neq_zero)
  then have "v / value_flow Δ f ≤ 1" using v
    by (auto intro!: divide_le_posI_ennreal)
  ultimately have "flow Δ ?f" by (rule scale_flow)
  moreover {
    have "value_flow Δ ?f = v * (value_flow Δ f / value_flow Δ f)"
      by(subst value_scale_flow)(simp add: divide_ennreal_def ac_simps)
    also have "… = v" using * by(subst ennreal_divide_self) (auto simp: less_top[symmetric])
    also note calculation }
  ultimately show ?thesis by(rule that)
qed

theorem ex_max_flow':
  assumes wf: "wf_residual_network"
  assumes source_out: "⋀y. edge Δ (source Δ) y ⟷ y = x"
  and nontrivial: "V - {source Δ, sink Δ} ≠ {}"
  and real: "α = ennreal α'" and α'_nonneg[simp]: "0 ≤ α'"
  shows "∃f. flow Δ f ∧ value_flow Δ f = α ∧ (∀x. d_IN f x ≤ value_flow Δ f)"
proof -
  have α'_not_neg[simp]: "¬ α' < 0"
    using α'_nonneg by linarith

  let ?v = "λi. (1 - (1 / 2) ^ i) * α"
  let ?v_r = "λi. ennreal ((1 - (1 / 2) ^ i) * α')"
  have v_eq: "?v i = ?v_r i" for i
    by (auto simp: real ennreal_mult power_le_one ennreal_lessI ennreal_minus[symmetric]
                   ennreal_power[symmetric] divide_ennreal_def)
  have "∃f. flow Δ f ∧ value_flow Δ f = ?v i" for i :: nat
  proof(cases "α = 0")
    case True thus ?thesis by(auto intro!: exI[where x="λ_. 0"])
  next
    case False
    then have "?v i < α"
      unfolding v_eq by (auto simp: real field_simps intro!: ennreal_lessI) (simp_all add: less_le)
    then obtain f where "flow Δ f" and "value_flow Δ f = ?v i"
      by(rule flow_by_value)(simp add: real)
    thus ?thesis by blast
  qed
  then obtain f_aux where f_aux: "⋀i. flow Δ (f_aux i)"
    and value_aux: "⋀i. value_flow Δ (f_aux i) = ?v_r i"
    unfolding v_eq by moura

  def f_i  "rec_nat (λ_. 0) (λi f_i.
    let g = f_aux (Suc i) ⊖ f_i;
      k_i = SOME k. k ≤ g ∧ flow (residual_network f_i) k ∧ value_flow (residual_network f_i) k = value_flow (residual_network f_i) g ∧ (∀x. d_IN k x ≤ value_flow (residual_network f_i) k)
    in f_i ⊕ k_i)"
  let ?P = "λi k. k ≤ f_aux (Suc i) ⊖ f_i i ∧ flow (residual_network (f_i i)) k ∧ value_flow (residual_network (f_i i)) k = value_flow (residual_network (f_i i)) (f_aux (Suc i) ⊖ f_i i) ∧ (∀x. d_IN k x ≤ value_flow (residual_network (f_i i)) k)"
  def k_i  "λi. Eps (?P i)"

  have f_i_simps [simp]: "f_i 0 = (λ_. 0)" "f_i (Suc i) = f_i i ⊕ k_i i" for i
    by(simp_all add: f_i_def Let_def k_i_def)

  have k_i: "flow (residual_network (f_i i)) (k_i i)" (is ?k_i)
    and value_k_i: "value_flow (residual_network (f_i i)) (k_i i) = value_flow (residual_network (f_i i)) (f_aux (Suc i) ⊖ f_i i)" (is "?value_k_i")
    and IN_k_i: "d_IN (k_i i) x ≤ value_flow (residual_network (f_i i)) (k_i i)" (is "?IN_k_i")
    and value_diff: "value_flow (residual_network (f_i i)) (f_aux (Suc i) ⊖ f_i i) = value_flow Δ (f_aux (Suc i)) - value_flow Δ (f_i i)" (is "?value_diff")
    if "flow_network Δ (f_i i)" and value_f_i: "value_flow Δ (f_i i) = value_flow Δ (f_aux i)" for i x
  proof -
    let ?RES = "residual_network (f_i i)"
    interpret fn: flow_network Δ "f_i i" by(rule that)
    interpret RES: flow_attainability "?RES" using wf fn.g by(rule residual_flow_attainability)
    have le: "value_flow Δ (f_i i) ≤ value_flow Δ (f_aux (Suc i))"
      unfolding value_aux value_f_i
      unfolding v_eq by (rule ennreal_leI) (auto simp: field_simps)
    with wf f_aux fn.g have res_flow: "flow ?RES (f_aux (Suc i) ⊖ f_i i)"
      using flowD_finite[OF f_aux] source_out
      by(rule minus_flow)
    show value': ?value_diff by(simp add: value_minus_flow[OF f_aux fn.g le source_out])
    also have "… < ⊤"
      unfolding value_aux v_eq by (auto simp: less_top[symmetric])
    finally have "value_flow ?RES (f_aux (Suc i) ⊖ f_i i) ≠ ⊤" by simp
    then have fn': "flow_network ?RES (f_aux (Suc i) ⊖ f_i i)"
      using nontrivial res_flow by(unfold_locales) simp_all
    then interpret fn': flow_network "?RES" "f_aux (Suc i) ⊖ f_i i" .
    from fn'.flow_cleanup show ?k_i ?value_k_i ?IN_k_i unfolding k_i_def by(rule someI2_ex; blast)+
  qed

  have fn_i: "flow_network Δ (f_i i)"
    and value_f_i: "value_flow Δ (f_i i) = value_flow Δ (f_aux i)"
    and d_IN_i: "d_IN (f_i i) x ≤ value_flow Δ (f_i i)"  for i x
  proof(induction i)
    case 0
    { case 1 show ?case using nontrivial by(unfold_locales)(simp_all add: f_aux value_aux) }
    { case 2 show ?case by(simp add: value_aux) }
    { case 3 show ?case by(simp) }
  next
    case (Suc i)
    interpret fn: flow_network Δ "f_i i" using Suc.IH(1) .
    let ?RES = "residual_network (f_i i)"

    have k_i: "flow ?RES (k_i i)"
      and value_k_i: "value_flow ?RES (k_i i) = value_flow ?RES (f_aux (Suc i) ⊖ f_i i)"
      and d_IN_k_i: "d_IN (k_i i) x ≤ value_flow ?RES (k_i i)" for x
      using Suc.IH(1-2) by(rule k_i value_k_i IN_k_i)+

    interpret RES: flow_attainability "?RES" using wf fn.g by(rule residual_flow_attainability)
    have le: "value_flow Δ (f_i i) ≤ value_flow Δ (f_aux (Suc i))"
      unfolding value_aux Suc.IH(2) v_eq using α'_nonneg by(intro ennreal_leI)(simp add: real field_simps)
    { case 1 show ?case unfolding f_i_simps
      proof
        show "flow Δ (f_i i ⊕ k_i i)" using wf fn.g k_i by(rule flow_residual_add)
        with RES.flowD_finite[OF k_i] show "value_flow Δ (f_i i ⊕ k_i i) ≠ ⊤"
          by(simp add: value_flow[OF _ source_out])
      qed(rule nontrivial) }
    from value_k_i have value_k: "value_flow ?RES (k_i i) = value_flow Δ (f_aux (Suc i)) - value_flow Δ (f_aux i)"
      by(simp add: value_minus_flow[OF f_aux fn.g le source_out] Suc.IH)
    { case 2 show ?case using value_k
        by(auto simp add: source_out value_plus_flow[OF wf fn.g k_i] Suc.IH value_aux field_simps intro!: ennreal_leI) }
    note value_f = this
    { case 3
      have "d_IN (f_i i ⊕ k_i i) x ≤ d_IN (f_i i) x + d_IN (k_i i) x"
        using fn.g k_i by(rule d_IN_plus_flow[OF wf])
      also have "… ≤ value_flow Δ (f_i i) + d_IN (k_i i) x" using Suc.IH(3) by(rule add_right_mono)
      also have "… ≤ value_flow Δ (f_i i) + value_flow ?RES (k_i i)" using d_IN_k_i by(rule add_left_mono)
      also have "… = value_flow Δ (f_i (Suc i))" unfolding value_f Suc.IH(2) value_k
        by(auto simp add: value_aux field_simps intro!: ennreal_leI)
      finally show ?case by simp }
  qed
  interpret fn: flow_network Δ "f_i i" for i by(rule fn_i)
  note k_i = k_i[OF fn_i value_f_i] and value_k_i = value_k_i[OF fn_i value_f_i]
    and IN_k_i = IN_k_i[OF fn_i value_f_i] and value_diff = value_diff[OF fn_i value_f_i]

  have "∃x≥0. f_i i e = ennreal x" for i e
    using fn.finite_g[of i e] by (cases "f_i i e") auto
  then obtain f_i' where f_i': "⋀i e. f_i i e = ennreal (f_i' i e)" and [simp]: "⋀i e. 0 ≤ f_i' i e"
    by metis

  { fix i e
    obtain x y :: 'v where e: "e = (x, y)" by(cases e)
    have "k_i i (x, y) ≤ d_IN (k_i i) y" by (rule d_IN_ge_point)
    also have "… ≤ value_flow (residual_network (f_i i)) (k_i i)" by(rule IN_k_i)
    also have "… < ⊤" using value_k_i[of i] value_diff[of i]
      by(simp add: value_k_i value_f_i value_aux real less_top[symmetric])
    finally have "∃x≥0. k_i i e = ennreal x"
      by(cases "k_i i e")(auto simp add: e) }
  then obtain k_i' where k_i': "⋀i e. k_i i e = ennreal (k_i' i e)" and k_i'_nonneg[simp]: "⋀i e. 0 ≤ k_i' i e"
    by metis

  have wf_k: "(x, y) ∈ E ⟹ k_i i (y, x) ≤ f_i i (x, y) + k_i i (x, y)" for i x y
    using flowD_capacity[OF k_i, of i "(y, x)"]
    by (auto split: if_split_asm dest: wf_residual_networkD[OF wf] elim: order_trans)

  have f_i'_0[simp]: "f_i' 0 = (λ_. 0)" using f_i_simps(1) by (simp del: f_i_simps add: fun_eq_iff f_i')

  have f_i'_Suc[simp]: "f_i' (Suc i) e =  (if e ∈ E then f_i' i e + (k_i' i e - k_i' i (prod.swap e)) else 0)" for i e
    using f_i_simps(2)[of i, unfolded fun_eq_iff, THEN spec, of e] wf_k[of "fst e" "snd e" i]
    by (auto simp del: f_i_simps ennreal_plus
             simp add: fun_eq_iff f_i' k_i' ennreal_plus[symmetric] ennreal_minus split: if_split_asm)

  have k_i'_le: "k_i' i e ≤ α' / 2 ^ (Suc i)" for i e
  proof -
    obtain x y where e: "e = (x, y)" by(cases e)
    have "k_i' i (x, y) ≤ d_IN (k_i' i) y" by (rule d_IN_ge_point)
    also have "… ≤ value_flow (residual_network (f_i i)) (k_i' i)"
      using IN_k_i[of i y] by(simp add: k_i'[abs_def])
    also have "… = α' / 2 ^ (Suc i)" using value_k_i[of i] value_diff[of i]
      by(simp add: value_f_i value_aux real k_i'[abs_def] field_simps ennreal_minus mult_le_cancel_left1)
    finally show ?thesis using e by simp
  qed

  have convergent: "convergent (λi. f_i' i e)" for e
  proof(cases "α' > 0")
    case False
    obtain x y where [simp]: "e = (x, y)" by(cases e)

    { fix i
      from False α'_nonneg have "α' = 0" by simp
      moreover have "f_i i (x, y) ≤ d_IN (f_i i) y" by (rule d_IN_ge_point)
      ultimately have "f_i i (x, y) = 0" using d_IN_i[of i y]
        by(simp add: value_f_i value_aux real) }
    thus ?thesis by(simp add: f_i' convergent_const)
  next
    case α'_pos: True
    show ?thesis
    proof(rule real_Cauchy_convergent Cauchy_real_Suc_diff)+
      fix n
      have "¦k_i' n e - k_i' n (prod.swap e)¦ ≤ ¦k_i' n e¦ + ¦k_i' n (prod.swap e)¦"
        by (rule abs_triangle_ineq4)
      then have "¦k_i' n e - k_i' n (prod.swap e)¦ ≤ α' / 2 ^ n"
        using k_i'_le[of n e] k_i'_le[of n "prod.swap e"] by simp
      then have "¦f_i' (Suc n) e - f_i' n e¦ ≤ α' / 2 ^ n"
        using flowD_outside[OF fn.g] by (cases e) (auto simp: f_i')
      thus "¦f_i' (Suc n) e - f_i' n e¦ ≤ α' / 2 ^ n" by simp
    qed simp
  qed
  then obtain f' where f': "⋀e. (λi. f_i' i e) ⇢ f' e" unfolding convergent_def by metis
  hence f: "⋀e. (λi. f_i i e) ⇢ ennreal (f' e)" unfolding f_i' by simp

  have f'_nonneg: "0 ≤ f' e" for e
    by (rule LIMSEQ_le_const[OF f']) auto

  let ?k = "λi x y. (k_i' i (x, y) - k_i' i (y, x)) * indicator E (x, y)"
  have sum_i': "f_i' i (x, y) = (∑j<i. ?k j x y)" for x y i
    by (induction i) auto

  have summable_nk: "summable (λi. ¦?k i x y¦)" for x y
  proof(rule summable_rabs_comparison_test)
    show "∃N. ∀i≥N. ¦?k i x y¦ ≤ α' * (1 / 2) ^ i"
    proof (intro exI allI impI)
      fix i have "¦?k i x y¦ ≤ k_i' i (x, y) + k_i' i (y, x)"
        by (auto intro!: abs_triangle_ineq4[THEN order_trans] split: split_indicator)
      also have "… ≤ α' * (1 / 2) ^ i"
        using k_i'_le[of i "(x, y)"] k_i'_le[of i "(y, x)"] α'_nonneg k_i'_nonneg[of i "(x, y)"] k_i'_nonneg[of i "(y, x)"]
        by(auto simp add: abs_real_def power_divide split: split_indicator)
      finally show "¦?k i x y¦ ≤ α' * (1 / 2) ^ i"
        by simp
    qed
    show "summable (λi. α' * (1 / 2) ^ i)"
      by(rule summable_mult complete_algebra_summable_geometric)+ simp
  qed
  hence summable_k: "summable (λi. ?k i x y)" for x y by(auto intro: summable_norm_cancel)

  have suminf: "(∑i. (k_i' i (x, y) - k_i' i (y, x)) * indicator E (x, y)) = f' (x, y)" for x y
    by(rule LIMSEQ_unique[OF summable_LIMSEQ])(simp_all add: sum_i'[symmetric] f' summable_k)

  have flow: "flow Δ f'"
  proof
    fix e
    have "f' e ≤ Sup {..capacity Δ e}" using _ f
      by(rule ereal_Sup_lim)(simp add: flowD_capacity[OF fn.g])
    then show "f' e ≤ capacity Δ e" by simp
  next
    fix x
    assume x: "x ≠ source Δ" "x ≠ sink Δ"

    have integrable_f_i: "integrable (count_space UNIV) (λy. f_i' i (x, y))" for i
      using flowD_finite_OUT[OF fn.g x, of i] by(auto intro!: integrableI_bounded simp add: f_i' d_OUT_def less_top)
    have integrable_f_i': "integrable (count_space UNIV) (λy. f_i' i (y, x))" for i
      using flowD_finite_IN[OF fn.g, of x i] x by(auto intro!: integrableI_bounded simp add: f_i' d_IN_def less_top)

    have integral_k_bounded: "(∑+ y. norm (?k i x y)) ≤ α' / 2 ^ i" (is ?thesis1)
      and integral_k'_bounded: "(∑+ y. norm (?k i y x)) ≤ α' / 2 ^ i" (is ?thesis2) for i
    proof -
      def b  "∑+ y. k_i i (x, y) + k_i i (y, x)"
      have "b = d_OUT (k_i i) x + d_IN (k_i i) x" unfolding b_def
        by(subst nn_integral_add)(simp_all add: d_OUT_def d_IN_def)
      also have "d_OUT (k_i i) x = d_IN (k_i i) x" using k_i by(rule flowD_KIR)(simp_all add: x)
      also have "… + … ≤ value_flow Δ (k_i i) + value_flow Δ (k_i i)"
        using IN_k_i[of i x, simplified] by-(rule add_mono)
      also have "… ≤ α' / 2 ^ i" using value_k_i[of i] value_diff[of i]
        by(simp add: value_aux value_f_i  field_simps ennreal_minus_if ennreal_plus_if mult_le_cancel_left1
                del: ennreal_plus)
      also have "(∑+ y∈UNIV. norm (?k i x y)) ≤ b" and "(∑+ y. norm (?k i y x)) ≤ b" unfolding b_def
        by(rule nn_integral_mono; simp add: abs_real_def k_i' ennreal_plus_if del:  ennreal_plus split: split_indicator)+
      ultimately show ?thesis1 ?thesis2 by(auto)
    qed

    have integrable_k: "integrable (count_space UNIV) (λy. ?k i x y)"
      and integrable_k': "integrable (count_space UNIV) (λy. ?k i y x)" for i
      using integral_k_bounded[of i] integral_k'_bounded[of i] real
      by(auto intro!: integrableI_bounded simp: less_top[symmetric] top_unique ennreal_divide_eq_top_iff)

    have summable'_k: "summable (λi. ∫ y. ¦?k i x y¦ ∂count_space UNIV)"
    proof(rule summable_comparison_test)
      have "¦∫ y. ¦?k i x y¦ ∂count_space UNIV¦ ≤ α' * (1 / 2) ^ i" for i
        using integral_norm_bound_ennreal[OF integrable_norm, OF integrable_k, of i] integral_k_bounded[of i]
        by(bestsimp simp add: real power_divide dest: order_trans)
      thus "∃N. ∀i≥N. norm (∫ y. ¦?k i x y¦ ∂count_space UNIV) ≤ α' * (1 / 2) ^ i"
        by auto
      show "summable (λi. α' * (1 / 2) ^ i)"
        by(rule summable_mult complete_algebra_summable_geometric)+ simp
    qed
    have summable'_k': "summable (λi. ∫ y. ¦?k i y x¦ ∂count_space UNIV)"
    proof(rule summable_comparison_test)
      have "¦∫ y. ¦?k i y x¦ ∂count_space UNIV¦ ≤ α' * (1 / 2) ^ i" for i
        using integral_norm_bound_ennreal[OF integrable_norm, OF integrable_k', of i] integral_k'_bounded[of i]
        by(bestsimp simp add: real power_divide dest: order_trans)
      thus "∃N. ∀i≥N. norm (∫ y. ¦?k i y x¦ ∂count_space UNIV) ≤ α' * (1 / 2) ^ i" by auto
      show "summable (λi. α' * (1 / 2) ^ i)"
        by(rule summable_mult complete_algebra_summable_geometric)+ simp
    qed

    have "(λi. ∫ y. ?k i x y ∂count_space UNIV) sums ∫ y. (∑i. ?k i x y) ∂count_space UNIV"
      using integrable_k by(rule sums_integral)(simp_all add: summable_nk summable'_k)
    also have "… = ∫ y. f' (x, y) ∂count_space UNIV" by(rule Bochner_Integration.integral_cong[OF refl])(rule suminf)
    finally have "(λi. ∑j<i. ∫ y. ?k j x y ∂count_space UNIV) ⇢ …" unfolding sums_def .
    also have "(λi. ∑j<i. ∫ y. ?k j x y ∂count_space UNIV) = (λi. ∫ y. f_i' i (x, y) ∂count_space UNIV)"
      unfolding sum_i' by(rule ext Bochner_Integration.integral_sum[symmetric] integrable_k)+
    finally have "(λi. ennreal (∫ y. f_i' i (x, y) ∂count_space UNIV)) ⇢ ennreal (∫ y. f' (x, y) ∂count_space UNIV)" by simp
    also have "(λi. ennreal (∫ y. f_i' i (x, y) ∂count_space UNIV)) = (λi. d_OUT (f_i i) x)"
      unfolding d_OUT_def f_i' by(rule ext nn_integral_eq_integral[symmetric] integrable_f_i)+ simp
    also have "ennreal (∫ y. f' (x, y) ∂count_space UNIV) = d_OUT f' x"
      unfolding d_OUT_def by(rule nn_integral_eq_integral[symmetric])(simp_all add: f'_nonneg, simp add: suminf[symmetric] integrable_suminf integrable_k summable_nk summable'_k)
    also have "(λi. d_OUT (f_i i) x) = (λi. d_IN (f_i i) x)"
      using flowD_KIR[OF fn.g x] by(simp)
    finally have *: "(λi. d_IN (f_i i) x) ⇢ d_OUT (λx. ennreal (f' x)) x" .

    have "(λi. ∫ y. ?k i y x ∂count_space UNIV) sums ∫ y. (∑i. ?k i y x) ∂count_space UNIV"
      using integrable_k' by(rule sums_integral)(simp_all add: summable_nk summable'_k')
    also have "… = ∫ y. f' (y, x) ∂count_space UNIV" by(rule Bochner_Integration.integral_cong[OF refl])(rule suminf)
    finally have "(λi. ∑j<i. ∫ y. ?k j y x ∂count_space UNIV) ⇢ …" unfolding sums_def .
    also have "(λi. ∑j<i. ∫ y. ?k j y x ∂count_space UNIV) = (λi. ∫ y. f_i' i (y, x) ∂count_space UNIV)"
      unfolding sum_i' by(rule ext Bochner_Integration.integral_sum[symmetric] integrable_k')+
    finally have "(λi. ennreal (∫ y. f_i' i (y, x) ∂count_space UNIV)) ⇢ ennreal (∫ y. f' (y, x) ∂count_space UNIV)" by simp
    also have "(λi. ennreal (∫ y. f_i' i (y, x) ∂count_space UNIV)) = (λi. d_IN (f_i i) x)"
      unfolding d_IN_def f_i' by(rule ext nn_integral_eq_integral[symmetric] integrable_f_i')+ simp
    also have "ennreal (∫ y. f' (y, x) ∂count_space UNIV) = d_IN f' x"
      unfolding d_IN_def by(rule nn_integral_eq_integral[symmetric])(simp_all add: f'_nonneg, simp add: suminf[symmetric] integrable_suminf integrable_k' summable_nk summable'_k')
    finally show "d_OUT f' x = d_IN f' x" using * by(blast intro: LIMSEQ_unique)
  qed
  moreover
  { have "incseq (λi. value_flow Δ (f_i i))"
      by(rule incseq_SucI)(simp add: value_aux value_f_i real field_simps α'_nonneg ennreal_leI del: f_i_simps)
    then have "(λi. value_flow Δ (f_i i)) ⇢ (SUP i. value_flow Δ (f_i i))" by(rule LIMSEQ_SUP)
    also have "(SUP i. value_flow Δ (f_i i)) = α"
    proof -
      have "α - (SUP i. value_flow Δ (f_i i)) = (INF i. α - value_flow Δ (f_i i))"
        by(simp add: ennreal_SUP_const_minus real)
      also have "α - value_flow Δ (f_i i) = α' / 2 ^ i" for i
        by(simp add: value_f_i value_aux real ennreal_minus_if field_simps mult_le_cancel_left1)
      hence "(INF i. α - value_flow Δ (f_i i)) = (INF i. ennreal (α' / 2  ^ i))"
        by(auto intro: INF_cong)
      also have "… = 0"
      proof(rule LIMSEQ_unique)
        show "(λi. α' / 2 ^ i) ⇢ (INF i. ennreal (α' / 2  ^ i))"
          by(rule LIMSEQ_INF)(simp add: field_simps real decseq_SucI)
      qed(simp add: LIMSEQ_divide_realpow_zero real ennreal_0[symmetric] del: ennreal_0)
      finally show "(SUP i. value_flow Δ (f_i i)) = α"
        apply (intro antisym)
        apply (auto simp: α_def intro!: SUP_mono fn.g) []
        apply (rule ennreal_minus_eq_0)
        apply assumption
        done
    qed
    also have "(λi. value_flow Δ (f_i i)) ⇢ value_flow Δ f'"
      by(simp add: value_flow[OF flow source_out] value_flow[OF fn.g source_out] f)
    ultimately have "value_flow Δ f' = α" by(blast intro: LIMSEQ_unique) }
  note value_f = this
  moreover {
    fix x
    have "d_IN f' x = ∫+ y. liminf (λi. f_i i (y, x)) ∂count_space UNIV" unfolding d_IN_def using f
      by(simp add: tendsto_iff_Liminf_eq_Limsup)
    also have "… ≤ liminf (λi. d_IN (f_i i) x)" unfolding d_IN_def
      by(rule nn_integral_liminf)(simp_all add:)
    also have "… ≤ liminf (λi. α)" using d_IN_i[of _ x] fn.g
      by(auto intro!: Liminf_mono SUP_upper2 eventually_sequentiallyI simp add: α_def)
    also have "… = value_flow Δ f'" using value_f by(simp add: Liminf_const)
    also note calculation
  } ultimately show ?thesis by blast
qed

theorem ex_max_flow'': -- ‹eliminate assumption of no antiparallel edges using locale @{const wf_residual_network}›
  assumes source_out: "⋀y. edge Δ (source Δ) y ⟷ y = x"
  and nontrivial: "E ≠ {}"
  and real: "α = ennreal α'" and nn[simp]: "0 ≤ α'"
  shows "∃f. flow Δ f ∧ value_flow Δ f = α ∧ (∀x. d_IN f x ≤ value_flow Δ f)"
proof -
  interpret antiparallel_edges Δ ..
  interpret Δ'': flow_attainability Δ''
    by(rule Δ''_flow_attainability flow_attainability.axioms(2))+unfold_locales
  have wf_Δ'': "Δ''.wf_residual_network"
    by(rule Δ''_wf_residual_network; simp add: no_loop)

  have source_out': "edge Δ'' (source Δ'') y ⟷ y = Edge (source Δ) x" for y
    by(auto simp add: source_out)
  have nontrivial': "VΔ'' - {source Δ'', sink Δ''} ≠ {}" using nontrivial by(auto simp add: "V_Δ''")

  have "(SUP g : {g. flow Δ'' g}. value_flow Δ'' g) = (SUP g : {g. flow Δ g}. value_flow Δ g)" (is "?lhs = ?rhs")
  proof(intro antisym SUP_least; unfold mem_Collect_eq)
    fix g
    assume g: "flow Δ'' g"
    hence "value_flow Δ'' g = value_flow Δ (collect g)" by(simp add: value_collect)
    also { from g have "flow Δ (collect g)" by simp }
    then have "… ≤ ?rhs" by(blast intro: SUP_upper2)
    finally show "value_flow Δ'' g ≤ …" .
  next
    fix g
    assume g: "flow Δ g"
    hence "value_flow Δ g = value_flow Δ'' (split g)" by simp
    also { from g have "flow Δ'' (split g)" by simp }
    then have "… ≤ ?lhs" by(blast intro: SUP_upper2)
    finally show "value_flow Δ g ≤ ?lhs" .
  qed
  with real have eq: "(SUP g : {g. flow Δ'' g}. value_flow Δ'' g) = ennreal α'" by(simp add: α_def)
  from Δ''.ex_max_flow'[OF wf_Δ'' source_out' nontrivial' eq]
  obtain f where f: "flow Δ'' f"
    and "value_flow Δ'' f = α"
    and IN: "⋀x. d_IN f x ≤ value_flow Δ'' f" unfolding eq real using nn by blast
  hence "flow Δ (collect f)" and "value_flow Δ (collect f) = α" by(simp_all add: value_collect)
  moreover {
    fix x
    have "d_IN (collect f) x = (∑+ y∈range (λy. Edge y x). f (y, Vertex x))"
      by(simp add: nn_integral_count_space_reindex d_IN_def)
    also have "… ≤ d_IN f (Vertex x)" unfolding d_IN_def
      by (auto intro!: nn_integral_mono simp add: nn_integral_count_space_indicator split: split_indicator)
    also have "… ≤ value_flow Δ (collect f)" using IN[of "Vertex x"] f by(simp add: value_collect)
    also note calculation }
  ultimately show ?thesis by blast
qed

context begin -- ‹We eliminate the assumption of only one edge leaving the source by introducing a new source vertex.›
private datatype (plugins del: transfer size) 'v' node = SOURCE | Inner (inner: 'v')

private lemma not_Inner_conv: "x ∉ range Inner ⟷ x = SOURCE"
by(cases x) auto

private lemma inj_on_Inner [simp]: "inj_on Inner A"
by(simp add: inj_on_def)

private inductive edge' :: "'v node ⇒ 'v node ⇒ bool"
where
  SOURCE: "edge' SOURCE (Inner (source Δ))"
| Inner: "edge Δ x y ⟹ edge' (Inner x) (Inner y)"

private inductive_simps edge'_simps [simp]:
  "edge' SOURCE x"
  "edge' (Inner y) x"
  "edge' y SOURCE"
  "edge' y (Inner x)"

private fun capacity' :: "'v node flow"
where
  "capacity' (SOURCE, Inner x) = (if x = source Δ then α else 0)"
| "capacity' (Inner x, Inner y) = capacity Δ (x, y)"
| "capacity' _ = 0"

private lemma capacity'_source_in [simp]: "capacity' (y, Inner (source Δ)) = (if y = SOURCE then α else 0)"
by(cases y)(simp_all add: capacity_outside source_in)

private definition Δ' :: "'v node network"
where "Δ' = ⦇edge = edge', capacity = capacity', source = SOURCE, sink = Inner (sink Δ)⦈"

private lemma Δ'_sel [simp]:
  "edge Δ' = edge'"
  "capacity Δ' = capacity'"
  "source Δ' = SOURCE"
  "sink Δ' = Inner (sink Δ)"
by(simp_all add: Δ'_def)

private lemma "E_Δ'": "EΔ' = {(SOURCE, Inner (source Δ))} ∪ (λ(x, y). (Inner x, Inner y)) ` E"
by(auto elim: edge'.cases)

private lemma Δ'_countable_network:
  assumes "α ≠ ⊤"
  shows "countable_network Δ'"
proof
  show "countable EΔ'" unfolding "E_Δ'" by simp
  show "source Δ' ≠ sink Δ'" by simp
  show "capacity Δ' e = 0" if "e ∉ EΔ'" for e using that unfolding "E_Δ'"
    by(cases e rule: capacity'.cases)(auto intro: capacity_outside)
  show "capacity Δ' e ≠ ⊤" for e by(cases e rule: capacity'.cases)(simp_all add: assms)
qed

private lemma Δ'_flow_attainability:
  assumes "α ≠ ⊤"
  shows "flow_attainability Δ'"
proof -
  interpret Δ': countable_network Δ' using assms by(rule Δ'_countable_network)
  show ?thesis
  proof
    show "d_IN (capacity Δ') x ≠ ⊤ ∨ d_OUT (capacity Δ') x ≠ ⊤" if sink: "x ≠ sink Δ'" for x
    proof(cases x)
      case (Inner x')
      consider (source) "x' = source Δ" | (IN) "x' ≠ source Δ" "d_IN (capacity Δ) x' ≠ ⊤" | (OUT) "d_OUT (capacity Δ) x' ≠ ⊤"
        using finite_capacity[of x'] sink Inner by(auto)
      thus ?thesis
      proof(cases)
        case source
        with Inner have "d_IN (capacity Δ') x = (∑+ y. α * indicator {SOURCE :: 'v node} y)"
          unfolding d_IN_def by(intro nn_integral_cong)(simp split: split_indicator)
        also have "… = α" by(simp add: max_def)
        finally show ?thesis using assms by simp
      next
        case IN
        with Inner have "d_IN (capacity Δ') x = (∑+ y∈range Inner. capacity Δ (node.inner y, x'))"
          by(auto simp add: d_IN_def nn_integral_count_space_indicator not_Inner_conv intro!: nn_integral_cong split: split_indicator)
        also have "… = d_IN (capacity Δ) x'" unfolding d_IN_def
          by(simp add: nn_integral_count_space_reindex)
        finally show ?thesis using Inner sink IN by(simp)
      next
        case OUT
        from Inner have "d_OUT (capacity Δ') x = (∑+ y∈range Inner. capacity Δ (x', node.inner y))"
          by(auto simp add: d_OUT_def nn_integral_count_space_indicator not_Inner_conv intro!: nn_integral_cong split: split_indicator)
        also have "… = d_OUT (capacity Δ) x'" by(simp add: d_OUT_def nn_integral_count_space_reindex)
        finally show ?thesis using OUT by auto
      qed
    qed(simp add: d_IN_def)
    show "¬ edge Δ' x x" for x by(cases x)(simp_all add: no_loop)
    show "¬ edge Δ' x (source Δ')" for x by simp
  qed
qed

private fun lift :: "'v flow ⇒ 'v node flow"
where
  "lift f (SOURCE, Inner y) = (if y = source Δ then value_flow Δ f else 0)"
| "lift f (Inner x, Inner y) = f (x, y)"
| "lift f _ = 0"

private lemma d_OUT_lift_Inner [simp]: "d_OUT (lift f) (Inner x) = d_OUT f x" (is "?lhs = ?rhs")
proof -
  have "?lhs = (∑+ y∈range Inner. lift f (Inner x, y))"
    by(auto simp add: d_OUT_def nn_integral_count_space_indicator not_Inner_conv intro!: nn_integral_cong split: split_indicator)
  also have "… = ?rhs" by(simp add: nn_integral_count_space_reindex d_OUT_def)
  finally show ?thesis .
qed

private lemma d_OUT_lift_SOURCE [simp]: "d_OUT (lift f) SOURCE = value_flow Δ f" (is "?lhs = ?rhs")
proof -
  have "?lhs = (∑+ y. lift f (SOURCE, y) * indicator {Inner (source Δ)} y)"
    unfolding d_OUT_def by(rule nn_integral_cong)(case_tac x; simp)
  also have "… = ?rhs" by(simp add: nn_integral_count_space_indicator max_def)
  finally show ?thesis .
qed

private lemma d_IN_lift_Inner [simp]:
  assumes "x ≠ source Δ"
  shows "d_IN (lift f) (Inner x) = d_IN f x" (is "?lhs = ?rhs")
proof -
  have "?lhs = (∑+ y∈range Inner. lift f (y, Inner x))" using assms
    by(auto simp add: d_IN_def nn_integral_count_space_indicator not_Inner_conv intro!: nn_integral_cong split: split_indicator)
  also have "… = ?rhs" by(simp add: nn_integral_count_space_reindex d_IN_def)
  finally show ?thesis .
qed

private lemma d_IN_lift_source [simp]: "d_IN (lift f) (Inner (source Δ)) = value_flow Δ f + d_IN f (source Δ)" (is "?lhs = ?rhs")
proof -
  have "?lhs = (∑+ y. lift f (y, Inner (source Δ)) * indicator {SOURCE} y) + (∑+ y∈range Inner. lift f (y, Inner (source Δ)))"
    (is "_ = ?SOURCE + ?rest")
    unfolding d_IN_def
    apply(subst nn_integral_count_space_indicator, simp)
    apply(subst nn_integral_add[symmetric])
    apply(auto simp add: AE_count_space max_def not_Inner_conv split: split_indicator intro!: nn_integral_cong)
    done
  also have "?rest = d_IN f (source Δ)" by(simp add: nn_integral_count_space_reindex d_IN_def)
  also have "?SOURCE = value_flow Δ f"
    by(simp add: max_def one_ennreal_def[symmetric] )
  finally show ?thesis .
qed

private lemma flow_lift [simp]:
  assumes "flow Δ f"
  shows "flow Δ' (lift f)"
proof
  show "lift f e ≤ capacity Δ' e" for e
    by(cases e rule: capacity'.cases)(auto intro: flowD_capacity[OF assms] simp add: α_def intro: SUP_upper2 assms)

  fix x
  assume x: "x ≠ source Δ'" "x ≠ sink Δ'"
  then obtain x' where x': "x = Inner x'" by(cases x) auto
  then show "KIR (lift f) x" using x
    by(cases "x' = source Δ")(auto simp add: flowD_source_IN[OF assms] dest: flowD_KIR[OF assms])
qed

private abbreviation (input) unlift :: "'v node flow ⇒ 'v flow"
where "unlift f ≡ (λ(x, y). f (Inner x, Inner y))"

private lemma flow_unlift [simp]:
  assumes f: "flow Δ' f"
  shows "flow Δ (unlift f)"
proof
  show "unlift f e ≤ capacity Δ e" for e using flowD_capacity[OF f, of "map_prod Inner Inner e"]
    by(cases e)(simp)
next
  fix x
  assume x: "x ≠ source Δ" "x ≠ sink Δ"
  have "d_OUT (unlift f) x = (∑+ y∈range Inner. f (Inner x, y))"
    by(simp add: nn_integral_count_space_reindex d_OUT_def)
  also have "… = d_OUT f (Inner x)" using flowD_capacity[OF f, of "(Inner x, SOURCE)"]
    by(auto simp add: nn_integral_count_space_indicator d_OUT_def not_Inner_conv intro!: nn_integral_cong split: split_indicator)
  also have "… = d_IN f (Inner x)" using x flowD_KIR[OF f, of "Inner x"] by(simp)
  also have "… = (∑+ y∈range Inner. f (y, Inner x))"
    using x flowD_capacity[OF f, of "(SOURCE, Inner x)"]
    by(auto simp add: nn_integral_count_space_indicator d_IN_def not_Inner_conv intro!: nn_integral_cong split: split_indicator)
  also have "… = d_IN (unlift f) x" by(simp add: nn_integral_count_space_reindex d_IN_def)
  finally show "KIR (unlift f) x" .
qed

private lemma value_unlift:
  assumes f: "flow Δ' f"
  shows "value_flow Δ (unlift f) = value_flow Δ' f"
proof -
  have "value_flow Δ (unlift f) = (∑+ y∈range Inner. f (Inner (source Δ), y))"
    by(simp add: nn_integral_count_space_reindex d_OUT_def)
  also have "… = d_OUT f (Inner (source Δ))" using flowD_capacity[OF f, of "(Inner (source Δ), SOURCE)"]
    by(auto simp add: nn_integral_count_space_indicator d_OUT_def not_Inner_conv intro!: nn_integral_cong split: split_indicator)
  also have "… = d_IN f (Inner (source Δ))" using flowD_KIR[OF f, of "Inner (source Δ)"] by(simp)
  also have "… = (∑+ y. f (y, Inner (source Δ)) * indicator {SOURCE} y)"
    unfolding d_IN_def using flowD_capacity[OF f, of "(x, Inner (source Δ))" for x]
    by(intro nn_integral_cong)(auto intro!: antisym split: split_indicator if_split_asm elim: meta_allE)
  also have "… = f (SOURCE, Inner (source Δ))" by simp
  also have "… = (∑+ y. f (SOURCE, y) * indicator {Inner (source Δ)} y)"
    by(simp add: one_ennreal_def[symmetric])
  also have "… = value_flow Δ' f" unfolding d_OUT_def
    unfolding d_OUT_def using flowD_capacity[OF f, of "(SOURCE, Inner x)" for x] flowD_capacity[OF f, of "(SOURCE, SOURCE)"]
    apply(intro nn_integral_cong)
    apply(case_tac x)
    apply(auto intro!: antisym split: split_indicator if_split_asm elim: meta_allE)
    done
  finally show ?thesis .
qed

theorem ex_max_flow:
  "∃f. flow Δ f ∧ value_flow Δ f = α ∧ (∀x. d_IN f x ≤ value_flow Δ f)"
proof(cases "α")
  case (real α')
  hence α: "α ≠ ⊤" by simp
  then interpret Δ': flow_attainability Δ' by(rule Δ'_flow_attainability)

  have source_out: "edge Δ' (source Δ') y ⟷ y = Inner (source Δ)" for y by(auto)
  have nontrivial: "EΔ' ≠ {}" by(auto intro: edge'.intros)

  have eq: "(SUP g : {g. flow Δ' g}. value_flow Δ' g) = (SUP g : {g. flow Δ g}. value_flow Δ g)" (is "?lhs = ?rhs")
  proof(intro antisym SUP_least; unfold mem_Collect_eq)
    fix g
    assume g: "flow Δ' g"
    hence "value_flow Δ' g = value_flow Δ (unlift g)" by(simp add: value_unlift)
    also { from g have "flow Δ (unlift g)" by simp }
    then have "… ≤ ?rhs" by(blast intro: SUP_upper2)
    finally show "value_flow Δ' g ≤ …" .
  next
    fix g
    assume g: "flow Δ g"
    hence "value_flow Δ g = value_flow Δ' (lift g)" by simp
    also { from g have "flow Δ' (lift g)" by simp }
    then have "… ≤ ?lhs" by(blast intro: SUP_upper2)
    finally show "value_flow Δ g ≤ ?lhs" .
  qed
  also have "… = ennreal α'" using real by(simp add: α_def)
  finally obtain f where f: "flow Δ' f"
    and value_f: "value_flow Δ' f = (⨆g∈{g. flow Δ' g}. value_flow Δ' g)"
    and IN_f: "⋀x. d_IN f x ≤ value_flow Δ' f"
    using ‹0 ≤ α'› by(blast dest: Δ'.ex_max_flow''[OF source_out nontrivial])
  have "flow Δ (unlift f)" using f by simp
  moreover have "value_flow Δ (unlift f) = α" using f eq value_f by(simp add: value_unlift α_def)
  moreover {
    fix x
    have "d_IN (unlift f) x = (∑+ y∈range Inner. f (y, Inner x))"
      by(simp add: nn_integral_count_space_reindex d_IN_def)
    also have "… ≤ d_IN f (Inner x)" unfolding d_IN_def
      by(auto intro!: nn_integral_mono simp add: nn_integral_count_space_indicator split: split_indicator)
    also have "… ≤ value_flow Δ (unlift f)" using IN_f[of "Inner x"] f by(simp add: value_unlift)
    also note calculation }
  ultimately show ?thesis by blast
next
  case top
  show ?thesis
  proof(cases "∃f. flow Δ f ∧ value_flow Δ f = ⊤")
    case True
    with top show ?thesis by auto
  next
    case False
    hence real: "∀f. α = ⊤ ⟶ flow Δ f ⟶ value_flow Δ f < α" using top by (auto simp: less_top)
    { fix i
      have "2 * 2 ^ i < α" using top by (simp_all add: ennreal_mult_less_top power_less_top_ennreal)
      from flow_by_value[OF this real] have "∃f. flow Δ f ∧ value_flow Δ f = 2 * 2 ^ i" by blast }
    then obtain f_i where f_i: "⋀i. flow Δ (f_i i)"
      and value_i: "⋀i. value_flow Δ (f_i i) = 2 * 2 ^ i" by metis
    def f  "λe. ∑+ i. f_i i e / (2 * 2 ^ i)"
    have "flow Δ f"
    proof
      fix e
      have "f e ≤ (∑+ i. (SUP i. f_i i e) / (2 * 2 ^ i))" unfolding f_def
        by(rule nn_integral_mono)(auto intro!: divide_right_mono_ennreal SUP_upper)
      also have "… = (SUP i. f_i i e) / 2 * (∑+ i. 1 / 2 ^ i)"
        apply(subst nn_integral_cmult[symmetric])
        apply(auto intro!: nn_integral_cong intro: SUP_upper2
          simp: divide_ennreal_def ennreal_inverse_mult power_less_top_ennreal mult_ac)
        done
      also have "(∑+ i. 1 / 2 ^ i) = (∑i. ennreal ((1 / 2) ^ i))"
        by(simp add: nn_integral_count_space_nat power_divide divide_ennreal[symmetric] ennreal_power[symmetric])
      also have "… = ennreal (∑i. (1 / 2) ^ i)"
        by(intro suminf_ennreal2 complete_algebra_summable_geometric) simp_all
      also have "… = 2" by(subst suminf_geometric; simp)
      also have "(SUP i. f_i i e) / 2 * 2 = (SUP i. f_i i e)"
        by (simp add: ennreal_divide_times)
      also have "… ≤ capacity Δ e" by(rule SUP_least)(rule flowD_capacity[OF f_i])
      finally show "f e ≤ capacity Δ e" .

      fix x
      assume x: "x ≠ source Δ" "x ≠ sink Δ"
      have "d_OUT f x = (∑+ i∈UNIV. ∑+ y. f_i i (x, y) / (2 * 2 ^ i))"
        unfolding d_OUT_def f_def
        by(subst nn_integral_snd_count_space[where f="case_prod _", simplified])
          (simp add: nn_integral_fst_count_space[where f="case_prod _", simplified])
      also have "… = (∑+ i. d_OUT (f_i i) x / (2 * 2 ^ i))" unfolding d_OUT_def
        by(simp add: nn_integral_divide)
      also have "… = (∑+ i. d_IN (f_i i) x / (2 * 2 ^ i))" by(simp add: flowD_KIR[OF f_i, OF x])
      also have "… = (∑+ i∈UNIV. ∑+ y. f_i i (y, x) / (2 * 2 ^ i))"
        by(simp add: nn_integral_divide d_IN_def)
      also have "… = d_IN f x" unfolding d_IN_def f_def
        by(subst nn_integral_snd_count_space[where f="case_prod _", simplified])
          (simp add: nn_integral_fst_count_space[where f="case_prod _", simplified])
      finally show "KIR f x" .
    qed
    moreover {
      have "value_flow Δ f = (∑+ i. value_flow Δ (f_i i) / (2 * 2 ^ i))"
        unfolding d_OUT_def f_def
        by(subst nn_integral_snd_count_space[where f="case_prod _", simplified])
          (simp add: nn_integral_fst_count_space[where f="case_prod _", simplified] nn_integral_divide[symmetric])
      also have "… = ⊤"
        by(simp add: value_i ennreal_mult_less_top power_less_top_ennreal)
      finally have "value_flow Δ f = ⊤" .
    }
    ultimately show ?thesis using top by auto
  qed
qed

end

end

end

section ‹Webs and currents›

record 'v web = "'v graph" +
  weight :: "'v ⇒ ennreal"
  A :: "'v set"
  B :: "'v set"

lemma vertex_weight_update [simp]: "vertex (weight_update f Γ) = vertex Γ"
by(simp add: vertex_def fun_eq_iff)

type_synonym 'v current = "'v edge ⇒ ennreal"

inductive current :: "('v, 'more) web_scheme ⇒ 'v current ⇒ bool"
  for Γ f
where
  current:
  "⟦ ⋀x. d_OUT f x ≤ weight Γ x;
     ⋀x. d_IN f x ≤ weight Γ x;
     ⋀x. x ∉ A Γ ⟹ d_OUT f x ≤ d_IN f x;
     ⋀a. a ∈ A Γ ⟹ d_IN f a = 0;
     ⋀b. b ∈ B Γ ⟹ d_OUT f b = 0;
     ⋀e. e ∉ EΓ ⟹ f e = 0 ⟧
  ⟹ current Γ f"

lemma currentD_weight_OUT: "current Γ f ⟹ d_OUT f x ≤ weight Γ x"
by(simp add: current.simps)

lemma currentD_weight_IN: "current Γ f ⟹ d_IN f x ≤ weight Γ x"
by(simp add: current.simps)

lemma currentD_OUT_IN: "⟦ current Γ f; x ∉ A Γ ⟧ ⟹ d_OUT f x ≤ d_IN f x"
by(simp add: current.simps)

lemma currentD_IN: "⟦ current Γ f; a ∈ A Γ ⟧ ⟹ d_IN f a = 0"
by(simp add: current.simps)

lemma currentD_OUT: "⟦ current Γ f; b ∈ B Γ ⟧ ⟹ d_OUT f b = 0"
by(simp add: current.simps)

lemma currentD_outside: "⟦ current Γ f; ¬ edge Γ x y ⟧ ⟹ f (x, y) = 0"
by(blast elim: current.cases)

lemma currentD_outside': "⟦ current Γ f; e ∉ EΓ ⟧ ⟹ f e = 0"
by(blast elim: current.cases)

lemma currentD_OUT_eq_0:
  assumes "current Γ f"
  shows "d_OUT f x = 0 ⟷ (∀y. f (x, y) = 0)"
by(simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0)

lemma currentD_IN_eq_0:
  assumes "current Γ f"
  shows "d_IN f x = 0 ⟷ (∀y. f (y, x) = 0)"
by(simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0)

lemma current_support_flow:
  fixes Γ (structure)
  assumes "current Γ f"
  shows "support_flow f ⊆ E"
using currentD_outside[OF assms] by(auto simp add: support_flow.simps intro: ccontr)

lemma currentD_outside_IN: "⟦ current Γ f; x ∉ VΓ ⟧ ⟹ d_IN f x = 0"
by(auto simp add: d_IN_def vertex_def nn_integral_0_iff  AE_count_space emeasure_count_space_eq_0 dest: currentD_outside)

lemma currentD_outside_OUT: "⟦ current Γ f; x ∉ VΓ ⟧ ⟹ d_OUT f x = 0"
by(auto simp add: d_OUT_def vertex_def nn_integral_0_iff  AE_count_space emeasure_count_space_eq_0 dest: currentD_outside)

lemma currentD_weight_in: "current Γ h ⟹ h (x, y) ≤ weight Γ y"
  by (metis order_trans d_IN_ge_point currentD_weight_IN)

lemma currentD_weight_out: "current Γ h ⟹ h (x, y) ≤ weight Γ x"
  by (metis order_trans d_OUT_ge_point currentD_weight_OUT)

lemma current_leI:
  fixes Γ (structure)
  assumes f: "current Γ f"
  and le: "⋀e. g e ≤ f e"
  and OUT_IN: "⋀x. x ∉ A Γ ⟹ d_OUT g x ≤ d_IN g x"
  shows "current Γ g"
proof
  show "d_OUT g x ≤ weight Γ x" for x
    using d_OUT_mono[of g x f, OF le] currentD_weight_OUT[OF f] by(rule order_trans)
  show "d_IN g x ≤ weight Γ x" for x
    using d_IN_mono[of g x f, OF le] currentD_weight_IN[OF f] by(rule order_trans)
  show "d_IN g a = 0" if "a ∈ A Γ" for a
    using d_IN_mono[of g a f, OF le] currentD_IN[OF f that] by auto
  show "d_OUT g b = 0" if "b ∈ B Γ" for b
    using d_OUT_mono[of g b f, OF le] currentD_OUT[OF f that] by auto
  show "g e = 0" if "e ∉ E" for e
    using currentD_outside'[OF f that] le[of e] by simp
qed(blast intro: OUT_IN)+

lemma current_weight_mono:
  "⟦ current Γ f; edge Γ = edge Γ'; A Γ = A Γ'; B Γ = B Γ'; ⋀x. weight Γ x ≤ weight Γ' x ⟧
  ⟹ current Γ' f"
by(auto 4 3 elim!: current.cases intro!: current.intros intro: order_trans)

abbreviation (input) zero_current :: "'v current"
where "zero_current ≡ λ_. 0"

lemma SINK_0 [simp]: "SINK zero_current = UNIV"
by(auto simp add: SINK.simps)

lemma current_0 [simp]: "current Γ zero_current"
by(auto simp add: current.simps)

inductive web_flow :: "('v, 'more) web_scheme ⇒ 'v current ⇒ bool"
  for Γ (structure) and f
where
  web_flow: "⟦ current Γ f; ⋀x. ⟦ x ∈ V; x ∉ A Γ; x ∉ B Γ ⟧ ⟹ KIR f x ⟧ ⟹ web_flow Γ f"

lemma web_flowD_current: "web_flow Γ f ⟹ current Γ f"
by(erule web_flow.cases)

lemma web_flowD_KIR: "⟦ web_flow Γ f; x ∉ A Γ; x ∉ B Γ ⟧ ⟹ KIR f x"
apply(cases "x ∈ VΓ")
 apply(fastforce elim!: web_flow.cases)
apply(auto simp add: vertex_def d_OUT_def d_IN_def elim!: web_flow.cases)
apply(subst (1 2) currentD_outside[of _ f]; auto)
done

subsection ‹Saturated and terminal vertices›

inductive_set SAT :: "('v, 'more) web_scheme ⇒ 'v current ⇒ 'v set"
  for Γ f
where
  A: "x ∈ A Γ ⟹ x ∈ SAT Γ f"
| IN: "d_IN f x ≥ weight Γ x ⟹ x ∈ SAT Γ f"
  -- ‹We use @{text "≥ weight"} such that @{text SAT} is monotone w.r.t. increasing currents›

lemma SAT_0 [simp]: "SAT Γ zero_current = A Γ ∪ {x. weight Γ x ≤ 0}"
by(auto simp add: SAT.simps)

lemma SAT_mono:
  assumes "⋀e. f e ≤ g e"
  shows "SAT Γ f ⊆ SAT Γ g"
proof
  fix x
  assume "x ∈ SAT Γ f"
  thus "x ∈ SAT Γ g"
  proof cases
    case IN
    also have "d_IN f x ≤ d_IN g x" using assms by(rule d_IN_mono)
    finally show ?thesis ..
  qed(rule SAT.A)
qed

lemma SAT_Sup_upper: "f ∈ Y ⟹ SAT Γ f ⊆ SAT Γ (Sup Y)"
by(rule SAT_mono)(rule Sup_upper[THEN le_funD])

lemma currentD_SAT:
  assumes "current Γ f"
  shows "x ∈ SAT Γ f ⟷ x ∈ A Γ ∨ d_IN f x = weight Γ x"
using currentD_weight_IN[OF assms, of x] by(auto simp add: SAT.simps)

abbreviation terminal :: "('v, 'more) web_scheme ⇒ 'v current ⇒ 'v set" ("TERı")
where "terminal Γ f ≡ SAT Γ f ∩ SINK f"

subsection ‹Separation›

inductive separating_gen :: "('v, 'more) graph_scheme ⇒ 'v set ⇒ 'v set ⇒ 'v set ⇒ bool"
  for G A B S
where separating:
  "(⋀x y p. ⟦ x ∈ A; y ∈ B; path G x p y ⟧ ⟹ (∃z ∈ set p. z ∈ S) ∨ x ∈ S)
  ⟹ separating_gen G A B S"

abbreviation separating :: "('v, 'more) web_scheme ⇒ 'v set ⇒ bool"
where "separating Γ ≡ separating_gen Γ (A Γ) (B Γ)"

abbreviation separating_network :: "('v, 'more) network_scheme ⇒ 'v set ⇒ bool"
where "separating_network Δ ≡ separating_gen Δ {source Δ} {sink Δ}"

lemma separating_networkI [intro?]:
  "(⋀p. path Δ (source Δ) p (sink Δ) ⟹ (∃z ∈ set p. z ∈ S) ∨ source Δ ∈ S)
  ⟹ separating_network Δ S"
by(auto intro: separating)

lemma separatingD:
  "⋀A B. ⟦ separating_gen G A B S; path G x p y; x ∈ A; y ∈ B ⟧ ⟹ (∃z ∈ set p. z ∈ S) ∨ x ∈ S"
by(blast elim: separating_gen.cases)

lemma separating_left [simp]: "⋀A B. A ⊆ A' ⟹ separating_gen Γ A B A'"
by(auto simp add: separating_gen.simps)

lemma separating_weakening:
  "⋀A B. ⟦ separating_gen G A B S; S ⊆ S' ⟧ ⟹ separating_gen G A B S'"
by(rule separating; drule (3) separatingD; blast)

definition essential :: "('v, 'more) graph_scheme ⇒ 'v set ⇒ 'v set ⇒ 'v ⇒ bool"
where -- ‹Should we allow only simple paths here?›
  "⋀B. essential G B S x ⟷ (∃p. ∃y∈B. path G x p y ∧ (x ≠ y ⟶ (∀z∈set p. z = x ∨ z ∉ S)))"

abbreviation essential_web :: "('v, 'more) web_scheme ⇒ 'v set ⇒ 'v set" ("ℰı")
where "essential_web Γ S ≡ {x∈S. essential Γ (B Γ) S x}"

lemma essential_weight_update [simp]:
  "essential (weight_update f G) = essential G"
by(simp add: essential_def fun_eq_iff)

lemma not_essentialD:
  "⋀B. ⟦ ¬ essential G B S x; path G x p y; y ∈ B ⟧ ⟹ x ≠ y ∧ (∃z∈set p. z ≠ x ∧ z ∈ S)"
by(simp add: essential_def)

lemma essentialE [elim?, consumes 1, case_names essential, cases pred: essential]:
  "⋀B. ⟦ essential G B S x; ⋀p y. ⟦ path G x p y; y ∈ B; ⋀z. ⟦ x ≠ y; z ∈ set p ⟧ ⟹ z = x ∨ z ∉ S ⟧ ⟹ thesis ⟧ ⟹ thesis"
by(auto simp add: essential_def)

lemma essentialI [intro?]:
  "⋀B. ⟦ path G x p y; y ∈ B; ⋀z. ⟦ x ≠ y; z ∈ set p ⟧ ⟹ z = x ∨ z ∉ S ⟧ ⟹ essential G B S x"
by(auto simp add: essential_def)

lemma essential_vertex: "⋀B. ⟦ essential G B S x; x ∉ B ⟧ ⟹vertex G x"
by(auto elim!: essentialE simp add: vertex_def elim: rtrancl_path.cases)

lemma essential_BI: "⋀B. x ∈ B ⟹ essential G B S x"
by(auto simp add: essential_def intro: rtrancl_path.base)

lemma ℰ_E [elim?, consumes 1, case_names , cases set: essential_web]:
  fixes Γ (structure)
  assumes "x ∈ ℰ S"
  obtains p y where "path Γ x p y" "y ∈ B Γ" "⋀z. ⟦ x ≠ y; z ∈ set p ⟧ ⟹ z = x ∨ z ∉ S"
using assms by(auto elim: essentialE)

lemma essential_mono: "⋀B. ⟦ essential G B S x; S' ⊆ S ⟧ ⟹ essential G B S' x"
by(auto simp add: essential_def)

lemma separating_essential: -- ‹Lem. 3.4 (cf. Lem. 2.14 in [5])›
  fixes G A B S
  assumes "separating_gen G A B S"
  shows "separating_gen G A B {x∈S. essential G B S x}" (is "separating_gen _ _ _ ?E")
proof
  fix x y p
  assume x: "x ∈ A" and y: "y ∈ B" and p: "path G x p y"
  from separatingD[OF assms p x y] have "∃z ∈ set (x # p). z ∈ S" by auto
  from split_list_last_prop[OF this] obtain ys z zs where decomp: "x # p = ys @ z # zs"
    and z: "z ∈ S" and last: "⋀z. z ∈ set zs ⟹ z ∉ S" by auto
  from decomp consider (empty) "ys = []" "x = z" "p = zs"
    | (Cons) ys' where "ys = x # ys'" "p = ys' @ z # zs"
    by(auto simp add: Cons_eq_append_conv)
  then show "(∃z∈set p. z ∈ ?E) ∨ x ∈ ?E"
  proof(cases)
    case empty
    hence "x ∈ ?E" using z p last y by(auto simp add: essential_def)
    thus ?thesis ..
  next
    case (Cons ys')
    from p have "path G z zs y" unfolding Cons by(rule rtrancl_path_appendE)
    hence "z ∈ ?E" using z y last by(auto simp add: essential_def)
    thus ?thesis using Cons by auto
  qed
qed

definition roofed_gen :: "('v, 'more) graph_scheme ⇒ 'v set ⇒ 'v set ⇒ 'v set"
where roofed_def: "⋀B. roofed_gen G B S = {x. ∀p. ∀y∈B. path G x p y ⟶ (∃z∈set p. z ∈ S) ∨ x ∈ S}"

abbreviation roofed :: "('v, 'more) web_scheme ⇒ 'v set ⇒ 'v set" ("RFı")
where "roofed Γ ≡ roofed_gen Γ (B Γ)"

abbreviation roofed_network :: "('v, 'more) network_scheme ⇒ 'v set ⇒ 'v set" ("RFNı")
where "roofed_network Δ ≡ roofed_gen Δ {sink Δ}"

lemma roofedI [intro?]:
  "⋀B. (⋀p y. ⟦ path G x p y; y ∈ B ⟧ ⟹ (∃z∈set p. z ∈ S) ∨ x ∈ S) ⟹ x ∈ roofed_gen G B S"
by(auto simp add: roofed_def)

lemma not_roofedE: fixes B
  assumes "x ∉ roofed_gen G B S"
  obtains p y where "path G x p y" "y ∈ B" "⋀z. z ∈ set (x # p) ⟹ z ∉ S"
using assms by(auto simp add: roofed_def)

lemma roofed_greater: "⋀B. S ⊆ roofed_gen G B S"
by(auto simp add: roofed_def)

lemma roofed_greaterI: "⋀B. x ∈ S ⟹ x ∈ roofed_gen G B S"
using roofed_greater[of S G] by blast

lemma roofed_mono: "⋀B. S ⊆ S' ⟹ roofed_gen G B S ⊆ roofed_gen G B S'"
by(fastforce simp add: roofed_def)

lemma in_roofed_mono: "⋀B. ⟦ x ∈ roofed_gen G B S; S ⊆ S' ⟧ ⟹ x ∈ roofed_gen G B S'"
using roofed_mono[THEN subsetD] .

lemma roofedD: "⋀B. ⟦ x ∈ roofed_gen G B S; path G x p y; y ∈ B ⟧ ⟹ (∃z∈set p. z ∈ S) ∨ x ∈ S"
unfolding roofed_def by blast

lemma separating_RF_A:
  fixes A B
  assumes "separating_gen G A B X"
  shows "A ⊆ roofed_gen G B X"
by(rule subsetI roofedI)+(erule separatingD[OF assms])

lemma roofed_idem: fixes B shows "roofed_gen G B (roofed_gen G B S) = roofed_gen G B S"
proof(rule equalityI subsetI roofedI)+
  fix x p y
  assume x: "x ∈ roofed_gen G B (roofed_gen G B S)" and p: "path G x p y" and y: "y ∈ B"
  from roofedD[OF x p y] obtain z where *: "z ∈ set (x # p)" and z: "z ∈ roofed_gen G B S" by auto
  from split_list[OF *] obtain ys zs where split: "x # p = ys @ z # zs" by blast
  with p have p': "path G z zs y" by(auto simp add: Cons_eq_append_conv elim: rtrancl_path_appendE)
  from roofedD[OF z p' y] split show "(∃z∈set p. z ∈ S) ∨ x ∈ S"
    by(auto simp add: Cons_eq_append_conv)
qed(rule roofed_mono roofed_greater)+

lemma in_roofed_mono': "⋀B. ⟦ x ∈ roofed_gen G B S; S ⊆ roofed_gen G B S' ⟧ ⟹ x ∈ roofed_gen G B S'"
by(subst roofed_idem[symmetric])(erule in_roofed_mono)

lemma roofed_mono': "⋀B. S ⊆ roofed_gen G B S' ⟹ roofed_gen G B S ⊆ roofed_gen G B S'"
by(rule subsetI)(rule in_roofed_mono')

lemma roofed_idem_Un1: fixes B shows "roofed_gen G B (roofed_gen G B S ∪ T) = roofed_gen G B (S ∪ T)"
proof -
  have "S ⊆ T ∪ roofed_gen G B S"
    by (metis (no_types) UnCI roofed_greater subsetCE subsetI)
  then have "S ∪ T ⊆ T ∪ roofed_gen G B S ∧ T ∪ roofed_gen G B S ⊆ roofed_gen G B (S ∪ T)"
    by (metis (no_types) Un_subset_iff Un_upper2 roofed_greater roofed_mono sup.commute)
  then show ?thesis
    by (metis (no_types) roofed_idem roofed_mono subset_antisym sup.commute)
qed

lemma roofed_UN: fixes A B
  shows "roofed_gen G B (⋃i∈A. roofed_gen G B (X i)) = roofed_gen G B (⋃i∈A. X i)" (is "?lhs = ?rhs")
proof(rule equalityI)
  show "?rhs ⊆ ?lhs" by(rule roofed_mono)(blast intro: roofed_greaterI)
  show "?lhs ⊆ ?rhs" by(rule roofed_mono')(blast intro: in_roofed_mono)
qed

lemma RF_essential: fixes Γ (structure) shows "RF (ℰ S) = RF S"
proof(intro set_eqI iffI)
  fix x
  assume RF: "x ∈ RF S"
  show "x ∈ RF (ℰ S)"
  proof
    fix p y
    assume p: "path Γ x p y" and y: "y ∈ B Γ"
    from roofedD[OF RF this] have "∃z∈set (x # p). z ∈ S" by auto
    from split_list_last_prop[OF this] obtain ys z zs where decomp: "x # p = ys @ z # zs"
      and z: "z ∈ S" and last: "⋀z. z ∈ set zs ⟹ z ∉ S" by auto
    from decomp consider (empty) "ys = []" "x = z" "p = zs"
      | (Cons) ys' where "ys = x # ys'" "p = ys' @ z # zs"
      by(auto simp add: Cons_eq_append_conv)
    then show "(∃z∈set p. z ∈ ℰ S) ∨ x ∈ ℰ S"
    proof(cases)
      case empty
      hence "x ∈ ℰ S" using z p last y by(auto simp add: essential_def)
      thus ?thesis ..
    next
      case (Cons ys')
      from p have "path Γ z zs y" unfolding Cons by(rule rtrancl_path_appendE)
      hence "z ∈ ℰ S" using z y last by(auto simp add: essential_def)
      thus ?thesis using Cons by auto
    qed
  qed
qed(blast intro: in_roofed_mono)

lemma essentialE_RF:
  fixes Γ (structure) and B
  assumes "essential Γ B S x"
  obtains p y where "path Γ x p y" "y ∈ B" "distinct (x # p)" "⋀z. z ∈ set p ⟹ z ∉ roofed_gen Γ B S"
proof -
  from assms obtain p y where p: "path Γ x p y" and y: "y ∈ B"
    and bypass: "⋀z. ⟦ x ≠ y; z ∈ set p ⟧ ⟹ z = x ∨ z ∉ S" by(rule essentialE) blast
  from p obtain p' where p': "path Γ x p' y" and distinct: "distinct (x # p')"
    and subset: "set p' ⊆ set p" by(rule rtrancl_path_distinct)
  { fix z
    assume z: "z ∈ set p'"
    hence "y ∈ set p'" using rtrancl_path_last[OF p', symmetric] p'
      by(auto elim: rtrancl_path.cases intro: last_in_set)
    with distinct z subset have neq: "x ≠ y" and "z ∈ set p" by(auto)
    from bypass[OF this] z distinct have "z ∉ S" by auto

    have "z ∉ roofed_gen Γ B S"
    proof
      assume z': "z ∈ roofed_gen Γ B S"
      from split_list[OF z] obtain ys zs where decomp: "p' = ys @ z # zs" by blast
      with p' have "path Γ z zs y" by(auto elim: rtrancl_path_appendE)
      from roofedD[OF z' this y] ‹z ∉ S› obtain z' where "z' ∈ set zs" "z' ∈ S" by auto
      with bypass[of z'] neq decomp subset distinct show False by auto
    qed }
  with p' y distinct show thesis ..
qed

lemma ℰ_E_RF:
  fixes Γ (structure)
  assumes "x ∈ ℰ S"
  obtains p y where "path Γ x p y" "y ∈ B Γ" "distinct (x # p)" "⋀z. z ∈ set p ⟹ z ∉ RF S"
using assms by(auto elim: essentialE_RF)

lemma in_roofed_essentialD:
  fixes Γ (structure)
  assumes RF: "x ∈ RF S"
  and ess: "essential Γ (B Γ) S x"
  shows "x ∈ S"
proof -
  from ess obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ" and "distinct (x # p)"
    and bypass: "⋀z. z ∈ set p ⟹ z ∉ S" by(rule essentialE_RF)(auto intro: roofed_greaterI)
  from roofedD[OF RF p y] bypass show "x ∈ S" by auto
qed

lemma separating_RF: fixes Γ (structure) shows "separating Γ (RF S) ⟷ separating Γ S"
proof
  assume sep: "separating Γ (RF S)"
  show "separating Γ S"
  proof
    fix x y p
    assume p: "path Γ x p y" and x: "x ∈ A Γ" and y: "y ∈ B Γ"
    from separatingD[OF sep p x y] have "∃z ∈ set (x # p). z ∈ RF S" by auto
    from split_list_last_prop[OF this] obtain ys z zs where split: "x # p = ys @ z # zs"
      and z: "z ∈ RF S" and bypass: "⋀z'. z' ∈ set zs ⟹ z' ∉ RF S" by auto
    from p split have "path Γ z zs y" by(cases ys)(auto elim: rtrancl_path_appendE)
    hence "essential Γ (B Γ) S z" using y
      by(rule essentialI)(auto dest: bypass intro: roofed_greaterI)
    with z have "z ∈ S" by(rule in_roofed_essentialD)
    with split show "(∃z∈set p. z ∈ S) ∨ x ∈ S" by(cases ys)auto
  qed
qed(blast intro: roofed_greaterI separating_weakening)

definition roofed_circ :: "('v, 'more) web_scheme ⇒ 'v set ⇒ 'v set" ("RFı")
where "roofed_circ Γ S = roofed Γ S - ℰΓ S"

lemma roofed_circI: fixes Γ (structure) shows
  "⟦ x ∈ RF T; x ∈ T ⟹ ¬ essential Γ (B Γ) T x ⟧ ⟹ x ∈ RF T"
by(simp add: roofed_circ_def)

lemma roofed_circE:
  fixes Γ (structure)
  assumes "x ∈ RF T"
  obtains "x ∈ RF T" "¬ essential Γ (B Γ) T x"
using assms by(auto simp add: roofed_circ_def intro: in_roofed_essentialD)

lemma ℰ_ℰ: fixes Γ (structure) shows "ℰ (ℰ S) = ℰ S"
by(auto intro: essential_mono)

lemma roofed_circ_essential: fixes Γ (structure) shows "RF (ℰ S) = RF S"
unfolding roofed_circ_def RF_essential ℰ_ℰ ..

lemma essential_RF: fixes B
  shows "essential G B (roofed_gen G B S) = essential G B S"  (is "essential _ _ ?RF = _")
proof(intro ext iffI)
  show "essential G B S x" if "essential G B ?RF x" for x using that
    by(rule essential_mono)(blast intro: roofed_greaterI)
  show "essential G B ?RF x" if "essential G B S x" for x
    using that by(rule essentialE_RF)(erule (1) essentialI, blast)
qed

lemma ℰ_RF: fixes Γ (structure) shows "ℰ (RF S) = ℰ S"
by(auto dest: in_roofed_essentialD simp add: essential_RF intro: roofed_greaterI)

lemma essential_ℰ: fixes Γ (structure) shows "essential Γ (B Γ) (ℰ S) = essential Γ (B Γ) S"
by(subst essential_RF[symmetric])(simp only: RF_essential essential_RF)

lemma RF_in_B: fixes Γ (structure) shows "x ∈ B Γ ⟹ x ∈ RF S ⟷ x ∈ S"
by(auto intro: roofed_greaterI dest: roofedD[OF _ rtrancl_path.base])

lemma RF_circ_edge_forward:
  fixes Γ (structure)
  assumes x: "x ∈ RF S"
  and edge: "edge Γ x y"
  shows "y ∈ RF S"
proof
  fix p z
  assume p: "path Γ y p z" and z: "z ∈ B Γ"
  from x have rf: "x ∈ RF S" and ness: "x ∉ ℰ S" by(auto elim: roofed_circE)
  show "(∃z∈set p. z ∈ S) ∨ y ∈ S"
  proof(cases "∃z'∈set (y # p). z' ∈ S")
    case False
    from edge p have p': "path Γ x (y # p) z" ..
    from roofedD[OF rf this z] False have "x ∈ S" by auto
    moreover have "essential Γ (B Γ) S x" using p' False z by(auto intro!: essentialI)
    ultimately have "x ∈ ℰ S" by simp
    with ness show ?thesis by contradiction
  qed auto
qed

subsection ‹Waves›

inductive wave :: "('v, 'more) web_scheme ⇒ 'v current ⇒ bool"
  for Γ (structure) and f
where
  wave:
  "⟦ separating Γ (TER f);
     ⋀x. x ∉ RF (TER f) ⟹ d_OUT f x = 0 ⟧
  ⟹ wave Γ f"

lemma wave_0 [simp]: "wave Γ zero_current"
by rule simp_all

lemma waveD_separating: "wave Γ f ⟹ separating Γ (TERΓ f)"
by(simp add: wave.simps)

lemma waveD_OUT: "⟦ wave Γ f; x ∉ RFΓ (TERΓ f) ⟧ ⟹ d_OUT f x = 0"
by(simp add: wave.simps)

lemma wave_A_in_RF: fixes Γ (structure)
  shows "⟦ wave Γ f; x ∈ A Γ ⟧ ⟹ x ∈ RF (TER f)"
by(auto intro!: roofedI dest!: waveD_separating separatingD)

lemma wave_not_RF_IN_zero:
  fixes Γ (structure)
  assumes f: "current Γ f"
  and w: "wave Γ f"
  and x: "x ∉ RF (TER f)"
  shows "d_IN f x = 0"
proof -
  from x obtain p z where z: "z ∈ B Γ" and p: "path Γ x p z"
    and bypass: "⋀z. z ∈ set p ⟹ z ∉ TER f" "x ∉ TER f"
    by(clarsimp simp add: roofed_def)
  have "f (y, x) = 0" for y
  proof(cases "edge Γ y x")
    case edge: True
    have "d_OUT f y = 0"
    proof(cases "y ∈ TER f")
      case False
      with z p bypass edge have "y ∉ RF (TER f)"
        by(auto simp add: roofed_def intro: rtrancl_path.step intro!: exI rev_bexI)
      thus "d_OUT f y = 0" by(rule waveD_OUT[OF w])
    qed(auto simp add: SINK.simps)
    moreover have "f (y, x) ≤ d_OUT f y" by (rule d_OUT_ge_point)
    ultimately show ?thesis by simp
  qed(simp add: currentD_outside[OF f])
  then show "d_IN f x = 0" unfolding d_IN_def
    by(simp add: nn_integral_0_iff emeasure_count_space_eq_0)
qed

lemma current_Sup:
  fixes Γ (structure)
  assumes chain: "Complete_Partial_Order.chain op ≤ Y"
  and Y: "Y ≠ {}"
  and current: "⋀f. f ∈ Y ⟹ current Γ f"
  and countable [simp]: "countable (support_flow (Sup Y))"
  shows "current Γ (Sup Y)"
proof(rule, goal_cases)
  case (1 x)
  have "d_OUT (Sup Y) x = (SUP f:Y. d_OUT f x)" using chain Y by(simp add: d_OUT_Sup)
  also have "… ≤ weight Γ x" using 1
    by(intro SUP_least)(auto dest!: current currentD_weight_OUT)
  finally show ?case .
next
  case (2 x)
  have "d_IN (Sup Y) x = (SUP f:Y. d_IN f x)" using chain Y by(simp add: d_IN_Sup)
  also have "… ≤ weight Γ x" using 2
    by(intro SUP_least)(auto dest!: current currentD_weight_IN)
  finally show ?case .
next
  case (3 x)
  have "d_OUT (Sup Y) x = (SUP f:Y. d_OUT f x)" using chain Y by(simp add: d_OUT_Sup)
  also have "… ≤ (SUP f:Y. d_IN f x)" using 3
    by(intro SUP_mono)(auto dest: current currentD_OUT_IN)
  also have "… = d_IN (Sup Y) x" using chain Y by(simp add: d_IN_Sup)
  finally show ?case .
next
  case (4 a)
  have "d_IN (Sup Y) a = (SUP f:Y. d_IN f a)" using chain Y by(simp add: d_IN_Sup)
  also have "… = (SUP f:Y. 0)" using 4 by(intro SUP_cong)(auto dest!: current currentD_IN)
  also have "… = 0" using Y by simp
  finally show ?case .
next
  case (5 b)
  have "d_OUT (Sup Y) b = (SUP f:Y. d_OUT f b)" using chain Y by(simp add: d_OUT_Sup)
  also have "… = (SUP f:Y. 0)" using 5 by(intro SUP_cong)(auto dest!: current currentD_OUT)
  also have "… = 0" using Y by simp
  finally show ?case .
next
  fix e
  assume "e ∉ E"
  from currentD_outside'[OF current this] have "f e = 0" if "f ∈ Y" for f using that by simp
  hence "Sup Y e = (SUP _:Y. 0)" by(auto intro: SUP_cong)
  then show "Sup Y e = 0" using Y by(simp)
qed

lemma wave_lub: -- ‹Lemma 4.3›
  fixes Γ (structure)
  assumes chain: "Complete_Partial_Order.chain op ≤ Y"
  and Y: "Y ≠ {}"
  and wave: "⋀f. f ∈ Y ⟹ wave Γ f"
  and countable [simp]: "countable (support_flow (Sup Y))"
  shows "wave Γ (Sup Y)"
proof
  { fix x y p
    assume p: "path Γ x p y" and y: "y ∈ B Γ"
    def P  "{x} ∪ set p"

    let ?f = "λf. SINK f ∩ P"
    have "Complete_Partial_Order.chain op ⊇ (?f ` Y)" using chain
      by(rule chain_imageI)(auto dest: SINK_mono')
    moreover have "… ⊆ Pow P" by auto
    hence "finite (?f ` Y)" by(rule finite_subset)(simp add: P_def)
    ultimately have "(⋂(?f ` Y)) ∈ ?f ` Y"
      by(rule ccpo.in_chain_finite[OF complete_lattice_ccpo_dual])(simp add: Y)
    then obtain f where f: "f ∈ Y" and eq: "⋂(?f ` Y) = ?f f" by clarify
    hence *: "(⋂f∈Y. SINK f) ∩ P = SINK f ∩ P" by(clarsimp simp add: prod_lub_def Y)+
    { fix g
      assume "g ∈ Y" "f ≤ g"
      with * have "(⋂f∈Y. SINK f) ∩ P = SINK g ∩ P" by(blast dest: SINK_mono')
      then have "TER (Sup Y) ∩ P ⊇ TER g ∩ P"
        using SAT_Sup_upper[OF ‹g ∈ Y›, of Γ] SINK_Sup[OF chain Y countable] by blast }
    with f have "∃f∈Y. ∀g∈Y. g ≥ f ⟶ TER g ∩ P ⊆ TER (Sup Y) ∩ P" by blast }
  note subset = this

  show "separating Γ (TER (Sup Y))"
  proof
    fix x y p
    assume *: "path Γ x p y" "y ∈ B Γ" and "x ∈ A Γ"
    let ?P = "{x} ∪ set p"
    from subset[OF *] obtain f where f:"f ∈ Y"
      and subset: "TER f ∩ ?P ⊆ TER (Sup Y) ∩ ?P" by blast
    from wave[OF f] have "TER f ∩ ?P ≠ {}" using * ‹x ∈ A Γ›
      by(auto simp add: wave.simps dest: separatingD)
    with subset show " (∃z∈set p. z ∈ TER (Sup Y)) ∨ x ∈ TER (Sup Y)" by blast
  qed

  fix x
  assume "x ∉ RF (TER (Sup Y))"
  then obtain p y where y: "y ∈ B Γ"
    and p: "path Γ x p y"
    and ter: "TER (Sup Y) ∩ ({x} ∪ set p) = {}" by(auto simp add: roofed_def)
  let ?P = "{x} ∪ set p"
  from subset[OF p y] obtain f where f: "f ∈ Y"
    and subset: "⋀g. ⟦ g ∈ Y; f ≤ g ⟧ ⟹ TER g ∩ ?P ⊆ TER (Sup Y) ∩ ?P" by blast

  { fix g
    assume g: "g ∈ Y"
    with chain f have "f ≤ g ∨ g ≤ f" by(rule chainD)
    hence "d_OUT g x = 0"
    proof
      assume "f ≤ g"
      from subset[OF g this] ter have "TER g ∩ ?P = {}" by blast
      with p y have "x ∉ RF (TER g)" by(auto simp add: roofed_def)
      with wave[OF g] show ?thesis by(blast elim: wave.cases)
    next
      assume "g ≤ f"
      from subset ter f have "TER f ∩ ?P = {}" by blast
      with y p have "x ∉ RF (TER f)" by(auto simp add: roofed_def)
      with wave[OF f] have "d_OUT f x = 0" by(blast elim: wave.cases)
      moreover have "d_OUT g x ≤ d_OUT f x" using ‹g ≤ f›[THEN le_funD] by(rule d_OUT_mono)
      ultimately show ?thesis by simp
    qed }
  thus "d_OUT (Sup Y) x = 0" using chain Y by(simp add: d_OUT_Sup)
qed

lemma ex_maximal_wave: -- ‹Corollary 4.4›
  fixes Γ (structure)
  assumes countable: "countable E"
  shows "∃f. current Γ f ∧ wave Γ f ∧ (∀w. current Γ w ∧ wave Γ w ∧ f ≤ w ⟶ f = w)"
proof -
  def Field_r  "{f. current Γ f ∧ wave Γ f}"
  def r  "{(f, g). f ∈ Field_r ∧ g ∈ Field_r ∧ f ≤ g}"
  have Field_r: "Field r = Field_r" by(auto simp add: Field_def r_def)

  have "Partial_order r" unfolding order_on_defs
    by(auto intro!: refl_onI transI antisymI simp add: Field_r r_def Field_def)
  hence "∃m∈Field r. ∀a∈Field r. (m, a) ∈ r ⟶ a = m"
  proof(rule Zorns_po_lemma; intro strip)
    fix Y
    assume "Y ∈ Chains r"
    hence Y: "Complete_Partial_Order.chain op ≤ Y"
      and w: "⋀f. f ∈ Y ⟹ wave Γ f"
      and f: "⋀f. f ∈ Y ⟹ current Γ f"
      by(auto simp add: Chains_def r_def chain_def Field_r_def)
    show "∃w ∈ Field r. ∀f ∈ Y. (f, w) ∈ r"
    proof(cases "Y = {}")
      case True
      have "zero_current ∈ Field r" by(simp add: Field_r Field_r_def)
      with True show ?thesis by blast
    next
      case False
      have "support_flow (Sup Y) ⊆ E" by(auto simp add: support_flow_Sup elim!: support_flow.cases dest!: f dest: currentD_outside)
      hence c: "countable (support_flow (Sup Y))" using countable  by(rule countable_subset)
      with Y False f w have "Sup Y ∈ Field r" unfolding Field_r Field_r_def
        by(blast intro: wave_lub current_Sup)
      moreover then have "(f, Sup Y) ∈ r" if "f ∈ Y" for f using w[OF that] f[OF that] that unfolding Field_r
        by(auto simp add: r_def Field_r_def intro: Sup_upper)
      ultimately show ?thesis by blast
    qed
  qed
  thus ?thesis by(simp add: Field_r Field_r_def)(auto simp add: r_def Field_r_def)
qed

lemma essential_leI:
  fixes Γ (structure)
  assumes g: "current Γ g" and w: "wave Γ g"
  and le: "⋀e. f e ≤ g e"
  and x: "x ∈ ℰ (TER g)"
  shows "essential Γ (B Γ) (TER f) x"
proof -
  from x obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ" and distinct: "distinct (x # p)"
    and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TER g)" by(rule ℰ_E_RF) blast
  show ?thesis using p y
  proof
    fix z
    assume "z ∈ set p"
    hence z: "z ∉ RF (TER g)" by(auto dest: bypass)
    with w have OUT: "d_OUT g z = 0" and IN: "d_IN g z = 0" by(rule waveD_OUT wave_not_RF_IN_zero[OF g])+
    with z have "z ∉ A Γ" "weight Γ z > 0" by(auto intro!: roofed_greaterI simp add: SAT.simps SINK.simps)
    moreover from IN d_IN_mono[of f z g, OF le] have "d_IN f z ≤ 0" by(simp)
    ultimately have "z ∉ TER f" by(auto simp add: SAT.simps)
    then show "z = x ∨ z ∉ TER f" by simp
  qed
qed

lemma essential_eq_leI:
  fixes Γ (structure)
  assumes g: "current Γ g" and w: "wave Γ g"
  and le: "⋀e. f e ≤ g e"
  and subset: "ℰ (TER g) ⊆ TER f"
  shows "ℰ (TER f) = ℰ (TER g)"
proof
  show subset: "ℰ (TER g) ⊆ ℰ (TER f)"
  proof
    fix x
    assume x: "x ∈ ℰ (TER g)"
    hence "x ∈ TER f" using subset by blast
    moreover have "essential Γ (B Γ) (TER f) x" using g w le x by(rule essential_leI)
    ultimately show "x ∈ ℰ (TER f)" by simp
  qed

  show "… ⊆ ℰ (TER g)"
  proof
    fix x
    assume x: "x ∈ ℰ (TER f)"
    hence "x ∈ TER f" by auto
    hence "x ∈ RF (TER g)"
    proof(rule contrapos_pp)
      assume x: "x ∉ RF (TER g)"
      with w have OUT: "d_OUT g x = 0" and IN: "d_IN g x = 0" by(rule waveD_OUT wave_not_RF_IN_zero[OF g])+
      with x have "x ∉ A Γ" "weight Γ x > 0" by(auto intro!: roofed_greaterI simp add: SAT.simps SINK.simps)
      moreover from IN d_IN_mono[of f x g, OF le] have "d_IN f x ≤ 0" by(simp)
      ultimately show "x ∉ TER f" by(auto simp add: SAT.simps)
    qed
    moreover have "x ∉ RF (TER g)"
    proof
      assume "x ∈ RF (TER g)"
      hence RF: "x ∈ RF (ℰ (TER g))" and not_E: "x ∉ ℰ (TER g)"
        unfolding RF_essential by(simp_all add: roofed_circ_def)
      from x obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ" and distinct: "distinct (x # p)"
        and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TER f)" by(rule ℰ_E_RF) blast
      from roofedD[OF RF p y] not_E obtain z where "z ∈ set p" "z ∈ ℰ (TER g)" by blast
      with subset bypass[of z] show False by(auto intro: roofed_greaterI)
    qed
    ultimately show "x ∈ ℰ (TER g)" by(simp add: roofed_circ_def)
  qed
qed

subsection ‹Hindrances and looseness›

inductive hindrance_by :: "('v, 'more) web_scheme ⇒ 'v current ⇒ ennreal ⇒ bool"
  for Γ (structure) and f and ε
where
  hindrance_by:
  "⟦ a ∈ A Γ; a ∉ ℰ (TER f); d_OUT f a < weight Γ a; ε < weight Γ a - d_OUT f a ⟧ ⟹ hindrance_by Γ f ε"

inductive hindrance :: "('v, 'more) web_scheme ⇒ 'v current ⇒ bool"
  for Γ (structure) and f
where
  hindrance:
  "⟦ a ∈ A Γ; a ∉ ℰ (TER f); d_OUT f a < weight Γ a ⟧ ⟹ hindrance Γ f"

inductive hindered :: "('v, 'more) web_scheme ⇒ bool"
  for Γ (structure)
where hindered: "⟦ hindrance Γ f; current Γ f; wave Γ f ⟧ ⟹ hindered Γ"

inductive hindered_by :: "('v, 'more) web_scheme ⇒ ennreal ⇒ bool"
  for Γ (structure) and ε
where hindered_by: "⟦ hindrance_by Γ f ε; current Γ f; wave Γ f ⟧ ⟹ hindered_by Γ ε"

lemma hindrance_into_hindrance_by:
  assumes "hindrance Γ f"
  shows "∃ε>0. hindrance_by Γ f ε"
using assms
proof cases
  case (hindrance a)
  let  = "if weight Γ a = ⊤ then 1 else (weight Γ a - d_OUT f a) / 2"
  from ‹d_OUT f a < weight Γ a› have "weight Γ a - d_OUT f a > 0" "weight Γ a ≠ ⊤ ⟹ weight Γ a - d_OUT f a < ⊤"
    by(simp_all add: diff_gr0_ennreal less_top diff_less_top_ennreal)
  from ennreal_mult_strict_left_mono[of 1 2, OF _ this]
  have "0 < ?ε" and "?ε < weight Γ a - d_OUT f a" using ‹d_OUT f a < weight Γ a›
    by(auto intro!: diff_gr0_ennreal simp: ennreal_zero_less_divide divide_less_ennreal)
  with hindrance show ?thesis by(auto intro!: hindrance_by.intros)
qed

lemma hindrance_by_into_hindrance: "hindrance_by Γ f ε ⟹ hindrance Γ f"
by(blast elim: hindrance_by.cases intro: hindrance.intros)

lemma hindrance_conv_hindrance_by: "hindrance Γ f ⟷ (∃ε>0. hindrance_by Γ f ε)"
by(blast intro: hindrance_into_hindrance_by hindrance_by_into_hindrance)

lemma hindered_into_hindered_by: "hindered Γ ⟹ ∃ε>0. hindered_by Γ ε"
by(blast intro: hindered_by.intros elim: hindered.cases dest: hindrance_into_hindrance_by)

lemma hindered_by_into_hindered: "hindered_by Γ ε ⟹ hindered Γ"
by(blast elim: hindered_by.cases intro: hindered.intros dest: hindrance_by_into_hindrance)

lemma hindered_conv_hindered_by: "hindered Γ ⟷ (∃ε>0. hindered_by Γ ε)"
by(blast intro: hindered_into_hindered_by hindered_by_into_hindered)

inductive loose :: "('v, 'more) web_scheme ⇒ bool"
  for Γ
where
  loose: "⟦ ⋀f. ⟦ current Γ f; wave Γ f ⟧ ⟹ f = zero_current; ¬ hindrance Γ zero_current ⟧
  ⟹ loose Γ"

lemma looseD_hindrance: "loose Γ ⟹ ¬ hindrance Γ zero_current"
by(simp add: loose.simps)

lemma looseD_wave:
  "⟦ loose Γ; current Γ f; wave Γ f ⟧ ⟹ f = zero_current"
by(simp add: loose.simps)

lemma loose_unhindered:
  fixes Γ (structure)
  assumes "loose Γ"
  shows "¬ hindered Γ"
apply auto
  apply(erule hindered.cases)
apply(frule (1) looseD_wave[OF assms])
apply simp
using looseD_hindrance[OF assms]
by simp

context
  fixes Γ Γ' :: "('v, 'more) web_scheme"
  assumes [simp]: "edge Γ = edge Γ'" "A Γ = A Γ'" "B Γ = B Γ'"
  and weight_eq: "⋀x. x ∉ A Γ' ⟹ weight Γ x = weight Γ' x"
  and weight_le: "⋀a. a ∈ A Γ' ⟹ weight Γ a ≥ weight Γ' a"
begin

private lemma essential_eq: "essential Γ = essential Γ'"
by(simp add: fun_eq_iff essential_def)

qualified lemma TER_eq: "TERΓ f = TERΓ' f"
apply(auto simp add: SINK.simps SAT.simps)
apply(erule contrapos_np; drule weight_eq; simp)+
done

qualified lemma separating_eq: "separating_gen Γ = separating_gen Γ'"
by(intro ext iffI; rule separating_gen.intros; drule separatingD; simp)

qualified lemma roofed_eq: "⋀B. roofed_gen Γ B S = roofed_gen Γ' B S"
by(simp add: roofed_def)

lemma wave_eq_web: -- ‹Observation 4.6›
  "wave Γ f ⟷ wave Γ' f"
by(simp add: wave.simps separating_eq TER_eq roofed_eq)

lemma current_mono_web: "current Γ' f ⟹ current Γ f"
apply(rule current, simp_all add: currentD_OUT_IN currentD_IN currentD_OUT  currentD_outside')
subgoal for x by(cases "x ∈ A Γ'")(auto dest!: weight_eq weight_le dest: currentD_weight_OUT intro: order_trans)
subgoal for x by(cases "x ∈ A Γ'")(auto dest!: weight_eq weight_le dest: currentD_weight_IN intro: order_trans)
done

lemma hindrance_mono_web: "hindrance Γ' f ⟹ hindrance Γ f"
apply(erule hindrance.cases)
apply(rule hindrance)
  apply simp
 apply(unfold TER_eq, simp add: essential_eq)
apply(auto dest!: weight_le)
done

lemma hindered_mono_web: "hindered Γ' ⟹ hindered Γ"
apply(erule hindered.cases)
apply(rule hindered.intros)
  apply(erule hindrance_mono_web)
 apply(erule current_mono_web)
apply(simp add: wave_eq_web)
done

end

subsection ‹Linkage›

text ‹
  The following definition of orthogonality is stronger than the original definition 3.5 in
  @{cite AharoniBergerGeorgakopoulusPerlsteinSpruessel2011JCT} in that the outflow from any
  @{text "A"}-vertices in the set must saturate the vertex; @{term "S ⊆ SAT Γ f"} is not enough.

  With the original definition of orthogonal current, the reduction from networks to webs fails because
  the induced flow need not saturate edges going out of the source. Consider the network with three
  nodes @{text s}, @{text x}, and @{text t} and edges @{text "(s, x)"} and @{text "(x, t)"} with
  capacity @{text 1}. Then, the corresponding web has the vertices @{text "(s, x)"} and
  @{text "(x, t)"} and one edge from @{text "(s, x)"} to @{text "(x, t)"}. Clearly, the zero current
  @{term [source] zero_current} is a web-flow and @{text "TER zero_current = {(s, x)}"}, which is essential.
  Moreover, @{term [source] zero_current} and @{text "{(s, x)}"} are orthogonal because
  @{term [source] zero_current} trivially saturates @{text "(s, x)"} as this is a vertex in @{text A}.
›
inductive orthogonal_current :: "('v, 'more) web_scheme ⇒ 'v current ⇒ 'v set ⇒ bool"
  for Γ (structure) and f S
where orthogonal_current:
  "⟦ ⋀x. ⟦ x ∈ S; x ∉ A Γ ⟧ ⟹ weight Γ x ≤ d_IN f x;
     ⋀x. ⟦ x ∈ S; x ∈ A Γ; x ∉ B Γ ⟧ ⟹ d_OUT f x = weight Γ x;
    ⋀u v. ⟦ v ∈ RF S; u ∉ RF S ⟧ ⟹ f (u, v) = 0 ⟧
  ⟹ orthogonal_current Γ f S"

lemma orthogonal_currentD_SAT: "⟦ orthogonal_current Γ f S; x ∈ S ⟧ ⟹ x ∈ SAT Γ f"
by(auto elim!: orthogonal_current.cases intro: SAT.intros)

lemma orthogonal_currentD_A: "⟦ orthogonal_current Γ f S; x ∈ S; x ∈ A Γ; x ∉ B Γ ⟧ ⟹ d_OUT f x = weight Γ x"
by(auto elim: orthogonal_current.cases)

lemma orthogonal_currentD_in: "⟦ orthogonal_current Γ f S; v ∈ RFΓ S; u ∉ RFΓ S ⟧ ⟹ f (u, v) = 0"
by(auto elim: orthogonal_current.cases)

inductive linkage :: "('v, 'more) web_scheme ⇒ 'v current ⇒ bool"
  for Γ f
where -- ‹Omit the condition @{const web_flow}›
  linkage: "(⋀x. x ∈ A Γ ⟹ d_OUT f x = weight Γ x) ⟹ linkage Γ f"

lemma linkageD: "⟦ linkage Γ f; x ∈ A Γ ⟧ ⟹ d_OUT f x = weight Γ x"
by(rule linkage.cases)

abbreviation linkable :: "('v, 'more) web_scheme ⇒ bool"
where "linkable Γ ≡ ∃f. web_flow Γ f ∧ linkage Γ f"

subsection ‹Trimming›

context
  fixes Γ :: "('v, 'more) web_scheme" (structure)
  and f :: "'v current"
begin

inductive trimming :: "'v current ⇒ bool"
  for g
where
  trimming:
  -- ‹omits the condition that @{term f} is a wave›
  "⟦ current Γ g; wave Γ g; g ≤ f; ⋀x. ⟦ x ∈ RF (TER f); x ∉ A Γ ⟧ ⟹ KIR g x; ℰ (TER g) - A Γ = ℰ (TER f) - A Γ ⟧
  ⟹ trimming g"

lemma assumes "trimming g"
  shows trimmingD_current: "current Γ g"
  and trimmingD_wave: "wave Γ g"
  and trimmingD_le: "⋀e. g e ≤ f e"
  and trimmingD_KIR: "⟦ x ∈ RF (TER f); x ∉ A Γ ⟧ ⟹ KIR g x"
  and trimmingD_ℰ: "ℰ (TER g) - A Γ = ℰ (TER f) - A Γ"
using assms by(blast elim: trimming.cases dest: le_funD)+

lemma ex_trimming: -- ‹Lemma 4.8›
  assumes f: "current Γ f"
  and w: "wave Γ f"
  and countable: "countable E"
  and weight_finite: "⋀x. weight Γ x ≠ ⊤"
  shows "∃g. trimming g"
proof -
  def F  "{g. current Γ g ∧ wave Γ g ∧ g ≤ f ∧ ℰ (TER g) = ℰ (TER f)}"
  def leq  "restrict_rel F {(g, g'). g' ≤ g}"
  have in_F [simp]: "g ∈ F ⟷ current Γ g ∧ wave Γ g ∧ (∀e. g e ≤ f e) ∧ ℰ (TER g) = ℰ (TER f)" for g
    by(simp add: F_def le_fun_def)

  have f_finite [simp]: "f e ≠ ⊤" for e
  proof(cases e)
    case (Pair x y)
    have "f (x, y) ≤ d_IN f y" by (rule d_IN_ge_point)
    also have "… ≤ weight Γ y" by(rule currentD_weight_IN[OF f])
    also have "… < ⊤" by(simp add: weight_finite less_top[symmetric])
    finally show ?thesis using Pair by simp
  qed

  have chainD: "Inf M ∈ F" if M: "M ∈ Chains leq" and nempty: "M ≠ {}" for M
  proof -
    from nempty obtain g0 where g0: "g0 ∈ M" by auto
    have g0_le_f: "g0 e ≤ f e" and g: "current Γ g0" and w0: "wave Γ g0" for e
      using Chains_FieldD[OF M g0] by(cases e, auto simp add: leq_def)

    have finite_OUT: "d_OUT f x ≠ ⊤" for x using weight_finite[of x]
      by(rule neq_top_trans)(rule currentD_weight_OUT[OF f])
    have finite_IN: "d_IN f x ≠ ⊤" for x using weight_finite[of x]
      by(rule neq_top_trans)(rule currentD_weight_IN[OF f])

    from M have "M ∈ Chains {(g, g'). g' ≤ g}"
      by(rule mono_Chains[THEN subsetD, rotated])(auto simp add: leq_def in_restrict_rel_iff)
    then have chain: "Complete_Partial_Order.chain op ≥ M" by(rule Chains_into_chain)
    hence chain': "Complete_Partial_Order.chain op ≤ M" by(simp add: chain_dual)

    have countable': "countable (support_flow f)"
      using current_support_flow[OF f] by(rule countable_subset)(rule countable)

    have OUT_M: "d_OUT (Inf M) x = (INF g:M. d_OUT g x)" for x using chain' nempty countable' _ finite_OUT
      by(rule d_OUT_Inf)(auto dest!: Chains_FieldD[OF M]  simp add: leq_def)
    have IN_M: "d_IN (Inf M) x = (INF g:M. d_IN g x)" for x using chain' nempty countable' _ finite_IN
      by(rule d_IN_Inf)(auto dest!: Chains_FieldD[OF M]  simp add: leq_def)

    have c: "current Γ (Inf M)" using g
    proof(rule current_leI)
      show "(Inf M) e ≤ g0 e" for e using g0 by(auto intro: INF_lower)
      show "d_OUT (⨅M) x ≤ d_IN (⨅M) x" if "x ∉ A Γ" for x
        by(auto 4 4 simp add: IN_M OUT_M leq_def intro!: INF_mono dest: Chains_FieldD[OF M] intro: currentD_OUT_IN[OF _ that])
    qed

    have INF_le_f: "Inf M e ≤ f e" for e using g0 by(auto intro!: INF_lower2 g0_le_f)
    have eq: "ℰ (TER (Inf M)) = ℰ (TER f)" using f w INF_le_f
    proof(rule essential_eq_leI; intro subsetI)
      fix x
      assume x: "x ∈ ℰ (TER f)"
      hence "x ∈ SINK (Inf M)" using d_OUT_mono[of "Inf M" x f, OF INF_le_f]
        by(auto simp add: SINK.simps)
      moreover from x have "x ∈ SAT Γ g" if "g ∈ M" for g using Chains_FieldD[OF M that] by(auto simp add: leq_def)
      hence "x ∈ SAT Γ (Inf M)" by(auto simp add: SAT.simps IN_M intro!: INF_greatest)
      ultimately show "x ∈ TER (Inf M)" by auto
    qed

    have w': "wave Γ (Inf M)"
    proof
      have "separating Γ (ℰ (TER f))" by(rule separating_essential)(rule waveD_separating[OF w])
      then show "separating Γ (TER (Inf M))" unfolding eq[symmetric] by(rule separating_weakening) auto

      fix x
      assume "x ∉ RF (TER (Inf M))"
      hence "x ∉ RF (ℰ (TER (Inf M)))" unfolding RF_essential .
      hence "x ∉ RF (TER f)" unfolding eq RF_essential .
      hence "d_OUT f x = 0" by(rule waveD_OUT[OF w])
      with d_OUT_mono[of _ x f, OF INF_le_f]
      show "d_OUT (Inf M) x = 0" by (metis le_zero_eq)
    qed
    from c w' INF_le_f eq show ?thesis by simp
  qed

  def trim1 
   "λg. if trimming g then g
        else let z = SOME z. z ∈ RF (TER g) ∧ z ∉ A Γ ∧ ¬ KIR g z;
            factor = d_OUT g z / d_IN g z
          in (λ(y, x). (if x = z then factor else 1) * g (y, x))"

  have increasing: "trim1 g ≤ g ∧ trim1 g ∈ F" if "g ∈ F" for g
  proof(cases "trimming g")
    case True
    thus ?thesis using that by(simp add: trim1_def)
  next
    case False
    let ?P = "λz. z ∈ RF (TER g) ∧ z ∉ A Γ ∧ ¬ KIR g z"
    def z  "Eps ?P"
    from that have g: "current Γ g" and w': "wave Γ g" and le_f: "⋀e. g e ≤ f e"
      and : "ℰ (TER g) = ℰ (TER f)" by(auto simp add: le_fun_def)
    { with False obtain z where z: "z ∈ RF (TER f)" and A: "z ∉ A Γ" and neq: "d_OUT g z ≠ d_IN g z"
        by(auto simp add: trimming.simps le_fun_def)
      from z have "z ∈ RF (ℰ (TER f))" unfolding roofed_circ_essential .
      with  roofed_circ_essential[of Γ "TER g"] have "z ∈ RF (TER g)" by simp
      with A neq have "∃x. ?P x" by auto }
    hence "?P z" unfolding z_def by(rule someI_ex)
    hence RF: "z ∈ RF (TER g)" and A: "z ∉ A Γ" and neq: "d_OUT g z ≠ d_IN g z" by simp_all
    let ?factor = "d_OUT g z / d_IN g z"
    have trim1 [simp]: "trim1 g (y, x) = (if x = z then ?factor else 1) * g (y, x)" for x y
      using False by(auto simp add: trim1_def z_def Let_def)

    from currentD_OUT_IN[OF g A] neq have less: "d_OUT g z < d_IN g z" by auto
    hence "?factor ≤ 1" (is "?factor ≤ _")
      by (auto intro!: divide_le_posI_ennreal simp: zero_less_iff_neq_zero)
    hence le': "?factor * g (y, x) ≤ 1 * g (y, x)" for y x
      by(rule mult_right_mono) simp
    hence le: "trim1 g e ≤ g e" for e by(cases e)simp
    moreover {
      have c: "current Γ (trim1 g)" using g le
      proof(rule current_leI)
        fix x
        assume x: "x ∉ A Γ"
        have "d_OUT (trim1 g) x ≤ d_OUT g x" unfolding d_OUT_def using le' by(auto intro: nn_integral_mono)
        also have "… ≤ d_IN (trim1 g) x"
        proof(cases "x = z")
          case True
          have "d_OUT g x = d_IN (trim1 g) x" unfolding d_IN_def
            using True currentD_weight_IN[OF g, of x] currentD_OUT_IN[OF g x]
            apply (cases "d_IN g x = 0")
            apply(auto simp add: nn_integral_divide nn_integral_cmult d_IN_def[symmetric] ennreal_divide_times)
            apply (subst ennreal_divide_self)
            apply (auto simp: less_top[symmetric] top_unique weight_finite)
            done
          thus ?thesis by simp
        next
          case False
          have "d_OUT g x ≤ d_IN g x" using x by(rule currentD_OUT_IN[OF g])
          also have "… ≤ d_IN (trim1 g) x" unfolding d_IN_def using False by(auto intro!: nn_integral_mono)
          finally show ?thesis .
        qed
        finally show "d_OUT (trim1 g) x ≤ d_IN (trim1 g) x" .
      qed
      moreover have le_f: "trim1 g ≤ f" using le le_f by(blast intro: le_funI order_trans)
      moreover have eq: "ℰ (TER (trim1 g)) = ℰ (TER f)" unfolding [symmetric] using g w' le
      proof(rule essential_eq_leI; intro subsetI)
        fix x
        assume x: "x ∈ ℰ (TER g)"
        hence "x ∈ SINK (trim1 g)" using d_OUT_mono[of "trim1 g" x g, OF le]
          by(auto simp add: SINK.simps)
        moreover from x have "x ≠ z" using RF by(auto simp add: roofed_circ_def)
        hence "d_IN (trim1 g) x = d_IN g x" unfolding d_IN_def by simp
        with ‹x ∈ ℰ (TER g)› have "x ∈ SAT Γ (trim1 g)" by(auto simp add: SAT.simps)
        ultimately show "x ∈ TER (trim1 g)" by auto
      qed
      moreover have "wave Γ (trim1 g)"
      proof
        have "separating Γ (ℰ (TER f))" by(rule separating_essential)(rule waveD_separating[OF w])
        then show "separating Γ (TER (trim1 g))" unfolding eq[symmetric] by(rule separating_weakening) auto

        fix x
        assume "x ∉ RF (TER (trim1 g))"
        hence "x ∉ RF (ℰ (TER (trim1 g)))" unfolding RF_essential .
        hence "x ∉ RF (TER f)" unfolding eq RF_essential .
        hence "d_OUT f x = 0" by(rule waveD_OUT[OF w])
        with d_OUT_mono[of _ x f, OF le_f[THEN le_funD]]
        show "d_OUT (trim1 g) x = 0" by (metis le_zero_eq)
      qed
      ultimately have "trim1 g ∈ F" by(simp add: F_def) }
    ultimately show ?thesis using that by(simp add: le_fun_def del: trim1)
  qed

  have "bourbaki_witt_fixpoint Inf leq trim1" using chainD increasing unfolding leq_def
    by(intro bourbaki_witt_fixpoint_restrict_rel)(auto intro: Inf_greatest Inf_lower)
  then interpret bourbaki_witt_fixpoint Inf leq trim1 .

  have f_Field: "f ∈ Field leq" using f w by(simp add: leq_def)

  def g  "fixp_above f"

  have "g ∈ Field leq" using f_Field unfolding g_def by(rule fixp_above_Field)
  hence le_f: "g ≤ f"
    and g: "current Γ g"
    and w': "wave Γ g"
    and TER: "ℰ (TER g) = ℰ (TER f)" by(auto simp add: leq_def intro: le_funI)

  have "trimming g"
  proof(rule ccontr)
    let ?P = "λx. x ∈ RF (TER g) ∧ x ∉ A Γ ∧ ¬ KIR g x"
    def x  "Eps ?P"
    assume False: "¬ ?thesis"
    hence "∃x. ?P x" using le_f g w' TER
      by(auto simp add: trimming.simps roofed_circ_essential[of Γ "TER g", symmetric] roofed_circ_essential[of Γ "TER f", symmetric])
    hence "?P x" unfolding x_def by(rule someI_ex)
    hence x: "x ∈ RF (TER g)" and A: "x ∉ A Γ" and neq: "d_OUT g x ≠ d_IN g x" by simp_all
    from neq have "∃y. edge Γ y x ∧ g (y, x) > 0"
    proof(rule contrapos_np)
      assume "¬ ?thesis"
      hence "d_IN g x = 0" using currentD_outside[OF g, of _ x]
        by(force simp add: d_IN_def nn_integral_0_iff_AE AE_count_space not_less)
      with currentD_OUT_IN[OF g A] show "KIR g x" by simp
    qed
    then obtain y where y: "edge Γ y x" and gr0: "g (y, x) > 0" by blast

    have [simp]: "g (y, x) ≠ ⊤"
    proof -
      have "g (y, x) ≤ d_OUT g y" by (rule d_OUT_ge_point)
      also have "… ≤ weight Γ y" by(rule currentD_weight_OUT[OF g])
      also have "… < ⊤" by(simp add: weight_finite less_top[symmetric])
      finally show ?thesis by simp
    qed

    from neq have factor: "d_OUT g x / d_IN g x ≠ 1"
      by (simp add: divide_eq_1_ennreal)

    have "trim1 g (y, x) = g (y, x) * (d_OUT g x / d_IN g x)"
      by(clarsimp simp add: False trim1_def Let_def x_def[symmetric] mult.commute)
    moreover have "… ≠ g (y, x) * 1" unfolding ennreal_mult_cancel_left using gr0 factor by auto
    ultimately have "trim1 g (y, x) ≠ g (y, x)" by auto
    hence "trim1 g ≠ g" by(auto simp add: fun_eq_iff)
    moreover have "trim1 g = g" using f_Field unfolding g_def by(rule fixp_above_unfold[symmetric])
    ultimately show False by contradiction
  qed
  then show ?thesis by blast
qed

end

lemma trimming_ℰ:
  fixes Γ (structure)
  assumes w: "wave Γ f" and trimming: "trimming Γ f g"
  shows "ℰ (TER f) = ℰ (TER g)"
proof(rule set_eqI)
  show "x ∈ ℰ (TER f) ⟷ x ∈ ℰ (TER g)" for x
  proof(cases "x ∈ A Γ")
    case False
    thus ?thesis using trimmingD_ℰ[OF trimming] by blast
  next
    case True
    show ?thesis
    proof
      assume x: "x ∈ ℰ (TER f)"
      hence "x ∈ TER g" using d_OUT_mono[of g x f, OF trimmingD_le[OF trimming]] True
        by(simp add: SINK.simps SAT.A)
      moreover from x have "essential Γ (B Γ) (TER f) x" by simp
      then obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ"
        and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TER f)" by(rule essentialE_RF) blast
      from p y have "essential Γ (B Γ) (ℰ (TER g)) x"
      proof(rule essentialI)
        fix z
        assume "z ∈ set p"
        hence z: "z ∉ RF (TER f)" by(rule bypass)
        with waveD_separating[OF w, THEN separating_RF_A] have "z ∉ A Γ" by blast
        with z have "z ∉ ℰ (TER g)" using trimmingD_ℰ[OF trimming] by(auto intro: roofed_greaterI)
        thus "z = x ∨ z ∉ ℰ (TER g)" ..
      qed
      ultimately show "x ∈ ℰ (TER g)" unfolding essential_ℰ by simp
    next
      assume "x ∈ ℰ (TER g)"
      then obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ"
        and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TER g)" by(rule ℰ_E_RF) blast
      have z: "z ∉ ℰ (TER f)" if "z ∈ set p" for z
      proof -
        from that have z: "z ∉ RF (TER g)" by(rule bypass)
        with waveD_separating[OF trimmingD_wave[OF trimming], THEN separating_RF_A] have "z ∉ A Γ" by blast
        with z show "z ∉ ℰ (TER f)" using trimmingD_ℰ[OF trimming] by(auto intro: roofed_greaterI)
      qed
      then have "essential Γ (B Γ) (ℰ (TER f)) x" by(intro essentialI[OF p y]) auto
      moreover have "x ∈ TER f"
        using waveD_separating[THEN separating_essential, THEN separatingD, OF w p True y] z
        by auto
      ultimately show "x ∈ ℰ (TER f)" unfolding essential_ℰ by simp
    qed
  qed
qed

subsection ‹Composition of waves via quotients›

definition quotient_web :: "('v, 'more) web_scheme ⇒ 'v current ⇒ ('v, 'more) web_scheme"
where -- ‹Modifications to original Definition 4.9: No incoming edges to nodes in @{const A},
  @{term "B Γ - A Γ"} is not part of @{const A} such that @{const A} contains only vertices
  is disjoint from @{const B}. The weight of vertices in @{const B} saturated by @{term f} is
  therefore set to @{term "0 :: ennreal"}.›
  "quotient_web Γ f =
   ⦇edge = λx y. edge Γ x y ∧ x ∉ roofed_circ Γ (TERΓ f) ∧ y ∉ roofed Γ (TERΓ f),
    weight = λx. if x ∈ RFΓ (TERΓ f) ∨ x ∈ TERΓ f ∩ B Γ then 0 else weight Γ x,
    A = ℰΓ (TERΓ f) - (B Γ - A Γ),
    B = B Γ,
    … = web.more Γ⦈"

lemma quotient_web_sel [simp]:
  fixes Γ (structure) shows
  "edge (quotient_web Γ f) x y ⟷ edge Γ x y ∧ x ∉ RF (TER f) ∧ y ∉ RF (TER f)"
  "weight (quotient_web Γ f) x = (if x ∈ RF (TER f) ∨ x ∈ TERΓ f ∩ B Γ then 0 else weight Γ x)"
  "A (quotient_web Γ f) = ℰ (TER f)- (B Γ - A Γ)"
  "B (quotient_web Γ f) = B Γ"
  "web.more (quotient_web Γ f) = web.more Γ"
by(simp_all add: quotient_web_def)

lemma vertex_quotient_webD: fixes Γ (structure) shows
  "vertex (quotient_web Γ f) x ⟹ vertex Γ x ∧ x ∉ RF (TER f)"
by(auto simp add: vertex_def roofed_circ_def)

lemma path_quotient_web:
  fixes Γ (structure)
  assumes "path Γ x p y"
  and "x ∉ RF (TER f)"
  and "⋀z. z ∈ set p ⟹ z ∉ RF (TER f)"
  shows "path (quotient_web Γ f) x p y"
using assms by(induction)(auto intro: rtrancl_path.intros simp add: roofed_circ_def)

definition restrict_current :: "('v, 'more) web_scheme ⇒ 'v current ⇒ 'v current ⇒ 'v current"
where "restrict_current Γ f g = (λ(x, y). g (x, y) * indicator (- RFΓ (TERΓ f)) x * indicator (- RFΓ (TERΓ f)) y)"

abbreviation restrict_curr :: "'v current ⇒ ('v, 'more) web_scheme ⇒ 'v current ⇒ 'v current" ("_ ↿ _ '/ _" [100, 0, 100] 100)
where "restrict_curr g Γ f ≡ restrict_current Γ f g"

lemma restrict_current_simps [simp]: fixes Γ (structure) shows
  "(g ↿ Γ / f) (x, y) = (g (x, y) * indicator (- RF (TER f)) x * indicator (- RF (TER f)) y)"
by(simp add: restrict_current_def)

lemma d_OUT_restrict_current_outside: fixes Γ (structure) shows
  "x ∈ RF (TER f) ⟹ d_OUT (g ↿ Γ / f) x = 0"
by(simp add: d_OUT_def)

lemma d_IN_restrict_current_outside: fixes Γ (structure) shows
  "x ∈ RF (TER f) ⟹ d_IN (g ↿ Γ / f) x = 0"
by(simp add: d_IN_def)

lemma restrict_current_le: "(g ↿ Γ / f) e ≤ g e"
by(cases e)(clarsimp split: split_indicator)

lemma d_OUT_restrict_current_le: "d_OUT (g ↿ Γ / f) x ≤ d_OUT g x"
unfolding d_OUT_def by(rule nn_integral_mono, simp split: split_indicator)

lemma d_IN_restrict_current_le: "d_IN (g ↿ Γ / f) x ≤ d_IN g x"
unfolding d_IN_def by(rule nn_integral_mono, simp split: split_indicator)

lemma restrict_current_IN_not_RF:
  fixes Γ (structure)
  assumes g: "current Γ g"
  and x: "x ∉ RF (TER f)"
  shows "d_IN (g ↿ Γ / f) x = d_IN g x"
proof -
  {
    fix y
    assume y: "y ∈ RF (TER f)"
    have "g (y, x) = 0"
    proof(cases "edge Γ y x")
      case True
      from y have y': "y ∈ RF (TER f)" and essential: "y ∉ ℰ (TER f)" by(simp_all add: roofed_circ_def)
      moreover from x obtain p z where z: "z ∈ B Γ" and p: "path Γ x p z"
        and bypass: "⋀z. z ∈ set p ⟹ z ∉ TER f" "x ∉ TER f"
        by(clarsimp simp add: roofed_def)
      from roofedD[OF y' rtrancl_path.step, OF True p z] bypass have "x ∈ TER f ∨ y ∈ TER f" by auto
      with roofed_greater[THEN subsetD, of x "TER f" Γ] x have "x ∉ TER f" "y ∈ TER f" by auto
      with essential bypass have False
        by(auto dest!: not_essentialD[OF _ rtrancl_path.step, OF _ True p z])
      thus ?thesis ..
    qed(simp add: currentD_outside[OF g]) }
  then show ?thesis unfolding d_IN_def
    using x by(auto intro!: nn_integral_cong split: split_indicator)
qed

lemma restrict_current_IN_A:
  "a ∈ A (quotient_web Γ f) ⟹ d_IN (g ↿ Γ / f) a = 0"
by(simp add: d_IN_restrict_current_outside roofed_greaterI)

lemma restrict_current_nonneg: "0 ≤ g e ⟹ 0 ≤ (g ↿ Γ / f) e"
by(cases e) simp

lemma in_SINK_restrict_current: "x ∈ SINK g ⟹ x ∈ SINK (g ↿ Γ / f)"
using d_OUT_restrict_current_le[of Γ f g x]
by(simp add: SINK.simps)

lemma SAT_restrict_current:
  fixes Γ (structure)
  assumes f: "current Γ f"
  and g: "current Γ g"
  shows "SAT (quotient_web Γ f) (g ↿ Γ / f) = RF (TER f) ∪ (SAT Γ g - A Γ)" (is "SAT ?Γ ?g = ?rhs")
proof(intro set_eqI iffI; (elim UnE DiffE)?)
  show "x ∈ ?rhs" if "x ∈ SAT ?Γ ?g" for x using that
  proof cases
    case IN
    thus ?thesis using currentD_weight_OUT[OF f, of x]
      by(cases "x ∈ RF (TER f)")(auto simp add: d_IN_restrict_current_outside roofed_circ_def restrict_current_IN_not_RF[OF g] SAT.IN currentD_IN[OF g] roofed_greaterI SAT.A SINK.simps RF_in_B split: if_split_asm intro: essentialI[OF rtrancl_path.base])
  qed(simp add: roofed_greaterI)
  show "x ∈ SAT ?Γ ?g" if "x ∈ RF (TER f)" for x using that
    by(auto simp add: SAT.simps roofed_circ_def d_IN_restrict_current_outside)
  show "x ∈ SAT ?Γ ?g" if "x ∈ SAT Γ g" "x ∉ A Γ" for x using that
    by(auto simp add: SAT.simps roofed_circ_def d_IN_restrict_current_outside restrict_current_IN_not_RF[OF g])
qed

lemma current_restrict_current:
  fixes Γ (structure)
  assumes w: "wave Γ f"
  and g: "current Γ g"
  shows "current (quotient_web Γ f) (g ↿ Γ / f)" (is "current ?Γ ?g")
proof
  show "d_OUT ?g x ≤ weight ?Γ x" for x
    using d_OUT_restrict_current_le[of Γ f g x] currentD_weight_OUT[OF g, of x] currentD_OUT[OF g, of x]
    by(auto simp add: d_OUT_restrict_current_outside)
  show "d_IN ?g x ≤ weight ?Γ x" for x
    using d_IN_restrict_current_le[of Γ f g x] currentD_weight_IN[OF g, of x]
    by(auto simp add: d_IN_restrict_current_outside roofed_circ_def)
      (subst d_IN_restrict_current_outside[of x Γ f g]; simp add: roofed_greaterI)

  fix x
  assume "x ∉ A ?Γ"
  hence x: "x ∉ ℰ (TER f) - B Γ" by simp
  show "d_OUT ?g x ≤ d_IN ?g x"
  proof(cases "x ∈ RF (TER f)")
    case True
    with x have "x ∈ RF (TER f) ∪ B Γ" by(simp add: roofed_circ_def)
    with True show ?thesis using currentD_OUT[OF g, of x] d_OUT_restrict_current_le[of Γ f g x]
      by(auto simp add: d_OUT_restrict_current_outside d_IN_restrict_current_outside)
  next
    case False
    then obtain p z where z: "z ∈ B Γ" and p: "path Γ x p z"
      and bypass: "⋀z. z ∈ set p ⟹ z ∉ TER f" "x ∉ TER f"
      by(clarsimp simp add: roofed_def)
    from g False have "d_IN ?g x = d_IN g x" by(rule restrict_current_IN_not_RF)
    moreover have "d_OUT ?g x ≤ d_OUT g x"
      by(rule d_OUT_mono restrict_current_le)+
    moreover have "x ∉ A Γ"
      using separatingD[OF waveD_separating[OF w] p _ z] bypass by blast
    note currentD_OUT_IN[OF g this]
    ultimately show ?thesis by simp
  qed
next
  show "d_IN ?g a = 0" if "a ∈ A ?Γ" for a using that by(rule restrict_current_IN_A)
  show "d_OUT ?g b = 0" if "b ∈ B ?Γ" for b
    using d_OUT_restrict_current_le[of Γ f g b] currentD_OUT[OF g, of b] that by simp
  show "?g e = 0" if "e ∉ E" for e using that currentD_outside'[OF g, of e]
    by(cases e)(auto split: split_indicator)
qed

lemma TER_restrict_current:
  fixes Γ (structure)
  assumes f: "current Γ f"
  and w: "wave Γ f"
  and g: "current Γ g"
  shows "TER g ⊆ TERquotient_web Γ f (g ↿ Γ / f)" (is "_ ⊆ ?TER" is "_ ⊆ TER ?g")
proof
  fix x
  assume x: "x ∈ TER g"
  hence "x ∈ SINK ?g" by(simp add: in_SINK_restrict_current)
  moreover have "x ∈ RF (TER f)" if "x ∈ A Γ"
    using waveD_separating[OF w, THEN separatingD, OF _ that] by(rule roofedI)
  then have "x ∈ SAT ?Γ ?g" using SAT_restrict_current[OF f g] x by auto
  ultimately show "x ∈ ?TER" by simp
qed

lemma wave_restrict_current:
  fixes Γ (structure)
  assumes f: "current Γ f"
  and w: "wave Γ f"
  and g: "current Γ g"
  and w': "wave Γ g"
  shows "wave (quotient_web Γ f) (g ↿ Γ / f)" (is "wave ?Γ ?g")
proof
  show "separating ?Γ (TER ?g)" (is "separating _ ?TER")
  proof
    fix x y p
    assume "x ∈ A ?Γ" "y ∈ B ?Γ" and p: "path ?Γ x p y"
    hence x: "x ∈ ℰ (TER f)" and y: "y ∈ B Γ" and SAT: "x ∈ SAT ?Γ ?g" by(simp_all add: SAT.A)
    from p have p': "path Γ x p y" by(rule rtrancl_path_mono) simp

    { assume "x ∉ ?TER"
      hence "x ∉ SINK ?g" using SAT by(simp)
      hence "x ∉ SINK g" using d_OUT_restrict_current_le[of Γ f g x]
        by(auto simp add: SINK.simps)
      hence "x ∈ RF (TER g)" using waveD_OUT[OF w'] by(auto simp add: SINK.simps)
      from roofedD[OF this p' y] ‹x ∉ SINK g› have "(∃z∈set p. z ∈ ?TER)"
        using TER_restrict_current[OF f w g] by blast }
    then show "(∃z∈set p. z ∈ ?TER) ∨ x ∈ ?TER" by blast
  qed

  fix x
  assume "x ∉ RF ?TER"
  hence "x ∉ RF (TER g)"
  proof(rule contrapos_nn)
    assume *: "x ∈ RF (TER g)"
    show "x ∈ RF ?TER"
    proof
      fix p y
      assume "path ?Γ x p y" "y ∈ B ?Γ"
      hence "path Γ x p y" "y ∈ B Γ" by(auto elim: rtrancl_path_mono)
      from roofedD[OF * this] show "(∃z∈set p. z ∈ ?TER) ∨ x ∈ ?TER"
        using TER_restrict_current[OF f w g] by blast
    qed
  qed
  with w' have "d_OUT g x = 0" by(rule waveD_OUT)
  with d_OUT_restrict_current_le[of Γ f g x]
  show "d_OUT ?g x = 0" by simp
qed

definition plus_current :: "'v current ⇒ 'v current ⇒ 'v current"
where "plus_current f g = (λe. f e + g e)"

lemma plus_current_simps [simp]: "plus_current f g e = f e + g e"
by(simp add: plus_current_def)

lemma plus_zero_current [simp]: "plus_current f zero_current = f"
by(simp add: fun_eq_iff)

lemma support_flow_plus_current: "support_flow (plus_current f g) ⊆ support_flow f ∪ support_flow g"
by(clarsimp simp add: support_flow.simps)

lemma SINK_plus_current: "SINK (plus_current f g) = SINK f ∩ SINK g"
by(auto simp add: SINK.simps set_eq_iff d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 add_eq_0_iff_both_eq_0)

abbreviation plus_web :: "('v, 'more) web_scheme ⇒ 'v current ⇒ 'v current ⇒ 'v current" ("_ ⌢ı _" [66, 66] 65)
where "plus_web Γ f g ≡ plus_current f (g ↿ Γ / f)"

context
  fixes Γ :: "('v, 'more) web_scheme" (structure) and f g
  assumes f: "current Γ f"
  and w: "wave Γ f"
  and g: "current (quotient_web Γ f) g"
begin

lemma OUT_plus_current: "d_OUT (plus_current f g) x = (if x ∈ RF (TER f) then d_OUT f x else d_OUT g x)" (is "d_OUT ?g _ = _")
proof -
  have "d_OUT ?g x = d_OUT f x + d_OUT g x" unfolding plus_current_def
    by(subst d_OUT_add) simp_all
  also have "… = (if x ∈ RF (TER f) then d_OUT f x else d_OUT g x)"
  proof(cases "x ∈ RF (TER f)")
    case True
    hence "d_OUT g x = 0" by(intro currentD_outside_OUT[OF g])(auto dest: vertex_quotient_webD)
    thus ?thesis using True by simp
  next
    case False
    hence "d_OUT f x = 0" by(auto simp add: roofed_circ_def SINK.simps dest: waveD_OUT[OF w])
    with False show ?thesis by simp
  qed
  finally show ?thesis .
qed

lemma IN_plus_current: "d_IN (plus_current f g) x = (if x ∈ RF (TER f) then d_IN f x else d_IN g x)" (is "d_IN ?g _ = _")
proof -
  have "d_IN ?g x = d_IN f x + d_IN g x" unfolding plus_current_def
    by(subst d_IN_add) simp_all
  also consider (RF) "x ∈ RF (TER f) - (B Γ - A Γ)" | (B) "x ∈ RF (TER f)" "x ∈ B Γ - A Γ" | (beyond) "x ∉ RF (TER f)" by blast
  then have "d_IN f x + d_IN g x = (if x ∈ RF (TER f) then d_IN f x else d_IN g x)"
  proof(cases)
    case RF
    hence "d_IN g x = 0"
      by(cases "x ∈ ℰ (TER f)")(auto intro: currentD_outside_IN[OF g] currentD_IN[OF g] dest!: vertex_quotient_webD simp add: roofed_circ_def)
    thus ?thesis using RF by simp
  next
    case B
    hence "d_IN g x = 0" using currentD_outside_IN[OF g, of x] currentD_weight_IN[OF g, of x]
      by(auto dest: vertex_quotient_webD simp add: roofed_circ_def)
    with B show ?thesis by simp
  next
    case beyond
    from f w beyond have "d_IN f x = 0" by(rule wave_not_RF_IN_zero)
    with beyond show ?thesis by simp
  qed
  finally show ?thesis .
qed

lemma current_plus_current: "current Γ (plus_current f g)" (is "current _ ?g")
proof
  show "d_OUT ?g x ≤ weight Γ x" for x
    using currentD_weight_OUT[OF g, of x] currentD_weight_OUT[OF f, of x]
    by(auto simp add: OUT_plus_current split: if_split_asm elim: order_trans)
  show "d_IN ?g x ≤ weight Γ x" for x
    using currentD_weight_IN[OF f, of x] currentD_weight_IN[OF g, of x]
    by(auto simp add: IN_plus_current roofed_circ_def split: if_split_asm elim: order_trans)
  show "d_OUT ?g x ≤ d_IN ?g x" if "x ∉ A Γ" for x
  proof(cases "x ∈ ℰ (TER f)")
    case False
    thus ?thesis
      using currentD_OUT_IN[OF f that] currentD_OUT_IN[OF g, of x] that
      by(auto simp add: OUT_plus_current IN_plus_current roofed_circ_def SINK.simps)
  next
    case True
    with that have "d_OUT f x = 0" "weight Γ x ≤ d_IN f x"
      by(auto simp add: SINK.simps elim: SAT.cases)
    thus ?thesis using that True currentD_OUT_IN[OF g, of x] currentD_weight_OUT[OF g, of x]
      by(auto simp add: OUT_plus_current IN_plus_current roofed_circ_def intro: roofed_greaterI split: if_split_asm)
  qed
  show "d_IN ?g a = 0" if "a ∈ A Γ" for a
    using wave_A_in_RF[OF w that] currentD_IN[OF f that] by(simp add: IN_plus_current)
  show "d_OUT ?g b = 0" if "b ∈ B Γ" for b
    using that currentD_OUT[OF f that] currentD_OUT[OF g, of b] that
    by(auto simp add: OUT_plus_current SINK.simps roofed_circ_def intro: roofed_greaterI)
  show "?g e = 0" if "e ∉ E" for e using currentD_outside'[OF f, of e] currentD_outside'[OF g, of e] that
    by(cases e) auto
qed

lemma in_TER_plus_current:
  assumes RF: "x ∉ RF (TER f)"
  and x: "x ∈ TERquotient_web Γ f g" (is "_ ∈ ?TER _")
  shows "x ∈ TER (plus_current f g)"  (is "_ ∈ TER ?g")
proof(cases "x ∈ ℰ (TER f) - (B Γ - A Γ)")
  case True
  with x show ?thesis using currentD_IN[OF g, of x]
    by(fastforce intro: roofed_greaterI SAT.intros simp add: SINK.simps OUT_plus_current IN_plus_current elim!: SAT.cases)
next
  case *: False
  have "x ∈ SAT Γ ?g"
  proof(cases "x ∈ B Γ - A Γ")
    case False
    with x RF * have "weight Γ x ≤ d_IN g x"
      by(auto elim!: SAT.cases split: if_split_asm simp add: essential_BI)
    also have "… ≤ d_IN ?g x" unfolding plus_current_def by(intro d_IN_mono) simp
    finally show ?thesis ..
  next
    case True
    with * x have "weight Γ x ≤ d_IN ?g x" using currentD_OUT[OF f, of x]
      by(auto simp add: IN_plus_current RF_in_B SINK.simps roofed_circ_def elim!: SAT.cases split: if_split_asm)
    thus ?thesis ..
  qed
  moreover have "x ∈ SINK ?g" using x by(simp add: SINK.simps OUT_plus_current RF)
  ultimately show ?thesis by simp
qed

context
  assumes w': "wave (quotient_web Γ f) g"
begin

lemma separating_TER_plus_current:
  assumes x: "x ∈ RF (TER f)" and y: "y ∈ B Γ" and p: "path Γ x p y"
  shows "(∃z∈set p. z ∈ TER (plus_current f g)) ∨ x ∈ TER (plus_current f g)" (is "_ ∨ _ ∈ TER ?g")
proof -
  from x have "x ∈ RF (ℰ (TER f))" unfolding RF_essential .
  from roofedD[OF this p y] have "∃z∈set (x # p). z ∈ ℰ (TER f)" by auto
  from split_list_last_prop[OF this] obtain ys z zs
    where decomp: "x # p = ys @ z # zs" and z: "z ∈ ℰ (TER f)"
    and outside: "⋀z. z ∈ set zs ⟹ z ∉ ℰ (TER f)" by auto
  have zs: "path Γ z zs y" using decomp p
    by(cases ys)(auto elim: rtrancl_path_appendE)
  moreover have "z ∉ RF (TER f)" using z by(simp add: roofed_circ_def)
  moreover have RF: "z' ∉ RF (TER f)" if "z' ∈ set zs" for z'
  proof
    assume "z' ∈ RF (TER f)"
    hence z': "z' ∈ RF (ℰ (TER f))" by(simp only: RF_essential)
    from split_list[OF that] obtain ys' zs' where decomp': "zs = ys' @ z' # zs'" by blast
    with zs have "path Γ z' zs' y" by(auto elim: rtrancl_path_appendE)
    from roofedD[OF z' this y] outside decomp' show False by auto
  qed
  ultimately have p': "path (quotient_web Γ f) z zs y" by(rule path_quotient_web)
  show ?thesis
  proof(cases "z ∈ B Γ - A Γ")
    case False
    with separatingD[OF waveD_separating[OF w'] p'] z y
    obtain z' where z': "z' ∈ set (z # zs)" and TER: "z' ∈ TERquotient_web Γ f g" by auto
    hence "z' ∈ TER ?g" using in_TER_plus_current[of z'] RF[of z'] ‹z ∉ RF (TER f)› by(auto simp add: roofed_circ_def)
    with decomp z' show ?thesis by(cases ys) auto
  next
    case True
    hence "z ∈ TER ?g" using currentD_OUT[OF current_plus_current, of z] z
      by(auto simp add: SINK.simps SAT.simps IN_plus_current intro: roofed_greaterI)
    then show ?thesis using decomp by(cases ys) auto
  qed
qed

lemma wave_plus_current: "wave Γ (plus_current f g)" (is "wave _ ?g")
proof
  let  = "quotient_web Γ f"
  let ?TER = "TER"

  show "separating Γ (TER ?g)" using separating_TER_plus_current[OF wave_A_in_RF[OF w]] by(rule separating)

  fix x
  assume x: "x ∉ RF (TER ?g)"
  hence "x ∉ RF (TER f)" by(rule contrapos_nn)(rule roofedI, rule separating_TER_plus_current)
  hence *: "x ∉ RF (TER f)" by(simp add: roofed_circ_def)
  moreover have "x ∉ RF (?TER g)"
  proof
    assume RF': "x ∈ RF (?TER g)"
    from x obtain p y where y: "y ∈ B Γ" and p: "path Γ x p y"
      and bypass: "⋀z. z ∈ set p ⟹ z ∉ TER ?g" and x': "x ∉ TER ?g"
      by(auto simp add: roofed_def)
    have RF: "z ∉ RF (TER f)" if "z ∈ set p" for z
    proof
      assume z: "z ∈ RF (TER f)"
      from split_list[OF that] obtain ys' zs' where decomp: "p = ys' @ z # zs'" by blast
      with p have "path Γ z zs' y" by(auto elim: rtrancl_path_appendE)
      from separating_TER_plus_current[OF z y this] decomp bypass show False by auto
    qed
    with p have "path ?Γ x p y" using *
      by(induction)(auto intro: rtrancl_path.intros simp add: roofed_circ_def)
    from roofedD[OF RF' this] y consider (x) "x ∈ ?TER g" | (z) z where "z ∈ set p" "z ∈ ?TER g" by auto
    then show False
    proof(cases)
      case x
      with * have "x ∈ TER ?g" by(rule in_TER_plus_current)
      with x' show False by contradiction
    next
      case (z z)
      from z(1) have "z ∉ RF (TER f)" by(rule RF)
      hence "z ∉ RF (TER f)" by(simp add: roofed_circ_def)
      hence "z ∈ TER ?g" using z(2) by(rule in_TER_plus_current)
      moreover from z(1) have "z ∉ TER ?g" by(rule bypass)
      ultimately show False by contradiction
    qed
  qed
  with w' have "d_OUT g x = 0" by(rule waveD_OUT)
  ultimately show "d_OUT ?g x = 0" by(simp add: OUT_plus_current)
qed

end

end

lemma d_OUT_plus_web:
  fixes Γ (structure)
  shows "d_OUT (f ⌢ g) x = d_OUT f x + d_OUT (g ↿ Γ / f) x" (is "?lhs = ?rhs")
proof -
  have "?lhs = d_OUT f x + (∑+ y. (if x ∈ RF (TER f) then 0 else g (x, y) * indicator (- RF (TER f)) y))"
    unfolding d_OUT_def by(subst nn_integral_add[symmetric])(auto intro!: nn_integral_cong split: split_indicator)
  also have "… = ?rhs" by(auto simp add: d_OUT_def intro!: arg_cong2[where f="op +"] nn_integral_cong)
  finally show "?thesis" .
qed

lemma d_IN_plus_web:
  fixes Γ (structure)
  shows "d_IN (f ⌢ g) y = d_IN f y + d_IN (g ↿ Γ / f) y" (is "?lhs = ?rhs")
proof -
  have "?lhs = d_IN f y + (∑+ x. (if y ∈ RF (TER f) then 0 else g (x, y) * indicator (- RF (TER f)) x))"
    unfolding d_IN_def by(subst nn_integral_add[symmetric])(auto intro!: nn_integral_cong split: split_indicator)
  also have "… = ?rhs" by(auto simp add: d_IN_def intro!: arg_cong2[where f="op +"] nn_integral_cong)
  finally show ?thesis .
qed

lemma plus_web_greater: "f e ≤ (f ⌢Γ g) e"
by(cases e)(auto split: split_indicator)

lemma current_plus_web:
  fixes Γ (structure)
  shows "⟦ current Γ f; wave Γ f; current Γ g ⟧ ⟹ current Γ (f ⌢ g)"
by(blast intro: current_plus_current current_restrict_current)

context
  fixes Γ :: "('v, 'more) web_scheme" (structure)
  and f g :: "'v current"
  assumes f: "current Γ f"
  and w: "wave Γ f"
  and g: "current Γ g"
begin

context
  fixes x :: "'v"
  assumes x: "x ∈ ℰ (TER f ∪ TER g)"
begin

qualified lemma RF_f: "x ∉ RF (TER f)"
proof
  assume *: "x ∈ RF (TER f)"

  from x obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ"
    and bypass: "⋀z. ⟦x ≠ y; z ∈ set p⟧ ⟹ z = x ∨ z ∉ TER f ∪ TER g" by(rule ℰ_E) blast
  from rtrancl_path_distinct[OF p] obtain p'
    where p: "path Γ x p' y" and p': "set p' ⊆ set p" and distinct: "distinct (x # p')" .

  from * have x': "x ∈ RF (TER f)" and : "x ∉ ℰ (TER f)" by(auto simp add: roofed_circ_def)
  hence "x ∉ TER f" using not_essentialD[OF _ p y] p' bypass by blast
  with roofedD[OF x' p y] obtain z where z: "z ∈ set p'" "z ∈ TER f" by auto
  with p have "y ∈ set p'" by(auto dest!: rtrancl_path_last intro: last_in_set)
  with distinct have "x ≠ y" by auto
  with bypass z p' distinct show False by auto
qed

private lemma RF_g: "x ∉ RF (TER g)"
proof
  assume *: "x ∈ RF (TER g)"

  from x obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ"
    and bypass: "⋀z. ⟦x ≠ y; z ∈ set p⟧ ⟹ z = x ∨ z ∉ TER f ∪ TER g" by(rule ℰ_E) blast
  from rtrancl_path_distinct[OF p] obtain p'
    where p: "path Γ x p' y" and p': "set p' ⊆ set p" and distinct: "distinct (x # p')" .

  from * have x': "x ∈ RF (TER g)" and : "x ∉ ℰ (TER g)" by(auto simp add: roofed_circ_def)
  hence "x ∉ TER g" using not_essentialD[OF _ p y] p' bypass by blast
  with roofedD[OF x' p y] obtain z where z: "z ∈ set p'" "z ∈ TER g" by auto
  with p have "y ∈ set p'" by(auto dest!: rtrancl_path_last intro: last_in_set)
  with distinct have "x ≠ y" by auto
  with bypass z p' distinct show False by auto
qed

lemma TER_plus_web_aux:
  assumes SINK: "x ∈ SINK (g ↿ Γ / f)" (is "_ ∈ SINK ?g")
  shows "x ∈ TER (f ⌢ g)"
proof
  from x obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ"
    and bypass: "⋀z. ⟦x ≠ y; z ∈ set p⟧ ⟹ z = x ∨ z ∉ TER f ∪ TER g" by(rule ℰ_E) blast
  from rtrancl_path_distinct[OF p] obtain p'
    where p: "path Γ x p' y" and p': "set p' ⊆ set p" and distinct: "distinct (x # p')" .

  from RF_f have "x ∈ SINK f"
    by(auto simp add: roofed_circ_def SINK.simps dest: waveD_OUT[OF w])
  thus "x ∈ SINK (f ⌢ g)" using SINK
    by(simp add: SINK.simps d_OUT_plus_web)
  show "x ∈ SAT Γ (f ⌢ g)"
  proof(cases "x ∈ TER f")
    case True
    hence "x ∈ SAT Γ f" by simp
    moreover have "… ⊆ SAT Γ (f ⌢ g)" by(rule SAT_mono plus_web_greater)+
    ultimately show ?thesis by blast
  next
    case False
    with x have "x ∈ TER g" by auto
    from False RF_f have "x ∉ RF (TER f)" by(auto simp add: roofed_circ_def)
    moreover { fix z
      assume z: "z ∈ RF (TER f)"
      have "(z, x) ∉ E"
      proof
        assume "(z, x) ∈ E"
        hence path': "path Γ z (x # p') y" using p by(simp add: rtrancl_path.step)
        from z have "z ∈ RF (TER f)" by(simp add: roofed_circ_def)
        from roofedD[OF this path' y] False
        consider (path) z' where  "z' ∈ set p'" "z' ∈ TER f" | (TER) "z ∈ TER f" by auto
        then show False
        proof cases
          { case (path z')
            with p distinct have "x ≠ y"
              by(auto 4 3 intro: last_in_set elim: rtrancl_path.cases dest: rtrancl_path_last[symmetric])
            from bypass[OF this, of z'] path False p' show False by auto }
          note that = this
          case TER
          with z have "¬ essential Γ (B Γ) (TER f) z" by(simp add: roofed_circ_def)
          from not_essentialD[OF this path' y] False obtain z' where "z' ∈ set p'" "z' ∈ TER f" by auto
          thus False by(rule that)
        qed
      qed }
    ultimately have "d_IN ?g x = d_IN g x" unfolding d_IN_def
      by(intro nn_integral_cong)(clarsimp split: split_indicator simp add: currentD_outside[OF g])
    hence "d_IN (f ⌢ g) x ≥ d_IN g x"
      by(simp add: d_IN_plus_web)
    with ‹x ∈ TER g› show ?thesis by(auto elim!: SAT.cases intro: SAT.intros)
  qed
qed

qualified lemma SINK_TER_in'':
  assumes "⋀x. x ∉ RF (TER g) ⟹ d_OUT g x = 0"
  shows "x ∈ SINK g"
using RF_g by(auto simp add: roofed_circ_def SINK.simps assms)

end

lemma wave_plus: "wave (quotient_web Γ f) (g ↿ Γ / f) ⟹ wave Γ (f ⌢ g)"
using f w by(rule wave_plus_current)(rule current_restrict_current[OF w g])

lemma TER_plus_web'':
  assumes "⋀x. x ∉ RF (TER g) ⟹ d_OUT g x = 0"
  shows "ℰ (TER f ∪ TER g) ⊆ TER (f ⌢ g)"
proof
  fix x
  assume *: "x ∈ ℰ (TER f ∪ TER g)"
  moreover have "x ∈ SINK (g ↿ Γ / f)"
    by(rule in_SINK_restrict_current)(rule Max_Flow_Min_Cut_Countable.SINK_TER_in''[OF f w g * assms])
  ultimately show "x ∈ TER (f ⌢ g)" by(rule TER_plus_web_aux)
qed

lemma TER_plus_web': "wave Γ g ⟹ ℰ (TER f ∪ TER g) ⊆ TER (f ⌢ g)"
by(rule TER_plus_web'')(rule waveD_OUT)

lemma wave_plus': "wave Γ g ⟹ wave Γ (f ⌢ g)"
by(rule wave_plus)(rule wave_restrict_current[OF f w g])

end

lemma RF_TER_plus_web:
  fixes Γ (structure)
  assumes f: "current Γ f"
  and w: "wave Γ f"
  and g: "current Γ g"
  and w': "wave Γ g"
  shows "RF (TER (f ⌢ g)) = RF (TER f ∪ TER g)"
proof
  have "RF (ℰ (TER f ∪ TER g)) ⊆ RF (TER (f ⌢ g))"
    by(rule roofed_mono)(rule TER_plus_web'[OF f w g w'])
  also have "RF (ℰ (TER f ∪ TER g)) = RF (TER f ∪ TER g)" by(rule RF_essential)
  finally show "… ⊆ RF (TER (f ⌢ g))" .
next
  have fg: "current Γ (f ⌢ g)" using f w g by(rule current_plus_web)
  show "RF (TER (f ⌢ g)) ⊆ RF (TER f ∪ TER g)"
  proof(intro subsetI roofedI)
    fix x p y
    assume RF: "x ∈ RF (TER (f ⌢ g))" and p: "path Γ x p y" and y: "y ∈ B Γ"
    from roofedD[OF RF p y] obtain z where z: "z ∈ set (x # p)" and TER: "z ∈ TER (f ⌢ g)" by auto
    from TER have SINK: "z ∈ SINK f"
      by(auto simp add: SINK.simps d_OUT_plus_web add_eq_0_iff_both_eq_0)
    from TER have "z ∈ SAT Γ (f ⌢ g)" by simp
    hence SAT: "z ∈ SAT Γ f ∪ SAT Γ g"
      by(cases "z ∈ RF (TER f)")(auto simp add: currentD_SAT[OF f] currentD_SAT[OF g] currentD_SAT[OF fg] d_IN_plus_web d_IN_restrict_current_outside restrict_current_IN_not_RF[OF g] wave_not_RF_IN_zero[OF f w])

    show "(∃z∈set p. z ∈ TER f ∪ TER g) ∨ x ∈ TER f ∪ TER g"
    proof(cases "z ∈ RF (TER g)")
      case False
      hence "z ∈ SINK g" by(simp add: SINK.simps waveD_OUT[OF w'])
      with SINK SAT have "z ∈ TER f ∪ TER g" by auto
      thus ?thesis using z by auto
    next
      case True
      from split_list[OF z] obtain ys zs where split: "x # p = ys @ z # zs" by blast
      with p have "path Γ z zs y" by(auto elim: rtrancl_path_appendE simp add: Cons_eq_append_conv)
      from roofedD[OF True this y] split show ?thesis by(auto simp add: Cons_eq_append_conv)
    qed
  qed
qed

lemma RF_TER_Sup:
  fixes Γ (structure)
  assumes f: "⋀f. f ∈ Y ⟹ current Γ f"
  and w: "⋀f. f ∈ Y ⟹ wave Γ f"
  and Y: "Complete_Partial_Order.chain op ≤ Y" "Y ≠ {}" "countable (support_flow (Sup Y))"
  shows "RF (TER (Sup Y)) = RF (⋃f∈Y. TER f)"
proof(rule set_eqI iffI)+
  fix x
  assume x: "x ∈ RF (TER (Sup Y))"
  have "x ∈ RF (RF (⋃f∈Y. TER f))"
  proof
    fix p y
    assume p: "path Γ x p y" and y: "y ∈ B Γ"
    from roofedD[OF x p y] obtain z where z: "z ∈ set (x # p)" and TER: "z ∈ TER (Sup Y)" by auto
    from TER have SINK: "z ∈ SINK f" if "f ∈ Y" for f using that by(auto simp add: SINK_Sup[OF Y])

    from Y(2) obtain f where y: "f ∈ Y" by blast

    show "(∃z∈set p. z ∈ RF (⋃f∈Y. TER f)) ∨ x ∈ RF (⋃f∈Y. TER f)"
    proof(cases "∃f∈Y. z ∈ RF (TER f)")
      case True
      then obtain f where fY: "f ∈ Y" and zf: "z ∈ RF (TER f)" by blast
      from zf have "z ∈ RF (⋃f∈Y. TER f)" by(rule in_roofed_mono)(auto intro: fY)
      with z show ?thesis by auto
    next
      case False
      hence *: "d_IN f z = 0" if "f ∈ Y" for f using that by(auto intro: wave_not_RF_IN_zero[OF f w])
      hence "d_IN (Sup Y) z = 0" using Y(2) by(simp add: d_IN_Sup[OF Y])
      with TER have "z ∈ SAT Γ f" using *[OF y]
        by(simp add: SAT.simps)
      with SINK[OF y] have "z ∈ TER f" by simp
      with z y show ?thesis by(auto intro: roofed_greaterI)
    qed
  qed
  then show "x ∈ RF (⋃f∈Y. TER f)" unfolding roofed_idem .
next
  fix x
  assume x: "x ∈ RF (⋃f∈Y. TER f)"
  have "x ∈ RF (RF (TER (⨆Y)))"
  proof(rule roofedI)
    fix p y
    assume p: "path Γ x p y" and y: "y ∈ B Γ"
    from roofedD[OF x p y] obtain z f where *: "z ∈ set (x # p)"
      and **: "f ∈ Y" and TER: "z ∈ TER f" by auto
    have "z ∈ RF (TER (Sup Y))"
    proof(rule ccontr)
      assume z: "z ∉ RF (TER (Sup Y))"
      have "wave Γ (Sup Y)" using Y(1-2) w Y(3) by(rule wave_lub)
      hence "d_OUT (Sup Y) z = 0" using z by(rule waveD_OUT)
      hence "z ∈ SINK (Sup Y)" by(simp add: SINK.simps)
      moreover have "z ∈ SAT Γ (Sup Y)" using TER SAT_Sup_upper[OF **, of Γ] by blast
      ultimately have "z ∈ TER (Sup Y)" by simp
      hence "z ∈ RF (TER (Sup Y))" by(rule roofed_greaterI)
      with z show False by contradiction
    qed
    thus "(∃z∈set p. z ∈ RF (TER (Sup Y))) ∨ x ∈ RF (TER (Sup Y))" using * by auto
  qed
  then show "x ∈ RF (TER (⨆Y))" unfolding roofed_idem .
qed

lemma loose_quotient_web:
  fixes Γ :: "('v, 'more) web_scheme" (structure)
  assumes weight_finite: "⋀x. weight Γ x ≠ ⊤"
  and f: "current Γ f"
  and w: "wave Γ f"
  and maximal: "⋀w. ⟦ current Γ w; wave Γ w; f ≤ w ⟧ ⟹ f = w"
  shows "loose (quotient_web Γ f)" (is "loose ?Γ")
proof
  fix g
  assume g: "current ?Γ g" and w': "wave ?Γ g"
  let ?g = "plus_current f g"
  from f w g have "current Γ ?g" "wave Γ ?g" by(rule current_plus_current wave_plus_current)+ (rule w')
  moreover have "f ≤ ?g" by(clarsimp simp add: le_fun_def add_eq_0_iff_both_eq_0)
  ultimately have eq: "f = ?g" by(rule maximal)
  have "g e = 0" for e
  proof(cases e)
    case (Pair x y)
    have "f e ≤ d_OUT f x" unfolding Pair by (rule d_OUT_ge_point)
    also have "… ≤ weight Γ x" by(rule currentD_weight_OUT[OF f])
    also have "… < ⊤" by(simp add: weight_finite less_top[symmetric])
    finally show "g e = 0" using Pair eq[THEN fun_cong, of e]
      by(cases "f e" "g e" rule: ennreal2_cases)(simp_all add: fun_eq_iff)
  qed
  thus "g = (λ_. 0)" by(simp add: fun_eq_iff)
next
  have 0: "current ?Γ zero_current" by(simp add: )
  show "¬ hindrance ?Γ zero_current"
  proof
    assume "hindrance ?Γ zero_current"
    then obtain x where a: "x ∈ A ?Γ" and : "x ∉ ℰ (TER zero_current)"
      and "d_OUT zero_current x < weight ?Γ x" by cases
    from a have "x ∈ ℰ (TER f)" by simp
    then obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ"
      and bypass: "⋀z. ⟦x ≠ y; z ∈ set p⟧ ⟹ z = x ∨ z ∉ TER f" by(rule ℰ_E) blast
    from p obtain p' where p': "path Γ x p' y" and distinct: "distinct (x # p')"
      and subset: "set p' ⊆ set p" by(auto elim: rtrancl_path_distinct)
    note p'
    moreover have RF: "z ∉ RF (TER f)" if "z ∈ set p'" for z
    proof
      assume z: "z ∈ RF (TER f)"
      from split_list[OF that] obtain ys zs where decomp: "p' = ys @ z # zs" by blast
      with p' have "y ∈ set p'" by(auto dest!: rtrancl_path_last intro: last_in_set)
      with distinct have neq: "x ≠ y" by auto
      from decomp p' have "path Γ z zs y" by(auto elim: rtrancl_path_appendE)
      from roofedD[OF z this y] obtain z' where "z' ∈ set (z # zs)" "z' ∈ TER f" by auto
      with distinct decomp subset bypass[OF neq] show False by auto
    qed
    moreover have "x ∉ RF (TER f)" using ‹x ∈ ℰ (TER f)› by(simp add: roofed_circ_def)
    ultimately have p'': "path ?Γ x p' y"
      by(induction)(auto intro: rtrancl_path.intros simp add: roofed_circ_def)
    from a  have "¬ essential ?Γ (B ?Γ) (TER zero_current) x" by simp
    from not_essentialD[OF this p''] y obtain z where neq: "x ≠ y"
      and "z ∈ set p'" "z ≠ x" "z ∈ TER zero_current" by auto
    moreover with subset RF[of z] have "z ∈ TER f"
      using currentD_weight_OUT[OF f, of z] currentD_weight_IN[OF f, of z]
      by(auto simp add: roofed_circ_def SINK.simps intro: SAT.IN split: if_split_asm)
    ultimately show False using bypass[of z] subset by auto
  qed
qed

lemma quotient_web_trimming:
  fixes Γ (structure)
  assumes w: "wave Γ f"
  and trimming: "trimming Γ f g"
  shows "quotient_web Γ f = quotient_web Γ g" (is "?lhs = ?rhs")
proof(rule web.equality)
  from trimming have : "ℰ (TER g) - A Γ = ℰ (TER f) - A Γ" by cases

  have RF: "RF (TER g) = RF (TER f)"
    by(subst (1 2) RF_essential[symmetric])(simp only: trimming_ℰ[OF w trimming])
  have RFc: "RF (TER g) = RF (TER f)"
    by(subst (1 2) roofed_circ_essential[symmetric])(simp only: trimming_ℰ[OF w trimming])

  show "edge ?lhs = edge ?rhs" by(rule ext)+(simp add: RF RFc)
  have "weight ?lhs = (λx. if x ∈ RF (TER g) ∨ x ∈ RF (TER g) ∩ B Γ then 0 else weight Γ x)"
    unfolding RF RFc by(auto simp add: fun_eq_iff RF_in_B)
  also have "… = weight ?rhs" by(auto simp add: fun_eq_iff RF_in_B)
  finally show "weight ?lhs = weight ?rhs" .

  show "A ?lhs = A ?rhs" unfolding quotient_web_sel trimming_ℰ[OF w trimming] ..
qed simp_all

subsection ‹Well-formed webs›

locale web =
  fixes Γ :: "('v, 'more) web_scheme" (structure)
  assumes A_in: "x ∈ A Γ ⟹ ¬ edge Γ y x"
  and B_out: "x ∈ B Γ ⟹ ¬ edge Γ x y"
  and A_vertex: "A Γ ⊆ V"
  and disjoint: "A Γ ∩ B Γ = {}"
  and no_loop: "⋀x. ¬ edge Γ x x"
  and weight_outside: "⋀x. x ∉ V ⟹ weight Γ x = 0"
  and weight_finite [simp]: "⋀x. weight Γ x ≠ ⊤"
begin

lemma web_weight_update:
  assumes "⋀x. ¬ vertex Γ x ⟹ w x = 0"
  and "⋀x. w x ≠ ⊤"
  shows "web (Γ⦇weight := w⦈)"
by unfold_locales(simp_all add: A_in B_out A_vertex disjoint no_loop assms)

lemma currentI [intro?]:
  assumes "⋀x. d_OUT f x ≤ weight Γ x"
  and "⋀x. d_IN f x ≤ weight Γ x"
  and OUT_IN: "⋀x. ⟦ x ∉ A Γ; x ∉ B Γ ⟧ ⟹ d_OUT f x ≤ d_IN f x"
  and outside: "⋀e. e ∉ E ⟹ f e = 0"
  shows "current Γ f"
proof
  show "d_IN f a = 0" if "a ∈ A Γ" for a using that
    by(auto simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0 A_in intro: outside)
  show "d_OUT f b = 0" if "b ∈ B Γ" for b using that
    by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 B_out intro: outside)
  then show "d_OUT f x ≤ d_IN f x" if "x ∉ A Γ" for x using OUT_IN[OF that]
    by(cases "x ∈ B Γ") auto
qed(blast intro: assms)+

lemma currentD_finite_IN:
  assumes f: "current Γ f"
  shows "d_IN f x ≠ ⊤"
proof(cases "x ∈ V")
  case True
  have "d_IN f x ≤ weight Γ x" using f by(rule currentD_weight_IN)
  also have "… < ⊤" using True weight_finite[of x] by (simp add: less_top[symmetric])
  finally show ?thesis by simp
next
  case False
  then have "d_IN f x = 0"
    by(auto simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0 vertex_def intro: currentD_outside[OF f])
  thus ?thesis by simp
qed

lemma currentD_finite_OUT:
  assumes f: "current Γ f"
  shows "d_OUT f x ≠ ⊤"
proof(cases "x ∈ V")
  case True
  have "d_OUT f x ≤ weight Γ x" using f by(rule currentD_weight_OUT)
  also have "… < ⊤" using True weight_finite[of x] by (simp add: less_top[symmetric])
  finally show ?thesis by simp
next
  case False
  then have "d_OUT f x = 0"
    by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 vertex_def intro: currentD_outside[OF f])
  thus ?thesis by simp
qed

lemma currentD_finite:
  assumes f: "current Γ f"
  shows "f e ≠ ⊤"
proof(cases e)
  case (Pair x y)
  have "f (x, y) ≤ d_OUT f x" by (rule d_OUT_ge_point)
  also have "… < ⊤" using currentD_finite_OUT[OF f] by (simp add: less_top[symmetric])
  finally show ?thesis by(simp add: Pair)
qed

lemma web_quotient_web: "web (quotient_web Γ f)" (is "web ?Γ")
proof
  show "¬ edge ?Γ y x" if "x ∈ A ?Γ" for x y using that by(auto intro: roofed_greaterI)
  show "¬ edge ?Γ x y" if "x ∈ B ?Γ" for x y using that by(auto simp add: B_out)
  show "A ?Γ ∩ B ?Γ = {}" using disjoint by auto
  show "A ?Γ ⊆ V"
  proof
    fix x
    assume "x ∈ A ?Γ"
    hence : "x ∈ ℰ (TER f)" and x: "x ∉ B Γ" using disjoint by auto
    from this(1) obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ" and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TER f)"
      by(rule ℰ_E_RF) blast
    from p y x have "p ≠ []" by(auto simp add: rtrancl_path_simps)
    with rtrancl_path_nth[OF p, of 0] have "edge Γ x (p ! 0)" by simp
    moreover have "x ∉ RF (TER f)" using  by(simp add: roofed_circ_def)
    moreover have "p ! 0 ∉ RF (TER f)" using bypass ‹p ≠ []› by auto
    ultimately have "edge ?Γ x (p ! 0)" by simp
    thus "x ∈ V" by(auto intro: vertexI1)
  qed
  show "¬ edge ?Γ x x" for x by(simp add: no_loop)
  show "weight ?Γ x = 0" if "x ∉ V" for x
  proof(cases "x ∈ RF (TER f) ∨ x ∈ TER f ∩ B Γ")
    case True thus ?thesis by simp
  next
    case False
    hence RF: "x ∉ RF (TER f)" and B: "x ∈ B Γ ⟹ x ∉ TER f" by auto
    from RF obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ" and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TER f)"
      apply(cases "x ∉ RF (RF (TER f))")
       apply(auto elim!: not_roofedE)[1]
      apply(auto simp add: roofed_circ_def roofed_idem elim: essentialE_RF)
      done
    from that have "p = []" using p y B RF bypass
      by(auto 4 3 simp add: vertex_def dest!: rtrancl_path_nth[where i=0])
    with p have xy: "x = y" by(simp add: rtrancl_path_simps)
    with B y have "x ∉ TER f" by simp
    hence RF': "x ∉ RF (TER f)" using xy y by(subst RF_in_B) auto
    have "¬ vertex Γ x"
    proof
      assume "vertex Γ x"
      then obtain x' where "edge Γ x' x" using xy y by(auto simp add: vertex_def B_out)
      moreover hence "x' ∉ RF (TER f)" using RF' by(auto dest: RF_circ_edge_forward)
      ultimately have "edge ?Γ x' x" using RF' by simp
      hence "vertex ?Γ x" by(rule vertexI2)
      with that show False by simp
    qed
    thus ?thesis by(simp add: weight_outside)
  qed
  show "weight ?Γ x ≠ ⊤" for x by simp
qed

end

locale countable_web = web Γ
  for Γ :: "('v, 'more) web_scheme" (structure)
  +
  assumes edge_antiparallel: "edge Γ x y ⟹ ¬ edge Γ y x"
  and countable [simp]: "countable E"
begin

lemma countable_V [simp]: "countable V"
by(simp add: "V_def")

lemma countable_web_quotient_web: "countable_web (quotient_web Γ f)" (is "countable_web ?Γ")
proof -
  interpret r: web  by(rule web_quotient_web)
  show ?thesis
  proof
    have "EE" by auto
    then show "countable E" by(rule countable_subset) simp
  qed(simp add: edge_antiparallel)
qed

end

subsection ‹Subtraction of a wave›

definition minus_web :: "('v, 'more) web_scheme ⇒ 'v current ⇒ ('v, 'more) web_scheme" (infixl "⊖" 65) -- ‹Definition 6.6›
where "Γ ⊖ f = Γ⦇weight := λx. if x ∈ A Γ then weight Γ x - d_OUT f x else weight Γ x + d_OUT f x - d_IN f x⦈"

lemma minus_web_sel [simp]:
  "edge (Γ ⊖ f) = edge Γ"
  "weight (Γ ⊖ f) x = (if x ∈ A Γ then weight Γ x - d_OUT f x else weight Γ x + d_OUT f x - d_IN f x)"
  "A (Γ ⊖ f) = A Γ"
  "B (Γ ⊖ f) = B Γ"
  "VΓ ⊖ f = VΓ"
  "EΓ ⊖ f = EΓ"
  "web.more (Γ ⊖ f) = web.more Γ"
by(auto simp add: minus_web_def vertex_def)

lemma vertex_minus_web [simp]: "vertex (Γ ⊖ f) = vertex Γ"
by(simp add: vertex_def fun_eq_iff)

lemma roofed_gen_minus_web [simp]: "roofed_gen (Γ ⊖ f) = roofed_gen Γ"
by(simp add: fun_eq_iff roofed_def)

lemma minus_zero_current [simp]: "Γ ⊖ zero_current = Γ"
by(rule web.equality)(simp_all add: fun_eq_iff)

lemma (in web) web_minus_web:
  assumes f: "current Γ f"
  shows "web (Γ ⊖ f)"
unfolding minus_web_def
apply(rule web_weight_update)
apply(auto simp: weight_outside currentD_weight_IN[OF f] currentD_outside_OUT[OF f]
                 currentD_outside_IN[OF f] currentD_weight_OUT[OF f] currentD_finite_OUT[OF f])
done

section ‹Bipartite webs›

locale countable_bipartite_web =
  fixes Γ :: "('v, 'more) web_scheme" (structure)
  assumes bipartite_V: "V ⊆ A Γ ∪ B Γ"
  and A_vertex: "A Γ ⊆ V"
  and bipartite_E: "edge Γ x y ⟹ x ∈ A Γ ∧ y ∈ B Γ"
  and disjoint: "A Γ ∩ B Γ = {}"
  and weight_outside: "⋀x. x ∉ V ⟹ weight Γ x = 0"
  and weight_finite [simp]: "⋀x. weight Γ x ≠ ⊤"
  and countable_E [simp]: "countable E"
begin

lemma not_vertex: "⟦ x ∉ A Γ; x ∉ B Γ ⟧ ⟹ ¬ vertex Γ x"
using bipartite_V by blast

lemma no_loop: "¬ edge Γ x x"
using disjoint by(auto dest: bipartite_E)

lemma edge_antiparallel: "edge Γ x y ⟹ ¬ edge Γ y x"
using disjoint by(auto dest: bipartite_E)

lemma A_in: "x ∈ A Γ ⟹ ¬ edge Γ y x"
using disjoint by(auto dest: bipartite_E)

lemma B_out: "x ∈ B Γ ⟹ ¬ edge Γ x y"
using disjoint by(auto dest: bipartite_E)

sublocale countable_web using disjoint
by(unfold_locales)(auto simp add: A_in B_out A_vertex no_loop weight_outside  edge_antiparallel)

lemma currentD_OUT':
  assumes f: "current Γ f"
  and x: "x ∉ A Γ"
  shows "d_OUT f x = 0"
using currentD_outside_OUT[OF f, of x] x currentD_OUT[OF f, of x] bipartite_V by auto

lemma currentD_IN':
  assumes f: "current Γ f"
  and x: "x ∉ B Γ"
  shows "d_IN f x = 0"
using currentD_outside_IN[OF f, of x] x currentD_IN[OF f, of x] bipartite_V by auto

lemma current_bipartiteI [intro?]:
  assumes OUT: "⋀x. d_OUT f x ≤ weight Γ x"
  and IN: "⋀x. d_IN f x ≤ weight Γ x"
  and outside: "⋀e. e ∉ E ⟹ f e = 0"
  shows "current Γ f"
proof
  fix x
  assume "x ∉ A Γ" "x ∉ B Γ"
  hence "d_OUT f x = 0" by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 intro!: outside dest: bipartite_E)
  then show "d_OUT f x ≤ d_IN f x" by simp
qed(rule OUT IN outside)+

lemma wave_bipartiteI [intro?]:
  assumes sep: "separating Γ (TER f)"
  and f: "current Γ f"
  shows "wave Γ f"
using sep
proof(rule wave.intros)
  fix x
  assume "x ∉ RF (TER f)"
  then consider "x ∉ V" | "x ∈ V" "x ∈ B Γ" using separating_RF_A[OF sep] bipartite_V by auto
  then show "d_OUT f x = 0" using currentD_OUT[OF f, of x] currentD_outside_OUT[OF f, of x]
    by cases auto
qed

subsection ‹Hindered webs with reduced weights›

context
  fixes u :: "'v ⇒ ennreal"
  and ε
  defines "ε ≡ (∫+ y. u y ∂count_space (B Γ))"
  assumes u_outside: "⋀x. x ∉ B Γ ⟹ u x = 0"
  and finite_ε: "ε ≠ ⊤"
begin

private lemma u_A: "x ∈ A Γ ⟹ u x = 0"
using u_outside[of x] disjoint by auto

private lemma u_finite: "u y ≠ ⊤"
proof(cases "y ∈ B Γ")
  case True
  have "u y ≤ ε" unfolding ε_def by(rule nn_integral_ge_point)(simp add: True)
  also have "… < ⊤" using finite_ε by (simp add: less_top[symmetric])
  finally show ?thesis by simp
qed(simp add: u_outside)

lemma hindered_reduce: -- ‹Lemma 6.7›
  assumes u: "u ≤ weight Γ"
  assumes hindered_by: "hindered_by (Γ⦇weight := weight Γ - u⦈) ε" (is "hindered_by ?Γ _")
  shows "hindered Γ"
proof -
  note [simp] = u_finite
  let ?TER = "TER"
  from hindered_by obtain f
    where hindrance_by: "hindrance_by ?Γ f ε"
    and f: "current ?Γ f"
    and w: "wave ?Γ f" by cases
  from hindrance_by obtain a where a: "a ∈ A Γ" "a ∉ ℰ (?TER f)"
    and a_le: "d_OUT f a < weight Γ a"
    and ε_less: "weight Γ a - d_OUT f a > ε"
    and ε_nonneg: "ε ≥ 0" by(auto simp add: u_A hindrance_by.simps)

  from f have f': "current Γ f" by(rule current_weight_mono)(auto intro: diff_le_self_ennreal)

  write Some ("⟨_⟩")

  def edge'  "λxo yo. case (xo, yo) of
      (None, Some y) ⇒ y ∈ V ∧ y ∉ A Γ
    | (Some x, Some y) ⇒ edge Γ x y ∨ edge Γ y x
    | _ ⇒ False"
  def cap  "λe. case e of
      (None, Some y) ⇒ if y ∈ V then u y else 0
    | (Some x, Some y) ⇒ if edge Γ x y ∧ x ≠ a then f (x, y) else if edge Γ y x then max (weight Γ x) (weight Γ y) + 1 else 0
    | _ ⇒ 0"
  def Ψ  "⦇edge = edge', capacity = cap, source = None, sink = Some a⦈"

  have edge'_simps [simp]:
    "edge' None ⟨y⟩ ⟷ y ∈ V ∧ y ∉ A Γ"
    "edge' xo None ⟷ False"
    "edge' ⟨x⟩ ⟨y⟩ ⟷ edge Γ x y ∨ edge Γ y x"
    for xo x y by(simp_all add: edge'_def split: option.split)

  have edge_None1E [elim!]: thesis if "edge' None y" "⋀z. ⟦ y = ⟨z⟩; z ∈ V; z ∉ A Γ ⟧ ⟹ thesis" for y thesis
    using that by(simp add: edge'_def split: option.split_asm sum.split_asm)
  have edge_Some1E [elim!]: thesis if "edge' ⟨x⟩ y" "⋀z. ⟦ y = ⟨z⟩; edge Γ x z ∨ edge Γ z x ⟧ ⟹ thesis" for x y thesis
    using that by(simp add: edge'_def split: option.split_asm sum.split_asm)
  have edge_Some2E [elim!]: thesis if "edge' x ⟨y⟩" "⟦ x = None; y ∈ V; y ∉ A Γ ⟧ ⟹ thesis" "⋀z. ⟦ x = ⟨z⟩; edge Γ z y ∨ edge Γ y z ⟧ ⟹ thesis" for x y thesis
    using that by(simp add: edge'_def split: option.split_asm sum.split_asm)

  have cap_simps [simp]:
    "cap (None, ⟨y⟩) = (if y ∈ V then u y else 0)"
    "cap (xo, None) = 0"
    "cap (⟨x⟩, ⟨y⟩) =
    (if edge Γ x y ∧ x ≠ a then f (x, y) else if edge Γ y x then max (weight Γ x) (weight Γ y) + 1 else 0)"
    for xo x y by(simp_all add: cap_def split: option.split)

  have Ψ_sel [simp]:
    "edge Ψ = edge'"
    "capacity Ψ = cap"
    "source Ψ = None"
    "sink Ψ = ⟨a⟩"
    by(simp_all add: Ψ_def)

  have cap_outside1: "¬ vertex Γ x ⟹ cap (⟨x⟩, y) = 0" for x y
    by(cases y)(auto simp add: vertex_def)
  have capacity_A_weight: "d_OUT cap ⟨x⟩ ≤ weight Γ x" if "x ∈ A Γ" for x
  proof -
    have "d_OUT cap ⟨x⟩ ≤ (∑+ y∈range Some. f (x, the y))"
      using that disjoint a(1) unfolding d_OUT_def
      by(auto 4 4 intro!: nn_integral_mono simp add: nn_integral_count_space_indicator notin_range_Some currentD_outside[OF f] split: split_indicator dest: edge_antiparallel bipartite_E)
    also have "… = d_OUT f x" by(simp add: d_OUT_def nn_integral_count_space_reindex)
    also have "… ≤ weight Γ x" using f' by(rule currentD_weight_OUT)
    finally show ?thesis .
  qed
  have flow_attainability: "flow_attainability Ψ"
  proof
    have "EΨ = (λ(x, y). (⟨x⟩, ⟨y⟩)) ` E ∪ (λ(x, y). (⟨y⟩, ⟨x⟩)) ` E ∪ (λx. (None, ⟨x⟩)) ` (V ∩ - A Γ)"
      by(auto simp add: edge'_def split: option.split_asm)
    thus "countable EΨ" by simp
  next
    fix v
    assume "v ≠ sink Ψ"
    consider (sink) "v = None" | (A) x where "v = ⟨x⟩" "x ∈ A Γ"
      | (B) y where "v = ⟨y⟩" "y ∉ A Γ" "y ∈ V" | (outside) x where "v = ⟨x⟩" "x ∉ V"
      by(cases v) auto
    then show "d_IN (capacity Ψ) v ≠ ⊤ ∨ d_OUT (capacity Ψ) v ≠ ⊤"
    proof cases
      case sink thus ?thesis by(simp add: d_IN_def)
    next
      case (A x)
      thus ?thesis using capacity_A_weight[of x] by (auto simp: top_unique)
    next
      case (B y)
      have "d_IN (capacity Ψ) v ≤ (∑+ x. f (the x, y) * indicator (range Some) x + u y * indicator {None} x)"
        using B disjoint bipartite_V a(1) unfolding d_IN_def
        by(auto 4 4 intro!: nn_integral_mono simp add: nn_integral_count_space_indicator notin_range_Some currentD_outside[OF f] split: split_indicator dest: edge_antiparallel bipartite_E)
      also have "… = (∑+ x∈range Some. f (the x, y)) + u y"
        by(subst nn_integral_add)(simp_all add: nn_integral_count_space_indicator)
      also have "… = d_IN f y + u y" by(simp add: d_IN_def nn_integral_count_space_reindex)
      also have "d_IN f y ≤ weight Γ y" using f' by(rule currentD_weight_IN)
      finally show ?thesis by(auto simp add: add_right_mono top_unique split: if_split_asm)
    next
      case outside
      hence "d_OUT (capacity Ψ) v = 0"
        by(auto simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space cap_def vertex_def split: option.split)
      thus ?thesis by simp
    qed
  next
    show "capacity Ψ e ≠ ⊤" for e using weight_finite
      by(auto simp add: cap_def max_def vertex_def currentD_finite[OF f'] split: option.split prod.split simp del: weight_finite)
    show "capacity Ψ e = 0" if "e ∉ EΨ" for e
      using that bipartite_V disjoint
      by(auto simp add: cap_def max_def intro: u_outside split: option.split prod.split)
    show "¬ edge Ψ x (source Ψ)" for x  by simp
    show "¬ edge Ψ x x" for x by(cases x)(simp_all add: no_loop)
    show "source Ψ ≠ sink Ψ" by simp
  qed
  then interpret Ψ: flow_attainability "Ψ" .

  def α  "(⨆g∈{g. flow Ψ g}. value_flow Ψ g)"
  have α_le: "α ≤ ε"
  proof -
    have "α ≤ d_OUT cap None" unfolding α_def by(rule SUP_least)(auto intro!: d_OUT_mono dest: flowD_capacity)
    also have "… ≤ ∫+ y. cap (None, y) ∂count_space (range Some)" unfolding d_OUT_def
      by(auto simp add: nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_mono split: split_indicator)
    also have "… ≤ ε" unfolding ε_def
      by (subst (2) nn_integral_count_space_indicator, auto simp add: nn_integral_count_space_reindex u_outside intro!: nn_integral_mono split: split_indicator)
    finally show ?thesis by simp
  qed
  then have finite_flow: "α ≠ ⊤" using finite_ε by (auto simp: top_unique)

  from Ψ.ex_max_flow
  obtain j where j: "flow Ψ j"
    and value_j: "value_flow Ψ j = α"
    and IN_j: "⋀x. d_IN j x ≤ α"
    unfolding α_def by auto

  have j_le_f: "j (Some x, Some y) ≤ f (x, y)" if "edge Γ x y" for x y
    using that flowD_capacity[OF j, of "(Some x, Some y)"] a(1) disjoint
    by(auto split: if_split_asm dest: bipartite_E intro: order_trans)

  have IN_j_finite [simp]: "d_IN j x ≠ ⊤" for x using finite_flow by(rule neq_top_trans)(simp add: IN_j)

  have j_finite[simp]: "j (x, y) < ⊤" for x y
    by (rule le_less_trans[OF d_IN_ge_point]) (simp add: IN_j_finite[of y] less_top[symmetric])

  have OUT_j_finite: "d_OUT j x ≠ ⊤" for x
  proof(cases "x = source Ψ ∨ x = sink Ψ")
    case True thus ?thesis
    proof cases
      case left thus ?thesis using finite_flow value_j by simp
    next
      case right
      have "d_OUT (capacity Ψ) ⟨a⟩ ≠ ⊤" using capacity_A_weight[of a] a(1) by(auto simp: top_unique)
      thus ?thesis unfolding right[simplified]
        by(rule neq_top_trans)(rule d_OUT_mono flowD_capacity[OF j])+
    qed
  next
    case False then show ?thesis by(simp add: flowD_KIR[OF j])
  qed

  have IN_j_le_weight: "d_IN j ⟨x⟩ ≤ weight Γ x" for x
  proof(cases "x ∈ A Γ")
    case xA: True
    show ?thesis
    proof(cases "x = a")
      case True
      have "d_IN j ⟨x⟩ ≤ α" by(rule IN_j)
      also have "… ≤ ε" by(rule α_le)
      also have "ε < weight Γ a" using ε_less diff_le_self_ennreal less_le_trans by blast
      finally show ?thesis using True by(auto intro: order.strict_implies_order)
    next
      case False
      have "d_IN j ⟨x⟩ = d_OUT j ⟨x⟩" using flowD_KIR[OF j, of "Some x"] False by simp
      also have "… ≤ d_OUT cap ⟨x⟩" using flowD_capacity[OF j] by(auto intro: d_OUT_mono)
      also have "… ≤ weight Γ x" using xA  by(rule capacity_A_weight)
      finally show ?thesis .
    qed
  next
    case xA: False
    show ?thesis
    proof(cases "x ∈ B Γ")
      case True
      have "d_IN j ⟨x⟩ ≤ d_IN cap ⟨x⟩" using flowD_capacity[OF j] by(auto intro: d_IN_mono)
      also have "… ≤ (∑+ z. f (the z, x) * indicator (range Some) z) + (∑+ z :: 'v option. u x * indicator {None} z)"
        using True disjoint
        by(subst nn_integral_add[symmetric])(auto simp add: vertex_def currentD_outside[OF f] d_IN_def B_out intro!: nn_integral_mono split: split_indicator)
      also have "… = d_IN f x + u x"
        by(simp add: nn_integral_count_space_indicator[symmetric] nn_integral_count_space_reindex d_IN_def)
      also have "… ≤ weight Γ x" using currentD_weight_IN[OF f, of x] u_finite[of x]
        using ε_less u by (auto simp add: ennreal_le_minus_iff le_fun_def)
      finally show ?thesis .
    next
      case False
      with xA have "x ∉ V" using bipartite_V by blast
      then have "d_IN j ⟨x⟩ = 0" using False
        by(auto simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0 vertex_def edge'_def split: option.split_asm intro!: Ψ.flowD_outside[OF j])
      then show ?thesis
        by simp
    qed
  qed

  let ?j = "j ∘ map_prod Some Some ∘ prod.swap"

  have finite_j_OUT: "(∑+ y∈OUT x. j (⟨x⟩, ⟨y⟩)) ≠ ⊤" (is "?j_OUT ≠ _") if "x ∈ A Γ" for x
    using currentD_finite_OUT[OF f', of x]
    by(rule neq_top_trans)(auto intro!: nn_integral_mono j_le_f simp add: d_OUT_def nn_integral_count_space_indicator outgoing_def split: split_indicator)
  have j_OUT_eq: "?j_OUT x = d_OUT j ⟨x⟩" if "x ∈ A Γ" for x
  proof -
    have "?j_OUT x = (∑+ y∈range Some. j (Some x, y))" using that disjoint
      by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator outgoing_def intro!: nn_integral_cong Ψ.flowD_outside[OF j] dest: bipartite_E split: split_indicator)
    also have "… = d_OUT j ⟨x⟩"
      by(auto simp add: d_OUT_def nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_cong Ψ.flowD_outside[OF j] split: split_indicator)
    finally show ?thesis .
  qed

  def g  "f ⊕ ?j"
  have g_simps: "g (x, y) = (f ⊕ ?j) (x, y)" for x y by(simp add: g_def)

  have OUT_g_A: "d_OUT g x = d_OUT f x + d_IN j ⟨x⟩ - d_OUT j ⟨x⟩" if "x ∈ A Γ" for x
  proof -
    have "d_OUT g x = (∑+ y∈OUT x. f (x, y) + j (⟨y⟩, ⟨x⟩) - j (⟨x⟩, ⟨y⟩))"
      by(auto simp add: d_OUT_def g_simps currentD_outside[OF f'] outgoing_def nn_integral_count_space_indicator intro!: nn_integral_cong)
    also have "… = (∑+ y∈OUT x. f (x, y) + j (⟨y⟩, ⟨x⟩)) - (∑+ y∈OUT x. j (⟨x⟩, ⟨y⟩))"
      (is "_ = _ - ?j_OUT") using finite_j_OUT[OF that]
      by(subst nn_integral_diff)(auto simp add: AE_count_space outgoing_def intro!: order_trans[OF j_le_f])
    also have "… = (∑+ y∈OUT x. f (x, y)) + (∑+ y∈OUT x. j (Some y, Some x)) - ?j_OUT"
      (is "_ = ?f + ?j_IN - _") by(subst nn_integral_add) simp_all
    also have "?f = d_OUT f x" by(subst d_OUT_alt_def[where G=Γ])(simp_all add: currentD_outside[OF f])
    also have "?j_OUT = d_OUT j ⟨x⟩" using that by(rule j_OUT_eq)
    also have "?j_IN = (∑+ y∈range Some. j (y, ⟨x⟩))" using that disjoint
      by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator outgoing_def intro!: nn_integral_cong Ψ.flowD_outside[OF j] split: split_indicator dest: bipartite_E)
    also have "… = d_IN j (Some x)" using that disjoint
      by(auto 4 3 simp add: d_IN_def nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_cong Ψ.flowD_outside[OF j] split: split_indicator)
    finally show ?thesis by simp
  qed

  have OUT_g_B: "d_OUT g x = 0" if "x ∉ A Γ" for x
    using disjoint that
    by(auto simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space g_simps dest: bipartite_E)

  have OUT_g_a: "d_OUT g a < weight Γ a" using a(1)
  proof -
    have "d_OUT g a = d_OUT f a + d_IN j ⟨a⟩ - d_OUT j ⟨a⟩" using a(1) by(rule OUT_g_A)
    also have "… ≤ d_OUT f a + d_IN j ⟨a⟩"
      by(rule diff_le_self_ennreal)
    also have "… < weight Γ a + d_IN j ⟨a⟩ - ε"
      using finite_ε ε_less currentD_finite_OUT[OF f']
      by (simp add: less_diff_eq_ennreal less_top ac_simps)
    also have "… ≤ weight Γ a"
      using IN_j[THEN order_trans, OF α_le] by (simp add: ennreal_minus_le_iff)
    finally show ?thesis .
  qed

  have OUT_jj: "d_OUT ?j x = d_IN j ⟨x⟩ - j (None, ⟨x⟩)" for x
  proof -
    have "d_OUT ?j x = (∑+ y∈range Some. j (y, ⟨x⟩))" by(simp add: d_OUT_def nn_integral_count_space_reindex)
    also have "… = d_IN j ⟨x⟩ - (∑+ y. j (y, ⟨x⟩) * indicator {None} y)" unfolding d_IN_def
      by(subst nn_integral_diff[symmetric])(auto simp add: max_def Ψ.flowD_finite[OF j] AE_count_space nn_integral_count_space_indicator split: split_indicator intro!: nn_integral_cong)
    also have "… = d_IN j ⟨x⟩ - j (None, ⟨x⟩)" by(simp add: max_def)
    finally show ?thesis .
  qed

  have OUT_jj_finite [simp]: "d_OUT ?j x ≠ ⊤" for x
    by(simp add: OUT_jj)

  have IN_g: "d_IN g x = d_IN f x + j (None, ⟨x⟩)" for x
  proof(cases "x ∈ B Γ")
    case True
    have finite: "(∑+ y∈IN x. j (Some y, Some x)) ≠ ⊤" using currentD_finite_IN[OF f', of x]
      by(rule neq_top_trans)(auto intro!: nn_integral_mono j_le_f simp add: d_IN_def nn_integral_count_space_indicator incoming_def split: split_indicator)

    have "d_IN g x = d_IN (f ⊕ ?j) x" by(simp add: g_def)
    also have "… = (∑+ y∈IN x. f (y, x) + j (Some x, Some y) - j (Some y, Some x))"
      by(auto simp add: d_IN_def currentD_outside[OF f'] incoming_def nn_integral_count_space_indicator intro!: nn_integral_cong)
    also have "… = (∑+ y∈IN x. f (y, x) + j (Some x, Some y)) - (∑+ y∈IN x. j (Some y, Some x))"
      (is "_ = _ - ?j_IN") using finite
      by(subst nn_integral_diff)(auto simp add: AE_count_space incoming_def intro!: order_trans[OF j_le_f])
    also have "… = (∑+ y∈IN x. f (y, x)) + (∑+ y∈IN x. j (Some x, Some y)) - ?j_IN"
      (is "_ = ?f + ?j_OUT - _") by(subst nn_integral_add) simp_all
    also have "?f = d_IN f x" by(subst d_IN_alt_def[where G=Γ])(simp_all add: currentD_outside[OF f])
    also have "?j_OUT = (∑+ y∈range Some. j (Some x, y))" using True disjoint
      by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator incoming_def intro!: nn_integral_cong Ψ.flowD_outside[OF j] split: split_indicator dest: bipartite_E)
    also have "… = d_OUT j (Some x)" using disjoint
      by(auto 4 3 simp add: d_OUT_def nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_cong Ψ.flowD_outside[OF j] split: split_indicator)
    also have "… = d_IN j (Some x)" using flowD_KIR[OF j, of "Some x"] True a disjoint by auto
    also have "?j_IN = (∑+ y∈range Some. j (y, Some x))" using True disjoint
      by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator incoming_def intro!: nn_integral_cong Ψ.flowD_outside[OF j] dest: bipartite_E split: split_indicator)
    also have "… = d_IN j (Some x) - (∑+ y :: 'v option. j (None, Some x) * indicator {None} y)"
      unfolding d_IN_def using flowD_capacity[OF j, of "(None, Some x)"]
      by(subst nn_integral_diff[symmetric])
        (auto simp add: nn_integral_count_space_indicator AE_count_space top_unique image_iff
              intro!: nn_integral_cong ennreal_diff_self split: split_indicator if_split_asm)
    also have "d_IN f x + d_IN j (Some x) - … = d_IN f x + j (None, Some x)"
      using ennreal_add_diff_cancel_right[OF IN_j_finite[of "Some x"], of "d_IN f x + j (None, Some x)"]
      apply(subst diff_diff_ennreal')
      apply(auto simp add: d_IN_def intro!: nn_integral_ge_point ennreal_diff_le_mono_left)
      apply(simp add: ac_simps)
      done
    finally show ?thesis .
  next
    case False
    hence "d_IN g x = 0" "d_IN f x = 0" "j (None, ⟨x⟩) = 0"
      using disjoint currentD_IN[OF f', of x] bipartite_V currentD_outside_IN[OF f'] u_outside[OF False] flowD_capacity[OF j, of "(None, ⟨x⟩)"]
      by(cases "vertex Γ x"; auto simp add: d_IN_def nn_integral_0_iff_AE AE_count_space g_simps dest: bipartite_E split: if_split_asm)+
    thus ?thesis by simp
  qed

  have g: "current Γ g"
  proof
    show "d_OUT g x ≤ weight Γ x" for x
    proof(cases "x ∈ A Γ")
      case False
      thus ?thesis by(simp add: OUT_g_B)
    next
      case True
      with OUT_g_a show ?thesis
        by(cases "x = a")(simp_all add: OUT_g_A flowD_KIR[OF j] currentD_weight_OUT[OF f'])
    qed

    show "d_IN g x ≤ weight Γ x" for x
    proof(cases "x ∈ B Γ")
      case False
      hence "d_IN g x = 0" using disjoint
        by(auto simp add: d_IN_def nn_integral_0_iff_AE AE_count_space g_simps dest: bipartite_E)
      thus ?thesis by simp
    next
      case True
      have "d_IN g x ≤ (weight Γ x - u x) + u x" unfolding IN_g
        using currentD_weight_IN[OF f, of x] flowD_capacity[OF j, of "(None, Some x)"] True bipartite_V
        by(intro add_mono)(simp_all split: if_split_asm)
      also have "… = weight Γ x"
        using u by (intro diff_add_cancel_ennreal) (simp add: le_fun_def)
      finally show ?thesis .
    qed
    show "g e = 0" if "e ∉ E" for e using that
      by(cases e)(auto simp add: g_simps)
  qed

  def cap'  "λ(x, y). if edge Γ x y then g (x, y) else if edge Γ y x then 1 else 0"

  have cap'_simps [simp]: "cap' (x, y) = (if edge Γ x y then g (x, y) else if edge Γ y x then 1 else 0)"
    for x y by(simp add: cap'_def)

  def G  "⦇edge = λx y. cap' (x, y) > 0⦈"
  have G_sel [simp]: "edge G x y ⟷ cap' (x, y) > 0" for x y by(simp add: G_def)
  def reachable  "λx. (edge G)** x a"
  have reachable_alt_def: "reachable ≡ λx. ∃p. path G x p a"
    by(simp add: reachable_def rtranclp_eq_rtrancl_path)

  have [simp]: "reachable a" by(auto simp add: reachable_def)

  have AB_edge: "edge G x y" if "edge Γ y x" for x y
    using that
    by(auto dest: edge_antiparallel simp add: min_def le_neq_trans add_eq_0_iff_both_eq_0)
  have reachable_AB: "reachable y" if "reachable x" "(x, y) ∈ E" for x y
    using that by(auto simp add: reachable_def simp del: G_sel dest!: AB_edge intro: rtrancl_path.step)
  have reachable_BA: "g (x, y) = 0" if "reachable y" "(x, y) ∈ E" "¬ reachable x" for x y
  proof(rule ccontr)
    assume "g (x, y) ≠ 0"
    then have "g (x, y) > 0" by (simp add: zero_less_iff_neq_zero)
    hence "edge G x y" using that by simp
    then have "reachable x" using ‹reachable y›
      unfolding reachable_def by(rule converse_rtranclp_into_rtranclp)
    with ‹¬ reachable x› show False by contradiction
  qed
  have reachable_V: "vertex Γ x" if "reachable x" for x
  proof -
    from that obtain p where p: "path G x p a" unfolding reachable_alt_def ..
    then show ?thesis using rtrancl_path_nth[OF p, of 0] a(1) A_vertex
      by(cases "p = []")(auto 4 3 simp add: vertex_def elim: rtrancl_path.cases split: if_split_asm)
  qed

  have finite_j_IN: "(∫+ y. j (Some y, Some x) ∂count_space (IN x)) ≠ ⊤" for x
  proof -
    have "(∫+ y. j (Some y, Some x) ∂count_space (IN x)) ≤ d_IN f x"
      by(auto intro!: nn_integral_mono j_le_f simp add: d_IN_def nn_integral_count_space_indicator incoming_def split: split_indicator)
    thus ?thesis using currentD_finite_IN[OF f', of x] by (auto simp: top_unique)
  qed
  have j_outside: "j (x, y) = 0" if "¬ edge Ψ x y" for x y
    using that flowD_capacity[OF j, of "(x, y)"] Ψ.capacity_outside[of "(x, y)"]
    by(auto)

  def h  "λ(x, y). if reachable x ∧ reachable y then g (x, y) else 0"
  have h_simps [simp]: "h (x, y) = (if reachable x ∧ reachable y then g (x, y) else 0)" for x y
    by(simp add: h_def)
  have h_le_g: "h e ≤ g e" for e by(cases e) simp

  have OUT_h: "d_OUT h x = (if reachable x then d_OUT g x else 0)" for x
  proof -
    have "d_OUT h x = (∑+ y∈OUT x. h (x, y))" using h_le_g currentD_outside[OF g]
      by(intro d_OUT_alt_def) auto
    also have "… = (∑+ y∈OUT x. if reachable x then g (x, y) else 0)"
      by(auto intro!: nn_integral_cong simp add: outgoing_def dest: reachable_AB)
    also have "… = (if reachable x then d_OUT g x else 0)"
      by(auto intro!: d_OUT_alt_def[symmetric] currentD_outside[OF g])
    finally show ?thesis .
  qed
  have IN_h: "d_IN h x = (if reachable x then d_IN g x else 0)" for x
  proof -
    have "d_IN h x = (∑+ y∈IN x. h (y, x))"
      using h_le_g currentD_outside[OF g] by(intro d_IN_alt_def) auto
    also have "… = (∑+ y∈IN x. if reachable x then g (y, x) else 0)"
      by(auto intro!: nn_integral_cong simp add: incoming_def dest: reachable_BA)
    also have "… = (if reachable x then d_IN g x else 0)"
      by(auto intro!: d_IN_alt_def[symmetric] currentD_outside[OF g])
    finally show ?thesis .
  qed

  have h: "current Γ h" using g h_le_g
  proof(rule current_leI)
    show "d_OUT h x ≤ d_IN h x" if "x ∉ A Γ" for x
      by(simp add: OUT_h IN_h currentD_OUT_IN[OF g that])
  qed

  have reachable_full: "j (None, ⟨y⟩) = u y" if reach: "reachable y" for y
  proof(rule ccontr)
    assume "j (None, ⟨y⟩) ≠ u y"
    with flowD_capacity[OF j, of "(None, ⟨y⟩)"]
    have le: "j (None, ⟨y⟩) < u y" by(auto split: if_split_asm simp add: u_outside Ψ.flowD_outside[OF j] zero_less_iff_neq_zero)
    then obtain y: "y ∈ B Γ" and uy: "u y > 0" using u_outside[of y]
      by(cases "y ∈ B Γ"; cases "u y = 0") (auto simp add: zero_less_iff_neq_zero)

    from reach obtain q where q: "path G y q a" and distinct: "distinct (y # q)"
      unfolding reachable_alt_def by(auto elim: rtrancl_path_distinct)
    have q_Nil: "q ≠ []" using q a(1) disjoint y by(auto elim!: rtrancl_path.cases)

    let ?E = "zip (y # q) q"
    def E  "(None, Some y) # map (map_prod Some Some) ?E"
    def ζ  "Min (insert (u y - j (None, Some y)) (cap' ` set ?E))"
    let ?j' = "λe. (if e ∈ set E then ζ else 0) + j e"
    def j'  "cleanup ?j'"

    have j_free: "0 < cap' e" if "e ∈ set ?E" for e using that unfolding E_def list.sel
    proof -
      from that obtain i where e: "e = ((y # q) ! i, q ! i)"
        and i: "i < length q" by(auto simp add: set_zip)
      have e': "edge G ((y # q) ! i) (q ! i)" using q i by(rule rtrancl_path_nth)
      thus ?thesis using e by(simp)
    qed
    have ζ_pos: "0 < ζ" unfolding ζ_def using le
      by(auto intro: Min_grI j_free diff_gr0_ennreal)
    have ζ_le: "ζ ≤ cap' e" if "e ∈ set ?E" for e using that unfolding ζ_def by auto
    have finite_ζ [simplified]: "ζ < ⊤" unfolding ζ_def
      by(intro Min_less_iff[THEN iffD2])(auto simp add: less_top[symmetric])

    have E_antiparallel: "(x', y') ∈ set ?E ⟹ (y', x') ∉ set ?E" for x' y'
      using distinct
      apply(auto simp add: in_set_zip nth_Cons in_set_conv_nth)
      apply(auto simp add: distinct_conv_nth split: nat.split_asm)
      by (metis Suc_lessD less_Suc_eq less_irrefl_nat)

    have OUT_j': "d_OUT ?j' x' = ζ * card (set [(x'', y) ← E. x'' = x']) + d_OUT j x'" for x'
    proof -
      have "d_OUT ?j' x' = d_OUT (λe. if e ∈ set E then ζ else 0) x' + d_OUT j x'"
        using ζ_pos by(intro d_OUT_add)
      also have "d_OUT (λe. if e ∈ set E then ζ else 0) x' = ∫+ y. ζ * indicator (set E) (x', y) ∂count_space UNIV"
        unfolding d_OUT_def by(rule nn_integral_cong)(simp)
      also have "… = (∫+ e. ζ * indicator (set E) e ∂embed_measure (count_space UNIV) (Pair x'))"
        by(simp add: measurable_embed_measure1 nn_integral_embed_measure)
      also have "… = (∫+ e. ζ * indicator (set [(x'', y) ← E. x'' = x']) e ∂count_space UNIV)"
        by(auto simp add: embed_measure_count_space' nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
      also have "… = ζ * card (set [(x'', y) ← E. x'' = x'])" using ζ_pos by(simp add: nn_integral_cmult)
      finally show ?thesis .
    qed
    have IN_j': "d_IN ?j' x' = ζ * card (set [(y, x'') ← E. x'' = x']) + d_IN j x'" for x'
    proof -
      have "d_IN ?j' x' = d_IN (λe. if e ∈ set E then ζ else 0) x' + d_IN j x'"
        using ζ_pos by(intro d_IN_add)
      also have "d_IN (λe. if e ∈ set E then ζ else 0) x' = ∫+ y. ζ * indicator (set E) (y, x') ∂count_space UNIV"
        unfolding d_IN_def by(rule nn_integral_cong)(simp)
      also have "… = (∫+ e. ζ * indicator (set E) e ∂embed_measure (count_space UNIV) (λy. (y, x')))"
        by(simp add: measurable_embed_measure1 nn_integral_embed_measure)
      also have "… = (∫+ e. ζ * indicator (set [(y, x'') ← E. x'' = x']) e ∂count_space UNIV)"
        by(auto simp add: embed_measure_count_space' nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
      also have "… = ζ * card (set [(y, x'') ← E. x'' = x'])"
        using ζ_pos by(auto simp add: nn_integral_cmult)
      finally show ?thesis .
    qed

    have j': "flow Ψ j'"
    proof
      fix e :: "'v option edge"
      consider (None) "e = (None, Some y)"
        | (Some) x y where "e = (Some x, Some y)" "(x, y) ∈ set ?E"
        | (old) x y where "e = (Some x, Some y)" "(x, y) ∉ set ?E"
        | y' where "e = (None, Some y')" "y ≠ y'"
        | "e = (None, None)" | x where "e = (Some x, None)"
        by(cases e; case_tac a; case_tac b)(auto)
      then show "j' e ≤ capacity Ψ e" using uy ζ_pos flowD_capacity[OF j, of e]
      proof(cases)
        case None
        have "ζ ≤ u y - j (None, Some y)" by(simp add: ζ_def)
        then have "ζ + j (None, Some y) ≤ u y"
          using ζ_pos by (auto simp add: ennreal_le_minus_iff)
        thus ?thesis using reachable_V[OF reach] None Ψ.flowD_outside[OF j, of "(Some y, None)"] uy
          by(auto simp add: j'_def E_def)
      next
        case (Some x' y')
        have e: "ζ ≤ cap' (x', y')" using Some(2) by(rule ζ_le)
        then consider (backward) "edge Γ x' y'" "x' ≠ a" | (forward) "edge Γ y' x'" "¬ edge Γ x' y'"
          | (a') "edge Γ x' y'" "x' = a"
          using Some ζ_pos by(auto split: if_split_asm)
        then show ?thesis
        proof cases
          case backward
          have "ζ ≤ f (x', y') + j (Some y', Some x') - j (Some x', Some y')"
            using e backward Some(1) by(simp add: g_simps)
          hence "ζ + j (Some x', Some y') - j (Some y', Some x') ≤ (f (x', y') + j (Some y', Some x') - j (Some x', Some y')) + j (Some x', Some y') - j (Some y', Some x')"
            by(intro ennreal_minus_mono add_right_mono) simp_all
          also have "… = f (x', y')"
            using j_le_f[OF ‹edge Γ x' y'›]
            by(simp_all add: add_increasing2 less_top diff_add_assoc2_ennreal)
          finally show ?thesis using Some backward
            by(auto simp add: j'_def E_def dest: in_set_tlD E_antiparallel)
        next
          case forward
          have "ζ + j (Some x', Some y') - j (Some y', Some x') ≤ ζ + j (Some x', Some y')"
            by(rule diff_le_self_ennreal)
          also have "j (Some x', Some y') ≤ d_IN j (Some y')"
            by(rule d_IN_ge_point)
          also have "… ≤ weight Γ y'" by(rule IN_j_le_weight)
          also have "ζ ≤ 1" using e forward by simp
          finally have "ζ + j (Some x', Some y') - j (Some y', Some x') ≤ max (weight Γ x') (weight Γ y') + 1"
            by(simp add: add_left_mono add_right_mono max_def)(metis (no_types, lifting) add.commute add_right_mono less_imp_le less_le_trans not_le)
          then show ?thesis using Some forward e
            by(auto simp add: j'_def E_def max_def dest: in_set_tlD E_antiparallel)
        next
          case a'
          with Some have "a ∈ set (map fst (zip (y # q) q))" by(auto intro: rev_image_eqI)
          also have "map fst (zip (y # q) q) = butlast (y # q)" by(induction q arbitrary: y) auto
          finally have False using rtrancl_path_last[OF q q_Nil] distinct q_Nil
            by(cases q rule: rev_cases) auto
          then show ?thesis ..
        qed
      next
        case (old x' y')
        hence "j' e ≤ j e" using ζ_pos
          by(auto simp add: j'_def E_def intro!: diff_le_self_ennreal)
        also have "j e ≤ capacity Ψ e" using j by(rule flowD_capacity)
        finally show ?thesis .
      qed(auto simp add: j'_def E_def Ψ.flowD_outside[OF j] uy)
    next
      fix x'
      assume x': "x' ≠ source Ψ" "x' ≠ sink Ψ"
      then obtain x'' where x'': "x' = Some x''" by auto
      have "d_OUT ?j' x' = ζ * card (set [(x'', y) ← E. x'' = x']) + d_OUT j x'" by(rule OUT_j')
      also have "card (set [(x'', y) ← E. x'' = x']) = card (set [(y, x'') ← E. x'' = x'])" (is "?lhs = ?rhs")
      proof -
        have "?lhs = length [(x'', y) ← E. x'' = x']" using distinct
          by(subst distinct_card)(auto simp add: E_def filter_map distinct_map inj_map_prod' distinct_zipI1)
        also have "… = length [x''' ← map fst ?E. x''' = x'']"
          by(simp add: E_def x'' split_beta cong: filter_cong)
        also have "map fst ?E = butlast (y # q)" by(induction q arbitrary: y) simp_all
        also have "[x''' ← butlast (y # q). x''' = x''] = [x''' ← y # q. x''' = x'']"
          using q_Nil rtrancl_path_last[OF q q_Nil] x' x''
          by(cases q rule: rev_cases) simp_all
        also have "q = map snd ?E" by(induction q arbitrary: y) auto
        also have "length [x''' ← y # …. x''' = x''] = length [x'' ← map snd E. x'' = x']" using x''
          by(simp add: E_def cong: filter_cong)
        also have "… = length [(y, x'') ← E. x'' = x']" by(simp cong: filter_cong add: split_beta)
        also have "… = ?rhs" using distinct
          by(subst distinct_card)(auto simp add: E_def filter_map distinct_map inj_map_prod' distinct_zipI1)
        finally show ?thesis .
      qed
      also have "ζ * … + d_OUT j x' =  d_IN ?j' x'"
        unfolding flowD_KIR[OF j x'] by(rule IN_j'[symmetric])
      also have "d_IN ?j' x' ≠ ⊤"
        using Ψ.flowD_finite_IN[OF j x'(2)] finite_ζ IN_j'[of x'] by (auto simp: top_add ennreal_mult_eq_top_iff)
      ultimately show "KIR j' x'" unfolding j'_def by(rule KIR_cleanup)
    qed
    hence "value_flow Ψ j' ≤ α" unfolding α_def by(auto intro: SUP_upper)
    moreover have "value_flow Ψ j' > value_flow Ψ j"
    proof -
      have "value_flow Ψ j + 0 < value_flow Ψ j + ζ * 1"
        using ζ_pos value_j finite_flow by simp
      also have "[(x', y') ← E. x' = None] = [(None, Some y)]"
        using q_Nil by(cases q)(auto simp add: E_def filter_map cong: filter_cong split_beta)
      hence "ζ * 1 ≤ ζ * card (set [(x', y') ← E. x' = None])" using ζ_pos
        by(intro mult_left_mono)(auto simp add: E_def real_of_nat_ge_one_iff neq_Nil_conv card_insert)
      also have "value_flow Ψ j + … = value_flow Ψ ?j'"
        using OUT_j' by(simp add: add.commute)
      also have "… = value_flow Ψ j'" unfolding j'_def
        by(subst value_flow_cleanup)(auto simp add: E_def Ψ.flowD_outside[OF j])
      finally show ?thesis by(simp add: add_left_mono)
    qed
    ultimately show False using finite_flow ζ_pos value_j
      by(cases "value_flow Ψ j" ζ rule: ennreal2_cases) simp_all
  qed

  have sep_h: "y ∈ TER h" if reach: "reachable y" and y: "y ∈ B Γ" and TER: "y ∈ ?TER f" for y
  proof(rule ccontr)
    assume y': "y ∉ TER h"
    from y a(1) disjoint have yna: "y ≠ a" by auto

    from reach obtain p' where "path G y p' a" unfolding reachable_alt_def ..
    then obtain p' where p': "path G y p' a" and distinct: "distinct (y # p')" by(rule rtrancl_path_distinct)

    have SINK: "y ∈ SINK h" using y disjoint
      by(auto simp add: SINK.simps d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 intro: currentD_outside[OF g] dest: bipartite_E)
    have hg: "d_IN h y = d_IN g y" using reach by(simp add: IN_h)
    also have "… = d_IN f y + j (None, Some y)" by(simp add: IN_g)
    also have "d_IN f y = weight Γ y - u y" using currentD_weight_IN[OF f, of y] y disjoint TER
      by(auto elim!: SAT.cases)
    also have "d_IN h y < weight Γ y" using y' currentD_weight_IN[OF g, of y] y disjoint SINK
      by(auto intro: SAT.intros)
    ultimately have le: "j (None, Some y) < u y"
      by(cases "weight Γ y" "u y" "j (None, Some y)" rule: ennreal3_cases; cases "u y ≤ weight Γ y")
        (auto simp: ennreal_minus ennreal_plus[symmetric] add_top ennreal_less_iff ennreal_neg simp del: ennreal_plus)
    moreover from reach have "j (None, ⟨y⟩) = u y" by(rule reachable_full)
    ultimately show False by simp
  qed

  have w': "wave Γ h"
  proof
    show sep: "separating Γ (TER h)"
    proof(rule ccontr)
      assume "¬ ?thesis"
      then obtain x p y where x: "x ∈ A Γ" and y: "y ∈ B Γ" and p: "path Γ x p y"
        and x': "x ∉ TER h" and bypass: "⋀z. z ∈ set p ⟹ z ∉ TER h"
        by(auto simp add: separating_gen.simps)
      from p disjoint x y have p_eq: "p = [y]" and edge: "(x, y) ∈ E"
        by -(erule rtrancl_path.cases, auto dest: bipartite_E)+
      from p_eq bypass have y': "y ∉ TER h" by simp
      have "reachable x" using x' by(rule contrapos_np)(simp add: SINK.simps d_OUT_def SAT.A x)
      hence reach: "reachable y" using edge by(rule reachable_AB)

      have *: "x ∉ ℰ (?TER f)" using x'
      proof(rule contrapos_nn)
        assume *: "x ∈ ℰ (?TER f)"
        have "d_OUT h x ≤ d_OUT g x" using h_le_g by(rule d_OUT_mono)
        also from * have "x ≠ a" using a by auto
        then have "d_OUT j (Some x) = d_IN j (Some x)" by(auto intro: flowD_KIR[OF j])
        hence "d_OUT g x ≤ d_OUT f x" using OUT_g_A[OF x] IN_j[of "Some x"] finite_flow
          by(auto split: if_split_asm)
        also have "… = 0" using * by(auto elim: SINK.cases)
        finally have "x ∈ SINK h" by(simp add: SINK.simps)
        with x show "x ∈ TER h" by(simp add: SAT.A)
      qed
      from p p_eq x y have "path ?Γ x [y] y" "x ∈ A ?Γ" "y ∈ B ?Γ" by simp_all
      from * separatingD[OF separating_essential, OF waveD_separating, OF w this]
      have "y ∈ ?TER f" by auto
      with reach y have "y ∈ TER h" by(rule sep_h)
      with y' show False by contradiction
    qed
  qed(rule h)

  have OUT_g_a: "d_OUT g a = d_OUT h a" by(simp add: OUT_h)
  have "a ∉ ℰ (TER h)"
  proof
    assume *: "a ∈ ℰ (TER h)"

    have "j (Some a, Some y) = 0" for y
      using flowD_capacity[OF j, of "(Some a, Some y)"] a(1) disjoint
      by(auto split: if_split_asm dest: bipartite_E)
    then have "d_OUT f a ≤ d_OUT g a" unfolding d_OUT_def
      -- ‹This step requires that @{term j} does not decrease the outflow of @{term a}. That's
          why we set the capacity of the outgoing edges from @{term "Some a"} in @{term Ψ} to @{term "0 :: ennreal"}›
      by(intro nn_integral_mono)(auto simp add: g_simps currentD_outside[OF f] intro: )
    then have "a ∈ SINK f" using OUT_g_a * by(simp add: SINK.simps)
    with a(1) have "a ∈ ?TER f" by(auto intro: SAT.A)
    with a(2) have a': "¬ essential Γ (B Γ) (?TER f) a" by simp

    from * obtain y where ay: "edge Γ a y" and y: "y ∈ B Γ" and y': "y ∉ TER h" using disjoint a(1)
      by(auto 4 4 simp add: essential_def elim: rtrancl_path.cases dest: bipartite_E)
    from not_essentialD[OF a' rtrancl_path.step, OF ay rtrancl_path.base y]
    have TER: "y ∈ ?TER f" by simp

    have "reachable y" using ‹reachable a› by(rule reachable_AB)(simp add: ay)
    hence "y ∈ TER h" using y TER by(rule sep_h)
    with y' show False by contradiction
  qed
  with ‹a ∈ A Γ› have "hindrance Γ h"
  proof
    have "d_OUT h a = d_OUT g a" by(simp add: OUT_g_a)
    also have "… ≤ d_OUT f a + ∫+ y. j (Some y, Some a) ∂count_space UNIV"
      unfolding d_OUT_def d_IN_def
      by(subst nn_integral_add[symmetric])(auto simp add: g_simps intro!: nn_integral_mono diff_le_self_ennreal)
    also have "(∫+ y. j (Some y, Some a) ∂count_space UNIV) = (∫+ y. j (y, Some a) ∂embed_measure (count_space UNIV) Some)"
      by(simp add: nn_integral_embed_measure measurable_embed_measure1)
    also have "… ≤ d_IN j (Some a)" unfolding d_IN_def
      by(auto simp add: embed_measure_count_space nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator)
    also have "… ≤ α" by(rule IN_j)
    also have "… ≤ ε" by(rule α_le)
    also have "d_OUT f a + … < d_OUT f a + (weight Γ a - d_OUT f a)" using ε_less
      using currentD_finite_OUT[OF f'] by (simp add: ennreal_add_left_cancel_less)
    also have "… = weight Γ a"
      using a_le by simp
    finally show "d_OUT h a < weight Γ a" by(simp add: add_left_mono)
  qed
  then show ?thesis using h w' by(blast intro: hindered.intros)
qed

end

corollary hindered_reduce_current: -- ‹Corollary 6.8›
  fixes ε g
  defines "ε ≡ ∑+ x∈B Γ. d_IN g x - d_OUT g x"
  assumes g: "current Γ g"
  and ε_finite: "ε ≠ ⊤"
  and hindered: "hindered_by (Γ ⊖ g) ε"
  shows "hindered Γ"
proof -
  def Γ'  "Γ⦇weight := λx. if x ∈ A Γ then weight Γ x - d_OUT g x else weight Γ x⦈"
  have Γ'_sel [simp]:
    "edge Γ' = edge Γ"
    "A Γ' = A Γ"
    "B Γ' = B Γ"
    "weight Γ' x = (if x ∈ A Γ then weight Γ x - d_OUT g x else weight Γ x)"
    "vertex Γ' = vertex Γ"
    "web.more Γ' = web.more Γ"
    for x by(simp_all add: Γ'_def)
  have "countable_bipartite_web Γ'"
    by unfold_locales(simp_all add: A_in B_out A_vertex disjoint bipartite_V no_loop weight_outside currentD_outside_OUT[OF g]  currentD_weight_OUT[OF g] edge_antiparallel, rule bipartite_E)
  then interpret Γ': countable_bipartite_web Γ' .
  let ?u = "λx. (d_IN g x - d_OUT g x) * indicator (- A Γ) x"

  have "hindered Γ'"
  proof(rule Γ'.hindered_reduce)
    show "?u x = 0" if "x ∉ B Γ'" for x using that bipartite_V
      by(cases "vertex Γ' x")(auto simp add: currentD_outside_OUT[OF g] currentD_outside_IN[OF g])

    have *: "(∑+ x∈B Γ'. ?u x) = ε" using disjoint
      by(auto intro!: nn_integral_cong simp add: ε_def nn_integral_count_space_indicator currentD_outside_OUT[OF g] currentD_outside_IN[OF g] not_vertex split: split_indicator)
    thus "(∑+ x∈B Γ'. ?u x) ≠ ⊤" using ε_finite by simp

    have **: "Γ'⦇weight := weight Γ' - ?u⦈ = Γ ⊖ g"
      using currentD_weight_IN[OF g] currentD_OUT_IN[OF g] currentD_IN[OF g] currentD_finite_OUT[OF g]
      by(intro web.equality)(simp_all add: fun_eq_iff diff_diff_ennreal' ennreal_diff_le_mono_left)
    show "hindered_by (Γ'⦇weight := weight Γ' - ?u⦈) (∑+ x∈B Γ'. ?u x)"
      unfolding * ** by(fact hindered)
    show "(λx. (d_IN g x - d_OUT g x) * indicator (- A Γ) x) ≤ weight Γ'"
      using currentD_weight_IN[OF g]
      by (simp add: le_fun_def ennreal_diff_le_mono_left)
  qed
  then show ?thesis
    by(rule hindered_mono_web[rotated -1]) simp_all
qed

end

subsection ‹Reduced weight in a loose web›

definition reduce_weight :: "('v, 'more) web_scheme ⇒ 'v ⇒ real ⇒ ('v, 'more) web_scheme"
where "reduce_weight Γ x r = Γ⦇weight := λy. weight Γ y - (if x = y then r else 0)⦈"

lemma reduce_weight_sel [simp]:
  "edge (reduce_weight Γ x r) = edge Γ"
  "A (reduce_weight Γ x r) = A Γ"
  "B (reduce_weight Γ x r) = B Γ"
  "vertex (reduce_weight Γ x r) = vertex Γ"
  "weight (reduce_weight Γ x r) y = (if x = y then weight Γ x - r else weight Γ y)"
  "web.more (reduce_weight Γ x r) = web.more Γ"
by(simp_all add: reduce_weight_def zero_ennreal_def[symmetric] vertex_def fun_eq_iff)

lemma essential_reduce_weight [simp]: "essential (reduce_weight Γ x r) = essential Γ"
by(simp add: fun_eq_iff essential_def)

lemma roofed_reduce_weight [simp]: "roofed_gen (reduce_weight Γ x r) = roofed_gen Γ"
by(simp add: fun_eq_iff roofed_def)

context countable_bipartite_web begin

context begin
private datatype (plugins del: transfer size) 'a vertex = SOURCE | SINK | Inner (inner: 'a)

private lemma notin_range_Inner:  "x ∉ range Inner ⟷ x = SOURCE ∨ x = SINK"
by(cases x) auto

private lemma inj_Inner [simp]: "⋀A. inj_on Inner A"
by(simp add: inj_on_def)

lemma unhinder_bipartite:
  assumes h: "⋀n :: nat. current Γ (h n)"
  and SAT: "⋀n. (B Γ ∩ V) - {b} ⊆ SAT Γ (h n)"
  and b: "b ∈ B Γ"
  and IN: "(SUP n. d_IN (h n) b) = weight Γ b"
  and h0_b: "⋀n. d_IN (h 0) b ≤ d_IN (h n) b"
  and b_V: "b ∈ V"
  shows "∃h'. current Γ h' ∧ wave Γ h' ∧ B Γ ∩ V ⊆ SAT Γ h'"
proof -
  write Inner ("⟨_⟩")
  def edge'  "λxo yo. case (xo, yo) of
      (⟨x⟩, ⟨y⟩) ⇒ edge Γ x y ∨ edge Γ y x
    | (⟨x⟩, SINK) ⇒ x ∈ A Γ
    | (SOURCE, ⟨y⟩) ⇒ y = b
    | (SINK, ⟨x⟩) ⇒ x ∈ A Γ
    | _ ⇒ False"
  have edge'_simps [simp]:
    "edge' ⟨x⟩ ⟨y⟩ ⟷ edge Γ x y ∨ edge Γ y x"
    "edge' ⟨x⟩ SINK ⟷ x ∈ A Γ"
    "edge' SOURCE yo ⟷ yo = ⟨b⟩"
    "edge' SINK ⟨x⟩ ⟷ x ∈ A Γ"
    "edge' SINK SINK ⟷ False"
    "edge' xo SOURCE ⟷ False"
    for x y yo xo by(simp_all add: edge'_def split: vertex.split)
  have edge'E: "thesis" if "edge' xo yo"
    "⋀x y. ⟦ xo = ⟨x⟩; yo = ⟨y⟩; edge Γ x y ∨ edge Γ y x ⟧ ⟹ thesis"
    "⋀x. ⟦ xo = ⟨x⟩; yo = SINK; x ∈ A Γ ⟧ ⟹ thesis"
    "⋀x. ⟦ xo = SOURCE; yo = ⟨b⟩ ⟧ ⟹ thesis"
    "⋀y. ⟦ xo = SINK; yo = ⟨y⟩; y ∈ A Γ ⟧ ⟹ thesis"
    for xo yo thesis using that by(auto simp add: edge'_def split: option.split_asm vertex.split_asm)
  have edge'_Inner1 [elim!]: "thesis" if "edge' ⟨x⟩ yo"
    "⋀y. ⟦ yo = ⟨y⟩; edge Γ x y ∨ edge Γ y x ⟧ ⟹ thesis"
    "⟦ yo = SINK; x ∈ A Γ ⟧ ⟹ thesis"
    for x yo thesis using that by(auto elim: edge'E)
  have edge'_Inner2 [elim!]: "thesis" if "edge' xo ⟨y⟩"
    "⋀x. ⟦ xo = ⟨x⟩; edge Γ x y ∨ edge Γ y x ⟧ ⟹ thesis"
    "⟦ xo = SOURCE; y = b ⟧ ⟹ thesis"
    "⟦ xo = SINK; y ∈ A Γ ⟧ ⟹ thesis"
    for xo y thesis using that by(auto elim: edge'E)
  have edge'_SINK1 [elim!]: "thesis" if "edge' SINK yo"
    "⋀y. ⟦ yo = ⟨y⟩; y ∈ A Γ ⟧ ⟹ thesis"
    for yo thesis using that by(auto elim: edge'E)
  have edge'_SINK2 [elim!]: "thesis" if "edge' xo SINK"
    "⋀x. ⟦ xo = ⟨x⟩; x ∈ A Γ ⟧ ⟹ thesis"
    for xo thesis using that by(auto elim: edge'E)

  def cap  "λxoyo. case xoyo of
      (⟨x⟩, ⟨y⟩) ⇒ if edge Γ x y then h 0 (x, y) else if edge Γ y x then max (weight Γ x) (weight Γ y) else 0
    | (⟨x⟩, SINK) ⇒ if x ∈ A Γ then weight Γ x - d_OUT (h 0) x else 0
    | (SOURCE, yo) ⇒ if yo = ⟨b⟩ then weight Γ b - d_IN (h 0) b else 0
    | (SINK, ⟨y⟩) ⇒ if y ∈ A Γ then weight Γ y else 0
    | _ ⇒ 0"
  have cap_simps [simp]:
    "cap (⟨x⟩, ⟨y⟩) = (if edge Γ x y then h 0 (x, y) else if edge Γ y x then max (weight Γ x) (weight Γ y) else 0)"
    "cap (⟨x⟩, SINK) = (if x ∈ A Γ then weight Γ x - d_OUT (h 0) x else 0)"
    "cap (SOURCE, yo) = (if yo = ⟨b⟩ then weight Γ b - d_IN (h 0) b else 0)"
    "cap (SINK, ⟨y⟩) = (if y ∈ A Γ then weight Γ y else 0)"
    "cap (SINK, SINK) = 0"
    "cap (xo, SOURCE) = 0"
    for x y yo xo by(simp_all add: cap_def split: vertex.split)
  def Ψ  "⦇edge = edge', capacity = cap, source = SOURCE, sink = SINK⦈"
  have Ψ_sel [simp]:
    "edge Ψ = edge'"
    "capacity Ψ = cap"
    "source Ψ = SOURCE"
    "sink Ψ = SINK"
    by(simp_all add: Ψ_def)

  have cap_outside1: "¬ vertex Γ x ⟹ cap (⟨x⟩, y) = 0" for x y using A_vertex
    by(cases y)(auto simp add: vertex_def)
  have capacity_A_weight: "d_OUT cap ⟨x⟩ ≤ 2 * weight Γ x" if "x ∈ A Γ" for x
  proof -
    have "d_OUT cap ⟨x⟩ ≤ (∑+ y. h 0 (x, inner y) * indicator (range Inner) y + weight Γ x * indicator {SINK} y)"
      using that disjoint unfolding d_OUT_def
      by(auto intro!: nn_integral_mono diff_le_self_ennreal simp add: A_in notin_range_Inner  split: split_indicator)
    also have "… = (∑+ y∈range Inner. h 0 (x, inner y)) + weight Γ x"
      by(auto simp add: nn_integral_count_space_indicator nn_integral_add)
    also have "(∑+ y∈range Inner. h 0 (x, inner y)) = d_OUT (h 0) x"
      by(simp add: d_OUT_def nn_integral_count_space_reindex)
    also have "… ≤ weight Γ x" using h by(rule currentD_weight_OUT)
    finally show ?thesis unfolding one_add_one[symmetric] distrib_right by(simp add: add_right_mono)
  qed
  have flow_attainability: "flow_attainability Ψ"
  proof
    have "EΨ ⊆ (λ(x, y). (⟨x⟩, ⟨y⟩)) ` E ∪ (λ(x, y). (⟨y⟩, ⟨x⟩)) ` E ∪ (λx. (⟨x⟩, SINK)) ` A Γ ∪ (λx. (SINK, ⟨x⟩)) ` A Γ ∪ {(SOURCE, ⟨b⟩)}"
      by(auto simp add: edge'_def split: vertex.split_asm)
    moreover have "countable (A Γ)" using A_vertex by(rule countable_subset) simp
    ultimately show "countable EΨ" by(auto elim: countable_subset)
  next
    fix v
    assume "v ≠ sink Ψ"
    then consider (source) "v = SOURCE" | (A) x where "v = ⟨x⟩" "x ∈ A Γ"
      | (B) y where "v = ⟨y⟩" "y ∉ A Γ" "y ∈ V" | (outside) x where "v = ⟨x⟩" "x ∉ V"
      by(cases v) auto
    then show "d_IN (capacity Ψ) v ≠ ⊤ ∨ d_OUT (capacity Ψ) v ≠ ⊤"
    proof cases
      case source thus ?thesis by(simp add: d_IN_def)
    next
      case (A x)
      thus ?thesis using capacity_A_weight[of x] by (auto simp: top_unique ennreal_mult_eq_top_iff)
    next
      case (B y)
      have "d_IN (capacity Ψ) v ≤ (∑+ x. h 0 (inner x, y) * indicator (range Inner) x + weight Γ b * indicator {SOURCE} x)"
        using B bipartite_V
        by(auto 4 4 intro!: nn_integral_mono simp add: diff_le_self_ennreal   d_IN_def notin_range_Inner nn_integral_count_space_indicator currentD_outside[OF h] split: split_indicator dest: bipartite_E)
      also have "… = (∑+ x∈range Inner. h 0 (inner x, y)) + weight Γ b"
        by(simp add: nn_integral_add nn_integral_count_space_indicator)
      also have "(∑+ x∈range Inner. h 0 (inner x, y)) = d_IN (h 0) y"
        by(simp add: d_IN_def nn_integral_count_space_reindex)
      also have "d_IN (h 0) y ≤ weight Γ y" using h by(rule currentD_weight_IN)
      finally show ?thesis by(auto simp add: top_unique add_right_mono split: if_split_asm)
    next
      case outside
      hence "d_OUT (capacity Ψ) v = 0" using A_vertex
        by(auto simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space cap_def vertex_def split: vertex.split)
      thus ?thesis by simp
    qed
  next
    show "capacity Ψ e ≠ ⊤" for e
      by(auto simp add: cap_def max_def vertex_def currentD_finite[OF h] split: vertex.split prod.split)
    show "capacity Ψ e = 0" if "e ∉ EΨ" for e using that
      by(auto simp add: cap_def max_def split: prod.split; split vertex.split)+
    show "¬ edge Ψ x (source Ψ)" for x using b by(auto simp add: B_out)
    show "¬ edge Ψ x x" for x by(cases x)(simp_all add: no_loop)
    show "source Ψ ≠ sink Ψ" by simp
  qed
  then interpret Ψ: flow_attainability "Ψ" .
  def α  "SUP f:{f. flow Ψ f}. value_flow Ψ f"

  def f  "λn xoyo. case xoyo of
      (⟨x⟩, ⟨y⟩) ⇒ if edge Γ x y then h 0 (x, y) - h n (x, y) else if edge Γ y x then h n (y, x) - h 0 (y, x) else 0
    | (SOURCE, ⟨y⟩) ⇒ if y = b then d_IN (h n) b - d_IN (h 0) b else 0
    | (⟨x⟩, SINK) ⇒ if x ∈ A Γ then d_OUT (h n) x - d_OUT (h 0) x else 0
    | (SINK, ⟨y⟩) ⇒ if y ∈ A Γ then d_OUT (h 0) y - d_OUT (h n) y else 0
    | _ ⇒ 0"
  have f_cases: thesis if "⋀x y. e = (⟨x⟩, ⟨y⟩) ⟹ thesis" "⋀y. e = (SOURCE, ⟨y⟩) ⟹ thesis"
    "⋀x. e = (⟨x⟩, SINK) ⟹ thesis" "⋀y. e = (SINK, ⟨y⟩) ⟹ thesis" "e = (SINK, SINK) ⟹ thesis"
    "⋀xo. e = (xo, SOURCE) ⟹ thesis" "e = (SOURCE, SINK) ⟹ thesis"
    for e :: "'v vertex edge" and thesis
    using that by(cases e; cases "fst e" "snd e" rule: vertex.exhaust[case_product vertex.exhaust]) simp_all
  have f_simps [simp]:
    "f n (⟨x⟩, ⟨y⟩) = (if edge Γ x y then h 0 (x, y) - h n (x, y) else if edge Γ y x then h n (y, x) - h 0 (y, x) else 0)"
    "f n (SOURCE, ⟨y⟩) = (if y = b then d_IN (h n) b - d_IN (h 0) b else 0)"
    "f n (⟨x⟩, SINK) = (if x ∈ A Γ then d_OUT (h n) x - d_OUT (h 0) x else 0)"
    "f n (SINK, ⟨y⟩) = (if y ∈ A Γ then d_OUT (h 0) y - d_OUT (h n) y else 0)"
    "f n (SOURCE, SINK) = 0"
    "f n (SINK, SINK) = 0"
    "f n (xo, SOURCE) = 0"
    for n x y xo by(simp_all add: f_def split: vertex.split)
  have OUT_f_SOURCE: "d_OUT (f n) SOURCE = d_IN (h n) b - d_IN (h 0) b" for n
  proof(rule trans)
    show "d_OUT (f n) SOURCE = (∑+ y. f n (SOURCE, y) * indicator {⟨b⟩} y)" unfolding d_OUT_def
      apply(rule nn_integral_cong) subgoal for x by(cases x) auto done
    show "… = d_IN (h n) b - d_IN (h 0) b" using h0_b[of n]
      by(auto simp add: max_def)
  qed

  have OUT_f_outside: "d_OUT (f n) ⟨x⟩ = 0" if "x ∉ V" for x n using A_vertex that
    apply(clarsimp simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0)
    subgoal for y by(cases y)(auto simp add: vertex_def)
    done
  have IN_f_outside: "d_IN (f n) ⟨x⟩ = 0" if "x ∉ V" for x n using b_V that
    apply(clarsimp simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0)
    subgoal for y by(cases y)(auto simp add: currentD_outside_OUT[OF h] vertex_def)
    done

  have f: "flow Ψ (f n)" for n
  proof
    show f_le: "f n e ≤ capacity Ψ e" for e
      using currentD_weight_out[OF h] currentD_weight_IN[OF h] currentD_weight_OUT[OF h]
      by(cases e rule: f_cases)
        (auto dest: edge_antiparallel simp add: not_le le_max_iff_disj intro: ennreal_minus_mono ennreal_diff_le_mono_left)

    fix xo
    assume "xo ≠ source Ψ" "xo ≠ sink Ψ"
    then consider (A) x where "xo = ⟨x⟩" "x ∈ A Γ" | (B) x where "xo = ⟨x⟩" "x ∈ B Γ" "x ∈ V"
      | (outside) x where "xo = ⟨x⟩" "x ∉ V" using bipartite_V by(cases xo) auto
    then show "KIR (f n) xo"
    proof cases
      case outside
      thus ?thesis by(simp add: OUT_f_outside IN_f_outside)
    next
      case A

      have finite1: "(∑+ y. h n (x, y) * indicator A y) ≠ ⊤" for A n
        using currentD_finite_OUT[OF h, of n x, unfolded d_OUT_def]
        by(rule neq_top_trans)(auto intro!: nn_integral_mono simp add: split: split_indicator)

      let ?h0_ge_hn = "{y. h 0 (x, y) ≥ h n (x, y)}"
      let ?h0_lt_hn = "{y. h 0 (x, y) < h n (x, y)}"

      have "d_OUT (f n) ⟨x⟩ = (∑+ y. f n (⟨x⟩, y) * indicator (range Inner) y + f n (⟨x⟩, y) * indicator {SINK} y)"
        unfolding d_OUT_def by(intro nn_integral_cong)(auto split: split_indicator simp add: notin_range_Inner)
      also have "… = (∑+ y∈range Inner. f n (⟨x⟩, y)) + f n (⟨x⟩, SINK)"
        by(simp add: nn_integral_add nn_integral_count_space_indicator max.left_commute max.commute)
      also have "(∑+ y∈range Inner. f n (⟨x⟩, y)) = (∑+ y. h 0 (x, y) - h n (x, y))" using A
        apply(simp add: nn_integral_count_space_reindex cong: nn_integral_cong_simp outgoing_def)
        apply(auto simp add: nn_integral_count_space_indicator outgoing_def A_in max.absorb1 currentD_outside[OF h] intro!: nn_integral_cong split: split_indicator dest: edge_antiparallel)
        done
      also have "… = (∑+ y. h 0 (x, y) * indicator ?h0_ge_hn y) - (∑+ y. h n (x, y) * indicator ?h0_ge_hn y)"
        apply(subst nn_integral_diff[symmetric])
        apply(simp_all add: AE_count_space finite1 split: split_indicator)
        apply(rule nn_integral_cong; auto simp add: max_def not_le split: split_indicator)
        by (metis diff_eq_0_ennreal le_less not_le top_greatest)
      also have "(∑+ y. h n (x, y) * indicator ?h0_ge_hn y) = d_OUT (h n) x - (∑+ y. h n (x, y) * indicator ?h0_lt_hn y)"
        unfolding d_OUT_def
        apply(subst nn_integral_diff[symmetric])
        apply(auto simp add: AE_count_space finite1 currentD_finite[OF h] split: split_indicator intro!: nn_integral_cong)
        done
      also have "(∑+ y. h 0 (x, y) * indicator ?h0_ge_hn y) - … + f n (⟨x⟩, SINK) =
        (∑+ y. h 0 (x, y) * indicator ?h0_ge_hn y) + (∑+ y. h n (x, y) * indicator ?h0_lt_hn y) - min (d_OUT (h n) x) (d_OUT (h 0) x)"
        using finite1[of n "{_}"] A finite1[of n UNIV]
        apply (subst diff_diff_ennreal')
        apply (auto simp: d_OUT_def finite1 AE_count_space nn_integral_diff[symmetric] top_unique nn_integral_add[symmetric]
                    split: split_indicator intro!: nn_integral_mono ennreal_diff_self)
        apply (simp add: min_def not_le diff_eq_0_ennreal finite1 less_top[symmetric])
        apply (subst diff_add_assoc2_ennreal)
        apply (auto simp: add_diff_eq_ennreal intro!: nn_integral_mono split: split_indicator)
        apply (subst diff_diff_commute_ennreal)
        apply (simp add: ennreal_add_diff_cancel )
        done
      also have "… = (∑+ y. h n (x, y) * indicator ?h0_lt_hn y) - (d_OUT (h 0) x - (∑+ y. h 0 (x, y) * indicator ?h0_ge_hn y)) + f n (SINK, ⟨x⟩)"
        apply(rule sym)
        using finite1[of 0 "{_}"] A finite1[of 0 UNIV]
        apply (subst diff_diff_ennreal')
        apply (auto simp: d_OUT_def finite1 AE_count_space nn_integral_diff[symmetric] top_unique nn_integral_add[symmetric]
                    split: split_indicator intro!: nn_integral_mono ennreal_diff_self)
        apply (simp add: min_def not_le diff_eq_0_ennreal finite1 less_top[symmetric])
        apply (subst diff_add_assoc2_ennreal)
        apply (auto simp: add_diff_eq_ennreal intro!: nn_integral_mono split: split_indicator)
        apply (subst diff_diff_commute_ennreal)
        apply (simp_all add: ennreal_add_diff_cancel ac_simps)
        done
      also have "d_OUT (h 0) x - (∑+ y. h 0 (x, y) * indicator ?h0_ge_hn y) = (∑+ y. h 0 (x, y) * indicator ?h0_lt_hn y)"
        unfolding d_OUT_def
        apply(subst nn_integral_diff[symmetric])
        apply(auto simp add: AE_count_space finite1 currentD_finite[OF h] split: split_indicator intro!: nn_integral_cong)
        done
      also have "(∑+ y. h n (x, y) * indicator ?h0_lt_hn y) - … = (∑+ y. h n (x, y) - h 0 (x, y))"
        apply(subst nn_integral_diff[symmetric])
        apply(simp_all add: AE_count_space finite1 order.strict_implies_order split: split_indicator)
        apply(rule nn_integral_cong; auto simp add: currentD_finite[OF h] top_unique less_top[symmetric] not_less split: split_indicator intro!: diff_eq_0_ennreal)
        done
      also have "… = (∑+ y∈range Inner. f n (y, ⟨x⟩))" using A
        apply(simp add: nn_integral_count_space_reindex cong: nn_integral_cong_simp outgoing_def)
        apply(auto simp add: nn_integral_count_space_indicator outgoing_def A_in max.commute currentD_outside[OF h] intro!: nn_integral_cong split: split_indicator dest: edge_antiparallel)
        done
      also have "… + f n (SINK, ⟨x⟩) = (∑+ y. f n (y, ⟨x⟩) * indicator (range Inner) y + f n (y, ⟨x⟩) * indicator {SINK} y)"
        by(simp add: nn_integral_add nn_integral_count_space_indicator)
      also have "… = d_IN (f n) ⟨x⟩"
        using A b disjoint unfolding d_IN_def
        by(intro nn_integral_cong)(auto split: split_indicator simp add: notin_range_Inner)
      finally show ?thesis using A by simp
    next
      case (B x)

      have finite1: "(∑+ y. h n (y, x) * indicator A y) ≠ ⊤" for A n
        using currentD_finite_IN[OF h, of n x, unfolded d_IN_def]
        by(rule neq_top_trans)(auto intro!: nn_integral_mono split: split_indicator)

      have finite_h[simp]: "h n (y, x) < ⊤" for y n
        using finite1[of n "{y}"] by (simp add: less_top)

      let ?h0_gt_hn = "{y. h 0 (y, x) > h n (y, x)}"
      let ?h0_le_hn = "{y. h 0 (y, x) ≤ h n (y, x)}"

      have eq: "d_IN (h 0) x + f n (SOURCE, ⟨x⟩) = d_IN (h n) x"
      proof(cases "x = b")
        case True with currentD_finite_IN[OF h, of _ b] show ?thesis
          by(simp add: add_diff_self_ennreal h0_b)
      next
        case False
        with B SAT have "x ∈ SAT Γ (h n)" "x ∈ SAT Γ (h 0)" by auto
        with B disjoint have "d_IN (h n) x = d_IN (h 0) x" by(auto simp add: currentD_SAT[OF h])
        thus ?thesis using False by(simp add: currentD_finite_IN[OF h])
      qed

      have "d_IN (f n) ⟨x⟩ = (∑+ y. f n (y, ⟨x⟩) * indicator (range Inner) y + f n (y, ⟨x⟩) * indicator {SOURCE} y)"
        using B disjoint unfolding d_IN_def
        by(intro nn_integral_cong)(auto split: split_indicator simp add: notin_range_Inner)
      also have "… = (∑+ y∈range Inner. f n (y, ⟨x⟩)) + f n (SOURCE, ⟨x⟩)" using h0_b[of n]
        by(simp add: nn_integral_add nn_integral_count_space_indicator max_def)
      also have "(∑+ y∈range Inner. f n (y, ⟨x⟩)) = (∑+ y. h 0 (y, x) - h n (y, x))"
        using B disjoint
        apply(simp add: nn_integral_count_space_reindex cong: nn_integral_cong_simp outgoing_def)
        apply(auto simp add: nn_integral_count_space_indicator outgoing_def B_out max.commute currentD_outside[OF h] intro!: nn_integral_cong split: split_indicator dest: edge_antiparallel)
        done
      also have "… = (∑+ y. h 0 (y, x) * indicator ?h0_gt_hn y) - (∑+ y. h n (y, x) * indicator ?h0_gt_hn y)"
        apply(subst nn_integral_diff[symmetric])
        apply(simp_all add: AE_count_space finite1 order.strict_implies_order split: split_indicator)
        apply(rule nn_integral_cong; auto simp add: currentD_finite[OF h] top_unique less_top[symmetric] not_less split: split_indicator intro!: diff_eq_0_ennreal)
        done
      also have eq_h_0: "(∑+ y. h 0 (y, x) * indicator ?h0_gt_hn y) = d_IN (h 0) x - (∑+ y. h 0 (y, x) * indicator ?h0_le_hn y)"
        unfolding d_IN_def
        apply(subst nn_integral_diff[symmetric])
        apply(auto simp add: AE_count_space finite1 currentD_finite[OF h] split: split_indicator intro!: nn_integral_cong)
        done
      also have eq_h_n: "(∑+ y. h n (y, x) * indicator ?h0_gt_hn y) = d_IN (h n) x - (∑+ y. h n (y, x) * indicator ?h0_le_hn y)"
        unfolding d_IN_def
        apply(subst nn_integral_diff[symmetric])
        apply(auto simp add: AE_count_space finite1 currentD_finite[OF h] split: split_indicator intro!: nn_integral_cong)
        done
      also have "d_IN (h 0) x - (∑+ y. h 0 (y, x) * indicator ?h0_le_hn y) - (d_IN (h n) x - (∑+ y. h n (y, x) * indicator ?h0_le_hn y)) + f n (SOURCE, ⟨x⟩) =
                (∑+ y. h n (y, x) * indicator ?h0_le_hn y) - (∑+ y. h 0 (y, x) * indicator ?h0_le_hn y)"
        apply (subst diff_add_assoc2_ennreal)
        subgoal by (auto simp add: eq_h_0[symmetric] eq_h_n[symmetric] split: split_indicator intro!: nn_integral_mono)
        apply (subst diff_add_assoc2_ennreal)
        subgoal by (auto simp: d_IN_def split: split_indicator intro!: nn_integral_mono)
        apply (subst diff_diff_commute_ennreal)
        apply (subst diff_diff_ennreal')
        subgoal
          by (auto simp: d_IN_def split: split_indicator intro!: nn_integral_mono) []
        subgoal
          unfolding eq_h_n[symmetric]
          by (rule add_increasing2)
             (auto simp add: d_IN_def split: split_indicator intro!: nn_integral_mono)
        apply (subst diff_add_assoc2_ennreal[symmetric])
        unfolding eq
        using currentD_finite_IN[OF h]
        apply simp_all
        done
      also have "(∑+ y. h n (y, x) * indicator ?h0_le_hn y) - (∑+ y. h 0 (y, x) * indicator ?h0_le_hn y) = (∑+ y. h n (y, x) - h 0 (y, x))"
        apply(subst nn_integral_diff[symmetric])
        apply(simp_all add: AE_count_space max_def finite1 split: split_indicator)
        apply(rule nn_integral_cong; auto simp add: not_le split: split_indicator)
        by (metis diff_eq_0_ennreal le_less not_le top_greatest)
      also have "… = (∑+ y∈range Inner. f n (⟨x⟩, y))" using B disjoint
        apply(simp add: nn_integral_count_space_reindex cong: nn_integral_cong_simp outgoing_def)
        apply(auto simp add: B_out currentD_outside[OF h] max.commute intro!: nn_integral_cong split: split_indicator dest: edge_antiparallel)
        done
      also have "… = (∑+ y. f n (⟨x⟩, y) * indicator (range Inner) y)"
        by(simp add: nn_integral_add nn_integral_count_space_indicator max.left_commute max.commute)
      also have "… = d_OUT (f n) ⟨x⟩"  using B disjoint
        unfolding d_OUT_def by(intro nn_integral_cong)(auto split: split_indicator simp add: notin_range_Inner)
      finally show ?thesis using B by(simp)
    qed
  qed

  have "weight Γ b - d_IN (h 0) b = (SUP n. value_flow Ψ (f n))"
    using OUT_f_SOURCE currentD_finite_IN[OF h, of 0 b] IN
    by (simp add: SUP_diff_ennreal less_top)
  also have "(SUP n. value_flow Ψ (f n)) ≤ α" unfolding α_def
    apply(rule SUP_least)
    apply(rule SUP_upper)
    apply(simp add: f)
    done
  also have "α ≤ weight Γ b - d_IN (h 0) b" unfolding α_def
  proof(rule SUP_least; clarsimp)
    fix f
    assume f: "flow Ψ f"
    have "d_OUT f SOURCE = (∑+ y. f (SOURCE, y) * indicator {⟨b⟩} y)" unfolding d_OUT_def
      apply(rule nn_integral_cong)
      subgoal for x using flowD_capacity[OF f, of "(SOURCE, x)"]
        by(auto split: split_indicator)
      done
    also have "… = f (SOURCE, ⟨b⟩)" by(simp add: max_def)
    also have "… ≤ weight Γ b - d_IN (h 0) b" using flowD_capacity[OF f, of "(SOURCE, ⟨b⟩)"] by simp
    finally show "d_OUT f SOURCE ≤ …" .
  qed
  ultimately have α: "α = weight Γ b - d_IN (h 0) b" by(rule antisym[rotated])
  hence α_finite: "α ≠ ⊤" by simp

  from Ψ.ex_max_flow
  obtain g where g: "flow Ψ g"
    and value_g: "value_flow Ψ g = α"
    and IN_g: "⋀x. d_IN g x ≤ value_flow Ψ g" unfolding α_def by blast

  have g_le_h0: "g (⟨x⟩, ⟨y⟩) ≤ h 0 (x, y)" if "edge Γ x y" for x y
    using flowD_capacity[OF g, of "(⟨x⟩, ⟨y⟩)"] that by simp
  note [simp] = Ψ.flowD_finite[OF g]

  have g_SOURCE: "g (SOURCE, ⟨x⟩) = (if x = b then α else 0)" for x
  proof(cases "x = b")
    case True
    have "g (SOURCE, ⟨x⟩) = (∑+ y. g (SOURCE, y) * indicator {⟨x⟩} y)" by(simp add: max_def)
    also have "… = value_flow Ψ g" unfolding d_OUT_def using True
      by(intro nn_integral_cong)(auto split: split_indicator intro: Ψ.flowD_outside[OF g])
    finally show ?thesis using value_g by(simp add: True)
  qed(simp add: Ψ.flowD_outside[OF g])

  let ?g = "λ(x, y). g (⟨y⟩, ⟨x⟩)"

  def h'  "h 0 ⊕ ?g"
  have h'_simps: "h' (x, y) = (if edge Γ x y then h 0 (x, y) + g (⟨y⟩, ⟨x⟩) - g (⟨x⟩, ⟨y⟩) else 0)" for x y
    by(simp add: h'_def)

  have OUT_h'_B [simp]: "d_OUT h' x = 0" if "x ∈ B Γ" for x using that unfolding d_OUT_def
    by(simp add: nn_integral_0_iff emeasure_count_space_eq_0)(simp add: h'_simps B_out)
  have IN_h'_A [simp]: "d_IN h' x = 0" if "x ∈ A Γ" for x using that unfolding d_IN_def
    by(simp add: nn_integral_0_iff emeasure_count_space_eq_0)(simp add: h'_simps A_in)
  have h'_outside: "h' e = 0" if "e ∉ E" for e unfolding h'_def using that by(rule plus_flow_outside)
  have OUT_h'_outside: "d_OUT h' x = 0" and IN_h'_outside: "d_IN h' x = 0" if "x ∉ V" for x using that
    by(auto simp add: d_OUT_def d_IN_def nn_integral_0_iff emeasure_count_space_eq_0 vertex_def intro: h'_outside)

  have g_le_OUT: "g (SINK, ⟨x⟩) ≤ d_OUT g ⟨x⟩" for x
    by (subst flowD_KIR[OF g]) (simp_all add: d_IN_ge_point)

  have OUT_g_A: "d_OUT ?g x = d_OUT g ⟨x⟩ - g (SINK, ⟨x⟩)" if "x ∈ A Γ" for x
  proof -
    have "d_OUT ?g x = (∑+ y∈range Inner. g (y, ⟨x⟩))"
      by(simp add: nn_integral_count_space_reindex d_OUT_def)
    also have "… = d_IN g ⟨x⟩ - (∑+ y. g (y, ⟨x⟩) * indicator {SINK} y)" unfolding d_IN_def
      using that b disjoint flowD_capacity[OF g, of "(SOURCE, ⟨x⟩)"]
      by(subst nn_integral_diff[symmetric])
        (auto simp add: nn_integral_count_space_indicator notin_range_Inner max_def intro!: nn_integral_cong split: split_indicator if_split_asm)
    also have "… = d_OUT g ⟨x⟩ - g (SINK, ⟨x⟩)" by(simp add: flowD_KIR[OF g] max_def)
    finally show ?thesis .
  qed
  have IN_g_A: "d_IN ?g x = d_OUT g ⟨x⟩ - g (⟨x⟩, SINK)" if "x ∈ A Γ" for x
  proof -
    have "d_IN ?g x = (∑+ y∈range Inner. g (⟨x⟩, y))"
      by(simp add: nn_integral_count_space_reindex d_IN_def)
    also have "… = d_OUT g ⟨x⟩ - (∑+ y. g (⟨x⟩, y) * indicator {SINK} y)" unfolding d_OUT_def
      using that b disjoint flowD_capacity[OF g, of "(⟨x⟩, SOURCE)"]
      by(subst nn_integral_diff[symmetric])
        (auto simp add: nn_integral_count_space_indicator notin_range_Inner max_def intro!: nn_integral_cong split: split_indicator if_split_asm)
    also have "… = d_OUT g ⟨x⟩ - g (⟨x⟩, SINK)" by(simp add: max_def)
    finally show ?thesis .
  qed
  have OUT_g_B: "d_OUT ?g x = d_IN g ⟨x⟩ - g (SOURCE, ⟨x⟩)" if "x ∈ B Γ" for x
  proof -
    have "d_OUT ?g x = (∑+ y∈range Inner. g (y, ⟨x⟩))"
      by(simp add: nn_integral_count_space_reindex d_OUT_def)
    also have "… = d_IN g ⟨x⟩ - (∑+ y. g (y, ⟨x⟩) * indicator {SOURCE} y)" unfolding d_IN_def
      using that b disjoint flowD_capacity[OF g, of "(SINK, ⟨x⟩)"]
      by(subst nn_integral_diff[symmetric])
        (auto simp add: nn_integral_count_space_indicator notin_range_Inner max_def intro!: nn_integral_cong split: split_indicator if_split_asm)
    also have "… = d_IN g ⟨x⟩ - g (SOURCE, ⟨x⟩)" by(simp add: max_def)
    finally show ?thesis .
  qed
  have IN_g_B: "d_IN ?g x = d_OUT g ⟨x⟩" if "x ∈ B Γ" for x
  proof -
    have "d_IN ?g x = (∑+ y∈range Inner. g (⟨x⟩, y))"
      by(simp add: nn_integral_count_space_reindex d_IN_def)
    also have "… = d_OUT g ⟨x⟩" unfolding d_OUT_def using that disjoint
      by(auto 4 3 simp add: nn_integral_count_space_indicator notin_range_Inner intro!: nn_integral_cong Ψ.flowD_outside[OF g] split: split_indicator)
    finally show ?thesis .
  qed

  have finite_g_IN: "d_IN ?g x ≠ ⊤" for x using α_finite
  proof(rule neq_top_trans)
    have "d_IN ?g x = (∑+ y∈range Inner. g (⟨x⟩, y))"
      by(auto simp add: d_IN_def nn_integral_count_space_reindex)
    also have "… ≤ d_OUT g ⟨x⟩" unfolding d_OUT_def
      by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator)
    also have "… = d_IN g ⟨x⟩" by(rule flowD_KIR[OF g]) simp_all
    also have "… ≤ α" using IN_g value_g by simp
    finally show "d_IN ?g x ≤ α" .
  qed

  have OUT_h'_A: "d_OUT h' x = d_OUT (h 0) x + g (⟨x⟩, SINK) - g (SINK, ⟨x⟩)" if "x ∈ A Γ" for x
  proof -
    have "d_OUT h' x = d_OUT (h 0) x + (∑+ y. ?g (x, y) * indicator E (x, y)) - (∑+ y. ?g (y, x) * indicator E (x, y))"
      unfolding h'_def
      apply(subst OUT_plus_flow[of Γ "h 0" ?g, OF currentD_outside'[OF h]])
      apply(auto simp add: g_le_h0 finite_g_IN)
      done
    also have "(∑+ y. ?g (x, y) * indicator E (x, y)) = d_OUT ?g x" unfolding d_OUT_def using that
      by(auto simp add: A_in split: split_indicator intro!: nn_integral_cong Ψ.flowD_outside[OF g])
    also have "…  = d_OUT g ⟨x⟩ - g (SINK, ⟨x⟩)" using that by(rule OUT_g_A)
    also have "(∑+ y. ?g (y, x) * indicator E (x, y)) = d_IN ?g x" using that unfolding d_IN_def
      by(auto simp add: A_in split: split_indicator intro!: nn_integral_cong Ψ.flowD_outside[OF g])
    also have "… = d_OUT g ⟨x⟩ - g (⟨x⟩, SINK)" using that by(rule IN_g_A)
    also have "d_OUT (h 0) x + (d_OUT g ⟨x⟩ - g (SINK, ⟨x⟩)) - … = d_OUT (h 0) x + g (⟨x⟩, SINK) - g (SINK, ⟨x⟩)"
      apply(simp add: g_le_OUT add_diff_eq_ennreal d_OUT_ge_point)
      apply(subst diff_diff_commute_ennreal)
      apply(simp add: add_increasing d_OUT_ge_point g_le_OUT diff_diff_ennreal')
      apply(subst add.assoc)
      apply(subst (2) add.commute)
      apply(subst add.assoc[symmetric])
      apply(subst ennreal_add_diff_cancel_right)
      apply(simp_all add: Ψ.flowD_finite_OUT[OF g])
      done
    finally show ?thesis .
  qed

  have finite_g_OUT: "d_OUT ?g x ≠ ⊤" for x using α_finite
  proof(rule neq_top_trans)
    have "d_OUT ?g x = (∑+ y∈range Inner. g (y, ⟨x⟩))"
      by(auto simp add: d_OUT_def nn_integral_count_space_reindex)
    also have "… ≤ d_IN g ⟨x⟩" unfolding d_IN_def
      by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator)
    also have "… ≤ α" using IN_g value_g by simp
    finally show "d_OUT ?g x ≤ α" .
  qed

  have IN_h'_B: "d_IN h' x = d_IN (h 0) x + g (SOURCE, ⟨x⟩)" if "x ∈ B Γ" for x
  proof -
    have g_le: "g (SOURCE, ⟨x⟩) ≤ d_IN g ⟨x⟩"
      by (rule d_IN_ge_point)

    have "d_IN h' x = d_IN (h 0) x + (∑+ y. g (⟨x⟩, ⟨y⟩) * indicator E (y, x)) - (∑+ y. g (⟨y⟩, ⟨x⟩) * indicator E (y, x))"
      unfolding h'_def
      by(subst IN_plus_flow[of Γ "h 0" ?g, OF currentD_outside'[OF h]])
        (auto simp add: g_le_h0 finite_g_OUT)
    also have "(∑+ y. g (⟨x⟩, ⟨y⟩) * indicator E (y, x)) = d_IN ?g x" unfolding d_IN_def using that
      by(intro nn_integral_cong)(auto split: split_indicator intro!: Ψ.flowD_outside[OF g] simp add: B_out)
    also have "… = d_OUT g ⟨x⟩" using that by(rule IN_g_B)
    also have "(∑+ y. g (⟨y⟩, ⟨x⟩) * indicator E (y, x)) = d_OUT ?g x" unfolding d_OUT_def using that
      by(intro nn_integral_cong)(auto split: split_indicator intro!: Ψ.flowD_outside[OF g] simp add: B_out)
    also have "… = d_IN g ⟨x⟩ - g (SOURCE, ⟨x⟩)" using that by(rule OUT_g_B)
    also have "d_IN (h 0) x + d_OUT g ⟨x⟩ - … = d_IN (h 0) x + g (SOURCE, ⟨x⟩)"
      using Ψ.flowD_finite_IN[OF g] g_le
      by(cases "d_IN (h 0) x"; cases "d_IN g ⟨x⟩"; cases "d_IN g ⟨x⟩"; cases "g (SOURCE, ⟨x⟩)")
        (auto simp: flowD_KIR[OF g] top_add ennreal_minus_if ennreal_plus_if simp del: ennreal_plus)
    finally show ?thesis .
  qed

  have h': "current Γ h'"
  proof
    fix x
    consider (A) "x ∈ A Γ" | (B) "x ∈ B Γ" | (outside) "x ∉ V" using bipartite_V by auto
    note cases = this

    show "d_OUT h' x ≤ weight Γ x"
    proof(cases rule: cases)
      case A
      then have "d_OUT h' x = d_OUT (h 0) x + g (⟨x⟩, SINK) - g (SINK, ⟨x⟩)" by(simp add: OUT_h'_A)
      also have "… ≤ d_OUT (h 0) x + g (⟨x⟩, SINK)" by(rule diff_le_self_ennreal)
      also have "g (⟨x⟩, SINK) ≤ weight Γ x - d_OUT (h 0) x"
        using flowD_capacity[OF g, of "(⟨x⟩, SINK)"] A by simp
      also have "d_OUT (h 0) x + … = weight Γ x"
        by(simp add: add_diff_eq_ennreal add_diff_inverse_ennreal  currentD_finite_OUT[OF h] currentD_weight_OUT[OF h])
      finally show ?thesis by(simp add: add_left_mono)
    qed(simp_all add: OUT_h'_outside )

    show "d_IN h' x ≤ weight Γ x"
    proof(cases rule: cases)
      case B
      hence "d_IN h' x = d_IN (h 0) x + g (SOURCE, ⟨x⟩)" by(rule IN_h'_B)
      also have "… ≤ weight Γ x"
        by(simp add: g_SOURCE α currentD_weight_IN[OF h] add_diff_eq_ennreal add_diff_inverse_ennreal currentD_finite_IN[OF h])
      finally show ?thesis .
    qed(simp_all add:  IN_h'_outside)
  next
    show "h' e = 0" if "e ∉ E" for e using that by(simp split: prod.split_asm add: h'_simps)
  qed
  moreover
  have SAT_h': "B Γ ∩ V ⊆ SAT Γ h'"
  proof
    show "x ∈ SAT Γ h'" if "x ∈ B Γ ∩ V" for x using that
    proof(cases "x = b")
      case True
      have "d_IN h' x = weight Γ x" using that True
        by(simp add: IN_h'_B g_SOURCE α currentD_weight_IN[OF h] add_diff_eq_ennreal add_diff_inverse_ennreal currentD_finite_IN[OF h])
      thus ?thesis by(simp add: SAT.simps)
    next
      case False
      have "d_IN h' x = d_IN (h 0) x" using that False by(simp add: IN_h'_B g_SOURCE)
      also have "… = weight Γ x"
        using SAT[of 0, THEN subsetD, of x] False that currentD_SAT[OF h, of x 0] disjoint by auto
      finally show ?thesis by(simp add: SAT.simps)
    qed
  qed
  moreover
  have "wave Γ h'"
  proof
    have "separating Γ (B Γ ∩ V)"
    proof
      fix x y p
      assume x: "x ∈ A Γ" and y: "y ∈ B Γ" and p: "path Γ x p y"
      hence Nil: "p ≠ []" using disjoint by(auto simp add: rtrancl_path_simps)
      from rtrancl_path_last[OF p Nil] last_in_set[OF Nil] y rtrancl_path_Range[OF p, of y]
      show "(∃z∈set p. z ∈ B Γ ∩ V) ∨ x ∈ B Γ ∩ V" by(auto intro: vertexI2)
    qed
    moreover have TER: "B Γ ∩ V ⊆ TER h'" using SAT_h' by(auto simp add: SINK)
    ultimately show "separating Γ (TER h')" by(rule separating_weakening)
  qed(rule h')
  ultimately show ?thesis by blast
qed

end

lemma countable_bipartite_web_reduce_weight:
  assumes "weight Γ x ≥ w"
  shows "countable_bipartite_web (reduce_weight Γ x w)"
using bipartite_V A_vertex bipartite_E disjoint assms
by unfold_locales (auto 4 3 simp add: weight_outside )

lemma unhinder: -- "Lemma 6.9"
  assumes loose: "loose Γ"
  and b: "b ∈ B Γ"
  and wb: "weight Γ b > 0"
  and δ: "δ > 0"
  shows "∃ε>0. ε < δ ∧ ¬ hindered (reduce_weight Γ b ε)"
proof(rule ccontr)
  assume "¬ ?thesis"
  hence hindered: "hindered (reduce_weight Γ b ε)" if "ε > 0" "ε < δ" for ε using that by simp

  from b disjoint have bnA: "b ∉ A Γ" by blast

  def wb  "enn2real (weight Γ b)"
  have wb_conv: "weight Γ b = ennreal wb" by(simp add: wb_def less_top[symmetric])
  have wb_pos: "wb > 0" using wb by(simp add: wb_conv)

  def ε  "λn :: nat. min δ wb / (n + 2)"
  have ε_pos: "ε n > 0" for n using wb_pos δ by(simp add: ε_def)
  have ε_nonneg: "0 ≤ ε n" for n using ε_pos[of n] by simp
  have *: "ε n ≤ min wb δ / 2" for n using wb_pos δ
    by(auto simp add: ε_def field_simps min_def)
  have ε_le: "ε n ≤ wb" and ε_less: "ε n < wb" and ε_less_δ: "ε n < δ" and ε_le': "ε n ≤ wb / 2" for n
    using *[of n] ε_pos[of n] by(auto)

  def Γ'  "λn. reduce_weight Γ b (ε n)"
  have Γ'_sel [simp]:
    "edge (Γ' n) = edge Γ"
    "A (Γ' n) = A Γ"
    "B (Γ' n) = B Γ"
    "weight (Γ' n) x = weight Γ x - (if x = b then ε n else 0)"
    "essential (Γ' n) = essential Γ"
    "roofed_gen (Γ' n) = roofed_gen Γ"
    for n x by(simp_all add: Γ'_def)

  have vertex_Γ' [simp]: "vertex (Γ' n) = vertex Γ" for n
    by(simp add: vertex_def fun_eq_iff)

  from wb have "b ∈ V" using weight_outside[of b] by(auto intro: ccontr)
  interpret Γ': countable_bipartite_web "Γ' n" for n unfolding Γ'_def
    using wb_pos by(intro countable_bipartite_web_reduce_weight)(simp_all add: wb_conv ε_le ε_nonneg)

  obtain g where g: "⋀n. current (Γ' n) (g n)"
    and w: "⋀n. wave (Γ' n) (g n)"
    and hind: "⋀n. hindrance (Γ' n) (g n)" using hindered[OF ε_pos, unfolded wb_conv ennreal_less_iff, OF ε_less_δ]
    unfolding hindered.simps Γ'_def by atomize_elim metis
  from g have : "current Γ (g n)" for n
    by(rule current_weight_mono)(auto simp add: ε_nonneg diff_le_self_ennreal)
  note [simp] = currentD_finite[OF ]

  have b_TER: "b ∈ TERΓ' n (g n)" for n
  proof(rule ccontr)
    assume b': "b ∉ TERΓ' n (g n)"
    then have TER: "TERΓ' n (g n) = TER (g n)" using b ε_nonneg[of n]
      by(auto simp add: SAT.simps split: if_split_asm intro: ennreal_diff_le_mono_left)
    from w[of n] TER have "wave Γ (g n)" by(simp add: wave.simps separating_gen.simps)
    moreover have "hindrance Γ (g n)" using hind[of n] TER bnA b'
      by(auto simp add: hindrance.simps split: if_split_asm)
    ultimately show False using loose_unhindered[OF loose] [of n] by(auto intro: hindered.intros)
  qed

  have IN_g_b: "d_IN (g n) b = weight Γ b - ε n" for n using b_TER[of n] bnA
    by(auto simp add: currentD_SAT[OF g])

  def factor  "λn :: nat. (wb - ε 0) / (wb - ε n)"
  have factor_le_1: "factor n ≤ 1" for n using wb_pos δ ε_less[of n]
    by(auto simp add: factor_def field_simps ε_def min_def)
  have factor_pos: "0 < factor n" for n using wb_pos δ * ε_less by(simp add: factor_def field_simps)
  have factor: "(wb - ε n) * factor n = wb - ε 0" for n using ε_less[of n]
    by(simp add: factor_def field_simps)

  def g'  "λn (x, y). if y = b then g n (x, y) * factor n else g n (x, y)"
  have g'_simps: "g' n (x, y) = (if y = b then g n (x, y) * factor n else g n (x, y))" for n x y by(simp add: g'_def)
  have g'_le_g: "g' n e ≤ g n e" for n e using factor_le_1[of n]
    by(cases e "g n e" rule: prod.exhaust[case_product ennreal_cases])
      (auto simp add: g'_simps field_simps mult_left_le)

  have "4 + (n * 6 + n * (n * 2)) ≠ (0 :: real)" for n :: nat
    by(metis (mono_tags, hide_lams) add_is_0 of_nat_eq_0_iff of_nat_numeral zero_neq_numeral)
  then have IN_g': "d_IN (g' n) x = (if x = b then weight Γ b - ε 0 else d_IN (g n) x)" for x n
    using b_TER[of n] bnA factor_pos[of n] factor[of n] wb_pos δ
    by(auto simp add: d_IN_def g'_simps nn_integral_divide nn_integral_cmult currentD_SAT[OF g] wb_conv ε_def field_simps
                      ennreal_minus ennreal_mult'[symmetric] intro!: arg_cong[where f=ennreal])
  have OUT_g': "d_OUT (g' n) x = d_OUT (g n) x - g n (x, b) * (1 - factor n)" for n x
  proof -
    have "d_OUT (g' n) x = (∑+ y. g n (x, y)) - (∑+ y. (g n (x, y) * (1 - factor n)) * indicator {b} y)"
      using factor_le_1[of n] factor_pos[of n]
      apply(cases "g n (x, b)")
      apply(subst nn_integral_diff[symmetric])
      apply(auto simp add: g'_simps nn_integral_divide d_OUT_def AE_count_space mult_left_le ennreal_mult_eq_top_iff
                           ennreal_mult'[symmetric] ennreal_minus_if
                 intro!: nn_integral_cong split: split_indicator)
      apply(simp_all add: field_simps)
      done
    also have "… = d_OUT (g n) x - g n (x, b) * (1 - factor n)" using factor_le_1[of n]
      by(subst nn_integral_indicator_singleton)(simp_all add: d_OUT_def field_simps)
    finally show ?thesis .
  qed

  have g': "current (Γ' 0) (g' n)" for n
  proof
    show "d_OUT (g' n) x ≤ weight (Γ' 0) x" for x
      using b_TER[of n] currentD_weight_OUT[OF g, of n x] ε_le[of 0] factor_le_1[of n]
      by(auto simp add: OUT_g' SINK.simps ennreal_diff_le_mono_left)
    show "d_IN (g' n) x ≤ weight (Γ' 0) x" for x
      using d_IN_mono[of "g' n" x, OF g'_le_g] currentD_weight_IN[OF g, of n x] b_TER[of n] b
      by(auto simp add: IN_g' SAT.simps wb_conv ε_def)
    show "g' n e = 0" if "e ∉ EΓ' 0" for e using that by(cases e)(clarsimp simp add: g'_simps currentD_outside[OF g])
  qed

  have SINK_g': "SINK (g n) = SINK (g' n)" for n using factor_pos[of n]
    by(auto simp add: SINK.simps currentD_OUT_eq_0[OF g] currentD_OUT_eq_0[OF g'] g'_simps split: if_split_asm)
  have SAT_g': "SAT (Γ' n) (g n) = SAT (Γ' 0) (g' n)" for n using b_TER[of n] ε_le'[of 0]
    by(auto simp add: SAT.simps wb_conv IN_g' IN_g_b)
  have TER_g': "TERΓ' n (g n) = TERΓ' 0 (g' n)" for n
    using b_TER[of n] by(auto simp add: SAT.simps SINK_g' OUT_g' IN_g' wb_conv ε_def)

  have w': "wave (Γ' 0) (g' n)" for n
  proof
    have "separating (Γ' 0) (TERΓ' n (g n))" using waveD_separating[OF w, of n]
      by(simp add: separating_gen.simps)
    then show "separating (Γ' 0) (TERΓ' 0 (g' n))" unfolding TER_g' .
  qed(rule g')

  def f  "rec_nat (g 0) (λn rec. rec ⌢Γ' 0 g' (n + 1))"
  have f_simps [simp]:
    "f 0 = g 0"
    "f (Suc n) = f n ⌢Γ' 0 g' (n + 1)"
    for n by(simp_all add: f_def)

  have f: "current (Γ' 0) (f n)" and fw: "wave (Γ' 0) (f n)" for n
  proof(induction n)
    case (Suc n)
    { case 1 show ?case unfolding f_simps using Suc.IH g' by(rule current_plus_web) }
    { case 2 show ?case unfolding f_simps using Suc.IH g' w' by(rule wave_plus') }
  qed(simp_all add: g w)

  have f_inc: "n ≤ m ⟹ f n ≤ f m" for n m
  proof(induction m rule: dec_induct)
    case (step k)
    note step.IH
    also have "f k ≤ (f k ⌢Γ' 0 g' (k + 1))"
      by(rule le_funI plus_web_greater)+
    also have "… = f (Suc k)" by simp
    finally show ?case .
  qed simp
  have chain_f: "Complete_Partial_Order.chain op ≤ (range f)"
    by(rule chain_imageI[where le_a="op ≤"])(simp_all add: f_inc)
  have "countable (support_flow (f n))" for n using current_support_flow[OF f, of n]
    by(rule countable_subset) simp
  hence supp_f: "countable (support_flow (SUP n. f n))" by(subst support_flow_Sup)simp

  have RF_f: "RF (TERΓ' 0 (f n)) = RF (⋃i≤n. TERΓ' 0 (g' i))" for n
  proof(induction n)
    case 0 show ?case by(simp add: TER_g')
  next
    case (Suc n)
    have "RF (TERΓ' 0 (f (Suc n))) = RFΓ' 0 (TERΓ' 0 (f n ⌢Γ' 0 g' (n + 1)))" by simp
    also have "… = RFΓ' 0 (TERΓ' 0 (f n) ∪ TERΓ' 0 (g' (n + 1)))" using f fw g' w'
      by(rule RF_TER_plus_web)
    also have "… = RFΓ' 0 (RFΓ' 0 (TERΓ' 0 (f n)) ∪ TERΓ' 0 (g' (n + 1)))"
      by(simp add: roofed_idem_Un1)
    also have "RFΓ' 0 (TERΓ' 0 (f n)) = RFΓ' 0 (⋃i≤n. TERΓ' 0 (g' i))" by(simp add: Suc.IH)
    also have "RFΓ' 0 (… ∪ TERΓ' 0 (g' (n + 1))) = RFΓ' 0 ((⋃i≤n. TERΓ' 0 (g' i)) ∪ TERΓ' 0 (g' (n + 1)))"
      by(simp add: roofed_idem_Un1)
    also have "(⋃i≤n. TERΓ' 0 (g' i)) ∪ TERΓ' 0 (g' (n + 1)) = (⋃i≤Suc n. TERΓ' 0 (g' i))"
      unfolding atMost_Suc UN_insert by(simp add: Un_commute)
    finally show ?case by simp
  qed

  def   "SUP n. f n"
  have : "current (Γ' 0) gω" unfolding gω_def using chain_f
    by(rule current_Sup)(auto simp add: f supp_f)
  have : "wave (Γ' 0) gω" unfolding gω_def using chain_f
    by(rule wave_lub)(auto simp add: fw  supp_f)
  from  have gω': "current (Γ' n) gω" for n using wb_pos δ
    by(elim current_weight_mono)(auto simp add: ε_le wb_conv ε_def field_simps ennreal_minus_if min_le_iff_disj)

  have SINK_gω: "SINK gω = (⋂n. SINK (f n))" unfolding gω_def
    by(subst SINK_Sup[OF chain_f])(simp_all add: supp_f)
  have SAT_gω: "SAT (Γ' 0) (f n) ⊆ SAT (Γ' 0) gω" for n
    unfolding gω_def by(rule SAT_Sup_upper) simp

  have g_b_out: "g n (b, x) = 0" for n x using b_TER[of n] by(simp add: SINK.simps currentD_OUT_eq_0[OF g])
  have g'_b_out: "g' n (b, x) = 0" for n x by(simp add: g'_simps g_b_out)
  have "f n (b, x) = 0" for n x by(induction n)(simp_all add: g_b_out g'_b_out)
  hence b_SINK_f: "b ∈ SINK (f n)" for n by(simp add: SINK.simps d_OUT_def)
  hence b_SINK_gω: "b ∈ SINK gω" by(simp add: SINK_gω)

  have RF_circ: "RFΓ' n (TERΓ' 0 (g' n)) = RFΓ' 0 (TERΓ' 0 (g' n))" for n by(simp add: roofed_circ_def)
  have edge_restrict_Γ': "edge (quotient_web (Γ' 0) (g' n)) = edge (quotient_web (Γ' n) (g n))" for n
    by(simp add: fun_eq_iff TER_g' RF_circ)
  have restrict_curr_g': "f ↿ Γ' 0 / g' n = f ↿ Γ' n / g n" for n f
    by(simp add: restrict_current_def RF_circ TER_g')

  have RF_restrict: "roofed_gen (quotient_web (Γ' n) (g n)) = roofed_gen (quotient_web (Γ' 0) (g' n))" for n
    by(simp add: roofed_def fun_eq_iff edge_restrict_Γ')

  have gωr': "current (quotient_web (Γ' 0) (g' n)) (gω ↿ Γ' 0 / g' n)" for n using w' 
    by(rule current_restrict_current)
  have gωr: "current (quotient_web (Γ' n) (g n)) (gω ↿ Γ' n / g n)" for n using w gω'
    by(rule current_restrict_current)
  have wωr: "wave (quotient_web (Γ' n) (g n)) (gω ↿ Γ' n / g n)" (is "wave ?Γ' ?gω") for n
  proof
    have *: "wave (quotient_web (Γ' 0) (g' n)) (gω ↿ Γ' 0 / g' n)"
      using g' w'   by(rule wave_restrict_current)
    have "d_IN (gω ↿ Γ' n / g n) b = 0"
      by(rule d_IN_restrict_current_outside roofed_greaterI b_TER)+
    hence SAT_subset: "SAT (quotient_web (Γ' 0) (g' n)) (gω ↿ Γ' n / g n) ⊆ SAT ?Γ' (gω ↿ Γ' n / g n)"
      using b_TER[of n] wb_pos
      by(auto simp add: SAT.simps TER_g' RF_circ wb_conv ε_def field_simps
                        ennreal_minus_if split: if_split_asm)
    hence TER_subset: "TERquotient_web (Γ' 0) (g' n) (gω ↿ Γ' n / g n) ⊆ TER?Γ' (gω ↿ Γ' n / g n)"
      using SINK_g' by(auto simp add: restrict_curr_g')

    show "separating ?Γ' (TER?Γ' ?gω)" (is "separating _ ?TER")
    proof
      fix x y p
      assume xy: "x ∈ A ?Γ'" "y ∈ B ?Γ'" and p: "path ?Γ' x p y"
      from p have p': "path (quotient_web (Γ' 0) (g' n)) x p y" by(simp add: edge_restrict_Γ')
      with waveD_separating[OF *, THEN separatingD, simplified, OF p'] TER_g'[of n] SINK_g' SAT_g' restrict_curr_g' SAT_subset xy
      show "(∃z∈set p. z ∈ ?TER) ∨ x ∈ ?TER" by auto
    qed

    show "d_OUT (gω ↿ Γ' n / g n) x = 0" if "x ∉ RF?Γ' ?TER" for x
      unfolding restrict_curr_g'[symmetric] using TER_subset that
      by(intro waveD_OUT[OF *])(auto simp add: TER_g' restrict_curr_g' RF_restrict intro: in_roofed_mono)
  qed

  have RF_gω: "RF (TERΓ' 0 gω) = RF (⋃n. TERΓ' 0 (g' n))"
  proof -
    have "RFΓ' 0 (TERΓ' 0 gω) = RF (⋃i. TERΓ' 0 (f i))"
      unfolding gω_def by(subst RF_TER_Sup[OF _ _ chain_f])(auto simp add: f fw supp_f)
    also have "… = RF (⋃i. RF (TERΓ' 0 (f i)))" by(simp add: roofed_UN)
    also have "… = RF (⋃i. ⋃j≤i. TERΓ' 0 (g' j))" unfolding RF_f roofed_UN by simp
    also have "(⋃i. ⋃j≤i. TERΓ' 0 (g' j)) = (⋃i. TERΓ' 0 (g' i))" by auto
    finally show ?thesis by simp
  qed

  have SAT_plus_ω: "SAT (Γ' n) (g n ⌢Γ' n gω) = SAT (Γ' 0) (g' n ⌢Γ' 0 gω)" for n
    apply(intro set_eqI)
    apply(simp add: SAT.simps IN_plus_current[OF g w gωr] IN_plus_current[OF g' w' gωr'] TER_g')
    apply(cases "d_IN (gω ↿ Γ' n / g n) b")
    apply(auto simp add: SAT.simps wb_conv d_IN_plus_web IN_g')
    apply(simp_all add: wb_conv IN_g_b restrict_curr_g' ε_def field_simps)
    apply(metis TER_g' b_TER roofed_greaterI)+
    done
  have SINK_plus_ω: "SINK (g n ⌢Γ' n gω) = SINK (g' n ⌢Γ' 0 gω)" for n
    apply(rule set_eqI; simp add: SINK.simps OUT_plus_current[OF g w gωr] OUT_plus_current[OF g' w'] current_restrict_current[OF w' ])
    using factor_pos[of n]
    by(auto simp add: RF_circ TER_g' restrict_curr_g' currentD_OUT_eq_0[OF g] currentD_OUT_eq_0[OF g'] g'_simps split: if_split_asm)
  have TER_plus_ω: "TERΓ' n (g n ⌢Γ' n gω) = TERΓ' 0 (g' n ⌢Γ' 0 gω)" for n
    by(rule set_eqI iffI)+(simp_all add: SAT_plus_ω SINK_plus_ω)

  def h  "λn. g n ⌢Γ' n gω"
  have h: "current (Γ' n) (h n)" for n unfolding h_def using g w
    by(rule current_plus_current)(rule current_restrict_current[OF w gω'])
  have hw: "wave (Γ' n) (h n)" for n unfolding h_def using g w gω' wωr by(rule wave_plus)

  def T  "TERΓ' 0 gω"
  have RF_h: "RF (TERΓ' n (h n)) = RF T" for n
  proof -
    have "RFΓ' 0 (TERΓ' n (h n)) = RFΓ' 0 (RFΓ' 0 (TERΓ' 0 gω) ∪ TERΓ' 0 (g' n))"
      unfolding h_def TER_plus_ω RF_TER_plus_web[OF g' w'  ] roofed_idem_Un1 by(simp add: Un_commute)
    also have "… = RF ((⋃n. TERΓ' 0 (g' n)) ∪ TERΓ' 0 (g' n))"
      by(simp add: RF_gω roofed_idem_Un1)
    also have "… = RFΓ' 0 T" unfolding T_def
      by(auto simp add: RF_gω intro!: arg_cong2[where f=roofed] del: equalityI) auto
    finally show ?thesis by simp
  qed
  have OUT_h_nT: "d_OUT (h n) x = 0" if "x ∉ RF T" for n x
    by(rule waveD_OUT[OF hw])(simp add: RF_h that)
  have IN_h_nT: "d_IN (h n) x = 0" if "x ∉ RF T" for n x
    by(rule wave_not_RF_IN_zero[OF h hw])(simp add: RF_h that)
  have OUT_h_b: "d_OUT (h n) b = 0" for n using b_TER[of n] b_SINK_gω[THEN in_SINK_restrict_current]
    by(auto simp add: h_def OUT_plus_current[OF g w gωr] SINK.simps)
  have OUT_h_ℰ: "d_OUT (h n) x = 0" if "x ∈ ℰ T" for x n using that
    apply(subst (asm) ℰ_RF[symmetric])
    apply(subst (asm) (1 2) RF_h[symmetric, of n])
    apply(subst (asm) ℰ_RF)
    apply(simp add: SINK.simps)
    done
  have IN_h_ℰ: "d_IN (h n) x = weight (Γ' n) x" if "x ∈ ℰ T" "x ∉ A Γ" for x n using that
    apply(subst (asm) ℰ_RF[symmetric])
    apply(subst (asm) (1 2) RF_h[symmetric, of n])
    apply(subst (asm) ℰ_RF)
    apply(simp add: currentD_SAT[OF h])
    done

  have b_SAT: "b ∈ SAT (Γ' 0) (h 0)" using b_TER[of 0]
    by(auto simp add: h_def SAT.simps d_IN_plus_web intro: order_trans)
  have b_T: "b ∈ T" using b_SINK_gω b_TER by(simp add: T_def)(metis SAT_gω subsetD f_simps(1))

  have essential: "b ∈ ℰ T"
  proof(rule ccontr)
    assume "b ∉ ℰ T"
    hence b: "b ∉ ℰ (TERΓ' 0 (h 0))"
    proof(rule contrapos_nn)
      assume "b ∈ ℰ (TERΓ' 0 (h 0))"
      then obtain p y where p: "path Γ b p y" and y: "y ∈ B Γ" and distinct: "distinct (b # p)"
        and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TERΓ' 0 (h 0))" by(rule ℰ_E_RF) auto
      from bypass have bypass': "⋀z. z ∈ set p ⟹ z ∉ T" unfolding RF_h by(auto intro: roofed_greaterI)
      have "essential Γ (B Γ) T b" using p y by(rule essentialI)(auto dest: bypass')
      then show "b ∈ ℰ T" using b_T by simp
    qed

    have h0: "current Γ (h 0)" using h[of 0] by(rule current_weight_mono)(simp_all add: wb_conv ε_nonneg)
    moreover have "wave Γ (h 0)"
    proof
      have "separating (Γ' 0) (ℰΓ' 0 (TERΓ' 0 (h 0)))" by(rule separating_essential)(rule waveD_separating[OF hw])
      then have "separating Γ (ℰ (TERΓ' 0 (h 0)))" by(simp add: separating_gen.simps)
      moreover have subset: "ℰ (TERΓ' 0 (h 0)) ⊆ TER (h 0)" using ε_nonneg[of 0] b
        by(auto simp add: SAT.simps wb_conv split: if_split_asm)
      ultimately show "separating Γ (TER (h 0))" by(rule separating_weakening)
    qed(rule h0)
    ultimately have "h 0 = zero_current" by(rule looseD_wave[OF loose])
    then have "d_IN (h 0) b = 0" by(simp)
    with b_SAT wb ‹b ∉ A Γ› show False by(simp add: SAT.simps wb_conv ε_def ennreal_minus_if split: if_split_asm)
  qed

  def S  "{x ∈ RF (T ∩ B Γ) ∩ A Γ. essential Γ (T ∩ B Γ) (RF (T ∩ B Γ) ∩ A Γ) x}"
  def Γ_h  "⦇ edge = λx y. edge Γ x y ∧ x ∈ S ∧ y ∈ T ∧ y ∈ B Γ, weight = λx. weight Γ x * indicator (S ∪ T ∩ B Γ) x, A = S, B = T ∩ B Γ⦈"
  have Γ_h_sel [simp]:
    "edge Γ_h x y ⟷ edge Γ x y ∧ x ∈ S ∧ y ∈ T ∧ y ∈ B Γ"
    "A Γ_h = S"
    "B Γ_h = T ∩ B Γ"
    "weight Γ_h x = weight Γ x * indicator (S ∪ T ∩ B Γ) x"
    for x y
    by(simp_all add: Γ_h_def)

  have vertex_Γ_hD: "x ∈ S ∪ (T ∩ B Γ)" if "vertex Γ_h x" for x
    using that by(auto simp add: vertex_def)
  have S_vertex: "vertex Γ_h x" if "x ∈ S" for x
  proof -
    from that have a: "x ∈ A Γ" and RF: "x ∈ RF (T ∩ B Γ)" and ess: "essential Γ (T ∩ B Γ) (RF (T ∩ B Γ) ∩ A Γ) x"
      by(simp_all add: S_def)
    from ess obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ" and yT: "y ∈ T"
      and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (T ∩ B Γ) ∩ A Γ" by(rule essentialE_RF)(auto intro: roofed_greaterI)
    from p a y disjoint have "edge Γ x y"
      by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E)
    with that y yT show ?thesis by(auto intro!: vertexI1)
  qed
  have OUT_not_S: "d_OUT (h n) x = 0" if "x ∉ S" for x n
  proof(rule classical)
    assume *: "d_OUT (h n) x ≠ 0"
    consider (A) "x ∈ A Γ" | (B) "x ∈ B Γ" | (outside) "x ∉ A Γ" "x ∉ B Γ" by blast
    then show ?thesis
    proof cases
      case B with currentD_OUT[OF h, of x n] show ?thesis by simp
    next
      case outside with currentD_outside_OUT[OF h, of x n] show ?thesis by(simp add: not_vertex)
    next
      case A
      from * obtain y where xy: "h n (x, y) ≠ 0" using currentD_OUT_eq_0[OF h, of n x] by auto
      then have edge: "edge Γ x y" using currentD_outside[OF h] by(auto)
      hence p: "path Γ x [y] y" by(simp add: rtrancl_path_simps)

      from bipartite_E[OF edge] have x: "x ∈ A Γ" and y: "y ∈ B Γ" by simp_all
      moreover have "x ∈ RF (RF (T ∩ B Γ))"
      proof
        fix p y'
        assume p: "path Γ x p y'" and y': "y' ∈ B Γ"
        from p x y' disjoint have py: "p = [y']"
          by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E)
        have "separating (Γ' 0) (RFΓ' 0 (TERΓ' 0 (h 0)))" unfolding separating_RF
          by(rule waveD_separating[OF hw])
        from separatingD[OF this, of x p y'] py p x y'
        have "x ∈ RF T ∨ y' ∈ RF T" by(auto simp add: RF_h)
        thus "(∃z∈set p. z ∈ RF (T ∩ B Γ)) ∨ x ∈ RF (T ∩ B Γ)"
        proof cases
          case right with y' py show ?thesis by(simp add: RF_in_B)
        next
          case left
          have "x ∉ ℰ T" using OUT_h_ℰ[of x n] xy by(auto simp add: currentD_OUT_eq_0[OF h])
          with left have "x ∈ RF T" by(simp add: roofed_circ_def)
          from RF_circ_edge_forward[OF this, of y'] p py have "y' ∈ RF T" by(simp add: rtrancl_path_simps)
          with y' have "y' ∈ T" by(simp add: RF_in_B)
          with y' show ?thesis using py by(auto intro: roofed_greaterI)
        qed
      qed
      moreover have "y ∈ T" using IN_h_nT[of y n] y xy by(auto simp add: RF_in_B currentD_IN_eq_0[OF h])
      with p x y disjoint have "essential Γ (T ∩ B Γ) (RF (T ∩ B Γ) ∩ A Γ) x" by(auto intro!: essentialI)
      ultimately have "x ∈ S" unfolding roofed_idem by(simp add: S_def)
      with that show ?thesis by contradiction
    qed
  qed

  have B_vertex: "vertex Γ_h y" if T: "y ∈ T" and B: "y ∈ B Γ" and w: "weight Γ y > 0" for y
  proof -
    from T B disjoint ε_less[of 0] w
    have "d_IN (h 0) y > 0" using IN_h_ℰ[of y 0] by(cases "y ∈ A Γ")(auto simp add: essential_BI wb_conv ennreal_minus_if)
    then obtain x where xy: "h 0 (x, y) ≠ 0" using currentD_IN_eq_0[OF h, of 0 y] by(auto)
    then have edge: "edge Γ x y" using currentD_outside[OF h] by(auto)
    from xy have "d_OUT (h 0) x ≠ 0" by(auto simp add: currentD_OUT_eq_0[OF h])
    hence "x ∈ S" using OUT_not_S[of x 0] by(auto)
    with edge T B show ?thesis by(simp add: vertexI2)
  qed

  have Γ_h: "countable_bipartite_web Γ_h"
  proof
    show "VΓ_h ⊆ A Γ_h ∪ B Γ_h" by(auto simp add: vertex_def)
    show "A Γ_h ⊆ VΓ_h" using S_vertex by auto
    show "x ∈ A Γ_h ∧ y ∈ B Γ_h" if "edge Γ_h x y" for x y using that by auto
    show "A Γ_h ∩ B Γ_h = {}" using disjoint by(auto simp add: S_def)
    have "EΓ_hE" by auto
    thus "countable EΓ_h" by(rule countable_subset) simp
    show "weight Γ_h x ≠ ⊤" for x by(simp split: split_indicator)
    show "weight Γ_h x = 0" if "x ∉ VΓ_h" for x
      using that S_vertex B_vertex[of x]
      by(cases "weight Γ_h x > 0")(auto split: split_indicator)
  qed
  then interpret Γ_h: countable_bipartite_web Γ_h .

  have essential_T: "essential Γ (B Γ) T = essential Γ (B Γ) (TERΓ' 0 (h 0))"
  proof(rule ext iffI)+
    fix x
    assume "essential Γ (B Γ) T x"
    then obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ" and distinct: "distinct (x # p)"
      and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF T" by(rule essentialE_RF)auto
    from bypass have bypass': "⋀z. z ∈ set p ⟹ z ∉ TERΓ' 0 (h 0)"
      unfolding RF_h[of 0, symmetric] by(blast intro: roofed_greaterI)
    show "essential Γ (B Γ) (TERΓ' 0 (h 0)) x" using p y
      by(blast intro: essentialI dest: bypass')
  next
    fix x
    assume "essential Γ (B Γ) (TERΓ' 0 (h 0)) x"
    then obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ" and distinct: "distinct (x # p)"
      and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TERΓ' 0 (h 0))" by(rule essentialE_RF)auto
    from bypass have bypass': "⋀z. z ∈ set p ⟹ z ∉ T"
      unfolding RF_h[of 0] by(blast intro: roofed_greaterI)
    show "essential Γ (B Γ) T x" using p y
      by(blast intro: essentialI dest: bypass')
  qed

  have h': "current Γ_h (h n)" for n
  proof
    show "d_OUT (h n) x ≤ weight Γ_h x" for x
      using currentD_weight_OUT[OF h, of n x] ε_nonneg[of n] Γ'.currentD_OUT'[OF h, of x n] OUT_not_S
      by(auto split: split_indicator if_split_asm elim: order_trans intro: diff_le_self_ennreal in_roofed_mono simp add: OUT_h_b  roofed_circ_def)

    show "d_IN (h n) x ≤ weight Γ_h x" for x
      using currentD_weight_IN[OF h, of n x] currentD_IN[OF h, of x n] ε_nonneg[of n] b_T b Γ'.currentD_IN'[OF h, of x n] IN_h_nT[of x n]
      by(cases "x ∈ B Γ")(auto 4 3 split: split_indicator split: if_split_asm elim: order_trans intro: diff_le_self_ennreal simp add: S_def  roofed_circ_def RF_in_B)

    show "h n e = 0" if "e ∉ EΓ_h" for e
      using that OUT_not_S[of "fst e" n] currentD_outside'[OF h, of e n] Γ'.currentD_IN'[OF h, of "snd e" n] disjoint
      apply(cases "e ∈ E")
      apply(auto split: prod.split_asm simp add: currentD_OUT_eq_0[OF h] currentD_IN_eq_0[OF h])
      apply(cases "fst e ∈ S"; clarsimp simp add: S_def)
      apply(frule RF_circ_edge_forward[rotated])
      apply(erule roofed_circI, blast)
      apply(drule bipartite_E)
      apply(simp add: RF_in_B)
      done
  qed

  have SAT_h': "B Γ_h ∩ VΓ_h - {b} ⊆ SAT Γ_h (h n)" for n
  proof
    fix x
    assume "x ∈ B Γ_h ∩ VΓ_h - {b}"
    then have x: "x ∈ T" and B: "x ∈ B Γ" and b: "x ≠ b" and vertex: "x ∈ VΓ_h" by auto
    from B disjoint have xnA: "x ∉ A Γ" by blast
    from x B have "x ∈ ℰ T" by(simp add: essential_BI)
    hence "d_IN (h n) x = weight (Γ' n) x" using xnA by(rule IN_h_ℰ)
    with xnA b x B show "x ∈ SAT Γ_h (h n)" by(simp add: currentD_SAT[OF h'])
  qed
  moreover have "b ∈ B Γ_h" using b essential by simp
  moreover have "(λn. min δ wb * (1 / (real (n + 2)))) ⇢ 0"
    apply(rule LIMSEQ_ignore_initial_segment)
    apply(rule tendsto_mult_right_zero)
    apply(rule lim_1_over_real_power[where s=1, simplified])
    done
  then have "(INF n. ennreal (ε n)) = 0" using wb_pos δ
    apply(simp add: ε_def)
    apply(rule INF_Lim_ereal)
    apply(rule decseq_SucI)
    apply(simp add: field_simps min_def)
    apply(simp add: add.commute ennreal_0[symmetric] del: ennreal_0)
    done
  then have "(SUP n. d_IN (h n) b) = weight Γ_h b" using essential b bnA wb IN_h_ℰ[of b]
    by(simp add: SUP_const_minus_ennreal)
  moreover have "d_IN (h 0) b ≤ d_IN (h n) b" for n using essential b bnA wb_pos δ IN_h_ℰ[of b]
    by(simp add: wb_conv ε_def field_simps ennreal_minus_if min_le_iff_disj)
  moreover have b_V: "b ∈ VΓ_h" using b wb essential by(auto dest: B_vertex)
  ultimately have "∃h'. current Γ_h h' ∧ wave Γ_h h' ∧ B Γ_h ∩ VΓ_h ⊆ SAT Γ_h h'"
    by(rule Γ_h.unhinder_bipartite[OF h'])
  then obtain h' where h': "current Γ_h h'" and h'w: "wave Γ_h h'"
    and B_SAT': "B Γ_h ∩ VΓ_h ⊆ SAT Γ_h h'" by blast

  have h'': "current Γ h'"
  proof
    show "d_OUT h' x ≤ weight Γ x" for x using currentD_weight_OUT[OF h', of x]
      by(auto split: split_indicator_asm elim: order_trans intro: )
    show "d_IN h' x ≤ weight Γ x" for x using currentD_weight_IN[OF h', of x]
      by(auto split: split_indicator_asm elim: order_trans intro: )
    show "h' e = 0" if "e ∉ E" for e using currentD_outside'[OF h', of e] that by auto
  qed
  moreover have "wave Γ h'"
  proof
    have "separating (Γ' 0) T" unfolding T_def by(rule waveD_separating[OF ])
    hence "separating Γ T" by(simp add: separating_gen.simps)
    hence *: "separating Γ (ℰ T)" by(rule separating_essential)
    show "separating Γ (TER h')"
    proof
      fix x p y
      assume x: "x ∈ A Γ" and p: "path Γ x p y" and y: "y ∈ B Γ"
      from p x y disjoint have py: "p = [y]"
        by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E)
      from separatingD[OF * p x y] py have "x ∈ ℰ T ∨ y ∈ ℰ T" by auto
      then show "(∃z∈set p. z ∈ TER h') ∨ x ∈ TER h'"
      proof cases
        case left
        then have "x ∉ VΓ_h" using x disjoint
          by(auto 4 4 dest!: vertex_Γ_hD simp add: S_def elim: essentialE_RF intro!: roofed_greaterI dest: roofedD)
        hence "d_OUT h' x = 0" by(intro currentD_outside_OUT[OF h'])
        with x have "x ∈ TER h'" by(auto simp add: SAT.A SINK.simps)
        thus ?thesis ..
      next
        case right
        have "y ∈ SAT Γ h'"
        proof(cases "weight Γ y > 0")
          case True
          with py x y right have "vertex Γ_h y" by(auto intro: B_vertex)
          hence "y ∈ SAT Γ_h h'" using B_SAT' right y by auto
          with right y disjoint show ?thesis
            by(auto simp add: currentD_SAT[OF h'] currentD_SAT[OF h''] S_def)
        qed(auto simp add: SAT.simps)
        with currentD_OUT[OF h', of y] y right have "y ∈ TER h'" by(auto simp add: SINK)
        thus ?thesis using py by simp
      qed
    qed
  qed(rule h'')
  ultimately have "h' = zero_current" by(rule looseD_wave[OF loose])
  hence "d_IN h' b = 0" by simp
  moreover from essential b b_V B_SAT' have "b ∈ SAT Γ_h h'" by(auto)
  ultimately show False using wb b essential disjoint by(auto simp add: SAT.simps S_def)
qed

end

subsection ‹Single-vertex saturation in unhindered bipartite webs›

text ‹
  The proof of lemma 6.10 in @{cite "AharoniBergerGeorgakopoulusPerlsteinSpruessel2011JCT"} is flawed.
  The transfinite steps (taking the least upper bound) only preserves unhinderedness, but not looseness.
  However, the single steps to non-limit ordinals assumes that @{text "Ω - fi"} is loose in order to
  apply Lemma 6.9.

  Counterexample: The bipartite web with three nodes @{text a1}, @{text a2}, @{text a3} in @{text A}
  and two nodes @{text b1}, @{text b2} in @{text B} and edges @{text "(a1, b1)"}, @{text "(a2, b1)"},
  @{text "(a2, b2)"}, @{text "(a3, b2)"} and weights @{text "a1 = a3 = 1"} and @{text "a2 = 2"} and
  @{text "b1 = 3"} and @{text "b2 = 2"}.
  Then, we can get a sequence of weight reductions on @{text b2} from @{text 2} to @{text "1.5"},
  @{text "1.25"}, @{text "1.125"}, etc. with limit @{text 1}.
  All maximal waves in the restricted webs in the sequence are @{term [source] zero_current}, so in
  the limit, we get @{text "k = 0"} and @{text "ε = 1"} for @{text "a2"} and @{text "b2"}. Now, the
  restricted web for the two is not loose because it contains the wave which assigns 1 to @{text "(a3, b2)"}.

  We prove a stronger version which only assumes and ensures on unhinderedness.
›
context countable_bipartite_web begin

lemma web_flow_iff: "web_flow Γ f ⟷ current Γ f"
using bipartite_V by(auto simp add: web_flow.simps)

lemma countable_bipartite_web_minus_web:
  assumes f: "current Γ f"
  shows "countable_bipartite_web (Γ ⊖ f)"
using bipartite_V A_vertex bipartite_E disjoint currentD_finite_OUT[OF f] currentD_weight_OUT[OF f] currentD_weight_IN[OF f] currentD_outside_OUT[OF f] currentD_outside_IN[OF f]
by unfold_locales (auto simp add:  weight_outside)

lemma current_plus_current_minus:
  assumes f: "current Γ f"
  and g: "current (Γ ⊖ f) g"
  shows "current Γ (plus_current f g)" (is "current _ ?fg")
proof
  interpret Γ: countable_bipartite_web "Γ ⊖ f" using f by(rule countable_bipartite_web_minus_web)
  show "d_OUT ?fg x ≤ weight Γ x" for x
    using currentD_weight_OUT[OF g, of x] currentD_OUT[OF g, of x] currentD_finite_OUT[OF f, of x] currentD_OUT[OF f, of x] currentD_outside_IN[OF f, of x] currentD_outside_OUT[OF f, of x] currentD_weight_OUT[OF f, of x]
    by(cases "x ∈ A Γ ∨ x ∈ B Γ")(auto simp add: add.commute d_OUT_def nn_integral_add not_vertex ennreal_le_minus_iff split: if_split_asm)
  show "d_IN ?fg x ≤ weight Γ x" for x
    using currentD_weight_IN[OF g, of x] currentD_IN[OF g, of x] currentD_finite_IN[OF f, of x] currentD_OUT[OF f, of x] currentD_outside_IN[OF f, of x] currentD_outside_OUT[OF f, of x] currentD_weight_IN[OF f, of x]
    by(cases "x ∈ A Γ ∨ x ∈ B Γ")(auto simp add: add.commute  d_IN_def nn_integral_add not_vertex ennreal_le_minus_iff split: if_split_asm)
  show "?fg e = 0" if "e ∉ E" for e using that currentD_outside'[OF f, of e] currentD_outside'[OF g, of e] by(cases e) simp
qed

lemma wave_plus_current_minus:
  assumes f: "current Γ f"
  and w: "wave Γ f"
  and g: "current (Γ ⊖ f) g"
  and w': "wave (Γ ⊖ f) g"
  shows "wave Γ (plus_current f g)" (is "wave _ ?fg")
proof
  show fg: "current Γ ?fg" using f g by(rule current_plus_current_minus)
  show sep: "separating Γ (TER ?fg)"
  proof
    fix x p y
    assume x: "x ∈ A Γ" and p: "path Γ x p y" and y: "y ∈ B Γ"
    from p x y disjoint have py: "p = [y]"
      by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E)
    with waveD_separating[THEN separatingD, OF w p x y] have "x ∈ TER f ∨ y ∈ TER f" by auto
    thus "(∃z∈set p. z ∈ TER ?fg) ∨ x ∈ TER ?fg"
    proof cases
      case right
      with y disjoint have "y ∈ TER ?fg" using currentD_OUT[OF fg y]
        by(auto simp add: SAT.simps SINK.simps d_IN_def nn_integral_add not_le add_increasing2)
      thus ?thesis using py by simp
    next
      case x': left
      from p have "path (Γ ⊖ f) x p y" by simp
      from waveD_separating[THEN separatingD, OF w' this] x y py
      have "x ∈ TERΓ ⊖ f g ∨ y ∈ TERΓ ⊖ f g" by auto
      thus ?thesis
      proof cases
        case left
        hence "x ∈ TER ?fg" using x x'
          by(auto simp add: SAT.simps SINK.simps d_OUT_def nn_integral_add)
        thus ?thesis ..
      next
        case right
        hence "y ∈ TER ?fg" using disjoint y currentD_OUT[OF fg y] currentD_OUT[OF f y] currentD_finite_IN[OF f, of y]
          by(auto simp add: add.commute SINK.simps SAT.simps d_IN_def nn_integral_add ennreal_minus_le_iff split: if_split_asm)
        with py show ?thesis by auto
      qed
    qed
  qed
qed

lemma minus_plus_current:
  assumes f: "current Γ f"
  and g: "current (Γ ⊖ f) g"
  shows "Γ ⊖ plus_current f g = Γ ⊖ f ⊖ g" (is "?lhs = ?rhs")
proof(rule web.equality)
  have "weight ?lhs x = weight ?rhs x" for x
    using currentD_weight_IN[OF f, of x] currentD_weight_IN[OF g, of x]
    by (auto simp add: d_IN_def d_OUT_def nn_integral_add diff_add_eq_diff_diff_swap_ennreal add_increasing2 diff_add_assoc2_ennreal add.assoc)
  thus "weight ?lhs = weight ?rhs" ..
qed simp_all

lemma unhindered_minus_web:
  assumes unhindered: "¬ hindered Γ"
  and f: "current Γ f"
  and w: "wave Γ f"
  shows "¬ hindered (Γ ⊖ f)"
proof
  assume "hindered (Γ ⊖ f)"
  then obtain g where g: "current (Γ ⊖ f) g"
    and w': "wave (Γ ⊖ f) g"
    and hind: "hindrance (Γ ⊖ f) g" by cases
  let ?fg = "plus_current f g"
  have fg: "current Γ ?fg" using f g by(rule current_plus_current_minus)
  moreover have "wave Γ ?fg" using f w g w' by(rule wave_plus_current_minus)
  moreover from hind obtain a where a: "a ∈ A Γ" and nℰ: "a ∉ ℰΓ ⊖ f (TERΓ ⊖ f g)"
    and wa: "d_OUT g a < weight (Γ ⊖ f) a" by cases auto
  from a have "hindrance Γ ?fg"
  proof
    show "a ∉ ℰ (TER ?fg)"
    proof
      assume : "a ∈ ℰ (TER ?fg)"
      then obtain p y where p: "path Γ a p y" and y: "y ∈ B Γ"
        and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TER ?fg)" by(rule ℰ_E_RF) blast
      from p a y disjoint have py: "p = [y]"
        by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E)
      from bypass[of y] py have "y ∉ TER ?fg" by(auto intro: roofed_greaterI)
      with currentD_OUT[OF fg y] have "y ∉ SAT Γ ?fg" by(auto simp add: SINK.simps)
      hence "y ∉ SAT (Γ ⊖ f) g" using y currentD_OUT[OF f y] currentD_finite_IN[OF f, of y]
        by(auto simp add: SAT.simps d_IN_def nn_integral_add ennreal_minus_le_iff add.commute)
      hence "essential (Γ ⊖ f) (B (Γ ⊖ f)) (TERΓ ⊖ f g) a" using p py y
        by(auto intro!: essentialI)
      moreover from  a have "a ∈ TERΓ ⊖ f g"
        by(auto simp add: SAT.A SINK_plus_current)
      ultimately have "a ∈ ℰΓ ⊖ f (TERΓ ⊖ f g)" by blast
      thus False using nℰ by contradiction
    qed
    show "d_OUT ?fg a < weight Γ a" using a wa currentD_finite_OUT[OF f, of a]
      by(simp add: d_OUT_def less_diff_eq_ennreal less_top add.commute nn_integral_add)
  qed
  ultimately have "hindered Γ" by(blast intro: hindered.intros)
  with unhindered show False by contradiction
qed

lemma loose_minus_web:
  assumes unhindered: "¬ hindered Γ"
  and f: "current Γ f"
  and w: "wave Γ f"
  and maximal: "⋀w. ⟦ current Γ w; wave Γ w; f ≤ w ⟧ ⟹ f = w"
  shows "loose (Γ ⊖ f)" (is "loose ?Γ")
proof
  fix g
  assume g: "current ?Γ g" and w': "wave ?Γ g"
  let ?g = "plus_current f g"
  from f g have "current Γ ?g" by(rule current_plus_current_minus)
  moreover from f w g w' have "wave Γ ?g" by(rule wave_plus_current_minus)
  moreover have "f ≤ ?g" by(clarsimp simp add: le_fun_def)
  ultimately have eq: "f = ?g" by(rule maximal)
  have "g e = 0" for e
  proof(cases e)
    case (Pair x y)
    have "f e ≤ d_OUT f x" unfolding d_OUT_def Pair by(rule nn_integral_ge_point) simp
    also have "… ≤ weight Γ x" by(rule currentD_weight_OUT[OF f])
    also have "… < ⊤" by(simp add: less_top[symmetric])
    finally show "g e = 0" using Pair eq[THEN fun_cong, of e]
      by(cases "f e" "g e" rule: ennreal2_cases)(simp_all add: fun_eq_iff)
  qed
  thus "g = (λ_. 0)" by(simp add: fun_eq_iff)
next
  show "¬ hindrance ?Γ zero_current" using unhindered
  proof(rule contrapos_nn)
    assume "hindrance ?Γ zero_current"
    then obtain x where a: "x ∈ A ?Γ" and : "x ∉ ℰ (TER zero_current)"
      and weight: "d_OUT zero_current x < weight ?Γ x" by cases
    have "hindrance Γ f"
    proof
      show a': "x ∈ A Γ" using a by simp
      with weight show "d_OUT f x < weight Γ x"
        by(simp add: less_diff_eq_ennreal less_top[symmetric] currentD_finite_OUT[OF f])
      show "x ∉ ℰ (TER f)"
      proof
        assume "x ∈ ℰ (TER f)"
        then obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ"
          and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TER f)" by(rule ℰ_E_RF) auto
        from p a' y disjoint have py: "p = [y]"
          by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E)
        hence "y ∉ (TER f)" using bypass[of y] by(auto intro: roofed_greaterI)
        hence "weight ?Γ y > 0" using currentD_OUT[OF f y] disjoint y
          by(auto simp add: SINK.simps SAT.simps diff_gr0_ennreal)
        hence "y ∉ TER zero_current" using y disjoint by(auto)
        hence "essential ?Γ (B ?Γ) (TER zero_current) x" using p y py by(auto intro!: essentialI)
        with a have "x ∈ ℰ (TER zero_current)" by simp
        with  show False by contradiction
      qed
    qed
    thus "hindered Γ" using f w ..
  qed
qed

lemma weight_minus_web:
  assumes f: "current Γ f"
  shows "weight (Γ ⊖ f) x = (if x ∈ A Γ then weight Γ x - d_OUT f x else weight Γ x - d_IN f x)"
proof(cases "x ∈ B Γ")
  case True
  with currentD_OUT[OF f True] disjoint show ?thesis by auto
next
  case False
  hence "d_IN f x = 0" "d_OUT f x = 0" if "x ∉ A Γ"
    using currentD_outside_OUT[OF f, of x] currentD_outside_IN[OF f, of x] bipartite_V that by auto
  then show ?thesis by simp
qed


lemma (in -) separating_minus_web [simp]: "separating_gen (G ⊖ f) = separating_gen G"
by(simp add: separating_gen.simps fun_eq_iff)

lemma current_minus:
  assumes f: "current Γ f"
  and g: "current Γ g"
  and le: "⋀e. g e ≤ f e"
  shows "current (Γ ⊖ g) (f - g)"
proof -
  interpret Γ: countable_bipartite_web "Γ ⊖ g" using g by(rule countable_bipartite_web_minus_web)
  note [simp del] = minus_web_sel(2)
    and [simp] = weight_minus_web[OF g]
  show ?thesis
  proof
    show "d_OUT (f - g) x ≤ weight (Γ ⊖ g) x" for x unfolding fun_diff_def
      using currentD_weight_OUT[OF f, of x] currentD_weight_IN[OF g, of x]
      by(subst d_OUT_diff)(simp_all add: le currentD_finite_OUT[OF g] currentD_OUT'[OF f] currentD_OUT'[OF g] ennreal_minus_mono)
    show "d_IN (f - g) x ≤ weight (Γ ⊖ g) x" for x unfolding fun_diff_def
      using currentD_weight_IN[OF f, of x] currentD_weight_OUT[OF g, of x]
      by(subst d_IN_diff)(simp_all add: le currentD_finite_IN[OF g] currentD_IN[OF f] currentD_IN[OF g] ennreal_minus_mono)
    show "(f - g) e = 0" if "e ∉ EΓ ⊖ g" for e using that currentD_outside'[OF f] currentD_outside'[OF g] by simp
  qed
qed

lemma
  assumes w: "wave Γ f"
  and g: "current Γ g"
  and le: "⋀e. g e ≤ f e"
  shows wave_minus: "wave (Γ ⊖ g) (f - g)"
  and TER_minus: "TER f ⊆ TERΓ ⊖ g (f - g)"
proof
  have "x ∈ SINK f ⟹ x ∈ SINK (f - g)" for x using d_OUT_mono[of g _ f, OF le, of x]
    by(auto simp add: SINK.simps fun_diff_def d_OUT_diff le currentD_finite_OUT[OF g])
  moreover have "x ∈ SAT Γ f ⟹ x ∈ SAT (Γ ⊖ g) (f - g)" for x
    by(auto simp add: SAT.simps currentD_OUT'[OF g] fun_diff_def d_IN_diff le currentD_finite_IN[OF g] ennreal_minus_mono)
  ultimately show TER: "TER f ⊆ TERΓ ⊖ g (f - g)" by(auto)

  from w have "separating Γ (TER f)" by(rule waveD_separating)
  thus "separating (Γ ⊖ g) (TERΓ ⊖ g (f - g))" using TER by(simp add: separating_weakening)

  fix x
  assume "x ∉ RFΓ ⊖ g (TERΓ ⊖ g (f - g))"
  hence "x ∉ RF (TER f)" using TER by(auto intro: in_roofed_mono)
  hence "d_OUT f x = 0" by(rule waveD_OUT[OF w])
  moreover have "0 ≤ f e" for e using le[of e] by simp
  ultimately show "d_OUT (f - g) x = 0" unfolding d_OUT_def
    by(simp add: nn_integral_0_iff emeasure_count_space_eq_0)
qed

lemma (in -) essential_minus_web [simp]: "essential (Γ ⊖ f) = essential Γ"
by(simp add: essential_def fun_eq_iff)

lemma (in -) RF_in_essential: fixes B shows "essential Γ B S x ⟹ x ∈ roofed_gen Γ B S ⟷ x ∈ S"
by(auto intro: roofed_greaterI elim!: essentialE_RF dest: roofedD)

lemma (in -) d_OUT_fun_upd:
  assumes "f (x, y) ≠ ⊤" "f (x, y) ≥ 0" "k ≠ ⊤" "k ≥ 0"
  shows "d_OUT (f((x, y) := k)) x' = (if x = x' then d_OUT f x - f (x, y) + k else d_OUT f x')"
  (is "?lhs = ?rhs")
proof(cases "x = x'")
  case True
  have "?lhs = (∑+ y'. f (x, y') + k * indicator {y} y') - (∑+ y'. f (x, y') * indicator {y} y')"
    unfolding d_OUT_def using assms True
    by(subst nn_integral_diff[symmetric])
      (auto intro!: nn_integral_cong simp add: AE_count_space split: split_indicator)
  also have "(∑+ y'. f (x, y') + k * indicator {y} y') = d_OUT f x + (∑+ y'. k * indicator {y} y')"
    unfolding d_OUT_def using assms
    by(subst nn_integral_add[symmetric])
      (auto intro!: nn_integral_cong split: split_indicator)
  also have "… - (∑+ y'. f (x, y') * indicator {y} y') = ?rhs" using True assms
    by(subst diff_add_assoc2_ennreal[symmetric])(auto simp add: d_OUT_def intro!: nn_integral_ge_point)
  finally show ?thesis .
qed(simp add: d_OUT_def)

lemma unhindered_saturate1: -- "Lemma 6.10"
  assumes unhindered: "¬ hindered Γ"
  and a: "a ∈ A Γ"
  shows "∃f. current Γ f ∧ d_OUT f a = weight Γ a ∧ ¬ hindered (Γ ⊖ f)"
proof -
  from a A_vertex have a_vertex: "vertex Γ a" by auto
  from unhindered have "¬ hindrance Γ zero_current" by(auto intro!: hindered.intros simp add: )
  then have a_ℰ: "a ∈ ℰ (A Γ)" if "weight Γ a > 0"
  proof(rule contrapos_np)
    assume "a ∉ ℰ (A Γ)"
    with a have "¬ essential Γ (B Γ) (A Γ) a" by simp
    hence "¬ essential Γ (B Γ) (A Γ ∪ {x. weight Γ x ≤ 0}) a"
      by(rule contrapos_nn)(erule essential_mono; simp)
    with a that show "hindrance Γ zero_current" by(auto intro: hindrance)
  qed

  def F  "λ(ε, h :: 'v current). plus_current ε h"
  have F_simps: "F (ε, h) = plus_current ε h" for ε h by(simp add: F_def)
  def Fld  "{(ε, h).
     current Γ ε ∧ (∀x. x ≠ a ⟶ d_OUT ε x = 0) ∧
     current (Γ ⊖ ε) h ∧ wave (Γ ⊖ ε) h ∧
     ¬ hindered (Γ ⊖ F (ε, h))}"
  def leq  "restrict_rel Fld {(f, f'). f ≤ f'}"
  have Fld: "Field leq = Fld" by(auto simp add: leq_def)
  have F_I [intro?]: "(ε, h) ∈ Field leq"
    if "current Γ ε" and "⋀x. x ≠ a ⟹ d_OUT ε x = 0"
    and "current (Γ ⊖ ε) h" and "wave (Γ ⊖ ε) h"
    and "¬ hindered (Γ ⊖ F (ε, h))"
    for ε h using that by(simp add: Fld Fld_def)
  have ε_curr: "current Γ ε" if "(ε, h) ∈ Field leq" for ε h using that by(simp add: Fld Fld_def)
  have OUT_ε: "⋀x. x ≠ a ⟹ d_OUT ε x = 0" if "(ε, h) ∈ Field leq" for ε h using that by(simp add: Fld Fld_def)
  have h: "current (Γ ⊖ ε) h" if "(ε, h) ∈ Field leq" for ε h using that by(simp add: Fld Fld_def)
  have h_w: "wave (Γ ⊖ ε) h" if "(ε, h) ∈ Field leq" for ε h using that by(simp add: Fld Fld_def)
  have unhindered': "¬ hindered (Γ ⊖ F εh)" if "εh ∈ Field leq" for εh using that by(simp add: Fld Fld_def split: prod.split_asm)
  have f: "current Γ (F εh)" if "εh ∈ Field leq" for εh using ε_curr h that
    by(cases εh)(simp add: F_simps current_plus_current_minus)
  have out_ε: "ε (x, y) = 0" if "(ε, h) ∈ Field leq" "x ≠ a" for ε h x y
  proof(rule antisym)
    have "ε (x, y) ≤ d_OUT ε x" unfolding d_OUT_def by(rule nn_integral_ge_point) simp
    with OUT_ε[OF that] show "ε (x, y) ≤ 0" by simp
  qed simp
  have IN_ε: "d_IN ε x = ε (a, x)" if "(ε, h) ∈ Field leq" for ε h x
  proof(rule trans)
    show "d_IN ε x = (∑+ y. ε (y, x) * indicator {a} y)" unfolding d_IN_def
      by(rule nn_integral_cong)(simp add: out_ε[OF that] split: split_indicator)
  qed(simp add: max_def ε_curr[OF that])
  have leqI: "((ε, h), (ε', h')) ∈ leq" if "ε ≤ ε'" "h ≤ h'" "(ε, h) ∈ Field leq" "(ε', h') ∈ Field leq" for ε h ε' h'
    using that unfolding Fld by(simp add: leq_def in_restrict_rel_iff)

  have chain_Field: "Sup M ∈ Field leq" if M: "M ∈ Chains leq" and nempty: "M ≠ {}" for M
    unfolding Sup_prod_def
  proof
    from nempty obtain ε h where in_M: "(ε, h) ∈ M" by auto
    with M have Field: "(ε, h) ∈ Field leq" by(rule Chains_FieldD)

    from M have chain: "Complete_Partial_Order.chain (λε ε'. (ε, ε') ∈ leq) M"
      by(intro Chains_into_chain) simp
    hence chain': "Complete_Partial_Order.chain op ≤ M"
      by(auto simp add: chain_def leq_def in_restrict_rel_iff)
    hence chain1: "Complete_Partial_Order.chain op ≤ (fst ` M)"
      and chain2: "Complete_Partial_Order.chain op ≤ (snd ` M)"
      by(rule chain_imageI; auto)+

    have outside1: "Sup (fst ` M) (x, y) = 0" if "¬ edge Γ x y" for x y using that
      by(auto intro!: SUP_eq_const simp add: nempty dest!: Chains_FieldD[OF M] ε_curr currentD_outside)
    then have "support_flow (Sup (fst ` M)) ⊆ E" by(auto elim!: support_flow.cases intro: ccontr)
    hence supp_flow1: "countable (support_flow (Sup (fst ` M)))" by(rule countable_subset) simp

    show SM1: "current Γ (Sup (fst ` M))"
      by(rule current_Sup[OF chain1 _ _ supp_flow1])(auto dest: Chains_FieldD[OF M, THEN ε_curr] simp add: nempty)
    show OUT1_na: "d_OUT (Sup (fst ` M)) x = 0" if "x ≠ a" for x using that
      by(subst d_OUT_Sup[OF chain1 _ supp_flow1])(auto simp add: nempty intro!: SUP_eq_const dest: Chains_FieldD[OF M, THEN OUT_ε])

    interpret SM1: countable_bipartite_web "Γ ⊖ Sup (fst ` M)"
      using SM1 by(rule countable_bipartite_web_minus_web)

    let ?h = "Sup (snd ` M)"
    have outside2: "?h (x, y) = 0" if "¬ edge Γ x y" for x y using that
      by(auto intro!: SUP_eq_const simp add: nempty dest!: Chains_FieldD[OF M] h currentD_outside)
    then have "support_flow ?h ⊆ E" by(auto elim!: support_flow.cases intro: ccontr)
    hence supp_flow2: "countable (support_flow ?h)" by(rule countable_subset) simp

    have OUT1: "d_OUT (Sup (fst ` M)) x = (SUP (ε, h):M. d_OUT ε x)" for x
      by(subst d_OUT_Sup[OF chain1 _ supp_flow1])(simp_all add: nempty split_beta)
    have OUT1': "d_OUT (Sup (fst ` M)) x = (if x = a then SUP (ε, h):M. d_OUT ε a else 0)" for x
      unfolding OUT1 by(auto intro!: SUP_eq_const simp add: nempty OUT_ε dest!: Chains_FieldD[OF M])
    have OUT1_le: "(⨆εh∈M. d_OUT (fst εh) x) ≤ weight Γ x" for x
      using currentD_weight_OUT[OF SM1, of x] OUT1[of x] by(simp add: split_beta)
    have OUT1_nonneg: "0 ≤ (⨆εh∈M. d_OUT (fst εh) x)" for x using in_M by(rule SUP_upper2)(simp add: )
    have IN1: "d_IN (Sup (fst ` M)) x = (SUP (ε, h):M. d_IN ε x)" for x
      by(subst d_IN_Sup[OF chain1 _ supp_flow1])(simp_all add: nempty split_beta)
    have IN1_le: "(⨆εh∈M. d_IN (fst εh) x) ≤ weight Γ x" for x
      using currentD_weight_IN[OF SM1, of x] IN1[of x] by(simp add: split_beta)
    have IN1_nonneg: "0 ≤ (⨆εh∈M. d_IN (fst εh) x)" for x using in_M by(rule SUP_upper2) simp
    have IN1': "d_IN (Sup (fst ` M)) x = (SUP (ε, h):M. ε (a, x))" for x
      unfolding IN1 by(rule SUP_cong[OF refl])(auto dest!: Chains_FieldD[OF M] IN_ε)

    have directed: "∃εk''∈M. F (snd εk) + F (fst εk') ≤ F (snd εk'') + F (fst εk'')"
      if mono: "⋀f g. (⋀z. f z ≤ g z) ⟹ F f ≤ F g" "εk ∈ M" "εk' ∈ M"
      for εk εk' and F :: "_ ⇒ ennreal"
      using chainD[OF chain that(2-3)]
    proof cases
      case left
      hence "snd εk ≤ snd εk'" by(simp add: leq_def less_eq_prod_def in_restrict_rel_iff)
      hence "F (snd εk) + F (fst εk') ≤ F (snd εk') + F (fst εk')"
        by(intro add_right_mono mono)(clarsimp simp add: le_fun_def)
      with that show ?thesis by blast
    next
      case right
      hence "fst εk' ≤ fst εk" by(simp add: leq_def less_eq_prod_def in_restrict_rel_iff)
      hence "F (snd εk) + F (fst εk') ≤ F (snd εk) + F (fst εk)"
        by(intro add_left_mono mono)(clarsimp simp add: le_fun_def)
      with that show ?thesis by blast
    qed
    have directed_OUT: "∃εk''∈M. d_OUT (snd εk) x + d_OUT (fst εk') x ≤ d_OUT (snd εk'') x + d_OUT (fst εk'') x"
      if "εk ∈ M" "εk' ∈ M" for x εk εk' by(rule directed; rule d_OUT_mono that)
    have directed_IN: "∃εk''∈M. d_IN (snd εk) x + d_IN (fst εk') x ≤ d_IN (snd εk'') x + d_IN (fst εk'') x"
      if "εk ∈ M" "εk' ∈ M" for x εk εk' by(rule directed; rule d_IN_mono that)

    let  = "Γ ⊖ Sup (fst ` M)"

    have hM2: "current ?Γ h" if εh: "(ε, h) ∈ M" for ε h
    proof
      from εh have Field: "(ε, h) ∈ Field leq" by(rule Chains_FieldD[OF M])
      then have H: "current (Γ ⊖ ε) h" and ε_curr': "current Γ ε" by(rule h ε_curr)+
      from ε_curr' interpret Γ: countable_bipartite_web "Γ ⊖ ε" by(rule countable_bipartite_web_minus_web)

      fix x
      have "d_OUT h x ≤ d_OUT ?h x" using εh by(intro d_OUT_mono)(auto intro: SUP_upper2)
      also have OUT: "… = (SUP h:snd ` M. d_OUT h x)" using chain2 _ supp_flow2
        by(rule d_OUT_Sup)(simp_all add: nempty)
      also have "… = … + (SUP ε:fst ` M. d_OUT ε x) - (SUP ε:fst ` M. d_OUT ε x)"
        using OUT1_le[of x]
        by (intro ennreal_add_diff_cancel_right[symmetric] neq_top_trans[OF weight_finite, of _ x]) simp
      also have "… = (SUP (ε, k):M. d_OUT k x + d_OUT ε x) - (SUP ε:fst ` M. d_OUT ε x)" unfolding split_def
        by(subst SUP_add_directed_ennreal[OF directed_OUT])(simp_all add: )
      also have "(SUP (ε, k):M. d_OUT k x + d_OUT ε x) ≤ weight Γ x"
        apply(clarsimp dest!: Chains_FieldD[OF M] intro!: SUP_least)
        subgoal premises that for ε h
          using currentD_weight_OUT[OF h[OF that], of x] currentD_weight_OUT[OF ε_curr[OF that], of x]
             countable_bipartite_web_minus_web[OF ε_curr, THEN countable_bipartite_web.currentD_OUT', OF that h[OF that], where x=x]
          by (auto simp add: ennreal_le_minus_iff split: if_split_asm)
        done
      also have "(SUP ε:fst ` M. d_OUT ε x) = d_OUT (Sup (fst ` M)) x" using OUT1 by(simp add: split_beta)
      finally show "d_OUT h x ≤ weight ?Γ x"
        using Γ.currentD_OUT'[OF h[OF Field], of x] currentD_weight_IN[OF SM1, of x] by(auto simp add: ennreal_minus_mono)

      have "d_IN h x ≤ d_IN ?h x" using εh by(intro d_IN_mono)(auto intro: SUP_upper2)
      also have IN: "… = (SUP h:snd ` M. d_IN h x)" using chain2 _ supp_flow2
        by(rule d_IN_Sup)(simp_all add: nempty)
      also have "… = … + (SUP ε:fst ` M. d_IN ε x) - (SUP ε:fst ` M. d_IN ε x)"
        using IN1_le[of x]
        by (intro ennreal_add_diff_cancel_right[symmetric] neq_top_trans[OF weight_finite, of _ x]) simp
      also have "… = (SUP (ε, k):M. d_IN k x + d_IN ε x) - (SUP ε:fst ` M. d_IN ε x)" unfolding split_def
        by(subst SUP_add_directed_ennreal[OF directed_IN])simp_all
      also have "(SUP (ε, k):M. d_IN k x + d_IN ε x) ≤ weight Γ x"
        apply(clarsimp dest!: Chains_FieldD[OF M] intro!: SUP_least)
        subgoal premises that for ε h
          using currentD_weight_OUT[OF h, OF that, where x=x] currentD_weight_IN[OF h, OF that, where x=x]
            countable_bipartite_web_minus_web[OF ε_curr, THEN countable_bipartite_web.currentD_OUT', OF that h[OF that], where x=x]
            currentD_OUT'[OF ε_curr, OF that, where x=x] currentD_IN[OF ε_curr, OF that, of x] currentD_weight_IN[OF ε_curr, OF that, where x=x]
          by(auto simp add: ennreal_le_minus_iff
                     split: if_split_asm intro: add_increasing2 order_trans[rotated])
        done
      also have "(SUP ε:fst ` M. d_IN ε x) = d_IN (Sup (fst ` M)) x" using IN1 by(simp add: split_beta)
      finally show "d_IN h x ≤ weight ?Γ x"
        using currentD_IN[OF h[OF Field], of x] currentD_weight_OUT[OF SM1, of x]
        by(auto simp add: ennreal_minus_mono)
          (auto simp add: ennreal_le_minus_iff add_increasing2)
      show "h e = 0" if "e ∉ E" for e using currentD_outside'[OF H, of e] that by simp
    qed

    from nempty have "snd ` M ≠ {}" by simp
    from chain2 this _ supp_flow2 show current: "current ?Γ ?h"
      by(rule current_Sup)(clarify; rule hM2; simp)

    have wM: "wave ?Γ h" if "(ε, h) ∈ M" for ε h
    proof
      let ?Γ' = "Γ ⊖ ε"
      have subset: "TER?Γ' h ⊆ TER h"
        using currentD_OUT'[OF SM1] currentD_OUT'[OF ε_curr[OF Chains_FieldD[OF M that]]] that
        by(auto 4 7 elim!: SAT.cases intro: SAT.intros elim!: order_trans[rotated] intro: ennreal_minus_mono d_IN_mono intro!: SUP_upper2 split: if_split_asm)

      from h_w[OF Chains_FieldD[OF M], OF that] have "separating ?Γ' (TER?Γ' h)" by(rule waveD_separating)
      then show "separating ?Γ (TER h)" using subset by(auto intro: separating_weakening)
    qed(rule hM2[OF that])
    show wave: "wave ?Γ ?h" using chain2 ‹snd ` M ≠ {}› _ supp_flow2
      by(rule wave_lub)(clarify; rule wM; simp)

    def f  "F (Sup (fst ` M), Sup (snd ` M))"
    have supp_flow: "countable (support_flow f)"
      using supp_flow1 supp_flow2 support_flow_plus_current[of "Sup (fst ` M)" ?h]
      unfolding f_def F_simps by(blast intro: countable_subset)
    have f_alt: "f = Sup ((λ(ε, h). plus_current ε h) ` M)"
      apply(simp add: fun_eq_iff split_def f_def nempty F_def)
      apply(subst (1 2) add.commute)
      apply(subst SUP_add_directed_ennreal)
      apply(rule directed)
      apply(auto dest!: Chains_FieldD[OF M])
      done
    have f_curr: "current Γ f" unfolding f_def F_simps using SM1 current by(rule current_plus_current_minus)
    have IN_f: "d_IN f x = d_IN (Sup (fst ` M)) x + d_IN (Sup (snd ` M)) x" for x
      unfolding f_def F_simps plus_current_def
      by(rule d_IN_add SM1 current)+
    have OUT_f: "d_OUT f x = d_OUT (Sup (fst ` M)) x + d_OUT (Sup (snd ` M)) x" for x
      unfolding f_def F_simps plus_current_def
      by(rule d_OUT_add SM1 current)+

    show "¬ hindered (Γ ⊖ f)" (is "¬ hindered ?Ω") -- ‹Assertion 6.11›
    proof
      assume hindered: "hindered ?Ω"
      then obtain g where g: "current ?Ω g" and g_w: "wave ?Ω g" and hindrance: "hindrance ?Ω g" by cases
      from hindrance obtain z where z: "z ∈ A Γ" and zℰ: "z ∉ ℰ (TER g)"
        and OUT_z: "d_OUT g z < weight ?Ω z" by cases auto
      def δ  "weight ?Ω z - d_OUT g z"
      have δ_pos: "δ > 0" using OUT_z by(simp add: δ_def diff_gr0_ennreal del: minus_web_sel)
      then have δ_finite[simp]: "δ ≠ ⊤" using z
        by(simp_all add: δ_def)

      have "∃(ε, h) ∈ M. d_OUT f a < d_OUT (plus_current ε h) a + δ"
      proof(rule ccontr)
        assume "¬ ?thesis"
        hence greater: "d_OUT (plus_current ε h) a + δ ≤ d_OUT f a" if "(ε, h) ∈ M" for ε h using that by auto

        have chain'': "Complete_Partial_Order.chain op ≤ ((λ(ε, h). plus_current ε h) ` M)"
          using chain' by(rule chain_imageI)(auto simp add: le_fun_def add_mono)

        have "d_OUT f a + 0 < d_OUT f a + δ"
          using currentD_finite_OUT[OF f_curr, of a] by (simp add: δ_pos)
        also have "d_OUT f a + δ = (SUP (ε, h):M. d_OUT (plus_current ε h) a) + δ"
          using chain'' nempty supp_flow
          unfolding f_alt by(subst d_OUT_Sup)(simp_all add: plus_current_def[abs_def] split_def)
        also have "… ≤ d_OUT f a"
          unfolding ennreal_SUP_add_left[symmetric, OF nempty]
        proof(rule SUP_least, clarify)
          show "d_OUT (plus_current ε h) a + δ ≤ d_OUT f a" if "(ε, h) ∈ M" for ε h
            using greater[OF that] currentD_finite_OUT[OF Chains_FieldD[OF M that, THEN f], of a]
            by(auto simp add: ennreal_le_minus_iff add.commute F_def)
        qed
        finally show False by simp
      qed
      then obtain ε h where hM: "(ε, h) ∈ M" and close: "d_OUT f a < d_OUT (plus_current ε h) a + δ" by blast
      have Field: "(ε, h) ∈ Field leq" using hM by(rule Chains_FieldD[OF M])
      then have ε: "current Γ ε"
        and unhindered_h: "¬ hindered (Γ ⊖ F (ε, h))"
        and h_curr: "current (Γ ⊖ ε) h"
        and h_w: "wave (Γ ⊖ ε) h"
        and OUT_ε: "x ≠ a ⟹ d_OUT ε x = 0" for x
        by(rule ε_curr OUT_ε h h_w unhindered')+
      let ?εh = "plus_current ε h"
      have εh_curr: "current Γ ?εh" using Field unfolding F_simps[symmetric] by(rule f)

      interpret h: countable_bipartite_web "Γ ⊖ ε" using ε by(rule countable_bipartite_web_minus_web)
      interpret εh: countable_bipartite_web "Γ ⊖ ?εh" using εh_curr by(rule countable_bipartite_web_minus_web)
      note [simp del] = minus_web_sel(2)
        and [simp] = weight_minus_web[OF εh_curr] weight_minus_web[OF SM1, simplified]

      def k  "λe. Sup (fst ` M) e - ε e"
      have k_simps: "k (x, y) = Sup (fst ` M) (x, y) - ε (x, y)" for x y by(simp add: k_def)
      have k_alt: "k (x, y) = (if x = a ∧ edge Γ x y then Sup (fst ` M) (a, y) - ε (a, y) else 0)" for x y
        by(cases "x = a")(auto simp add: k_simps out_ε[OF Field] currentD_outside[OF ε] intro!: SUP_eq_const[OF nempty] dest!: Chains_FieldD[OF M] intro: currentD_outside[OF ε_curr] out_ε)
      have OUT_k: "d_OUT k x = (if x = a then d_OUT (Sup (fst ` M)) a - d_OUT ε a else 0)" for x
      proof -
        have "d_OUT k x = (if x = a then (∑+ y. Sup (fst ` M) (a, y) - ε (a, y)) else 0)"
          using currentD_outside[OF SM1] currentD_outside[OF ε]
          by(auto simp add: k_alt d_OUT_def intro!: nn_integral_cong)
        also have "(∑+ y. Sup (fst ` M) (a, y) - ε (a, y)) = d_OUT (Sup (fst `M)) a - d_OUT ε a"
          using currentD_finite_OUT[OF ε, of a] hM unfolding d_OUT_def
          by(subst nn_integral_diff[symmetric])(auto simp add: AE_count_space intro!: SUP_upper2)
        finally show ?thesis .
      qed
      have IN_k: "d_IN k y = (if edge Γ a y then Sup (fst ` M) (a, y) - ε (a, y) else 0)" for y
      proof -
        have "d_IN k y = (∑+ x. (if edge Γ x y then Sup (fst ` M) (a, y) - ε (a, y) else 0) * indicator {a} x)"
          unfolding d_IN_def by(rule nn_integral_cong)(auto simp add: k_alt outgoing_def split: split_indicator)
        also have "… = (if edge Γ a y then Sup (fst ` M) (a, y) - ε (a, y) else 0)" using hM
          by(auto simp add: max_def intro!: SUP_upper2)
        finally show ?thesis .
      qed

      have OUT_εh: "d_OUT ?εh x = d_OUT ε x + d_OUT h x" for x
        unfolding plus_current_def by(rule d_OUT_add)+
      have IN_εh: "d_IN ?εh x = d_IN ε x + d_IN h x" for x
        unfolding plus_current_def by(rule d_IN_add)+

      have OUT1_le': "d_OUT (Sup (fst`M)) x ≤ weight Γ x" for x
        using OUT1_le[of x] unfolding OUT1 by (simp add: split_beta')

      have k: "current (Γ ⊖ ?εh) k"
      proof
        fix x
        show "d_OUT k x ≤ weight (Γ ⊖ ?εh) x"
          using a OUT1_na[of x] currentD_weight_OUT[OF hM2[OF hM], of x] currentD_weight_IN[OF εh_curr, of x]
            currentD_weight_IN[OF ε, of x] OUT1_le'[of x]
          apply(auto simp add: diff_add_eq_diff_diff_swap_ennreal diff_add_assoc2_ennreal[symmetric]
                               OUT_k OUT_ε OUT_εh IN_εh currentD_OUT'[OF ε] IN_ε[OF Field] h.currentD_OUT'[OF h_curr])
          apply(subst diff_diff_commute_ennreal)
          apply(intro ennreal_minus_mono)
          apply(auto simp add: ennreal_le_minus_iff ac_simps less_imp_le OUT1)
          done

        have *: "(⨆xa∈M. fst xa (a, x)) ≤ d_IN (Sup (fst`M)) x"
          unfolding IN1 by (intro SUP_subset_mono) (auto simp: split_beta' d_IN_ge_point)
        also have "… ≤ weight Γ x"
          using IN1_le[of x] IN1 by (simp add: split_beta')
        finally show "d_IN k x ≤ weight (Γ ⊖ ?εh) x"
          using currentD_weight_IN[OF εh_curr, of x] currentD_weight_OUT[OF εh_curr, of x]
            currentD_weight_IN[OF hM2[OF hM], of x] IN_ε[OF Field, of x] *
          apply(auto simp add: IN_k outgoing_def IN_εh IN_ε A_in diff_add_eq_diff_diff_swap_ennreal)
          apply(subst diff_diff_commute_ennreal)
          apply(intro ennreal_minus_mono[OF _ order_refl])
          apply(auto simp add: ennreal_le_minus_iff ac_simps intro: order_trans add_mono)
          done
        show "k e = 0" if "e ∉ EΓ ⊖ ?εh" for e using that by(cases e)(simp add: k_alt)
      qed

      def q  "∑+ y∈B (Γ ⊖ ?εh). d_IN k y - d_OUT k y"
      have q_alt: "q = (∑+ y∈- A (Γ ⊖ ?εh). d_IN k y - d_OUT k y)" using disjoint
        by(auto simp add: q_def nn_integral_count_space_indicator currentD_outside_OUT[OF k] currentD_outside_IN[OF k] not_vertex split: split_indicator intro!: nn_integral_cong)
      have q_simps: "q = d_OUT (Sup (fst ` M)) a - d_OUT ε a"
      proof -
        have "q = (∑+ y. d_IN k y)" using a IN1 OUT1 OUT1_na unfolding q_alt
          by(auto simp add: nn_integral_count_space_indicator OUT_k IN_ε[OF Field] OUT_ε currentD_outside[OF ε] outgoing_def no_loop A_in IN_k intro!: nn_integral_cong split: split_indicator)
        also have "… = d_OUT (Sup (fst ` M)) a - d_OUT ε a" using currentD_finite_OUT[OF ε, of a] hM currentD_outside[OF SM1] currentD_outside[OF ε]
          by(subst d_OUT_diff[symmetric])(auto simp add: d_OUT_def IN_k intro!: SUP_upper2 nn_integral_cong)
        finally show ?thesis .
      qed
      have q_finite: "q ≠ ⊤" using currentD_finite_OUT[OF SM1, of a]
        by(simp add: q_simps)
      have q_nonneg: "0 ≤ q" using hM by(auto simp add: q_simps intro!: d_OUT_mono SUP_upper2)
      have q_less_δ: "q < δ" using close
        unfolding q_simps δ_def OUT_εh OUT_f
      proof -
        let ?F = "d_OUT (Sup (fst`M)) a" and ?S = "d_OUT (Sup (snd`M)) a"
          and  = "d_OUT ε a" and ?h = "d_OUT h a" and ?w = "weight (Γ ⊖ f) z - d_OUT g z"
        have "?F + ?h ≤ ?F + ?S"
          using hM by (auto intro!: add_mono d_OUT_mono SUP_upper2)
        also assume "?F + ?S < ?ε + ?h + ?w"
        finally have "?h + ?F < ?h + (?w + ?ε)"
          by (simp add: ac_simps)
        then show "?F - ?ε < ?w"
          using currentD_finite_OUT[OF ε, of a] hM unfolding ennreal_add_left_cancel_less
          by (subst minus_less_iff_ennreal) (auto intro!: d_OUT_mono SUP_upper2 simp: less_top)
      qed

      def g'  "plus_current g (Sup (snd ` M) - h)"
      have g'_simps: "g' e = g e + Sup (snd ` M) e - h e" for e
        using hM by(auto simp add: g'_def intro!: add_diff_eq_ennreal intro: SUP_upper2)
      have OUT_g': "d_OUT g' x = d_OUT g x + (d_OUT (Sup (snd ` M)) x - d_OUT h x)" for x
        unfolding g'_simps[abs_def] using εh.currentD_finite_OUT[OF k] hM h.currentD_finite_OUT[OF h_curr] hM
        apply(subst d_OUT_diff)
         apply(auto simp add: add_diff_eq_ennreal[symmetric] k_simps intro: add_increasing intro!: SUP_upper2)
        apply(subst d_OUT_add)
         apply(auto simp add: add_diff_eq_ennreal[symmetric] k_simps intro: add_increasing intro!:)
        apply(simp add: add_diff_eq_ennreal SUP_apply[abs_def])
        apply(auto simp add: g'_def intro!: add_diff_eq_ennreal[symmetric] d_OUT_mono intro: SUP_upper2)
        done
      have IN_g': "d_IN g' x = d_IN g x + (d_IN (Sup (snd ` M)) x - d_IN h x)" for x
        unfolding g'_simps[abs_def] using εh.currentD_finite_IN[OF k] hM h.currentD_finite_IN[OF h_curr] hM
        apply(subst d_IN_diff)
         apply(auto simp add: add_diff_eq_ennreal[symmetric] k_simps intro: add_increasing intro!: SUP_upper2)
        apply(subst d_IN_add)
         apply(auto simp add: add_diff_eq_ennreal[symmetric] k_simps intro: add_increasing intro!: SUP_upper)
        apply(auto simp add: g'_def SUP_apply[abs_def] intro!: add_diff_eq_ennreal[symmetric] d_IN_mono intro: SUP_upper2)
        done

      have h': "current (Γ ⊖ Sup (fst ` M)) h" using hM by(rule hM2)

      let  = "Γ ⊖ ?εh ⊖ k"
      interpret Γ: web  using k by(rule εh.web_minus_web)
      note [simp] = εh.weight_minus_web[OF k] h.weight_minus_web[OF h_curr]
        weight_minus_web[OF f_curr] SM1.weight_minus_web[OF h', simplified]

      interpret Ω: countable_bipartite_web "Γ ⊖ f" using f_curr by(rule countable_bipartite_web_minus_web)

      have *: "Γ ⊖ f = Γ ⊖ Sup (fst ` M) ⊖ Sup (snd ` M)" unfolding f_def F_simps
        using SM1 current by(rule minus_plus_current)
      have OUT_εk: "d_OUT (Sup (fst ` M)) x = d_OUT ε x + d_OUT k x" for x
        using OUT1'[of x] currentD_finite_OUT[OF ε] hM
        by(auto simp add: OUT_k OUT_ε add_diff_self_ennreal SUP_upper2)
      have IN_εk: "d_IN (Sup (fst ` M)) x = d_IN ε x + d_IN k x" for x
        using IN1'[of x] currentD_finite_IN[OF ε] currentD_outside[OF ε] currentD_outside[OF ε_curr]
        by(auto simp add: IN_k IN_ε[OF Field] add_diff_self_ennreal split_beta nempty
                dest!: Chains_FieldD[OF M] intro!: SUP_eq_const intro: SUP_upper2[OF hM])
      have **: "?Γ = Γ ⊖ Sup (fst ` M) ⊖ h"
      proof(rule web.equality)
        show "weight ?Γ = weight (Γ ⊖ Sup (fst ` M) ⊖ h)"
          using OUT_εk OUT_εh currentD_finite_OUT[OF ε] IN_εk IN_εh currentD_finite_IN[OF ε]
          by(auto simp add: diff_add_eq_diff_diff_swap_ennreal diff_diff_commute_ennreal)
      qed simp_all
      have g'_alt: "g' = plus_current (Sup (snd ` M)) g - h"
        by(simp add: fun_eq_iff g'_simps add_diff_eq_ennreal add.commute)

      have "current (Γ ⊖ Sup (fst ` M)) (plus_current (Sup (snd ` M)) g)" using current g unfolding *
        by(rule SM1.current_plus_current_minus)
      hence g': "current ?Γ g'" unfolding * ** g'_alt using hM2[OF hM]
        by(rule SM1.current_minus)(auto intro!: add_increasing2 SUP_upper2 hM)

      have "wave (Γ ⊖ Sup (fst ` M)) (plus_current (Sup (snd ` M)) g)" using current wave g g_w
        unfolding * by(rule SM1.wave_plus_current_minus)
      then have g'_w: "wave ?Γ g'" unfolding * ** g'_alt using hM2[OF hM]
        by(rule SM1.wave_minus)(auto intro!: add_increasing2 SUP_upper2 hM)

      have "hindrance_by ?Γ g' q"
      proof
        show "z ∈ A ?Γ" using z by simp
        show "z ∉ ℰ (TER g')"
        proof
          assume "z ∈ ℰ (TER g')"
          hence OUT_z: "d_OUT g' z = 0"
            and ess: "essential ?Γ (B Γ) (TER g') z" by(simp_all add: SINK.simps)
          from ess obtain p y where p: "path Γ z p y" and y: "y ∈ B Γ"
            and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF (TER g')" by(rule essentialE_RF) auto
          from y have y': "y ∉ A Γ" using disjoint by blast
          from p z y obtain py: "p = [y]" and edge: "edge Γ z y" using disjoint
            by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E)
          hence yRF: "y ∉ RF (TER g')" using bypass[of y] by(auto)
          with wave_not_RF_IN_zero[OF g' g'_w, of y] have IN_g'_y: "d_IN g' y = 0"
            by(auto intro: roofed_greaterI)
          with yRF y y' have w_y: "weight ?Γ y > 0" using currentD_OUT[OF g', of y]
            by(auto simp add: RF_in_B currentD_SAT[OF g'] SINK.simps zero_less_iff_neq_zero)
          have "y ∉ SAT (Γ ⊖ f) g"
          proof
            assume "y ∈ SAT (Γ ⊖ f) g"
            with y disjoint have IN_g_y: "d_IN g y = weight (Γ ⊖ f) y" by(auto simp add: currentD_SAT[OF g])
            have "0 < weight Γ y - d_IN (⨆x∈M. fst x) y - d_IN h y"
              using y' w_y unfolding ** by auto
            have "d_IN g' y > 0"
              using y' w_y hM unfolding **
              apply(simp add: IN_g' IN_f IN_g_y diff_add_eq_diff_diff_swap_ennreal)
              apply(subst add_diff_eq_ennreal)
              apply(auto intro!: SUP_upper2 d_IN_mono simp: diff_add_self_ennreal diff_gt_0_iff_gt_ennreal)
              done
            with IN_g'_y show False by simp
          qed
          then have "y ∉ TERΓ ⊖ f g" by simp
          with p y py have "essential Γ (B Γ) (TERΓ ⊖ f g) z" by(auto intro: essentialI)
          moreover with z waveD_separating[OF g_w, THEN separating_RF_A] have "z ∈ ℰ (TER g)"
            by(auto simp add: RF_in_essential)
          with zℰ show False by contradiction
        qed
        have "δ ≤ weight ?Γ z - d_OUT g' z"
          unfolding ** OUT_g' using z
          apply (simp add: δ_def OUT_f diff_add_eq_diff_diff_swap_ennreal)
          apply (subst (5) diff_diff_commute_ennreal)
          apply (rule ennreal_minus_mono[OF _ order_refl])
          apply (auto simp add: ac_simps diff_add_eq_diff_diff_swap_ennreal[symmetric] add_diff_self_ennreal
                      intro!: ennreal_minus_mono[OF order_refl] SUP_upper2[OF hM] d_OUT_mono)
          done
        then show q_z: "q < weight ?Γ z - d_OUT g' z" using q_less_δ by simp
        then show "d_OUT g' z < weight ?Γ z" using q_nonneg z
          by(auto simp add: less_diff_eq_ennreal less_top[symmetric] ac_simps Γ.currentD_finite_OUT[OF g']
                  intro: le_less_trans[rotated] add_increasing)
      qed
      then have hindered_by: "hindered_by (Γ ⊖ ?εh ⊖ k) q" using g' g'_w by(rule hindered_by.intros)
      then have "hindered (Γ ⊖ ?εh)" using q_finite unfolding q_def by -(rule εh.hindered_reduce_current[OF k])
      with unhindered_h show False unfolding F_simps by contradiction
    qed
  qed

  def sat  "λ(ε, h). let
      f = F (ε, h);
      k = SOME k. current (Γ ⊖ f) k ∧ wave (Γ ⊖ f) k ∧ (∀k'. current (Γ ⊖ f) k' ∧ wave (Γ ⊖ f) k' ∧ k ≤ k' ⟶ k = k')
    in
      if d_OUT (plus_current f k) a < weight Γ a then
        let
          Ω = Γ ⊖ f ⊖ k;
          y = SOME y. y ∈ OUTΩ a ∧ weight Ω y > 0;
          δ = SOME δ. δ > 0 ∧ δ < enn2real (min (weight Ω a) (weight Ω y)) ∧ ¬ hindered (reduce_weight Ω y δ)
        in
          (plus_current ε (zero_current((a, y) := δ)), plus_current h k)
      else (ε, h)"

  have zero: "(zero_current, zero_current) ∈ Field leq"
    by(rule F_I)(simp_all add: unhindered  F_def)

  have a_TER: "a ∈ TERΓ ⊖ F εh k"
    if that: "εh ∈ Field leq"
    and k: "current (Γ ⊖ F εh) k" and k_w: "wave (Γ ⊖ F εh) k"
    and less: "d_OUT (plus_current (F εh) k) a < weight Γ a" for εh k
  proof(rule ccontr)
    assume "¬ ?thesis"
    hence : "a ∉ ℰΓ ⊖ F εh (TERΓ ⊖ F εh k)" by auto
    from that have f: "current Γ (F εh)" and unhindered: "¬ hindered (Γ ⊖ F εh)"
       by(cases εh; simp add: f unhindered'; fail)+

    from less have "d_OUT k a < weight (Γ ⊖ F εh) a" using a currentD_finite_OUT[OF f, of a]
      by(simp add: d_OUT_def nn_integral_add less_diff_eq_ennreal add.commute less_top[symmetric])
    with _  have "hindrance (Γ ⊖ F εh) k" by(rule hindrance)(simp add: a)
    then have "hindered (Γ ⊖ F εh)" using k k_w ..
    with unhindered show False by contradiction
  qed

  note minus_web_sel(2)[simp del]

  let ?P_y = "λεh k y. y ∈ OUTΓ ⊖ F εh ⊖ k a ∧ weight (Γ ⊖ F εh ⊖ k) y > 0"
  have Ex_y: "Ex (?P_y εh k)"
    if that: "εh ∈ Field leq"
    and k: "current (Γ ⊖ F εh) k" and k_w: "wave (Γ ⊖ F εh) k"
    and less: "d_OUT (plus_current (F εh) k) a < weight Γ a" for εh k
  proof(rule ccontr)
    let  = "Γ ⊖ F εh ⊖ k"
    assume *: "¬ ?thesis"

    interpret Γ: countable_bipartite_web "Γ ⊖ F εh" using f[OF that] by(rule countable_bipartite_web_minus_web)
    note [simp] = weight_minus_web[OF f[OF that]] Γ.weight_minus_web[OF k]

    have "hindrance ?Ω zero_current"
    proof
      show "a ∈ A ?Ω" using a by simp
      show "a ∉ ℰ (TER zero_current)" (is "a ∉ ℰ_ ?TER")
      proof
        assume "a ∈ ℰ ?TER"
        then obtain p y where p: "path ?Ω a p y" and y: "y ∈ B ?Ω"
          and bypass: "⋀z. z ∈ set p ⟹ z ∉ RF ?TER" by(rule ℰ_E_RF)(auto)
        from a p y disjoint have Nil: "p ≠ []" by(auto simp add: rtrancl_path_simps)
        hence "edge ?Ω a (p ! 0)" "p ! 0 ∉ RF ?TER"
          using rtrancl_path_nth[OF p, of 0] bypass by auto
        with * show False by(auto simp add: not_less outgoing_def intro: roofed_greaterI)
      qed

      have "d_OUT (plus_current (F εh) k) x = d_OUT (F εh) x + d_OUT k x" for x
        by(simp add: d_OUT_def nn_integral_add)
      then show "d_OUT zero_current a < weight ?Ω a" using less a_TER[OF that k k_w less] a
        by(simp add: SINK.simps diff_gr0_ennreal)
    qed
    hence "hindered ?Ω"
      by(auto intro!: hindered.intros order_trans[OF currentD_weight_OUT[OF k]] order_trans[OF currentD_weight_IN[OF k]])
    moreover have "¬ hindered ?Ω" using unhindered'[OF that] k k_w by(rule Γ.unhindered_minus_web)
    ultimately show False by contradiction
  qed

  have increasing: "εh ≤ sat εh ∧ sat εh ∈ Field leq" if "εh ∈ Field leq" for εh
  proof(cases εh)
    case (Pair ε h)
    with that have that: "(ε, h) ∈ Field leq" by simp
    have f: "current Γ (F (ε, h))" and unhindered: "¬ hindered (Γ ⊖ F (ε, h))"
      and ε: "current Γ ε"
      and h: "current (Γ ⊖ ε) h" and h_w: "wave (Γ ⊖ ε) h" and OUT_ε: "x ≠ a ⟹ d_OUT ε x = 0" for x
      using that by(rule f unhindered' ε_curr OUT_ε h h_w)+
    interpret Γ: countable_bipartite_web "Γ ⊖ F (ε, h)" using f by(rule countable_bipartite_web_minus_web)
    note [simp] = weight_minus_web[OF f]

    let ?P_k = "λk. current (Γ ⊖ F (ε, h)) k ∧ wave (Γ ⊖ F (ε, h)) k ∧ (∀k'. current (Γ ⊖ F (ε, h)) k' ∧ wave (Γ ⊖ F (ε, h)) k' ∧ k ≤ k' ⟶ k = k')"
    def k  "Eps ?P_k"
    have "Ex ?P_k" by(intro ex_maximal_wave)(simp_all)
    hence "?P_k k" unfolding k_def by(rule someI_ex)
    hence k: "current (Γ ⊖ F (ε, h)) k" and k_w: "wave (Γ ⊖ F (ε, h)) k"
      and maximal: "⋀k'. ⟦ current (Γ ⊖ F (ε, h)) k'; wave (Γ ⊖ F (ε, h)) k'; k ≤ k' ⟧ ⟹ k = k'" by blast+
    note [simp] = Γ.weight_minus_web[OF k]

    let ?fk = "plus_current (F (ε, h)) k"
    have IN_fk: "d_IN ?fk x = d_IN (F (ε, h)) x + d_IN k x" for x
      by(simp add: d_IN_def nn_integral_add)
    have OUT_fk: "d_OUT ?fk x = d_OUT (F (ε, h)) x + d_OUT k x" for x
      by(simp add: d_OUT_def nn_integral_add)
    have fk: "current Γ ?fk" using f k by(rule current_plus_current_minus)

    show ?thesis
    proof(cases "d_OUT ?fk a < weight Γ a")
      case less: True

      def Ω  "Γ ⊖ F (ε, h) ⊖ k"
      have B_Ω [simp]: "B Ω = B Γ" by(simp add: Ω_def)

      have loose: "loose Ω" unfolding Ω_def using unhindered k k_w maximal by(rule Γ.loose_minus_web)
      interpret Ω: countable_bipartite_web Ω using k unfolding Ω_def
        by(rule Γ.countable_bipartite_web_minus_web)

      have a_ℰ: "a ∈ TERΓ ⊖ F (ε, h) k" using that k k_w less by(rule a_TER)
      then have weight_Ω_a: "weight Ω a = weight Γ a - d_OUT (F (ε, h)) a"
        using a disjoint by(auto simp add: roofed_circ_def Ω_def SINK.simps)
      then have weight_a: "0 < weight Ω a" using less a_ℰ
        by(simp add: OUT_fk SINK.simps diff_gr0_ennreal)

      let ?P_y = "λy. y ∈ OUTΩ a ∧ weight Ω y > 0"
      def y  "Eps ?P_y"
      have "Ex ?P_y" using that k k_w less unfolding Ω_def by(rule Ex_y)
      hence "?P_y y" unfolding y_def by(rule someI_ex)
      hence y_OUT: "y ∈ OUTΩ a" and weight_y: "weight Ω y > 0" by blast+
      from y_OUT have y_B: "y ∈ B Ω" by(auto simp add: outgoing_def Ω_def dest: bipartite_E)
      with weight_y have yRF: "y ∉ RFΓ ⊖ F (ε, h) (TERΓ ⊖ F (ε, h) k)"
        unfolding Ω_def using currentD_OUT[OF k, of y] disjoint
        by(auto split: if_split_asm simp add: SINK.simps currentD_SAT[OF k] roofed_circ_def RF_in_B Γ.currentD_finite_IN[OF k])
      hence IN_k_y: "d_IN k y = 0" by(rule wave_not_RF_IN_zero[OF k k_w])

      def bound  "enn2real (min (weight Ω a) (weight Ω y))"
      have bound_pos: "bound > 0" using weight_y weight_a using Ω.weight_finite
        by(cases "weight Ω a" "weight Ω y" rule: ennreal2_cases)
          (simp_all add: bound_def min_def split: if_split_asm)

      let ?P_δ = "λδ. δ > 0 ∧ δ < bound ∧ ¬ hindered (reduce_weight Ω y δ)"
      def δ  "Eps ?P_δ"
      let  = "reduce_weight Ω y δ"

      from Ω.unhinder[OF loose _ weight_y bound_pos] y_B disjoint
      have "Ex ?P_δ" by(auto simp add: Ω_def)
      hence "?P_δ δ" unfolding δ_def by(rule someI_ex)
      hence δ_pos: "0 < δ" and δ_le_bound: "δ < bound" and unhindered': "¬ hindered ?Ω" by blast+
      from δ_pos have δ_nonneg: "0 ≤ δ" by simp
      from δ_le_bound δ_pos have δ_le_a: "δ < weight Ω a" and δ_le_y: "δ < weight Ω y"
        by(cases "weight Ω a" "weight Ω y" rule: ennreal2_cases;
           simp add: bound_def min_def ennreal_less_iff split: if_split_asm)+

      let  = "Γ ⊖ ?fk"
      interpret Γ': countable_bipartite_web  by(rule countable_bipartite_web_minus_web fk)+
      note [simp] = weight_minus_web[OF fk]

      let ?g = "zero_current((a, y) := δ)"
      have OUT_g: "d_OUT ?g x = (if x = a then δ else 0)" for x
      proof(rule trans)
        show "d_OUT ?g x = (∑+ z. (if x = a then δ else 0) * indicator {y} z)" unfolding d_OUT_def
          by(rule nn_integral_cong) simp
        show "… = (if x = a then δ else 0)" using δ_pos by(simp add: max_def)
      qed
      have IN_g: "d_IN ?g x = (if x = y then δ else 0)" for x
      proof(rule trans)
        show "d_IN ?g x = (∑+ z. (if x = y then δ else 0) * indicator {a} z)" unfolding d_IN_def
          by(rule nn_integral_cong) simp
        show "… = (if x = y then δ else 0)" using δ_pos by(simp add: max_def)
      qed

      have g: "current ?Γ ?g"
      proof
        show "d_OUT ?g x ≤ weight ?Γ x" for x
        proof(cases "x = a")
          case False
          then show ?thesis using currentD_weight_OUT[OF fk, of x] currentD_weight_IN[OF fk, of x]
            by(auto simp add: OUT_g zero_ennreal_def[symmetric])
        next
          case True
          then show ?thesis using δ_le_a a a_ℰ δ_pos unfolding OUT_g
            by(simp add: OUT_g Ω_def SINK.simps OUT_fk split: if_split_asm)
        qed
        show "d_IN ?g x ≤ weight ?Γ x" for x
        proof(cases "x = y")
          case False
          then show ?thesis using currentD_weight_OUT[OF fk, of x] currentD_weight_IN[OF fk, of x]
            by(auto simp add: IN_g zero_ennreal_def[symmetric])
        next
          case True
          then show ?thesis using δ_le_y y_B a_ℰ δ_pos currentD_OUT[OF k, of y] IN_k_y
            by(simp add: OUT_g Ω_def SINK.simps OUT_fk IN_fk IN_g split: if_split_asm)
        qed
        show "?g e = 0" if "e ∉ E" for e using y_OUT that by(auto simp add: Ω_def outgoing_def)
      qed
      interpret Γ'': web "Γ ⊖ ?fk ⊖ ?g" using g by(rule Γ'.web_minus_web)

      let ?ε' = "plus_current ε (zero_current((a, y) := δ))"
      let ?h' = "plus_current h k"
      have F': "F (?ε', ?h') = plus_current (plus_current (F (ε, h)) k) (zero_current((a, y) := δ))" (is "_ = ?f'")
        by(auto simp add: F_simps fun_eq_iff add_ac)
      have sat: "sat (ε, h) = (?ε', ?h')" using less
        by(simp add: sat_def k_def Ω_def Let_def y_def bound_def δ_def)

      have le: "(ε, h) ≤ (?ε', ?h')" using δ_pos
        by(auto simp add: le_fun_def add_increasing2 add_increasing)

      have "current (Γ ⊖ ε) ((λ_. 0)((a, y) := ennreal δ))" using g
        by(rule current_weight_mono)(auto simp add: weight_minus_web[OF ε] intro!: ennreal_minus_mono d_OUT_mono d_IN_mono, simp_all add: F_def add_increasing2)
      with ε have ε': "current Γ ?ε'" by(rule current_plus_current_minus)
      moreover have eq_0: "d_OUT ?ε' x = 0" if "x ≠ a" for x unfolding plus_current_def using that
        by(subst d_OUT_add)(simp_all add: δ_nonneg d_OUT_fun_upd OUT_ε)
      moreover
      from ε' interpret ε': countable_bipartite_web "Γ ⊖ ?ε'" by(rule countable_bipartite_web_minus_web)
      from ε interpret ε: countable_bipartite_web "Γ ⊖ ε" by(rule countable_bipartite_web_minus_web)
      have g': "current (Γ ⊖ ε) ?g" using g
        apply(rule current_weight_mono)
        apply(auto simp add: weight_minus_web[OF ε] intro!: ennreal_minus_mono d_OUT_mono d_IN_mono)
        apply(simp_all add: F_def add_increasing2)
        done
      have k': "current (Γ ⊖ ε ⊖ h) k" using k unfolding F_simps minus_plus_current[OF ε h] .
      with h have "current (Γ ⊖ ε) (plus_current h k)" by(rule ε.current_plus_current_minus)
      hence "current (Γ ⊖ ε) (plus_current (plus_current h k) ?g)" using g unfolding minus_plus_current[OF f k]
        unfolding F_simps minus_plus_current[OF ε h] ε.minus_plus_current[OF h k', symmetric]
        by(rule ε.current_plus_current_minus)
      then have "current (Γ ⊖ ε ⊖ ?g) (plus_current (plus_current h k) ?g - ?g)" using g'
        by(rule ε.current_minus)(auto simp add: add_increasing)
      then have h'': "current (Γ ⊖ ?ε') ?h'"
        by(rule arg_cong2[where f=current, THEN iffD1, rotated -1])
          (simp_all add: minus_plus_current[OF ε g'] fun_eq_iff add_diff_eq_ennreal[symmetric])
      moreover have "wave (Γ ⊖ ?ε') ?h'"
      proof
        have "separating (Γ ⊖ ε) (TERΓ ⊖ ε (plus_current h k))"
          using k k_w unfolding F_simps minus_plus_current[OF ε h]
          by(intro waveD_separating ε.wave_plus_current_minus[OF h h_w])
        moreover have "TERΓ ⊖ ε (plus_current h k) ⊆ TERΓ ⊖ ?ε' (plus_current h k)"
          by(auto 4 4 simp add: SAT.simps weight_minus_web[OF ε] weight_minus_web[OF ε'] split: if_split_asm elim: order_trans[rotated] intro!: ennreal_minus_mono d_IN_mono add_increasing2 δ_nonneg)
        ultimately show sep: "separating (Γ ⊖ ?ε') (TERΓ ⊖ ?ε' ?h')"
          by(simp add: minus_plus_current[OF ε g'] separating_weakening)
      qed(rule h'')
      moreover
      have "¬ hindered (Γ ⊖ F (?ε', ?h'))" using unhindered'
      proof(rule contrapos_nn)
        assume "hindered (Γ ⊖ F (?ε', ?h'))"
        thus "hindered ?Ω"
        proof(rule hindered_mono_web[rotated -1])
          show "weight ?Ω z = weight (Γ ⊖ F (?ε', ?h')) z" if "z ∉ A (Γ ⊖ F (?ε', ?h'))" for z
            using that unfolding F'
            apply(cases "z = y")
            apply(simp_all add: Ω_def minus_plus_current[OF fk g] Γ'.weight_minus_web[OF g] IN_g)
            apply(simp_all add: plus_current_def d_IN_add diff_add_eq_diff_diff_swap_ennreal currentD_finite_IN[OF f])
            done
          have "y ≠ a" using y_B a disjoint by auto
          then show "weight (Γ ⊖ F (?ε', ?h')) z ≤ weight ?Ω z" if "z ∈ A (Γ ⊖ F (?ε', ?h'))" for z
            using that y_B disjoint δ_nonneg unfolding F'
            apply(cases "z = a")
            apply(simp_all add: Ω_def minus_plus_current[OF fk g] Γ'.weight_minus_web[OF g] OUT_g)
            apply(auto simp add: plus_current_def d_OUT_add diff_add_eq_diff_diff_swap_ennreal currentD_finite_OUT[OF f])
            done
        qed(simp_all add: Ω_def)
      qed
      ultimately have "(?ε', ?h') ∈ Field leq" by-(rule F_I)
      with Pair le sat that show ?thesis by(auto)
    next
      case False
      with currentD_weight_OUT[OF fk, of a] have "d_OUT ?fk a = weight Γ a" by simp
      have "sat εh = εh" using False Pair by(simp add: sat_def k_def)
      thus ?thesis using that Pair by(auto)
    qed
  qed

  have "bourbaki_witt_fixpoint Sup leq sat" using increasing chain_Field unfolding leq_def
    by(intro bourbaki_witt_fixpoint_restrict_rel)(auto intro: Sup_upper Sup_least)
  then interpret bourbaki_witt_fixpoint Sup leq sat .

  def f  "fixp_above (zero_current, zero_current)"
  have Field: "f ∈ Field leq" using fixp_above_Field[OF zero] unfolding f_def .
  then have f: "current Γ (F f)" and unhindered: "¬ hindered (Γ ⊖ F f)"
    by(cases f; simp add: f unhindered'; fail)+
  interpret Γ: countable_bipartite_web "Γ ⊖ F f" using f by(rule countable_bipartite_web_minus_web)
  note [simp] = weight_minus_web[OF f]
  have Field': "(fst f, snd f) ∈ Field leq" using Field by simp

  let ?P_k = "λk. current (Γ ⊖ F f) k ∧ wave (Γ ⊖ F f) k ∧ (∀k'. current (Γ ⊖ F f) k' ∧ wave (Γ ⊖ F f) k' ∧ k ≤ k' ⟶ k = k')"
  def k  "Eps ?P_k"
  have "Ex ?P_k" by(intro ex_maximal_wave)(simp_all)
  hence "?P_k k" unfolding k_def by(rule someI_ex)
  hence k: "current (Γ ⊖ F f) k" and k_w: "wave (Γ ⊖ F f) k"
    and maximal: "⋀k'. ⟦ current (Γ ⊖ F f) k'; wave (Γ ⊖ F f) k'; k ≤ k' ⟧ ⟹ k = k'" by blast+
  note [simp] = Γ.weight_minus_web[OF k]

  let ?fk = "plus_current (F f) k"
  have IN_fk: "d_IN ?fk x = d_IN (F f) x + d_IN k x" for x
    by(simp add: d_IN_def nn_integral_add)
  have OUT_fk: "d_OUT ?fk x = d_OUT (F f) x + d_OUT k x" for x
    by(simp add: d_OUT_def nn_integral_add)
  have fk: "current Γ ?fk" using f k by(rule current_plus_current_minus)

  have "d_OUT ?fk a ≥ weight Γ a"
  proof(rule ccontr)
    assume "¬ ?thesis"
    hence less: "d_OUT ?fk a < weight Γ a" by simp

    def Ω  "Γ ⊖ F f ⊖ k"
    have B_Ω [simp]: "B Ω = B Γ" by(simp add: Ω_def)

    have loose: "loose Ω" unfolding Ω_def using unhindered k k_w maximal by(rule Γ.loose_minus_web)
    interpret Ω: countable_bipartite_web Ω using k unfolding Ω_def
      by(rule Γ.countable_bipartite_web_minus_web)

    have a_ℰ: "a ∈ TERΓ ⊖ F f k" using Field k k_w less by(rule a_TER)
    then have "weight Ω a = weight Γ a - d_OUT (F f) a"
      using a disjoint by(auto simp add: roofed_circ_def Ω_def SINK.simps)
    then have weight_a: "0 < weight Ω a" using less a_ℰ
      by(simp add: OUT_fk SINK.simps diff_gr0_ennreal)

    let ?P_y = "λy. y ∈ OUTΩ a ∧ weight Ω y > 0"
    def y  "Eps ?P_y"
    have "Ex ?P_y" using Field k k_w less unfolding Ω_def by(rule Ex_y)
    hence "?P_y y" unfolding y_def by(rule someI_ex)
    hence "y ∈ OUTΩ a" and weight_y: "weight Ω y > 0" by blast+
    then have y_B: "y ∈ B Ω" by(auto simp add: outgoing_def Ω_def dest: bipartite_E)

    def bound  "enn2real (min (weight Ω a) (weight Ω y))"
    have bound_pos: "bound > 0" using weight_y weight_a Ω.weight_finite
      by(cases "weight Ω a" "weight Ω y" rule: ennreal2_cases)
        (simp_all add: bound_def min_def split: if_split_asm)

    let ?P_δ = "λδ. δ > 0 ∧ δ < bound ∧ ¬ hindered (reduce_weight Ω y δ)"
    def δ  "Eps ?P_δ"
    from Ω.unhinder[OF loose _ weight_y bound_pos] y_B disjoint have "Ex ?P_δ" by(auto simp add: Ω_def)
    hence "?P_δ δ" unfolding δ_def by(rule someI_ex)
    hence δ_pos: "0 < δ" by blast+

    let ?f' = "(plus_current (fst f) (zero_current((a, y) := δ)), plus_current (snd f) k)"
    have sat: "?f' = sat f" using less by(simp add: sat_def k_def Ω_def Let_def y_def bound_def δ_def split_def)
    also have "… = f" unfolding f_def using fixp_above_unfold[OF zero] by simp
    finally have "fst ?f' (a, y) = fst f (a, y)" by simp
    hence "δ = 0" using currentD_finite[OF ε_curr[OF Field']] δ_pos
      by(cases "fst f (a, y)") simp_all
    with δ_pos show False by simp
  qed

  with currentD_weight_OUT[OF fk, of a] have "d_OUT ?fk a = weight Γ a" by simp
  moreover have "current Γ ?fk" using f k by(rule current_plus_current_minus)
  moreover have "¬ hindered (Γ ⊖ ?fk)" unfolding minus_plus_current[OF f k]
    using unhindered k k_w by(rule Γ.unhindered_minus_web)
  ultimately show ?thesis by blast
qed

end

subsection ‹Linkability of unhindered bipartite webs›

context countable_bipartite_web begin

theorem unhindered_linkable:
  assumes unhindered: "¬ hindered Γ"
  shows "linkable Γ"
proof(cases "A Γ = {}")
  case True
  thus ?thesis by(auto intro!: exI[where x="zero_current"] linkage.intros simp add: web_flow_iff )
next
  case nempty: False

  let ?P = "λf a f'. current (Γ ⊖ f) f' ∧ d_OUT f' a = weight (Γ ⊖ f) a ∧ ¬ hindered (Γ ⊖ f ⊖ f')"

  def enum  "from_nat_into (A Γ)"
  have enum_A: "enum n ∈ A Γ" for n using from_nat_into[OF nempty, of n] by(simp add: enum_def)
  have vertex_enum [simp]: "vertex Γ (enum n)" for n using enum_A[of n] A_vertex by blast

  def f  "rec_nat zero_current (λn f. let f' = SOME f'. ?P f (enum n) f' in plus_current f f')"
  have f_0 [simp]: "f 0 = zero_current" by(simp add: f_def)
  have f_Suc: "f (Suc n) = plus_current (f n) (Eps (?P (f n) (enum n)))" for n by(simp add: f_def)

  have f: "current Γ (f n)"
    and sat: "⋀m. m < n ⟹ d_OUT (f n) (enum m) = weight Γ (enum m)"
    and unhindered: "¬ hindered (Γ ⊖ f n)" for n
  proof(induction n)
    case 0
    { case 1 thus ?case by(simp add: ) }
    { case 2 thus ?case by simp }
    { case 3 thus ?case using unhindered by simp }
  next
    case (Suc n)
    interpret Γ: countable_bipartite_web "Γ ⊖ f n" using Suc.IH(1) by(rule countable_bipartite_web_minus_web)

    def f'  "Eps (?P (f n) (enum n))"
    have "Ex (?P (f n) (enum n))" using Suc.IH(3) by(rule Γ.unhindered_saturate1)(simp add: enum_A)
    hence "?P (f n) (enum n) f'" unfolding f'_def by(rule someI_ex)
    hence f': "current (Γ ⊖ f n) f'"
      and OUT: "d_OUT f' (enum n) = weight (Γ ⊖ f n) (enum n)"
      and unhindered': "¬ hindered (Γ ⊖ f n ⊖ f')" by blast+
    have f_Suc: "f (Suc n) = plus_current (f n) f'" by(simp add: f'_def f_Suc)
    { case 1 show ?case unfolding f_Suc using Suc.IH(1) f' by(rule current_plus_current_minus) }
    note f'' = this
    { case (2 m)
      have "d_OUT (f (Suc n)) (enum m) ≤ weight Γ (enum m)" using f'' by(rule currentD_weight_OUT)
      moreover have "weight Γ (enum m) ≤ d_OUT (f (Suc n)) (enum m)"
      proof(cases "m = n")
        case True
        then show ?thesis unfolding f_Suc using OUT True
          by(simp add: d_OUT_def nn_integral_add enum_A add_diff_self_ennreal less_imp_le)
      next
        case False
        hence "m < n" using 2 by simp
        thus ?thesis using Suc.IH(2)[OF ‹m < n›] unfolding f_Suc
          by(simp add: d_OUT_def nn_integral_add add_increasing2 )
      qed
      ultimately show ?case by(rule antisym) }
    { case 3 show ?case unfolding f_Suc minus_plus_current[OF Suc.IH(1) f'] by(rule unhindered') }
  qed
  interpret Γ: countable_bipartite_web "Γ ⊖ f n" for n using f by(rule countable_bipartite_web_minus_web)

  have Ex_P: "Ex (?P (f n) (enum n))" for n using unhindered by(rule Γ.unhindered_saturate1)(simp add: enum_A)
  have f_mono: "f n ≤ f (Suc n)" for n using someI_ex[OF Ex_P, of n]
    by(auto simp add: le_fun_def f_Suc enum_A intro: add_increasing2 dest: )
  hence incseq: "incseq f" by(rule incseq_SucI)
  hence chain: "Complete_Partial_Order.chain op ≤ (range f)" by(rule incseq_chain_range)

  def g  "Sup (range f)"
  have "support_flow g ⊆ E"
    by(auto simp add: g_def support_flow.simps currentD_outside[OF f] elim: contrapos_pp)
  then have countable_g: "countable (support_flow g)" by(rule countable_subset) simp
  with chain _ _ have g: "current Γ g" unfolding g_def  by(rule current_Sup)(auto simp add: f)
  moreover
  have "d_OUT g x = weight Γ x" if "x ∈ A Γ" for x
  proof(rule antisym)
    show "d_OUT g x ≤ weight Γ x" using g by(rule currentD_weight_OUT)
    have "countable (A Γ)" using A_vertex by(rule countable_subset) simp
    from that subset_range_from_nat_into[OF this] obtain n where "x = enum n" unfolding enum_def by blast
    with sat[of n "Suc n"] have "d_OUT (f (Suc n)) x ≥ weight Γ x" by simp
    then show "weight Γ x ≤ d_OUT g x" using countable_g unfolding g_def
      by(subst d_OUT_Sup[OF chain])(auto intro: SUP_upper2)
  qed
  ultimately show ?thesis by(auto simp add: web_flow_iff linkage.simps)
qed

end

section ‹The Max-Flow Min-Cut theorem›

subsection ‹Reduction to bipartite webs›

definition bipartite_web_of :: "('v, 'more) web_scheme ⇒ ('v + 'v, 'more) web_scheme"
where
  "bipartite_web_of Γ =
  ⦇edge = λuv uv'. case (uv, uv') of (Inl u, Inr v) ⇒ edge Γ u v ∨ u = v ∧ u ∈ vertices Γ ∧ u ∉ A Γ ∧ v ∉ B Γ | _ ⇒ False,
   weight = λuv. case uv of Inl u ⇒ if u ∈ B Γ then 0 else weight Γ u | Inr u ⇒ if u ∈ A Γ then 0 else weight Γ u,
   A = Inl ` (vertices Γ - B Γ),
   B = Inr ` (- A Γ),
   … = web.more Γ⦈"

lemma bipartite_web_of_sel [simp]: fixes Γ (structure) shows
  "edge (bipartite_web_of Γ) (Inl u) (Inr v) ⟷ edge Γ u v ∨ u = v ∧ u ∈ V ∧ u ∉ A Γ ∧ v ∉ B Γ"
  "edge (bipartite_web_of Γ) uv (Inl u) ⟷ False"
  "edge (bipartite_web_of Γ) (Inr v) uv ⟷ False"
  "weight (bipartite_web_of Γ) (Inl u) = (if u ∈ B Γ then 0 else weight Γ u)"
  "weight (bipartite_web_of Γ) (Inr v) = (if v ∈ A Γ then 0 else weight Γ v)"
  "A (bipartite_web_of Γ) = Inl ` (V - B Γ)"
  "B (bipartite_web_of Γ) = Inr ` (- A Γ)"
by(simp_all add: bipartite_web_of_def split: sum.split)

lemma edge_bipartite_webI1: "edge Γ u v ⟹ edge (bipartite_web_of Γ) (Inl u) (Inr v)"
by(auto)

lemma edge_bipartite_webI2:
  "⟦ u ∈ VΓ; u ∉ A Γ; u ∉ B Γ ⟧ ⟹ edge (bipartite_web_of Γ) (Inl u) (Inr u)"
by(auto)

lemma edge_bipartite_webE:
  fixes Γ (structure)
  assumes "edge (bipartite_web_of Γ) uv uv'"
  obtains u v where "uv = Inl u" "uv' = Inr v" "edge Γ u v"
    | u where "uv = Inl u" "uv' = Inr u" "u ∈ V" "u ∉ A Γ" "u ∉ B Γ"
using assms by(cases uv uv' rule: sum.exhaust[case_product sum.exhaust]) auto

lemma E_bipartite_web:
  fixes Γ (structure) shows
  "Ebipartite_web_of Γ = (λ(x, y). (Inl x, Inr y)) ` E ∪ (λx. (Inl x, Inr x)) ` (V - A Γ - B Γ)"
by(auto elim: edge_bipartite_webE)

context web begin

lemma vertex_bipartite_web [simp]:
  "vertex (bipartite_web_of Γ) (Inl x) ⟷ vertex Γ x ∧ x ∉ B Γ"
  "vertex (bipartite_web_of Γ) (Inr x) ⟷ vertex Γ x ∧ x ∉ A Γ"
by(auto 4 4 simp add: vertex_def dest: B_out A_in intro: edge_bipartite_webI1 edge_bipartite_webI2 elim!: edge_bipartite_webE)

definition separating_of_bipartite :: "('v + 'v) set ⇒ 'v set"
where
  "separating_of_bipartite S =
  (let A_S = Inl -` S; B_S = Inr -` S in (A_S ∩ B_S) ∪ (A Γ ∩ A_S) ∪ (B Γ ∩ B_S))"

context
  fixes S :: "('v + 'v) set"
  assumes sep: "separating (bipartite_web_of Γ) S"
begin

text ‹Proof of separation follows @{cite Aharoni1983EJC}›

lemma separating_of_bipartite_aux:
  assumes p: "path Γ x p y" and y: "y ∈ B Γ"
  and x: "x ∈ A Γ ∨ Inr x ∈ S"
  shows "(∃z∈set p. z ∈ separating_of_bipartite S) ∨ x ∈ separating_of_bipartite S"
proof(cases "p = []")
  case True
  with p have "x = y" by cases auto
  with y x have "x ∈ separating_of_bipartite S" using disjoint
    by(auto simp add: separating_of_bipartite_def Let_def)
  thus ?thesis ..
next
  case nNil: False
  def inmarked  "λx. x ∈ A Γ ∨ Inr x ∈ S"
  def outmarked  "λx. x ∈ B Γ ∨ Inl x ∈ S"
  let  = "bipartite_web_of Γ"
  let ?double = "λx. inmarked x ∧ outmarked x"
  def tailmarked  "λ(x, y :: 'v). outmarked x"
  def headmarked  "λ(x :: 'v, y). inmarked y"

  have marked_E: "tailmarked e ∨ headmarked e" if "e ∈ E" for e -- ‹Lemma 1b›
  proof(cases e)
    case (Pair x y)
    with that have "path ?Γ (Inl x) [Inr y] (Inr y)" by(auto intro!: rtrancl_path.intros)
    from separatingD[OF sep this] that Pair show ?thesis
      by(fastforce simp add: vertex_def inmarked_def outmarked_def tailmarked_def headmarked_def)
  qed

  have "∃z∈set (x # p). ?double z" -- ‹Lemma 2›
  proof -
    have "inmarked ((x # p) ! (i + 1)) ∨ outmarked ((x # p) ! i)" if "i < length p" for i
      using rtrancl_path_nth[OF p that] marked_E[of "((x # p) ! i, p ! i)"] that
      by(auto simp add: tailmarked_def headmarked_def nth_Cons split: nat.split)
    hence "{i. i < length p ∧ inmarked (p ! i)} ∪ {i. i < length (x # butlast p) ∧ outmarked ((x # butlast p) ! i)} = {i. i < length p}"
      (is "?in ∪ ?out = _") using nNil
      by(force simp add: nth_Cons' nth_butlast elim: meta_allE[where x=0] cong del: old.nat.case_cong_weak)
    hence "length p + 2 = card (?in ∪ ?out) + 2" by simp
    also have "… ≤ (card ?in + 1) + (card ?out + 1)" by(simp add: card_Un_le)
    also have "card ?in = card ((λi. Inl (i + 1) :: _ + nat) ` ?in)"
      by(rule card_image[symmetric])(simp add: inj_on_def)
    also have "… + 1 = card (insert (Inl 0) {Inl (Suc i) :: _ + nat|i. i < length p ∧ inmarked (p ! i)})"
      by(subst card_insert_if)(auto intro!: arg_cong[where f=card])
    also have "… = card {Inl i :: nat + nat|i. i < length (x # p) ∧ inmarked ((x # p) ! i)}" (is "_ = card ?in")
      using x by(intro arg_cong[where f=card])(auto simp add: nth_Cons inmarked_def split: nat.split_asm)
    also have "card ?out = card ((Inr :: _ ⇒ nat + _) ` ?out)" by(simp add: card_image)
    also have "… + 1 = card (insert (Inr (length p)) {Inr i :: nat + _|i. i < length p ∧ outmarked ((x # p) ! i)})"
      using nNil by(subst card_insert_if)(auto intro!: arg_cong[where f=card] simp add: nth_Cons nth_butlast cong: nat.case_cong)
    also have "… = card {Inr i :: nat + _|i. i < length (x # p) ∧ outmarked ((x # p) ! i)}" (is "_ = card ?out")
      using nNil rtrancl_path_last[OF p nNil] y
      by(intro arg_cong[where f=card])(auto simp add: outmarked_def last_conv_nth)
    also have "card ?in + card ?out = card (?in ∪ ?out)"
      by(rule card_Un_disjoint[symmetric]) auto
    also let ?f = "case_sum id id"
    have "?f ` (?in ∪ ?out) ⊆ {i. i < length (x # p)}" by auto
    from card_mono[OF _ this] have "card (?f ` (?in ∪ ?out)) ≤ length p + 1" by simp
    ultimately have "¬ inj_on ?f (?in ∪ ?out)" by(intro pigeonhole) simp
    then obtain i where "i < length (x # p)" "?double ((x # p) ! i)"
      by(auto simp add: inj_on_def)
    thus ?thesis by(auto simp add: set_conv_nth)
  qed
  moreover have "z ∈ separating_of_bipartite S" if "?double z" for z using that disjoint
    by(auto simp add: separating_of_bipartite_def Let_def inmarked_def outmarked_def)
  ultimately show ?thesis by auto
qed

lemma separating_of_bipartite:
  "separating Γ (separating_of_bipartite S)"
by(rule separating_gen.intros)(erule (1) separating_of_bipartite_aux; simp)

end

lemma current_bipartite_web_finite:
  assumes f: "current (bipartite_web_of Γ) f" (is "current ?Γ _")
  shows "f e ≠ ⊤"
proof(cases e)
  case (Pair x y)
  have "f e ≤ d_OUT f x" unfolding Pair d_OUT_def by(rule nn_integral_ge_point) simp
  also have "… ≤ weight ?Γ x" by(rule currentD_weight_OUT[OF f])
  also have "… < ⊤" by(cases x)(simp_all add: less_top[symmetric])
  finally show ?thesis by simp
qed

definition current_of_bipartite :: "('v + 'v) current ⇒ 'v current"
where "current_of_bipartite f = (λ(x, y). f (Inl x, Inr y) * indicator E (x, y))"

lemma current_of_bipartite_simps [simp]: "current_of_bipartite f (x, y) = f (Inl x, Inr y) * indicator E (x, y)"
by(simp add: current_of_bipartite_def)

lemma d_OUT_current_of_bipartite:
  assumes f: "current (bipartite_web_of Γ) f"
  shows "d_OUT (current_of_bipartite f) x = d_OUT f (Inl x) - f (Inl x, Inr x)"
proof -
  have "d_OUT (current_of_bipartite f) x = ∫+ y. f (Inl x, y) * indicator E (x, projr y) ∂count_space (range Inr)"
    by(simp add: d_OUT_def nn_integral_count_space_reindex)
  also have "… = d_OUT f (Inl x) - ∫+ y. f (Inl x, y) * indicator {Inr x} y ∂count_space UNIV" (is "_ = _ - ?rest")
    unfolding d_OUT_def by(subst nn_integral_diff[symmetric])(auto 4 4 simp add: current_bipartite_web_finite[OF f] AE_count_space nn_integral_count_space_indicator no_loop split: split_indicator intro!: nn_integral_cong intro: currentD_outside[OF f] elim: edge_bipartite_webE)
  finally show ?thesis by simp
qed

lemma d_IN_current_of_bipartite:
  assumes f: "current (bipartite_web_of Γ) f"
  shows "d_IN (current_of_bipartite f) x = d_IN f (Inr x) - f (Inl x, Inr x)"
proof -
  have "d_IN (current_of_bipartite f) x = ∫+ y. f (y, Inr x) * indicator E (projl y, x) ∂count_space (range Inl)"
    by(simp add: d_IN_def nn_integral_count_space_reindex)
  also have "… = d_IN f (Inr x) - ∫+ y. f (y, Inr x) * indicator {Inl x} y ∂count_space UNIV" (is "_ = _ - ?rest")
    unfolding d_IN_def by(subst nn_integral_diff[symmetric])(auto 4 4 simp add: current_bipartite_web_finite[OF f] AE_count_space nn_integral_count_space_indicator no_loop split: split_indicator intro!: nn_integral_cong intro: currentD_outside[OF f] elim: edge_bipartite_webE)
  finally show ?thesis by simp
qed

lemma current_current_of_bipartite: -- ‹Lemma 6.3›
  assumes f: "current (bipartite_web_of Γ) f" (is "current ?Γ _")
  and w: "wave (bipartite_web_of Γ) f"
  shows "current Γ (current_of_bipartite f)" (is "current _ ?f")
proof
  fix x
  have "d_OUT ?f x ≤ d_OUT f (Inl x)"
    by(simp add: d_OUT_current_of_bipartite[OF f] diff_le_self_ennreal)
  also have "… ≤ weight Γ x"
    using currentD_weight_OUT[OF f, of "Inl x"]
    by(simp split: if_split_asm)
  finally show "d_OUT ?f x ≤ weight Γ x" .
next
  fix x
  have "d_IN ?f x ≤ d_IN f (Inr x)"
    by(simp add: d_IN_current_of_bipartite[OF f] diff_le_self_ennreal)
  also have "… ≤ weight Γ x"
    using currentD_weight_IN[OF f, of "Inr x"]
    by(simp split: if_split_asm)
  finally show "d_IN ?f x ≤ weight Γ x" .
next
  have OUT: "d_OUT ?f b = 0" if "b ∈ B Γ" for b using that
    by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 intro!: currentD_outside[OF f] dest: B_out)
  show "d_OUT ?f x ≤ d_IN ?f x" if A: "x ∉ A Γ" for x
  proof(cases "x ∈ B Γ ∨ x ∉ V")
    case True
    then show ?thesis
    proof
      assume "x ∈ B Γ"
      with OUT[OF this] show ?thesis by auto
    next
      assume "x ∉ V"
      hence "d_OUT ?f x = 0" by(auto simp add: d_OUT_def vertex_def nn_integral_0_iff emeasure_count_space_eq_0 intro!: currentD_outside[OF f])
      thus ?thesis by simp
    qed
  next
    case B [simplified]: False
    have "d_OUT ?f x = d_OUT f (Inl x) - f (Inl x, Inr x)" (is "_ = _ - ?rest")
      by(simp add: d_OUT_current_of_bipartite[OF f])
    also have "d_OUT f (Inl x) ≤ d_IN f (Inr x)"
    proof(rule ccontr)
      assume "¬ ?thesis"
      hence *: "d_IN f (Inr x) < d_OUT f (Inl x)" by(simp add: not_less)
      also have "… ≤ weight Γ x" using currentD_weight_OUT[OF f, of "Inl x"] B by simp
      finally have "Inr x ∉ TER f" using A by(auto elim!: SAT.cases)
      moreover have "Inl x ∉ TER f" using * by(auto simp add: SINK.simps)
      moreover have "path ?Γ (Inl x) [Inr x] (Inr x)"
        by(rule rtrancl_path.step)(auto intro!: rtrancl_path.base simp add: no_loop A B)
      ultimately show False using waveD_separating[OF w] A B by(auto dest!: separatingD)
    qed
    hence "d_OUT f (Inl x) - ?rest ≤ d_IN f (Inr x) - ?rest" by(rule ennreal_minus_mono) simp
    also have "… =  d_IN ?f x" by(simp add: d_IN_current_of_bipartite[OF f])
    finally show ?thesis .
  qed
  show "?f e = 0" if "e ∉ E" for e using that by(cases e)(auto)
qed

lemma TER_current_of_bipartite: -- ‹Lemma 6.3›
  assumes f: "current (bipartite_web_of Γ) f" (is "current ?Γ _")
  and w: "wave (bipartite_web_of Γ) f"
  shows "TER (current_of_bipartite f) = separating_of_bipartite (TERbipartite_web_of Γ f)"
    (is "TER ?f = separating_of_bipartite ?TER")
proof(rule set_eqI)
  fix x
  consider (A) "x ∈ A Γ" "x ∈ V" | (B) "x ∈ B Γ" "x ∈ V"
    | (inner) "x ∉ A Γ" "x ∉ B Γ" "x ∈ V" | (outside) "x ∉ V" by auto
  thus "x ∈ TER ?f ⟷ x ∈ separating_of_bipartite ?TER"
  proof cases
    case A
    hence "d_OUT ?f x = d_OUT f (Inl x)" using currentD_outside[OF f, of "Inl x" "Inr x"]
      by(simp add: d_OUT_current_of_bipartite[OF f] no_loop)
    thus ?thesis using A disjoint
      by(auto simp add: separating_of_bipartite_def Let_def SINK.simps intro!: SAT.A imageI)
  next
    case B
    then have "d_IN ?f x = d_IN f (Inr x)"
      using currentD_outside[OF f, of "Inl x" "Inr x"]
      by(simp add: d_IN_current_of_bipartite[OF f] no_loop)
    moreover have "d_OUT ?f x = 0" using B currentD_outside[OF f, of "Inl x" "Inr x"]
      by(simp add: d_OUT_current_of_bipartite[OF f] no_loop)(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 intro!: currentD_outside[OF f] elim!: edge_bipartite_webE dest: B_out)
    moreover have "d_OUT f (Inr x) = 0" using B disjoint by(intro currentD_OUT[OF f]) auto
    ultimately show ?thesis using B
      by(auto simp add: separating_of_bipartite_def Let_def SINK.simps SAT.simps)
  next
    case outside
    with current_current_of_bipartite[OF f w] have "d_OUT ?f x = 0" "d_IN ?f x = 0"
      by(rule currentD_outside_OUT currentD_outside_IN)+
    moreover from outside have "Inl x ∉ vertices ?Γ" "Inr x ∉ vertices ?Γ" by auto
    hence "d_OUT f (Inl x) = 0" "d_IN f (Inl x) = 0" "d_OUT f (Inr x) = 0" "d_IN f (Inr x) = 0"
      by(blast intro: currentD_outside_OUT[OF f] currentD_outside_IN[OF f])+
    ultimately show ?thesis using outside weight_outside[of x]
      by(auto simp add: separating_of_bipartite_def Let_def SINK.simps SAT.simps not_le)
  next
    case inner
    show ?thesis
    proof
      assume "x ∈ separating_of_bipartite ?TER"
      with inner have l: "Inl x ∈ ?TER" and r: "Inr x ∈ ?TER"
        by(auto simp add: separating_of_bipartite_def Let_def)
      have "f (Inl x, Inr x) ≤ d_OUT f (Inl x)"
        unfolding d_OUT_def by(rule nn_integral_ge_point) simp
      with l have 0: "f (Inl x, Inr x) = 0"
        by(auto simp add: SINK.simps)
      with l have "x ∈ SINK ?f"  by(simp add: SINK.simps d_OUT_current_of_bipartite[OF f])
      moreover from r have "Inr x ∈ SAT ?Γ f" by auto
      with inner have "x ∈ SAT Γ ?f"
        by(auto elim!: SAT.cases intro!: SAT.IN simp add: 0 d_IN_current_of_bipartite[OF f])
      ultimately show "x ∈ TER ?f" by simp
    next
      assume *: "x ∈ TER ?f"
      have "d_IN f (Inr x) ≤ weight ?Γ (Inr x)" using f by(rule currentD_weight_IN)
      also have "… ≤ weight Γ x" using inner by simp
      also have "… ≤ d_IN ?f x" using inner * by(auto elim: SAT.cases)
      also have "… ≤ d_IN f (Inr x)"
        by(simp add: d_IN_current_of_bipartite[OF f] max_def diff_le_self_ennreal)
      ultimately have eq: "d_IN ?f x = d_IN f (Inr x)" by simp
      hence 0: "f (Inl x, Inr x) = 0"
        using ennreal_minus_cancel_iff[of "d_IN f (Inr x)" "f (Inl x, Inr x)" 0] currentD_weight_IN[OF f, of "Inr x"] inner
          d_IN_ge_point[of f "Inl x" "Inr x"]
        by(auto simp add: d_IN_current_of_bipartite[OF f] top_unique)
      have "Inl x ∈ ?TER" "Inr x ∈ ?TER" using inner * currentD_OUT[OF f, of "Inr x"]
        by(auto simp add: SAT.simps SINK.simps d_OUT_current_of_bipartite[OF f] 0 eq)
      thus "x ∈ separating_of_bipartite ?TER" unfolding separating_of_bipartite_def Let_def by blast
    qed
  qed
qed

lemma wave_current_of_bipartite: -- ‹Lemma 6.3›
  assumes f: "current (bipartite_web_of Γ) f" (is "current ?Γ _")
  and w: "wave (bipartite_web_of Γ) f"
  shows "wave Γ (current_of_bipartite f)" (is "wave _ ?f")
proof
  have sep': "separating Γ (separating_of_bipartite (TER f))"
    by(rule separating_of_bipartite)(rule waveD_separating[OF w])
  then show sep: "separating Γ (TER (current_of_bipartite f))"
    by(simp add: TER_current_of_bipartite[OF f w])

  fix x
  assume "x ∉ RF (TER ?f)"
  then obtain p y where p: "path Γ x p y" and y: "y ∈ B Γ" and x: "x ∉ TER ?f"
    and bypass: "⋀z. z ∈ set p ⟹ z ∉ TER ?f"
    by(auto simp add: roofed_def elim: rtrancl_path_distinct)
  from p obtain p' where p': "path Γ x p' y" and distinct: "distinct (x # p')"
    and subset: "set p' ⊆ set p" by(auto elim: rtrancl_path_distinct)
  consider (outside) "x ∉ V" | (A) "x ∈ A Γ" | (B) "x ∈ B Γ" | (inner) "x ∈ V" "x ∉ A Γ" "x ∉ B Γ" by auto
  then show "d_OUT ?f x = 0"
  proof cases
    case outside
    with f w show ?thesis by(rule currentD_outside_OUT[OF current_current_of_bipartite])
  next
    case A
    from separatingD[OF sep p A y] bypass have "x ∈ TER ?f" by blast
    thus ?thesis by(simp add: SINK.simps)
  next
    case B
    with f w show ?thesis by(rule currentD_OUT[OF current_current_of_bipartite])
  next
    case inner
    hence "path ?Γ (Inl x) [Inr x] (Inr x)" by(auto intro!: rtrancl_path.intros)
    from inner waveD_separating[OF w, THEN separatingD, OF this]
    consider (Inl) "Inl x ∈ TER f" | (Inr) "Inr x ∈ TER f" by auto
    then show ?thesis
    proof cases
      case Inl
      thus ?thesis
        by(auto simp add: SINK.simps d_OUT_current_of_bipartite[OF f] max_def)
    next
      case Inr
      with separating_of_bipartite_aux[OF waveD_separating[OF w] p y] x bypass
      have False unfolding TER_current_of_bipartite[OF f w] by blast
      thus ?thesis ..
    qed
  qed
qed

end

context countable_web begin

lemma countable_bipartite_web_of: "countable_bipartite_web (bipartite_web_of Γ)" (is "countable_bipartite_web ?Γ")
proof
  show "V ⊆ A ?Γ ∪ B ?Γ"
    apply(rule subsetI)
    subgoal for x by(cases x) auto
    done
  show "A ?Γ ⊆ V" by auto
  show "x ∈ A ?Γ ∧ y ∈ B ?Γ" if "edge ?Γ x y" for x y using that
    by(cases x y rule: sum.exhaust[case_product sum.exhaust])(auto simp add: inj_image_mem_iff vertex_def B_out A_in)
  show "A ?Γ ∩ B ?Γ = {}" by auto
  show "countable E" by(simp add: E_bipartite_web)
  show "weight ?Γ x ≠ ⊤" for x by(cases x) simp_all
  show "weight (bipartite_web_of Γ) x = 0" if "x ∉ V" for x using that
    by(cases x)(auto simp add: weight_outside)
qed

end

context web begin

lemma unhindered_bipartite_web_of:
  assumes loose: "loose Γ"
  shows "¬ hindered (bipartite_web_of Γ)"
proof
  assume "hindered (bipartite_web_of Γ)" (is "hindered ?Γ")
  then obtain f where f: "current ?Γ f" and w: "wave ?Γ f" and hind: "hindrance ?Γ f" by cases
  from f w have "current Γ (current_of_bipartite f)" by(rule current_current_of_bipartite)
  moreover from f w have "wave Γ (current_of_bipartite f)" by(rule wave_current_of_bipartite)
  ultimately have *: "current_of_bipartite f = zero_current" by(rule looseD_wave[OF loose])
  have zero: "f (Inl x, Inr y) = 0" if "x ≠ y" for x y using that *[THEN fun_cong, of "(x, y)"]
     by(cases "edge Γ x y")(auto intro: currentD_outside[OF f])

  have OUT: "d_OUT f (Inl x) = f (Inl x, Inr x)" for x
  proof -
    have "d_OUT f (Inl x) = (∑+ y. f (Inl x, y) * indicator {Inr x} y)" unfolding d_OUT_def
      using zero currentD_outside[OF f]
      apply(intro nn_integral_cong)
      subgoal for y by(cases y)(auto split: split_indicator)
      done
    also have "… = f (Inl x, Inr x)" by simp
    finally show ?thesis .
  qed
  have IN: "d_IN f (Inr x) = f (Inl x, Inr x)" for x
  proof -
    have "d_IN f (Inr x) = (∑+ y. f (y, Inr x) * indicator {Inl x} y)" unfolding d_IN_def
      using zero currentD_outside[OF f]
      apply(intro nn_integral_cong)
      subgoal for y by(cases y)(auto split: split_indicator)
      done
    also have "… = f (Inl x, Inr x)" by simp
    finally show ?thesis .
  qed

  let ?TER = "TER f"
  from hind obtain a where a: "a ∈ A ?Γ" and nℰ: "a ∉ ℰ (TER f)"
    and OUT_a: "d_OUT f a < weight ?Γ a" by cases
  from a obtain a' where a': "a = Inl a'" and v: "vertex Γ a'" and b: "a' ∉ B Γ" by auto
  have A: "a' ∈ A Γ"
  proof(rule ccontr)
    assume A: "a' ∉ A Γ"
    hence "edge ?Γ (Inl a') (Inr a')" using v b by simp
    hence p: "path ?Γ (Inl a') [Inr a'] (Inr a')" by(simp add: rtrancl_path_simps)
    from separatingD[OF waveD_separating[OF w] this] b v A
    have "Inl a' ∈ ?TER ∨ Inr a' ∈ ?TER" by auto
    thus False
    proof cases
      case left
      hence "d_OUT f (Inl a') = 0" by(simp add: SINK.simps)
      moreover hence "d_IN f (Inr a') = 0" by(simp add: OUT IN)
      ultimately have "Inr a' ∉ ?TER" using v b A OUT_a a' by(auto simp add: SAT.simps)
      then have "essential ?Γ (B ?Γ) ?TER (Inl a')" using A
        by(intro essentialI[OF p]) simp_all
      with nℰ left a' show False by simp
    next
      case right
      hence "d_IN f (Inr a') = weight Γ a'" using A by(auto simp add: currentD_SAT[OF f])
      hence "d_OUT f (Inl a') = weight Γ a'" by(simp add: IN OUT)
      with OUT_a a' b show False by simp
    qed
  qed
  moreover

  from A have "d_OUT f (Inl a') = 0" using currentD_outside[OF f, of "Inl a'" "Inr a'"]
    by(simp add: OUT no_loop)
  with b v have TER: "Inl a' ∈ ?TER" by(simp add: SAT.A SINK.simps)
  with nℰ a' have ness: "¬ essential ?Γ (B ?Γ) ?TER (Inl a')" by simp

  have "a' ∉ ℰ (TER zero_current)"
  proof
    assume "a' ∈ ℰ (TER zero_current)"
    then obtain p y where p: "path Γ a' p y" and y: "y ∈ B Γ"
      and bypass: "⋀z. z ∈ set p ⟹ z ∉ TER zero_current"
      by(rule ℰ_E_RF)(auto intro: roofed_greaterI)

    from p show False
    proof cases
      case base with y A disjoint show False by auto
    next
      case (step x p')
      from step(2) have "path ?Γ (Inl a') [Inr x] (Inr x)" by(simp add: rtrancl_path_simps)
      from not_essentialD[OF ness this] bypass[of x] step(1)
      have "Inr x ∈ ?TER" by simp
      with bypass[of x] step(1) have "d_IN f (Inr x) > 0"
        by(auto simp add: currentD_SAT[OF f] zero_less_iff_neq_zero)
      hence x: "Inl x ∉ ?TER" by(auto simp add: SINK.simps OUT IN)
      from step(1) have "set (x # p') ⊆ set p" by auto
      with ‹path Γ x p' y› x y show False
      proof induction
        case (base x)
        thus False using currentD_outside_IN[OF f, of "Inl x"] currentD_outside_OUT[OF f, of "Inl x"]
          by(auto simp add: currentD_SAT[OF f] SINK.simps dest!: bypass)
      next
        case (step x z p' y)
        from step.prems(3) bypass[of x] weight_outside[of x] have x: "vertex Γ x" by(auto)
        from ‹edge Γ x z› have "path ?Γ (Inl x) [Inr z] (Inr z)" by(simp add: rtrancl_path_simps)
        from separatingD[OF waveD_separating[OF w] this] step.prems(1) step.prems(3) bypass[of z] x ‹edge Γ x z›
        have "Inr z ∈ ?TER" by(force simp add: B_out inj_image_mem_iff)
        with bypass[of z] step.prems(3) ‹edge Γ x z› have "d_IN f (Inr z) > 0"
          by(auto simp add: currentD_SAT[OF f] A_in zero_less_iff_neq_zero)
        hence x: "Inl z ∉ ?TER" by(auto simp add: SINK.simps OUT IN)
        with step.IH[OF this] step.prems(2,3) show False by auto
      qed
    qed
  qed
  moreover have "d_OUT zero_current a' < weight Γ a'"
    using OUT_a a' b by (auto simp: zero_less_iff_neq_zero)
  ultimately have "hindrance Γ zero_current" by(rule hindrance)
  with looseD_hindrance[OF loose] show False by contradiction
qed

lemma (in -) divide_less_1_iff_ennreal: "a / b < (1::ennreal) ⟷ (0 < b ∧ a < b ∨ b = 0 ∧ a = 0 ∨ b = top)"
  by (cases a; cases b; cases "b = 0")
     (auto simp: divide_ennreal ennreal_less_iff ennreal_top_divide)

lemma linkable_bipartite_web_ofD:
  assumes link: "linkable (bipartite_web_of Γ)" (is "linkable ?Γ")
  and countable: "countable E"
  shows "linkable Γ"
proof -
  from link obtain f where wf: "web_flow ?Γ f" and link: "linkage ?Γ f" by blast
  from wf have f: "current ?Γ f" by(rule web_flowD_current)
  def f'  "current_of_bipartite f"

  have IN_le_OUT: "d_IN f' x ≤ d_OUT f' x" if "x ∉ B Γ" for x
  proof(cases "x ∈ V")
    case True
    have "d_IN f' x = d_IN f (Inr x) - f (Inl x, Inr x)" (is "_ = _ - ?rest")
      by(simp add: f'_def d_IN_current_of_bipartite[OF f])
    also have "… ≤ weight ?Γ (Inr x) - ?rest"
      using currentD_weight_IN[OF f, of "Inr x"] by(rule ennreal_minus_mono) simp
    also have "… ≤ weight ?Γ (Inl x) - ?rest" using that ennreal_minus_mono by(auto)
    also have "weight ?Γ (Inl x) = d_OUT f (Inl x)" using that linkageD[OF link, of "Inl x"] True by auto
    also have "… - ?rest = d_OUT f' x"
      by(simp add: f'_def d_OUT_current_of_bipartite[OF f])
    finally show ?thesis .
  next
    case False
    with currentD_outside_OUT[OF f, of "Inl x"] currentD_outside_IN[OF f, of "Inr x"]
    show ?thesis by(simp add: f'_def d_IN_current_of_bipartite[OF f] d_OUT_current_of_bipartite[OF f])
  qed
  have link: "linkage Γ f'"
  proof
    show "d_OUT f' a = weight Γ a" if "a ∈ A Γ" for a
    proof(cases "a ∈ V")
      case True
      from that have "a ∉ B Γ" using disjoint by auto
      with that True linkageD[OF link, of "Inl a"] ennreal_minus_cancel_iff[of _ _ 0] currentD_outside[OF f, of "Inl a" "Inr a"]
      show ?thesis by(clarsimp simp add: f'_def d_OUT_current_of_bipartite[OF f] max_def no_loop)
    next
      case False
      with weight_outside[OF this] currentD_outside[OF f, of "Inl a" "Inr a"] currentD_outside_OUT[OF f, of "Inl a"]
      show ?thesis by(simp add: f'_def d_OUT_current_of_bipartite[OF f] no_loop)
    qed
  qed

  def F  "{g. (∀e. 0 ≤ g e) ∧ (∀e. e ∉ E ⟶ g e = 0) ∧
    (∀x. x ∉ B Γ ⟶ d_IN g x ≤ d_OUT g x) ∧
    linkage Γ g ∧
    (∀x∈A Γ. d_IN g x = 0) ∧
    (∀x. d_OUT g x ≤ weight Γ x) ∧
    (∀x. d_IN g x ≤ weight Γ x) ∧
    (∀x∈B Γ. d_OUT g x = 0) ∧ g ≤ f'}"
  def leq  "restrict_rel F {(f, f'). f' ≤ f}"
  have F: "Field leq = F" by(auto simp add: leq_def)
  have F_I [intro?]: "f ∈ Field leq" if "⋀e. 0 ≤ f e" and "⋀e. e ∉ E ⟹ f e = 0"
    and "⋀x. x ∉ B Γ ⟹ d_IN f x ≤ d_OUT f x" and "linkage Γ f"
    and "⋀x. x ∈ A Γ ⟹ d_IN f x = 0" and "⋀x. d_OUT f x ≤ weight Γ x"
    and "⋀x. d_IN f x ≤ weight Γ x" and "⋀x. x ∈ B Γ ⟹ d_OUT f x = 0"
    and "f ≤ f'" for f using that by(simp add: F F_def)
  have F_nonneg: "0 ≤ f e" if "f ∈ Field leq" for f e using that by(cases e)(simp add: F F_def)
  have F_outside: "f e = 0" if "f ∈ Field leq" "e ∉ E" for f e using that by(cases e)(simp add: F F_def)
  have F_IN_OUT: "d_IN f x ≤ d_OUT f x" if "f ∈ Field leq" "x ∉ B Γ" for f x using that by(simp add: F F_def)
  have F_link: "linkage Γ f" if "f ∈ Field leq" for f using that by(simp add: F F_def)
  have F_IN: "d_IN f x = 0" if "f ∈ Field leq" "x ∈ A Γ" for f x using that by(simp add: F F_def)
  have F_OUT: "d_OUT f x = 0" if "f ∈ Field leq" "x ∈ B Γ" for f x using that by(simp add: F F_def)
  have F_weight_OUT: "d_OUT f x ≤ weight Γ x" if "f ∈ Field leq" for f x using that by(simp add: F F_def)
  have F_weight_IN: "d_IN f x ≤ weight Γ x" if "f ∈ Field leq" for f x using that by(simp add: F F_def)
  have F_le: "f e ≤ f' e" if "f ∈ Field leq" for f e using that by(cases e)(simp add: F F_def le_fun_def)

  have F_finite_OUT: "d_OUT f x ≠ ⊤" if "f ∈ Field leq" for f x
  proof -
    have "d_OUT f x ≤ weight Γ x" by(rule F_weight_OUT[OF that])
    also have "… < ⊤" by(simp add: less_top[symmetric])
    finally show ?thesis by simp
  qed

  have F_finite: "f e ≠ ⊤" if "f ∈ Field leq" for f e
  proof(cases e)
    case (Pair x y)
    have "f e ≤ d_OUT f x" unfolding Pair d_OUT_def by(rule nn_integral_ge_point) simp
    also have "… < ⊤" by(simp add: F_finite_OUT[OF that] less_top[symmetric])
    finally show ?thesis by simp
  qed

  have f': "f' ∈ Field leq"
  proof
    show "0 ≤ f' e" for e by(cases e)(simp add: f'_def)
    show "f' e = 0" if "e ∉ E" for e using that by(clarsimp split: split_indicator_asm simp add: f'_def)
    show "d_IN f' x ≤ d_OUT f' x" if "x ∉ B Γ" for x using that by(rule IN_le_OUT)
    show "linkage Γ f'" by(rule link)
    show "d_IN f' x = 0" if "x ∈ A Γ" for x using that currentD_IN[OF f, of "Inl x"] disjoint
      currentD_outside[OF f, of "Inl x" "Inr x"] currentD_outside_IN[OF f, of "Inr x"]
      by(cases "x ∈ V")(auto simp add: d_IN_current_of_bipartite[OF f] no_loop f'_def)
    show "d_OUT f' x = 0" if "x ∈ B Γ" for x using that currentD_OUT[OF f, of "Inr x"] disjoint
      currentD_outside[OF f, of "Inl x" "Inr x"] currentD_outside_OUT[OF f, of "Inl x"]
      by(cases "x ∈ V")(auto simp add: d_OUT_current_of_bipartite[OF f] no_loop f'_def)
    show "d_OUT f' x ≤ weight Γ x" for x using currentD_weight_OUT[OF f, of "Inl x"]
      by(simp add: d_OUT_current_of_bipartite[OF f] ennreal_diff_le_mono_left f'_def split: if_split_asm)
    show "d_IN f' x ≤ weight Γ x" for x using currentD_weight_IN[OF f, of "Inr x"]
      by(simp add: d_IN_current_of_bipartite[OF f] ennreal_diff_le_mono_left f'_def split: if_split_asm)
  qed simp

  have F_leI: "g ∈ Field leq" if f: "f ∈ Field leq" and le: "⋀e. g e ≤ f e"
    and nonneg: "⋀e. 0 ≤ g e" and IN_OUT: "⋀x. x ∉ B Γ ⟹ d_IN g x ≤ d_OUT g x"
    and link: "linkage Γ g"
    for f g
  proof
    show "g e = 0" if "e ∉ E" for e using nonneg[of e] F_outside[OF f that] le[of e] by simp
    show "d_IN g a = 0" if "a ∈ A Γ" for a
      using d_IN_mono[of g a f, OF le] F_IN[OF f that] by auto
    show "d_OUT g b = 0" if "b ∈ B Γ" for b
      using d_OUT_mono[of g b f, OF le] F_OUT[OF f that] by auto
    show "d_OUT g x ≤ weight Γ x" for x
      using d_OUT_mono[of g x f, OF le] F_weight_OUT[OF f] by(rule order_trans)
    show "d_IN g x ≤ weight Γ x" for x
      using d_IN_mono[of g x f, OF le] F_weight_IN[OF f] by(rule order_trans)
    show "g ≤ f'" using order_trans[OF le F_le[OF f]] by(auto simp add: le_fun_def)
  qed(blast intro: IN_OUT link nonneg)+

  have chain_Field: "Inf M ∈ F" if M: "M ∈ Chains leq" and nempty: "M ≠ {}" for M
  proof -
    from nempty obtain g0 where g0_in_M: "g0 ∈ M" by auto
    with M have g0: "g0 ∈ Field leq" by(rule Chains_FieldD)

    from M have "M ∈ Chains {(g, g'). g' ≤ g}"
      by(rule mono_Chains[THEN subsetD, rotated])(auto simp add: leq_def in_restrict_rel_iff)
    then have "Complete_Partial_Order.chain op ≥ M" by(rule Chains_into_chain)
    hence chain': "Complete_Partial_Order.chain op ≤ M" by(simp add: chain_dual)

    have "support_flow f' ⊆ E" using F_outside[OF f'] by(auto intro: ccontr simp add: support_flow.simps)
    then have countable': "countable (support_flow f')"
      by(rule countable_subset)(simp add: E_bipartite_web countable "V_def")

    have finite_OUT: "d_OUT f' x ≠ ⊤" for x using weight_finite[of x]
      by(rule neq_top_trans)(rule F_weight_OUT[OF f'])
    have finite_IN: "d_IN f' x ≠ ⊤" for x using weight_finite[of x]
      by(rule neq_top_trans)(rule F_weight_IN[OF f'])
    have OUT_M: "d_OUT (Inf M) x = (INF g:M. d_OUT g x)" for x using chain' nempty countable' _ finite_OUT
      by(rule d_OUT_Inf)(auto dest!: Chains_FieldD[OF M] simp add: leq_def F_nonneg F_le)
    have IN_M: "d_IN (Inf M) x = (INF g:M. d_IN g x)" for x using chain' nempty countable' _ finite_IN
      by(rule d_IN_Inf)(auto dest!: Chains_FieldD[OF M] simp add: leq_def F_nonneg F_le)

    show "Inf M ∈ F" using g0 unfolding F[symmetric]
    proof(rule F_leI)
      show "(Inf M) e ≤ g0 e" for e using g0_in_M by(auto intro: INF_lower)
      show "0 ≤ (Inf M) e" for e by(auto intro!: INF_greatest dest: F_nonneg dest!: Chains_FieldD[OF M])
      show "d_IN (Inf M) x ≤ d_OUT (Inf M) x" if "x ∉ B Γ" for x using that
        by(auto simp add: IN_M OUT_M intro!: INF_mono dest: Chains_FieldD[OF M] intro: F_IN_OUT[OF _ that])
      show "linkage Γ (Inf M)" using nempty
        by(simp add: linkage.simps OUT_M F_link[THEN linkageD] Chains_FieldD[OF M] cong: INF_cong)
    qed
  qed

  let ?P = "λg z. z ∉ A Γ ∧ z ∉ B Γ ∧ d_OUT g z > d_IN g z"

  def link 
   "λg. if ∃z. ?P g z then
          let z = SOME z. ?P g z; factor = d_IN g z / d_OUT g z
          in (λ(x, y). (if x = z then factor else 1) * g (x, y))
        else g"
  have increasing: "link g ≤ g ∧ link g ∈ Field leq" if g: "g ∈ Field leq" for g
  proof(cases "∃z. ?P g z")
    case False
    thus ?thesis using that by(auto simp add: link_def leq_def)
  next
    case True
    def z  "Eps (?P g)"
    from True have "?P g z" unfolding z_def by(rule someI_ex)
    hence A: "z ∉ A Γ" and B: "z ∉ B Γ" and less: "d_IN g z < d_OUT g z" by simp_all
    let ?factor = "d_IN g z / d_OUT g z"
    have link [simp]: "link g (x, y) = (if x = z then ?factor else 1) * g (x, y)" for x y
      using True by(auto simp add: link_def z_def Let_def)

    have "?factor ≤ 1" (is "?factor ≤ _") using less
      by(cases "d_OUT g z" "d_IN g z" rule: ennreal2_cases) (simp_all add: ennreal_less_iff divide_ennreal)
    hence le': "?factor * g (x, y) ≤ 1 * g (x, y)" for y x
      by(rule mult_right_mono)(simp add: F_nonneg[OF g])
    hence le: "link g e ≤ g e" for e by(cases e)simp
    have "link g ∈ Field leq" using g le
    proof(rule F_leI)
      show nonneg: "0 ≤ link g e" for e
        using F_nonneg[OF g, of e] by(cases e) simp
      show "linkage Γ (link g)" using that A F_link[OF g] by(clarsimp simp add: linkage.simps d_OUT_def)

      fix x
      assume x: "x ∉ B Γ"
      have "d_IN (link g) x ≤ d_IN g x" unfolding d_IN_def using le' by(auto intro: nn_integral_mono)
      also have "… ≤ d_OUT (link g) x"
      proof(cases "x = z")
        case True
        have "d_IN g x = d_OUT (link g) x" unfolding d_OUT_def
          using True F_weight_IN[OF g, of x] F_IN_OUT[OF g x] F_finite_OUT F_finite_OUT[OF g, of x]
          by(cases "d_OUT g z = 0")
            (auto simp add: nn_integral_divide nn_integral_cmult d_OUT_def[symmetric] ennreal_divide_times less_top)
        thus ?thesis by simp
      next
        case False
        have "d_IN g x ≤ d_OUT g x" using x by(rule F_IN_OUT[OF g])
        also have "… ≤ d_OUT (link g) x" unfolding d_OUT_def using False by(auto intro!: nn_integral_mono)
        finally show ?thesis .
      qed
      finally show "d_IN (link g) x ≤ d_OUT (link g) x" .
    qed
    with le g show ?thesis unfolding F by(simp add: leq_def le_fun_def del: link)
  qed

  have "bourbaki_witt_fixpoint Inf leq link" using chain_Field increasing unfolding leq_def
    by(intro bourbaki_witt_fixpoint_restrict_rel)(auto intro: Inf_greatest Inf_lower)
  then interpret bourbaki_witt_fixpoint Inf leq link .

  def g  "fixp_above f'"

  have g: "g ∈ Field leq" using f' unfolding g_def by(rule fixp_above_Field)
  hence "linkage Γ g" by(rule F_link)
  moreover
  have "web_flow Γ g"
  proof(intro web_flow.intros current.intros)
    show "d_OUT g x ≤ weight Γ x" for x using g by(rule F_weight_OUT)
    show "d_IN g x ≤ weight Γ x" for x using g by(rule F_weight_IN)
    show "d_IN g x = 0" if "x ∈ A Γ" for x using g that by(rule F_IN)
    show B: "d_OUT g x = 0" if "x ∈ B Γ" for x using g that by(rule F_OUT)
    show "g e = 0" if "e ∉ E" for e using g that by(rule F_outside)

    show KIR: "KIR g x" if A: "x ∉ A Γ" and B: "x ∉ B Γ" for x
    proof(rule ccontr)
      def z  "Eps (?P g)"
      assume "¬ KIR g x"
      with F_IN_OUT[OF g B] have "d_OUT g x > d_IN g x" by simp
      with A B have Ex: "∃x. ?P g x" by blast
      then have "?P g z" unfolding z_def by(rule someI_ex)
      hence A: "z ∉ A Γ" and B: "z ∉ B Γ" and less: "d_IN g z < d_OUT g z" by simp_all
      let ?factor = "d_IN g z / d_OUT g z"
      have "∃y. edge Γ z y ∧ g (z, y) > 0"
      proof(rule ccontr)
        assume "¬ ?thesis"
        hence "d_OUT g z = 0" using F_outside[OF g]
          by(force simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space not_less)
        with less show False by simp
      qed
      then obtain y where y: "edge Γ z y" and gr0: "g (z, y) > 0" by blast
      have "?factor < 1" (is "?factor < _") using less
        by(cases "d_OUT g z" "d_IN g z" rule: ennreal2_cases)
          (auto simp add: ennreal_less_iff divide_ennreal)

      hence le': "?factor * g (z, y) < 1 * g (z, y)" using gr0 F_finite[OF g]
        by(intro ennreal_mult_strict_right_mono) (auto simp: less_top)
      hence "link g (z, y) ≠ g (z, y)" using Ex by(auto simp add: link_def z_def Let_def)
      hence "link g ≠ g" by(auto simp add: fun_eq_iff)
      moreover have "link g = g" using f' unfolding g_def by(rule fixp_above_unfold[symmetric])
      ultimately show False by contradiction
    qed
    show "d_OUT g x ≤ d_IN g x" if "x ∉ A Γ" for x using KIR[of x] that B[of x]
      by(cases "x ∈ B Γ") auto
  qed
  ultimately show ?thesis by blast
qed

end

subsection ‹Loose webs are linkable›

lemma linkage_quotient_webD:
  fixes Γ :: "('v, 'more) web_scheme" (structure) and h g
  defines "k ≡ plus_current h g"
  assumes f: "current Γ f"
  and w: "wave Γ f"
  and wg: "web_flow (quotient_web Γ f) g" (is "web_flow ?Γ _")
  and link: "linkage (quotient_web Γ f) g"
  and trim: "trimming Γ f h"
  shows "web_flow Γ k"
  and "orthogonal_current Γ k (ℰ (TER f))"
proof -
  from wg have g: "current ?Γ g" by(rule web_flowD_current)

  from trim obtain h: "current Γ h" and w': "wave Γ h" and h_le_f: "⋀e. h e ≤ f e"
    and KIR: "⋀x. ⟦x ∈ RF (TER f); x ∉ A Γ⟧ ⟹ KIR h x"
    and TER: "TER h ⊇ ℰ (TER f) - A Γ"
    by(cases)(auto simp add: le_fun_def)

  have eq: "quotient_web Γ f = quotient_web Γ h" using w trim by(rule quotient_web_trimming)

  let ?T = "ℰ (TER f)"

  have RFc: "RF (TER h) = RF (TER f)"
    by(subst (1 2) roofed_circ_essential[symmetric])(simp only: trimming_ℰ[OF w trim])
  have OUT_k: "d_OUT k x = (if x ∈ RF (TER f) then d_OUT h x else d_OUT g x)" for x
    using OUT_plus_current[OF h w', of g] web_flowD_current[OF wg] unfolding eq k_def RFc by simp
  have RF: "RF (TER h) = RF (TER f)"
    by(subst (1 2) RF_essential[symmetric])(simp only: trimming_ℰ[OF w trim])
  have IN_k: "d_IN k x = (if x ∈ RF (TER f) then d_IN h x else d_IN g x)" for x
    using IN_plus_current[OF h w', of g] web_flowD_current[OF wg] unfolding eq k_def RF by simp

  have k: "current Γ k" unfolding k_def using h w' g unfolding eq by(rule current_plus_current)
  then show wk: "web_flow Γ k"
  proof(rule web_flow)
    fix x
    assume "x ∈ V" and A: "x ∉ A Γ" and B: "x ∉ B Γ"
    show "KIR k x"
    proof(cases "x ∈ ℰ (TER f)")
      case False
      thus ?thesis using A KIR[of x] web_flowD_KIR[OF wg, of x] B by(auto simp add: OUT_k IN_k roofed_circ_def)
    next
      case True
      with A TER have [simp]: "d_OUT h x = 0" and "d_IN h x ≥ weight Γ x"
        by(auto simp add: SINK.simps elim: SAT.cases)
      with currentD_weight_IN[OF h, of x] have [simp]: "d_IN h x = weight Γ x" by auto
      from linkageD[OF link, of x] True currentD_IN[OF g, of x] B
      have "d_OUT g x = weight Γ x" "d_IN g x = 0" by(auto simp add: roofed_circ_def)
      thus ?thesis using True by(auto simp add: IN_k OUT_k roofed_circ_def intro: roofed_greaterI)
    qed
  qed

  have h_le_k: "h e ≤ k e" for e unfolding k_def plus_current_def by(rule add_increasing2) simp_all
  hence "SAT Γ h ⊆ SAT Γ k" by(rule SAT_mono)
  hence SAT: "?T ⊆ SAT Γ k" using TER by(auto simp add: elim!: SAT.cases intro: SAT.intros)
  show "orthogonal_current Γ k ?T"
  proof(rule orthogonal_current)
    show "weight Γ x ≤ d_IN k x" if "x ∈ ?T" "x ∉ A Γ" for x
      using subsetD[OF SAT, of x] that by(auto simp add: currentD_SAT[OF k])
  next
    fix x
    assume x: "x ∈ ?T" and A: "x ∈ A Γ" and B: "x ∉ B Γ"
    with d_OUT_mono[of h x f, OF h_le_f] have "d_OUT h x = 0" by(auto simp add: SINK.simps)
    moreover from linkageD[OF link, of x] x A have "d_OUT g x = weight ?Γ x" by simp
    ultimately show "d_OUT k x = weight Γ x" using x A currentD_IN[OF f A] B
      by(auto simp add: d_OUT_add roofed_circ_def k_def plus_current_def )
  next
    fix u v
    assume v: "v ∈ RF ?T" and u: "u ∉ RF ?T"
    have "h (u, v) ≤ f (u, v)" by(rule h_le_f)
    also have "… ≤ d_OUT f u" unfolding d_OUT_def by(rule nn_integral_ge_point) simp
    also have "… = 0" using u using RF_essential[of Γ "TER f"]
      by(auto simp add: roofed_circ_def SINK.simps intro: waveD_OUT[OF w])
    finally have "h (u, v) = 0" by simp
    moreover have "g (u, v) = 0" using g v RF_essential[of Γ "TER f"]
      by(auto intro: currentD_outside simp add: roofed_circ_def)
    ultimately show "k (u, v) = 0" by(simp add: k_def)
  qed
qed

context countable_web begin

theorem loose_linkable: -- ‹Theorem 6.2›
  assumes "loose Γ"
  shows "linkable Γ"
proof -
  interpret bi: countable_bipartite_web "bipartite_web_of Γ" by(rule countable_bipartite_web_of)
  have "¬ hindered (bipartite_web_of Γ)" using assms by(rule unhindered_bipartite_web_of)
  then have "linkable (bipartite_web_of Γ)"
    by(rule bi.unhindered_linkable)
  then show ?thesis by(rule linkable_bipartite_web_ofD) simp
qed

lemma ex_orthogonal_current: -- ‹Lemma 4.15›
  "∃f S. web_flow Γ f ∧ separating Γ S ∧ orthogonal_current Γ f S"
proof -
  from ex_maximal_wave[OF countable]
  obtain f where f: "current Γ f"
    and w: "wave Γ f"
    and maximal: "⋀w. ⟦ current Γ w; wave Γ w; f ≤ w ⟧ ⟹ f = w" by blast
  from ex_trimming[OF f w countable weight_finite] obtain h where h: "trimming Γ f h" ..

  let  = "quotient_web Γ f"
  interpret Γ: countable_web "?Γ" by(rule countable_web_quotient_web)
  have "loose ?Γ" using f w maximal by(rule loose_quotient_web[OF  weight_finite])
  then have "linkable ?Γ" by(rule Γ.loose_linkable)
  then obtain g where wg: "web_flow ?Γ g" and link: "linkage ?Γ g" by blast

  let ?k = "plus_current h g"
  have "web_flow Γ ?k" "orthogonal_current Γ ?k (ℰ (TER f))"
    by(rule linkage_quotient_webD[OF f w wg link h])+
  moreover have "separating Γ (ℰ (TER f))"
    using waveD_separating[OF w] by(rule separating_essential)
  ultimately show ?thesis by blast
qed

end

subsection ‹Reduction of networks to webs›

definition web_of_network :: "('v, 'more) network_scheme ⇒ ('v edge, 'more) web_scheme"
where
  "web_of_network Δ =
   ⦇edge = λ(x, y) (y', z). y' = y ∧ edge Δ x y ∧ edge Δ y z,
    weight = capacity Δ,
    A = {(source Δ, x)|x. edge Δ (source Δ) x},
    B = {(x, sink Δ)|x. edge Δ x (sink Δ)},
    … = network.more Δ⦈"

lemma web_of_network_sel [simp]:
  fixes Δ (structure) shows
  "edge (web_of_network Δ) e e' ⟷ e ∈ E ∧ e' ∈ E ∧ snd e = fst e'"
  "weight (web_of_network Δ) e = capacity Δ e"
  "A (web_of_network Δ) = {(source Δ, x)|x. edge Δ (source Δ) x}"
  "B (web_of_network Δ) = {(x, sink Δ)|x. edge Δ x (sink Δ)}"
by(auto simp add: web_of_network_def split: prod.split)

lemma vertex_web_of_network [simp]:
  "vertex (web_of_network Δ) (x, y) ⟷ edge Δ x y ∧ (∃z. edge Δ y z ∨ edge Δ z x)"
by(auto simp add: vertex_def Domainp.simps Rangep.simps)

definition flow_of_current :: "('v, 'more) network_scheme ⇒ 'v edge current ⇒ 'v flow"
where "flow_of_current Δ f e = max (d_OUT f e) (d_IN f e)"

lemma flow_flow_of_current:
  fixes Δ (structure) and Γ
  defines [simp]: "Γ ≡ web_of_network Δ"
  assumes fw: "web_flow Γ f"
  shows "flow Δ (flow_of_current Δ f)" (is "flow _ ?f")
proof
  from fw have f: "current Γ f" and KIR: "⋀x. ⟦ x ∉ A Γ; x ∉ B Γ ⟧ ⟹ KIR f x"
    by(auto 4 3 dest: web_flowD_current web_flowD_KIR)

  show "?f e ≤ capacity Δ e" for e
    using currentD_weight_OUT[OF f, of e] currentD_weight_IN[OF f, of e]
    by(simp add: flow_of_current_def)

  fix x
  assume x: "x ≠ source Δ" "x ≠ sink Δ"
  have "d_OUT ?f x = (∑+ y. d_IN f (x, y))" unfolding d_OUT_def
    by(simp add: flow_of_current_def max_absorb2 currentD_OUT_IN[OF f] x)
  also have "… = (∑+ y. ∑+ e∈range (λz. (z, x)). f (e, x, y))"
    by(auto simp add: nn_integral_count_space_indicator d_IN_def intro!: nn_integral_cong currentD_outside[OF f] split: split_indicator)
  also have "… = (∑+ z∈UNIV. ∑+ y. f ((z, x), x, y))"
    by(subst nn_integral_snd_count_space[of "case_prod _", simplified])
      (simp add: nn_integral_count_space_reindex nn_integral_fst_count_space[of "case_prod _", simplified])
  also have "… = (∑+ z. ∑+ e∈range (Pair x). f ((z, x), e))"
    by(simp add: nn_integral_count_space_reindex)
  also have "… = (∑+ z. d_OUT f (z, x))"
    by(auto intro!: nn_integral_cong currentD_outside[OF f] simp add: d_OUT_def nn_integral_count_space_indicator split: split_indicator)
  also have "… = (∑+ z∈{z. edge Δ z x}. d_OUT f (z, x))"
    by(auto intro!: nn_integral_cong currentD_outside_OUT[OF f] simp add: nn_integral_count_space_indicator split: split_indicator)
  also have "… = (∑+ z∈{z. edge Δ z x}. max (d_OUT f (z, x)) (d_IN f (z, x)))"
  proof(rule nn_integral_cong)
    show "d_OUT f (z, x) = max (d_OUT f (z, x)) (d_IN f (z, x))"
      if "z ∈ space (count_space {z. edge Δ z x})" for z using currentD_IN[OF f] that
      by(cases "z = source Δ")(simp_all add: max_absorb1  currentD_IN[OF f] KIR x)
  qed
  also have "… = (∑+ z. max (d_OUT f (z, x)) (d_IN f (z, x)))"
    by(auto intro!: nn_integral_cong currentD_outside_OUT[OF f] currentD_outside_IN[OF f] simp add: nn_integral_count_space_indicator max_def split: split_indicator)
  also have "… = d_IN ?f x" by(simp add: flow_of_current_def d_IN_def)
  finally show "KIR ?f x" .
qed

text ‹
  The reduction of Conjecture 1.2 to Conjecture 3.6 is flawed in @{cite "AharoniBergerGeorgakopoulusPerlsteinSpruessel2011JCT"}.
  Not every essential A-B separating set of vertices in @{term "web_of_network Δ"} is an s-t-cut in
  @{term Δ}, as the following counterexample shows.

  The network @{term Δ} has five nodes @{term "s"}, @{term "t"}, @{term "x"}, @{term "y"} and @{term "z"}
  and edges @{term "(s, x)"}, @{term "(x, y)"}, @{term "(y, z)"}, @{term "(y, t)"} and @{term "(z, t)"}.
  For @{term "web_of_network Δ"}, the set @{term "S = {(x, y), (y, z)}"} is essential and A-B separating.
  (@{term "(x, y)"} is essential due to the path @{term "[(y, z)]"} and @{term "(y, z)"} is essential
  due to the path @{term "[(z, t)]"}). However, @{term S} is not a cut in @{term Δ} because the node @{term y}
  has an outgoing edge that is in @{term S} and one that is not in @{term S}.

  However, this can be remedied if all edges carry positive capacity. Then, orthogonality of the current
  rules out the above possibility.
›
lemma cut_RF_separating:
  fixes Δ (structure)
  assumes sep: "separating_network Δ V"
  and sink: "sink Δ ∉ V"
  shows "cut Δ (RFN V)"
proof
  show "source Δ ∈ RFN V" by(rule roofedI)(auto dest: separatingD[OF sep])
  show "sink Δ ∉ RFN V" using sink by(auto dest: roofedD[OF _ rtrancl_path.base])
qed

context
  fixes Δ :: "('v, 'more) network_scheme" and Γ (structure)
  defines Γ_def: "Γ ≡ web_of_network Δ"
begin

lemma separating_network_cut_of_sep:
  assumes sep: "separating Γ S"
  and source_sink: "source Δ ≠ sink Δ"
  shows "separating_network Δ (fst ` ℰ S)"
proof
  def s  "source Δ" and t  "sink Δ"
  fix p
  assume p: "path Δ s p t"
  with p source_sink have "p ≠ []" by cases(auto simp add: s_def t_def)
  with p have p': "path Γ (s, hd p) (zip p (tl p)) (last (s # butlast p), t)"
  proof(induction)
    case (step x y p z)
    then show ?case by(cases p)(auto elim: rtrancl_path.cases intro: rtrancl_path.intros simp add: Γ_def)
  qed simp
  from sep have "separating Γ (ℰ S)" by(rule separating_essential)
  from this p' have "(∃z∈set (zip p (tl p)). z ∈ ℰ S) ∨ (s, hd p) ∈ ℰ S"
    apply(rule separatingD)
    using rtrancl_path_nth[OF p, of 0] rtrancl_path_nth[OF p, of "length p - 1"] ‹p ≠ []› rtrancl_path_last[OF p]
    apply(auto simp add: Γ_def s_def t_def hd_conv_nth last_conv_nth nth_butlast nth_Cons' cong: if_cong split: if_split_asm)
    apply(metis One_nat_def Suc_leI cancel_comm_monoid_add_class.diff_cancel le_antisym length_butlast length_greater_0_conv list.size(3))+
    done
  then show "(∃z∈set p. z ∈ fst ` ℰ S) ∨ s ∈ fst ` ℰ S"
    by(auto dest!: set_zip_leftD intro: rev_image_eqI)
qed

definition cut_of_sep :: "('v × 'v) set ⇒ 'v set"
where "cut_of_sep S = RFNΔ (fst ` ℰ S)"

lemma separating_cut:
  assumes sep: "separating Γ S"
  and neq: "source Δ ≠ sink Δ"
  and sink_out: "⋀x. ¬ edge Δ (sink Δ) x"
  shows "cut Δ (cut_of_sep S)"
  unfolding cut_of_sep_def
proof(rule cut_RF_separating)
  show "separating_network Δ (fst ` ℰ S)" using sep neq by(rule separating_network_cut_of_sep)

  show "sink Δ ∉ fst ` ℰ S"
  proof
    assume "sink Δ ∈ fst ` ℰ S"
    then obtain x where "(sink Δ, x) ∈ ℰ S" by auto
    hence "(sink Δ, x) ∈ V" by(auto simp add: Γ_def dest!: essential_vertex)
    then show False by(simp add: Γ_def sink_out)
  qed
qed

context
  fixes f :: "'v edge current" and S
  assumes wf: "web_flow Γ f"
  and ortho: "orthogonal_current Γ f S"
  and sep: "separating Γ S"
  and capacity_pos: "⋀e. e ∈ EΔ ⟹ capacity Δ e > 0"
begin

private lemma f: "current Γ f" using wf by(rule web_flowD_current)

lemma orthogonal_leave_RF:
  assumes e: "edge Δ x y"
  and x: "x ∈ (cut_of_sep S)"
  and y: "y ∉ (cut_of_sep S)"
  shows "(x, y) ∈ S"
proof -
  from y obtain p where p: "path Δ y p (sink Δ)" and y': "y ∉ fst ` ℰ S"
    and bypass: "⋀z. z ∈ set p ⟹ z ∉ fst ` ℰ S" by(auto simp add: roofed_def cut_of_sep_def Γ_def[symmetric])
  from e p have p': "path Δ x (y # p) (sink Δ)" ..
  from roofedD[OF x[unfolded cut_of_sep_def] this] y' bypass have "x ∈ fst ` ℰ S" by(auto simp add: Γ_def[symmetric])
  then obtain z where xz: "(x, z) ∈ ℰ S" by auto
  then obtain q b where q: "path Γ (x, z) q b" and b: "b ∈ B Γ"
    and distinct: "distinct ((x, z) # q)" and bypass': "⋀z. z ∈ set q ⟹ z ∉ RF S"
    by(rule ℰ_E_RF) blast

  def p'  "y # p"
  hence "p' ≠ []" by simp
  with p' have "path Γ (x, hd p') (zip p' (tl p')) (last (x # butlast p'), sink Δ)"
    unfolding p'_def[symmetric]
  proof(induction)
    case (step x y p z)
    then show ?case
      by(cases p)(auto elim: rtrancl_path.cases intro: rtrancl_path.intros simp add: Γ_def)
  qed simp
  then have p'': "path Γ (x, y) (zip (y # p) p) (last (x # butlast (y # p)), sink Δ)" (is "path _ ?y ?p ?t")
    by(simp add: p'_def)
  have "(?y # ?p) ! length p = ?t" using rtrancl_path_last[OF p'] p rtrancl_path_last[OF p]
    apply(auto split: if_split_asm simp add: nth_Cons butlast_conv_take take_Suc_conv_app_nth split: nat.split elim: rtrancl_path.cases)
    apply(simp add: last_conv_nth)
    done
  moreover have "length p < length (?y # ?p)" by simp
  ultimately have t: "?t ∈ B Γ" using rtrancl_path_nth[OF p'', of "length p - 1"] e
    by(cases p)(simp_all add: Γ_def split: if_split_asm)

  show S: "(x, y) ∈ S"
  proof(cases "x = source Δ")
    case True
    from separatingD[OF separating_essential, OF sep p'' _ t] e True
    consider (z) z z' where "(z, z') ∈ set ?p" "(z, z') ∈ ℰ S" | "(x, y) ∈ S" by(auto simp add: Γ_def)
    thus ?thesis
    proof cases
      case (z z)
      hence "z ∈ set p" "z ∈ fst ` ℰ S"
        using y' by(auto dest!: set_zip_leftD intro: rev_image_eqI)
      hence False by(auto dest: bypass)
      thus ?thesis ..
    qed
  next
    case False
    have "∃e. edge Γ e (x, z) ∧ f (e, (x, z)) > 0"
    proof(rule ccontr)
      assume "¬ ?thesis"
      then have "d_IN f (x, z) = 0" unfolding d_IN_def using currentD_outside[OF f, of _ "(x, z)"]
        by(force simp add: nn_integral_0_iff_AE AE_count_space not_less)
      moreover
      from xz have "(x, z) ∈ S" by auto
      hence "(x, z) ∈ SAT Γ f" by(rule orthogonal_currentD_SAT[OF ortho])
      with False have "d_IN f (x, z) ≥ capacity Δ (x, z)" by(auto simp add: SAT.simps Γ_def)
      ultimately have "¬ capacity Δ (x, z) > 0" by auto
      hence "¬ edge Δ x z" using capacity_pos[of "(x, z)"] by auto
      moreover with q have "b = (x, z)" by cases(auto simp add: Γ_def)
      with b have "edge Δ x z" by(simp add: Γ_def)
      ultimately show False by contradiction
    qed
    then obtain u where ux: "edge Δ u x" and xz': "edge Δ x z" and uxz: "edge Γ (u, x) (x, z)"
      and gt0: "f ((u, x), (x, z)) > 0" by(auto simp add: Γ_def)
    have "(u, x) ∈ RF S" using orthogonal_currentD_in[OF ortho, of "(x, z)" "(u, x)"] gt0 xz
      by(fastforce intro: roofed_greaterI)
    hence ux_RF: "(u, x) ∈ RF (ℰ S)" and ux_ℰ: "(u, x) ∉ ℰ S" unfolding RF_essential by(simp_all add: roofed_circ_def)

    from ux e have "edge Γ (u, x) (x, y)" by(simp add: Γ_def)
    hence "path Γ (u, x) ((x, y) # ?p) ?t" using p'' ..
    from roofedD[OF ux_RF this t] ux_ℰ
    consider "(x, y) ∈ S" | (z) z z' where "(z, z') ∈ set ?p" "(z, z') ∈ ℰ S" by auto
    thus ?thesis
    proof cases
      case (z z)
      with bypass[of z] y' have False by(fastforce dest!: set_zip_leftD intro: rev_image_eqI)
      thus ?thesis ..
    qed
  qed
qed

lemma orthogonal_flow_of_current:
  assumes source_sink: "source Δ ≠ sink Δ"
  and sink_out: "⋀x. ¬ edge Δ (sink Δ) x"
  and no_direct_edge: "¬ edge Δ (source Δ) (sink Δ)" -- ‹Otherwise, @{const A} and @{const B} of the web would not be disjoint.›
  shows "orthogonal Δ (flow_of_current Δ f) (cut_of_sep S)" (is "orthogonal _ ?f ?S")
proof
  fix x y
  assume e: "edge Δ x y" and "x ∈ ?S" and "y ∉ ?S"
  then have S: "(x, y) ∈ S"by(rule orthogonal_leave_RF)

  show "?f (x, y) = capacity Δ (x, y)"
  proof(cases "x = source Δ")
    case False
    with orthogonal_currentD_SAT[OF ortho S]
    have "weight Γ (x, y) ≤ d_IN f (x, y)" by cases(simp_all add: Γ_def)
    with False currentD_OUT_IN[OF f, of "(x, y)"] currentD_weight_IN[OF f, of "(x, y)"]
    show ?thesis by(simp add: flow_of_current_def Γ_def max_def)
  next
    case True
    with orthogonal_currentD_A[OF ortho S] e currentD_weight_IN[OF f, of "(x, y)"] no_direct_edge
    show ?thesis by(auto simp add: flow_of_current_def Γ_def max_def)
  qed
next
  from sep source_sink sink_out have cut: "cut Δ ?S" by(rule separating_cut)

  fix x y
  assume xy: "edge Δ x y"
    and x: "x ∉ ?S"
    and y: "y ∈ ?S"
  from x obtain p where p: "path Δ x p (sink Δ)" and x': "x ∉ fst ` ℰ S"
    and bypass: "⋀z. z ∈ set p ⟹ z ∉ fst ` ℰ S" by(auto simp add: roofed_def cut_of_sep_def)
  have source: "x ≠ source Δ"
  proof
    assume "x = source Δ"
    have "separating_network Δ (fst ` ℰ S)" using sep source_sink by(rule separating_network_cut_of_sep)
    from separatingD[OF this p] ‹x = source Δ› x show False
      by(auto dest: bypass intro: roofed_greaterI simp add: cut_of_sep_def)
  qed
  hence A: "(x, y) ∉ A Γ" by(simp add: Γ_def)

  have "f ((u, v), x, y) = 0" for u v
  proof(cases "edge Γ (u, v) (x, y)")
    case False with f show ?thesis by(rule currentD_outside)
  next
    case True
    hence [simp]: "v = x" and ux: "edge Δ u x" by(auto simp add: Γ_def)
    have "(x, y) ∈ RF S"
    proof
      fix q b
      assume q: "path Γ (x, y) q b" and b: "b ∈ B Γ"
      def xy  "(x, y)"
      from q have "path Δ (snd xy) (map snd q) (snd b)" unfolding xy_def[symmetric]
        by(induction)(auto intro: rtrancl_path.intros simp add: Γ_def)
      with b have "path Δ y (map snd q) (sink Δ)" by(auto simp add: xy_def Γ_def)
      from roofedD[OF y[unfolded cut_of_sep_def] this] have "∃z∈set (y # map snd q). z ∈ ?S"
        by(auto intro: roofed_greaterI simp add: cut_of_sep_def)
      from split_list_last_prop[OF this] obtain xs z ys where decomp: "y # map snd q = xs @ z # ys"
        and z: "z ∈ ?S" and last: "⋀z. z ∈ set ys ⟹ z ∉ ?S" by auto
      from decomp obtain x' xs' z'' ys' where decomp': "(x', y) # q = xs' @ (z'', z) # ys'"
        and "xs = map snd xs'" and ys: "ys = map snd ys'" and x': "xs' = [] ⟹ x' = x"
        by(fastforce simp add: Cons_eq_append_conv map_eq_append_conv)
      from cut z have z_sink: "z ≠ sink Δ" by cases(auto)
      then have "ys' ≠ []" using rtrancl_path_last[OF q] decomp' b x' q
        by(auto simp add: Cons_eq_append_conv Γ_def elim: rtrancl_path.cases)
      then obtain w z''' ys'' where ys': "ys' = (w, z''') # ys''" by(auto simp add: neq_Nil_conv)
      from q[THEN rtrancl_path_nth, of "length xs'"] decomp' ys' x' have "edge Γ (z'', z) (w, z''')"
        by(auto simp add: Cons_eq_append_conv nth_append)
      hence w: "w = z" and zz''': "edge Δ z z'''" by(auto simp add: Γ_def)
      from ys' ys last[of z'''] have "z''' ∉ ?S" by simp
      with zz''' z have "(z, z''') ∈ S" by(rule orthogonal_leave_RF)
      moreover have "(z, z''') ∈ set q" using decomp' ys' w by(auto simp add: Cons_eq_append_conv)
      ultimately show "(∃z∈set q. z ∈ S) ∨ (x, y) ∈ S" by blast
    qed
    moreover
    have "(u, x) ∉ RF S"
    proof
      assume "(u, x) ∈ RF S"
      hence ux_RF: "(u, x) ∈ RF (ℰ S)" and ux_ℰ: "(u, x) ∉ ℰ S"
        unfolding RF_essential by(simp_all add: roofed_circ_def)

      have "x ≠ sink Δ" using p xy by cases(auto simp add: sink_out)
      with p have nNil: "p ≠ []" by(auto elim: rtrancl_path.cases)
      with p have "edge Δ x (hd p)" by cases auto
      with ux have "edge Γ (u, x) (x, hd p)" by(simp add: Γ_def)
      moreover
      from p nNil have "path Γ (x, hd p) (zip p (tl p)) (last (x # butlast p), sink Δ)" (is "path _ ?x ?p ?t")
      proof(induction)
        case (step x y p z)
        then show ?case
          by(cases p)(auto elim: rtrancl_path.cases intro: rtrancl_path.intros simp add: Γ_def)
      qed simp
      ultimately have p': "path Γ (u, x) (?x # ?p) ?t" ..

      have "(?x # ?p) ! (length p - 1) = ?t" using rtrancl_path_last[OF p] p nNil
        apply(auto split: if_split_asm simp add: nth_Cons butlast_conv_take take_Suc_conv_app_nth not_le split: nat.split elim: rtrancl_path.cases)
        apply(simp add: last_conv_nth nth_tl)
        done
      moreover have "length p - 1 < length (?x # ?p)" by simp
      ultimately have t: "?t ∈ B Γ" using rtrancl_path_nth[OF p', of "length p - 1"]
        by(cases p)(simp_all add: Γ_def split: if_split_asm)
      from roofedD[OF ux_RF p' t] ux_ℰ consider (X) "(x, hd p) ∈ ℰ S"
        | (z) z z' where "(z, z') ∈ set (zip p (tl p))" "(z, z') ∈ ℰ S" by auto
      thus False
      proof cases
        case X with x' show False by(auto simp add: cut_of_sep_def intro: rev_image_eqI)
      next
        case (z z)
        with bypass[of z] show False by(auto 4 3 simp add: cut_of_sep_def intro: rev_image_eqI dest!: set_zip_leftD)
      qed
    qed
    ultimately show ?thesis unfolding ‹v = x› by(rule orthogonal_currentD_in[OF ortho])
  qed
  then have "d_IN f (x, y) = 0" unfolding d_IN_def
    by(simp add: nn_integral_0_iff emeasure_count_space_eq_0)
  with currentD_OUT_IN[OF f A] show "flow_of_current Δ f (x, y) = 0"
    by(simp add: flow_of_current_def max_def)
qed

end

end

subsection ‹The theorem›

context antiparallel_edges begin

abbreviation cut' :: "'a vertex set ⇒ 'a set" where "cut' S ≡ Vertex -` S"

lemma cut_cut': "cut Δ'' S ⟹ cut Δ (cut' S)"
by(auto simp add: cut.simps)

lemma IN_Edge: "INΔ'' (Edge x y) = (if edge Δ x y then {Vertex x} else {})"
by(auto simp add: incoming_def)

lemma OUT_Edge: "OUTΔ'' (Edge x y) = (if edge Δ x y then {Vertex y} else {})"
by(auto simp add: outgoing_def)

interpretation Δ'': countable_network Δ'' by(rule Δ''_countable_network)

lemma d_IN_Edge:
  assumes f: "flow Δ'' f"
  shows "d_IN f (Edge x y) = f (Vertex x, Edge x y)"
by(subst d_IN_alt_def[OF Δ''.flowD_outside[OF f], of _ Δ''])(simp_all add: IN_Edge nn_integral_count_space_indicator max_def Δ''.flowD_outside[OF f])

lemma d_OUT_Edge:
  assumes f: "flow Δ'' f"
  shows "d_OUT f (Edge x y) = f (Edge x y, Vertex y)"
by(subst d_OUT_alt_def[OF Δ''.flowD_outside[OF f], of _ Δ''])(simp_all add: OUT_Edge nn_integral_count_space_indicator max_def Δ''.flowD_outside[OF f])

lemma orthogonal_cut':
  assumes ortho: "orthogonal Δ'' f S"
  and f: "flow Δ'' f"
  shows "orthogonal Δ (collect f) (cut' S)"
proof
  show "collect f (x, y) = capacity Δ (x, y)" if edge: "edge Δ x y" and x: "x ∈ cut' S" and y: "y ∉ cut' S" for x y
  proof(cases "Edge x y ∈ S")
    case True
    from y have "Vertex y ∉ S" by auto
    from orthogonalD_out[OF ortho _ True this] edge show ?thesis by simp
  next
    case False
    from x have "Vertex x ∈ S" by auto
    from orthogonalD_out[OF ortho _ this False] edge
    have "capacity Δ (x, y) = d_IN f (Edge x y)" by(simp add: d_IN_Edge[OF f])
    also have "… = d_OUT f (Edge x y)" by(simp add: flowD_KIR[OF f])
    also have "… = f (Edge x y, Vertex y)" using edge by(simp add: d_OUT_Edge[OF f])
    finally show ?thesis by simp
  qed

  show "collect f (x, y) = 0" if edge: "edge Δ x y" and x: "x ∉ cut' S" and y: "y ∈ cut' S" for x y
  proof(cases "Edge x y ∈ S")
    case True
    from x have "Vertex x ∉ S" by auto
    from orthogonalD_in[OF ortho _ this True] edge have "0 = d_IN f (Edge x y)" by(simp add: d_IN_Edge[OF f])
    also have "… = d_OUT f (Edge x y)" by(simp add: flowD_KIR[OF f])
    also have "… = f (Edge x y, Vertex y)" using edge by(simp add: d_OUT_Edge[OF f])
    finally show ?thesis by simp
  next
    case False
    from y have "Vertex y ∈ S" by auto
    from orthogonalD_in[OF ortho _ False this] edge show ?thesis by simp
  qed
qed

end

context countable_network begin

context begin

qualified lemma max_flow_min_cut':
  assumes source_in: "⋀x. ¬ edge Δ x (source Δ)"
  and sink_out: "⋀y. ¬ edge Δ (sink Δ) y"
  and undead: "⋀x y. edge Δ x y ⟹ (∃z. edge Δ y z) ∨ (∃z. edge Δ z x)"
  and source_sink: "¬ edge Δ (source Δ) (sink Δ)"
  and no_loop: "⋀x. ¬ edge Δ x x"
  and edge_antiparallel: "⋀x y. edge Δ x y ⟹ ¬ edge Δ y x"
  and capacity_pos: "⋀e. e ∈ E ⟹ capacity Δ e > 0"
  shows "∃f S. flow Δ f ∧ cut Δ S ∧ orthogonal Δ f S"
proof -
  let  = "web_of_network Δ"
  interpret web: countable_web 
  proof
    show "¬ edge ?Γ y x" if "x ∈ A ?Γ" for x y using that by(clarsimp simp add: source_in)
    show "¬ edge ?Γ x y" if "x ∈ B ?Γ" for x y using that by(clarsimp simp add: sink_out)
    show "A ?Γ ⊆ V" by(auto 4 3 dest: undead)
    show "A ?Γ ∩ B ?Γ = {}" using source_sink by auto
    show "¬ edge ?Γ x x" for x by(auto simp add: no_loop)
    show "weight ?Γ x = 0" if "x ∉ V" for x using that undead
      by(cases x)(auto intro!: capacity_outside)
    show "weight ?Γ x ≠ ⊤" for x using capacity_finite[of x] by(cases x) simp
    show "¬ edge ?Γ y x" if "edge ?Γ x y" for x y using that by(auto simp add: edge_antiparallel)
    have "EE × E" by auto
    thus "countable E" by(rule countable_subset)(simp)
  qed
  from web.ex_orthogonal_current obtain f S
    where f: "web_flow (web_of_network Δ) f"
    and S: "separating (web_of_network Δ) S"
    and ortho: "orthogonal_current (web_of_network Δ) f S" by blast+
  let ?f = "flow_of_current Δ f" and ?S = "cut_of_sep Δ S"
  from f have "flow Δ ?f" by(rule flow_flow_of_current)
  moreover have "cut Δ ?S" using S source_neq_sink sink_out by(rule separating_cut)
  moreover have "orthogonal Δ ?f ?S" using f ortho S capacity_pos source_neq_sink sink_out source_sink
    by(rule orthogonal_flow_of_current)
  ultimately show ?thesis by blast
qed

qualified lemma max_flow_min_cut'':
  assumes sink_out: "⋀y. ¬ edge Δ (sink Δ) y"
  and source_in: "⋀x. ¬ edge Δ x (source Δ)"
  and no_loop: "⋀x. ¬ edge Δ x x"
  and capacity_pos: "⋀e. e ∈ E ⟹ capacity Δ e > 0"
  shows "∃f S. flow Δ f ∧ cut Δ S ∧ orthogonal Δ f S"
proof -
  interpret antiparallel_edges Δ ..
  interpret Δ'': countable_network Δ'' by(rule Δ''_countable_network)
  have "∃f S. flow Δ'' f ∧ cut Δ'' S ∧ orthogonal Δ'' f S"
    by(rule Δ''.max_flow_min_cut')(auto simp add: sink_out source_in no_loop capacity_pos elim: edg.cases)
  then obtain f S where f: "flow Δ'' f" and cut: "cut Δ'' S" and ortho: "orthogonal Δ'' f S" by blast
  have "flow Δ (collect f)" using f by(rule flow_collect)
  moreover have "cut Δ (cut' S)" using cut by(rule cut_cut')
  moreover have "orthogonal Δ (collect f) (cut' S)" using ortho f by(rule orthogonal_cut')
  ultimately show ?thesis by blast
qed

qualified lemma max_flow_min_cut''':
  assumes sink_out: "⋀y. ¬ edge Δ (sink Δ) y"
  and source_in: "⋀x. ¬ edge Δ x (source Δ)"
  and capacity_pos: "⋀e. e ∈ E ⟹ capacity Δ e > 0"
  shows "∃f S. flow Δ f ∧ cut Δ S ∧ orthogonal Δ f S"
proof -
  interpret antiparallel_edges Δ ..
  interpret Δ'': countable_network Δ'' by(rule Δ''_countable_network)
  have "∃f S. flow Δ'' f ∧ cut Δ'' S ∧ orthogonal Δ'' f S"
    by(rule Δ''.max_flow_min_cut'')(auto simp add: sink_out source_in capacity_pos elim: edg.cases)
  then obtain f S where f: "flow Δ'' f" and cut: "cut Δ'' S" and ortho: "orthogonal Δ'' f S" by blast
  have "flow Δ (collect f)" using f by(rule flow_collect)
  moreover have "cut Δ (cut' S)" using cut by(rule cut_cut')
  moreover have "orthogonal Δ (collect f) (cut' S)" using ortho f by(rule orthogonal_cut')
  ultimately show ?thesis by blast
qed

theorem max_flow_min_cut:
  "∃f S. flow Δ f ∧ cut Δ S ∧ orthogonal Δ f S"
proof -
  def Δ'  "⦇edge = λx y. edge Δ x y ∧ capacity Δ (x, y) > 0 ∧ y ≠ source Δ ∧ x ≠ sink Δ,
    capacity = λ(x, y). if x = sink Δ ∨ y = source Δ then 0 else capacity Δ (x, y),
    source = source Δ,
    sink = sink Δ⦈"
  have Δ'_sel [simp]:
    "edge Δ' x y ⟷ edge Δ x y ∧ capacity Δ (x, y) > 0 ∧ y ≠ source Δ ∧ x ≠ sink Δ"
    "capacity Δ' (x, y) = (if x = sink Δ ∨ y = source Δ then 0 else capacity Δ (x, y))"
    "source Δ' = source Δ"
    "sink Δ' = sink Δ"
    for x y by(simp_all add: Δ'_def)

  interpret Δ': countable_network Δ'
  proof(unfold_locales)
    have "EΔ'E" by auto
    then show "countable EΔ'" by(rule countable_subset) simp
    show "capacity Δ' e = 0" if "e ∉ EΔ'" for e
      using capacity_outside[of e] that by(auto split: if_split_asm intro: ccontr)
  qed(auto simp add: split: if_split_asm)

  have "∃f S. flow Δ' f ∧ cut Δ' S ∧ orthogonal Δ' f S" by(rule Δ'.max_flow_min_cut''') auto
  then obtain f S where f: "flow Δ' f" and cut: "cut Δ' S" and ortho: "orthogonal Δ' f S" by blast
  have "flow Δ f"
  proof
    show "f e ≤ capacity Δ e" for e using flowD_capacity[OF f, of e]
      by(cases e)(simp split: if_split_asm)
    show "KIR f x" if "x ≠ source Δ" "x ≠ sink Δ" for x using flowD_KIR[OF f, of x] that by simp
  qed
  moreover have "cut Δ S" using cut by(simp add: cut.simps)
  moreover have "orthogonal Δ f S"
  proof
    show "f (x, y) = capacity Δ (x, y)" if edge: "edge Δ x y" and x: "x ∈ S" and y: "y ∉ S" for x y
    proof(cases "edge Δ' x y")
      case True
      with orthogonalD_out[OF ortho this x y] show ?thesis by simp
    next
      case False
      from cut y x have xy: "y ≠ source Δ ∧ x ≠ sink Δ" by(cases) auto
      with xy edge False have "capacity Δ (x, y) = 0" by simp
      with Δ'.flowD_outside[OF f, of "(x, y)"] False show ?thesis by simp
    qed

    show "f (x, y) = 0" if edge: "edge Δ x y" and x: "x ∉ S" and y: "y ∈ S" for x y
      using orthogonalD_in[OF ortho _ x y] Δ'.flowD_outside[OF f, of "(x, y)"]
      by(cases "edge Δ' x y")simp_all
  qed
  ultimately show ?thesis by blast
qed

end

end

hide_const (open) A B weight

end