Theory Feedback_Arcs
section ‹Find a Feedback Arc Set›
theory Feedback_Arcs
imports
"../DFS_Framework"
CAVA_Automata.Digraph_Impl
Reachable_Nodes
begin
text ‹A feedback arc set is a set of edges that breaks all reachable cycles.
In this theory, we define an algorithm to find a feedback arc set.›
definition is_fas :: "('v, 'more) graph_rec_scheme ⇒ 'v rel ⇒ bool" where
"is_fas G EC ≡ ¬(∃ u ∈ (g_E G)⇧* `` g_V0 G. (u, u) ∈ (g_E G - EC)⇧+)"
lemma is_fas_alt:
"is_fas G EC = acyclic ((g_E G ∩ ((g_E G)⇧* `` g_V0 G × UNIV) - EC))"
unfolding is_fas_def acyclic_def
proof (clarsimp, safe)
fix u
assume A: "(u,u) ∈ (g_E G ∩ (g_E G)⇧* `` g_V0 G × UNIV - EC)⇧+"
hence "(u,u)∈(g_E G - EC)⇧+" by (rule trancl_mono) blast
moreover from A have "u ∈ (g_E G)⇧* `` g_V0 G" by (cases rule: converse_tranclE) auto
moreover assume "∀u∈(g_E G)⇧* `` g_V0 G. (u, u) ∉ (g_E G - EC)⇧+"
ultimately show False by blast
next
fix u v0
assume 1: "v0∈g_V0 G" and 2: "(v0,u)∈(g_E G)⇧*" and 3: "(u,u)∈(g_E G - EC)⇧+"
have "(u, u) ∈ (Restr (g_E G - EC) ((g_E G)⇧* `` g_V0 G))⇧+"
apply (rule trancl_restrict_reachable[OF 3, where S="(g_E G)⇧* `` g_V0 G"])
apply (rule order_trans[OF _ rtrancl_image_unfold_right])
using 1 2 by auto
hence "(u, u) ∈ (g_E G ∩ (g_E G)⇧* `` g_V0 G × UNIV - EC)⇧+"
by (rule trancl_mono) auto
moreover assume "∀x. (x, x) ∉ (g_E G ∩ (g_E G)⇧* `` g_V0 G × UNIV - EC)⇧+"
ultimately show False by blast
qed
subsection ‹Instantiation of the DFS-Framework›
record 'v fas_state = "'v state" +
fas :: "('v×'v) set"
lemma fas_more_cong: "state.more s = state.more s' ⟹ fas s = fas s'"
by (cases s, cases s', simp)
lemma [simp]: "s⦇ state.more := ⦇ fas = foo ⦈ ⦈ = s ⦇ fas := foo ⦈"
by (cases s) simp
definition fas_params :: "('v,('v,unit) fas_state_ext) parameterization"
where "fas_params ≡ dflt_parametrization state.more
(RETURN ⦇ fas = {} ⦈) ⦇
on_back_edge := λu v s. RETURN ⦇ fas = insert (u,v) (fas s) ⦈
⦈"
lemmas fas_params_simp[simp] =
gen_parameterization.simps[mk_record_simp, OF fas_params_def[simplified]]