Theory MFMC_Unbounded

(* Author: Andreas Lochbihler, ETH Zurich *)

section ‹The max-flow min-cut theorems in unbounded networks›

theory MFMC_Unbounded imports
  MFMC_Web
  MFMC_Flow_Attainability
  MFMC_Reduction
begin

subsection ‹More about waves›

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)"

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="(+)"] 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="(+)"] 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 MFMC_Unbounded.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 "(zset 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 (≤) Y" "Y  {}" "countable (support_flow (Sup Y))"
  shows "RF (TER (Sup Y)) = RF (fY. TER f)"
proof(rule set_eqI iffI)+
  fix x
  assume x: "x  RF (TER (Sup Y))"
  have "x  RF (RF (fY. 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 "(zset p. z  RF (fY. TER f))  x  RF (fY. TER f)"
    proof(cases "fY. 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 (fY. 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 (fY. TER f)" unfolding roofed_idem .
next
  fix x
  assume x: "x  RF (fY. 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 "(zset 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

subsection ‹Hindered webs with reduced weights›

context countable_bipartite_web begin

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 ("_")

  define edge'
    where "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)" for xo yo
  define cap
    where "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)" for e
  define Ψ where "Ψ = 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  (+ yrange 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 " = (+ xrange 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 "Ψ" .

  define α where "α = (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: "(+ yOUT 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 = (+ yrange 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

  define g where "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 = (+ yOUT 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 " = (+ yOUT x. f (x, y) + j (y, x)) - (+ yOUT 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 " = (+ yOUT x. f (x, y)) + (+ yOUT 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 = (+ yrange 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 = (+ yrange 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: "(+ yIN 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 " = (+ yIN 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 " = (+ yIN x. f (y, x) + j (Some x, Some y)) - (+ yIN 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 " = (+ yIN x. f (y, x)) + (+ yIN 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 = (+ yrange 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 = (+ yrange 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

  define cap' where "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)

  define G where "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)
  define reachable where "reachable x = (edge G)** x a" for x
  have reachable_alt_def: "reachable  λx. p. path G x p a"
    by(simp add: reachable_def [abs_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)

  define h where "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 = (+ yOUT x. h (x, y))" using h_le_g currentD_outside[OF g]
      by(intro d_OUT_alt_def) auto
    also have " = (+ yOUT 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 = (+ yIN x. h (y, x))"
      using h_le_g currentD_outside[OF g] by(intro d_IN_alt_def) auto
    also have " = (+ yIN 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"
    define E where "E = (None, Some y) # map (map_prod Some Some) ?E"
    define ζ where "ζ = Min (insert (u y - j (None, Some y)) (cap' ` set ?E))"
    let ?j' = "λe. (if e  set E then ζ else 0) + j e"
    define j' where "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: 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_remove)
      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 "ε  + xB Γ. d_IN g x - d_OUT g x"
  assumes g: "current Γ g"
  and ε_finite: "ε  "
  and hindered: "hindered_by (Γ  g) ε"
  shows "hindered Γ"
proof -
  define Γ' where "Γ' = Γ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 *: "(+ xB Γ'. ?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 "(+ xB Γ'. ?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) (+ xB Γ'. ?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 ("_")
  define edge'
    where "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)" for xo yo
  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)

  define cap
    where "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)" for xoyo
  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)
  define Ψ where "Ψ = 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 " = (+ yrange Inner. h 0 (x, inner y)) + weight Γ x"
      by(auto simp add: nn_integral_count_space_indicator nn_integral_add)
    also have "(+ yrange 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 " = (+ xrange Inner. h 0 (inner x, y)) + weight Γ b"
        by(simp add: nn_integral_add nn_integral_count_space_indicator)
      also have "(+ xrange 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 "Ψ" .
  define α where "α = (SUP f{f. flow Ψ f}. value_flow Ψ f)"

  define f
    where "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)" for n xoyo
  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 " = (+ yrange 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 "(+ yrange 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 " = (+ yrange 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 " = (+ yrange 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 "(+ yrange 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 " = (+ yrange 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)"

  define h' where "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 = (+ yrange 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 = (+ yrange 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 = (+ yrange 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 = (+ yrange 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 = (+ yrange 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 = (+ yrange 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 "(zset 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

  define wb where "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)

  define ε where "ε n = min δ wb / (n + 2)" for n :: nat
  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)

  define Γ' where "Γ' n = reduce_weight Γ b (ε n)" for n :: nat
  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])

  define factor where "factor n = (wb - ε 0) / (wb - ε n)" for 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)

  define g' where "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, opaque_lifting) 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')

  define f where "f = rec_nat (g 0) (λn rec. rec Γ' 0g' (n + 1))"
  have f_simps [simp]:
    "f 0 = g 0"
    "f (Suc n) = f n Γ' 0g' (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 Γ' 0g' (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 (≤) (range f)"
    by(rule chain_imageI[where le_a="(≤)"])(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 (in. 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 Γ' 0g' (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(in. TERΓ' 0(g' i))" by(simp add: Suc.IH)
    also have "RFΓ' 0(  TERΓ' 0(g' (n + 1))) = RFΓ' 0((in. TERΓ' 0(g' i))  TERΓ' 0(g' (n + 1)))"
      by(simp add: roofed_idem_Un1)
    also have "(in. TERΓ' 0(g' i))  TERΓ' 0(g' (n + 1)) = (iSuc n. TERΓ' 0(g' i))"
      unfolding atMost_Suc UN_insert by(simp add: Un_commute)
    finally show ?case by simp
  qed

  define  where " = (SUP n. f n)"
  have : "current (Γ' 0) " unfolding gω_def using chain_f
    by(rule current_Sup)(auto simp add: f supp_f)
  have : "wave (Γ' 0) " unfolding gω_def using chain_f
    by(rule wave_lub)(auto simp add: fw  supp_f)
  from  have gω': "current (Γ' n) " 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  = (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) " 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 " 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)) (  Γ' 0 / g' n)" for n using w' 
    by(rule current_restrict_current)
  have gωr: "current (quotient_web (Γ' n) (g n)) (  Γ' n / g n)" for n using w gω'
    by(rule current_restrict_current)
  have wωr: "wave (quotient_web (Γ' n) (g n)) (  Γ' n / g n)" (is "wave ?Γ' ?gω") for n
  proof
    have *: "wave (quotient_web (Γ' 0) (g' n)) (  Γ' 0 / g' n)"
      using g' w'   by(rule wave_restrict_current)
    have "d_IN (  Γ' 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)) (  Γ' n / g n)  SAT ?Γ' (  Γ' 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)(  Γ' n / g n)  TER?Γ'(  Γ' 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 "(zset p. z  ?TER)  x  ?TER" by auto
    qed

    show "d_OUT (  Γ' 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) = RF (n. TERΓ' 0(g' n))"
  proof -
    have "RFΓ' 0(TERΓ' 0) = 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. ji. TERΓ' 0(g' j))" unfolding RF_f roofed_UN by simp
    also have "(i. ji. 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) = SAT (Γ' 0) (g' n Γ' 0)" 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 (  Γ' 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) = SINK (g' n Γ' 0)" 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) = TERΓ' 0(g' n Γ' 0)" for n
    by(rule set_eqI iffI)+(simp_all add: SAT_plus_ω SINK_plus_ω)

  define h where "h n = g n Γ' n" for n
  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)

  define T where "T = TERΓ' 0"
  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)  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Γ' 0T" 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

  define S where "S = {x  RF (T  B Γ)  A Γ. essential Γ (T  B Γ) (RF (T  B Γ)  A Γ) x}"
  define Γ_h where "Γ_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 "(zset 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Γ_h E" 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)
    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 "(zset 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 Ω - fi is loose in order to
  apply Lemma 6.9.

  Counterexample: The bipartite web with three nodes a1, a2, a3 in A›
  and two nodes b1, b2 in B› and edges (a1, b1)›, (a2, b1)›,
  (a2, b2)›, (a3, b2)› and weights a1 = a3 = 1› and a2 = 2› and
  b1 = 3› and b2 = 2›.
  Then, we can get a sequence of weight reductions on b2 from 2› to 1.5›,
  1.25›, 1.125›, etc. with limit 1›.
  All maximal waves in the restricted webs in the sequence are @{term [source] zero_current}, so in
  the limit, we get k = 0› and ε = 1› for a2 and b2. Now, the
  restricted web for the two is not loose because it contains the wave which assigns 1 to (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 "(zset 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Γ  fg  y  TERΓ  fg" 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Γ  fg)"
    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Γ  fg) a" using p py y
        by(auto intro!: essentialI)
      moreover from  a have "a  TERΓ  fg"
        by(auto simp add: SAT.A SINK_plus_current)
      ultimately have "a  Γ  f(TERΓ  fg)" 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  (TERzero_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  TERzero_current" using y disjoint by(auto)
        hence "essential  (B ) (TERzero_current) x" using p y py by(auto intro!: essentialI)
        with a have "x  (TERzero_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

  define F where "F = (λ(ε, h :: 'v current). plus_current ε h)"
  have F_simps: "F (ε, h) = plus_current ε h" for ε h by(simp add: F_def)
  define Fld where "Fld = {(ε, h).
     current Γ ε  (x. x  a  d_OUT ε x = 0) 
     current (Γ  ε) h  wave (Γ  ε) h 
     ¬ hindered (Γ  F (ε, h))}"
  define leq where "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 (≤) M"
      by(auto simp add: chain_def leq_def in_restrict_rel_iff)
    hence chain1: "Complete_Partial_Order.chain (≤) (fst ` M)"
      and chain2: "Complete_Partial_Order.chain (≤) (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 image_comp)
    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: "(εhM. 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  (εhM. d_OUT (fst εh) x)" for x using in_M by(rule SUP_upper2)simp
    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 image_comp)
    have IN1_le: "(εhM. 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  (εhM. 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 hsnd ` 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 add: image_comp)
      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: image_comp)
      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 image_comp)
      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 hsnd ` 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 add: image_comp)
      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 add: image_comp)
      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 image_comp
                     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 image_comp)
      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  TERh"
        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  (TERh)" 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)

    define f where "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 image_comp)
      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  (TERg)"
        and OUT_z: "d_OUT g z < weight  z" by cases auto
      define δ where "δ = 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 (≤) ((λ(ε, 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 image_comp)
        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]

      define k where "k e = Sup (fst ` M) e - ε e" for 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 ε] image_comp
           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 *: "(xaM. 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 image_comp 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

      define q where "q = (+ yB (Γ  ?ε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

      define g' where "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 image_comp 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] image_comp 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 image_comp
                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  (TERg')"
        proof
          assume "z  (TERg')"
          hence OUT_z: "d_OUT g' z = 0"
            and ess: "essential  (B Γ) (TERg') 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 (TERg')" 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 (TERg')" 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 (xM. 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Γ  fg" by simp
          with p y py have "essential Γ (B Γ) (TERΓ  fg) z" by(auto intro: essentialI)
          moreover with z waveD_separating[OF g_w, THEN separating_RF_A] have "z  (TERg)"
            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 image_comp
                      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

  define sat where "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 εhk"
    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 εhk)" 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  ka  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  (TERzero_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')"
    define k where "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

      define Ω where "Ω = Γ  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"
      define y where "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])

      define bound where "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 δ)"
      define δ where "δ = 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 .

  define f where "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')"
  define k where "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

    define Ω where "Ω = Γ  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 fk" 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"
    define y where "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)

    define bound where "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 δ)"
    define δ where "δ = 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')"

  define enum where "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

  define f where "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 }
    { 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)

    define f' where "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 (≤) (range f)" by(rule incseq_chain_range)

  define g where "g = Sup (range f)"
  have "support_flow g  E"
    by (auto simp add: g_def support_flow.simps currentD_outside [OF f] image_comp 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


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"
  by(rule ex_orthogonal_current')(rule countable_web.loose_linkable[OF countable_web_quotient_web])

end

subsection ‹Glueing the reductions together›

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 capacity_pos: "e. e  E  capacity Δ e > 0"
  shows "f S. flow Δ f  cut Δ S  orthogonal Δ f S"
  by(rule max_flow_min_cut')(rule countable_web.ex_orthogonal_current[OF countable_web_web_of_network], fact+)

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 -
  interpret Δ''': countable_network Δ''' by(rule Δ'''_countable_network)
  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
  from flow_Δ'''[OF this] show ?thesis by blast
qed

end

end

end