diff --git a/thys/MFMC_Countable/MFMC_Bounded.thy b/thys/MFMC_Countable/MFMC_Bounded.thy --- a/thys/MFMC_Countable/MFMC_Bounded.thy +++ b/thys/MFMC_Countable/MFMC_Bounded.thy @@ -1,655 +1,653 @@ (* Author: Andreas Lochbihler, Digital Asset *) section \The max-flow min-cut theorem in bounded networks\ subsection \Linkages in unhindered bipartite webs\ theory MFMC_Bounded imports Matrix_For_Marginals MFMC_Reduction begin context countable_bipartite_web begin lemma countable_A [simp]: "countable (A \)" using A_vertex countable_V by(blast intro: countable_subset) lemma unhindered_criterion [rule_format]: assumes "\ hindered \" shows "\X \ A \. finite X \ (\\<^sup>+ x\X. weight \ x) \ (\\<^sup>+ y\\<^bold>E `` X. weight \ y)" using assms proof(rule contrapos_np) assume "\ ?thesis" then obtain X where "X \ {X. X \ A \ \ finite X \ (\\<^sup>+ y\\<^bold>E `` X. weight \ y) < (\\<^sup>+ x\X. weight \ x)}" (is "_ \ Collect ?P") by(auto simp add: not_le) from wf_eq_minimal[THEN iffD1, OF wf_finite_psubset, rule_format, OF this, simplified] obtain X where X_A: "X \ A \" and fin_X [simp]: "finite X" and less: "(\\<^sup>+ y\\<^bold>E `` X. weight \ y) < (\\<^sup>+ x\X. weight \ x)" and minimal: "\X'. X' \ X \ (\\<^sup>+ x\X'. weight \ x) \ (\\<^sup>+ y\\<^bold>E `` X'. weight \ y)" by(clarsimp simp add: not_less)(meson finite_subset order_trans psubset_imp_subset) have nonempty: "X \ {}" using less by auto then obtain xx where xx: "xx \ X" by auto define f where "f x = (if x = xx then (\\<^sup>+ y\\<^bold>E `` X. weight \ y) - (\\<^sup>+ x\X - {xx}. weight \ x) else if x \ X then weight \ x else 0)" for x define g where "g y = (if y \ \<^bold>E `` X then weight \ y else 0)" for y define E' where "E' \ \<^bold>E \ X \ UNIV" have Xxx: "X - {xx} \ X" using xx by blast have E [simp]: "E' `` X' = \<^bold>E `` X'" if "X' \ X" for X' using that by(auto simp add: E'_def) have in_E': "(x, y) \ E' \ x \ X \ (x, y) \ \<^bold>E" for x y by(auto simp add: E'_def) have "(\\<^sup>+ x\X. f x) = (\\<^sup>+ x\X - {xx}. f x) + (\\<^sup>+ x\{xx}. f x)" using xx by(auto simp add: nn_integral_count_space_indicator nn_integral_add[symmetric] simp del: nn_integral_indicator_singleton intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+ x\X - {xx}. weight \ x) + ((\\<^sup>+ y\\<^bold>E `` X. weight \ y) - (\\<^sup>+ x\X - {xx}. weight \ x))" by(rule arg_cong2[where f="(+)"])(auto simp add: f_def xx nn_integral_count_space_indicator intro!: nn_integral_cong) also have "\ = (\\<^sup>+ y\\<^bold>E `` X. g y)" using minimal[OF Xxx] xx by(subst add_diff_eq_iff_ennreal[THEN iffD2])(fastforce simp add: g_def[abs_def] nn_integral_count_space_indicator intro!: nn_integral_cong intro: nn_integral_mono elim: order_trans split: split_indicator)+ finally have sum_eq: "(\\<^sup>+ x\X. f x) = (\\<^sup>+ y\\<^bold>E `` X. g y)" . have "(\\<^sup>+ y\\<^bold>E `` X. weight \ y) = (\\<^sup>+ y\\<^bold>E `` X. g y)" by(auto simp add: nn_integral_count_space_indicator g_def intro!: nn_integral_cong) then have fin: "\ \ \" using less by auto have fin2: "(\\<^sup>+ x\X'. weight \ x) \ \" if "X' \ X" for X' proof - have "(\\<^sup>+ x\\<^bold>E `` X'. weight \ x) \ (\\<^sup>+ x\\<^bold>E `` X. weight \ x)" using that by(auto 4 3 simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator split_indicator_asm) then show ?thesis using minimal[OF that] less by(auto simp add: top_unique) qed have "f xx = (\\<^sup>+ y\\<^bold>E `` X. weight \ y) - (\\<^sup>+ x\X - {xx}. weight \ x)" by (simp add: f_def) also have "\ < (\\<^sup>+ x\X. weight \ x) - (\\<^sup>+ x\X - {xx}. weight \ x)" using less fin2[OF Xxx] minimal[OF Xxx] by(subst minus_less_iff_ennreal)(fastforce simp add: less_top[symmetric] nn_integral_count_space_indicator diff_add_self_ennreal intro: nn_integral_mono elim: order_trans split: split_indicator)+ also have "\ = (\\<^sup>+ x\{xx}. weight \ x)" using fin2[OF Xxx] xx apply(simp add: nn_integral_count_space_indicator del: nn_integral_indicator_singleton) apply(subst nn_integral_diff[symmetric]) apply(auto simp add: AE_count_space split: split_indicator simp del: nn_integral_indicator_singleton intro!: nn_integral_cong) done also have "\ = weight \ xx" by(simp add: nn_integral_count_space_indicator) finally have fxx: "f xx < weight \ xx" . have le: "(\\<^sup>+ x\X'. f x) \ (\\<^sup>+ y\\<^bold>E `` X'. g y)" if "X' \ X" for X' proof(cases "X' = X") case True then show ?thesis using sum_eq by simp next case False hence X': "X' \ X" using that by blast have "(\\<^sup>+ x\X'. f x) = (\\<^sup>+ x\X' - {xx}. f x) + (\\<^sup>+ x\{xx}. f x * indicator X' xx)" by(auto simp add: nn_integral_count_space_indicator nn_integral_add[symmetric] simp del: nn_integral_indicator_singleton intro!: nn_integral_cong split: split_indicator) also have "\ \ (\\<^sup>+ x\X' - {xx}. f x) + (\\<^sup>+ x\{xx}. weight \ x * indicator X' xx)" using fxx by(intro add_mono)(auto split: split_indicator simp add: nn_integral_count_space_indicator) also have "\ = (\\<^sup>+ x\X'. weight \ x)" using xx that by(auto simp add: nn_integral_count_space_indicator nn_integral_add[symmetric] f_def simp del: nn_integral_indicator_singleton intro!: nn_integral_cong split: split_indicator) also have "\ \ (\\<^sup>+ y\\<^bold>E `` X'. weight \ y)" by(rule minimal[OF X']) also have "\ = (\\<^sup>+ y\\<^bold>E `` X'. g y)" using that by(auto 4 3 intro!: nn_integral_cong simp add: g_def Image_iff) finally show ?thesis . qed have "countable X" using X_A A_vertex countable_V by(blast intro: countable_subset) moreover have "\<^bold>E `` X \ \<^bold>V" by(auto simp add: vertex_def) with countable_V have "countable (\<^bold>E `` X)" by(blast intro: countable_subset) moreover have "E' \ X \ \<^bold>E `` X" by(auto simp add: E'_def) ultimately obtain h' where h'_dom: "\x y. 0 < h' x y \ (x, y) \ E'" and h'_fin: "\x y. h' x y \ \" and h'_f: "\x. x \ X \ (\\<^sup>+ y\E' `` X. h' x y) = f x" and h'_g: "\y. y \ E' `` X \ (\\<^sup>+ x\X. h' x y) = g y" using bounded_matrix_for_marginals_ennreal[where f=f and g=g and A=X and B="E' `` X" and R=E' and thesis=thesis] sum_eq fin le by(auto) have h'_outside: "(x, y) \ E' \ h' x y = 0" for x y using h'_dom[of x y] not_gr_zero by(fastforce) define h where "h = (\(x, y). if x \ X \ edge \ x y then h' x y else 0)" have h_OUT: "d_OUT h x = (if x \ X then f x else 0)" for x by(auto 4 3 simp add: h_def d_OUT_def h'_f[symmetric] E'_def nn_integral_count_space_indicator intro!: nn_integral_cong intro: h'_outside split: split_indicator) have h_IN: "d_IN h y = (if y \ \<^bold>E `` X then weight \ y else 0)" for y using h'_g[of y, symmetric] by(auto 4 3 simp add: h_def d_IN_def g_def nn_integral_count_space_indicator nn_integral_0_iff_AE in_E' intro!: nn_integral_cong intro: h'_outside split: split_indicator split_indicator_asm) have h: "current \ h" proof show "d_OUT h x \ weight \ x" for x using fxx by(auto simp add: h_OUT f_def) show "d_IN h y \ weight \ y" for y by(simp add: h_IN) show "h e = 0" if "e \ \<^bold>E" for e using that by(cases e)(auto simp add: h_def) qed have "separating \ (TER h)" proof fix x y p assume x: "x \ A \" and y: "y \ B \" and p: "path \ x p y" then obtain [simp]: "p = [y]" and xy: "(x, y) \ \<^bold>E" using disjoint by -(erule rtrancl_path.cases; auto dest: bipartite_E)+ show "(\z\set p. z \ TER h) \ x \ TER h" proof(rule disjCI) assume "x \ TER h" hence "x \ X" using x by(auto simp add: SAT.simps SINK.simps h_OUT split: if_split_asm) hence "y \ TER h" using xy currentD_OUT[OF h y] by(auto simp add: SAT.simps h_IN SINK.simps) thus "\z\set p. z \ TER h" by simp qed qed then have w: "wave \ h" using h .. have "xx \ A \" using xx X_A by blast moreover have "xx \ \ (TER h)" proof assume "xx \ \ (TER h)" then obtain p y where y: "y \ B \" and p: "path \ xx p y" and bypass: "\z. \ xx \ y; z \ set p \ \ z = xx \ z \ TER h" by(rule \_E) auto from p obtain [simp]: "p = [y]" and xy: "edge \ xx y" and neq: "xx \ y" using disjoint X_A xx y by -(erule rtrancl_path.cases; auto dest: bipartite_E)+ from neq bypass[of y] have "y \ TER h" by simp moreover from xy xx currentD_OUT[OF h y] have "y \ TER h" by(auto simp add: SAT.simps h_IN SINK.simps) ultimately show False by contradiction qed moreover have "d_OUT h xx < weight \ xx" using fxx xx by(simp add: h_OUT) ultimately have "hindrance \ h" .. then show "hindered \" using h w .. qed end lemma nn_integral_count_space_top_approx: fixes f :: "nat => ennreal" and b :: ennreal assumes "nn_integral (count_space UNIV) f = top" and "b < top" obtains n where "b < sum f {.. of_nat x \ 1 \ x" by (metis of_nat_le_iff of_nat_1) locale bounded_countable_bipartite_web = countable_bipartite_web \ for \ :: "('v, 'more) web_scheme" (structure) + assumes bounded_B: "x \ A \ \ (\\<^sup>+ y \ \<^bold>E `` {x}. weight \ y) < \" begin theorem unhindered_linkable_bounded: assumes "\ hindered \" shows "linkable \" proof(cases "A \ = {}") case True hence "linkage \ (\_. 0)" by(auto simp add: linkage.simps) moreover have "web_flow \ (\_. 0)" by(auto simp add: web_flow.simps) ultimately show ?thesis by blast next case nonempty: False define A_n :: "nat \ 'v set" where "A_n n = from_nat_into (A \) ` {..n}" for n have fin_A_n [simp]: "finite (A_n n)" for n by(simp add: A_n_def) have A_n_A: "A_n n \ A \" for n by(auto simp add: A_n_def from_nat_into[OF nonempty]) have countable2: "countable (\<^bold>E `` A_n n)" for n using countable_V by(rule countable_subset[rotated])(auto simp add: vertex_def) have "\Y2. \n. \X \ A_n n. Y2 n X \ \<^bold>E `` X \ (\\<^sup>+ x\X. weight \ x) \ (\\<^sup>+ y\Y2 n X. weight \ y) \ (\\<^sup>+ y\Y2 n X. weight \ y) \ \" proof(rule choice strip ex_simps(6)[THEN iffD2])+ fix n X assume X: "X \ A_n n" then have [simp]: "finite X" by(rule finite_subset) simp have X_count: "countable (\<^bold>E `` X)" using countable2 by(rule countable_subset[rotated])(rule Image_mono[OF order_refl X]) show "\Y. Y \ \<^bold>E `` X \ (\\<^sup>+ x\X. weight \ x) \ (\\<^sup>+ y\Y. weight \ y) \ (\\<^sup>+ y\Y. weight \ y) \ \" (is "Ex ?P") proof(cases "(\\<^sup>+ y\\<^bold>E `` X. weight \ y) = \") case True define Y' where "Y' = to_nat_on (\<^bold>E `` X) ` (\<^bold>E `` X)" have inf: "infinite (\<^bold>E `` X)" using True by(intro notI)(auto simp add: nn_integral_count_space_finite) then have Y': "Y' = UNIV" using X_count by(auto simp add: Y'_def intro!: image_to_nat_on) have "(\\<^sup>+ y\\<^bold>E `` X. weight \ y) = (\\<^sup>+ y\from_nat_into (\<^bold>E `` X) ` Y'. weight \ y * indicator (\<^bold>E `` X) y)" using X_count by(auto simp add: nn_integral_count_space_indicator Y'_def image_image intro!: nn_integral_cong from_nat_into_to_nat_on[symmetric] rev_image_eqI split: split_indicator) also have "\ = (\\<^sup>+ y. weight \ (from_nat_into (\<^bold>E `` X) y) * indicator (\<^bold>E `` X) (from_nat_into (\<^bold>E `` X) y))" using X_count inf by(subst nn_integral_count_space_reindex)(auto simp add: inj_on_def Y') finally have "\ = \" using True by simp from nn_integral_count_space_top_approx[OF this, of "sum (weight \) X"] obtain yy where yy: "sum (weight \) X < (\ y (from_nat_into (\<^bold>E `` X) y) * indicator (\<^bold>E `` X) (from_nat_into (\<^bold>E `` X) y))" by(auto simp add: less_top[symmetric]) define Y where "Y = from_nat_into (\<^bold>E `` X) ` {.. \<^bold>E `` X" have [simp]: "finite Y" by(simp add: Y_def) have "(\\<^sup>+ x\X. weight \ x) = sum (weight \) X" by(simp add: nn_integral_count_space_finite) also have "\ \ (\ y (from_nat_into (\<^bold>E `` X) y) * indicator (\<^bold>E `` X) (from_nat_into (\<^bold>E `` X) y))" using yy by simp also have "\ = (\ y \ from_nat_into (\<^bold>E `` X) ` {.. y * indicator (\<^bold>E `` X) y)" using X_count inf by(subst sum.reindex)(auto simp add: inj_on_def) also have "\ = (\ y \ Y. weight \ y)" by(auto intro!: sum.cong simp add: Y_def) also have "\ = (\\<^sup>+ y\Y. weight \ y)" by(simp add: nn_integral_count_space_finite) also have "Y \ \<^bold>E `` X" by(simp add: Y_def) moreover have "(\\<^sup>+ y\Y. weight \ y) \ \" by(simp add: nn_integral_count_space_finite) ultimately show ?thesis by blast next case False with unhindered_criterion[OF assms, of X] X A_n_A[of n] have "?P (\<^bold>E `` X)" by auto then show ?thesis .. qed qed then obtain Y2 where Y2_A: "Y2 n X \ \<^bold>E `` X" and le: "(\\<^sup>+ x\X. weight \ x) \ (\\<^sup>+ y\Y2 n X. weight \ y)" and finY2: "(\\<^sup>+ y\Y2 n X. weight \ y) \ \" if "X \ A_n n" for n X by iprover define Y where "Y n = (\ X \ Pow (A_n n). Y2 n X)" for n define s where "s n = (\\<^sup>+ y\Y n. weight \ y)" for n have Y_vertex: "Y n \ \<^bold>V" for n by(auto 4 3 simp add: Y_def vertex_def dest!: Y2_A[of _ n]) have Y_B: "Y n \ B \" for n unfolding Y_def by(auto dest!: Y2_A[of _ n] dest: bipartite_E) have s_top [simp]: "s n \ \" for n proof - have "\x \ Y2 n X; X \ A_n n\ \ Suc 0 \ card {X. X \ A_n n \ x \ Y2 n X}" for x X by(subst card_le_Suc_iff)(auto intro!: exI[where x=X] exI[where x="{X. X \ A_n n \ x \ Y2 n X} - {X}"]) then have "(\\<^sup>+ y\Y n. weight \ y) \ (\\<^sup>+ y\Y n. \ X\Pow (A_n n). weight \ y * indicator (Y2 n X) y)" by(intro nn_integral_mono)(auto simp add: Y_def One_le_of_nat_ennreal intro!: mult_right_mono[of "1 :: ennreal", simplified]) also have "\ = (\ X\Pow (A_n n). \\<^sup>+ y\Y n. weight \ y * indicator (Y2 n X) y)" by(subst nn_integral_sum) auto also have "\ = (\ X\Pow (A_n n). \\<^sup>+ y\Y2 n X. weight \ y)" by(auto intro!: sum.cong nn_integral_cong simp add: nn_integral_count_space_indicator Y_def split: split_indicator) also have "\ < \" by(simp add: less_top[symmetric] finY2) finally show ?thesis by(simp add: less_top s_def) qed define f :: "nat \ 'v option \ real" where "f n xo = (case xo of Some x \ if x \ A_n n then enn2real (weight \ x) else 0 | None \ enn2real (s n - sum (weight \) (A_n n)))" for n xo define g :: "nat \ 'v \ real" where "g n y = enn2real (weight \ y * indicator (Y n) y)" for n y define R :: "nat \ ('v option \ 'v) set" where "R n = map_prod Some id ` (\<^bold>E \ A_n n \ Y n) \ {None} \ Y n" for n define A_n' where "A_n' n = Some ` A_n n \ {None}" for n have f_simps: "f n (Some x) = (if x \ A_n n then enn2real (weight \ x) else 0)" "f n None = enn2real (s n - sum (weight \) (A_n n))" for n x by(simp_all add: f_def) have g_s: "(\\<^sup>+ y\Y n. g n y) = s n" for n by(auto simp add: s_def g_def ennreal_enn2real_if intro!: nn_integral_cong) have "(\\<^sup>+ x\A_n' n. f n x) = (\\<^sup>+ x\Some`A_n n. weight \ (the x)) + (\\<^sup>+ x\{None}. f n x)" for n by(auto simp add: nn_integral_count_space_indicator nn_integral_add[symmetric] f_simps A_n'_def ennreal_enn2real_if simp del: nn_integral_indicator_singleton intro!: nn_integral_cong split: split_indicator) also have "\ n = sum (weight \) (A_n n) + (s n - sum (weight \) (A_n n))" for n by(subst nn_integral_count_space_reindex)(auto simp add: nn_integral_count_space_finite f_simps ennreal_enn2real_if) also have "\ n = s n" for n using le[OF order_refl, of n] by(simp add: s_def nn_integral_count_space_finite)(auto elim!: order_trans simp add: nn_integral_count_space_indicator Y_def intro!: nn_integral_mono split: split_indicator) finally have sum_eq: "(\\<^sup>+ x\A_n' n. f n x) = (\\<^sup>+ y\Y n. g n y)" for n using g_s by simp have "\h'. \n. (\x y. (x, y) \ R n \ h' n x y = 0) \ (\x y. h' n x y \ \) \ (\x\A_n' n. (\\<^sup>+ y\Y n. h' n x y) = f n x) \ (\y\Y n. (\\<^sup>+ x\A_n' n. h' n x y) = g n y)" (is "Ex (\h'. \n. ?Q n (h' n))") proof(rule choice allI)+ fix n note sum_eq moreover have "(\\<^sup>+ y\Y n. g n y) \ \" using g_s by simp moreover have le_fg: "(\\<^sup>+ x\X. f n x) \ (\\<^sup>+ y\R n `` X. g n y)" if "X \ A_n' n" for X proof(cases "None \ X") case True have "(\\<^sup>+ x\X. f n x) \ (\\<^sup>+ x\A_n' n. f n x)" using that by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ = (\\<^sup>+ y\Y n. g n y)" by(simp add: sum_eq) also have "R n `` X = Y n" using True by(auto simp add: R_def) ultimately show ?thesis by simp next case False then have *: "Some ` (the ` X) = X" by(auto simp add: image_image)(metis (no_types, lifting) image_iff notin_range_Some option.sel option.collapse)+ from False that have X: "the ` X \ A_n n" by(auto simp add: A_n'_def) from * have "(\\<^sup>+ x\X. f n x) = (\\<^sup>+ x\Some ` (the ` X). f n x)" by simp also have "\ = (\\<^sup>+ x\the ` X. f n (Some x))" by(rule nn_integral_count_space_reindex) simp also have "\ = (\\<^sup>+ x\the ` X. weight \ x)" using that False by(auto 4 3intro!: nn_integral_cong simp add: f_simps A_n'_def ennreal_enn2real_if) also have "\ \ (\\<^sup>+ y\Y2 n (the ` X). weight \ y)" using False that by(intro le)(auto simp add: A_n'_def) also have "\ \ (\\<^sup>+ y\R n `` X. weight \ y)" using False Y2_A[of "the ` X" n] X by(auto simp add: A_n'_def nn_integral_count_space_indicator R_def Image_iff Y_def intro: rev_image_eqI intro!: nn_integral_mono split: split_indicator) (drule (1) subsetD; clarify; drule (1) bspec; auto 4 3 intro: rev_image_eqI) also have "\ = (\\<^sup>+ y\R n `` X. g n y)" by(auto intro!: nn_integral_cong simp add: R_def g_def ennreal_enn2real_if) finally show ?thesis . qed moreover have "countable (A_n' n)" by(simp add: A_n'_def countable_finite) moreover have "countable (Y2 n X)" if "X \ A_n n" for X using Y2_A[OF that] by(rule countable_subset)(rule countable_subset[OF _ countable_V]; auto simp add: vertex_def) then have "countable (Y n)" unfolding Y_def by(intro countable_UN)(simp_all add: countable_finite) moreover have "R n \ A_n' n \ Y n" by(auto simp add: R_def A_n'_def) ultimately obtain h' where "\x y. 0 < h' x y \ (x, y) \ R n" "\x y. h' x y \ \" "\x. x \ A_n' n \ (\\<^sup>+ y\Y n. h' x y) = (f n x)" "\y. y \ Y n \ (\\<^sup>+ x\A_n' n. h' x y) = g n y" by(rule bounded_matrix_for_marginals_ennreal) blast+ hence "?Q n h'" by(auto)(use not_gr_zero in blast) thus "Ex (?Q n)" by blast qed then obtain h' where dom_h': "\x y. (x, y) \ R n \ h' n x y = 0" and fin_h' [simp]: "\x y. h' n x y \ \" and h'_f: "\x. x \ A_n' n \ (\\<^sup>+ y\Y n. h' n x y) = f n x" and h'_g: "\y. y \ Y n \ (\\<^sup>+ x\A_n' n. h' n x y) = g n y" for n by blast define h :: "nat \ 'v \ 'v \ real" where "h n = (\(x, y). if x \ A_n n \ y \ Y n then enn2real (h' n (Some x) y) else 0)" for n have h_nonneg: "0 \ h n xy" for n xy by(simp add: h_def split_def) have h_notB: "h n (x, y) = 0" if "y \ B \" for n x y using Y_B[of n] that by(auto simp add: h_def) have h_le_weight2: "h n (x, y) \ weight \ y" for n x y proof(cases "x \ A_n n \ y \ Y n") case True have "h' n (Some x) y \ (\\<^sup>+ x\A_n' n. h' n x y)" by(rule nn_integral_ge_point)(auto simp add: A_n'_def True) also have "\ \ weight \ y" using h'_g[of y n] True by(simp add: g_def ennreal_enn2real_if) finally show ?thesis using True by(simp add: h_def ennreal_enn2real_if) qed(auto simp add: h_def) have h_OUT: "d_OUT (h n) x = (if x \ A_n n then weight \ x else 0)" for n x using h'_f[of "Some x" n, symmetric] by(auto simp add: h_def d_OUT_def A_n'_def f_simps ennreal_enn2real_if nn_integral_count_space_indicator intro!: nn_integral_cong) have h_IN: "d_IN (h n) y = (if y \ Y n then enn2real (weight \ y - h' n None y) else 0)" for n y proof(cases "y \ Y n") case True have "d_IN (h n) y = (\\<^sup>+ x\Some ` A_n n. h' n x y)" by(subst nn_integral_count_space_reindex) (auto simp add: d_IN_def h_def nn_integral_count_space_indicator ennreal_enn2real_if R_def intro!: nn_integral_cong dom_h' split: split_indicator) also have "\ = (\\<^sup>+ x\A_n' n. h' n x y) - (\\<^sup>+ x\{None}. h' n x y)" apply(simp add: nn_integral_count_space_indicator del: nn_integral_indicator_singleton) apply(subst nn_integral_diff[symmetric]) apply(auto simp add: AE_count_space A_n'_def nn_integral_count_space_indicator split: split_indicator intro!: nn_integral_cong) done also have "\ = g n y - h' n None y" using h'_g[OF True] by(simp add: nn_integral_count_space_indicator) finally show ?thesis using True by(simp add: g_def ennreal_enn2real_if) qed(auto simp add: d_IN_def ennreal_enn2real_if nn_integral_0_iff_AE AE_count_space h_def g_def) let ?Q = "\<^bold>V \ \<^bold>V" have "bounded (range (\n. h n xy))" if "xy \ ?Q" for xy unfolding bounded_def dist_real_def proof(rule exI strip|erule imageE|hypsubst)+ fix n obtain x y where [simp]: "xy = (x, y)" by(cases xy) have "h n (x, y) \ d_OUT (h n) x" unfolding d_OUT_def by(rule nn_integral_ge_point) simp also have "\ \ weight \ x" by(simp add: h_OUT) finally show "\0 - h n xy\ \ enn2real (weight \ (fst xy))" by(simp add: h_nonneg)(metis enn2real_ennreal ennreal_cases ennreal_le_iff weight_finite) qed moreover have "countable ?Q" using countable_V by(simp) ultimately obtain k where k: "strict_mono k" and conv: "\xy. xy \ ?Q \ convergent (\n. h (k n) xy)" by(rule convergent_bounded_family) blast+ have h_outside: "h n xy = 0" if "xy \ ?Q" for xy n using that A_n_A[of n] A_vertex Y_vertex by(auto simp add: h_def split: prod.split) have h_outside_AB: "h n (x, y) = 0" if "x \ A \ \ y \ B \" for n x y using that A_n_A[of n] Y_B[of n] by(auto simp add: h_def) have h_outside_E: "h n (x, y) = 0" if "(x, y) \ \<^bold>E" for n x y using that unfolding h_def by(clarsimp)(subst dom_h', auto simp add: R_def) define H where "H xy = lim (\n. h (k n) xy)" for xy have H: "(\n. h (k n) xy) \ H xy" for xy using conv[of xy] unfolding H_def by(cases "xy \ ?Q")(auto simp add: convergent_LIMSEQ_iff h_outside) have H_outside: "H (x, y) = 0" if "x \ A \ \ y \ B \" for x y using that by(simp add: H_def convergent_LIMSEQ_iff h_outside_AB) have H': "(\n. ennreal (h (k n) xy)) \ H xy" for xy using H by(rule tendsto_ennrealI) have H_def': "H xy = lim (\n. ennreal (h (k n) xy))" for xy by (metis H' limI) have H_OUT: "d_OUT H x = weight \ x" if x: "x \ A \" for x proof - let ?w = "\y. if (x, y) \ \<^bold>E then weight \ y else 0" have sum_w: "(\\<^sup>+ y. if edge \ x y then weight \ y else 0) = (\\<^sup>+ y \ \<^bold>E `` {x}. weight \ y)" by(simp add: nn_integral_count_space_indicator indicator_def if_distrib cong: if_cong) have "(\n. d_OUT (h (k n)) x) \ d_OUT H x" unfolding d_OUT_def by(rule nn_integral_dominated_convergence[where w="?w"])(use bounded_B x in \simp_all add: AE_count_space H h_outside_E h_le_weight2 sum_w\) moreover define n_x where "n_x = to_nat_on (A \) x" have x': "x \ A_n (k n)" if "n \ n_x" for n using that seq_suble[OF k, of n] x unfolding A_n_def by(intro rev_image_eqI[where x=n_x])(simp_all add: A_n_def n_x_def) have "(\n. d_OUT (h (k n)) x) \ weight \ x" by(intro tendsto_eventually eventually_sequentiallyI[where c="n_x"])(simp add: h_OUT x') ultimately show ?thesis by(rule LIMSEQ_unique) qed then have "linkage \ H" .. moreover have "web_flow \ H" unfolding web_flow_iff proof show "d_OUT H x \ weight \ x" for x by(cases "x \ A \")(simp_all add: H_OUT[unfolded d_OUT_def] H_outside d_OUT_def) show "d_IN H y \ weight \ y" for y proof - have "d_IN H y = (\\<^sup>+ x. liminf (\n. ennreal (h (k n) (x, y))))" unfolding d_IN_def H_def' by(rule nn_integral_cong convergent_liminf_cl[symmetric] convergentI H')+ also have "\ \ liminf (\n. d_IN (h (k n)) y)" unfolding d_IN_def by(rule nn_integral_liminf) simp also have "\ \ liminf (\n. weight \ y)" unfolding h_IN by(rule Liminf_mono)(auto simp add: ennreal_enn2real_if) also have "\ = weight \ y" by(simp add: Liminf_const) finally show "?thesis" . qed show "ennreal (H e) = 0" if "e \ \<^bold>E" for e proof(rule LIMSEQ_unique[OF H']) obtain x y where [simp]: "e = (x, y)" by(cases e) have "ennreal (h (k n) (x, y)) = 0" for n using dom_h'[of "Some x" y "k n"] that by(auto simp add: h_def R_def enn2real_eq_0_iff elim: meta_mp) then show "(\n. ennreal (h (k n) e)) \ 0" using that by(intro tendsto_eventually eventually_sequentiallyI) simp qed qed ultimately show ?thesis by blast qed end subsection \Glueing the reductions together\ locale bounded_countable_web = countable_web \ for \ :: "('v, 'more) web_scheme" (structure) + assumes bounded_out: "x \ \<^bold>V - B \ \ (\\<^sup>+ y \ \<^bold>E `` {x}. weight \ y) < \" begin lemma bounded_countable_bipartite_web_of: "bounded_countable_bipartite_web (bipartite_web_of \)" (is "bounded_countable_bipartite_web ?\") proof - interpret bi: countable_bipartite_web ?\ by(rule countable_bipartite_web_of) show ?thesis proof fix x assume "x \ A ?\" then obtain x' where x: "x = Inl x'" and x': "vertex \ x'" "x' \ B \" by auto have "\<^bold>E\<^bsub>?\\<^esub> `` {x} \ Inr ` ({x'} \ (\<^bold>E `` {x'}))" using x by(auto simp add: bipartite_web_of_def vertex_def split: sum.split_asm) hence "(\\<^sup>+ y \ \<^bold>E\<^bsub>?\\<^esub> `` {x}. weight ?\ y) \ (\\<^sup>+ y \ \. weight ?\ y)" by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ = (\\<^sup>+ y\{x'} \ (\<^bold>E `` {x'}). weight (bipartite_web_of \) (Inr y))" by(rule nn_integral_count_space_reindex)(auto) also have "\ \ weight \ x' + (\\<^sup>+ y \ \<^bold>E `` {x'}. weight \ y)" apply(subst (1 2) nn_integral_count_space_indicator, simp, simp) apply(cases "\ edge \ x' x'") apply(subst nn_integral_disjoint_pair) apply(auto intro!: nn_integral_mono add_increasing split: split_indicator) done also have "\ < \" using bounded_out[of x'] x' using weight_finite[of x'] by(simp del: weight_finite add: less_top) finally show "(\\<^sup>+ y \ \<^bold>E\<^bsub>?\\<^esub> `` {x}. weight ?\ y) < \" . qed qed theorem loose_linkable_bounded: assumes "loose \" shows "linkable \" proof - interpret bi: bounded_countable_bipartite_web "bipartite_web_of \" by(rule bounded_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_bounded) then show ?thesis by(rule linkable_bipartite_web_ofD) simp qed lemma bounded_countable_web_quotient_web: "bounded_countable_web (quotient_web \ f)" (is "bounded_countable_web ?\") proof - interpret r: countable_web ?\ by(rule countable_web_quotient_web) show ?thesis proof fix x assume "x \ \<^bold>V\<^bsub>quotient_web \ f\<^esub> - B (quotient_web \ f)" then have x: "x \ \<^bold>V - B \" by(auto dest: vertex_quotient_webD) have "(\\<^sup>+ y \ \<^bold>E\<^bsub>?\\<^esub> `` {x}. weight ?\ y) \ (\\<^sup>+ y \ \<^bold>E `` {x}. weight \ y)" by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ < \" using x by(rule bounded_out) finally show "(\\<^sup>+ y \ \<^bold>E\<^bsub>?\\<^esub> `` {x}. weight ?\ y) < \" . qed qed lemma ex_orthogonal_current: "\f S. web_flow \ f \ separating \ S \ orthogonal_current \ f S" proof - from ex_maximal_wave[OF countable] obtain f where f: "current \ f" and w: "wave \ f" and maximal: "\w. \ current \ w; wave \ w; f \ w \ \ f = w" by blast from ex_trimming[OF f w countable weight_finite] obtain h where h: "trimming \ f h" .. let ?\ = "quotient_web \ f" interpret \: bounded_countable_web "?\" by(rule bounded_countable_web_quotient_web) have "loose ?\" using f w maximal by(rule loose_quotient_web[OF weight_finite]) then have "linkable ?\" by(rule \.loose_linkable_bounded) then obtain g where wg: "web_flow ?\ g" and link: "linkage ?\ g" by blast let ?k = "plus_current h g" have "web_flow \ ?k" "orthogonal_current \ ?k (\ (TER f))" by(rule linkage_quotient_webD[OF f w wg link h])+ moreover have "separating \ (\ (TER f))" using waveD_separating[OF w] by(rule separating_essential) ultimately show ?thesis by blast qed end locale bounded_countable_network = countable_network \ for \ :: "('v, 'more) network_scheme" (structure) + assumes out: "\ x \ \<^bold>V; x \ source \ \ \ d_OUT (capacity \) x < \" context antiparallel_edges begin lemma \''_bounded_countable_network: "bounded_countable_network \''" if "\x. \ x \ \<^bold>V; x \ source \ \ \ d_OUT (capacity \) x < \" proof - interpret ae: countable_network \'' by(rule \''_countable_network) show ?thesis proof fix x assume x: "x \ \<^bold>V\<^bsub>\''\<^esub>" and not_source: "x \ source \''" from x consider (Vertex) x' where "x = Vertex x'" "x' \ \<^bold>V" | (Edge) y z where "x = Edge y z" "edge \ y z" unfolding "\<^bold>V_\''" by auto then show "d_OUT (capacity \'') x < \" proof cases case Vertex then show ?thesis using x not_source that[of x'] by auto next case Edge then show ?thesis using capacity_finite[of "(y, z)"] by(simp del: capacity_finite add: less_top) qed qed qed end context bounded_countable_network begin lemma bounded_countable_web_web_of_network: assumes source_in: "\x. \ edge \ x (source \)" and sink_out: "\y. \ edge \ (sink \) y" and undead: "\x y. edge \ x y \ (\z. edge \ y z) \ (\z. edge \ z x)" and source_sink: "\ edge \ (source \) (sink \)" and no_loop: "\x. \ edge \ x x" - and edge_antiparallel: "\x y. edge \ x y \ \ edge \ y x" shows "bounded_countable_web (web_of_network \)" (is "bounded_countable_web ?\") proof - interpret web: countable_web ?\ by(rule countable_web_web_of_network) fact+ show ?thesis proof fix e assume "e \ \<^bold>V\<^bsub>?\\<^esub> - B ?\" then obtain x y where e: "e = (x, y)" and xy: "edge \ x y" by(cases e) simp from xy have y: "y \ source \" using source_in[of x] by auto have "\<^bold>E\<^bsub>?\\<^esub> `` {e} \ \<^bold>E \ {y} \ UNIV" using e by auto hence "(\\<^sup>+ e' \ \<^bold>E\<^bsub>?\\<^esub> `` {e}. weight ?\ e') \ (\\<^sup>+ e \ \<^bold>E \ {y} \ UNIV. capacity \ e)" using e by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ \ (\\<^sup>+ e \ Pair y ` UNIV. capacity \ e)" by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ = d_OUT (capacity \) y" unfolding d_OUT_def by(rule nn_integral_count_space_reindex) simp also have "\ < \" using out[of y] xy y by(auto simp add: vertex_def) finally show "(\\<^sup>+ e' \ \<^bold>E\<^bsub>?\\<^esub> `` {e}. weight ?\ e') < \" . qed qed context begin qualified lemma max_flow_min_cut'_bounded: assumes source_in: "\x. \ edge \ x (source \)" and sink_out: "\y. \ edge \ (sink \) y" and undead: "\x y. edge \ x y \ (\z. edge \ y z) \ (\z. edge \ z x)" and source_sink: "\ edge \ (source \) (sink \)" and no_loop: "\x. \ edge \ x x" - and edge_antiparallel: "\x y. edge \ x y \ \ edge \ y x" and capacity_pos: "\e. e \ \<^bold>E \ capacity \ e > 0" shows "\f S. flow \ f \ cut \ S \ orthogonal \ f S" by(rule max_flow_min_cut')(rule bounded_countable_web.ex_orthogonal_current[OF bounded_countable_web_web_of_network], fact+) qualified lemma max_flow_min_cut''_bounded: 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 \ \<^bold>E \ capacity \ e > 0" shows "\f S. flow \ f \ cut \ S \ orthogonal \ f S" proof - interpret antiparallel_edges \ .. interpret \'': bounded_countable_network \'' by(rule \''_bounded_countable_network)(rule out) have "\f S. flow \'' f \ cut \'' S \ orthogonal \'' f S" by(rule \''.max_flow_min_cut'_bounded)(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'''_bounded: assumes sink_out: "\y. \ edge \ (sink \) y" and source_in: "\x. \ edge \ x (source \)" and capacity_pos: "\e. e \ \<^bold>E \ capacity \ e > 0" shows "\f S. flow \ f \ cut \ S \ orthogonal \ f S" proof - interpret antiparallel_edges \ .. interpret \'': bounded_countable_network \'' by(rule \''_bounded_countable_network)(rule out) have "\f S. flow \'' f \ cut \'' S \ orthogonal \'' f S" by(rule \''.max_flow_min_cut''_bounded)(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 lemma \'''_bounded_countable_network: "bounded_countable_network \'''" proof - interpret \''': countable_network \''' by(rule \'''_countable_network) show ?thesis proof fix x assume x: "x \ \<^bold>V\<^bsub>\'''\<^esub>" and not_source: "x \ source \'''" from x have x': "x \ \<^bold>V" by(auto simp add: vertex_def) have "d_OUT (capacity \''') x \ d_OUT (capacity \) x" by(rule d_OUT_mono) simp also have "\ < \" using x' not_source by(simp add: out) finally show "d_OUT (capacity \''') x < \" . qed qed theorem max_flow_min_cut_bounded: "\f S. flow \ f \ cut \ S \ orthogonal \ f S" proof - interpret \': bounded_countable_network \''' by(rule \'''_bounded_countable_network) have "\f S. flow \''' f \ cut \''' S \ orthogonal \''' f S" by(rule \'.max_flow_min_cut'''_bounded) 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 diff --git a/thys/MFMC_Countable/MFMC_Reduction.thy b/thys/MFMC_Countable/MFMC_Reduction.thy --- a/thys/MFMC_Countable/MFMC_Reduction.thy +++ b/thys/MFMC_Countable/MFMC_Reduction.thy @@ -1,1379 +1,1376 @@ section \Reductions\ theory MFMC_Reduction imports MFMC_Web begin subsection \From a web to a bipartite web\ definition bipartite_web_of :: "('v, 'more) web_scheme \ ('v + 'v, 'more) web_scheme" where "bipartite_web_of \ = \edge = \uv uv'. case (uv, uv') of (Inl u, Inr v) \ edge \ u v \ u = v \ u \ vertices \ \ u \ A \ \ v \ B \ | _ \ False, weight = \uv. case uv of Inl u \ if u \ B \ then 0 else weight \ u | Inr u \ if u \ A \ then 0 else weight \ u, A = Inl ` (vertices \ - B \), B = Inr ` (- A \), \ = web.more \\" lemma bipartite_web_of_sel [simp]: fixes \ (structure) shows "edge (bipartite_web_of \) (Inl u) (Inr v) \ edge \ u v \ u = v \ u \ \<^bold>V \ u \ A \ \ v \ B \" "edge (bipartite_web_of \) uv (Inl u) \ False" "edge (bipartite_web_of \) (Inr v) uv \ False" "weight (bipartite_web_of \) (Inl u) = (if u \ B \ then 0 else weight \ u)" "weight (bipartite_web_of \) (Inr v) = (if v \ A \ then 0 else weight \ v)" "A (bipartite_web_of \) = Inl ` (\<^bold>V - B \)" "B (bipartite_web_of \) = Inr ` (- A \)" by(simp_all add: bipartite_web_of_def split: sum.split) lemma edge_bipartite_webI1: "edge \ u v \ edge (bipartite_web_of \) (Inl u) (Inr v)" by(auto) lemma edge_bipartite_webI2: "\ u \ \<^bold>V\<^bsub>\\<^esub>; u \ A \; u \ B \ \ \ edge (bipartite_web_of \) (Inl u) (Inr u)" by(auto) lemma edge_bipartite_webE: fixes \ (structure) assumes "edge (bipartite_web_of \) uv uv'" obtains u v where "uv = Inl u" "uv' = Inr v" "edge \ u v" | u where "uv = Inl u" "uv' = Inr u" "u \ \<^bold>V" "u \ A \" "u \ B \" using assms by(cases uv uv' rule: sum.exhaust[case_product sum.exhaust]) auto lemma E_bipartite_web: fixes \ (structure) shows "\<^bold>E\<^bsub>bipartite_web_of \\<^esub> = (\(x, y). (Inl x, Inr y)) ` \<^bold>E \ (\x. (Inl x, Inr x)) ` (\<^bold>V - A \ - B \)" by(auto elim: edge_bipartite_webE) context web begin lemma vertex_bipartite_web [simp]: "vertex (bipartite_web_of \) (Inl x) \ vertex \ x \ x \ B \" "vertex (bipartite_web_of \) (Inr x) \ vertex \ x \ x \ A \" by(auto 4 4 simp add: vertex_def dest: B_out A_in intro: edge_bipartite_webI1 edge_bipartite_webI2 elim!: edge_bipartite_webE) definition separating_of_bipartite :: "('v + 'v) set \ 'v set" where "separating_of_bipartite S = (let A_S = Inl -` S; B_S = Inr -` S in (A_S \ B_S) \ (A \ \ A_S) \ (B \ \ B_S))" context fixes S :: "('v + 'v) set" assumes sep: "separating (bipartite_web_of \) S" begin text \Proof of separation follows @{cite Aharoni1983EJC}\ lemma separating_of_bipartite_aux: assumes p: "path \ x p y" and y: "y \ B \" and x: "x \ A \ \ Inr x \ S" shows "(\z\set p. z \ separating_of_bipartite S) \ x \ separating_of_bipartite S" proof(cases "p = []") case True with p have "x = y" by cases auto with y x have "x \ separating_of_bipartite S" using disjoint by(auto simp add: separating_of_bipartite_def Let_def) thus ?thesis .. next case nNil: False define inmarked where "inmarked x \ x \ A \ \ Inr x \ S" for x define outmarked where "outmarked x \ x \ B \ \ Inl x \ S" for x let ?\ = "bipartite_web_of \" let ?double = "\x. inmarked x \ outmarked x" define tailmarked where "tailmarked = (\(x, y :: 'v). outmarked x)" define headmarked where "headmarked = (\(x :: 'v, y). inmarked y)" have marked_E: "tailmarked e \ headmarked e" if "e \ \<^bold>E" for e \ \Lemma 1b\ proof(cases e) case (Pair x y) with that have "path ?\ (Inl x) [Inr y] (Inr y)" by(auto intro!: rtrancl_path.intros) from separatingD[OF sep this] that Pair show ?thesis by(fastforce simp add: vertex_def inmarked_def outmarked_def tailmarked_def headmarked_def) qed have "\z\set (x # p). ?double z" \ \Lemma 2\ proof - have "inmarked ((x # p) ! (i + 1)) \ outmarked ((x # p) ! i)" if "i < length p" for i using rtrancl_path_nth[OF p that] marked_E[of "((x # p) ! i, p ! i)"] that by(auto simp add: tailmarked_def headmarked_def nth_Cons split: nat.split) hence "{i. i < length p \ inmarked (p ! i)} \ {i. i < length (x # butlast p) \ outmarked ((x # butlast p) ! i)} = {i. i < length p}" (is "?in \ ?out = _") using nNil by(force simp add: nth_Cons' nth_butlast elim: meta_allE[where x=0] cong del: old.nat.case_cong_weak) hence "length p + 2 = card (?in \ ?out) + 2" by simp also have "\ \ (card ?in + 1) + (card ?out + 1)" by(simp add: card_Un_le) also have "card ?in = card ((\i. Inl (i + 1) :: _ + nat) ` ?in)" by(rule card_image[symmetric])(simp add: inj_on_def) also have "\ + 1 = card (insert (Inl 0) {Inl (Suc i) :: _ + nat|i. i < length p \ inmarked (p ! i)})" by(subst card_insert_if)(auto intro!: arg_cong[where f=card]) also have "\ = card {Inl i :: nat + nat|i. i < length (x # p) \ inmarked ((x # p) ! i)}" (is "_ = card ?in") using x by(intro arg_cong[where f=card])(auto simp add: nth_Cons inmarked_def split: nat.split_asm) also have "card ?out = card ((Inr :: _ \ nat + _) ` ?out)" by(simp add: card_image) also have "\ + 1 = card (insert (Inr (length p)) {Inr i :: nat + _|i. i < length p \ outmarked ((x # p) ! i)})" using nNil by(subst card_insert_if)(auto intro!: arg_cong[where f=card] simp add: nth_Cons nth_butlast cong: nat.case_cong) also have "\ = card {Inr i :: nat + _|i. i < length (x # p) \ outmarked ((x # p) ! i)}" (is "_ = card ?out") using nNil rtrancl_path_last[OF p nNil] y by(intro arg_cong[where f=card])(auto simp add: outmarked_def last_conv_nth) also have "card ?in + card ?out = card (?in \ ?out)" by(rule card_Un_disjoint[symmetric]) auto also let ?f = "case_sum id id" have "?f ` (?in \ ?out) \ {i. i < length (x # p)}" by auto from card_mono[OF _ this] have "card (?f ` (?in \ ?out)) \ length p + 1" by simp ultimately have "\ inj_on ?f (?in \ ?out)" by(intro pigeonhole) simp then obtain i where "i < length (x # p)" "?double ((x # p) ! i)" by(auto simp add: inj_on_def) thus ?thesis by(auto simp add: set_conv_nth) qed moreover have "z \ separating_of_bipartite S" if "?double z" for z using that disjoint by(auto simp add: separating_of_bipartite_def Let_def inmarked_def outmarked_def) ultimately show ?thesis by auto qed lemma separating_of_bipartite: "separating \ (separating_of_bipartite S)" by(rule separating_gen.intros)(erule (1) separating_of_bipartite_aux; simp) end lemma current_bipartite_web_finite: assumes f: "current (bipartite_web_of \) f" (is "current ?\ _") shows "f e \ \" proof(cases e) case (Pair x y) have "f e \ d_OUT f x" unfolding Pair d_OUT_def by(rule nn_integral_ge_point) simp also have "\ \ weight ?\ x" by(rule currentD_weight_OUT[OF f]) also have "\ < \" by(cases x)(simp_all add: less_top[symmetric]) finally show ?thesis by simp qed definition current_of_bipartite :: "('v + 'v) current \ 'v current" where "current_of_bipartite f = (\(x, y). f (Inl x, Inr y) * indicator \<^bold>E (x, y))" lemma current_of_bipartite_simps [simp]: "current_of_bipartite f (x, y) = f (Inl x, Inr y) * indicator \<^bold>E (x, y)" by(simp add: current_of_bipartite_def) lemma d_OUT_current_of_bipartite: assumes f: "current (bipartite_web_of \) f" shows "d_OUT (current_of_bipartite f) x = d_OUT f (Inl x) - f (Inl x, Inr x)" proof - have "d_OUT (current_of_bipartite f) x = \\<^sup>+ y. f (Inl x, y) * indicator \<^bold>E (x, projr y) \count_space (range Inr)" by(simp add: d_OUT_def nn_integral_count_space_reindex) also have "\ = d_OUT f (Inl x) - \\<^sup>+ y. f (Inl x, y) * indicator {Inr x} y \count_space UNIV" (is "_ = _ - ?rest") unfolding d_OUT_def by(subst nn_integral_diff[symmetric])(auto 4 4 simp add: current_bipartite_web_finite[OF f] AE_count_space nn_integral_count_space_indicator no_loop split: split_indicator intro!: nn_integral_cong intro: currentD_outside[OF f] elim: edge_bipartite_webE) finally show ?thesis by simp qed lemma d_IN_current_of_bipartite: assumes f: "current (bipartite_web_of \) f" shows "d_IN (current_of_bipartite f) x = d_IN f (Inr x) - f (Inl x, Inr x)" proof - have "d_IN (current_of_bipartite f) x = \\<^sup>+ y. f (y, Inr x) * indicator \<^bold>E (projl y, x) \count_space (range Inl)" by(simp add: d_IN_def nn_integral_count_space_reindex) also have "\ = d_IN f (Inr x) - \\<^sup>+ y. f (y, Inr x) * indicator {Inl x} y \count_space UNIV" (is "_ = _ - ?rest") unfolding d_IN_def by(subst nn_integral_diff[symmetric])(auto 4 4 simp add: current_bipartite_web_finite[OF f] AE_count_space nn_integral_count_space_indicator no_loop split: split_indicator intro!: nn_integral_cong intro: currentD_outside[OF f] elim: edge_bipartite_webE) finally show ?thesis by simp qed lemma current_current_of_bipartite: \ \Lemma 6.3\ assumes f: "current (bipartite_web_of \) f" (is "current ?\ _") and w: "wave (bipartite_web_of \) f" shows "current \ (current_of_bipartite f)" (is "current _ ?f") proof fix x have "d_OUT ?f x \ d_OUT f (Inl x)" by(simp add: d_OUT_current_of_bipartite[OF f] diff_le_self_ennreal) also have "\ \ weight \ x" using currentD_weight_OUT[OF f, of "Inl x"] by(simp split: if_split_asm) finally show "d_OUT ?f x \ weight \ x" . next fix x have "d_IN ?f x \ d_IN f (Inr x)" by(simp add: d_IN_current_of_bipartite[OF f] diff_le_self_ennreal) also have "\ \ weight \ x" using currentD_weight_IN[OF f, of "Inr x"] by(simp split: if_split_asm) finally show "d_IN ?f x \ weight \ x" . next have OUT: "d_OUT ?f b = 0" if "b \ B \" for b using that by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 intro!: currentD_outside[OF f] dest: B_out) show "d_OUT ?f x \ d_IN ?f x" if A: "x \ A \" for x proof(cases "x \ B \ \ x \ \<^bold>V") case True then show ?thesis proof assume "x \ B \" with OUT[OF this] show ?thesis by auto next assume "x \ \<^bold>V" hence "d_OUT ?f x = 0" by(auto simp add: d_OUT_def vertex_def nn_integral_0_iff emeasure_count_space_eq_0 intro!: currentD_outside[OF f]) thus ?thesis by simp qed next case B [simplified]: False have "d_OUT ?f x = d_OUT f (Inl x) - f (Inl x, Inr x)" (is "_ = _ - ?rest") by(simp add: d_OUT_current_of_bipartite[OF f]) also have "d_OUT f (Inl x) \ d_IN f (Inr x)" proof(rule ccontr) assume "\ ?thesis" hence *: "d_IN f (Inr x) < d_OUT f (Inl x)" by(simp add: not_less) also have "\ \ weight \ x" using currentD_weight_OUT[OF f, of "Inl x"] B by simp finally have "Inr x \ TER\<^bsub>?\\<^esub> f" using A by(auto elim!: SAT.cases) moreover have "Inl x \ TER\<^bsub>?\\<^esub> f" using * by(auto simp add: SINK.simps) moreover have "path ?\ (Inl x) [Inr x] (Inr x)" by(rule rtrancl_path.step)(auto intro!: rtrancl_path.base simp add: no_loop A B) ultimately show False using waveD_separating[OF w] A B by(auto dest!: separatingD) qed hence "d_OUT f (Inl x) - ?rest \ d_IN f (Inr x) - ?rest" by(rule ennreal_minus_mono) simp also have "\ = d_IN ?f x" by(simp add: d_IN_current_of_bipartite[OF f]) finally show ?thesis . qed show "?f e = 0" if "e \ \<^bold>E" for e using that by(cases e)(auto) qed lemma TER_current_of_bipartite: \ \Lemma 6.3\ assumes f: "current (bipartite_web_of \) f" (is "current ?\ _") and w: "wave (bipartite_web_of \) f" shows "TER (current_of_bipartite f) = separating_of_bipartite (TER\<^bsub>bipartite_web_of \\<^esub> f)" (is "TER ?f = separating_of_bipartite ?TER") proof(rule set_eqI) fix x consider (A) "x \ A \" "x \ \<^bold>V" | (B) "x \ B \" "x \ \<^bold>V" | (inner) "x \ A \" "x \ B \" "x \ \<^bold>V" | (outside) "x \ \<^bold>V" by auto thus "x \ TER ?f \ x \ separating_of_bipartite ?TER" proof cases case A hence "d_OUT ?f x = d_OUT f (Inl x)" using currentD_outside[OF f, of "Inl x" "Inr x"] by(simp add: d_OUT_current_of_bipartite[OF f] no_loop) thus ?thesis using A disjoint by(auto simp add: separating_of_bipartite_def Let_def SINK.simps intro!: SAT.A imageI) next case B then have "d_IN ?f x = d_IN f (Inr x)" using currentD_outside[OF f, of "Inl x" "Inr x"] by(simp add: d_IN_current_of_bipartite[OF f] no_loop) moreover have "d_OUT ?f x = 0" using B currentD_outside[OF f, of "Inl x" "Inr x"] by(simp add: d_OUT_current_of_bipartite[OF f] no_loop)(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 intro!: currentD_outside[OF f] elim!: edge_bipartite_webE dest: B_out) moreover have "d_OUT f (Inr x) = 0" using B disjoint by(intro currentD_OUT[OF f]) auto ultimately show ?thesis using B by(auto simp add: separating_of_bipartite_def Let_def SINK.simps SAT.simps) next case outside with current_current_of_bipartite[OF f w] have "d_OUT ?f x = 0" "d_IN ?f x = 0" by(rule currentD_outside_OUT currentD_outside_IN)+ moreover from outside have "Inl x \ vertices ?\" "Inr x \ vertices ?\" by auto hence "d_OUT f (Inl x) = 0" "d_IN f (Inl x) = 0" "d_OUT f (Inr x) = 0" "d_IN f (Inr x) = 0" by(blast intro: currentD_outside_OUT[OF f] currentD_outside_IN[OF f])+ ultimately show ?thesis using outside weight_outside[of x] by(auto simp add: separating_of_bipartite_def Let_def SINK.simps SAT.simps not_le) next case inner show ?thesis proof assume "x \ separating_of_bipartite ?TER" with inner have l: "Inl x \ ?TER" and r: "Inr x \ ?TER" by(auto simp add: separating_of_bipartite_def Let_def) have "f (Inl x, Inr x) \ d_OUT f (Inl x)" unfolding d_OUT_def by(rule nn_integral_ge_point) simp with l have 0: "f (Inl x, Inr x) = 0" by(auto simp add: SINK.simps) with l have "x \ SINK ?f" by(simp add: SINK.simps d_OUT_current_of_bipartite[OF f]) moreover from r have "Inr x \ SAT ?\ f" by auto with inner have "x \ SAT \ ?f" by(auto elim!: SAT.cases intro!: SAT.IN simp add: 0 d_IN_current_of_bipartite[OF f]) ultimately show "x \ TER ?f" by simp next assume *: "x \ TER ?f" have "d_IN f (Inr x) \ weight ?\ (Inr x)" using f by(rule currentD_weight_IN) also have "\ \ weight \ x" using inner by simp also have "\ \ d_IN ?f x" using inner * by(auto elim: SAT.cases) also have "\ \ d_IN f (Inr x)" by(simp add: d_IN_current_of_bipartite[OF f] max_def diff_le_self_ennreal) ultimately have eq: "d_IN ?f x = d_IN f (Inr x)" by simp hence 0: "f (Inl x, Inr x) = 0" using ennreal_minus_cancel_iff[of "d_IN f (Inr x)" "f (Inl x, Inr x)" 0] currentD_weight_IN[OF f, of "Inr x"] inner d_IN_ge_point[of f "Inl x" "Inr x"] by(auto simp add: d_IN_current_of_bipartite[OF f] top_unique) have "Inl x \ ?TER" "Inr x \ ?TER" using inner * currentD_OUT[OF f, of "Inr x"] by(auto simp add: SAT.simps SINK.simps d_OUT_current_of_bipartite[OF f] 0 eq) thus "x \ separating_of_bipartite ?TER" unfolding separating_of_bipartite_def Let_def by blast qed qed qed lemma wave_current_of_bipartite: \ \Lemma 6.3\ assumes f: "current (bipartite_web_of \) f" (is "current ?\ _") and w: "wave (bipartite_web_of \) f" shows "wave \ (current_of_bipartite f)" (is "wave _ ?f") proof have sep': "separating \ (separating_of_bipartite (TER\<^bsub>?\\<^esub> f))" by(rule separating_of_bipartite)(rule waveD_separating[OF w]) then show sep: "separating \ (TER (current_of_bipartite f))" by(simp add: TER_current_of_bipartite[OF f w]) fix x assume "x \ RF (TER ?f)" then obtain p y where p: "path \ x p y" and y: "y \ B \" and x: "x \ TER ?f" and bypass: "\z. z \ set p \ z \ TER ?f" by(auto simp add: roofed_def elim: rtrancl_path_distinct) from p obtain p' where p': "path \ x p' y" and distinct: "distinct (x # p')" and subset: "set p' \ set p" by(auto elim: rtrancl_path_distinct) consider (outside) "x \ \<^bold>V" | (A) "x \ A \" | (B) "x \ B \" | (inner) "x \ \<^bold>V" "x \ A \" "x \ B \" by auto then show "d_OUT ?f x = 0" proof cases case outside with f w show ?thesis by(rule currentD_outside_OUT[OF current_current_of_bipartite]) next case A from separatingD[OF sep p A y] bypass have "x \ TER ?f" by blast thus ?thesis by(simp add: SINK.simps) next case B with f w show ?thesis by(rule currentD_OUT[OF current_current_of_bipartite]) next case inner hence "path ?\ (Inl x) [Inr x] (Inr x)" by(auto intro!: rtrancl_path.intros) from inner waveD_separating[OF w, THEN separatingD, OF this] consider (Inl) "Inl x \ TER\<^bsub>?\\<^esub> f" | (Inr) "Inr x \ TER\<^bsub>?\\<^esub> f" by auto then show ?thesis proof cases case Inl thus ?thesis by(auto simp add: SINK.simps d_OUT_current_of_bipartite[OF f] max_def) next case Inr with separating_of_bipartite_aux[OF waveD_separating[OF w] p y] x bypass have False unfolding TER_current_of_bipartite[OF f w] by blast thus ?thesis .. qed qed qed end context countable_web begin lemma countable_bipartite_web_of: "countable_bipartite_web (bipartite_web_of \)" (is "countable_bipartite_web ?\") proof show "\<^bold>V\<^bsub>?\\<^esub> \ A ?\ \ B ?\" apply(rule subsetI) subgoal for x by(cases x) auto done show "A ?\ \ \<^bold>V\<^bsub>?\\<^esub>" by auto show "x \ A ?\ \ y \ B ?\" if "edge ?\ x y" for x y using that by(cases x y rule: sum.exhaust[case_product sum.exhaust])(auto simp add: inj_image_mem_iff vertex_def B_out A_in) show "A ?\ \ B ?\ = {}" by auto show "countable \<^bold>E\<^bsub>?\\<^esub>" by(simp add: E_bipartite_web) show "weight ?\ x \ \" for x by(cases x) simp_all show "weight (bipartite_web_of \) x = 0" if "x \ \<^bold>V\<^bsub>?\\<^esub>" for x using that by(cases x)(auto simp add: weight_outside) qed end context web begin lemma unhindered_bipartite_web_of: assumes loose: "loose \" shows "\ hindered (bipartite_web_of \)" proof assume "hindered (bipartite_web_of \)" (is "hindered ?\") then obtain f where f: "current ?\ f" and w: "wave ?\ f" and hind: "hindrance ?\ f" by cases from f w have "current \ (current_of_bipartite f)" by(rule current_current_of_bipartite) moreover from f w have "wave \ (current_of_bipartite f)" by(rule wave_current_of_bipartite) ultimately have *: "current_of_bipartite f = zero_current" by(rule looseD_wave[OF loose]) have zero: "f (Inl x, Inr y) = 0" if "x \ y" for x y using that *[THEN fun_cong, of "(x, y)"] by(cases "edge \ x y")(auto intro: currentD_outside[OF f]) have OUT: "d_OUT f (Inl x) = f (Inl x, Inr x)" for x proof - have "d_OUT f (Inl x) = (\\<^sup>+ y. f (Inl x, y) * indicator {Inr x} y)" unfolding d_OUT_def using zero currentD_outside[OF f] apply(intro nn_integral_cong) subgoal for y by(cases y)(auto split: split_indicator) done also have "\ = f (Inl x, Inr x)" by simp finally show ?thesis . qed have IN: "d_IN f (Inr x) = f (Inl x, Inr x)" for x proof - have "d_IN f (Inr x) = (\\<^sup>+ y. f (y, Inr x) * indicator {Inl x} y)" unfolding d_IN_def using zero currentD_outside[OF f] apply(intro nn_integral_cong) subgoal for y by(cases y)(auto split: split_indicator) done also have "\ = f (Inl x, Inr x)" by simp finally show ?thesis . qed let ?TER = "TER\<^bsub>?\\<^esub> f" from hind obtain a where a: "a \ A ?\" and n\: "a \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> f)" and OUT_a: "d_OUT f a < weight ?\ a" by cases from a obtain a' where a': "a = Inl a'" and v: "vertex \ a'" and b: "a' \ B \" by auto have A: "a' \ A \" proof(rule ccontr) assume A: "a' \ A \" hence "edge ?\ (Inl a') (Inr a')" using v b by simp hence p: "path ?\ (Inl a') [Inr a'] (Inr a')" by(simp add: rtrancl_path_simps) from separatingD[OF waveD_separating[OF w] this] b v A have "Inl a' \ ?TER \ Inr a' \ ?TER" by auto thus False proof cases case left hence "d_OUT f (Inl a') = 0" by(simp add: SINK.simps) moreover hence "d_IN f (Inr a') = 0" by(simp add: OUT IN) ultimately have "Inr a' \ ?TER" using v b A OUT_a a' by(auto simp add: SAT.simps) then have "essential ?\ (B ?\) ?TER (Inl a')" using A by(intro essentialI[OF p]) simp_all with n\ left a' show False by simp next case right hence "d_IN f (Inr a') = weight \ a'" using A by(auto simp add: currentD_SAT[OF f]) hence "d_OUT f (Inl a') = weight \ a'" by(simp add: IN OUT) with OUT_a a' b show False by simp qed qed moreover from A have "d_OUT f (Inl a') = 0" using currentD_outside[OF f, of "Inl a'" "Inr a'"] by(simp add: OUT no_loop) with b v have TER: "Inl a' \ ?TER" by(simp add: SAT.A SINK.simps) with n\ a' have ness: "\ essential ?\ (B ?\) ?TER (Inl a')" by simp have "a' \ \ (TER zero_current)" proof assume "a' \ \ (TER zero_current)" then obtain p y where p: "path \ a' p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ TER zero_current" by(rule \_E_RF)(auto intro: roofed_greaterI) from p show False proof cases case base with y A disjoint show False by auto next case (step x p') from step(2) have "path ?\ (Inl a') [Inr x] (Inr x)" by(simp add: rtrancl_path_simps) from not_essentialD[OF ness this] bypass[of x] step(1) have "Inr x \ ?TER" by simp with bypass[of x] step(1) have "d_IN f (Inr x) > 0" by(auto simp add: currentD_SAT[OF f] zero_less_iff_neq_zero) hence x: "Inl x \ ?TER" by(auto simp add: SINK.simps OUT IN) from step(1) have "set (x # p') \ set p" by auto with \path \ x p' y\ x y show False proof induction case (base x) thus False using currentD_outside_IN[OF f, of "Inl x"] currentD_outside_OUT[OF f, of "Inl x"] by(auto simp add: currentD_SAT[OF f] SINK.simps dest!: bypass) next case (step x z p' y) from step.prems(3) bypass[of x] weight_outside[of x] have x: "vertex \ x" by(auto) from \edge \ x z\ have "path ?\ (Inl x) [Inr z] (Inr z)" by(simp add: rtrancl_path_simps) from separatingD[OF waveD_separating[OF w] this] step.prems(1) step.prems(3) bypass[of z] x \edge \ x z\ have "Inr z \ ?TER" by(force simp add: B_out inj_image_mem_iff) with bypass[of z] step.prems(3) \edge \ x z\ have "d_IN f (Inr z) > 0" by(auto simp add: currentD_SAT[OF f] A_in zero_less_iff_neq_zero) hence x: "Inl z \ ?TER" by(auto simp add: SINK.simps OUT IN) with step.IH[OF this] step.prems(2,3) show False by auto qed qed qed moreover have "d_OUT zero_current a' < weight \ a'" using OUT_a a' b by (auto simp: zero_less_iff_neq_zero) ultimately have "hindrance \ zero_current" by(rule hindrance) with looseD_hindrance[OF loose] show False by contradiction qed lemma (in -) divide_less_1_iff_ennreal: "a / b < (1::ennreal) \ (0 < b \ a < b \ b = 0 \ a = 0 \ b = top)" by (cases a; cases b; cases "b = 0") (auto simp: divide_ennreal ennreal_less_iff ennreal_top_divide) lemma linkable_bipartite_web_ofD: assumes link: "linkable (bipartite_web_of \)" (is "linkable ?\") and countable: "countable \<^bold>E" shows "linkable \" proof - from link obtain f where wf: "web_flow ?\ f" and link: "linkage ?\ f" by blast from wf have f: "current ?\ f" by(rule web_flowD_current) define f' where "f' = current_of_bipartite f" have IN_le_OUT: "d_IN f' x \ d_OUT f' x" if "x \ B \" for x proof(cases "x \ \<^bold>V") case True have "d_IN f' x = d_IN f (Inr x) - f (Inl x, Inr x)" (is "_ = _ - ?rest") by(simp add: f'_def d_IN_current_of_bipartite[OF f]) also have "\ \ weight ?\ (Inr x) - ?rest" using currentD_weight_IN[OF f, of "Inr x"] by(rule ennreal_minus_mono) simp also have "\ \ weight ?\ (Inl x) - ?rest" using that ennreal_minus_mono by(auto) also have "weight ?\ (Inl x) = d_OUT f (Inl x)" using that linkageD[OF link, of "Inl x"] True by auto also have "\ - ?rest = d_OUT f' x" by(simp add: f'_def d_OUT_current_of_bipartite[OF f]) finally show ?thesis . next case False with currentD_outside_OUT[OF f, of "Inl x"] currentD_outside_IN[OF f, of "Inr x"] show ?thesis by(simp add: f'_def d_IN_current_of_bipartite[OF f] d_OUT_current_of_bipartite[OF f]) qed have link: "linkage \ f'" proof show "d_OUT f' a = weight \ a" if "a \ A \" for a proof(cases "a \ \<^bold>V") case True from that have "a \ B \" using disjoint by auto with that True linkageD[OF link, of "Inl a"] ennreal_minus_cancel_iff[of _ _ 0] currentD_outside[OF f, of "Inl a" "Inr a"] show ?thesis by(clarsimp simp add: f'_def d_OUT_current_of_bipartite[OF f] max_def no_loop) next case False with weight_outside[OF this] currentD_outside[OF f, of "Inl a" "Inr a"] currentD_outside_OUT[OF f, of "Inl a"] show ?thesis by(simp add: f'_def d_OUT_current_of_bipartite[OF f] no_loop) qed qed define F where "F = {g. (\e. 0 \ g e) \ (\e. e \ \<^bold>E \ g e = 0) \ (\x. x \ B \ \ d_IN g x \ d_OUT g x) \ linkage \ g \ (\x\A \. d_IN g x = 0) \ (\x. d_OUT g x \ weight \ x) \ (\x. d_IN g x \ weight \ x) \ (\x\B \. d_OUT g x = 0) \ g \ f'}" define leq where "leq = restrict_rel F {(f, f'). f' \ f}" have F: "Field leq = F" by(auto simp add: leq_def) have F_I [intro?]: "f \ Field leq" if "\e. 0 \ f e" and "\e. e \ \<^bold>E \ f e = 0" and "\x. x \ B \ \ d_IN f x \ d_OUT f x" and "linkage \ f" and "\x. x \ A \ \ d_IN f x = 0" and "\x. d_OUT f x \ weight \ x" and "\x. d_IN f x \ weight \ x" and "\x. x \ B \ \ d_OUT f x = 0" and "f \ f'" for f using that by(simp add: F F_def) have F_nonneg: "0 \ f e" if "f \ Field leq" for f e using that by(cases e)(simp add: F F_def) have F_outside: "f e = 0" if "f \ Field leq" "e \ \<^bold>E" for f e using that by(cases e)(simp add: F F_def) have F_IN_OUT: "d_IN f x \ d_OUT f x" if "f \ Field leq" "x \ B \" for f x using that by(simp add: F F_def) have F_link: "linkage \ f" if "f \ Field leq" for f using that by(simp add: F F_def) have F_IN: "d_IN f x = 0" if "f \ Field leq" "x \ A \" for f x using that by(simp add: F F_def) have F_OUT: "d_OUT f x = 0" if "f \ Field leq" "x \ B \" for f x using that by(simp add: F F_def) have F_weight_OUT: "d_OUT f x \ weight \ x" if "f \ Field leq" for f x using that by(simp add: F F_def) have F_weight_IN: "d_IN f x \ weight \ x" if "f \ Field leq" for f x using that by(simp add: F F_def) have F_le: "f e \ f' e" if "f \ Field leq" for f e using that by(cases e)(simp add: F F_def le_fun_def) have F_finite_OUT: "d_OUT f x \ \" if "f \ Field leq" for f x proof - have "d_OUT f x \ weight \ x" by(rule F_weight_OUT[OF that]) also have "\ < \" by(simp add: less_top[symmetric]) finally show ?thesis by simp qed have F_finite: "f e \ \" if "f \ Field leq" for f e proof(cases e) case (Pair x y) have "f e \ d_OUT f x" unfolding Pair d_OUT_def by(rule nn_integral_ge_point) simp also have "\ < \" by(simp add: F_finite_OUT[OF that] less_top[symmetric]) finally show ?thesis by simp qed have f': "f' \ Field leq" proof show "0 \ f' e" for e by(cases e)(simp add: f'_def) show "f' e = 0" if "e \ \<^bold>E" for e using that by(clarsimp split: split_indicator_asm simp add: f'_def) show "d_IN f' x \ d_OUT f' x" if "x \ B \" for x using that by(rule IN_le_OUT) show "linkage \ f'" by(rule link) show "d_IN f' x = 0" if "x \ A \" for x using that currentD_IN[OF f, of "Inl x"] disjoint currentD_outside[OF f, of "Inl x" "Inr x"] currentD_outside_IN[OF f, of "Inr x"] by(cases "x \ \<^bold>V")(auto simp add: d_IN_current_of_bipartite[OF f] no_loop f'_def) show "d_OUT f' x = 0" if "x \ B \" for x using that currentD_OUT[OF f, of "Inr x"] disjoint currentD_outside[OF f, of "Inl x" "Inr x"] currentD_outside_OUT[OF f, of "Inl x"] by(cases "x \ \<^bold>V")(auto simp add: d_OUT_current_of_bipartite[OF f] no_loop f'_def) show "d_OUT f' x \ weight \ x" for x using currentD_weight_OUT[OF f, of "Inl x"] by(simp add: d_OUT_current_of_bipartite[OF f] ennreal_diff_le_mono_left f'_def split: if_split_asm) show "d_IN f' x \ weight \ x" for x using currentD_weight_IN[OF f, of "Inr x"] by(simp add: d_IN_current_of_bipartite[OF f] ennreal_diff_le_mono_left f'_def split: if_split_asm) qed simp have F_leI: "g \ Field leq" if f: "f \ Field leq" and le: "\e. g e \ f e" and nonneg: "\e. 0 \ g e" and IN_OUT: "\x. x \ B \ \ d_IN g x \ d_OUT g x" and link: "linkage \ g" for f g proof show "g e = 0" if "e \ \<^bold>E" for e using nonneg[of e] F_outside[OF f that] le[of e] by simp show "d_IN g a = 0" if "a \ A \" for a using d_IN_mono[of g a f, OF le] F_IN[OF f that] by auto show "d_OUT g b = 0" if "b \ B \" for b using d_OUT_mono[of g b f, OF le] F_OUT[OF f that] by auto show "d_OUT g x \ weight \ x" for x using d_OUT_mono[of g x f, OF le] F_weight_OUT[OF f] by(rule order_trans) show "d_IN g x \ weight \ x" for x using d_IN_mono[of g x f, OF le] F_weight_IN[OF f] by(rule order_trans) show "g \ f'" using order_trans[OF le F_le[OF f]] by(auto simp add: le_fun_def) qed(blast intro: IN_OUT link nonneg)+ have chain_Field: "Inf M \ F" if M: "M \ Chains leq" and nempty: "M \ {}" for M proof - from nempty obtain g0 where g0_in_M: "g0 \ M" by auto with M have g0: "g0 \ Field leq" by(rule Chains_FieldD) from M have "M \ Chains {(g, g'). g' \ g}" by(rule mono_Chains[THEN subsetD, rotated])(auto simp add: leq_def in_restrict_rel_iff) then have "Complete_Partial_Order.chain (\) M" by(rule Chains_into_chain) hence chain': "Complete_Partial_Order.chain (\) M" by(simp add: chain_dual) have "support_flow f' \ \<^bold>E" using F_outside[OF f'] by(auto intro: ccontr simp add: support_flow.simps) then have countable': "countable (support_flow f')" by(rule countable_subset)(simp add: E_bipartite_web countable "\<^bold>V_def") have finite_OUT: "d_OUT f' x \ \" for x using weight_finite[of x] by(rule neq_top_trans)(rule F_weight_OUT[OF f']) have finite_IN: "d_IN f' x \ \" for x using weight_finite[of x] by(rule neq_top_trans)(rule F_weight_IN[OF f']) have OUT_M: "d_OUT (Inf M) x = (INF g\M. d_OUT g x)" for x using chain' nempty countable' _ finite_OUT by(rule d_OUT_Inf)(auto dest!: Chains_FieldD[OF M] simp add: leq_def F_nonneg F_le) have IN_M: "d_IN (Inf M) x = (INF g\M. d_IN g x)" for x using chain' nempty countable' _ finite_IN by(rule d_IN_Inf)(auto dest!: Chains_FieldD[OF M] simp add: leq_def F_nonneg F_le) show "Inf M \ F" using g0 unfolding F[symmetric] proof(rule F_leI) show "(Inf M) e \ g0 e" for e using g0_in_M by(auto intro: INF_lower) show "0 \ (Inf M) e" for e by(auto intro!: INF_greatest dest: F_nonneg dest!: Chains_FieldD[OF M]) show "d_IN (Inf M) x \ d_OUT (Inf M) x" if "x \ B \" for x using that by(auto simp add: IN_M OUT_M intro!: INF_mono dest: Chains_FieldD[OF M] intro: F_IN_OUT[OF _ that]) show "linkage \ (Inf M)" using nempty by(simp add: linkage.simps OUT_M F_link[THEN linkageD] Chains_FieldD[OF M] cong: INF_cong) qed qed let ?P = "\g z. z \ A \ \ z \ B \ \ d_OUT g z > d_IN g z" define link where "link g = (if \z. ?P g z then let z = SOME z. ?P g z; factor = d_IN g z / d_OUT g z in (\(x, y). (if x = z then factor else 1) * g (x, y)) else g)" for g have increasing: "link g \ g \ link g \ Field leq" if g: "g \ Field leq" for g proof(cases "\z. ?P g z") case False thus ?thesis using that by(auto simp add: link_def leq_def) next case True define z where "z = Eps (?P g)" from True have "?P g z" unfolding z_def by(rule someI_ex) hence A: "z \ A \" and B: "z \ B \" and less: "d_IN g z < d_OUT g z" by simp_all let ?factor = "d_IN g z / d_OUT g z" have link [simp]: "link g (x, y) = (if x = z then ?factor else 1) * g (x, y)" for x y using True by(auto simp add: link_def z_def Let_def) have "?factor \ 1" (is "?factor \ _") using less by(cases "d_OUT g z" "d_IN g z" rule: ennreal2_cases) (simp_all add: ennreal_less_iff divide_ennreal) hence le': "?factor * g (x, y) \ 1 * g (x, y)" for y x by(rule mult_right_mono)(simp add: F_nonneg[OF g]) hence le: "link g e \ g e" for e by(cases e)simp have "link g \ Field leq" using g le proof(rule F_leI) show nonneg: "0 \ link g e" for e using F_nonneg[OF g, of e] by(cases e) simp show "linkage \ (link g)" using that A F_link[OF g] by(clarsimp simp add: linkage.simps d_OUT_def) fix x assume x: "x \ B \" have "d_IN (link g) x \ d_IN g x" unfolding d_IN_def using le' by(auto intro: nn_integral_mono) also have "\ \ d_OUT (link g) x" proof(cases "x = z") case True have "d_IN g x = d_OUT (link g) x" unfolding d_OUT_def using True F_weight_IN[OF g, of x] F_IN_OUT[OF g x] F_finite_OUT F_finite_OUT[OF g, of x] by(cases "d_OUT g z = 0") (auto simp add: nn_integral_divide nn_integral_cmult d_OUT_def[symmetric] ennreal_divide_times less_top) thus ?thesis by simp next case False have "d_IN g x \ d_OUT g x" using x by(rule F_IN_OUT[OF g]) also have "\ \ d_OUT (link g) x" unfolding d_OUT_def using False by(auto intro!: nn_integral_mono) finally show ?thesis . qed finally show "d_IN (link g) x \ d_OUT (link g) x" . qed with le g show ?thesis unfolding F by(simp add: leq_def le_fun_def del: link) qed have "bourbaki_witt_fixpoint Inf leq link" using chain_Field increasing unfolding leq_def by(intro bourbaki_witt_fixpoint_restrict_rel)(auto intro: Inf_greatest Inf_lower) then interpret bourbaki_witt_fixpoint Inf leq link . define g where "g = fixp_above f'" have g: "g \ Field leq" using f' unfolding g_def by(rule fixp_above_Field) hence "linkage \ g" by(rule F_link) moreover have "web_flow \ g" proof(intro web_flow.intros current.intros) show "d_OUT g x \ weight \ x" for x using g by(rule F_weight_OUT) show "d_IN g x \ weight \ x" for x using g by(rule F_weight_IN) show "d_IN g x = 0" if "x \ A \" for x using g that by(rule F_IN) show B: "d_OUT g x = 0" if "x \ B \" for x using g that by(rule F_OUT) show "g e = 0" if "e \ \<^bold>E" for e using g that by(rule F_outside) show KIR: "KIR g x" if A: "x \ A \" and B: "x \ B \" for x proof(rule ccontr) define z where "z = Eps (?P g)" assume "\ KIR g x" with F_IN_OUT[OF g B] have "d_OUT g x > d_IN g x" by simp with A B have Ex: "\x. ?P g x" by blast then have "?P g z" unfolding z_def by(rule someI_ex) hence A: "z \ A \" and B: "z \ B \" and less: "d_IN g z < d_OUT g z" by simp_all let ?factor = "d_IN g z / d_OUT g z" have "\y. edge \ z y \ g (z, y) > 0" proof(rule ccontr) assume "\ ?thesis" hence "d_OUT g z = 0" using F_outside[OF g] by(force simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space not_less) with less show False by simp qed then obtain y where y: "edge \ z y" and gr0: "g (z, y) > 0" by blast have "?factor < 1" (is "?factor < _") using less by(cases "d_OUT g z" "d_IN g z" rule: ennreal2_cases) (auto simp add: ennreal_less_iff divide_ennreal) hence le': "?factor * g (z, y) < 1 * g (z, y)" using gr0 F_finite[OF g] by(intro ennreal_mult_strict_right_mono) (auto simp: less_top) hence "link g (z, y) \ g (z, y)" using Ex by(auto simp add: link_def z_def Let_def) hence "link g \ g" by(auto simp add: fun_eq_iff) moreover have "link g = g" using f' unfolding g_def by(rule fixp_above_unfold[symmetric]) ultimately show False by contradiction qed show "d_OUT g x \ d_IN g x" if "x \ A \" for x using KIR[of x] that B[of x] by(cases "x \ B \") auto qed ultimately show ?thesis by blast qed end subsection \Extending a wave by a linkage\ lemma linkage_quotient_webD: fixes \ :: "('v, 'more) web_scheme" (structure) and h g defines "k \ plus_current h g" assumes f: "current \ f" and w: "wave \ f" and wg: "web_flow (quotient_web \ f) g" (is "web_flow ?\ _") and link: "linkage (quotient_web \ f) g" and trim: "trimming \ f h" shows "web_flow \ k" and "orthogonal_current \ k (\ (TER f))" proof - from wg have g: "current ?\ g" by(rule web_flowD_current) from trim obtain h: "current \ h" and w': "wave \ h" and h_le_f: "\e. h e \ f e" and KIR: "\x. \x \ RF\<^sup>\ (TER f); x \ A \\ \ KIR h x" and TER: "TER h \ \ (TER f) - A \" by(cases)(auto simp add: le_fun_def) have eq: "quotient_web \ f = quotient_web \ h" using w trim by(rule quotient_web_trimming) let ?T = "\ (TER f)" have RFc: "RF\<^sup>\ (TER h) = RF\<^sup>\ (TER f)" by(subst (1 2) roofed_circ_essential[symmetric])(simp only: trimming_\[OF w trim]) have OUT_k: "d_OUT k x = (if x \ RF\<^sup>\ (TER f) then d_OUT h x else d_OUT g x)" for x using OUT_plus_current[OF h w', of g] web_flowD_current[OF wg] unfolding eq k_def RFc by simp have RF: "RF (TER h) = RF (TER f)" by(subst (1 2) RF_essential[symmetric])(simp only: trimming_\[OF w trim]) have IN_k: "d_IN k x = (if x \ RF (TER f) then d_IN h x else d_IN g x)" for x using IN_plus_current[OF h w', of g] web_flowD_current[OF wg] unfolding eq k_def RF by simp have k: "current \ k" unfolding k_def using h w' g unfolding eq by(rule current_plus_current) then show wk: "web_flow \ k" proof(rule web_flow) fix x assume "x \ \<^bold>V" and A: "x \ A \" and B: "x \ B \" show "KIR k x" proof(cases "x \ \ (TER f)") case False thus ?thesis using A KIR[of x] web_flowD_KIR[OF wg, of x] B by(auto simp add: OUT_k IN_k roofed_circ_def) next case True with A TER have [simp]: "d_OUT h x = 0" and "d_IN h x \ weight \ x" by(auto simp add: SINK.simps elim: SAT.cases) with currentD_weight_IN[OF h, of x] have [simp]: "d_IN h x = weight \ x" by auto from linkageD[OF link, of x] True currentD_IN[OF g, of x] B have "d_OUT g x = weight \ x" "d_IN g x = 0" by(auto simp add: roofed_circ_def) thus ?thesis using True by(auto simp add: IN_k OUT_k roofed_circ_def intro: roofed_greaterI) qed qed have h_le_k: "h e \ k e" for e unfolding k_def plus_current_def by(rule add_increasing2) simp_all hence "SAT \ h \ SAT \ k" by(rule SAT_mono) hence SAT: "?T \ SAT \ k" using TER by(auto simp add: elim!: SAT.cases intro: SAT.intros) show "orthogonal_current \ k ?T" proof(rule orthogonal_current) show "weight \ x \ d_IN k x" if "x \ ?T" "x \ A \" for x using subsetD[OF SAT, of x] that by(auto simp add: currentD_SAT[OF k]) next fix x assume x: "x \ ?T" and A: "x \ A \" and B: "x \ B \" with d_OUT_mono[of h x f, OF h_le_f] have "d_OUT h x = 0" by(auto simp add: SINK.simps) moreover from linkageD[OF link, of x] x A have "d_OUT g x = weight ?\ x" by simp ultimately show "d_OUT k x = weight \ x" using x A currentD_IN[OF f A] B by(auto simp add: d_OUT_add roofed_circ_def k_def plus_current_def ) next fix u v assume v: "v \ RF ?T" and u: "u \ RF\<^sup>\ ?T" have "h (u, v) \ f (u, v)" by(rule h_le_f) also have "\ \ d_OUT f u" unfolding d_OUT_def by(rule nn_integral_ge_point) simp also have "\ = 0" using u using RF_essential[of \ "TER f"] by(auto simp add: roofed_circ_def SINK.simps intro: waveD_OUT[OF w]) finally have "h (u, v) = 0" by simp moreover have "g (u, v) = 0" using g v RF_essential[of \ "TER f"] by(auto intro: currentD_outside simp add: roofed_circ_def) ultimately show "k (u, v) = 0" by(simp add: k_def) qed qed context countable_web begin lemma ex_orthogonal_current': \ \Lemma 4.15\ assumes loose_linkable: "\f. \ current \ f; wave \ f; loose (quotient_web \ f) \ \ linkable (quotient_web \ f)" shows "\f S. web_flow \ f \ separating \ S \ orthogonal_current \ f S" proof - from ex_maximal_wave[OF countable] obtain f where f: "current \ f" and w: "wave \ f" and maximal: "\w. \ current \ w; wave \ w; f \ w \ \ f = w" by blast from ex_trimming[OF f w countable weight_finite] obtain h where h: "trimming \ f h" .. let ?\ = "quotient_web \ f" interpret \: countable_web "?\" by(rule countable_web_quotient_web) have "loose ?\" using f w maximal by(rule loose_quotient_web[OF weight_finite]) with f w have "linkable ?\" by(rule loose_linkable) then obtain g where wg: "web_flow ?\ g" and link: "linkage ?\ g" by blast let ?k = "plus_current h g" have "web_flow \ ?k" "orthogonal_current \ ?k (\ (TER f))" by(rule linkage_quotient_webD[OF f w wg link h])+ moreover have "separating \ (\ (TER f))" using waveD_separating[OF w] by(rule separating_essential) ultimately show ?thesis by blast qed end subsection \From a network to a web\ definition web_of_network :: "('v, 'more) network_scheme \ ('v edge, 'more) web_scheme" where "web_of_network \ = \edge = \(x, y) (y', z). y' = y \ edge \ x y \ edge \ y z, weight = capacity \, A = {(source \, x)|x. edge \ (source \) x}, B = {(x, sink \)|x. edge \ x (sink \)}, \ = network.more \\" lemma web_of_network_sel [simp]: fixes \ (structure) shows "edge (web_of_network \) e e' \ e \ \<^bold>E \ e' \ \<^bold>E \ snd e = fst e'" "weight (web_of_network \) e = capacity \ e" "A (web_of_network \) = {(source \, x)|x. edge \ (source \) x}" "B (web_of_network \) = {(x, sink \)|x. edge \ x (sink \)}" by(auto simp add: web_of_network_def split: prod.split) lemma vertex_web_of_network [simp]: "vertex (web_of_network \) (x, y) \ edge \ x y \ (\z. edge \ y z \ edge \ z x)" by(auto simp add: vertex_def Domainp.simps Rangep.simps) definition flow_of_current :: "('v, 'more) network_scheme \ 'v edge current \ 'v flow" where "flow_of_current \ f e = max (d_OUT f e) (d_IN f e)" lemma flow_flow_of_current: fixes \ (structure) and \ defines [simp]: "\ \ web_of_network \" assumes fw: "web_flow \ f" shows "flow \ (flow_of_current \ f)" (is "flow _ ?f") proof from fw have f: "current \ f" and KIR: "\x. \ x \ A \; x \ B \ \ \ KIR f x" by(auto 4 3 dest: web_flowD_current web_flowD_KIR) show "?f e \ capacity \ e" for e using currentD_weight_OUT[OF f, of e] currentD_weight_IN[OF f, of e] by(simp add: flow_of_current_def) fix x assume x: "x \ source \" "x \ sink \" have "d_OUT ?f x = (\\<^sup>+ y. d_IN f (x, y))" unfolding d_OUT_def by(simp add: flow_of_current_def max_absorb2 currentD_OUT_IN[OF f] x) also have "\ = (\\<^sup>+ y. \\<^sup>+ e\range (\z. (z, x)). f (e, x, y))" by(auto simp add: nn_integral_count_space_indicator d_IN_def intro!: nn_integral_cong currentD_outside[OF f] split: split_indicator) also have "\ = (\\<^sup>+ z\UNIV. \\<^sup>+ y. f ((z, x), x, y))" by(subst nn_integral_snd_count_space[of "case_prod _", simplified]) (simp add: nn_integral_count_space_reindex nn_integral_fst_count_space[of "case_prod _", simplified]) also have "\ = (\\<^sup>+ z. \\<^sup>+ e\range (Pair x). f ((z, x), e))" by(simp add: nn_integral_count_space_reindex) also have "\ = (\\<^sup>+ z. d_OUT f (z, x))" by(auto intro!: nn_integral_cong currentD_outside[OF f] simp add: d_OUT_def nn_integral_count_space_indicator split: split_indicator) also have "\ = (\\<^sup>+ z\{z. edge \ z x}. d_OUT f (z, x))" by(auto intro!: nn_integral_cong currentD_outside_OUT[OF f] simp add: nn_integral_count_space_indicator split: split_indicator) also have "\ = (\\<^sup>+ z\{z. edge \ z x}. max (d_OUT f (z, x)) (d_IN f (z, x)))" proof(rule nn_integral_cong) show "d_OUT f (z, x) = max (d_OUT f (z, x)) (d_IN f (z, x))" if "z \ space (count_space {z. edge \ z x})" for z using currentD_IN[OF f] that by(cases "z = source \")(simp_all add: max_absorb1 currentD_IN[OF f] KIR x) qed also have "\ = (\\<^sup>+ z. max (d_OUT f (z, x)) (d_IN f (z, x)))" by(auto intro!: nn_integral_cong currentD_outside_OUT[OF f] currentD_outside_IN[OF f] simp add: nn_integral_count_space_indicator max_def split: split_indicator) also have "\ = d_IN ?f x" by(simp add: flow_of_current_def d_IN_def) finally show "KIR ?f x" . qed text \ The reduction of Conjecture 1.2 to Conjecture 3.6 is flawed in @{cite "AharoniBergerGeorgakopoulusPerlsteinSpruessel2011JCT"}. Not every essential A-B separating set of vertices in @{term "web_of_network \"} is an s-t-cut in @{term \}, as the following counterexample shows. The network @{term \} has five nodes @{term "s"}, @{term "t"}, @{term "x"}, @{term "y"} and @{term "z"} and edges @{term "(s, x)"}, @{term "(x, y)"}, @{term "(y, z)"}, @{term "(y, t)"} and @{term "(z, t)"}. For @{term "web_of_network \"}, the set @{term "S = {(x, y), (y, z)}"} is essential and A-B separating. (@{term "(x, y)"} is essential due to the path @{term "[(y, z)]"} and @{term "(y, z)"} is essential due to the path @{term "[(z, t)]"}). However, @{term S} is not a cut in @{term \} because the node @{term y} has an outgoing edge that is in @{term S} and one that is not in @{term S}. However, this can be remedied if all edges carry positive capacity. Then, orthogonality of the current rules out the above possibility. \ lemma cut_RF_separating: fixes \ (structure) assumes sep: "separating_network \ V" and sink: "sink \ \ V" shows "cut \ (RF\<^sup>N V)" proof show "source \ \ RF\<^sup>N V" by(rule roofedI)(auto dest: separatingD[OF sep]) show "sink \ \ RF\<^sup>N V" using sink by(auto dest: roofedD[OF _ rtrancl_path.base]) qed context fixes \ :: "('v, 'more) network_scheme" and \ (structure) defines \_def: "\ \ web_of_network \" begin lemma separating_network_cut_of_sep: assumes sep: "separating \ S" and source_sink: "source \ \ sink \" shows "separating_network \ (fst ` \ S)" proof define s t where "s = source \" and "t = sink \" fix p assume p: "path \ s p t" with p source_sink have "p \ []" by cases(auto simp add: s_def t_def) with p have p': "path \ (s, hd p) (zip p (tl p)) (last (s # butlast p), t)" proof(induction) case (step x y p z) then show ?case by(cases p)(auto elim: rtrancl_path.cases intro: rtrancl_path.intros simp add: \_def) qed simp from sep have "separating \ (\ S)" by(rule separating_essential) from this p' have "(\z\set (zip p (tl p)). z \ \ S) \ (s, hd p) \ \ S" apply(rule separatingD) using rtrancl_path_nth[OF p, of 0] rtrancl_path_nth[OF p, of "length p - 1"] \p \ []\ rtrancl_path_last[OF p] apply(auto simp add: \_def s_def t_def hd_conv_nth last_conv_nth nth_butlast nth_Cons' cong: if_cong split: if_split_asm) apply(metis One_nat_def Suc_leI cancel_comm_monoid_add_class.diff_cancel le_antisym length_butlast length_greater_0_conv list.size(3))+ done then show "(\z\set p. z \ fst ` \ S) \ s \ fst ` \ S" by(auto dest!: set_zip_leftD intro: rev_image_eqI) qed definition cut_of_sep :: "('v \ 'v) set \ 'v set" where "cut_of_sep S = RF\<^sup>N\<^bsub>\\<^esub> (fst ` \ S)" lemma separating_cut: assumes sep: "separating \ S" and neq: "source \ \ sink \" and sink_out: "\x. \ edge \ (sink \) x" shows "cut \ (cut_of_sep S)" unfolding cut_of_sep_def proof(rule cut_RF_separating) show "separating_network \ (fst ` \ S)" using sep neq by(rule separating_network_cut_of_sep) show "sink \ \ fst ` \ S" proof assume "sink \ \ fst ` \ S" then obtain x where "(sink \, x) \ \ S" by auto hence "(sink \, x) \ \<^bold>V" by(auto simp add: \_def dest!: essential_vertex) then show False by(simp add: \_def sink_out) qed qed context fixes f :: "'v edge current" and S assumes wf: "web_flow \ f" and ortho: "orthogonal_current \ f S" and sep: "separating \ S" and capacity_pos: "\e. e \ \<^bold>E\<^bsub>\\<^esub> \ capacity \ e > 0" begin private lemma f: "current \ f" using wf by(rule web_flowD_current) lemma orthogonal_leave_RF: assumes e: "edge \ x y" and x: "x \ (cut_of_sep S)" and y: "y \ (cut_of_sep S)" shows "(x, y) \ S" proof - from y obtain p where p: "path \ y p (sink \)" and y': "y \ fst ` \ S" and bypass: "\z. z \ set p \ z \ fst ` \ S" by(auto simp add: roofed_def cut_of_sep_def \_def[symmetric]) from e p have p': "path \ x (y # p) (sink \)" .. from roofedD[OF x[unfolded cut_of_sep_def] this] y' bypass have "x \ fst ` \ S" by(auto simp add: \_def[symmetric]) then obtain z where xz: "(x, z) \ \ S" by auto then obtain q b where q: "path \ (x, z) q b" and b: "b \ B \" and distinct: "distinct ((x, z) # q)" and bypass': "\z. z \ set q \ z \ RF S" by(rule \_E_RF) blast define p' where "p' = y # p" hence "p' \ []" by simp with p' have "path \ (x, hd p') (zip p' (tl p')) (last (x # butlast p'), sink \)" unfolding p'_def[symmetric] proof(induction) case (step x y p z) then show ?case by(cases p)(auto elim: rtrancl_path.cases intro: rtrancl_path.intros simp add: \_def) qed simp then have p'': "path \ (x, y) (zip (y # p) p) (last (x # butlast (y # p)), sink \)" (is "path _ ?y ?p ?t") by(simp add: p'_def) have "(?y # ?p) ! length p = ?t" using rtrancl_path_last[OF p'] p rtrancl_path_last[OF p] apply(auto split: if_split_asm simp add: nth_Cons butlast_conv_take take_Suc_conv_app_nth split: nat.split elim: rtrancl_path.cases) apply(simp add: last_conv_nth) done moreover have "length p < length (?y # ?p)" by simp ultimately have t: "?t \ B \" using rtrancl_path_nth[OF p'', of "length p - 1"] e by(cases p)(simp_all add: \_def split: if_split_asm) show S: "(x, y) \ S" proof(cases "x = source \") case True from separatingD[OF separating_essential, OF sep p'' _ t] e True consider (z) z z' where "(z, z') \ set ?p" "(z, z') \ \ S" | "(x, y) \ S" by(auto simp add: \_def) thus ?thesis proof cases case (z z) hence "z \ set p" "z \ fst ` \ S" using y' by(auto dest!: set_zip_leftD intro: rev_image_eqI) hence False by(auto dest: bypass) thus ?thesis .. qed next case False have "\e. edge \ e (x, z) \ f (e, (x, z)) > 0" proof(rule ccontr) assume "\ ?thesis" then have "d_IN f (x, z) = 0" unfolding d_IN_def using currentD_outside[OF f, of _ "(x, z)"] by(force simp add: nn_integral_0_iff_AE AE_count_space not_less) moreover from xz have "(x, z) \ S" by auto hence "(x, z) \ SAT \ f" by(rule orthogonal_currentD_SAT[OF ortho]) with False have "d_IN f (x, z) \ capacity \ (x, z)" by(auto simp add: SAT.simps \_def) ultimately have "\ capacity \ (x, z) > 0" by auto hence "\ edge \ x z" using capacity_pos[of "(x, z)"] by auto moreover with q have "b = (x, z)" by cases(auto simp add: \_def) with b have "edge \ x z" by(simp add: \_def) ultimately show False by contradiction qed then obtain u where ux: "edge \ u x" and xz': "edge \ x z" and uxz: "edge \ (u, x) (x, z)" and gt0: "f ((u, x), (x, z)) > 0" by(auto simp add: \_def) have "(u, x) \ RF\<^sup>\ S" using orthogonal_currentD_in[OF ortho, of "(x, z)" "(u, x)"] gt0 xz by(fastforce intro: roofed_greaterI) hence ux_RF: "(u, x) \ RF (\ S)" and ux_\: "(u, x) \ \ S" unfolding RF_essential by(simp_all add: roofed_circ_def) from ux e have "edge \ (u, x) (x, y)" by(simp add: \_def) hence "path \ (u, x) ((x, y) # ?p) ?t" using p'' .. from roofedD[OF ux_RF this t] ux_\ consider "(x, y) \ S" | (z) z z' where "(z, z') \ set ?p" "(z, z') \ \ S" by auto thus ?thesis proof cases case (z z) with bypass[of z] y' have False by(fastforce dest!: set_zip_leftD intro: rev_image_eqI) thus ?thesis .. qed qed qed lemma orthogonal_flow_of_current: assumes source_sink: "source \ \ sink \" and sink_out: "\x. \ edge \ (sink \) x" and no_direct_edge: "\ edge \ (source \) (sink \)" \ \Otherwise, @{const A} and @{const B} of the web would not be disjoint.\ shows "orthogonal \ (flow_of_current \ f) (cut_of_sep S)" (is "orthogonal _ ?f ?S") proof fix x y assume e: "edge \ x y" and "x \ ?S" and "y \ ?S" then have S: "(x, y) \ S"by(rule orthogonal_leave_RF) show "?f (x, y) = capacity \ (x, y)" proof(cases "x = source \") case False with orthogonal_currentD_SAT[OF ortho S] have "weight \ (x, y) \ d_IN f (x, y)" by cases(simp_all add: \_def) with False currentD_OUT_IN[OF f, of "(x, y)"] currentD_weight_IN[OF f, of "(x, y)"] show ?thesis by(simp add: flow_of_current_def \_def max_def) next case True with orthogonal_currentD_A[OF ortho S] e currentD_weight_IN[OF f, of "(x, y)"] no_direct_edge show ?thesis by(auto simp add: flow_of_current_def \_def max_def) qed next from sep source_sink sink_out have cut: "cut \ ?S" by(rule separating_cut) fix x y assume xy: "edge \ x y" and x: "x \ ?S" and y: "y \ ?S" from x obtain p where p: "path \ x p (sink \)" and x': "x \ fst ` \ S" and bypass: "\z. z \ set p \ z \ fst ` \ S" by(auto simp add: roofed_def cut_of_sep_def) have source: "x \ source \" proof assume "x = source \" have "separating_network \ (fst ` \ S)" using sep source_sink by(rule separating_network_cut_of_sep) from separatingD[OF this p] \x = source \\ x show False by(auto dest: bypass intro: roofed_greaterI simp add: cut_of_sep_def) qed hence A: "(x, y) \ A \" by(simp add: \_def) have "f ((u, v), x, y) = 0" for u v proof(cases "edge \ (u, v) (x, y)") case False with f show ?thesis by(rule currentD_outside) next case True hence [simp]: "v = x" and ux: "edge \ u x" by(auto simp add: \_def) have "(x, y) \ RF S" proof fix q b assume q: "path \ (x, y) q b" and b: "b \ B \" define xy where "xy = (x, y)" from q have "path \ (snd xy) (map snd q) (snd b)" unfolding xy_def[symmetric] by(induction)(auto intro: rtrancl_path.intros simp add: \_def) with b have "path \ y (map snd q) (sink \)" by(auto simp add: xy_def \_def) from roofedD[OF y[unfolded cut_of_sep_def] this] have "\z\set (y # map snd q). z \ ?S" by(auto intro: roofed_greaterI simp add: cut_of_sep_def) from split_list_last_prop[OF this] obtain xs z ys where decomp: "y # map snd q = xs @ z # ys" and z: "z \ ?S" and last: "\z. z \ set ys \ z \ ?S" by auto from decomp obtain x' xs' z'' ys' where decomp': "(x', y) # q = xs' @ (z'', z) # ys'" and "xs = map snd xs'" and ys: "ys = map snd ys'" and x': "xs' = [] \ x' = x" by(fastforce simp add: Cons_eq_append_conv map_eq_append_conv) from cut z have z_sink: "z \ sink \" by cases(auto) then have "ys' \ []" using rtrancl_path_last[OF q] decomp' b x' q by(auto simp add: Cons_eq_append_conv \_def elim: rtrancl_path.cases) then obtain w z''' ys'' where ys': "ys' = (w, z''') # ys''" by(auto simp add: neq_Nil_conv) from q[THEN rtrancl_path_nth, of "length xs'"] decomp' ys' x' have "edge \ (z'', z) (w, z''')" by(auto simp add: Cons_eq_append_conv nth_append) hence w: "w = z" and zz''': "edge \ z z'''" by(auto simp add: \_def) from ys' ys last[of z'''] have "z''' \ ?S" by simp with zz''' z have "(z, z''') \ S" by(rule orthogonal_leave_RF) moreover have "(z, z''') \ set q" using decomp' ys' w by(auto simp add: Cons_eq_append_conv) ultimately show "(\z\set q. z \ S) \ (x, y) \ S" by blast qed moreover have "(u, x) \ RF\<^sup>\ S" proof assume "(u, x) \ RF\<^sup>\ S" hence ux_RF: "(u, x) \ RF (\ S)" and ux_\: "(u, x) \ \ S" unfolding RF_essential by(simp_all add: roofed_circ_def) have "x \ sink \" using p xy by cases(auto simp add: sink_out) with p have nNil: "p \ []" by(auto elim: rtrancl_path.cases) with p have "edge \ x (hd p)" by cases auto with ux have "edge \ (u, x) (x, hd p)" by(simp add: \_def) moreover from p nNil have "path \ (x, hd p) (zip p (tl p)) (last (x # butlast p), sink \)" (is "path _ ?x ?p ?t") proof(induction) case (step x y p z) then show ?case by(cases p)(auto elim: rtrancl_path.cases intro: rtrancl_path.intros simp add: \_def) qed simp ultimately have p': "path \ (u, x) (?x # ?p) ?t" .. have "(?x # ?p) ! (length p - 1) = ?t" using rtrancl_path_last[OF p] p nNil apply(auto split: if_split_asm simp add: nth_Cons butlast_conv_take take_Suc_conv_app_nth not_le split: nat.split elim: rtrancl_path.cases) apply(simp add: last_conv_nth nth_tl) done moreover have "length p - 1 < length (?x # ?p)" by simp ultimately have t: "?t \ B \" using rtrancl_path_nth[OF p', of "length p - 1"] by(cases p)(simp_all add: \_def split: if_split_asm) from roofedD[OF ux_RF p' t] ux_\ consider (X) "(x, hd p) \ \ S" | (z) z z' where "(z, z') \ set (zip p (tl p))" "(z, z') \ \ S" by auto thus False proof cases case X with x' show False by(auto simp add: cut_of_sep_def intro: rev_image_eqI) next case (z z) with bypass[of z] show False by(auto 4 3 simp add: cut_of_sep_def intro: rev_image_eqI dest!: set_zip_leftD) qed qed ultimately show ?thesis unfolding \v = x\ by(rule orthogonal_currentD_in[OF ortho]) qed then have "d_IN f (x, y) = 0" unfolding d_IN_def by(simp add: nn_integral_0_iff emeasure_count_space_eq_0) with currentD_OUT_IN[OF f A] show "flow_of_current \ f (x, y) = 0" by(simp add: flow_of_current_def max_def) qed end end subsection \Avoiding antiparallel edges and self-loops\ context antiparallel_edges begin abbreviation cut' :: "'a vertex set \ 'a set" where "cut' S \ Vertex -` S" lemma cut_cut': "cut \'' S \ cut \ (cut' S)" by(auto simp add: cut.simps) lemma IN_Edge: "\<^bold>I\<^bold>N\<^bsub>\''\<^esub> (Edge x y) = (if edge \ x y then {Vertex x} else {})" by(auto simp add: incoming_def) lemma OUT_Edge: "\<^bold>O\<^bold>U\<^bold>T\<^bsub>\''\<^esub> (Edge x y) = (if edge \ x y then {Vertex y} else {})" by(auto simp add: outgoing_def) interpretation \'': countable_network \'' by(rule \''_countable_network) lemma d_IN_Edge: assumes f: "flow \'' f" shows "d_IN f (Edge x y) = f (Vertex x, Edge x y)" by(subst d_IN_alt_def[OF \''.flowD_outside[OF f], of _ \''])(simp_all add: IN_Edge nn_integral_count_space_indicator max_def \''.flowD_outside[OF f]) lemma d_OUT_Edge: assumes f: "flow \'' f" shows "d_OUT f (Edge x y) = f (Edge x y, Vertex y)" by(subst d_OUT_alt_def[OF \''.flowD_outside[OF f], of _ \''])(simp_all add: OUT_Edge nn_integral_count_space_indicator max_def \''.flowD_outside[OF f]) lemma orthogonal_cut': assumes ortho: "orthogonal \'' f S" and f: "flow \'' f" shows "orthogonal \ (collect f) (cut' S)" proof show "collect f (x, y) = capacity \ (x, y)" if edge: "edge \ x y" and x: "x \ cut' S" and y: "y \ cut' S" for x y proof(cases "Edge x y \ S") case True from y have "Vertex y \ S" by auto from orthogonalD_out[OF ortho _ True this] edge show ?thesis by simp next case False from x have "Vertex x \ S" by auto from orthogonalD_out[OF ortho _ this False] edge have "capacity \ (x, y) = d_IN f (Edge x y)" by(simp add: d_IN_Edge[OF f]) also have "\ = d_OUT f (Edge x y)" by(simp add: flowD_KIR[OF f]) also have "\ = f (Edge x y, Vertex y)" using edge by(simp add: d_OUT_Edge[OF f]) finally show ?thesis by simp qed show "collect f (x, y) = 0" if edge: "edge \ x y" and x: "x \ cut' S" and y: "y \ cut' S" for x y proof(cases "Edge x y \ S") case True from x have "Vertex x \ S" by auto from orthogonalD_in[OF ortho _ this True] edge have "0 = d_IN f (Edge x y)" by(simp add: d_IN_Edge[OF f]) also have "\ = d_OUT f (Edge x y)" by(simp add: flowD_KIR[OF f]) also have "\ = f (Edge x y, Vertex y)" using edge by(simp add: d_OUT_Edge[OF f]) finally show ?thesis by simp next case False from y have "Vertex y \ S" by auto from orthogonalD_in[OF ortho _ False this] edge show ?thesis by simp qed qed end context countable_network begin lemma countable_web_web_of_network: assumes source_in: "\x. \ edge \ x (source \)" and sink_out: "\y. \ edge \ (sink \) y" and undead: "\x y. edge \ x y \ (\z. edge \ y z) \ (\z. edge \ z x)" and source_sink: "\ edge \ (source \) (sink \)" and no_loop: "\x. \ edge \ x x" - and edge_antiparallel: "\x y. edge \ x y \ \ edge \ y x" shows "countable_web (web_of_network \)" (is "countable_web ?\") proof show "\ edge ?\ y x" if "x \ A ?\" for x y using that by(clarsimp simp add: source_in) show "\ edge ?\ x y" if "x \ B ?\" for x y using that by(clarsimp simp add: sink_out) show "A ?\ \ \<^bold>V\<^bsub>?\\<^esub>" by(auto 4 3 dest: undead) show "A ?\ \ B ?\ = {}" using source_sink by auto show "\ edge ?\ x x" for x by(auto simp add: no_loop) show "weight ?\ x = 0" if "x \ \<^bold>V\<^bsub>?\\<^esub>" for x using that undead by(cases x)(auto intro!: capacity_outside) show "weight ?\ x \ \" for x using capacity_finite[of x] by(cases x) simp - show "\ edge ?\ y x" if "edge ?\ x y" for x y using that by(auto simp add: edge_antiparallel) have "\<^bold>E\<^bsub>?\\<^esub> \ \<^bold>E \ \<^bold>E" by auto thus "countable \<^bold>E\<^bsub>?\\<^esub>" by(rule countable_subset)(simp) qed lemma max_flow_min_cut': assumes ex_orthogonal_current: " \f S. web_flow (web_of_network \) f \ separating (web_of_network \) S \ orthogonal_current (web_of_network \) f S" and source_in: "\x. \ edge \ x (source \)" and sink_out: "\y. \ edge \ (sink \) y" and undead: "\x y. edge \ x y \ (\z. edge \ y z) \ (\z. edge \ z x)" and source_sink: "\ edge \ (source \) (sink \)" and no_loop: "\x. \ edge \ x x" - and edge_antiparallel: "\x y. edge \ x y \ \ edge \ y x" and capacity_pos: "\e. e \ \<^bold>E \ capacity \ e > 0" shows "\f S. flow \ f \ cut \ S \ orthogonal \ f S" proof - let ?\ = "web_of_network \" from ex_orthogonal_current obtain f S where f: "web_flow (web_of_network \) f" and S: "separating (web_of_network \) S" and ortho: "orthogonal_current (web_of_network \) f S" by blast+ let ?f = "flow_of_current \ f" and ?S = "cut_of_sep \ S" from f have "flow \ ?f" by(rule flow_flow_of_current) moreover have "cut \ ?S" using S source_neq_sink sink_out by(rule separating_cut) moreover have "orthogonal \ ?f ?S" using f ortho S capacity_pos source_neq_sink sink_out source_sink by(rule orthogonal_flow_of_current) ultimately show ?thesis by blast qed subsection \Eliminating zero edges and incoming edges to @{term source} and outgoing edges of @{term sink}\ definition \''' :: "'v network" where "\''' = \edge = \x y. edge \ x y \ capacity \ (x, y) > 0 \ y \ source \ \ x \ sink \, capacity = \(x, y). if x = sink \ \ y = source \ then 0 else capacity \ (x, y), source = source \, sink = sink \\" lemma \'''_sel [simp]: "edge \''' x y \ edge \ x y \ capacity \ (x, y) > 0 \ y \ source \ \ x \ sink \" "capacity \''' (x, y) = (if x = sink \ \ y = source \ then 0 else capacity \ (x, y))" "source \''' = source \" "sink \''' = sink \" for x y by(simp_all add: \'''_def) lemma \'''_countable_network: "countable_network \'''" proof(unfold_locales) have "\<^bold>E\<^bsub>\'''\<^esub> \ \<^bold>E" by auto then show "countable \<^bold>E\<^bsub>\'''\<^esub>" by(rule countable_subset) simp show "capacity \''' e = 0" if "e \ \<^bold>E\<^bsub>\'''\<^esub>" for e using capacity_outside[of e] that by(auto split: if_split_asm intro: ccontr) qed(auto simp add: split: if_split_asm) lemma flow_\''': assumes f: "flow \''' f" and cut: "cut \''' S" and ortho: "orthogonal \''' f S" shows "flow \ f" "cut \ S" "orthogonal \ f S" proof - interpret \''': countable_network \''' by(rule \'''_countable_network) show "flow \ f" proof show "f e \ capacity \ e" for e using flowD_capacity[OF f, of e] by(cases e)(simp split: if_split_asm) show "KIR f x" if "x \ source \" "x \ sink \" for x using flowD_KIR[OF f, of x] that by simp qed show "cut \ S" using cut by(simp add: cut.simps) show "orthogonal \ f S" proof show "f (x, y) = capacity \ (x, y)" if edge: "edge \ x y" and x: "x \ S" and y: "y \ S" for x y proof(cases "edge \''' x y") case True with orthogonalD_out[OF ortho this x y] show ?thesis by simp next case False from cut y x have xy: "y \ source \ \ x \ sink \" by(cases) auto with xy edge False have "capacity \ (x, y) = 0" by simp with \'''.flowD_outside[OF f, of "(x, y)"] False show ?thesis by simp qed show "f (x, y) = 0" if edge: "edge \ x y" and x: "x \ S" and y: "y \ S" for x y using orthogonalD_in[OF ortho _ x y] \'''.flowD_outside[OF f, of "(x, y)"] by(cases "edge \''' x y")simp_all qed qed end end \ No newline at end of file diff --git a/thys/MFMC_Countable/MFMC_Unbounded.thy b/thys/MFMC_Countable/MFMC_Unbounded.thy --- a/thys/MFMC_Countable/MFMC_Unbounded.thy +++ b/thys/MFMC_Countable/MFMC_Unbounded.thy @@ -1,3478 +1,3467 @@ (* 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) -context - fixes \ :: "('v, 'more) web_scheme" (structure) and f g - assumes f: "current \ f" - and w: "wave \ f" - and g: "current (quotient_web \ f) g" -begin - - -end - 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 + (\\<^sup>+ y. (if x \ RF\<^sup>\ (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 + (\\<^sup>+ x. (if y \ RF (TER f) then 0 else g (x, y) * indicator (- RF\<^sup>\ (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 \\<^bsub>\\<^esub> 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\<^sup>\ (TER f)" proof assume *: "x \ RF\<^sup>\ (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\<^sup>\ (TER g)" proof assume *: "x \ RF\<^sup>\ (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\<^sup>\ (TER f)" have "(z, x) \ \<^bold>E" proof assume "(z, x) \ \<^bold>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 "(\z\set p. z \ TER f \ TER g) \ x \ TER f \ TER g" proof(cases "z \ RF (TER g)") case False hence "z \ SINK g" by(simp add: SINK.simps waveD_OUT[OF w']) with SINK SAT have "z \ TER f \ TER g" by auto thus ?thesis using z by auto next case True from split_list[OF z] obtain ys zs where split: "x # p = ys @ z # zs" by blast with p have "path \ z zs y" by(auto elim: rtrancl_path_appendE simp add: Cons_eq_append_conv) from roofedD[OF True this y] split show ?thesis by(auto simp add: Cons_eq_append_conv) qed qed qed lemma RF_TER_Sup: fixes \ (structure) assumes f: "\f. f \ Y \ current \ f" and w: "\f. f \ Y \ wave \ f" and Y: "Complete_Partial_Order.chain (\) Y" "Y \ {}" "countable (support_flow (Sup Y))" shows "RF (TER (Sup Y)) = RF (\f\Y. TER f)" proof(rule set_eqI iffI)+ fix x assume x: "x \ RF (TER (Sup Y))" have "x \ RF (RF (\f\Y. TER f))" proof fix p y assume p: "path \ x p y" and y: "y \ B \" from roofedD[OF x p y] obtain z where z: "z \ set (x # p)" and TER: "z \ TER (Sup Y)" by auto from TER have SINK: "z \ SINK f" if "f \ Y" for f using that by(auto simp add: SINK_Sup[OF Y]) from Y(2) obtain f where y: "f \ Y" by blast show "(\z\set p. z \ RF (\f\Y. TER f)) \ x \ RF (\f\Y. TER f)" proof(cases "\f\Y. z \ RF (TER f)") case True then obtain f where fY: "f \ Y" and zf: "z \ RF (TER f)" by blast from zf have "z \ RF (\f\Y. TER f)" by(rule in_roofed_mono)(auto intro: fY) with z show ?thesis by auto next case False hence *: "d_IN f z = 0" if "f \ Y" for f using that by(auto intro: wave_not_RF_IN_zero[OF f w]) hence "d_IN (Sup Y) z = 0" using Y(2) by(simp add: d_IN_Sup[OF Y]) with TER have "z \ SAT \ f" using *[OF y] by(simp add: SAT.simps) with SINK[OF y] have "z \ TER f" by simp with z y show ?thesis by(auto intro: roofed_greaterI) qed qed then show "x \ RF (\f\Y. TER f)" unfolding roofed_idem . next fix x assume x: "x \ RF (\f\Y. TER f)" have "x \ RF (RF (TER (\Y)))" proof(rule roofedI) fix p y assume p: "path \ x p y" and y: "y \ B \" from roofedD[OF x p y] obtain z f where *: "z \ set (x # p)" and **: "f \ Y" and TER: "z \ TER f" by auto have "z \ RF (TER (Sup Y))" proof(rule ccontr) assume z: "z \ RF (TER (Sup Y))" have "wave \ (Sup Y)" using Y(1-2) w Y(3) by(rule wave_lub) hence "d_OUT (Sup Y) z = 0" using z by(rule waveD_OUT) hence "z \ SINK (Sup Y)" by(simp add: SINK.simps) moreover have "z \ SAT \ (Sup Y)" using TER SAT_Sup_upper[OF **, of \] by blast ultimately have "z \ TER (Sup Y)" by simp hence "z \ RF (TER (Sup Y))" by(rule roofed_greaterI) with z show False by contradiction qed thus "(\z\set p. z \ RF (TER (Sup Y))) \ x \ RF (TER (Sup Y))" using * by auto qed then show "x \ RF (TER (\Y))" unfolding roofed_idem . qed subsection \Hindered webs with reduced weights\ context countable_bipartite_web begin context fixes u :: "'v \ ennreal" and \ defines "\ \ (\\<^sup>+ 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\<^bsub>?\\<^esub>" 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 \ \\<^bsub>?\\<^esub> (?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 \ \<^bold>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 \ \<^bold>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 \ \<^bold>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 \ \<^bold>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 \ \<^bold>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 \ \<^bold>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\ \ (\\<^sup>+ y\range Some. f (x, the y))" using that disjoint a(1) unfolding d_OUT_def by(auto 4 4 intro!: nn_integral_mono simp add: nn_integral_count_space_indicator notin_range_Some currentD_outside[OF f] split: split_indicator dest: edge_antiparallel bipartite_E) also have "\ = d_OUT f x" by(simp add: d_OUT_def nn_integral_count_space_reindex) also have "\ \ weight \ x" using f' by(rule currentD_weight_OUT) finally show ?thesis . qed have flow_attainability: "flow_attainability \" proof have "\<^bold>E\<^bsub>\\<^esub> = (\(x, y). (\x\, \y\)) ` \<^bold>E \ (\(x, y). (\y\, \x\)) ` \<^bold>E \ (\x. (None, \x\)) ` (\<^bold>V \ - A \)" by(auto simp add: edge'_def split: option.split_asm) thus "countable \<^bold>E\<^bsub>\\<^esub>" 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 \ \<^bold>V" | (outside) x where "v = \x\" "x \ \<^bold>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 \ (\\<^sup>+ 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 "\ = (\\<^sup>+ x\range Some. f (the x, y)) + u y" by(subst nn_integral_add)(simp_all add: nn_integral_count_space_indicator) also have "\ = d_IN f y + u y" by(simp add: d_IN_def nn_integral_count_space_reindex) also have "d_IN f y \ weight \ y" using f' by(rule currentD_weight_IN) finally show ?thesis by(auto simp add: add_right_mono top_unique split: if_split_asm) next case outside hence "d_OUT (capacity \) v = 0" by(auto simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space cap_def vertex_def split: option.split) thus ?thesis by simp qed next show "capacity \ e \ \" for e using weight_finite by(auto simp add: cap_def max_def vertex_def currentD_finite[OF f'] split: option.split prod.split simp del: weight_finite) show "capacity \ e = 0" if "e \ \<^bold>E\<^bsub>\\<^esub>" 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 "\ \ \\<^sup>+ 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 "\ \ (\\<^sup>+ z. f (the z, x) * indicator (range Some) z) + (\\<^sup>+ 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 \ \<^bold>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: "(\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T 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 = (\\<^sup>+ y\range Some. j (Some x, y))" using that disjoint by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator outgoing_def intro!: nn_integral_cong \.flowD_outside[OF j] dest: bipartite_E split: split_indicator) also have "\ = d_OUT j \x\" by(auto simp add: d_OUT_def nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_cong \.flowD_outside[OF j] split: split_indicator) finally show ?thesis . qed 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 = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T 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 "\ = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. f (x, y) + j (\y\, \x\)) - (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T 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 "\ = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. f (x, y)) + (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T 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 = (\\<^sup>+ y\range Some. j (y, \x\))" using that disjoint by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator outgoing_def intro!: nn_integral_cong \.flowD_outside[OF j] split: split_indicator dest: bipartite_E) also have "\ = d_IN j (Some x)" using that disjoint by(auto 4 3 simp add: d_IN_def nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_cong \.flowD_outside[OF j] split: split_indicator) finally show ?thesis by simp qed have OUT_g_B: "d_OUT g x = 0" if "x \ A \" for x using disjoint that by(auto simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space g_simps dest: bipartite_E) have OUT_g_a: "d_OUT g a < weight \ a" using a(1) proof - have "d_OUT g a = d_OUT f a + d_IN j \a\ - d_OUT j \a\" using a(1) by(rule OUT_g_A) also have "\ \ d_OUT f a + d_IN j \a\" by(rule diff_le_self_ennreal) also have "\ < weight \ a + d_IN j \a\ - \" using finite_\ \_less currentD_finite_OUT[OF f'] by (simp add: less_diff_eq_ennreal less_top ac_simps) also have "\ \ weight \ a" using IN_j[THEN order_trans, OF \_le] by (simp add: ennreal_minus_le_iff) finally show ?thesis . qed have OUT_jj: "d_OUT ?j x = d_IN j \x\ - j (None, \x\)" for x proof - have "d_OUT ?j x = (\\<^sup>+ y\range Some. j (y, \x\))" by(simp add: d_OUT_def nn_integral_count_space_reindex) also have "\ = d_IN j \x\ - (\\<^sup>+ 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: "(\\<^sup>+ y\\<^bold>I\<^bold>N 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 "\ = (\\<^sup>+ y\\<^bold>I\<^bold>N 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 "\ = (\\<^sup>+ y\\<^bold>I\<^bold>N x. f (y, x) + j (Some x, Some y)) - (\\<^sup>+ y\\<^bold>I\<^bold>N 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 "\ = (\\<^sup>+ y\\<^bold>I\<^bold>N x. f (y, x)) + (\\<^sup>+ y\\<^bold>I\<^bold>N 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 = (\\<^sup>+ y\range Some. j (Some x, y))" using True disjoint by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator incoming_def intro!: nn_integral_cong \.flowD_outside[OF j] split: split_indicator dest: bipartite_E) also have "\ = d_OUT j (Some x)" using disjoint by(auto 4 3 simp add: d_OUT_def nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_cong \.flowD_outside[OF j] split: split_indicator) also have "\ = d_IN j (Some x)" using flowD_KIR[OF j, of "Some x"] True a disjoint by auto also have "?j_IN = (\\<^sup>+ y\range Some. j (y, Some x))" using True disjoint by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator incoming_def intro!: nn_integral_cong \.flowD_outside[OF j] dest: bipartite_E split: split_indicator) also have "\ = d_IN j (Some x) - (\\<^sup>+ 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 \ \<^bold>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)\<^sup>*\<^sup>* 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) \ \<^bold>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) \ \<^bold>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: "(\\<^sup>+ y. j (Some y, Some x) \count_space (\<^bold>I\<^bold>N x)) \ \" for x proof - have "(\\<^sup>+ y. j (Some y, Some x) \count_space (\<^bold>I\<^bold>N 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 = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. h (x, y))" using h_le_g currentD_outside[OF g] by(intro d_OUT_alt_def) auto also have "\ = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T 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 = (\\<^sup>+ y\\<^bold>I\<^bold>N x. h (y, x))" using h_le_g currentD_outside[OF g] by(intro d_IN_alt_def) auto also have "\ = (\\<^sup>+ y\\<^bold>I\<^bold>N 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' = \\<^sup>+ y. \ * indicator (set E) (x', y) \count_space UNIV" unfolding d_OUT_def by(rule nn_integral_cong)(simp) also have "\ = (\\<^sup>+ e. \ * indicator (set E) e \embed_measure (count_space UNIV) (Pair x'))" by(simp add: measurable_embed_measure1 nn_integral_embed_measure) also have "\ = (\\<^sup>+ 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' = \\<^sup>+ y. \ * indicator (set E) (y, x') \count_space UNIV" unfolding d_IN_def by(rule nn_integral_cong)(simp) also have "\ = (\\<^sup>+ 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 "\ = (\\<^sup>+ 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) \ \<^bold>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 \ \\<^bsub>?\\<^esub> (?TER f)" using x' proof(rule contrapos_nn) assume *: "x \ \\<^bsub>?\\<^esub> (?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 + \\<^sup>+ 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 "(\\<^sup>+ y. j (Some y, Some a) \count_space UNIV) = (\\<^sup>+ 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 "\ \ \\<^sup>+ x\B \. 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 *: "(\\<^sup>+ x\B \'. ?u x) = \" using disjoint by(auto intro!: nn_integral_cong simp add: \_def nn_integral_count_space_indicator currentD_outside_OUT[OF g] currentD_outside_IN[OF g] not_vertex split: split_indicator) thus "(\\<^sup>+ x\B \'. ?u x) \ \" using \_finite by simp have **: "\'\weight := weight \' - ?u\ = \ \ g" using currentD_weight_IN[OF g] currentD_OUT_IN[OF g] currentD_IN[OF g] currentD_finite_OUT[OF g] by(intro web.equality)(simp_all add: fun_eq_iff diff_diff_ennreal' ennreal_diff_le_mono_left) show "hindered_by (\'\weight := weight \' - ?u\) (\\<^sup>+ x\B \'. ?u x)" unfolding * ** by(fact hindered) show "(\x. (d_IN g x - d_OUT g x) * indicator (- A \) x) \ weight \'" using currentD_weight_IN[OF g] by (simp add: le_fun_def ennreal_diff_le_mono_left) qed then show ?thesis by(rule hindered_mono_web[rotated -1]) simp_all qed end subsection \Reduced weight in a loose web\ definition reduce_weight :: "('v, 'more) web_scheme \ 'v \ real \ ('v, 'more) web_scheme" where "reduce_weight \ x r = \\weight := \y. weight \ y - (if x = y then r else 0)\" lemma reduce_weight_sel [simp]: "edge (reduce_weight \ x r) = edge \" "A (reduce_weight \ x r) = A \" "B (reduce_weight \ x r) = B \" "vertex (reduce_weight \ x r) = vertex \" "weight (reduce_weight \ x r) y = (if x = y then weight \ x - r else weight \ y)" "web.more (reduce_weight \ x r) = web.more \" by(simp_all add: reduce_weight_def zero_ennreal_def[symmetric] vertex_def fun_eq_iff) lemma essential_reduce_weight [simp]: "essential (reduce_weight \ x r) = essential \" by(simp add: fun_eq_iff essential_def) lemma roofed_reduce_weight [simp]: "roofed_gen (reduce_weight \ x r) = roofed_gen \" by(simp add: fun_eq_iff roofed_def) context countable_bipartite_web begin context begin private datatype (plugins del: transfer size) 'a vertex = SOURCE | SINK | Inner (inner: 'a) private lemma notin_range_Inner: "x \ range Inner \ x = SOURCE \ x = SINK" by(cases x) auto private lemma inj_Inner [simp]: "\A. inj_on Inner A" by(simp add: inj_on_def) lemma unhinder_bipartite: assumes h: "\n :: nat. current \ (h n)" and SAT: "\n. (B \ \ \<^bold>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 \ \<^bold>V" shows "\h'. current \ h' \ wave \ h' \ B \ \ \<^bold>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\ \ (\\<^sup>+ 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 "\ = (\\<^sup>+ y\range Inner. h 0 (x, inner y)) + weight \ x" by(auto simp add: nn_integral_count_space_indicator nn_integral_add) also have "(\\<^sup>+ y\range Inner. h 0 (x, inner y)) = d_OUT (h 0) x" by(simp add: d_OUT_def nn_integral_count_space_reindex) also have "\ \ weight \ x" using h by(rule currentD_weight_OUT) finally show ?thesis unfolding one_add_one[symmetric] distrib_right by(simp add: add_right_mono) qed have flow_attainability: "flow_attainability \" proof have "\<^bold>E\<^bsub>\\<^esub> \ (\(x, y). (\x\, \y\)) ` \<^bold>E \ (\(x, y). (\y\, \x\)) ` \<^bold>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 \<^bold>E\<^bsub>\\<^esub>" 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 \ \<^bold>V" | (outside) x where "v = \x\" "x \ \<^bold>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 \ (\\<^sup>+ 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 "\ = (\\<^sup>+ x\range Inner. h 0 (inner x, y)) + weight \ b" by(simp add: nn_integral_add nn_integral_count_space_indicator) also have "(\\<^sup>+ x\range Inner. h 0 (inner x, y)) = d_IN (h 0) y" by(simp add: d_IN_def nn_integral_count_space_reindex) also have "d_IN (h 0) y \ weight \ y" using h by(rule currentD_weight_IN) finally show ?thesis by(auto simp add: top_unique add_right_mono split: if_split_asm) next case outside hence "d_OUT (capacity \) v = 0" using A_vertex by(auto simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space cap_def vertex_def split: vertex.split) thus ?thesis by simp qed next show "capacity \ e \ \" for e by(auto simp add: cap_def max_def vertex_def currentD_finite[OF h] split: vertex.split prod.split) show "capacity \ e = 0" if "e \ \<^bold>E\<^bsub>\\<^esub>" 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 = (\\<^sup>+ 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 \ \<^bold>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 \ \<^bold>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 \ \<^bold>V" | (outside) x where "xo = \x\" "x \ \<^bold>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: "(\\<^sup>+ 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\ = (\\<^sup>+ 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 "\ = (\\<^sup>+ y\range Inner. f n (\x\, y)) + f n (\x\, SINK)" by(simp add: nn_integral_add nn_integral_count_space_indicator max.left_commute max.commute) also have "(\\<^sup>+ y\range Inner. f n (\x\, y)) = (\\<^sup>+ 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 "\ = (\\<^sup>+ y. h 0 (x, y) * indicator ?h0_ge_hn y) - (\\<^sup>+ 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 "(\\<^sup>+ y. h n (x, y) * indicator ?h0_ge_hn y) = d_OUT (h n) x - (\\<^sup>+ 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 "(\\<^sup>+ y. h 0 (x, y) * indicator ?h0_ge_hn y) - \ + f n (\x\, SINK) = (\\<^sup>+ y. h 0 (x, y) * indicator ?h0_ge_hn y) + (\\<^sup>+ 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 "\ = (\\<^sup>+ y. h n (x, y) * indicator ?h0_lt_hn y) - (d_OUT (h 0) x - (\\<^sup>+ 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 - (\\<^sup>+ y. h 0 (x, y) * indicator ?h0_ge_hn y) = (\\<^sup>+ 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 "(\\<^sup>+ y. h n (x, y) * indicator ?h0_lt_hn y) - \ = (\\<^sup>+ 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 "\ = (\\<^sup>+ y\range Inner. f n (y, \x\))" using A apply(simp add: nn_integral_count_space_reindex cong: nn_integral_cong_simp outgoing_def) apply(auto simp add: nn_integral_count_space_indicator outgoing_def A_in max.commute currentD_outside[OF h] intro!: nn_integral_cong split: split_indicator dest: edge_antiparallel) done also have "\ + f n (SINK, \x\) = (\\<^sup>+ 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: "(\\<^sup>+ 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\ = (\\<^sup>+ 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 "\ = (\\<^sup>+ y\range Inner. f n (y, \x\)) + f n (SOURCE, \x\)" using h0_b[of n] by(simp add: nn_integral_add nn_integral_count_space_indicator max_def) also have "(\\<^sup>+ y\range Inner. f n (y, \x\)) = (\\<^sup>+ 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 "\ = (\\<^sup>+ y. h 0 (y, x) * indicator ?h0_gt_hn y) - (\\<^sup>+ 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: "(\\<^sup>+ y. h 0 (y, x) * indicator ?h0_gt_hn y) = d_IN (h 0) x - (\\<^sup>+ 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: "(\\<^sup>+ y. h n (y, x) * indicator ?h0_gt_hn y) = d_IN (h n) x - (\\<^sup>+ 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 - (\\<^sup>+ y. h 0 (y, x) * indicator ?h0_le_hn y) - (d_IN (h n) x - (\\<^sup>+ y. h n (y, x) * indicator ?h0_le_hn y)) + f n (SOURCE, \x\) = (\\<^sup>+ y. h n (y, x) * indicator ?h0_le_hn y) - (\\<^sup>+ 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 "(\\<^sup>+ y. h n (y, x) * indicator ?h0_le_hn y) - (\\<^sup>+ y. h 0 (y, x) * indicator ?h0_le_hn y) = (\\<^sup>+ 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 "\ = (\\<^sup>+ y\range Inner. f n (\x\, y))" using B disjoint apply(simp add: nn_integral_count_space_reindex cong: nn_integral_cong_simp outgoing_def) apply(auto simp add: B_out currentD_outside[OF h] max.commute intro!: nn_integral_cong split: split_indicator dest: edge_antiparallel) done also have "\ = (\\<^sup>+ 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 = (\\<^sup>+ 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\) = (\\<^sup>+ 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 \ \<^bold>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 \ \<^bold>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 = (\\<^sup>+ y\range Inner. g (y, \x\))" by(simp add: nn_integral_count_space_reindex d_OUT_def) also have "\ = d_IN g \x\ - (\\<^sup>+ 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 = (\\<^sup>+ y\range Inner. g (\x\, y))" by(simp add: nn_integral_count_space_reindex d_IN_def) also have "\ = d_OUT g \x\ - (\\<^sup>+ 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 = (\\<^sup>+ y\range Inner. g (y, \x\))" by(simp add: nn_integral_count_space_reindex d_OUT_def) also have "\ = d_IN g \x\ - (\\<^sup>+ 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 = (\\<^sup>+ y\range Inner. g (\x\, y))" by(simp add: nn_integral_count_space_reindex d_IN_def) also have "\ = d_OUT g \x\" unfolding d_OUT_def using that disjoint by(auto 4 3 simp add: nn_integral_count_space_indicator notin_range_Inner intro!: nn_integral_cong \.flowD_outside[OF g] split: split_indicator) finally show ?thesis . qed have finite_g_IN: "d_IN ?g x \ \" for x using \_finite proof(rule neq_top_trans) have "d_IN ?g x = (\\<^sup>+ y\range Inner. g (\x\, y))" by(auto simp add: d_IN_def nn_integral_count_space_reindex) also have "\ \ d_OUT g \x\" unfolding d_OUT_def by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ = d_IN g \x\" by(rule flowD_KIR[OF g]) simp_all also have "\ \ \" using IN_g value_g by simp finally show "d_IN ?g x \ \" . qed have OUT_h'_A: "d_OUT h' x = d_OUT (h 0) x + g (\x\, SINK) - g (SINK, \x\)" if "x \ A \" for x proof - have "d_OUT h' x = d_OUT (h 0) x + (\\<^sup>+ y. ?g (x, y) * indicator \<^bold>E (x, y)) - (\\<^sup>+ y. ?g (y, x) * indicator \<^bold>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 "(\\<^sup>+ y. ?g (x, y) * indicator \<^bold>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 "(\\<^sup>+ y. ?g (y, x) * indicator \<^bold>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 = (\\<^sup>+ y\range Inner. g (y, \x\))" by(auto simp add: d_OUT_def nn_integral_count_space_reindex) also have "\ \ d_IN g \x\" unfolding d_IN_def by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ \ \" using IN_g value_g by simp finally show "d_OUT ?g x \ \" . qed have IN_h'_B: "d_IN h' x = d_IN (h 0) x + g (SOURCE, \x\)" if "x \ B \" for x proof - have g_le: "g (SOURCE, \x\) \ d_IN g \x\" by (rule d_IN_ge_point) have "d_IN h' x = d_IN (h 0) x + (\\<^sup>+ y. g (\x\, \y\) * indicator \<^bold>E (y, x)) - (\\<^sup>+ y. g (\y\, \x\) * indicator \<^bold>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 "(\\<^sup>+ y. g (\x\, \y\) * indicator \<^bold>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 "(\\<^sup>+ y. g (\y\, \x\) * indicator \<^bold>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 \ \<^bold>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 \ \<^bold>E" for e using that by(simp split: prod.split_asm add: h'_simps) qed moreover have SAT_h': "B \ \ \<^bold>V \ SAT \ h'" proof show "x \ SAT \ h'" if "x \ B \ \ \<^bold>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 \ \ \<^bold>V)" proof fix x y p assume x: "x \ A \" and y: "y \ B \" and p: "path \ x p y" hence Nil: "p \ []" using disjoint by(auto simp add: rtrancl_path_simps) from rtrancl_path_last[OF p Nil] last_in_set[OF Nil] y rtrancl_path_Range[OF p, of y] show "(\z\set p. z \ B \ \ \<^bold>V) \ x \ B \ \ \<^bold>V" by(auto intro: vertexI2) qed moreover have TER: "B \ \ \<^bold>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 \ \<^bold>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 g\: "current \ (g n)" for n by(rule current_weight_mono)(auto simp add: \_nonneg diff_le_self_ennreal) note [simp] = currentD_finite[OF g\] have b_TER: "b \ TER\<^bsub>\' n\<^esub> (g n)" for n proof(rule ccontr) assume b': "b \ TER\<^bsub>\' n\<^esub> (g n)" then have TER: "TER\<^bsub>\' n\<^esub> (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] g\[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, hide_lams) add_is_0 of_nat_eq_0_iff of_nat_numeral zero_neq_numeral) then have IN_g': "d_IN (g' n) x = (if x = b then weight \ b - \ 0 else d_IN (g n) x)" for x n using b_TER[of n] bnA factor_pos[of n] factor[of n] wb_pos \ by(auto simp add: d_IN_def g'_simps nn_integral_divide nn_integral_cmult currentD_SAT[OF g] wb_conv \_def field_simps ennreal_minus ennreal_mult'[symmetric] intro!: arg_cong[where f=ennreal]) have OUT_g': "d_OUT (g' n) x = d_OUT (g n) x - g n (x, b) * (1 - factor n)" for n x proof - have "d_OUT (g' n) x = (\\<^sup>+ y. g n (x, y)) - (\\<^sup>+ 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 \ \<^bold>E\<^bsub>\' 0\<^esub>" 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\<^bsub>\' n\<^esub> (g n) = TER\<^bsub>\' 0\<^esub> (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\<^bsub>\' n\<^esub> (g n))" using waveD_separating[OF w, of n] by(simp add: separating_gen.simps) then show "separating (\' 0) (TER\<^bsub>\' 0\<^esub> (g' n))" unfolding TER_g' . qed(rule g') define f where "f = rec_nat (g 0) (\n rec. rec \\<^bsub>\' 0\<^esub> g' (n + 1))" have f_simps [simp]: "f 0 = g 0" "f (Suc n) = f n \\<^bsub>\' 0\<^esub> g' (n + 1)" for n by(simp_all add: f_def) have f: "current (\' 0) (f n)" and fw: "wave (\' 0) (f n)" for n proof(induction n) case (Suc n) { case 1 show ?case unfolding f_simps using Suc.IH g' by(rule current_plus_web) } { case 2 show ?case unfolding f_simps using Suc.IH g' w' by(rule wave_plus') } qed(simp_all add: g w) have f_inc: "n \ m \ f n \ f m" for n m proof(induction m rule: dec_induct) case (step k) note step.IH also have "f k \ (f k \\<^bsub>\' 0\<^esub> g' (k + 1))" by(rule le_funI plus_web_greater)+ also have "\ = f (Suc k)" by simp finally show ?case . qed simp have chain_f: "Complete_Partial_Order.chain (\) (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\<^bsub>\' 0\<^esub> (f n)) = RF (\i\n. TER\<^bsub>\' 0\<^esub> (g' i))" for n proof(induction n) case 0 show ?case by(simp add: TER_g') next case (Suc n) have "RF (TER\<^bsub>\' 0\<^esub> (f (Suc n))) = RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (f n \\<^bsub>\' 0\<^esub> g' (n + 1)))" by simp also have "\ = RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (f n) \ TER\<^bsub>\' 0\<^esub> (g' (n + 1)))" using f fw g' w' by(rule RF_TER_plus_web) also have "\ = RF\<^bsub>\' 0\<^esub> (RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (f n)) \ TER\<^bsub>\' 0\<^esub> (g' (n + 1)))" by(simp add: roofed_idem_Un1) also have "RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (f n)) = RF\<^bsub>\' 0\<^esub> (\i\n. TER\<^bsub>\' 0\<^esub> (g' i))" by(simp add: Suc.IH) also have "RF\<^bsub>\' 0\<^esub> (\ \ TER\<^bsub>\' 0\<^esub> (g' (n + 1))) = RF\<^bsub>\' 0\<^esub> ((\i\n. TER\<^bsub>\' 0\<^esub> (g' i)) \ TER\<^bsub>\' 0\<^esub> (g' (n + 1)))" by(simp add: roofed_idem_Un1) also have "(\i\n. TER\<^bsub>\' 0\<^esub> (g' i)) \ TER\<^bsub>\' 0\<^esub> (g' (n + 1)) = (\i\Suc n. TER\<^bsub>\' 0\<^esub> (g' i))" unfolding atMost_Suc UN_insert by(simp add: Un_commute) finally show ?case by simp qed define g\ where "g\ = (SUP n. f n)" have g\: "current (\' 0) g\" unfolding g\_def using chain_f by(rule current_Sup)(auto simp add: f supp_f) have w\: "wave (\' 0) g\" unfolding g\_def using chain_f by(rule wave_lub)(auto simp add: fw supp_f) from g\ have g\': "current (\' n) g\" for n using wb_pos \ by(elim current_weight_mono)(auto simp add: \_le wb_conv \_def field_simps ennreal_minus_if min_le_iff_disj) have SINK_g\: "SINK g\ = (\n. SINK (f n))" unfolding g\_def by(subst SINK_Sup[OF chain_f])(simp_all add: supp_f) have SAT_g\: "SAT (\' 0) (f n) \ SAT (\' 0) g\" for n unfolding g\_def by(rule SAT_Sup_upper) simp have g_b_out: "g n (b, x) = 0" for n x using b_TER[of n] by(simp add: SINK.simps currentD_OUT_eq_0[OF g]) have g'_b_out: "g' n (b, x) = 0" for n x by(simp add: g'_simps g_b_out) have "f n (b, x) = 0" for n x by(induction n)(simp_all add: g_b_out g'_b_out) hence b_SINK_f: "b \ SINK (f n)" for n by(simp add: SINK.simps d_OUT_def) hence b_SINK_g\: "b \ SINK g\" by(simp add: SINK_g\) have RF_circ: "RF\<^sup>\\<^bsub>\' n\<^esub> (TER\<^bsub>\' 0\<^esub> (g' n)) = RF\<^sup>\\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (g' n))" for n by(simp add: roofed_circ_def) have edge_restrict_\': "edge (quotient_web (\' 0) (g' n)) = edge (quotient_web (\' n) (g n))" for n by(simp add: fun_eq_iff TER_g' RF_circ) have restrict_curr_g': "f \ \' 0 / g' n = f \ \' n / g n" for n f by(simp add: restrict_current_def RF_circ TER_g') have RF_restrict: "roofed_gen (quotient_web (\' n) (g n)) = roofed_gen (quotient_web (\' 0) (g' n))" for n by(simp add: roofed_def fun_eq_iff edge_restrict_\') have g\r': "current (quotient_web (\' 0) (g' n)) (g\ \ \' 0 / g' n)" for n using w' g\ by(rule current_restrict_current) have g\r: "current (quotient_web (\' n) (g n)) (g\ \ \' n / g n)" for n using w g\' by(rule current_restrict_current) have w\r: "wave (quotient_web (\' n) (g n)) (g\ \ \' n / g n)" (is "wave ?\' ?g\") for n proof have *: "wave (quotient_web (\' 0) (g' n)) (g\ \ \' 0 / g' n)" using g' w' g\ w\ by(rule wave_restrict_current) have "d_IN (g\ \ \' n / g n) b = 0" by(rule d_IN_restrict_current_outside roofed_greaterI b_TER)+ hence SAT_subset: "SAT (quotient_web (\' 0) (g' n)) (g\ \ \' n / g n) \ SAT ?\' (g\ \ \' n / g n)" using b_TER[of n] wb_pos by(auto simp add: SAT.simps TER_g' RF_circ wb_conv \_def field_simps ennreal_minus_if split: if_split_asm) hence TER_subset: "TER\<^bsub>quotient_web (\' 0) (g' n)\<^esub> (g\ \ \' n / g n) \ TER\<^bsub>?\'\<^esub> (g\ \ \' n / g n)" using SINK_g' by(auto simp add: restrict_curr_g') show "separating ?\' (TER\<^bsub>?\'\<^esub> ?g\)" (is "separating _ ?TER") proof fix x y p assume xy: "x \ A ?\'" "y \ B ?\'" and p: "path ?\' x p y" from p have p': "path (quotient_web (\' 0) (g' n)) x p y" by(simp add: edge_restrict_\') with waveD_separating[OF *, THEN separatingD, simplified, OF p'] TER_g'[of n] SINK_g' SAT_g' restrict_curr_g' SAT_subset xy show "(\z\set p. z \ ?TER) \ x \ ?TER" by auto qed show "d_OUT (g\ \ \' n / g n) x = 0" if "x \ RF\<^bsub>?\'\<^esub> ?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\<^bsub>\' 0\<^esub> g\) = RF (\n. TER\<^bsub>\' 0\<^esub> (g' n))" proof - have "RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> g\) = RF (\i. TER\<^bsub>\' 0\<^esub> (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\<^bsub>\' 0\<^esub> (f i)))" by(simp add: roofed_UN) also have "\ = RF (\i. \j\i. TER\<^bsub>\' 0\<^esub> (g' j))" unfolding RF_f roofed_UN by simp also have "(\i. \j\i. TER\<^bsub>\' 0\<^esub> (g' j)) = (\i. TER\<^bsub>\' 0\<^esub> (g' i))" by auto finally show ?thesis by simp qed have SAT_plus_\: "SAT (\' n) (g n \\<^bsub>\' n\<^esub> g\) = SAT (\' 0) (g' n \\<^bsub>\' 0\<^esub> g\)" for n apply(intro set_eqI) apply(simp add: SAT.simps IN_plus_current[OF g w g\r] IN_plus_current[OF g' w' g\r'] TER_g') apply(cases "d_IN (g\ \ \' n / g n) b") apply(auto simp add: SAT.simps wb_conv d_IN_plus_web IN_g') apply(simp_all add: wb_conv IN_g_b restrict_curr_g' \_def field_simps) apply(metis TER_g' b_TER roofed_greaterI)+ done have SINK_plus_\: "SINK (g n \\<^bsub>\' n\<^esub> g\) = SINK (g' n \\<^bsub>\' 0\<^esub> g\)" for n apply(rule set_eqI; simp add: SINK.simps OUT_plus_current[OF g w g\r] OUT_plus_current[OF g' w'] current_restrict_current[OF w' g\]) 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\<^bsub>\' n\<^esub> (g n \\<^bsub>\' n\<^esub> g\) = TER\<^bsub>\' 0\<^esub> (g' n \\<^bsub>\' 0\<^esub> g\)" for n by(rule set_eqI iffI)+(simp_all add: SAT_plus_\ SINK_plus_\) define h where "h n = g n \\<^bsub>\' n\<^esub> g\" 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\<^bsub>\' 0\<^esub> g\" have RF_h: "RF (TER\<^bsub>\' n\<^esub> (h n)) = RF T" for n proof - have "RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' n\<^esub> (h n)) = RF\<^bsub>\' 0\<^esub> (RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> g\) \ TER\<^bsub>\' 0\<^esub> (g' n))" unfolding h_def TER_plus_\ RF_TER_plus_web[OF g' w' g\ w\] roofed_idem_Un1 by(simp add: Un_commute) also have "\ = RF ((\n. TER\<^bsub>\' 0\<^esub> (g' n)) \ TER\<^bsub>\' 0\<^esub> (g' n))" by(simp add: RF_g\ roofed_idem_Un1) also have "\ = RF\<^bsub>\' 0\<^esub> T" unfolding T_def by(auto simp add: RF_g\ intro!: arg_cong2[where f=roofed] del: equalityI) auto finally show ?thesis by simp qed have OUT_h_nT: "d_OUT (h n) x = 0" if "x \ RF T" for n x by(rule waveD_OUT[OF hw])(simp add: RF_h that) have IN_h_nT: "d_IN (h n) x = 0" if "x \ RF T" for n x by(rule wave_not_RF_IN_zero[OF h hw])(simp add: RF_h that) have OUT_h_b: "d_OUT (h n) b = 0" for n using b_TER[of n] b_SINK_g\[THEN in_SINK_restrict_current] by(auto simp add: h_def OUT_plus_current[OF g w g\r] SINK.simps) have OUT_h_\: "d_OUT (h n) x = 0" if "x \ \ T" for x n using that apply(subst (asm) \_RF[symmetric]) apply(subst (asm) (1 2) RF_h[symmetric, of n]) apply(subst (asm) \_RF) apply(simp add: SINK.simps) done have IN_h_\: "d_IN (h n) x = weight (\' n) x" if "x \ \ T" "x \ A \" for x n using that apply(subst (asm) \_RF[symmetric]) apply(subst (asm) (1 2) RF_h[symmetric, of n]) apply(subst (asm) \_RF) apply(simp add: currentD_SAT[OF h]) done have b_SAT: "b \ SAT (\' 0) (h 0)" using b_TER[of 0] by(auto simp add: h_def SAT.simps d_IN_plus_web intro: order_trans) have b_T: "b \ T" using b_SINK_g\ b_TER by(simp add: T_def)(metis SAT_g\ subsetD f_simps(1)) have essential: "b \ \ T" proof(rule ccontr) assume "b \ \ T" hence b: "b \ \ (TER\<^bsub>\' 0\<^esub> (h 0))" proof(rule contrapos_nn) assume "b \ \ (TER\<^bsub>\' 0\<^esub> (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\<^bsub>\' 0\<^esub> (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) (\\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (h 0)))" by(rule separating_essential)(rule waveD_separating[OF hw]) then have "separating \ (\ (TER\<^bsub>\' 0\<^esub> (h 0)))" by(simp add: separating_gen.simps) moreover have subset: "\ (TER\<^bsub>\' 0\<^esub> (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\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (h 0)))" unfolding separating_RF by(rule waveD_separating[OF hw]) from separatingD[OF this, of x p y'] py p x y' have "x \ RF T \ y' \ RF T" by(auto simp add: RF_h) thus "(\z\set p. z \ RF (T \ B \)) \ x \ RF (T \ B \)" proof cases case right with y' py show ?thesis by(simp add: RF_in_B) next case left have "x \ \ T" using OUT_h_\[of x n] xy by(auto simp add: currentD_OUT_eq_0[OF h]) with left have "x \ RF\<^sup>\ 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 "\<^bold>V\<^bsub>\_h\<^esub> \ A \_h \ B \_h" by(auto simp add: vertex_def) show "A \_h \ \<^bold>V\<^bsub>\_h\<^esub>" 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 "\<^bold>E\<^bsub>\_h\<^esub> \ \<^bold>E" by auto thus "countable \<^bold>E\<^bsub>\_h\<^esub>" by(rule countable_subset) simp show "weight \_h x \ \" for x by(simp split: split_indicator) show "weight \_h x = 0" if "x \ \<^bold>V\<^bsub>\_h\<^esub>" 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\<^bsub>\' 0\<^esub> (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\<^bsub>\' 0\<^esub> (h 0)" unfolding RF_h[of 0, symmetric] by(blast intro: roofed_greaterI) show "essential \ (B \) (TER\<^bsub>\' 0\<^esub> (h 0)) x" using p y by(blast intro: essentialI dest: bypass') next fix x assume "essential \ (B \) (TER\<^bsub>\' 0\<^esub> (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\<^bsub>\' 0\<^esub> (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 \ \<^bold>E\<^bsub>\_h\<^esub>" 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 \ \<^bold>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 \ \<^bold>V\<^bsub>\_h\<^esub> - {b} \ SAT \_h (h n)" for n proof fix x assume "x \ B \_h \ \<^bold>V\<^bsub>\_h\<^esub> - {b}" then have x: "x \ T" and B: "x \ B \" and b: "x \ b" and vertex: "x \ \<^bold>V\<^bsub>\_h\<^esub>" 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 \ \<^bold>V\<^bsub>\_h\<^esub>" using b wb essential by(auto dest: B_vertex) ultimately have "\h'. current \_h h' \ wave \_h h' \ B \_h \ \<^bold>V\<^bsub>\_h\<^esub> \ 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 \ \<^bold>V\<^bsub>\_h\<^esub> \ 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 \ \<^bold>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 w\]) hence "separating \ T" by(simp add: separating_gen.simps) hence *: "separating \ (\ T)" by(rule separating_essential) show "separating \ (TER h')" proof fix x p y assume x: "x \ A \" and p: "path \ x p y" and y: "y \ B \" from p x y disjoint have py: "p = [y]" by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) from separatingD[OF * p x y] py have "x \ \ T \ y \ \ T" by auto then show "(\z\set p. z \ TER h') \ x \ TER h'" proof cases case left then have "x \ \<^bold>V\<^bsub>\_h\<^esub>" 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 \\ - f\<^sub>i\ is loose in order to apply Lemma 6.9. Counterexample: The bipartite web with three nodes \a\<^sub>1\, \a\<^sub>2\, \a\<^sub>3\ in \A\ and two nodes \b\<^sub>1\, \b\<^sub>2\ in \B\ and edges \(a\<^sub>1, b\<^sub>1)\, \(a\<^sub>2, b\<^sub>1)\, \(a\<^sub>2, b\<^sub>2)\, \(a\<^sub>3, b\<^sub>2)\ and weights \a\<^sub>1 = a\<^sub>3 = 1\ and \a\<^sub>2 = 2\ and \b\<^sub>1 = 3\ and \b\<^sub>2 = 2\. Then, we can get a sequence of weight reductions on \b\<^sub>2\ 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 \a\<^sub>2\ and \b\<^sub>2\. Now, the restricted web for the two is not loose because it contains the wave which assigns 1 to \(a\<^sub>3, b\<^sub>2)\. 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 \ \<^bold>E" for e using that currentD_outside'[OF f, of e] currentD_outside'[OF g, of e] by(cases e) simp qed lemma wave_plus_current_minus: assumes f: "current \ f" and w: "wave \ f" and g: "current (\ \ f) g" and w': "wave (\ \ f) g" shows "wave \ (plus_current f g)" (is "wave _ ?fg") proof show fg: "current \ ?fg" using f g by(rule current_plus_current_minus) show sep: "separating \ (TER ?fg)" proof fix x p y assume x: "x \ A \" and p: "path \ x p y" and y: "y \ B \" from p x y disjoint have py: "p = [y]" by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) with waveD_separating[THEN separatingD, OF w p x y] have "x \ TER f \ y \ TER f" by auto thus "(\z\set p. z \ TER ?fg) \ x \ TER ?fg" proof cases case right with y disjoint have "y \ TER ?fg" using currentD_OUT[OF fg y] by(auto simp add: SAT.simps SINK.simps d_IN_def nn_integral_add not_le add_increasing2) thus ?thesis using py by simp next case x': left from p have "path (\ \ f) x p y" by simp from waveD_separating[THEN separatingD, OF w' this] x y py have "x \ TER\<^bsub>\ \ f\<^esub> g \ y \ TER\<^bsub>\ \ f\<^esub> g" by auto thus ?thesis proof cases case left hence "x \ TER ?fg" using x x' by(auto simp add: SAT.simps SINK.simps d_OUT_def nn_integral_add) thus ?thesis .. next case right hence "y \ TER ?fg" using disjoint y currentD_OUT[OF fg y] currentD_OUT[OF f y] currentD_finite_IN[OF f, of y] by(auto simp add: add.commute SINK.simps SAT.simps d_IN_def nn_integral_add ennreal_minus_le_iff split: if_split_asm) with py show ?thesis by auto qed qed qed qed lemma minus_plus_current: assumes f: "current \ f" and g: "current (\ \ f) g" shows "\ \ plus_current f g = \ \ f \ g" (is "?lhs = ?rhs") proof(rule web.equality) have "weight ?lhs x = weight ?rhs x" for x using currentD_weight_IN[OF f, of x] currentD_weight_IN[OF g, of x] by (auto simp add: d_IN_def d_OUT_def nn_integral_add diff_add_eq_diff_diff_swap_ennreal add_increasing2 diff_add_assoc2_ennreal add.assoc) thus "weight ?lhs = weight ?rhs" .. qed simp_all lemma unhindered_minus_web: assumes unhindered: "\ hindered \" and f: "current \ f" and w: "wave \ f" shows "\ hindered (\ \ f)" proof assume "hindered (\ \ f)" then obtain g where g: "current (\ \ f) g" and w': "wave (\ \ f) g" and hind: "hindrance (\ \ f) g" by cases let ?fg = "plus_current f g" have fg: "current \ ?fg" using f g by(rule current_plus_current_minus) moreover have "wave \ ?fg" using f w g w' by(rule wave_plus_current_minus) moreover from hind obtain a where a: "a \ A \" and n\: "a \ \\<^bsub>\ \ f\<^esub> (TER\<^bsub>\ \ f\<^esub> g)" and wa: "d_OUT g a < weight (\ \ f) a" by cases auto from a have "hindrance \ ?fg" proof show "a \ \ (TER ?fg)" proof assume \: "a \ \ (TER ?fg)" then obtain p y where p: "path \ a p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER ?fg)" by(rule \_E_RF) blast from p a y disjoint have py: "p = [y]" by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) from bypass[of y] py have "y \ TER ?fg" by(auto intro: roofed_greaterI) with currentD_OUT[OF fg y] have "y \ SAT \ ?fg" by(auto simp add: SINK.simps) hence "y \ SAT (\ \ f) g" using y currentD_OUT[OF f y] currentD_finite_IN[OF f, of y] by(auto simp add: SAT.simps d_IN_def nn_integral_add ennreal_minus_le_iff add.commute) hence "essential (\ \ f) (B (\ \ f)) (TER\<^bsub>\ \ f\<^esub> g) a" using p py y by(auto intro!: essentialI) moreover from \ a have "a \ TER\<^bsub>\ \ f\<^esub> g" by(auto simp add: SAT.A SINK_plus_current) ultimately have "a \ \\<^bsub>\ \ f\<^esub> (TER\<^bsub>\ \ f\<^esub> g)" by blast thus False using n\ by contradiction qed show "d_OUT ?fg a < weight \ a" using a wa currentD_finite_OUT[OF f, of a] by(simp add: d_OUT_def less_diff_eq_ennreal less_top add.commute nn_integral_add) qed ultimately have "hindered \" by(blast intro: hindered.intros) with unhindered show False by contradiction qed lemma loose_minus_web: assumes unhindered: "\ hindered \" and f: "current \ f" and w: "wave \ f" and maximal: "\w. \ current \ w; wave \ w; f \ w \ \ f = w" shows "loose (\ \ f)" (is "loose ?\") proof fix g assume g: "current ?\ g" and w': "wave ?\ g" let ?g = "plus_current f g" from f g have "current \ ?g" by(rule current_plus_current_minus) moreover from f w g w' have "wave \ ?g" by(rule wave_plus_current_minus) moreover have "f \ ?g" by(clarsimp simp add: le_fun_def) ultimately have eq: "f = ?g" by(rule maximal) have "g e = 0" for e proof(cases e) case (Pair x y) have "f e \ d_OUT f x" unfolding d_OUT_def Pair by(rule nn_integral_ge_point) simp also have "\ \ weight \ x" by(rule currentD_weight_OUT[OF f]) also have "\ < \" by(simp add: less_top[symmetric]) finally show "g e = 0" using Pair eq[THEN fun_cong, of e] by(cases "f e" "g e" rule: ennreal2_cases)(simp_all add: fun_eq_iff) qed thus "g = (\_. 0)" by(simp add: fun_eq_iff) next show "\ hindrance ?\ zero_current" using unhindered proof(rule contrapos_nn) assume "hindrance ?\ zero_current" then obtain x where a: "x \ A ?\" and \: "x \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> zero_current)" and weight: "d_OUT zero_current x < weight ?\ x" by cases have "hindrance \ f" proof show a': "x \ A \" using a by simp with weight show "d_OUT f x < weight \ x" by(simp add: less_diff_eq_ennreal less_top[symmetric] currentD_finite_OUT[OF f]) show "x \ \ (TER f)" proof assume "x \ \ (TER f)" then obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER f)" by(rule \_E_RF) auto from p a' y disjoint have py: "p = [y]" by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) hence "y \ (TER f)" using bypass[of y] by(auto intro: roofed_greaterI) hence "weight ?\ y > 0" using currentD_OUT[OF f y] disjoint y by(auto simp add: SINK.simps SAT.simps diff_gr0_ennreal) hence "y \ TER\<^bsub>?\\<^esub> zero_current" using y disjoint by(auto) hence "essential ?\ (B ?\) (TER\<^bsub>?\\<^esub> zero_current) x" using p y py by(auto intro!: essentialI) with a have "x \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> zero_current)" by simp with \ show False by contradiction qed qed thus "hindered \" using f w .. qed qed lemma weight_minus_web: assumes f: "current \ f" shows "weight (\ \ f) x = (if x \ A \ then weight \ x - d_OUT f x else weight \ x - d_IN f x)" proof(cases "x \ B \") case True with currentD_OUT[OF f True] disjoint show ?thesis by auto next case False hence "d_IN f x = 0" "d_OUT f x = 0" if "x \ A \" using currentD_outside_OUT[OF f, of x] currentD_outside_IN[OF f, of x] bipartite_V that by auto then show ?thesis by simp qed lemma (in -) separating_minus_web [simp]: "separating_gen (G \ f) = separating_gen G" by(simp add: separating_gen.simps fun_eq_iff) lemma current_minus: assumes f: "current \ f" and g: "current \ g" and le: "\e. g e \ f e" shows "current (\ \ g) (f - g)" proof - interpret \: countable_bipartite_web "\ \ g" using g by(rule countable_bipartite_web_minus_web) note [simp del] = minus_web_sel(2) and [simp] = weight_minus_web[OF g] show ?thesis proof show "d_OUT (f - g) x \ weight (\ \ g) x" for x unfolding fun_diff_def using currentD_weight_OUT[OF f, of x] currentD_weight_IN[OF g, of x] by(subst d_OUT_diff)(simp_all add: le currentD_finite_OUT[OF g] currentD_OUT'[OF f] currentD_OUT'[OF g] ennreal_minus_mono) show "d_IN (f - g) x \ weight (\ \ g) x" for x unfolding fun_diff_def using currentD_weight_IN[OF f, of x] currentD_weight_OUT[OF g, of x] by(subst d_IN_diff)(simp_all add: le currentD_finite_IN[OF g] currentD_IN[OF f] currentD_IN[OF g] ennreal_minus_mono) show "(f - g) e = 0" if "e \ \<^bold>E\<^bsub>\ \ g\<^esub>" 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\<^bsub>\ \ g\<^esub> (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\<^bsub>\ \ g\<^esub> (f - g)" by(auto) from w have "separating \ (TER f)" by(rule waveD_separating) thus "separating (\ \ g) (TER\<^bsub>\ \ g\<^esub> (f - g))" using TER by(simp add: separating_weakening) fix x assume "x \ RF\<^bsub>\ \ g\<^esub> (TER\<^bsub>\ \ g\<^esub> (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 = (\\<^sup>+ y'. f (x, y') + k * indicator {y} y') - (\\<^sup>+ 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 "(\\<^sup>+ y'. f (x, y') + k * indicator {y} y') = d_OUT f x + (\\<^sup>+ 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 "\ - (\\<^sup>+ 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 = (\\<^sup>+ 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)) \ \<^bold>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 \ \<^bold>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: "(\\h\M. d_OUT (fst \h) x) \ weight \ x" for x using currentD_weight_OUT[OF SM1, of x] OUT1[of x] by(simp add: split_beta) have OUT1_nonneg: "0 \ (\\h\M. d_OUT (fst \h) x)" for x using in_M by(rule SUP_upper2)(simp add: ) have IN1: "d_IN (Sup (fst ` M)) x = (SUP (\, h)\M. d_IN \ x)" for x by (subst d_IN_Sup [OF chain1 _ supp_flow1]) (simp_all add: nempty split_beta image_comp) have IN1_le: "(\\h\M. d_IN (fst \h) x) \ weight \ x" for x using currentD_weight_IN[OF SM1, of x] IN1[of x] by(simp add: split_beta) have IN1_nonneg: "0 \ (\\h\M. d_IN (fst \h) x)" for x using in_M by(rule SUP_upper2) simp have IN1': "d_IN (Sup (fst ` M)) x = (SUP (\, h)\M. \ (a, x))" for x unfolding IN1 by(rule SUP_cong[OF refl])(auto dest!: Chains_FieldD[OF M] IN_\) have directed: "\\k''\M. F (snd \k) + F (fst \k') \ F (snd \k'') + F (fst \k'')" if mono: "\f g. (\z. f z \ g z) \ F f \ F g" "\k \ M" "\k' \ M" for \k \k' and F :: "_ \ ennreal" using chainD[OF chain that(2-3)] proof cases case left hence "snd \k \ snd \k'" by(simp add: leq_def less_eq_prod_def in_restrict_rel_iff) hence "F (snd \k) + F (fst \k') \ F (snd \k') + F (fst \k')" by(intro add_right_mono mono)(clarsimp simp add: le_fun_def) with that show ?thesis by blast next case right hence "fst \k' \ fst \k" by(simp add: leq_def less_eq_prod_def in_restrict_rel_iff) hence "F (snd \k) + F (fst \k') \ F (snd \k) + F (fst \k)" by(intro add_left_mono mono)(clarsimp simp add: le_fun_def) with that show ?thesis by blast qed have directed_OUT: "\\k''\M. d_OUT (snd \k) x + d_OUT (fst \k') x \ d_OUT (snd \k'') x + d_OUT (fst \k'') x" if "\k \ M" "\k' \ M" for x \k \k' by(rule directed; rule d_OUT_mono that) have directed_IN: "\\k''\M. d_IN (snd \k) x + d_IN (fst \k') x \ d_IN (snd \k'') x + d_IN (fst \k'') x" if "\k \ M" "\k' \ M" for x \k \k' by(rule directed; rule d_IN_mono that) let ?\ = "\ \ Sup (fst ` M)" have hM2: "current ?\ h" if \h: "(\, h) \ M" for \ h proof from \h have Field: "(\, h) \ Field leq" by(rule Chains_FieldD[OF M]) then have H: "current (\ \ \) h" and \_curr': "current \ \" by(rule h \_curr)+ from \_curr' interpret \: countable_bipartite_web "\ \ \" by(rule countable_bipartite_web_minus_web) fix x have "d_OUT h x \ d_OUT ?h x" using \h by(intro d_OUT_mono)(auto intro: SUP_upper2) also have OUT: "\ = (SUP h\snd ` M. d_OUT h x)" using chain2 _ supp_flow2 by(rule d_OUT_Sup)(simp_all add: nempty) also have "\ = \ + (SUP \\fst ` M. d_OUT \ x) - (SUP \\fst ` M. d_OUT \ x)" using OUT1_le[of x] by (intro ennreal_add_diff_cancel_right[symmetric] neq_top_trans[OF weight_finite, of _ x]) (simp 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 h\snd ` M. d_IN h x)" using chain2 _ supp_flow2 by(rule d_IN_Sup)(simp_all add: nempty) also have "\ = \ + (SUP \\fst ` M. d_IN \ x) - (SUP \\fst ` M. d_IN \ x)" using IN1_le[of x] by (intro ennreal_add_diff_cancel_right [symmetric] neq_top_trans [OF weight_finite, of _ x]) (simp 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 \ \<^bold>E\<^bsub>?\\<^esub>" 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\<^bsub>?\'\<^esub> h \ TER\<^bsub>?\\<^esub> h" using currentD_OUT'[OF SM1] currentD_OUT'[OF \_curr[OF Chains_FieldD[OF M that]]] that by(auto 4 7 elim!: SAT.cases intro: SAT.intros elim!: order_trans[rotated] intro: ennreal_minus_mono d_IN_mono intro!: SUP_upper2 split: if_split_asm) from h_w[OF Chains_FieldD[OF M], OF that] have "separating ?\' (TER\<^bsub>?\'\<^esub> h)" by(rule waveD_separating) then show "separating ?\ (TER\<^bsub>?\\<^esub> h)" using subset by(auto intro: separating_weakening) qed(rule hM2[OF that]) show wave: "wave ?\ ?h" using chain2 \snd ` M \ {}\ _ supp_flow2 by(rule wave_lub)(clarify; rule wM; simp) 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 \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> g)" 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 (\\<^sup>+ 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 "(\\<^sup>+ 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 = (\\<^sup>+ x. (if edge \ x y then Sup (fst ` M) (a, y) - \ (a, y) else 0) * indicator {a} x)" unfolding d_IN_def by(rule nn_integral_cong)(auto simp add: k_alt outgoing_def split: split_indicator) also have "\ = (if edge \ a y then Sup (fst ` M) (a, y) - \ (a, y) else 0)" using hM by(auto simp add: max_def intro!: SUP_upper2) finally show ?thesis . qed have OUT_\h: "d_OUT ?\h x = d_OUT \ x + d_OUT h x" for x unfolding plus_current_def by(rule d_OUT_add)+ have IN_\h: "d_IN ?\h x = d_IN \ x + d_IN h x" for x unfolding plus_current_def by(rule d_IN_add)+ have OUT1_le': "d_OUT (Sup (fst`M)) x \ weight \ x" for x using OUT1_le[of x] unfolding OUT1 by (simp add: split_beta') have k: "current (\ \ ?\h) k" proof fix x show "d_OUT k x \ weight (\ \ ?\h) x" using a OUT1_na[of x] currentD_weight_OUT[OF hM2[OF hM], of x] currentD_weight_IN[OF \h_curr, of x] currentD_weight_IN[OF \, of x] OUT1_le'[of x] apply(auto simp add: diff_add_eq_diff_diff_swap_ennreal diff_add_assoc2_ennreal[symmetric] OUT_k OUT_\ OUT_\h IN_\h currentD_OUT'[OF \] IN_\[OF Field] h.currentD_OUT'[OF h_curr]) apply(subst diff_diff_commute_ennreal) apply(intro ennreal_minus_mono) apply(auto simp add: ennreal_le_minus_iff ac_simps less_imp_le OUT1) done have *: "(\xa\M. fst xa (a, x)) \ d_IN (Sup (fst`M)) x" unfolding IN1 by (intro SUP_subset_mono) (auto simp: split_beta' d_IN_ge_point) also have "\ \ weight \ x" using IN1_le[of x] IN1 by (simp add: split_beta') finally show "d_IN k x \ weight (\ \ ?\h) x" using currentD_weight_IN[OF \h_curr, of x] currentD_weight_OUT[OF \h_curr, of x] currentD_weight_IN[OF hM2[OF hM], of x] IN_\[OF Field, of x] * apply (auto simp add: IN_k outgoing_def IN_\h IN_\ A_in diff_add_eq_diff_diff_swap_ennreal) apply (subst diff_diff_commute_ennreal) apply (intro ennreal_minus_mono[OF _ order_refl]) apply (auto simp add: ennreal_le_minus_iff ac_simps image_comp intro: order_trans add_mono) done show "k e = 0" if "e \ \<^bold>E\<^bsub>\ \ ?\h\<^esub>" for e using that by (cases e) (simp add: k_alt) qed define q where "q = (\\<^sup>+ y\B (\ \ ?\h). d_IN k y - d_OUT k y)" have q_alt: "q = (\\<^sup>+ 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 = (\\<^sup>+ 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 \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> g')" proof assume "z \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> g')" hence OUT_z: "d_OUT g' z = 0" and ess: "essential ?\ (B \) (TER\<^bsub>?\\<^esub> g') z" by(simp_all add: SINK.simps) from ess obtain p y where p: "path \ z p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER\<^bsub>?\\<^esub> g')" by(rule essentialE_RF) auto from y have y': "y \ A \" using disjoint by blast from p z y obtain py: "p = [y]" and edge: "edge \ z y" using disjoint by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) hence yRF: "y \ RF (TER\<^bsub>?\\<^esub> g')" using bypass[of y] by(auto) with wave_not_RF_IN_zero[OF g' g'_w, of y] have IN_g'_y: "d_IN g' y = 0" by(auto intro: roofed_greaterI) with yRF y y' have w_y: "weight ?\ y > 0" using currentD_OUT[OF g', of y] by(auto simp add: RF_in_B currentD_SAT[OF g'] SINK.simps zero_less_iff_neq_zero) have "y \ SAT (\ \ f) g" proof assume "y \ SAT (\ \ f) g" with y disjoint have IN_g_y: "d_IN g y = weight (\ \ f) y" by(auto simp add: currentD_SAT[OF g]) have "0 < weight \ y - d_IN (\x\M. fst x) y - d_IN h y" using y' w_y unfolding ** by auto have "d_IN g' y > 0" using y' w_y hM unfolding ** apply(simp add: IN_g' IN_f IN_g_y diff_add_eq_diff_diff_swap_ennreal) apply(subst add_diff_eq_ennreal) apply(auto intro!: SUP_upper2 d_IN_mono simp: diff_add_self_ennreal diff_gt_0_iff_gt_ennreal) done with IN_g'_y show False by simp qed then have "y \ TER\<^bsub>\ \ f\<^esub> g" by simp with p y py have "essential \ (B \) (TER\<^bsub>\ \ f\<^esub> g) z" by(auto intro: essentialI) moreover with z waveD_separating[OF g_w, THEN separating_RF_A] have "z \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> g)" by(auto simp add: RF_in_essential) with z\ show False by contradiction qed have "\ \ weight ?\ z - d_OUT g' z" unfolding ** OUT_g' using z apply (simp add: \_def OUT_f diff_add_eq_diff_diff_swap_ennreal) apply (subst (5) diff_diff_commute_ennreal) apply (rule ennreal_minus_mono[OF _ order_refl]) apply (auto simp add: ac_simps diff_add_eq_diff_diff_swap_ennreal[symmetric] add_diff_self_ennreal 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 \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\\<^esub> 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\<^bsub>\ \ F \h\<^esub> k" if that: "\h \ Field leq" and k: "current (\ \ F \h) k" and k_w: "wave (\ \ F \h) k" and less: "d_OUT (plus_current (F \h) k) a < weight \ a" for \h k proof(rule ccontr) assume "\ ?thesis" hence \: "a \ \\<^bsub>\ \ F \h\<^esub> (TER\<^bsub>\ \ F \h\<^esub> k)" by auto from that have f: "current \ (F \h)" and unhindered: "\ hindered (\ \ F \h)" by(cases \h; simp add: f unhindered'; fail)+ from less have "d_OUT k a < weight (\ \ F \h) a" using a currentD_finite_OUT[OF f, of a] by(simp add: d_OUT_def nn_integral_add less_diff_eq_ennreal add.commute less_top[symmetric]) with _ \ have "hindrance (\ \ F \h) k" by(rule hindrance)(simp add: a) then have "hindered (\ \ F \h)" using k k_w .. with unhindered show False by contradiction qed note minus_web_sel(2)[simp del] let ?P_y = "\\h k y. y \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\ \ F \h \ k\<^esub> a \ weight (\ \ F \h \ k) y > 0" have Ex_y: "Ex (?P_y \h k)" if that: "\h \ Field leq" and k: "current (\ \ F \h) k" and k_w: "wave (\ \ F \h) k" and less: "d_OUT (plus_current (F \h) k) a < weight \ a" for \h k proof(rule ccontr) let ?\ = "\ \ F \h \ k" assume *: "\ ?thesis" interpret \: countable_bipartite_web "\ \ F \h" using f[OF that] by(rule countable_bipartite_web_minus_web) note [simp] = weight_minus_web[OF f[OF that]] \.weight_minus_web[OF k] have "hindrance ?\ zero_current" proof show "a \ A ?\" using a by simp show "a \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> zero_current)" (is "a \ \\<^bsub>_\<^esub> ?TER") proof assume "a \ \\<^bsub>?\\<^esub> ?TER" then obtain p y where p: "path ?\ a p y" and y: "y \ B ?\" and bypass: "\z. z \ set p \ z \ RF\<^bsub>?\\<^esub> ?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\<^bsub>?\\<^esub> ?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\<^bsub>\ \ F (\, h)\<^esub> 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 \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\\<^esub> 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 \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\\<^esub> 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\<^bsub>\ \ F (\, h)\<^esub> (TER\<^bsub>\ \ F (\, h)\<^esub> 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 = (\\<^sup>+ 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 = (\\<^sup>+ 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 \ \<^bold>E\<^bsub>?\\<^esub>" 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\<^bsub>\ \ \\<^esub> (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\<^bsub>\ \ \\<^esub> (plus_current h k) \ TER\<^bsub>\ \ ?\'\<^esub> (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\<^bsub>\ \ ?\'\<^esub> ?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\<^bsub>\ \ F f\<^esub> k" using Field k k_w less by(rule a_TER) then have "weight \ a = weight \ a - d_OUT (F f) a" using a disjoint by(auto simp add: roofed_circ_def \_def SINK.simps) then have weight_a: "0 < weight \ a" using less a_\ by(simp add: OUT_fk SINK.simps diff_gr0_ennreal) let ?P_y = "\y. y \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\\<^esub> 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 \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\\<^esub> 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 add: ) } { case 2 thus ?case by simp } { case 3 thus ?case using unhindered by simp } next case (Suc n) interpret \: countable_bipartite_web "\ \ f n" using Suc.IH(1) by(rule countable_bipartite_web_minus_web) 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 \ \<^bold>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 edge_antiparallel: "\x y. edge \ x y \ \ edge \ y x" and capacity_pos: "\e. e \ \<^bold>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 \ \<^bold>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 \ \<^bold>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 diff --git a/thys/MFMC_Countable/MFMC_Web.thy b/thys/MFMC_Countable/MFMC_Web.thy --- a/thys/MFMC_Countable/MFMC_Web.thy +++ b/thys/MFMC_Countable/MFMC_Web.thy @@ -1,1974 +1,1973 @@ theory MFMC_Web imports MFMC_Network begin section \Webs and currents\ record 'v web = "'v graph" + weight :: "'v \ ennreal" A :: "'v set" B :: "'v set" lemma vertex_weight_update [simp]: "vertex (weight_update f \) = vertex \" by(simp add: vertex_def fun_eq_iff) type_synonym 'v current = "'v edge \ ennreal" inductive current :: "('v, 'more) web_scheme \ 'v current \ bool" for \ f where current: "\ \x. d_OUT f x \ weight \ x; \x. d_IN f x \ weight \ x; \x. x \ A \ \ d_OUT f x \ d_IN f x; \a. a \ A \ \ d_IN f a = 0; \b. b \ B \ \ d_OUT f b = 0; \e. e \ \<^bold>E\<^bsub>\\<^esub> \ f e = 0 \ \ current \ f" lemma currentD_weight_OUT: "current \ f \ d_OUT f x \ weight \ x" by(simp add: current.simps) lemma currentD_weight_IN: "current \ f \ d_IN f x \ weight \ x" by(simp add: current.simps) lemma currentD_OUT_IN: "\ current \ f; x \ A \ \ \ d_OUT f x \ d_IN f x" by(simp add: current.simps) lemma currentD_IN: "\ current \ f; a \ A \ \ \ d_IN f a = 0" by(simp add: current.simps) lemma currentD_OUT: "\ current \ f; b \ B \ \ \ d_OUT f b = 0" by(simp add: current.simps) lemma currentD_outside: "\ current \ f; \ edge \ x y \ \ f (x, y) = 0" by(blast elim: current.cases) lemma currentD_outside': "\ current \ f; e \ \<^bold>E\<^bsub>\\<^esub> \ \ f e = 0" by(blast elim: current.cases) lemma currentD_OUT_eq_0: assumes "current \ f" shows "d_OUT f x = 0 \ (\y. f (x, y) = 0)" by(simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0) lemma currentD_IN_eq_0: assumes "current \ f" shows "d_IN f x = 0 \ (\y. f (y, x) = 0)" by(simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0) lemma current_support_flow: fixes \ (structure) assumes "current \ f" shows "support_flow f \ \<^bold>E" using currentD_outside[OF assms] by(auto simp add: support_flow.simps intro: ccontr) lemma currentD_outside_IN: "\ current \ f; x \ \<^bold>V\<^bsub>\\<^esub> \ \ d_IN f x = 0" by(auto simp add: d_IN_def vertex_def nn_integral_0_iff AE_count_space emeasure_count_space_eq_0 dest: currentD_outside) lemma currentD_outside_OUT: "\ current \ f; x \ \<^bold>V\<^bsub>\\<^esub> \ \ d_OUT f x = 0" by(auto simp add: d_OUT_def vertex_def nn_integral_0_iff AE_count_space emeasure_count_space_eq_0 dest: currentD_outside) lemma currentD_weight_in: "current \ h \ h (x, y) \ weight \ y" by (metis order_trans d_IN_ge_point currentD_weight_IN) lemma currentD_weight_out: "current \ h \ h (x, y) \ weight \ x" by (metis order_trans d_OUT_ge_point currentD_weight_OUT) lemma current_leI: fixes \ (structure) assumes f: "current \ f" and le: "\e. g e \ f e" and OUT_IN: "\x. x \ A \ \ d_OUT g x \ d_IN g x" shows "current \ g" proof show "d_OUT g x \ weight \ x" for x using d_OUT_mono[of g x f, OF le] currentD_weight_OUT[OF f] by(rule order_trans) show "d_IN g x \ weight \ x" for x using d_IN_mono[of g x f, OF le] currentD_weight_IN[OF f] by(rule order_trans) show "d_IN g a = 0" if "a \ A \" for a using d_IN_mono[of g a f, OF le] currentD_IN[OF f that] by auto show "d_OUT g b = 0" if "b \ B \" for b using d_OUT_mono[of g b f, OF le] currentD_OUT[OF f that] by auto show "g e = 0" if "e \ \<^bold>E" for e using currentD_outside'[OF f that] le[of e] by simp qed(blast intro: OUT_IN)+ lemma current_weight_mono: "\ current \ f; edge \ = edge \'; A \ = A \'; B \ = B \'; \x. weight \ x \ weight \' x \ \ current \' f" by(auto 4 3 elim!: current.cases intro!: current.intros intro: order_trans) abbreviation (input) zero_current :: "'v current" where "zero_current \ \_. 0" lemma SINK_0 [simp]: "SINK zero_current = UNIV" by(auto simp add: SINK.simps) lemma current_0 [simp]: "current \ zero_current" by(auto simp add: current.simps) inductive web_flow :: "('v, 'more) web_scheme \ 'v current \ bool" for \ (structure) and f where web_flow: "\ current \ f; \x. \ x \ \<^bold>V; x \ A \; x \ B \ \ \ KIR f x \ \ web_flow \ f" lemma web_flowD_current: "web_flow \ f \ current \ f" by(erule web_flow.cases) lemma web_flowD_KIR: "\ web_flow \ f; x \ A \; x \ B \ \ \ KIR f x" apply(cases "x \ \<^bold>V\<^bsub>\\<^esub>") apply(fastforce elim!: web_flow.cases) apply(auto simp add: vertex_def d_OUT_def d_IN_def elim!: web_flow.cases) apply(subst (1 2) currentD_outside[of _ f]; auto) done subsection \Saturated and terminal vertices\ inductive_set SAT :: "('v, 'more) web_scheme \ 'v current \ 'v set" for \ f where A: "x \ A \ \ x \ SAT \ f" | IN: "d_IN f x \ weight \ x \ x \ SAT \ f" \ \We use @{text "\ weight"} such that @{text SAT} is monotone w.r.t. increasing currents\ lemma SAT_0 [simp]: "SAT \ zero_current = A \ \ {x. weight \ x \ 0}" by(auto simp add: SAT.simps) lemma SAT_mono: assumes "\e. f e \ g e" shows "SAT \ f \ SAT \ g" proof fix x assume "x \ SAT \ f" thus "x \ SAT \ g" proof cases case IN also have "d_IN f x \ d_IN g x" using assms by(rule d_IN_mono) finally show ?thesis .. qed(rule SAT.A) qed lemma SAT_Sup_upper: "f \ Y \ SAT \ f \ SAT \ (Sup Y)" by(rule SAT_mono)(rule Sup_upper[THEN le_funD]) lemma currentD_SAT: assumes "current \ f" shows "x \ SAT \ f \ x \ A \ \ d_IN f x = weight \ x" using currentD_weight_IN[OF assms, of x] by(auto simp add: SAT.simps) abbreviation terminal :: "('v, 'more) web_scheme \ 'v current \ 'v set" ("TER\") where "terminal \ f \ SAT \ f \ SINK f" subsection \Separation\ inductive separating_gen :: "('v, 'more) graph_scheme \ 'v set \ 'v set \ 'v set \ bool" for G A B S where separating: "(\x y p. \ x \ A; y \ B; path G x p y \ \ (\z \ set p. z \ S) \ x \ S) \ separating_gen G A B S" abbreviation separating :: "('v, 'more) web_scheme \ 'v set \ bool" where "separating \ \ separating_gen \ (A \) (B \)" abbreviation separating_network :: "('v, 'more) network_scheme \ 'v set \ bool" where "separating_network \ \ separating_gen \ {source \} {sink \}" lemma separating_networkI [intro?]: "(\p. path \ (source \) p (sink \) \ (\z \ set p. z \ S) \ source \ \ S) \ separating_network \ S" by(auto intro: separating) lemma separatingD: "\A B. \ separating_gen G A B S; path G x p y; x \ A; y \ B \ \ (\z \ set p. z \ S) \ x \ S" by(blast elim: separating_gen.cases) lemma separating_left [simp]: "\A B. A \ A' \ separating_gen \ A B A'" by(auto simp add: separating_gen.simps) lemma separating_weakening: "\A B. \ separating_gen G A B S; S \ S' \ \ separating_gen G A B S'" by(rule separating; drule (3) separatingD; blast) definition essential :: "('v, 'more) graph_scheme \ 'v set \ 'v set \ 'v \ bool" where \ \Should we allow only simple paths here?\ "\B. essential G B S x \ (\p. \y\B. path G x p y \ (x \ y \ (\z\set p. z = x \ z \ S)))" abbreviation essential_web :: "('v, 'more) web_scheme \ 'v set \ 'v set" ("\\") where "essential_web \ S \ {x\S. essential \ (B \) S x}" lemma essential_weight_update [simp]: "essential (weight_update f G) = essential G" by(simp add: essential_def fun_eq_iff) lemma not_essentialD: "\B. \ \ essential G B S x; path G x p y; y \ B \ \ x \ y \ (\z\set p. z \ x \ z \ S)" by(simp add: essential_def) lemma essentialE [elim?, consumes 1, case_names essential, cases pred: essential]: "\B. \ essential G B S x; \p y. \ path G x p y; y \ B; \z. \ x \ y; z \ set p \ \ z = x \ z \ S \ \ thesis \ \ thesis" by(auto simp add: essential_def) lemma essentialI [intro?]: "\B. \ path G x p y; y \ B; \z. \ x \ y; z \ set p \ \ z = x \ z \ S \ \ essential G B S x" by(auto simp add: essential_def) lemma essential_vertex: "\B. \ essential G B S x; x \ B \ \vertex G x" by(auto elim!: essentialE simp add: vertex_def elim: rtrancl_path.cases) lemma essential_BI: "\B. x \ B \ essential G B S x" by(auto simp add: essential_def intro: rtrancl_path.base) lemma \_E [elim?, consumes 1, case_names \, cases set: essential_web]: fixes \ (structure) assumes "x \ \ S" obtains p y where "path \ x p y" "y \ B \" "\z. \ x \ y; z \ set p \ \ z = x \ z \ S" using assms by(auto elim: essentialE) lemma essential_mono: "\B. \ essential G B S x; S' \ S \ \ essential G B S' x" by(auto simp add: essential_def) lemma separating_essential: \ \Lem. 3.4 (cf. Lem. 2.14 in [5])\ fixes G A B S assumes "separating_gen G A B S" shows "separating_gen G A B {x\S. essential G B S x}" (is "separating_gen _ _ _ ?E") proof fix x y p assume x: "x \ A" and y: "y \ B" and p: "path G x p y" from separatingD[OF assms p x y] have "\z \ set (x # p). z \ S" by auto from split_list_last_prop[OF this] obtain ys z zs where decomp: "x # p = ys @ z # zs" and z: "z \ S" and last: "\z. z \ set zs \ z \ S" by auto from decomp consider (empty) "ys = []" "x = z" "p = zs" | (Cons) ys' where "ys = x # ys'" "p = ys' @ z # zs" by(auto simp add: Cons_eq_append_conv) then show "(\z\set p. z \ ?E) \ x \ ?E" proof(cases) case empty hence "x \ ?E" using z p last y by(auto simp add: essential_def) thus ?thesis .. next case (Cons ys') from p have "path G z zs y" unfolding Cons by(rule rtrancl_path_appendE) hence "z \ ?E" using z y last by(auto simp add: essential_def) thus ?thesis using Cons by auto qed qed definition roofed_gen :: "('v, 'more) graph_scheme \ 'v set \ 'v set \ 'v set" where roofed_def: "\B. roofed_gen G B S = {x. \p. \y\B. path G x p y \ (\z\set p. z \ S) \ x \ S}" abbreviation roofed :: "('v, 'more) web_scheme \ 'v set \ 'v set" ("RF\") where "roofed \ \ roofed_gen \ (B \)" abbreviation roofed_network :: "('v, 'more) network_scheme \ 'v set \ 'v set" ("RF\<^sup>N\") where "roofed_network \ \ roofed_gen \ {sink \}" lemma roofedI [intro?]: "\B. (\p y. \ path G x p y; y \ B \ \ (\z\set p. z \ S) \ x \ S) \ x \ roofed_gen G B S" by(auto simp add: roofed_def) lemma not_roofedE: fixes B assumes "x \ roofed_gen G B S" obtains p y where "path G x p y" "y \ B" "\z. z \ set (x # p) \ z \ S" using assms by(auto simp add: roofed_def) lemma roofed_greater: "\B. S \ roofed_gen G B S" by(auto simp add: roofed_def) lemma roofed_greaterI: "\B. x \ S \ x \ roofed_gen G B S" using roofed_greater[of S G] by blast lemma roofed_mono: "\B. S \ S' \ roofed_gen G B S \ roofed_gen G B S'" by(fastforce simp add: roofed_def) lemma in_roofed_mono: "\B. \ x \ roofed_gen G B S; S \ S' \ \ x \ roofed_gen G B S'" using roofed_mono[THEN subsetD] . lemma roofedD: "\B. \ x \ roofed_gen G B S; path G x p y; y \ B \ \ (\z\set p. z \ S) \ x \ S" unfolding roofed_def by blast lemma separating_RF_A: fixes A B assumes "separating_gen G A B X" shows "A \ roofed_gen G B X" by(rule subsetI roofedI)+(erule separatingD[OF assms]) lemma roofed_idem: fixes B shows "roofed_gen G B (roofed_gen G B S) = roofed_gen G B S" proof(rule equalityI subsetI roofedI)+ fix x p y assume x: "x \ roofed_gen G B (roofed_gen G B S)" and p: "path G x p y" and y: "y \ B" from roofedD[OF x p y] obtain z where *: "z \ set (x # p)" and z: "z \ roofed_gen G B S" by auto from split_list[OF *] obtain ys zs where split: "x # p = ys @ z # zs" by blast with p have p': "path G z zs y" by(auto simp add: Cons_eq_append_conv elim: rtrancl_path_appendE) from roofedD[OF z p' y] split show "(\z\set p. z \ S) \ x \ S" by(auto simp add: Cons_eq_append_conv) qed(rule roofed_mono roofed_greater)+ lemma in_roofed_mono': "\B. \ x \ roofed_gen G B S; S \ roofed_gen G B S' \ \ x \ roofed_gen G B S'" by(subst roofed_idem[symmetric])(erule in_roofed_mono) lemma roofed_mono': "\B. S \ roofed_gen G B S' \ roofed_gen G B S \ roofed_gen G B S'" by(rule subsetI)(rule in_roofed_mono') lemma roofed_idem_Un1: fixes B shows "roofed_gen G B (roofed_gen G B S \ T) = roofed_gen G B (S \ T)" proof - have "S \ T \ roofed_gen G B S" by (metis (no_types) UnCI roofed_greater subsetCE subsetI) then have "S \ T \ T \ roofed_gen G B S \ T \ roofed_gen G B S \ roofed_gen G B (S \ T)" by (metis (no_types) Un_subset_iff Un_upper2 roofed_greater roofed_mono sup.commute) then show ?thesis by (metis (no_types) roofed_idem roofed_mono subset_antisym sup.commute) qed lemma roofed_UN: fixes A B shows "roofed_gen G B (\i\A. roofed_gen G B (X i)) = roofed_gen G B (\i\A. X i)" (is "?lhs = ?rhs") proof(rule equalityI) show "?rhs \ ?lhs" by(rule roofed_mono)(blast intro: roofed_greaterI) show "?lhs \ ?rhs" by(rule roofed_mono')(blast intro: in_roofed_mono) qed lemma RF_essential: fixes \ (structure) shows "RF (\ S) = RF S" proof(intro set_eqI iffI) fix x assume RF: "x \ RF S" show "x \ RF (\ S)" proof fix p y assume p: "path \ x p y" and y: "y \ B \" from roofedD[OF RF this] have "\z\set (x # p). z \ S" by auto from split_list_last_prop[OF this] obtain ys z zs where decomp: "x # p = ys @ z # zs" and z: "z \ S" and last: "\z. z \ set zs \ z \ S" by auto from decomp consider (empty) "ys = []" "x = z" "p = zs" | (Cons) ys' where "ys = x # ys'" "p = ys' @ z # zs" by(auto simp add: Cons_eq_append_conv) then show "(\z\set p. z \ \ S) \ x \ \ S" proof(cases) case empty hence "x \ \ S" using z p last y by(auto simp add: essential_def) thus ?thesis .. next case (Cons ys') from p have "path \ z zs y" unfolding Cons by(rule rtrancl_path_appendE) hence "z \ \ S" using z y last by(auto simp add: essential_def) thus ?thesis using Cons by auto qed qed qed(blast intro: in_roofed_mono) lemma essentialE_RF: fixes \ (structure) and B assumes "essential \ B S x" obtains p y where "path \ x p y" "y \ B" "distinct (x # p)" "\z. z \ set p \ z \ roofed_gen \ B S" proof - from assms obtain p y where p: "path \ x p y" and y: "y \ B" and bypass: "\z. \ x \ y; z \ set p \ \ z = x \ z \ S" by(rule essentialE) blast from p obtain p' where p': "path \ x p' y" and distinct: "distinct (x # p')" and subset: "set p' \ set p" by(rule rtrancl_path_distinct) { fix z assume z: "z \ set p'" hence "y \ set p'" using rtrancl_path_last[OF p', symmetric] p' by(auto elim: rtrancl_path.cases intro: last_in_set) with distinct z subset have neq: "x \ y" and "z \ set p" by(auto) from bypass[OF this] z distinct have "z \ S" by auto have "z \ roofed_gen \ B S" proof assume z': "z \ roofed_gen \ B S" from split_list[OF z] obtain ys zs where decomp: "p' = ys @ z # zs" by blast with p' have "path \ z zs y" by(auto elim: rtrancl_path_appendE) from roofedD[OF z' this y] \z \ S\ obtain z' where "z' \ set zs" "z' \ S" by auto with bypass[of z'] neq decomp subset distinct show False by auto qed } with p' y distinct show thesis .. qed lemma \_E_RF: fixes \ (structure) assumes "x \ \ S" obtains p y where "path \ x p y" "y \ B \" "distinct (x # p)" "\z. z \ set p \ z \ RF S" using assms by(auto elim: essentialE_RF) lemma in_roofed_essentialD: fixes \ (structure) assumes RF: "x \ RF S" and ess: "essential \ (B \) S x" shows "x \ S" proof - from ess obtain p y where p: "path \ x p y" and y: "y \ B \" and "distinct (x # p)" and bypass: "\z. z \ set p \ z \ S" by(rule essentialE_RF)(auto intro: roofed_greaterI) from roofedD[OF RF p y] bypass show "x \ S" by auto qed lemma separating_RF: fixes \ (structure) shows "separating \ (RF S) \ separating \ S" proof assume sep: "separating \ (RF S)" show "separating \ S" proof fix x y p assume p: "path \ x p y" and x: "x \ A \" and y: "y \ B \" from separatingD[OF sep p x y] have "\z \ set (x # p). z \ RF S" by auto from split_list_last_prop[OF this] obtain ys z zs where split: "x # p = ys @ z # zs" and z: "z \ RF S" and bypass: "\z'. z' \ set zs \ z' \ RF S" by auto from p split have "path \ z zs y" by(cases ys)(auto elim: rtrancl_path_appendE) hence "essential \ (B \) S z" using y by(rule essentialI)(auto dest: bypass intro: roofed_greaterI) with z have "z \ S" by(rule in_roofed_essentialD) with split show "(\z\set p. z \ S) \ x \ S" by(cases ys)auto qed qed(blast intro: roofed_greaterI separating_weakening) definition roofed_circ :: "('v, 'more) web_scheme \ 'v set \ 'v set" ("RF\<^sup>\\") where "roofed_circ \ S = roofed \ S - \\<^bsub>\\<^esub> S" lemma roofed_circI: fixes \ (structure) shows "\ x \ RF T; x \ T \ \ essential \ (B \) T x \ \ x \ RF\<^sup>\ T" by(simp add: roofed_circ_def) lemma roofed_circE: fixes \ (structure) assumes "x \ RF\<^sup>\ T" obtains "x \ RF T" "\ essential \ (B \) T x" using assms by(auto simp add: roofed_circ_def intro: in_roofed_essentialD) lemma \_\: fixes \ (structure) shows "\ (\ S) = \ S" by(auto intro: essential_mono) lemma roofed_circ_essential: fixes \ (structure) shows "RF\<^sup>\ (\ S) = RF\<^sup>\ S" unfolding roofed_circ_def RF_essential \_\ .. lemma essential_RF: fixes B shows "essential G B (roofed_gen G B S) = essential G B S" (is "essential _ _ ?RF = _") proof(intro ext iffI) show "essential G B S x" if "essential G B ?RF x" for x using that by(rule essential_mono)(blast intro: roofed_greaterI) show "essential G B ?RF x" if "essential G B S x" for x using that by(rule essentialE_RF)(erule (1) essentialI, blast) qed lemma \_RF: fixes \ (structure) shows "\ (RF S) = \ S" by(auto dest: in_roofed_essentialD simp add: essential_RF intro: roofed_greaterI) lemma essential_\: fixes \ (structure) shows "essential \ (B \) (\ S) = essential \ (B \) S" by(subst essential_RF[symmetric])(simp only: RF_essential essential_RF) lemma RF_in_B: fixes \ (structure) shows "x \ B \ \ x \ RF S \ x \ S" by(auto intro: roofed_greaterI dest: roofedD[OF _ rtrancl_path.base]) lemma RF_circ_edge_forward: fixes \ (structure) assumes x: "x \ RF\<^sup>\ S" and edge: "edge \ x y" shows "y \ RF S" proof fix p z assume p: "path \ y p z" and z: "z \ B \" from x have rf: "x \ RF S" and ness: "x \ \ S" by(auto elim: roofed_circE) show "(\z\set p. z \ S) \ y \ S" proof(cases "\z'\set (y # p). z' \ S") case False from edge p have p': "path \ x (y # p) z" .. from roofedD[OF rf this z] False have "x \ S" by auto moreover have "essential \ (B \) S x" using p' False z by(auto intro!: essentialI) ultimately have "x \ \ S" by simp with ness show ?thesis by contradiction qed auto qed subsection \Waves\ inductive wave :: "('v, 'more) web_scheme \ 'v current \ bool" for \ (structure) and f where wave: "\ separating \ (TER f); \x. x \ RF (TER f) \ d_OUT f x = 0 \ \ wave \ f" lemma wave_0 [simp]: "wave \ zero_current" by rule simp_all lemma waveD_separating: "wave \ f \ separating \ (TER\<^bsub>\\<^esub> f)" by(simp add: wave.simps) lemma waveD_OUT: "\ wave \ f; x \ RF\<^bsub>\\<^esub> (TER\<^bsub>\\<^esub> f) \ \ d_OUT f x = 0" by(simp add: wave.simps) lemma wave_A_in_RF: fixes \ (structure) shows "\ wave \ f; x \ A \ \ \ x \ RF (TER f)" by(auto intro!: roofedI dest!: waveD_separating separatingD) lemma wave_not_RF_IN_zero: fixes \ (structure) assumes f: "current \ f" and w: "wave \ f" and x: "x \ RF (TER f)" shows "d_IN f x = 0" proof - from x obtain p z where z: "z \ B \" and p: "path \ x p z" and bypass: "\z. z \ set p \ z \ TER f" "x \ TER f" by(clarsimp simp add: roofed_def) have "f (y, x) = 0" for y proof(cases "edge \ y x") case edge: True have "d_OUT f y = 0" proof(cases "y \ TER f") case False with z p bypass edge have "y \ RF (TER f)" by(auto simp add: roofed_def intro: rtrancl_path.step intro!: exI rev_bexI) thus "d_OUT f y = 0" by(rule waveD_OUT[OF w]) qed(auto simp add: SINK.simps) moreover have "f (y, x) \ d_OUT f y" by (rule d_OUT_ge_point) ultimately show ?thesis by simp qed(simp add: currentD_outside[OF f]) then show "d_IN f x = 0" unfolding d_IN_def by(simp add: nn_integral_0_iff emeasure_count_space_eq_0) qed lemma current_Sup: fixes \ (structure) assumes chain: "Complete_Partial_Order.chain (\) Y" and Y: "Y \ {}" and current: "\f. f \ Y \ current \ f" and countable [simp]: "countable (support_flow (Sup Y))" shows "current \ (Sup Y)" proof(rule, goal_cases) case (1 x) have "d_OUT (Sup Y) x = (SUP f\Y. d_OUT f x)" using chain Y by(simp add: d_OUT_Sup) also have "\ \ weight \ x" using 1 by(intro SUP_least)(auto dest!: current currentD_weight_OUT) finally show ?case . next case (2 x) have "d_IN (Sup Y) x = (SUP f\Y. d_IN f x)" using chain Y by(simp add: d_IN_Sup) also have "\ \ weight \ x" using 2 by(intro SUP_least)(auto dest!: current currentD_weight_IN) finally show ?case . next case (3 x) have "d_OUT (Sup Y) x = (SUP f\Y. d_OUT f x)" using chain Y by(simp add: d_OUT_Sup) also have "\ \ (SUP f\Y. d_IN f x)" using 3 by(intro SUP_mono)(auto dest: current currentD_OUT_IN) also have "\ = d_IN (Sup Y) x" using chain Y by(simp add: d_IN_Sup) finally show ?case . next case (4 a) have "d_IN (Sup Y) a = (SUP f\Y. d_IN f a)" using chain Y by(simp add: d_IN_Sup) also have "\ = (SUP f\Y. 0)" using 4 by(intro SUP_cong)(auto dest!: current currentD_IN) also have "\ = 0" using Y by simp finally show ?case . next case (5 b) have "d_OUT (Sup Y) b = (SUP f\Y. d_OUT f b)" using chain Y by(simp add: d_OUT_Sup) also have "\ = (SUP f\Y. 0)" using 5 by(intro SUP_cong)(auto dest!: current currentD_OUT) also have "\ = 0" using Y by simp finally show ?case . next fix e assume "e \ \<^bold>E" from currentD_outside'[OF current this] have "f e = 0" if "f \ Y" for f using that by simp hence "Sup Y e = (SUP _\Y. 0)" by(auto intro: SUP_cong) then show "Sup Y e = 0" using Y by(simp) qed lemma wave_lub: \ \Lemma 4.3\ fixes \ (structure) assumes chain: "Complete_Partial_Order.chain (\) Y" and Y: "Y \ {}" and wave: "\f. f \ Y \ wave \ f" and countable [simp]: "countable (support_flow (Sup Y))" shows "wave \ (Sup Y)" proof { fix x y p assume p: "path \ x p y" and y: "y \ B \" define P where "P = {x} \ set p" let ?f = "\f. SINK f \ P" have "Complete_Partial_Order.chain (\) (?f ` Y)" using chain by(rule chain_imageI)(auto dest: SINK_mono') moreover have "\ \ Pow P" by auto hence "finite (?f ` Y)" by(rule finite_subset)(simp add: P_def) ultimately have "(\(?f ` Y)) \ ?f ` Y" by(rule ccpo.in_chain_finite[OF complete_lattice_ccpo_dual])(simp add: Y) then obtain f where f: "f \ Y" and eq: "\(?f ` Y) = ?f f" by clarify hence *: "(\f\Y. SINK f) \ P = SINK f \ P" by(clarsimp simp add: prod_lub_def Y)+ { fix g assume "g \ Y" "f \ g" with * have "(\f\Y. SINK f) \ P = SINK g \ P" by(blast dest: SINK_mono') then have "TER (Sup Y) \ P \ TER g \ P" using SAT_Sup_upper[OF \g \ Y\, of \] SINK_Sup[OF chain Y countable] by blast } with f have "\f\Y. \g\Y. g \ f \ TER g \ P \ TER (Sup Y) \ P" by blast } note subset = this show "separating \ (TER (Sup Y))" proof fix x y p assume *: "path \ x p y" "y \ B \" and "x \ A \" let ?P = "{x} \ set p" from subset[OF *] obtain f where f:"f \ Y" and subset: "TER f \ ?P \ TER (Sup Y) \ ?P" by blast from wave[OF f] have "TER f \ ?P \ {}" using * \x \ A \\ by(auto simp add: wave.simps dest: separatingD) with subset show " (\z\set p. z \ TER (Sup Y)) \ x \ TER (Sup Y)" by blast qed fix x assume "x \ RF (TER (Sup Y))" then obtain p y where y: "y \ B \" and p: "path \ x p y" and ter: "TER (Sup Y) \ ({x} \ set p) = {}" by(auto simp add: roofed_def) let ?P = "{x} \ set p" from subset[OF p y] obtain f where f: "f \ Y" and subset: "\g. \ g \ Y; f \ g \ \ TER g \ ?P \ TER (Sup Y) \ ?P" by blast { fix g assume g: "g \ Y" with chain f have "f \ g \ g \ f" by(rule chainD) hence "d_OUT g x = 0" proof assume "f \ g" from subset[OF g this] ter have "TER g \ ?P = {}" by blast with p y have "x \ RF (TER g)" by(auto simp add: roofed_def) with wave[OF g] show ?thesis by(blast elim: wave.cases) next assume "g \ f" from subset ter f have "TER f \ ?P = {}" by blast with y p have "x \ RF (TER f)" by(auto simp add: roofed_def) with wave[OF f] have "d_OUT f x = 0" by(blast elim: wave.cases) moreover have "d_OUT g x \ d_OUT f x" using \g \ f\[THEN le_funD] by(rule d_OUT_mono) ultimately show ?thesis by simp qed } thus "d_OUT (Sup Y) x = 0" using chain Y by(simp add: d_OUT_Sup) qed lemma ex_maximal_wave: \ \Corollary 4.4\ fixes \ (structure) assumes countable: "countable \<^bold>E" shows "\f. current \ f \ wave \ f \ (\w. current \ w \ wave \ w \ f \ w \ f = w)" proof - define Field_r where "Field_r = {f. current \ f \ wave \ f}" define r where "r = {(f, g). f \ Field_r \ g \ Field_r \ f \ g}" have Field_r: "Field r = Field_r" by(auto simp add: Field_def r_def) have "Partial_order r" unfolding order_on_defs by(auto intro!: refl_onI transI antisymI simp add: Field_r r_def Field_def) hence "\m\Field r. \a\Field r. (m, a) \ r \ a = m" proof(rule Zorns_po_lemma) fix Y assume "Y \ Chains r" hence Y: "Complete_Partial_Order.chain (\) Y" and w: "\f. f \ Y \ wave \ f" and f: "\f. f \ Y \ current \ f" by(auto simp add: Chains_def r_def chain_def Field_r_def) show "\w \ Field r. \f \ Y. (f, w) \ r" proof(cases "Y = {}") case True have "zero_current \ Field r" by(simp add: Field_r Field_r_def) with True show ?thesis by blast next case False have "support_flow (Sup Y) \ \<^bold>E" by(auto simp add: support_flow_Sup elim!: support_flow.cases dest!: f dest: currentD_outside) hence c: "countable (support_flow (Sup Y))" using countable by(rule countable_subset) with Y False f w have "Sup Y \ Field r" unfolding Field_r Field_r_def by(blast intro: wave_lub current_Sup) moreover then have "(f, Sup Y) \ r" if "f \ Y" for f using w[OF that] f[OF that] that unfolding Field_r by(auto simp add: r_def Field_r_def intro: Sup_upper) ultimately show ?thesis by blast qed qed thus ?thesis by(simp add: Field_r Field_r_def)(auto simp add: r_def Field_r_def) qed lemma essential_leI: fixes \ (structure) assumes g: "current \ g" and w: "wave \ g" and le: "\e. f e \ g e" and x: "x \ \ (TER g)" shows "essential \ (B \) (TER f) x" proof - from x obtain p y where p: "path \ x p y" and y: "y \ B \" and distinct: "distinct (x # p)" and bypass: "\z. z \ set p \ z \ RF (TER g)" by(rule \_E_RF) blast show ?thesis using p y proof fix z assume "z \ set p" hence z: "z \ RF (TER g)" by(auto dest: bypass) with w have OUT: "d_OUT g z = 0" and IN: "d_IN g z = 0" by(rule waveD_OUT wave_not_RF_IN_zero[OF g])+ with z have "z \ A \" "weight \ z > 0" by(auto intro!: roofed_greaterI simp add: SAT.simps SINK.simps) moreover from IN d_IN_mono[of f z g, OF le] have "d_IN f z \ 0" by(simp) ultimately have "z \ TER f" by(auto simp add: SAT.simps) then show "z = x \ z \ TER f" by simp qed qed lemma essential_eq_leI: fixes \ (structure) assumes g: "current \ g" and w: "wave \ g" and le: "\e. f e \ g e" and subset: "\ (TER g) \ TER f" shows "\ (TER f) = \ (TER g)" proof show subset: "\ (TER g) \ \ (TER f)" proof fix x assume x: "x \ \ (TER g)" hence "x \ TER f" using subset by blast moreover have "essential \ (B \) (TER f) x" using g w le x by(rule essential_leI) ultimately show "x \ \ (TER f)" by simp qed show "\ \ \ (TER g)" proof fix x assume x: "x \ \ (TER f)" hence "x \ TER f" by auto hence "x \ RF (TER g)" proof(rule contrapos_pp) assume x: "x \ RF (TER g)" with w have OUT: "d_OUT g x = 0" and IN: "d_IN g x = 0" by(rule waveD_OUT wave_not_RF_IN_zero[OF g])+ with x have "x \ A \" "weight \ x > 0" by(auto intro!: roofed_greaterI simp add: SAT.simps SINK.simps) moreover from IN d_IN_mono[of f x g, OF le] have "d_IN f x \ 0" by(simp) ultimately show "x \ TER f" by(auto simp add: SAT.simps) qed moreover have "x \ RF\<^sup>\ (TER g)" proof assume "x \ RF\<^sup>\ (TER g)" hence RF: "x \ RF (\ (TER g))" and not_E: "x \ \ (TER g)" unfolding RF_essential by(simp_all add: roofed_circ_def) from x obtain p y where p: "path \ x p y" and y: "y \ B \" and distinct: "distinct (x # p)" and bypass: "\z. z \ set p \ z \ RF (TER f)" by(rule \_E_RF) blast from roofedD[OF RF p y] not_E obtain z where "z \ set p" "z \ \ (TER g)" by blast with subset bypass[of z] show False by(auto intro: roofed_greaterI) qed ultimately show "x \ \ (TER g)" by(simp add: roofed_circ_def) qed qed subsection \Hindrances and looseness\ inductive hindrance_by :: "('v, 'more) web_scheme \ 'v current \ ennreal \ bool" for \ (structure) and f and \ where hindrance_by: "\ a \ A \; a \ \ (TER f); d_OUT f a < weight \ a; \ < weight \ a - d_OUT f a \ \ hindrance_by \ f \" inductive hindrance :: "('v, 'more) web_scheme \ 'v current \ bool" for \ (structure) and f where hindrance: "\ a \ A \; a \ \ (TER f); d_OUT f a < weight \ a \ \ hindrance \ f" inductive hindered :: "('v, 'more) web_scheme \ bool" for \ (structure) where hindered: "\ hindrance \ f; current \ f; wave \ f \ \ hindered \" inductive hindered_by :: "('v, 'more) web_scheme \ ennreal \ bool" for \ (structure) and \ where hindered_by: "\ hindrance_by \ f \; current \ f; wave \ f \ \ hindered_by \ \" lemma hindrance_into_hindrance_by: assumes "hindrance \ f" shows "\\>0. hindrance_by \ f \" using assms proof cases case (hindrance a) let ?\ = "if weight \ a = \ then 1 else (weight \ a - d_OUT f a) / 2" from \d_OUT f a < weight \ a\ have "weight \ a - d_OUT f a > 0" "weight \ a \ \ \ weight \ a - d_OUT f a < \" by(simp_all add: diff_gr0_ennreal less_top diff_less_top_ennreal) from ennreal_mult_strict_left_mono[of 1 2, OF _ this] have "0 < ?\" and "?\ < weight \ a - d_OUT f a" using \d_OUT f a < weight \ a\ by(auto intro!: diff_gr0_ennreal simp: ennreal_zero_less_divide divide_less_ennreal) with hindrance show ?thesis by(auto intro!: hindrance_by.intros) qed lemma hindrance_by_into_hindrance: "hindrance_by \ f \ \ hindrance \ f" by(blast elim: hindrance_by.cases intro: hindrance.intros) lemma hindrance_conv_hindrance_by: "hindrance \ f \ (\\>0. hindrance_by \ f \)" by(blast intro: hindrance_into_hindrance_by hindrance_by_into_hindrance) lemma hindered_into_hindered_by: "hindered \ \ \\>0. hindered_by \ \" by(blast intro: hindered_by.intros elim: hindered.cases dest: hindrance_into_hindrance_by) lemma hindered_by_into_hindered: "hindered_by \ \ \ hindered \" by(blast elim: hindered_by.cases intro: hindered.intros dest: hindrance_by_into_hindrance) lemma hindered_conv_hindered_by: "hindered \ \ (\\>0. hindered_by \ \)" by(blast intro: hindered_into_hindered_by hindered_by_into_hindered) inductive loose :: "('v, 'more) web_scheme \ bool" for \ where loose: "\ \f. \ current \ f; wave \ f \ \ f = zero_current; \ hindrance \ zero_current \ \ loose \" lemma looseD_hindrance: "loose \ \ \ hindrance \ zero_current" by(simp add: loose.simps) lemma looseD_wave: "\ loose \; current \ f; wave \ f \ \ f = zero_current" by(simp add: loose.simps) lemma loose_unhindered: fixes \ (structure) assumes "loose \" shows "\ hindered \" apply auto apply(erule hindered.cases) apply(frule (1) looseD_wave[OF assms]) apply simp using looseD_hindrance[OF assms] by simp context fixes \ \' :: "('v, 'more) web_scheme" assumes [simp]: "edge \ = edge \'" "A \ = A \'" "B \ = B \'" and weight_eq: "\x. x \ A \' \ weight \ x = weight \' x" and weight_le: "\a. a \ A \' \ weight \ a \ weight \' a" begin private lemma essential_eq: "essential \ = essential \'" by(simp add: fun_eq_iff essential_def) qualified lemma TER_eq: "TER\<^bsub>\\<^esub> f = TER\<^bsub>\'\<^esub> f" apply(auto simp add: SINK.simps SAT.simps) apply(erule contrapos_np; drule weight_eq; simp)+ done qualified lemma separating_eq: "separating_gen \ = separating_gen \'" by(intro ext iffI; rule separating_gen.intros; drule separatingD; simp) qualified lemma roofed_eq: "\B. roofed_gen \ B S = roofed_gen \' B S" by(simp add: roofed_def) lemma wave_eq_web: \ \Observation 4.6\ "wave \ f \ wave \' f" by(simp add: wave.simps separating_eq TER_eq roofed_eq) lemma current_mono_web: "current \' f \ current \ f" apply(rule current, simp_all add: currentD_OUT_IN currentD_IN currentD_OUT currentD_outside') subgoal for x by(cases "x \ A \'")(auto dest!: weight_eq weight_le dest: currentD_weight_OUT intro: order_trans) subgoal for x by(cases "x \ A \'")(auto dest!: weight_eq weight_le dest: currentD_weight_IN intro: order_trans) done lemma hindrance_mono_web: "hindrance \' f \ hindrance \ f" apply(erule hindrance.cases) apply(rule hindrance) apply simp apply(unfold TER_eq, simp add: essential_eq) apply(auto dest!: weight_le) done lemma hindered_mono_web: "hindered \' \ hindered \" apply(erule hindered.cases) apply(rule hindered.intros) apply(erule hindrance_mono_web) apply(erule current_mono_web) apply(simp add: wave_eq_web) done end subsection \Linkage\ text \ The following definition of orthogonality is stronger than the original definition 3.5 in @{cite AharoniBergerGeorgakopoulusPerlsteinSpruessel2011JCT} in that the outflow from any \A\-vertices in the set must saturate the vertex; @{term "S \ SAT \ f"} is not enough. With the original definition of orthogonal current, the reduction from networks to webs fails because the induced flow need not saturate edges going out of the source. Consider the network with three nodes \s\, \x\, and \t\ and edges \(s, x)\ and \(x, t)\ with capacity \1\. Then, the corresponding web has the vertices \(s, x)\ and \(x, t)\ and one edge from \(s, x)\ to \(x, t)\. Clearly, the zero current @{term [source] zero_current} is a web-flow and \TER zero_current = {(s, x)}\, which is essential. Moreover, @{term [source] zero_current} and \{(s, x)}\ are orthogonal because @{term [source] zero_current} trivially saturates \(s, x)\ as this is a vertex in \A\. \ inductive orthogonal_current :: "('v, 'more) web_scheme \ 'v current \ 'v set \ bool" for \ (structure) and f S where orthogonal_current: "\ \x. \ x \ S; x \ A \ \ \ weight \ x \ d_IN f x; \x. \ x \ S; x \ A \; x \ B \ \ \ d_OUT f x = weight \ x; \u v. \ v \ RF S; u \ RF\<^sup>\ S \ \ f (u, v) = 0 \ \ orthogonal_current \ f S" lemma orthogonal_currentD_SAT: "\ orthogonal_current \ f S; x \ S \ \ x \ SAT \ f" by(auto elim!: orthogonal_current.cases intro: SAT.intros) lemma orthogonal_currentD_A: "\ orthogonal_current \ f S; x \ S; x \ A \; x \ B \ \ \ d_OUT f x = weight \ x" by(auto elim: orthogonal_current.cases) lemma orthogonal_currentD_in: "\ orthogonal_current \ f S; v \ RF\<^bsub>\\<^esub> S; u \ RF\<^sup>\\<^bsub>\\<^esub> S \ \ f (u, v) = 0" by(auto elim: orthogonal_current.cases) inductive linkage :: "('v, 'more) web_scheme \ 'v current \ bool" for \ f where \ \Omit the condition @{const web_flow}\ linkage: "(\x. x \ A \ \ d_OUT f x = weight \ x) \ linkage \ f" lemma linkageD: "\ linkage \ f; x \ A \ \ \ d_OUT f x = weight \ x" by(rule linkage.cases) abbreviation linkable :: "('v, 'more) web_scheme \ bool" where "linkable \ \ \f. web_flow \ f \ linkage \ f" subsection \Trimming\ context fixes \ :: "('v, 'more) web_scheme" (structure) and f :: "'v current" begin inductive trimming :: "'v current \ bool" for g where trimming: \ \omits the condition that @{term f} is a wave\ "\ current \ g; wave \ g; g \ f; \x. \ x \ RF\<^sup>\ (TER f); x \ A \ \ \ KIR g x; \ (TER g) - A \ = \ (TER f) - A \ \ \ trimming g" lemma assumes "trimming g" shows trimmingD_current: "current \ g" and trimmingD_wave: "wave \ g" and trimmingD_le: "\e. g e \ f e" and trimmingD_KIR: "\ x \ RF\<^sup>\ (TER f); x \ A \ \ \ KIR g x" and trimmingD_\: "\ (TER g) - A \ = \ (TER f) - A \" using assms by(blast elim: trimming.cases dest: le_funD)+ lemma ex_trimming: \ \Lemma 4.8\ assumes f: "current \ f" and w: "wave \ f" and countable: "countable \<^bold>E" and weight_finite: "\x. weight \ x \ \" shows "\g. trimming g" proof - define F where "F = {g. current \ g \ wave \ g \ g \ f \ \ (TER g) = \ (TER f)}" define leq where "leq = restrict_rel F {(g, g'). g' \ g}" have in_F [simp]: "g \ F \ current \ g \ wave \ g \ (\e. g e \ f e) \ \ (TER g) = \ (TER f)" for g by(simp add: F_def le_fun_def) have f_finite [simp]: "f e \ \" for e proof(cases e) case (Pair x y) have "f (x, y) \ d_IN f y" by (rule d_IN_ge_point) also have "\ \ weight \ y" by(rule currentD_weight_IN[OF f]) also have "\ < \" by(simp add: weight_finite less_top[symmetric]) finally show ?thesis using Pair by simp qed have chainD: "Inf M \ F" if M: "M \ Chains leq" and nempty: "M \ {}" for M proof - from nempty obtain g0 where g0: "g0 \ M" by auto have g0_le_f: "g0 e \ f e" and g: "current \ g0" and w0: "wave \ g0" for e using Chains_FieldD[OF M g0] by(cases e, auto simp add: leq_def) have finite_OUT: "d_OUT f x \ \" for x using weight_finite[of x] by(rule neq_top_trans)(rule currentD_weight_OUT[OF f]) have finite_IN: "d_IN f x \ \" for x using weight_finite[of x] by(rule neq_top_trans)(rule currentD_weight_IN[OF f]) from M have "M \ Chains {(g, g'). g' \ g}" by(rule mono_Chains[THEN subsetD, rotated])(auto simp add: leq_def in_restrict_rel_iff) then have chain: "Complete_Partial_Order.chain (\) M" by(rule Chains_into_chain) hence chain': "Complete_Partial_Order.chain (\) M" by(simp add: chain_dual) have countable': "countable (support_flow f)" using current_support_flow[OF f] by(rule countable_subset)(rule countable) have OUT_M: "d_OUT (Inf M) x = (INF g\M. d_OUT g x)" for x using chain' nempty countable' _ finite_OUT by(rule d_OUT_Inf)(auto dest!: Chains_FieldD[OF M] simp add: leq_def) have IN_M: "d_IN (Inf M) x = (INF g\M. d_IN g x)" for x using chain' nempty countable' _ finite_IN by(rule d_IN_Inf)(auto dest!: Chains_FieldD[OF M] simp add: leq_def) have c: "current \ (Inf M)" using g proof(rule current_leI) show "(Inf M) e \ g0 e" for e using g0 by(auto intro: INF_lower) show "d_OUT (\M) x \ d_IN (\M) x" if "x \ A \" for x by(auto 4 4 simp add: IN_M OUT_M leq_def intro!: INF_mono dest: Chains_FieldD[OF M] intro: currentD_OUT_IN[OF _ that]) qed have INF_le_f: "Inf M e \ f e" for e using g0 by(auto intro!: INF_lower2 g0_le_f) have eq: "\ (TER (Inf M)) = \ (TER f)" using f w INF_le_f proof(rule essential_eq_leI; intro subsetI) fix x assume x: "x \ \ (TER f)" hence "x \ SINK (Inf M)" using d_OUT_mono[of "Inf M" x f, OF INF_le_f] by(auto simp add: SINK.simps) moreover from x have "x \ SAT \ g" if "g \ M" for g using Chains_FieldD[OF M that] by(auto simp add: leq_def) hence "x \ SAT \ (Inf M)" by(auto simp add: SAT.simps IN_M intro!: INF_greatest) ultimately show "x \ TER (Inf M)" by auto qed have w': "wave \ (Inf M)" proof have "separating \ (\ (TER f))" by(rule separating_essential)(rule waveD_separating[OF w]) then show "separating \ (TER (Inf M))" unfolding eq[symmetric] by(rule separating_weakening) auto fix x assume "x \ RF (TER (Inf M))" hence "x \ RF (\ (TER (Inf M)))" unfolding RF_essential . hence "x \ RF (TER f)" unfolding eq RF_essential . hence "d_OUT f x = 0" by(rule waveD_OUT[OF w]) with d_OUT_mono[of _ x f, OF INF_le_f] show "d_OUT (Inf M) x = 0" by (metis le_zero_eq) qed from c w' INF_le_f eq show ?thesis by simp qed define trim1 where "trim1 g = (if trimming g then g else let z = SOME z. z \ RF\<^sup>\ (TER g) \ z \ A \ \ \ KIR g z; factor = d_OUT g z / d_IN g z in (\(y, x). (if x = z then factor else 1) * g (y, x)))" for g have increasing: "trim1 g \ g \ trim1 g \ F" if "g \ F" for g proof(cases "trimming g") case True thus ?thesis using that by(simp add: trim1_def) next case False let ?P = "\z. z \ RF\<^sup>\ (TER g) \ z \ A \ \ \ KIR g z" define z where "z = Eps ?P" from that have g: "current \ g" and w': "wave \ g" and le_f: "\e. g e \ f e" and \: "\ (TER g) = \ (TER f)" by(auto simp add: le_fun_def) { with False obtain z where z: "z \ RF\<^sup>\ (TER f)" and A: "z \ A \" and neq: "d_OUT g z \ d_IN g z" by(auto simp add: trimming.simps le_fun_def) from z have "z \ RF\<^sup>\ (\ (TER f))" unfolding roofed_circ_essential . with \ roofed_circ_essential[of \ "TER g"] have "z \ RF\<^sup>\ (TER g)" by simp with A neq have "\x. ?P x" by auto } hence "?P z" unfolding z_def by(rule someI_ex) hence RF: "z \ RF\<^sup>\ (TER g)" and A: "z \ A \" and neq: "d_OUT g z \ d_IN g z" by simp_all let ?factor = "d_OUT g z / d_IN g z" have trim1 [simp]: "trim1 g (y, x) = (if x = z then ?factor else 1) * g (y, x)" for x y using False by(auto simp add: trim1_def z_def Let_def) from currentD_OUT_IN[OF g A] neq have less: "d_OUT g z < d_IN g z" by auto hence "?factor \ 1" (is "?factor \ _") by (auto intro!: divide_le_posI_ennreal simp: zero_less_iff_neq_zero) hence le': "?factor * g (y, x) \ 1 * g (y, x)" for y x by(rule mult_right_mono) simp hence le: "trim1 g e \ g e" for e by(cases e)simp moreover { have c: "current \ (trim1 g)" using g le proof(rule current_leI) fix x assume x: "x \ A \" have "d_OUT (trim1 g) x \ d_OUT g x" unfolding d_OUT_def using le' by(auto intro: nn_integral_mono) also have "\ \ d_IN (trim1 g) x" proof(cases "x = z") case True have "d_OUT g x = d_IN (trim1 g) x" unfolding d_IN_def using True currentD_weight_IN[OF g, of x] currentD_OUT_IN[OF g x] apply (cases "d_IN g x = 0") apply(auto simp add: nn_integral_divide nn_integral_cmult d_IN_def[symmetric] ennreal_divide_times) apply (subst ennreal_divide_self) apply (auto simp: less_top[symmetric] top_unique weight_finite) done thus ?thesis by simp next case False have "d_OUT g x \ d_IN g x" using x by(rule currentD_OUT_IN[OF g]) also have "\ \ d_IN (trim1 g) x" unfolding d_IN_def using False by(auto intro!: nn_integral_mono) finally show ?thesis . qed finally show "d_OUT (trim1 g) x \ d_IN (trim1 g) x" . qed moreover have le_f: "trim1 g \ f" using le le_f by(blast intro: le_funI order_trans) moreover have eq: "\ (TER (trim1 g)) = \ (TER f)" unfolding \[symmetric] using g w' le proof(rule essential_eq_leI; intro subsetI) fix x assume x: "x \ \ (TER g)" hence "x \ SINK (trim1 g)" using d_OUT_mono[of "trim1 g" x g, OF le] by(auto simp add: SINK.simps) moreover from x have "x \ z" using RF by(auto simp add: roofed_circ_def) hence "d_IN (trim1 g) x = d_IN g x" unfolding d_IN_def by simp with \x \ \ (TER g)\ have "x \ SAT \ (trim1 g)" by(auto simp add: SAT.simps) ultimately show "x \ TER (trim1 g)" by auto qed moreover have "wave \ (trim1 g)" proof have "separating \ (\ (TER f))" by(rule separating_essential)(rule waveD_separating[OF w]) then show "separating \ (TER (trim1 g))" unfolding eq[symmetric] by(rule separating_weakening) auto fix x assume "x \ RF (TER (trim1 g))" hence "x \ RF (\ (TER (trim1 g)))" unfolding RF_essential . hence "x \ RF (TER f)" unfolding eq RF_essential . hence "d_OUT f x = 0" by(rule waveD_OUT[OF w]) with d_OUT_mono[of _ x f, OF le_f[THEN le_funD]] show "d_OUT (trim1 g) x = 0" by (metis le_zero_eq) qed ultimately have "trim1 g \ F" by(simp add: F_def) } ultimately show ?thesis using that by(simp add: le_fun_def del: trim1) qed have "bourbaki_witt_fixpoint Inf leq trim1" using chainD increasing unfolding leq_def by(intro bourbaki_witt_fixpoint_restrict_rel)(auto intro: Inf_greatest Inf_lower) then interpret bourbaki_witt_fixpoint Inf leq trim1 . have f_Field: "f \ Field leq" using f w by(simp add: leq_def) define g where "g = fixp_above f" have "g \ Field leq" using f_Field unfolding g_def by(rule fixp_above_Field) hence le_f: "g \ f" and g: "current \ g" and w': "wave \ g" and TER: "\ (TER g) = \ (TER f)" by(auto simp add: leq_def intro: le_funI) have "trimming g" proof(rule ccontr) let ?P = "\x. x \ RF\<^sup>\ (TER g) \ x \ A \ \ \ KIR g x" define x where "x = Eps ?P" assume False: "\ ?thesis" hence "\x. ?P x" using le_f g w' TER by(auto simp add: trimming.simps roofed_circ_essential[of \ "TER g", symmetric] roofed_circ_essential[of \ "TER f", symmetric]) hence "?P x" unfolding x_def by(rule someI_ex) hence x: "x \ RF\<^sup>\ (TER g)" and A: "x \ A \" and neq: "d_OUT g x \ d_IN g x" by simp_all from neq have "\y. edge \ y x \ g (y, x) > 0" proof(rule contrapos_np) assume "\ ?thesis" hence "d_IN g x = 0" using currentD_outside[OF g, of _ x] by(force simp add: d_IN_def nn_integral_0_iff_AE AE_count_space not_less) with currentD_OUT_IN[OF g A] show "KIR g x" by simp qed then obtain y where y: "edge \ y x" and gr0: "g (y, x) > 0" by blast have [simp]: "g (y, x) \ \" proof - have "g (y, x) \ d_OUT g y" by (rule d_OUT_ge_point) also have "\ \ weight \ y" by(rule currentD_weight_OUT[OF g]) also have "\ < \" by(simp add: weight_finite less_top[symmetric]) finally show ?thesis by simp qed from neq have factor: "d_OUT g x / d_IN g x \ 1" by (simp add: divide_eq_1_ennreal) have "trim1 g (y, x) = g (y, x) * (d_OUT g x / d_IN g x)" by(clarsimp simp add: False trim1_def Let_def x_def[symmetric] mult.commute) moreover have "\ \ g (y, x) * 1" unfolding ennreal_mult_cancel_left using gr0 factor by auto ultimately have "trim1 g (y, x) \ g (y, x)" by auto hence "trim1 g \ g" by(auto simp add: fun_eq_iff) moreover have "trim1 g = g" using f_Field unfolding g_def by(rule fixp_above_unfold[symmetric]) ultimately show False by contradiction qed then show ?thesis by blast qed end lemma trimming_\: fixes \ (structure) assumes w: "wave \ f" and trimming: "trimming \ f g" shows "\ (TER f) = \ (TER g)" proof(rule set_eqI) show "x \ \ (TER f) \ x \ \ (TER g)" for x proof(cases "x \ A \") case False thus ?thesis using trimmingD_\[OF trimming] by blast next case True show ?thesis proof assume x: "x \ \ (TER f)" hence "x \ TER g" using d_OUT_mono[of g x f, OF trimmingD_le[OF trimming]] True by(simp add: SINK.simps SAT.A) moreover from x have "essential \ (B \) (TER f) x" by simp then obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER f)" by(rule essentialE_RF) blast from p y have "essential \ (B \) (\ (TER g)) x" proof(rule essentialI) fix z assume "z \ set p" hence z: "z \ RF (TER f)" by(rule bypass) with waveD_separating[OF w, THEN separating_RF_A] have "z \ A \" by blast with z have "z \ \ (TER g)" using trimmingD_\[OF trimming] by(auto intro: roofed_greaterI) thus "z = x \ z \ \ (TER g)" .. qed ultimately show "x \ \ (TER g)" unfolding essential_\ by simp next assume "x \ \ (TER g)" then obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER g)" by(rule \_E_RF) blast have z: "z \ \ (TER f)" if "z \ set p" for z proof - from that have z: "z \ RF (TER g)" by(rule bypass) with waveD_separating[OF trimmingD_wave[OF trimming], THEN separating_RF_A] have "z \ A \" by blast with z show "z \ \ (TER f)" using trimmingD_\[OF trimming] by(auto intro: roofed_greaterI) qed then have "essential \ (B \) (\ (TER f)) x" by(intro essentialI[OF p y]) auto moreover have "x \ TER f" using waveD_separating[THEN separating_essential, THEN separatingD, OF w p True y] z by auto ultimately show "x \ \ (TER f)" unfolding essential_\ by simp qed qed qed subsection \Composition of waves via quotients\ definition quotient_web :: "('v, 'more) web_scheme \ 'v current \ ('v, 'more) web_scheme" where \ \Modifications to original Definition 4.9: No incoming edges to nodes in @{const A}, @{term "B \ - A \"} is not part of @{const A} such that @{const A} contains only vertices is disjoint from @{const B}. The weight of vertices in @{const B} saturated by @{term f} is therefore set to @{term "0 :: ennreal"}.\ "quotient_web \ f = \edge = \x y. edge \ x y \ x \ roofed_circ \ (TER\<^bsub>\\<^esub> f) \ y \ roofed \ (TER\<^bsub>\\<^esub> f), weight = \x. if x \ RF\<^sup>\\<^bsub>\\<^esub> (TER\<^bsub>\\<^esub> f) \ x \ TER\<^bsub>\\<^esub> f \ B \ then 0 else weight \ x, A = \\<^bsub>\\<^esub> (TER\<^bsub>\\<^esub> f) - (B \ - A \), B = B \, \ = web.more \\" lemma quotient_web_sel [simp]: fixes \ (structure) shows "edge (quotient_web \ f) x y \ edge \ x y \ x \ RF\<^sup>\ (TER f) \ y \ RF (TER f)" "weight (quotient_web \ f) x = (if x \ RF\<^sup>\ (TER f) \ x \ TER\<^bsub>\\<^esub> f \ B \ then 0 else weight \ x)" "A (quotient_web \ f) = \ (TER f)- (B \ - A \)" "B (quotient_web \ f) = B \" "web.more (quotient_web \ f) = web.more \" by(simp_all add: quotient_web_def) lemma vertex_quotient_webD: fixes \ (structure) shows "vertex (quotient_web \ f) x \ vertex \ x \ x \ RF\<^sup>\ (TER f)" by(auto simp add: vertex_def roofed_circ_def) lemma path_quotient_web: fixes \ (structure) assumes "path \ x p y" and "x \ RF\<^sup>\ (TER f)" and "\z. z \ set p \ z \ RF (TER f)" shows "path (quotient_web \ f) x p y" using assms by(induction)(auto intro: rtrancl_path.intros simp add: roofed_circ_def) definition restrict_current :: "('v, 'more) web_scheme \ 'v current \ 'v current \ 'v current" where "restrict_current \ f g = (\(x, y). g (x, y) * indicator (- RF\<^sup>\\<^bsub>\\<^esub> (TER\<^bsub>\\<^esub> f)) x * indicator (- RF\<^bsub>\\<^esub> (TER\<^bsub>\\<^esub> f)) y)" abbreviation restrict_curr :: "'v current \ ('v, 'more) web_scheme \ 'v current \ 'v current" ("_ \ _ '/ _" [100, 0, 100] 100) where "restrict_curr g \ f \ restrict_current \ f g" lemma restrict_current_simps [simp]: fixes \ (structure) shows "(g \ \ / f) (x, y) = (g (x, y) * indicator (- RF\<^sup>\ (TER f)) x * indicator (- RF (TER f)) y)" by(simp add: restrict_current_def) lemma d_OUT_restrict_current_outside: fixes \ (structure) shows "x \ RF\<^sup>\ (TER f) \ d_OUT (g \ \ / f) x = 0" by(simp add: d_OUT_def) lemma d_IN_restrict_current_outside: fixes \ (structure) shows "x \ RF (TER f) \ d_IN (g \ \ / f) x = 0" by(simp add: d_IN_def) lemma restrict_current_le: "(g \ \ / f) e \ g e" by(cases e)(clarsimp split: split_indicator) lemma d_OUT_restrict_current_le: "d_OUT (g \ \ / f) x \ d_OUT g x" unfolding d_OUT_def by(rule nn_integral_mono, simp split: split_indicator) lemma d_IN_restrict_current_le: "d_IN (g \ \ / f) x \ d_IN g x" unfolding d_IN_def by(rule nn_integral_mono, simp split: split_indicator) lemma restrict_current_IN_not_RF: fixes \ (structure) assumes g: "current \ g" and x: "x \ RF (TER f)" shows "d_IN (g \ \ / f) x = d_IN g x" proof - { fix y assume y: "y \ RF\<^sup>\ (TER f)" have "g (y, x) = 0" proof(cases "edge \ y x") case True from y have y': "y \ RF (TER f)" and essential: "y \ \ (TER f)" by(simp_all add: roofed_circ_def) moreover from x obtain p z where z: "z \ B \" and p: "path \ x p z" and bypass: "\z. z \ set p \ z \ TER f" "x \ TER f" by(clarsimp simp add: roofed_def) from roofedD[OF y' rtrancl_path.step, OF True p z] bypass have "x \ TER f \ y \ TER f" by auto with roofed_greater[THEN subsetD, of x "TER f" \] x have "x \ TER f" "y \ TER f" by auto with essential bypass have False by(auto dest!: not_essentialD[OF _ rtrancl_path.step, OF _ True p z]) thus ?thesis .. qed(simp add: currentD_outside[OF g]) } then show ?thesis unfolding d_IN_def using x by(auto intro!: nn_integral_cong split: split_indicator) qed lemma restrict_current_IN_A: "a \ A (quotient_web \ f) \ d_IN (g \ \ / f) a = 0" by(simp add: d_IN_restrict_current_outside roofed_greaterI) lemma restrict_current_nonneg: "0 \ g e \ 0 \ (g \ \ / f) e" by(cases e) simp lemma in_SINK_restrict_current: "x \ SINK g \ x \ SINK (g \ \ / f)" using d_OUT_restrict_current_le[of \ f g x] by(simp add: SINK.simps) lemma SAT_restrict_current: fixes \ (structure) assumes f: "current \ f" and g: "current \ g" shows "SAT (quotient_web \ f) (g \ \ / f) = RF (TER f) \ (SAT \ g - A \)" (is "SAT ?\ ?g = ?rhs") proof(intro set_eqI iffI; (elim UnE DiffE)?) show "x \ ?rhs" if "x \ SAT ?\ ?g" for x using that proof cases case IN thus ?thesis using currentD_weight_OUT[OF f, of x] by(cases "x \ RF (TER f)")(auto simp add: d_IN_restrict_current_outside roofed_circ_def restrict_current_IN_not_RF[OF g] SAT.IN currentD_IN[OF g] roofed_greaterI SAT.A SINK.simps RF_in_B split: if_split_asm intro: essentialI[OF rtrancl_path.base]) qed(simp add: roofed_greaterI) show "x \ SAT ?\ ?g" if "x \ RF (TER f)" for x using that by(auto simp add: SAT.simps roofed_circ_def d_IN_restrict_current_outside) show "x \ SAT ?\ ?g" if "x \ SAT \ g" "x \ A \" for x using that by(auto simp add: SAT.simps roofed_circ_def d_IN_restrict_current_outside restrict_current_IN_not_RF[OF g]) qed lemma current_restrict_current: fixes \ (structure) assumes w: "wave \ f" and g: "current \ g" shows "current (quotient_web \ f) (g \ \ / f)" (is "current ?\ ?g") proof show "d_OUT ?g x \ weight ?\ x" for x using d_OUT_restrict_current_le[of \ f g x] currentD_weight_OUT[OF g, of x] currentD_OUT[OF g, of x] by(auto simp add: d_OUT_restrict_current_outside) show "d_IN ?g x \ weight ?\ x" for x using d_IN_restrict_current_le[of \ f g x] currentD_weight_IN[OF g, of x] by(auto simp add: d_IN_restrict_current_outside roofed_circ_def) (subst d_IN_restrict_current_outside[of x \ f g]; simp add: roofed_greaterI) fix x assume "x \ A ?\" hence x: "x \ \ (TER f) - B \" by simp show "d_OUT ?g x \ d_IN ?g x" proof(cases "x \ RF (TER f)") case True with x have "x \ RF\<^sup>\ (TER f) \ B \" by(simp add: roofed_circ_def) with True show ?thesis using currentD_OUT[OF g, of x] d_OUT_restrict_current_le[of \ f g x] by(auto simp add: d_OUT_restrict_current_outside d_IN_restrict_current_outside) next case False then obtain p z where z: "z \ B \" and p: "path \ x p z" and bypass: "\z. z \ set p \ z \ TER f" "x \ TER f" by(clarsimp simp add: roofed_def) from g False have "d_IN ?g x = d_IN g x" by(rule restrict_current_IN_not_RF) moreover have "d_OUT ?g x \ d_OUT g x" by(rule d_OUT_mono restrict_current_le)+ moreover have "x \ A \" using separatingD[OF waveD_separating[OF w] p _ z] bypass by blast note currentD_OUT_IN[OF g this] ultimately show ?thesis by simp qed next show "d_IN ?g a = 0" if "a \ A ?\" for a using that by(rule restrict_current_IN_A) show "d_OUT ?g b = 0" if "b \ B ?\" for b using d_OUT_restrict_current_le[of \ f g b] currentD_OUT[OF g, of b] that by simp show "?g e = 0" if "e \ \<^bold>E\<^bsub>?\\<^esub>" for e using that currentD_outside'[OF g, of e] by(cases e)(auto split: split_indicator) qed lemma TER_restrict_current: fixes \ (structure) assumes f: "current \ f" and w: "wave \ f" and g: "current \ g" shows "TER g \ TER\<^bsub>quotient_web \ f\<^esub> (g \ \ / f)" (is "_ \ ?TER" is "_ \ TER\<^bsub>?\\<^esub> ?g") proof fix x assume x: "x \ TER g" hence "x \ SINK ?g" by(simp add: in_SINK_restrict_current) moreover have "x \ RF (TER f)" if "x \ A \" using waveD_separating[OF w, THEN separatingD, OF _ that] by(rule roofedI) then have "x \ SAT ?\ ?g" using SAT_restrict_current[OF f g] x by auto ultimately show "x \ ?TER" by simp qed lemma wave_restrict_current: fixes \ (structure) assumes f: "current \ f" and w: "wave \ f" and g: "current \ g" and w': "wave \ g" shows "wave (quotient_web \ f) (g \ \ / f)" (is "wave ?\ ?g") proof show "separating ?\ (TER\<^bsub>?\\<^esub> ?g)" (is "separating _ ?TER") proof fix x y p assume "x \ A ?\" "y \ B ?\" and p: "path ?\ x p y" hence x: "x \ \ (TER f)" and y: "y \ B \" and SAT: "x \ SAT ?\ ?g" by(simp_all add: SAT.A) from p have p': "path \ x p y" by(rule rtrancl_path_mono) simp { assume "x \ ?TER" hence "x \ SINK ?g" using SAT by(simp) hence "x \ SINK g" using d_OUT_restrict_current_le[of \ f g x] by(auto simp add: SINK.simps) hence "x \ RF (TER g)" using waveD_OUT[OF w'] by(auto simp add: SINK.simps) from roofedD[OF this p' y] \x \ SINK g\ have "(\z\set p. z \ ?TER)" using TER_restrict_current[OF f w g] by blast } then show "(\z\set p. z \ ?TER) \ x \ ?TER" by blast qed fix x assume "x \ RF\<^bsub>?\\<^esub> ?TER" hence "x \ RF (TER g)" proof(rule contrapos_nn) assume *: "x \ RF (TER g)" show "x \ RF\<^bsub>?\\<^esub> ?TER" proof fix p y assume "path ?\ x p y" "y \ B ?\" hence "path \ x p y" "y \ B \" by(auto elim: rtrancl_path_mono) from roofedD[OF * this] show "(\z\set p. z \ ?TER) \ x \ ?TER" using TER_restrict_current[OF f w g] by blast qed qed with w' have "d_OUT g x = 0" by(rule waveD_OUT) with d_OUT_restrict_current_le[of \ f g x] show "d_OUT ?g x = 0" by simp qed definition plus_current :: "'v current \ 'v current \ 'v current" where "plus_current f g = (\e. f e + g e)" lemma plus_current_simps [simp]: "plus_current f g e = f e + g e" by(simp add: plus_current_def) lemma plus_zero_current [simp]: "plus_current f zero_current = f" by(simp add: fun_eq_iff) lemma support_flow_plus_current: "support_flow (plus_current f g) \ support_flow f \ support_flow g" by(clarsimp simp add: support_flow.simps) context fixes \ :: "('v, 'more) web_scheme" (structure) and f g assumes f: "current \ f" and w: "wave \ f" and g: "current (quotient_web \ f) g" begin lemma OUT_plus_current: "d_OUT (plus_current f g) x = (if x \ RF\<^sup>\ (TER f) then d_OUT f x else d_OUT g x)" (is "d_OUT ?g _ = _") proof - have "d_OUT ?g x = d_OUT f x + d_OUT g x" unfolding plus_current_def by(subst d_OUT_add) simp_all also have "\ = (if x \ RF\<^sup>\ (TER f) then d_OUT f x else d_OUT g x)" proof(cases "x \ RF\<^sup>\ (TER f)") case True hence "d_OUT g x = 0" by(intro currentD_outside_OUT[OF g])(auto dest: vertex_quotient_webD) thus ?thesis using True by simp next case False hence "d_OUT f x = 0" by(auto simp add: roofed_circ_def SINK.simps dest: waveD_OUT[OF w]) with False show ?thesis by simp qed finally show ?thesis . qed lemma IN_plus_current: "d_IN (plus_current f g) x = (if x \ RF (TER f) then d_IN f x else d_IN g x)" (is "d_IN ?g _ = _") proof - have "d_IN ?g x = d_IN f x + d_IN g x" unfolding plus_current_def by(subst d_IN_add) simp_all also consider (RF) "x \ RF (TER f) - (B \ - A \)" | (B) "x \ RF (TER f)" "x \ B \ - A \" | (beyond) "x \ RF (TER f)" by blast then have "d_IN f x + d_IN g x = (if x \ RF (TER f) then d_IN f x else d_IN g x)" proof(cases) case RF hence "d_IN g x = 0" by(cases "x \ \ (TER f)")(auto intro: currentD_outside_IN[OF g] currentD_IN[OF g] dest!: vertex_quotient_webD simp add: roofed_circ_def) thus ?thesis using RF by simp next case B hence "d_IN g x = 0" using currentD_outside_IN[OF g, of x] currentD_weight_IN[OF g, of x] by(auto dest: vertex_quotient_webD simp add: roofed_circ_def) with B show ?thesis by simp next case beyond from f w beyond have "d_IN f x = 0" by(rule wave_not_RF_IN_zero) with beyond show ?thesis by simp qed finally show ?thesis . qed lemma in_TER_plus_current: assumes RF: "x \ RF\<^sup>\ (TER f)" and x: "x \ TER\<^bsub>quotient_web \ f\<^esub> g" (is "_ \ ?TER _") shows "x \ TER (plus_current f g)" (is "_ \ TER ?g") proof(cases "x \ \ (TER f) - (B \ - A \)") case True with x show ?thesis using currentD_IN[OF g, of x] by(fastforce intro: roofed_greaterI SAT.intros simp add: SINK.simps OUT_plus_current IN_plus_current elim!: SAT.cases) next case *: False have "x \ SAT \ ?g" proof(cases "x \ B \ - A \") case False with x RF * have "weight \ x \ d_IN g x" by(auto elim!: SAT.cases split: if_split_asm simp add: essential_BI) also have "\ \ d_IN ?g x" unfolding plus_current_def by(intro d_IN_mono) simp finally show ?thesis .. next case True with * x have "weight \ x \ d_IN ?g x" using currentD_OUT[OF f, of x] by(auto simp add: IN_plus_current RF_in_B SINK.simps roofed_circ_def elim!: SAT.cases split: if_split_asm) thus ?thesis .. qed moreover have "x \ SINK ?g" using x by(simp add: SINK.simps OUT_plus_current RF) ultimately show ?thesis by simp qed lemma current_plus_current: "current \ (plus_current f g)" (is "current _ ?g") proof show "d_OUT ?g x \ weight \ x" for x using currentD_weight_OUT[OF g, of x] currentD_weight_OUT[OF f, of x] by(auto simp add: OUT_plus_current split: if_split_asm elim: order_trans) show "d_IN ?g x \ weight \ x" for x using currentD_weight_IN[OF f, of x] currentD_weight_IN[OF g, of x] by(auto simp add: IN_plus_current roofed_circ_def split: if_split_asm elim: order_trans) show "d_OUT ?g x \ d_IN ?g x" if "x \ A \" for x proof(cases "x \ \ (TER f)") case False thus ?thesis using currentD_OUT_IN[OF f that] currentD_OUT_IN[OF g, of x] that by(auto simp add: OUT_plus_current IN_plus_current roofed_circ_def SINK.simps) next case True with that have "d_OUT f x = 0" "weight \ x \ d_IN f x" by(auto simp add: SINK.simps elim: SAT.cases) thus ?thesis using that True currentD_OUT_IN[OF g, of x] currentD_weight_OUT[OF g, of x] by(auto simp add: OUT_plus_current IN_plus_current roofed_circ_def intro: roofed_greaterI split: if_split_asm) qed show "d_IN ?g a = 0" if "a \ A \" for a using wave_A_in_RF[OF w that] currentD_IN[OF f that] by(simp add: IN_plus_current) show "d_OUT ?g b = 0" if "b \ B \" for b using that currentD_OUT[OF f that] currentD_OUT[OF g, of b] that by(auto simp add: OUT_plus_current SINK.simps roofed_circ_def intro: roofed_greaterI) show "?g e = 0" if "e \ \<^bold>E" for e using currentD_outside'[OF f, of e] currentD_outside'[OF g, of e] that by(cases e) auto qed context assumes w': "wave (quotient_web \ f) g" begin lemma separating_TER_plus_current: assumes x: "x \ RF (TER f)" and y: "y \ B \" and p: "path \ x p y" shows "(\z\set p. z \ TER (plus_current f g)) \ x \ TER (plus_current f g)" (is "_ \ _ \ TER ?g") proof - from x have "x \ RF (\ (TER f))" unfolding RF_essential . from roofedD[OF this p y] have "\z\set (x # p). z \ \ (TER f)" by auto from split_list_last_prop[OF this] obtain ys z zs where decomp: "x # p = ys @ z # zs" and z: "z \ \ (TER f)" and outside: "\z. z \ set zs \ z \ \ (TER f)" by auto have zs: "path \ z zs y" using decomp p by(cases ys)(auto elim: rtrancl_path_appendE) moreover have "z \ RF\<^sup>\ (TER f)" using z by(simp add: roofed_circ_def) moreover have RF: "z' \ RF (TER f)" if "z' \ set zs" for z' proof assume "z' \ RF (TER f)" hence z': "z' \ RF (\ (TER f))" by(simp only: RF_essential) from split_list[OF that] obtain ys' zs' where decomp': "zs = ys' @ z' # zs'" by blast with zs have "path \ z' zs' y" by(auto elim: rtrancl_path_appendE) from roofedD[OF z' this y] outside decomp' show False by auto qed ultimately have p': "path (quotient_web \ f) z zs y" by(rule path_quotient_web) show ?thesis proof(cases "z \ B \ - A \") case False with separatingD[OF waveD_separating[OF w'] p'] z y obtain z' where z': "z' \ set (z # zs)" and TER: "z' \ TER\<^bsub>quotient_web \ f\<^esub> g" by auto hence "z' \ TER ?g" using in_TER_plus_current[of z'] RF[of z'] \z \ RF\<^sup>\ (TER f)\ by(auto simp add: roofed_circ_def) with decomp z' show ?thesis by(cases ys) auto next case True hence "z \ TER ?g" using currentD_OUT[OF current_plus_current, of z] z by(auto simp add: SINK.simps SAT.simps IN_plus_current intro: roofed_greaterI) then show ?thesis using decomp by(cases ys) auto qed qed lemma wave_plus_current: "wave \ (plus_current f g)" (is "wave _ ?g") proof let ?\ = "quotient_web \ f" let ?TER = "TER\<^bsub>?\\<^esub>" show "separating \ (TER ?g)" using separating_TER_plus_current[OF wave_A_in_RF[OF w]] by(rule separating) fix x assume x: "x \ RF (TER ?g)" hence "x \ RF (TER f)" by(rule contrapos_nn)(rule roofedI, rule separating_TER_plus_current) hence *: "x \ RF\<^sup>\ (TER f)" by(simp add: roofed_circ_def) moreover have "x \ RF\<^bsub>?\\<^esub> (?TER g)" proof assume RF': "x \ RF\<^bsub>?\\<^esub> (?TER g)" from x obtain p y where y: "y \ B \" and p: "path \ x p y" and bypass: "\z. z \ set p \ z \ TER ?g" and x': "x \ TER ?g" by(auto simp add: roofed_def) have RF: "z \ RF (TER f)" if "z \ set p" for z proof assume z: "z \ RF (TER f)" from split_list[OF that] obtain ys' zs' where decomp: "p = ys' @ z # zs'" by blast with p have "path \ z zs' y" by(auto elim: rtrancl_path_appendE) from separating_TER_plus_current[OF z y this] decomp bypass show False by auto qed with p have "path ?\ x p y" using * by(induction)(auto intro: rtrancl_path.intros simp add: roofed_circ_def) from roofedD[OF RF' this] y consider (x) "x \ ?TER g" | (z) z where "z \ set p" "z \ ?TER g" by auto then show False proof(cases) case x with * have "x \ TER ?g" by(rule in_TER_plus_current) with x' show False by contradiction next case (z z) from z(1) have "z \ RF (TER f)" by(rule RF) hence "z \ RF\<^sup>\ (TER f)" by(simp add: roofed_circ_def) hence "z \ TER ?g" using z(2) by(rule in_TER_plus_current) moreover from z(1) have "z \ TER ?g" by(rule bypass) ultimately show False by contradiction qed qed with w' have "d_OUT g x = 0" by(rule waveD_OUT) ultimately show "d_OUT ?g x = 0" by(simp add: OUT_plus_current) qed end end lemma loose_quotient_web: fixes \ :: "('v, 'more) web_scheme" (structure) assumes weight_finite: "\x. weight \ x \ \" and f: "current \ f" and w: "wave \ f" and maximal: "\w. \ current \ w; wave \ w; f \ w \ \ f = w" shows "loose (quotient_web \ f)" (is "loose ?\") proof fix g assume g: "current ?\ g" and w': "wave ?\ g" let ?g = "plus_current f g" from f w g have "current \ ?g" "wave \ ?g" by(rule current_plus_current wave_plus_current)+ (rule w') moreover have "f \ ?g" by(clarsimp simp add: le_fun_def add_eq_0_iff_both_eq_0) ultimately have eq: "f = ?g" by(rule maximal) have "g e = 0" for e proof(cases e) case (Pair x y) have "f e \ d_OUT f x" unfolding Pair by (rule d_OUT_ge_point) also have "\ \ weight \ x" by(rule currentD_weight_OUT[OF f]) also have "\ < \" by(simp add: weight_finite less_top[symmetric]) finally show "g e = 0" using Pair eq[THEN fun_cong, of e] by(cases "f e" "g e" rule: ennreal2_cases)(simp_all add: fun_eq_iff) qed thus "g = (\_. 0)" by(simp add: fun_eq_iff) next have 0: "current ?\ zero_current" by(simp add: ) show "\ hindrance ?\ zero_current" proof assume "hindrance ?\ zero_current" then obtain x where a: "x \ A ?\" and \: "x \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> zero_current)" and "d_OUT zero_current x < weight ?\ x" by cases from a have "x \ \ (TER f)" by simp then obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. \x \ y; z \ set p\ \ z = x \ z \ TER f" by(rule \_E) blast from p obtain p' where p': "path \ x p' y" and distinct: "distinct (x # p')" and subset: "set p' \ set p" by(auto elim: rtrancl_path_distinct) note p' moreover have RF: "z \ RF (TER f)" if "z \ set p'" for z proof assume z: "z \ RF (TER f)" from split_list[OF that] obtain ys zs where decomp: "p' = ys @ z # zs" by blast with p' have "y \ set p'" by(auto dest!: rtrancl_path_last intro: last_in_set) with distinct have neq: "x \ y" by auto from decomp p' have "path \ z zs y" by(auto elim: rtrancl_path_appendE) from roofedD[OF z this y] obtain z' where "z' \ set (z # zs)" "z' \ TER f" by auto with distinct decomp subset bypass[OF neq] show False by auto qed moreover have "x \ RF\<^sup>\ (TER f)" using \x \ \ (TER f)\ by(simp add: roofed_circ_def) ultimately have p'': "path ?\ x p' y" by(induction)(auto intro: rtrancl_path.intros simp add: roofed_circ_def) from a \ have "\ essential ?\ (B ?\) (TER\<^bsub>?\\<^esub> zero_current) x" by simp from not_essentialD[OF this p''] y obtain z where neq: "x \ y" and "z \ set p'" "z \ x" "z \ TER\<^bsub>?\\<^esub> zero_current" by auto moreover with subset RF[of z] have "z \ TER f" using currentD_weight_OUT[OF f, of z] currentD_weight_IN[OF f, of z] by(auto simp add: roofed_circ_def SINK.simps intro: SAT.IN split: if_split_asm) ultimately show False using bypass[of z] subset by auto qed qed lemma quotient_web_trimming: fixes \ (structure) assumes w: "wave \ f" and trimming: "trimming \ f g" shows "quotient_web \ f = quotient_web \ g" (is "?lhs = ?rhs") proof(rule web.equality) from trimming have \: "\ (TER g) - A \ = \ (TER f) - A \" by cases have RF: "RF (TER g) = RF (TER f)" by(subst (1 2) RF_essential[symmetric])(simp only: trimming_\[OF w trimming]) have RFc: "RF\<^sup>\ (TER g) = RF\<^sup>\ (TER f)" by(subst (1 2) roofed_circ_essential[symmetric])(simp only: trimming_\[OF w trimming]) show "edge ?lhs = edge ?rhs" by(rule ext)+(simp add: RF RFc) have "weight ?lhs = (\x. if x \ RF\<^sup>\ (TER g) \ x \ RF (TER g) \ B \ then 0 else weight \ x)" unfolding RF RFc by(auto simp add: fun_eq_iff RF_in_B) also have "\ = weight ?rhs" by(auto simp add: fun_eq_iff RF_in_B) finally show "weight ?lhs = weight ?rhs" . show "A ?lhs = A ?rhs" unfolding quotient_web_sel trimming_\[OF w trimming] .. qed simp_all subsection \Well-formed webs\ locale web = fixes \ :: "('v, 'more) web_scheme" (structure) assumes A_in: "x \ A \ \ \ edge \ y x" and B_out: "x \ B \ \ \ edge \ x y" and A_vertex: "A \ \ \<^bold>V" and disjoint: "A \ \ B \ = {}" and no_loop: "\x. \ edge \ x x" and weight_outside: "\x. x \ \<^bold>V \ weight \ x = 0" and weight_finite [simp]: "\x. weight \ x \ \" begin lemma web_weight_update: assumes "\x. \ vertex \ x \ w x = 0" and "\x. w x \ \" shows "web (\\weight := w\)" by unfold_locales(simp_all add: A_in B_out A_vertex disjoint no_loop assms) lemma currentI [intro?]: assumes "\x. d_OUT f x \ weight \ x" and "\x. d_IN f x \ weight \ x" and OUT_IN: "\x. \ x \ A \; x \ B \ \ \ d_OUT f x \ d_IN f x" and outside: "\e. e \ \<^bold>E \ f e = 0" shows "current \ f" proof show "d_IN f a = 0" if "a \ A \" for a using that by(auto simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0 A_in intro: outside) show "d_OUT f b = 0" if "b \ B \" for b using that by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 B_out intro: outside) then show "d_OUT f x \ d_IN f x" if "x \ A \" for x using OUT_IN[OF that] by(cases "x \ B \") auto qed(blast intro: assms)+ lemma currentD_finite_IN: assumes f: "current \ f" shows "d_IN f x \ \" proof(cases "x \ \<^bold>V") case True have "d_IN f x \ weight \ x" using f by(rule currentD_weight_IN) also have "\ < \" using True weight_finite[of x] by (simp add: less_top[symmetric]) finally show ?thesis by simp next case False then have "d_IN f x = 0" by(auto simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0 vertex_def intro: currentD_outside[OF f]) thus ?thesis by simp qed lemma currentD_finite_OUT: assumes f: "current \ f" shows "d_OUT f x \ \" proof(cases "x \ \<^bold>V") case True have "d_OUT f x \ weight \ x" using f by(rule currentD_weight_OUT) also have "\ < \" using True weight_finite[of x] by (simp add: less_top[symmetric]) finally show ?thesis by simp next case False then have "d_OUT f x = 0" by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 vertex_def intro: currentD_outside[OF f]) thus ?thesis by simp qed lemma currentD_finite: assumes f: "current \ f" shows "f e \ \" proof(cases e) case (Pair x y) have "f (x, y) \ d_OUT f x" by (rule d_OUT_ge_point) also have "\ < \" using currentD_finite_OUT[OF f] by (simp add: less_top[symmetric]) finally show ?thesis by(simp add: Pair) qed lemma web_quotient_web: "web (quotient_web \ f)" (is "web ?\") proof show "\ edge ?\ y x" if "x \ A ?\" for x y using that by(auto intro: roofed_greaterI) show "\ edge ?\ x y" if "x \ B ?\" for x y using that by(auto simp add: B_out) show "A ?\ \ B ?\ = {}" using disjoint by auto show "A ?\ \ \<^bold>V\<^bsub>?\\<^esub>" proof fix x assume "x \ A ?\" hence \: "x \ \ (TER f)" and x: "x \ B \" using disjoint by auto from this(1) obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER f)" by(rule \_E_RF) blast from p y x have "p \ []" by(auto simp add: rtrancl_path_simps) with rtrancl_path_nth[OF p, of 0] have "edge \ x (p ! 0)" by simp moreover have "x \ RF\<^sup>\ (TER f)" using \ by(simp add: roofed_circ_def) moreover have "p ! 0 \ RF (TER f)" using bypass \p \ []\ by auto ultimately have "edge ?\ x (p ! 0)" by simp thus "x \ \<^bold>V\<^bsub>?\\<^esub>" by(auto intro: vertexI1) qed show "\ edge ?\ x x" for x by(simp add: no_loop) show "weight ?\ x = 0" if "x \ \<^bold>V\<^bsub>?\\<^esub>" for x proof(cases "x \ RF\<^sup>\ (TER f) \ x \ TER f \ B \") case True thus ?thesis by simp next case False hence RF: "x \ RF\<^sup>\ (TER f)" and B: "x \ B \ \ x \ TER f" by auto from RF obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER f)" apply(cases "x \ RF (RF (TER f))") apply(auto elim!: not_roofedE)[1] apply(auto simp add: roofed_circ_def roofed_idem elim: essentialE_RF) done from that have "p = []" using p y B RF bypass by(auto 4 3 simp add: vertex_def dest!: rtrancl_path_nth[where i=0]) with p have xy: "x = y" by(simp add: rtrancl_path_simps) with B y have "x \ TER f" by simp hence RF': "x \ RF (TER f)" using xy y by(subst RF_in_B) auto have "\ vertex \ x" proof assume "vertex \ x" then obtain x' where "edge \ x' x" using xy y by(auto simp add: vertex_def B_out) moreover hence "x' \ RF\<^sup>\ (TER f)" using RF' by(auto dest: RF_circ_edge_forward) ultimately have "edge ?\ x' x" using RF' by simp hence "vertex ?\ x" by(rule vertexI2) with that show False by simp qed thus ?thesis by(simp add: weight_outside) qed show "weight ?\ x \ \" for x by simp qed end locale countable_web = web \ for \ :: "('v, 'more) web_scheme" (structure) + - assumes edge_antiparallel: "edge \ x y \ \ edge \ y x" - and countable [simp]: "countable \<^bold>E" + assumes countable [simp]: "countable \<^bold>E" begin lemma countable_V [simp]: "countable \<^bold>V" by(simp add: "\<^bold>V_def") lemma countable_web_quotient_web: "countable_web (quotient_web \ f)" (is "countable_web ?\") proof - interpret r: web ?\ by(rule web_quotient_web) show ?thesis proof have "\<^bold>E\<^bsub>?\\<^esub> \ \<^bold>E" by auto then show "countable \<^bold>E\<^bsub>?\\<^esub>" by(rule countable_subset) simp - qed(simp add: edge_antiparallel) + qed qed end subsection \Subtraction of a wave\ definition minus_web :: "('v, 'more) web_scheme \ 'v current \ ('v, 'more) web_scheme" (infixl "\" 65) \ \Definition 6.6\ where "\ \ f = \\weight := \x. if x \ A \ then weight \ x - d_OUT f x else weight \ x + d_OUT f x - d_IN f x\" lemma minus_web_sel [simp]: "edge (\ \ f) = edge \" "weight (\ \ f) x = (if x \ A \ then weight \ x - d_OUT f x else weight \ x + d_OUT f x - d_IN f x)" "A (\ \ f) = A \" "B (\ \ f) = B \" "\<^bold>V\<^bsub>\ \ f\<^esub> = \<^bold>V\<^bsub>\\<^esub>" "\<^bold>E\<^bsub>\ \ f\<^esub> = \<^bold>E\<^bsub>\\<^esub>" "web.more (\ \ f) = web.more \" by(auto simp add: minus_web_def vertex_def) lemma vertex_minus_web [simp]: "vertex (\ \ f) = vertex \" by(simp add: vertex_def fun_eq_iff) lemma roofed_gen_minus_web [simp]: "roofed_gen (\ \ f) = roofed_gen \" by(simp add: fun_eq_iff roofed_def) lemma minus_zero_current [simp]: "\ \ zero_current = \" by(rule web.equality)(simp_all add: fun_eq_iff) lemma (in web) web_minus_web: assumes f: "current \ f" shows "web (\ \ f)" unfolding minus_web_def apply(rule web_weight_update) apply(auto simp: weight_outside currentD_weight_IN[OF f] currentD_outside_OUT[OF f] currentD_outside_IN[OF f] currentD_weight_OUT[OF f] currentD_finite_OUT[OF f]) done subsection \Bipartite webs\ locale countable_bipartite_web = fixes \ :: "('v, 'more) web_scheme" (structure) assumes bipartite_V: "\<^bold>V \ A \ \ B \" and A_vertex: "A \ \ \<^bold>V" and bipartite_E: "edge \ x y \ x \ A \ \ y \ B \" and disjoint: "A \ \ B \ = {}" and weight_outside: "\x. x \ \<^bold>V \ weight \ x = 0" and weight_finite [simp]: "\x. weight \ x \ \" and countable_E [simp]: "countable \<^bold>E" begin lemma not_vertex: "\ x \ A \; x \ B \ \ \ \ vertex \ x" using bipartite_V by blast lemma no_loop: "\ edge \ x x" using disjoint by(auto dest: bipartite_E) lemma edge_antiparallel: "edge \ x y \ \ edge \ y x" using disjoint by(auto dest: bipartite_E) lemma A_in: "x \ A \ \ \ edge \ y x" using disjoint by(auto dest: bipartite_E) lemma B_out: "x \ B \ \ \ edge \ x y" using disjoint by(auto dest: bipartite_E) sublocale countable_web using disjoint -by(unfold_locales)(auto simp add: A_in B_out A_vertex no_loop weight_outside edge_antiparallel) +by(unfold_locales)(auto simp add: A_in B_out A_vertex no_loop weight_outside) lemma currentD_OUT': assumes f: "current \ f" and x: "x \ A \" shows "d_OUT f x = 0" using currentD_outside_OUT[OF f, of x] x currentD_OUT[OF f, of x] bipartite_V by auto lemma currentD_IN': assumes f: "current \ f" and x: "x \ B \" shows "d_IN f x = 0" using currentD_outside_IN[OF f, of x] x currentD_IN[OF f, of x] bipartite_V by auto lemma current_bipartiteI [intro?]: assumes OUT: "\x. d_OUT f x \ weight \ x" and IN: "\x. d_IN f x \ weight \ x" and outside: "\e. e \ \<^bold>E \ f e = 0" shows "current \ f" proof fix x assume "x \ A \" "x \ B \" hence "d_OUT f x = 0" by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 intro!: outside dest: bipartite_E) then show "d_OUT f x \ d_IN f x" by simp qed(rule OUT IN outside)+ lemma wave_bipartiteI [intro?]: assumes sep: "separating \ (TER f)" and f: "current \ f" shows "wave \ f" using sep proof(rule wave.intros) fix x assume "x \ RF (TER f)" then consider "x \ \<^bold>V" | "x \ \<^bold>V" "x \ B \" using separating_RF_A[OF sep] bipartite_V by auto then show "d_OUT f x = 0" using currentD_OUT[OF f, of x] currentD_outside_OUT[OF f, of x] by cases auto qed lemma web_flow_iff: "web_flow \ f \ current \ f" using bipartite_V by(auto simp add: web_flow.simps) end end