Theory SCC_Bloemen_Sequential

section ‹Overview›

text ‹
  Computing the maximal strongly connected components (SCCs) of a 
  finite directed graph is a celebrated problem in the
  theory of graph algorithms. Although Tarjan's algorithm~cite"tarjan:depth-first" 
  is perhaps the best-known solution, there are many others. In his PhD 
  thesis, Bloemen~cite"bloemen:strong" presents an algorithm that is itself based
  on earlier algorithms by Munro~cite"munro:efficient" and
  Dijkstra~cite"dijkstra:finding". Just like these algorithms, Bloemen's
  solution is based on enumerating SCCs in a depth-first traversal of the graph.
  Gabow's algorithm that has already been formalized in Isabelle~cite"lammich:gabow"
  also falls into this category of solutions.
  Nevertheless, Bloemen goes on to present a parallel variant of the algorithm
  suitable for execution on multi-core processors, based on clever data structures
  that minimize locking.

  In the following, we encode the sequential version of the algorithm in the
  proof assistant Isabelle/HOL, and prove its correctness. Bloemen's thesis
  briefly and informally explains why the algorithm is correct. Our proof expands
  on these arguments, making them completely formal. The encoding is based on
  a direct representation of the algorithm as a pair of mutually recursive
  functions; we are not aiming at extracting executable code.
›

theory SCC_Bloemen_Sequential
imports Main
begin

text ‹
  The record below represents the variables of the
  algorithm. Most variables correspond to those used in
  Bloemen's presentation. Thus, the variable @{text 𝒮}
  associates to every node the set of nodes that have
  already been determined to be part of the same SCC. A core
  invariant of the algorithm will be that this mapping represents
  equivalence classes of nodes: for all nodes @{text v} and @{text w},
  we maintain the relationship

  @{text "v ∈ 𝒮 w ⟷ 𝒮 v = 𝒮 w."}

  In an actual implementation of this algorithm, this variable
  could conveniently be represented by a union-find structure.
  Variable @{text stack} holds the list of roots of these 
  (not yet maximal) SCCs, in depth-first order,
  @{text visited} and @{text explored}
  represent the nodes that have already been seen, respectively
  that have been completely explored, by the algorithm, and
  @{text sccs} is the set of maximal SCCs that the algorithm
  has found so far.

  Additionally, the record holds some auxiliary variables that
  are used in the proof of correctness. In particular,
  @{text root} denotes the node on which the algorithm was called,
  @{text cstack} represents the call stack of the recursion of
  function @{text dfs},
  and @{text vsuccs} stores the successors of each node
  that have already been visited by the function @{text dfss}
  that loops over all successors of a given node.
›
record 'v env =
  root :: "'v"
  𝒮 :: "'v  'v set"
  explored :: "'v set"
  visited :: "'v set"
  vsuccs :: "'v  'v set"
  sccs :: "'v set set"
  stack :: "'v list"
  cstack :: "'v list"

text ‹
  The algorithm is initially called with an environment that
  initializes the root node and trivializes all other components.
›
definition init_env where
  "init_env v = 
    root = v,
    𝒮 = (λu. {u}),
    explored = {},
    visited = {},
    vsuccs = (λu. {}),
    sccs = {},
    stack = [],
    cstack = []
   "

― ‹Make the simplifier expand let-constructions automatically.›
declare Let_def[simp]



section ‹Auxiliary lemmas about lists›

text ‹
  We use the precedence order on the elements that appear
  in a list. In particular, stacks are represented as lists,
  and a node @{text x} precedes another node @{text y} on the
  stack if @{text x} was pushed on the stack later
  than @{text y}.
›

definition precedes ("_  _ in _" [100,100,100] 39) where
  "x  y in xs  l r. xs = l @ (x # r)  y  set (x # r)"

lemma precedes_mem:
  assumes "x  y in xs"
  shows "x  set xs" "y  set xs"
  using assms unfolding precedes_def by auto

lemma head_precedes:
  assumes "y  set (x # xs)"
  shows "x  y in (x # xs)"
  using assms unfolding precedes_def by force

lemma precedes_in_tail:
  assumes "x  z"
  shows "x  y in (z # zs)  x  y in zs"
  using assms unfolding precedes_def by (auto simp: Cons_eq_append_conv)

lemma tail_not_precedes:
  assumes "y  x in (x # xs)" "x  set xs"
  shows "x = y"
  using assms unfolding precedes_def
  by (metis Cons_eq_append_conv Un_iff list.inject set_append)

lemma split_list_precedes:
  assumes "y  set (ys @ [x])"
  shows "y  x in (ys @ x # xs)"
  using assms unfolding precedes_def
  by (metis append_Cons append_assoc in_set_conv_decomp
            rotate1.simps(2) set_ConsD set_rotate1)

lemma precedes_refl [simp]: "(x  x in xs) = (x  set xs)"
proof
  assume "x  x in xs" thus "x  set xs"
    by (simp add: precedes_mem)
next
  assume "x  set xs"
  from this[THEN split_list] show "x  x in xs"
    unfolding precedes_def by auto
qed

lemma precedes_append_left:
  assumes "x  y in xs"
  shows "x  y in (ys @ xs)"
  using assms unfolding precedes_def by (metis append.assoc)

lemma precedes_append_left_iff:
  assumes "x  set ys"
  shows "x  y in (ys @ xs)  x  y in xs" (is "?lhs = ?rhs")
proof
  assume "?lhs"
  then obtain l r where lr: "ys @ xs = l @ (x # r)" "y  set (x # r)"
    unfolding precedes_def by blast
  then obtain us where
    "(ys = l @ us  us @ xs = x # r)  (ys @ us = l  xs = us @ (x # r))"
    by (auto simp: append_eq_append_conv2)
  thus ?rhs
  proof
    assume us: "ys = l @ us  us @ xs = x # r"
    with assms have "us = []"
      by (metis Cons_eq_append_conv in_set_conv_decomp)
    with us lr show ?rhs
      unfolding precedes_def by auto
  next
    assume us: "ys @ us = l  xs = us @ (x # r)"
    with y  set (x # r) show ?rhs
      unfolding precedes_def by blast
  qed
next
  assume "?rhs" thus "?lhs" by (rule precedes_append_left)
qed

lemma precedes_append_right:
  assumes "x  y in xs"
  shows "x  y in (xs @ ys)"
  using assms unfolding precedes_def by force

lemma precedes_append_right_iff:
  assumes "y  set ys"
  shows "x  y in (xs @ ys)  x  y in xs" (is "?lhs = ?rhs")
proof
  assume ?lhs
  then obtain l r where lr: "xs @ ys = l @ (x # r)" "y  set (x # r)"
    unfolding precedes_def by blast
  then obtain us where
    "(xs = l @ us  us @ ys = x # r)  (xs @ us = l  ys = us @ (x # r))"
    by (auto simp: append_eq_append_conv2)
  thus ?rhs
  proof
    assume us: "xs = l @ us  us @ ys = x # r"
    with y  set (x # r) assms show ?rhs
      unfolding precedes_def by (metis Cons_eq_append_conv Un_iff set_append)
  next
    assume us: "xs @ us = l  ys = us @ (x # r)"
    with y  set (x # r) assms 
    show ?rhs by auto ― ‹contradiction›
  qed
next
  assume ?rhs thus ?lhs by (rule precedes_append_right)
qed

text ‹
  Precedence determines an order on the elements of a list,
  provided elements have unique occurrences. However, consider
  a list such as @{text "[2,3,1,2]"}: then $1$ precedes $2$ and
  $2$ precedes $3$, but $1$ does not precede $3$.
›
lemma precedes_trans:
  assumes "x  y in xs" and "y  z in xs" and "distinct xs"
  shows "x  z in xs"
  using assms unfolding precedes_def
  by (smt Un_iff append.assoc append_Cons_eq_iff distinct_append 
          not_distinct_conv_prefix set_append split_list_last)

lemma precedes_antisym:
  assumes "x  y in xs" and "y  x in xs" and "distinct xs"
  shows "x = y"
proof -
  from x  y in xs distinct xs obtain as bs where
    1: "xs = as @ (x # bs)" "y  set (x # bs)" "y  set as"
    unfolding precedes_def by force
  from y  x in xs distinct xs obtain cs ds where
    2: "xs = cs @ (y # ds)" "x  set (y # ds)" "x  set cs"
    unfolding precedes_def by force
  from 1 2 have "as @ (x # bs) = cs @ (y # ds)"
    by simp
  then obtain zs where
    "(as = cs @ zs  zs @ (x # bs) = y # ds) 
      (as @ zs = cs  x # bs = zs @ (y # ds))"  (is "?P  ?Q")
    by (auto simp: append_eq_append_conv2)
  then show ?thesis
  proof
    assume "?P" with y  set as show ?thesis 
      by (cases "zs") auto
  next
    assume "?Q" with x  set cs show ?thesis
      by (cases "zs") auto
  qed
qed


section ‹Finite directed graphs›

text ‹
  We represent a graph as an Isabelle locale that identifies a finite
  set of vertices (of some base type @{text "'v"}) and associates to
  each vertex its set of successor vertices.
›

locale graph =
  fixes vertices :: "'v set"
    and successors :: "'v  'v set"
  assumes vfin: "finite vertices"
    and sclosed: "x  vertices. successors x  vertices"

context graph
begin

abbreviation edge where
  "edge x y  y  successors x"

text ‹
  We inductively define reachability of nodes in the graph.
›
inductive reachable where
  reachable_refl[iff]: "reachable x x"
| reachable_succ[elim]: "edge x y; reachable y z  reachable x z"

lemma reachable_edge: "edge x y  reachable x y"
  by auto

lemma succ_reachable:
  assumes "reachable x y" and "edge y z"
  shows "reachable x z"
  using assms by induct auto

lemma reachable_trans:
  assumes y: "reachable x y" and z: "reachable y z"
  shows "reachable x z"
  using assms by induct auto

text ‹
  In order to derive a ``reverse'' induction rule for @{const "reachable"},
  we define an alternative reachability predicate and prove that the two
  coincide.
›
inductive reachable_end where
  re_refl[iff]: "reachable_end x x"
| re_succ: "reachable_end x y; edge y z  reachable_end x z"

lemma succ_re:
  assumes y: "edge x y" and z: "reachable_end y z"
  shows "reachable_end x z"
  using z y by (induction) (auto intro: re_succ)

lemma reachable_re:
  assumes "reachable x y"
  shows "reachable_end x y"
  using assms by (induction) (auto intro: succ_re)

lemma re_reachable:
  assumes "reachable_end x y"
  shows "reachable x y"
  using assms by (induction) (auto intro: succ_reachable)

lemma reachable_end_induct:
  assumes r: "reachable x y"
      and base: "x. P x x"
      and step: "x y z. P x y; edge y z  P x z"
  shows "P x y"
using r[THEN reachable_re] proof (induction)
  case (re_refl x)
  from base show ?case .
next
  case (re_succ x y z)
  with step show ?case by blast
qed

text ‹
  We also need the following variant of reachability avoiding
  certain edges. More precisely, @{text y} is reachable from @{text x}
  avoiding a set @{text E} of edges if there exists a path such that
  no edge from @{text E} appears along the path.
›
inductive reachable_avoiding where
  ra_refl[iff]: "reachable_avoiding x x E"
| ra_succ[elim]: "reachable_avoiding x y E; edge y z; (y,z)  E  reachable_avoiding x z E"

lemma edge_ra:
  assumes "edge x y" and "(x,y)  E"
  shows "reachable_avoiding x y E"
  using assms by (meson reachable_avoiding.simps)

lemma ra_trans:
  assumes 1: "reachable_avoiding x y E" and 2: "reachable_avoiding y z E"
  shows "reachable_avoiding x z E"
  using 2 1 by induction auto

lemma ra_cases:
  assumes "reachable_avoiding x y E"
  shows "x=y  (z. z  successors x  (x,z)  E  reachable_avoiding z y E)"
using assms proof (induction)
  case (ra_refl x S)
  then show ?case by simp
next
  case (ra_succ x y S z)
  then show ?case
    by (metis ra_refl reachable_avoiding.ra_succ)
qed

lemma ra_mono: 
  assumes "reachable_avoiding x y E" and "E'  E"
  shows "reachable_avoiding x y E'"
using assms by induction auto

lemma ra_add_edge:
  assumes "reachable_avoiding x y E"
  shows "reachable_avoiding x y (E  {(v,w)})
          (reachable_avoiding x v (E  {(v,w)})  reachable_avoiding w y (E  {(v,w)}))"
using assms proof (induction)
  case (ra_refl x E)
  then show ?case by simp
next
  case (ra_succ x y E z)
  then show ?case
    using reachable_avoiding.ra_succ by auto
qed


text ‹
  Reachability avoiding some edges obviously implies reachability.
  Conversely, reachability implies reachability avoiding the empty set.
›
lemma ra_reachable:
  "reachable_avoiding x y E  reachable x y"
  by (induction rule: reachable_avoiding.induct) (auto intro: succ_reachable)

lemma ra_empty:
  "reachable_avoiding x y {} = reachable x y"
proof
  assume "reachable_avoiding x y {}"
  thus "reachable x y"
    by (rule ra_reachable)
next
  assume "reachable x y"
  hence "reachable_end x y"
    by (rule reachable_re)
  thus "reachable_avoiding x y {}"
    by induction auto
qed


section ‹Strongly connected components›

text ‹
  A strongly connected component is a set @{text S} of nodes
  such that any two nodes in @{text S} are reachable from each other.
  This concept is represented by the predicate @{text "is_subscc"} below.
  We are ultimately interested in non-empty, maximal strongly connected
  components, represented by the predicate @{text "is_scc"}.
›

definition is_subscc where
  "is_subscc S  x  S. y  S. reachable x y"

definition is_scc where
  "is_scc S  S  {}  is_subscc S  (S'. S  S'  is_subscc S'  S' = S)"

lemma subscc_add:
  assumes "is_subscc S" and "x  S"
      and "reachable x y" and "reachable y x"
  shows "is_subscc (insert y S)"
using assms unfolding is_subscc_def by (metis insert_iff reachable_trans)

lemma sccE:
  ― ‹Two nodes that are reachable from each other are in the same SCC.›
  assumes "is_scc S" and "x  S"
      and "reachable x y" and "reachable y x"
  shows "y  S"
  using assms unfolding is_scc_def 
  by (metis insertI1 subscc_add subset_insertI)

lemma scc_partition:
  ― ‹Two SCCs that contain a common element are identical.›
  assumes "is_scc S" and "is_scc S'" and "x  S  S'"
  shows "S = S'"
  using assms unfolding is_scc_def is_subscc_def
  by (metis IntE assms(2) sccE subsetI)


section ‹Algorithm for computing strongly connected components›

text ‹
  We now introduce our representation of Bloemen's algorithm in Isabelle/HOL.
  The auxiliary function @{text unite} corresponds to the inner \textsf{while}
  loop in Bloemen's pseudo-code~cite‹p.32› in "bloemen:strong". It is applied to
  two nodes @{text v} and @{text w} (and the environment @{text e} holding the
  current values of the program variables) when a loop is found, i.e.\ when
  @{text w} is a successor of @{text v} in the graph that has already been
  visited in the depth-first search. In that case, the root of the SCC
  of node @{text w} determined so far must appear below the root of
  @{text v}'s SCC in the @{text stack} maintained by the algorithm. 
  The effect of the function is to merge the SCCs of all nodes on the
  top of the stack above (and including) @{text w}. Node @{text w}'s root
  will be the root of the merged SCC.
›

definition unite :: "'v  'v  'v env  'v env" where
  "unite v w e 
     let pfx = takeWhile (λx. w  𝒮 e x) (stack e);
         sfx = dropWhile (λx. w  𝒮 e x) (stack e);
         cc =  { 𝒮 e x | x . x  set pfx  {hd sfx} }
     in  e𝒮 := λx. if x  cc then cc else 𝒮 e x,
           stack := sfx"

text ‹
  We now represent the algorithm as two mutually recursive functions @{text dfs} and
  @{text dfss} in Isabelle/HOL. The function @{text dfs} corresponds to Bloemen's
  function \textsf{SetBased}, whereas @{text dfss} corresponds to the \textsf{forall}
  loop over the successors of the node on which @{text dfs} was called. Instead of
  using global program variables in imperative style, our functions explicitly pass
  environments that hold the current values of these variables.

  A technical complication in the development of the algorithm in Isabelle is the
  fact that the functions need not terminate when their pre-conditions (introduced
  below) are violated, for example when @{text dfs} is called for a node that was
  already visited previously. We therefore cannot prove termination at this point, 
  but will later show that the explicitly given pre-conditions ensure termination.
›

function (domintros) dfs :: "'v  'v env  'v env" 
and dfss :: "'v  'v env  'v env" where
  "dfs v e =
  (let e1 = evisited := visited e  {v}, 
              stack := (v # stack e), 
              cstack := (v # cstack e);
       e' = dfss v e1
   in if v = hd(stack e')
      then e'sccs := sccs e'  {𝒮 e' v}, 
              explored := explored e'  (𝒮 e' v), 
              stack := tl(stack e'), 
              cstack := tl(cstack e')
      else e'cstack := tl(cstack e'))"
| "dfss v e =
   (let vs = successors v - vsuccs e v
    in  if vs = {} then e
        else let w = SOME x. x  vs;
                 e' = (if w  explored e then e
                       else if w  visited e
                       then dfs w e
                       else unite v w e);
                 e'' = (e'vsuccs := 
                             (λx. if x=v then vsuccs e' v  {w} 
                                  else vsuccs e' x))
             in  dfss v e'')"
  by pat_completeness (force+)


section ‹Definition of the predicates used in the correctness proof›

text ‹
  Environments are partially ordered according to the following definition.
›
definition sub_env where
  "sub_env e e'  
     root e' = root e
    visited e  visited e'
    explored e  explored e'
    (v. vsuccs e v  vsuccs e' v)
    ( v. 𝒮 e v  𝒮 e' v)
    ( {𝒮 e v | v . v  set (stack e)})
      ( {𝒮 e' v | v . v  set (stack e')})
"

lemma sub_env_trans:
  assumes "sub_env e e'" and "sub_env e' e''"
  shows "sub_env e e''"
  using assms unfolding sub_env_def
  by (metis (no_types, lifting) subset_trans)

text ‹
  The set @{text "unvisited e u"} contains all edges @{text "(a,b)"}
  such that node @{text a} is in the same SCC as
  node @{text u} and the edge has not yet been followed, in the
  sense represented by variable @{text vsuccs}.
›
definition unvisited where
  "unvisited e u  
   {(a,b) | a b. a  𝒮 e u  b  successors a - vsuccs e a}"

subsection ‹Main invariant›

text ‹
  The following definition characterizes well-formed environments.
  This predicate will be shown to hold throughout the execution
  of the algorithm. In words, it asserts the following facts:
  \begin{itemize}
  \item Only nodes reachable from the root (for which the algorithm
    was originally called) are visited.
  \item The two stacks @{text stack} and @{text cstack} do not
    contain duplicate nodes, and @{text stack} contains a subset
    of the nodes on @{text cstack}, in the same order.
  \item Any node higher on the @{text stack} (i.e., that was pushed
    later) is reachable from nodes lower in the @{text stack}.
    This property also holds for nodes on the call stack,
    but this is not needed for the correctness proof.
  \item Every explored node, and every node on the call stack,
    has been visited.
  \item Nodes reachable from fully explored nodes have 
    themselves been fully explored.
  \item The set @{text "vsuccs e n"}, for any node @{text n},
    is a subset of @{text n}'s successors, and all these nodes
    are in @{text visited}. The set is empty if @{text "n ∉ visited"},
    and it contains all successors if @{text n} has been fully 
    explored or if @{text n} has been visited, but is no longer
    on the call stack.
  \item The sets @{text "𝒮 e n"} represent an equivalence relation.
    The equivalence classes of nodes that have not yet been visited 
    are singletons. Also, equivalence classes for two distinct nodes
    on the @{text stack} are disjoint because the stack only stores
    roots of SCCs, and the union of the equivalence classes for these
    root nodes corresponds to the set of live nodes, i.e. those nodes
    that have already been visited but not yet fully explored.
  \item More precisely, an equivalence class is represented on the
    stack by the oldest node in the sense of the call order: any
    node in the class that is still on the call stack precedes the
    representative on the call stack and was therefore pushed later.
  \item Equivalence classes represent the maximal available
    information about strong connectedness: nodes represented by
    some node @{text n} on the @{text stack} can reach some node
    @{text m} that is lower in the stack only by taking an
    edge from some node in @{text n}'s equivalence class that 
    has not yet been followed. (Remember that @{text m} can reach
    @{text n} by one of the previous conjuncts.)
  \item Equivalence classes represent partial SCCs in the sense
    of the predicate @{text is_subscc}. Variable @{text sccs}
    holds maximal SCCs in the sense of the predicate @{text is_scc},
    and their union corresponds to the set of explored nodes.
  \end{itemize}
›
definition wf_env where
  "wf_env e 
    (n  visited e. reachable (root e) n)
   distinct (stack e)
   distinct (cstack e)
   (n m. n  m in stack e  n  m in cstack e)
   (n m. n  m in stack e  reachable m n)
   explored e  visited e
   set (cstack e)  visited e
   (n  explored e. m. reachable n m  m  explored e)  
   (n. vsuccs e n  successors n  visited e)
   (n. n  visited e  vsuccs e n = {})
   (n  explored e. vsuccs e n = successors n)
   (n  visited e - set (cstack e). vsuccs e n = successors n)
   (n m. m  𝒮 e n  (𝒮 e n = 𝒮 e m))
   (n. n  visited e  𝒮 e n = {n})
   (n  set (stack e). m  set (stack e). n  m  𝒮 e n  𝒮 e m = {})
    {𝒮 e n | n. n  set (stack e)} = visited e - explored e
   (n  set (stack e). m  𝒮 e n. m  set (cstack e)  m  n in cstack e)
   (n m. n  m in stack e  n  m 
        (u  𝒮 e n. ¬ reachable_avoiding u m (unvisited e n)))
   (n. is_subscc (𝒮 e n))
   (S  sccs e. is_scc S)
    (sccs e) = explored e"

subsection ‹Consequences of the invariant›

text ‹
  Since every node on the call stack is an element
  of @{text visited} and every node on the @{text stack}
  also appears on @{text cstack}, all these nodes are
  also in @{text visited}.
›
lemma stack_visited:
  assumes "wf_env e" "n  set (stack e)"
  shows "n  visited e"
  using assms unfolding wf_env_def
  by (meson precedes_refl subset_iff)

text ‹
  Classes represented on the stack consist of visited nodes
  that have not yet been fully explored.
›
lemma stack_class:
  assumes "wf_env e" "n  set (stack e)" "m  𝒮 e n"
  shows "m  visited e - explored e"
  using assms unfolding wf_env_def by blast

text ‹
  Conversely, every such node belongs to some class
  represented on the stack.
›
lemma visited_unexplored:
  assumes "wf_env e" "m  visited e" "m  explored e"
  obtains n where "n  set (stack e)" "m  𝒮 e n"
  using assms unfolding wf_env_def
  by (smt (verit, ccfv_threshold) Diff_iff Union_iff mem_Collect_eq)

text ‹
  Every node belongs to its own equivalence class.
›
lemma S_reflexive:
  assumes "wf_env e"
  shows "n  𝒮 e n"
  using assms by (auto simp: wf_env_def)

text ‹
  No node on the stack has been fully explored.
›
lemma stack_unexplored:
  assumes 1: "wf_env e"
      and 2: "n  set (stack e)"
      and 3: "n  explored e"
  shows "P"
  using stack_class[OF 1 2] S_reflexive[OF 1] 3
  by blast

text ‹
  If @{text w} is reachable from visited node @{text v}, but
  no unvisited successor of a node reachable from @{text v}
  can reach @{text w}, then @{text w} must be visited.
›
lemma reachable_visited:
  assumes e: "wf_env e"
      and v: "v  visited e"
      and w: "reachable v w"
      and s: "n  visited e. m  successors n - vsuccs e n. 
                 reachable v n  ¬ reachable m w"
  shows "w  visited e"
using w v s proof (induction)
  case (reachable_refl x)
  then show ?case by simp
next
  case (reachable_succ x y z)
  then have "y  vsuccs e x" by blast
  with e have "y  visited e" 
    unfolding wf_env_def by (meson le_infE subset_eq)
  with reachable_succ reachable.reachable_succ show ?case
    by blast
qed

text ‹
  Edges towards explored nodes do not contribute to reachability
  of unexplored nodes avoiding some set of edges.
›
lemma avoiding_explored:
  assumes e: "wf_env e"
      and xy: "reachable_avoiding x y E"
      and y: "y  explored e"
      and w: "w  explored e"
    shows "reachable_avoiding x y (E  {(v,w)})"
using xy y proof (induction)
  case (ra_refl x E)
  then show ?case by simp
next
  case (ra_succ x y E z)
  from e z  successors y z  explored e
  have "y  explored e"
    unfolding wf_env_def by (meson reachable_edge)
  with ra_succ.IH have "reachable_avoiding x y (E  {(v,w)})" .
  moreover
  from w (y,z)  E z  explored e have "(y,z)  E  {(v,w)}"
    by auto
  ultimately show ?case 
    using  z  successors y by auto
qed

subsection ‹Pre- and post-conditions of function @{text dfs}

text ‹
  Function @{text dfs} should be called for a well-formed
  environment and a node @{text v} that has not yet been
  visited and that is reachable from the root node, 
  as well as from all nodes in the stack. No outgoing edges
  from node @{text v} have yet been followed.
›

definition pre_dfs where 
  "pre_dfs v e  
     wf_env e
    v  visited e
    reachable (root e) v
    vsuccs e v = {}
    (n  set (stack e). reachable n v)"

text ‹
  Function @{text dfs} maintains the invariant
  @{text wf_env} and returns an environment @{text e'} that
  extends the input environment @{text e}. Node @{text v} has been 
  visited and all its outgoing edges have been followed.
  Because the algorithm works in depth-first fashion, no 
  new outgoing edges of nodes that had already been
  visited in the input environment have been followed, and
  the stack of @{text e'} is a suffix of the one of @{text e}
  such that @{text v} is still reachable from all nodes on the
  stack. The stack may have been shortened because SCCs
  represented at the top of the stack may have been
  merged. The call stack is reestablished as it was in @{text e}.
  There are two possible outcomes of the algorithm:
  \begin{itemize}
  \item Either @{text v} has been fully explored, in which case
    the stacks of @{text e} and @{text e'} are the same, and
    the equivalence classes of all nodes represented on the
    stack are unchanged. This corresponds to the case where
    @{text v} is the root node of its (maximal) SCC.
  \item Alternatively, the stack of @{text e'} must be
    non-empty and @{text v} must be represented by the node at
    the top of the stack. The SCCs of the nodes
    lower on the stack are unchanged. This corresponds to the
    case where @{text v} is not the root node of its SCC, but
    some SCCs at the top of the stack may have been merged.
  \end{itemize}
›
definition post_dfs where 
  "post_dfs v e e'  
     wf_env e'
    v  visited e'
    sub_env e e'
    vsuccs e' v = successors v
    (w  visited e. vsuccs e' w = vsuccs e w)
    (n  set (stack e'). reachable n v)
    (ns. stack e = ns @ (stack e'))
    (  (v  explored e'  stack e' = stack e 
          (n  set (stack e'). 𝒮 e' n = 𝒮 e n)) 
      (stack e'  []  v  𝒮 e' (hd (stack e')) 
          (n  set (tl (stack e')). 𝒮 e' n = 𝒮 e n)))
    cstack e' = cstack e"

text ‹
  The initial environment is easily seen to satisfy @{text dfs}'s
  pre-condition.
›
lemma init_env_pre_dfs: "pre_dfs v (init_env v)"
  by (auto simp: pre_dfs_def wf_env_def init_env_def is_subscc_def 
           dest: precedes_mem)

text ‹
  Any node represented by the top stack element of the
  input environment is still represented by the top
  element of the output stack.
›
lemma dfs_S_hd_stack:
  assumes wf: "wf_env e"
      and post: "post_dfs v e e'"
      and n: "stack e  []" "n  𝒮 e (hd (stack e))"
    shows "stack e'  []" "n  𝒮 e' (hd (stack e'))"
proof -
  have 1: "stack e'  []  n  𝒮 e' (hd (stack e'))"
  proof (cases "stack e' = stack e  (n  set (stack e'). 𝒮 e' n = 𝒮 e n)")
    case True
    with n show ?thesis 
      by auto
  next
    case 2: False
    with post have "stack e'  []"
      by (simp add: post_dfs_def)
    from n have "hd (stack e)  set (stack e)"
      by simp
    with 2 n post obtain u where
      u: "u  set (stack e')" "n  𝒮 e' u"
      unfolding post_dfs_def sub_env_def by blast
    show ?thesis
    proof (cases "u = hd (stack e')")
      case True
      with u stack e'  [] show ?thesis
        by simp
    next
      case False
      with u have "u  set (tl (stack e'))"
        by (metis empty_set equals0D list.collapse set_ConsD)
      with u 2 post have "u  set (tl (stack e))  n  𝒮 e u"
        unfolding post_dfs_def
        by (metis Un_iff append_self_conv2 set_append tl_append2)
      with n wf hd (stack e)  set (stack e) show ?thesis
        unfolding wf_env_def
        by (metis (no_types, opaque_lifting) disjoint_iff_not_equal distinct.simps(2) list.collapse list.set_sel(2))
    qed
  qed
  from 1 show "stack e'  []" by simp
  from 1 show "n  𝒮 e' (hd (stack e'))" by simp
qed

text ‹
  Function @{text dfs} leaves the SCCs represented
  by elements in the (new) tail of the @{text stack} unchanged.
›
lemma dfs_S_tl_stack:
  assumes post: "post_dfs v e e'"
    and nempty: "stack e  []"
  shows "stack e'  []" "n  set (tl (stack e')). 𝒮 e' n = 𝒮 e n"
proof -
  have 1: "stack e'  []  (n  set (tl (stack e')). 𝒮 e' n = 𝒮 e n)"
  proof (cases "stack e' = stack e  (n  set (stack e'). 𝒮 e' n = 𝒮 e n)")
    case True
    with nempty show ?thesis
      by (simp add: list.set_sel(2))
  next
    case False
    with post show ?thesis
      by (auto simp: post_dfs_def)
  qed
  from 1 show "stack e'  []"
    by simp
  from 1 show "n  set (tl (stack e')). 𝒮 e' n = 𝒮 e n"
    by simp
qed

subsection ‹Pre- and post-conditions of function @{text dfss}

text ‹
  The pre- and post-conditions of function @{text dfss}
  correspond to the invariant of the loop over all outgoing
  edges from node @{text v}. The environment must be
  well-formed, node @{text v} must be visited and represented
  by the top element of the (non-empty) stack. Node @{text v}
  must be reachable from all nodes on the stack, and it must be
  the top node on the call stack. All outgoing
  edges of node @{text v} that have already been followed must
  either lead to completely explored nodes (that are no longer
  represented on the stack) or to nodes that are part of the
  same SCC as @{text v}.
›
definition pre_dfss where 
  "pre_dfss v e  
     wf_env e 
    v  visited e
    (stack e  [])
    (v  𝒮 e (hd (stack e)))
    (w  vsuccs e v. w  explored e  𝒮 e (hd (stack e)))
    (n  set (stack e). reachable n v)
    (ns. cstack e = v # ns)"

text ‹
  The post-condition establishes that all outgoing edges
  of node @{text v} have been followed. As for function
  @{text dfs}, no new outgoing edges of previously visited
  nodes have been followed. Also as before, the new stack
  is a suffix of the old one, and the call stack is restored.
  In case node @{text v} is still on the stack (and therefore
  is the root node of its SCC), no node that is lower on the stack
  can be reachable from @{text v}. This condition guarantees
  the maximality of the computed SCCs.
›
definition post_dfss where 
  "post_dfss v e e'  
     wf_env e'
    vsuccs e' v = successors v
    (w  visited e - {v}. vsuccs e' w = vsuccs e w)
    sub_env e e'
    (w  successors v. w  explored e'  𝒮 e' (hd (stack e')))
    (n  set (stack e'). reachable n v)
    (stack e'  [])
    (ns. stack e = ns @ (stack e'))
    v  𝒮 e' (hd (stack e'))
    (n  set (tl (stack e')). 𝒮 e' n = 𝒮 e n)
    (hd (stack e') = v  (n  set (tl (stack e')). ¬ reachable v n))
    cstack e' = cstack e"


section ‹Proof of partial correctness›

subsection ‹Lemmas about function @{text unite}

text ‹
  We start by establishing a few lemmas about function @{text unite}
  in the context where it is called.
›
lemma unite_stack:
  fixes e v w
  defines "e'  unite v w e"
  assumes wf: "wf_env e"
      and w: "w  successors v" "w  vsuccs e v" "w  visited e" "w  explored e"
  obtains pfx where "stack e = pfx @ (stack e')" 
                    "stack e'  []"
                    "let cc =  {𝒮 e n |n. n  set pfx  {hd (stack e')}}
                     in 𝒮 e' = (λx. if x  cc then cc else 𝒮 e x)"
                    "w  𝒮 e' (hd (stack e'))"
proof -
  define pfx where "pfx = takeWhile (λx. w  𝒮 e x) (stack e)"
  define sfx where "sfx = dropWhile (λx. w  𝒮 e x) (stack e)"
  define cc where "cc =  {𝒮 e x |x. x  set pfx  {hd sfx}}"

  have "stack e = pfx @ sfx"
    by (simp add: pfx_def sfx_def)
  moreover
  have "stack e' = sfx"
    by (simp add: e'_def unite_def sfx_def)
  moreover
  from wf w have "w   {𝒮 e n | n. n  set (stack e)}"
    by (simp add: wf_env_def)
  then obtain n where "n  set (stack e)" "w  𝒮 e n"
    by auto
  hence sfx: "sfx  []  w  𝒮 e (hd sfx)"
    unfolding sfx_def
    by (metis dropWhile_eq_Nil_conv hd_dropWhile)
  moreover
  have "𝒮 e' = (λx. if x  cc then cc else 𝒮 e x)"
    by (rule,
        auto simp add: e'_def unite_def pfx_def sfx_def cc_def)
  moreover
  from sfx have "w  cc"
    by (auto simp: cc_def)
  from S_reflexive[OF wf, of "hd sfx"]
  have "hd sfx  cc"
    by (auto simp: cc_def)
  with w  cc 𝒮 e' = (λx. if x  cc then cc else 𝒮 e x)
  have "w  𝒮 e' (hd sfx)"
    by simp
  ultimately show ?thesis
    using that e'_def unite_def pfx_def sfx_def cc_def
    by meson
qed

text ‹
  Function @{text unite} leaves intact the equivalence classes
  represented by the tail of the new stack.
›
lemma unite_S_tl:
  fixes e v w
  defines "e'  unite v w e"
  assumes wf: "wf_env e"
      and w: "w  successors v" "w  vsuccs e v" "w  visited e" "w  explored e"
      and n: "n  set (tl (stack e'))"
  shows "𝒮 e' n = 𝒮 e n"
proof -
  from assms obtain pfx where
    pfx: "stack e = pfx @ (stack e')" "stack e'  []"
         "let cc =  {𝒮 e n |n. n  set pfx  {hd (stack e')}}
          in 𝒮 e' = (λx. if x  cc then cc else 𝒮 e x)"
    by (blast dest: unite_stack)
  define cc where "cc   {𝒮 e n |n. n  set pfx  {hd (stack e')}}"

  have "n  cc"
  proof
    assume "n  cc"
    then obtain m where
      "m  set pfx  {hd (stack e')}" "n  𝒮 e m"
      by (auto simp: cc_def)
    with S_reflexive[OF wf, of n] n wf stack e = pfx @ stack e' stack e'  []
    show "False"
      unfolding wf_env_def
      by (smt (z3) Diff_triv Un_iff Un_insert_right append.right_neutral disjoint_insert(1) 
                   distinct.simps(2) distinct_append empty_set insertE insert_Diff list.exhaust_sel 
                   list.simps(15) set_append)
  qed
  with pfx show "𝒮 e' n = 𝒮 e n"
    by (auto simp add: cc_def)
qed

text ‹
  The stack of the result of @{text unite} represents the
  same vertices as the input stack, potentially in fewer
  equivalence classes.
›
lemma unite_S_equal:
  fixes e v w
  defines "e'  unite v w e"
  assumes wf: "wf_env e"
      and w: "w  successors v" "w  vsuccs e v" "w  visited e" "w  explored e"
  shows "( {𝒮 e' n | n. n  set (stack e')}) = ( {𝒮 e n | n. n  set (stack e)})"
proof -
  from assms obtain pfx where
    pfx: "stack e = pfx @ (stack e')" "stack e'  []"
         "let cc =  {𝒮 e n |n. n  set pfx  {hd (stack e')}}
          in 𝒮 e' = (λx. if x  cc then cc else 𝒮 e x)"
    by (blast dest: unite_stack)
  define cc where "cc   {𝒮 e n |n. n  set pfx  {hd (stack e')}}"

  from pfx have Se': "x. 𝒮 e' x = (if x  cc then cc else 𝒮 e x)"
    by (auto simp: cc_def)
  from S_reflexive[OF wf, of "hd (stack e')"]
  have S_hd: "𝒮 e' (hd (stack e')) = cc"
    by (auto simp: Se' cc_def)
  from stack e'  []
  have ste': "set (stack e') = {hd (stack e')}  set (tl (stack e'))"
    by (metis insert_is_Un list.exhaust_sel list.simps(15))

  from stack e = pfx @ stack e' stack e'  []
  have "stack e = pfx @ (hd (stack e') # tl (stack e'))"
    by auto
  hence " {𝒮 e n | n. n  set (stack e)}
        = cc  ( {𝒮 e n | n. n  set (tl (stack e'))})"
    by (auto simp add: cc_def)
  also from S_hd unite_S_tl[OF wf w]
  have " = 𝒮 e' (hd (stack e'))  ( {𝒮 e' n | n. n  set (tl (stack e'))})"
    by (auto simp: e'_def)
  also from ste'
  have " =  {𝒮 e' n | n. n  set (stack e')}"
    by auto
  finally show ?thesis
    by simp
qed

text ‹
  The head of the stack represents a (not necessarily maximal) SCC.
›
lemma unite_subscc:
  fixes e v w
  defines "e'  unite v w e"
  assumes pre: "pre_dfss v e"
      and w: "w  successors v" "w  vsuccs e v" "w  visited e" "w  explored e"
    shows "is_subscc (𝒮 e' (hd (stack e')))"
proof -
  from pre have wf: "wf_env e"
    by (simp add: pre_dfss_def)
  from assms obtain pfx where
    pfx: "stack e = pfx @ (stack e')" "stack e'  []"
         "let cc =  {𝒮 e n |n. n  set pfx  {hd (stack e')}}
          in 𝒮 e' = (λx. if x  cc then cc else 𝒮 e x)"
    by (blast dest: unite_stack[OF wf])

  define cc where "cc   {𝒮 e n |n. n  set pfx  {hd (stack e')}}"

  from wf w have "w   {𝒮 e n | n. n  set (stack e)}"
    by (simp add: wf_env_def)
  hence "w  𝒮 e (hd (stack e'))"
    apply (simp add: e'_def unite_def)
    by (metis dropWhile_eq_Nil_conv hd_dropWhile)

  have "is_subscc cc"
  proof (clarsimp simp: is_subscc_def)
    fix x y
    assume "x  cc" "y  cc"
    then obtain nx ny where
      nx: "nx  set pfx  {hd (stack e')}" "x  𝒮 e nx" and
      ny: "ny  set pfx  {hd (stack e')}" "y  𝒮 e ny"
      by (auto simp: cc_def)
    with wf have "reachable x nx" "reachable ny y"
      by (auto simp: wf_env_def is_subscc_def)
    from w pre have "reachable v w"
      by (auto simp: pre_dfss_def)
    from pre have "reachable (hd (stack e)) v"
      by (auto simp: pre_dfss_def wf_env_def is_subscc_def)
    from pre have "stack e = hd (stack e) # tl (stack e)"
      by (auto simp: pre_dfss_def)
    with nx stack e = pfx @ (stack e') stack e'  []
    have "hd (stack e)  nx in stack e"
      by (metis Un_iff Un_insert_right head_precedes list.exhaust_sel list.simps(15) 
                set_append sup_bot.right_neutral)
    with wf have "reachable nx (hd (stack e))"
      by (auto simp: wf_env_def)
    from stack e = pfx @ (stack e') stack e'  [] ny
    have "ny  hd (stack e') in stack e" 
      by (metis List.set_insert empty_set insert_Nil list.exhaust_sel set_append split_list_precedes)
    with wf have "reachable (hd (stack e')) ny"
      by (auto simp: wf_env_def is_subscc_def)
    from wf stack e'  [] w  𝒮 e (hd (stack e'))
    have "reachable w (hd (stack e'))"
      by (auto simp: wf_env_def is_subscc_def)

    from reachable x nx reachable nx (hd (stack e))
         reachable (hd (stack e)) v reachable v w
         reachable w (hd (stack e')) 
         reachable (hd (stack e')) ny reachable ny y
    show "reachable x y"
      using reachable_trans by meson
  qed
  with S_reflexive[OF wf, of "hd (stack e')"] pfx
  show ?thesis
    by (auto simp: cc_def)
qed

text ‹
  The environment returned by function @{text unite} extends the input environment.
›

lemma unite_sub_env:
  fixes e v w
  defines "e'  unite v w e"
  assumes pre: "pre_dfss v e"
        and w: "w  successors v" "w  vsuccs e v" "w  visited e" "w  explored e"
  shows "sub_env e e'"
proof -
  from pre have wf: "wf_env e"
    by (simp add: pre_dfss_def)
  from assms obtain pfx where
    pfx: "stack e = pfx @ (stack e')" "stack e'  []"
         "let cc =  {𝒮 e n |n. n  set pfx  {hd (stack e')}}
          in 𝒮 e' = (λx. if x  cc then cc else 𝒮 e x)"
    by (blast dest: unite_stack[OF wf])
  define cc where "cc   {𝒮 e n |n. n  set pfx  {hd (stack e')}}"
  have "n. 𝒮 e n  𝒮 e' n"
  proof (clarify)
    fix n u
    assume u: "u  𝒮 e n"
    show "u  𝒮 e' n"
    proof (cases "n  cc")
      case True
      then obtain m where
        m: "m  set pfx  {hd (stack e')}" "n  𝒮 e m"
        by (auto simp: cc_def)
      with wf S_reflexive[OF wf, of n] u have "u  𝒮 e m"
        by (auto simp: wf_env_def)
      with m pfx show ?thesis 
        by (auto simp: cc_def)
    next
      case False
      with pfx u show ?thesis
        by (auto simp: cc_def)
    qed
  qed
  moreover
  have "root e' = root e  visited e' = visited e
       explored e' = explored e  vsuccs e' = vsuccs e"
    by (simp add: e'_def unite_def)
  ultimately show ?thesis
    using unite_S_equal[OF wf w]
    by (simp add: e'_def sub_env_def)
qed

text ‹
  The environment returned by function @{text unite} is well-formed.
›
lemma unite_wf_env:
  fixes e v w
  defines "e'  unite v w e"
  assumes pre: "pre_dfss v e"
      and w: "w  successors v" "w  vsuccs e v" "w  visited e" "w  explored e"
  shows "wf_env e'"
proof -
  from pre have wf: "wf_env e"
    by (simp add: pre_dfss_def)
  from assms obtain pfx where
    pfx: "stack e = pfx @ (stack e')" "stack e'  []"
         "let cc =  {𝒮 e n |n. n  set pfx  {hd (stack e')}}
          in 𝒮 e' = (λx. if x  cc then cc else 𝒮 e x)"
    by (blast dest: unite_stack[OF wf])
  define cc where "cc   {𝒮 e n |n. n  set pfx  {hd (stack e')}}"

  from pfx have Se': "x. 𝒮 e' x = (if x  cc then cc else 𝒮 e x)"
    by (auto simp add: cc_def)

  have cc_Un: "cc =  {𝒮 e x | x. x  cc}"
  proof
    from S_reflexive[OF wf]
    show "cc   {𝒮 e x | x. x  cc}"
      by (auto simp: cc_def)
  next
    {
      fix n x
      assume "x  cc" "n  𝒮 e x"
      with wf have "n  cc"
        unfolding wf_env_def cc_def
        by (smt (verit) Union_iff mem_Collect_eq)
    }
    thus "( {𝒮 e x | x. x  cc})  cc"
      by blast
  qed

  from S_reflexive[OF wf, of "hd (stack e')"]
  have hd_cc: "𝒮 e' (hd (stack e')) = cc"
    by (auto simp: cc_def Se')

  {
    fix n m
    assume n: "n  set (tl (stack e'))"
       and m: "m  𝒮 e n  cc"
    from m obtain l where
      "l  set pfx  {hd (stack e')}" "m  𝒮 e l"
      by (auto simp: cc_def)
    with n m wf stack e = pfx @ stack e' stack e'  []
    have "False"
      unfolding wf_env_def
      by (metis (no_types, lifting) Int_iff UnCI UnE disjoint_insert(1) distinct.simps(2) 
                distinct_append emptyE hd_Cons_tl insert_iff list.set_sel(1) list.set_sel(2) 
                mk_disjoint_insert set_append)
  }
  hence tl_cc: "n  set (tl (stack e')). 𝒮 e n  cc = {}"
    by blast

  from wf 
  have "n  visited e'. reachable (root e') n"
       "distinct (cstack e')"
       "explored e'  visited e'"
       "set (cstack e')  visited e'"
       "n  explored e'. m. reachable n m  m  explored e'"
       "n. vsuccs e' n  successors n  visited e'"
       "n. n  visited e'  vsuccs e' n = {}"
       "n  explored e'. vsuccs e' n = successors n"
       "n  visited e' - set (cstack e'). vsuccs e' n = successors n"
       "S  sccs e'. is_scc S"
       " (sccs e') = explored e'"
    by (auto simp: wf_env_def e'_def unite_def)

  moreover
  from wf stack e = pfx @ stack e'
  have "distinct (stack e')"
    by (auto simp: wf_env_def)

  moreover
  have "n m. n  m in stack e'  n  m in cstack e'"
  proof (clarify)
    fix n m
    assume "n  m in stack e'"
    with stack e = pfx @ stack e' wf
    have "n  m in cstack e"
      unfolding wf_env_def
      by (metis precedes_append_left)
    thus "n  m in cstack e'"
      by (simp add: e'_def unite_def)
  qed

  moreover
  from wf stack e = pfx @ stack e'
  have "n m. n  m in stack e'  reachable m n"
    unfolding wf_env_def by (metis precedes_append_left)

  moreover
  have "n m. m  𝒮 e' n  (𝒮 e' n = 𝒮 e' m)"
  proof (clarify)
    fix n m
    show "m  𝒮 e' n  (𝒮 e' n = 𝒮 e' m)"
    proof
      assume l: "m  𝒮 e' n"
      show "𝒮 e' n = 𝒮 e' m"
      proof (cases "n  cc")
        case True
        with l show ?thesis
          by (simp add: Se')
      next
        case False
        with l wf have "𝒮 e n = 𝒮 e m"
          by (simp add: wf_env_def Se')
        with False cc_Un wf have "m  cc"
          unfolding wf_env_def e'_def
          by (smt (verit, best) Union_iff mem_Collect_eq)
        with 𝒮 e n = 𝒮 e m False show ?thesis
          by (simp add: Se')
      qed
    next
      assume r: "𝒮 e' n = 𝒮 e' m"
      show "m  𝒮 e' n"
      proof (cases "n  cc")
        case True
        with r pfx have "𝒮 e' m = cc"
          by (auto simp: cc_def)
        have "m  cc"
        proof (rule ccontr)
          assume "m  cc"
          with pfx have "𝒮 e' m = 𝒮 e m"
            by (auto simp: cc_def)
          with S_reflexive[OF wf, of m] 𝒮 e' m = cc m  cc
          show "False"
            by simp
        qed
        with pfx True show "m  𝒮 e' n"
          by (auto simp: cc_def)
      next
        case False
        hence "𝒮 e' n = 𝒮 e n"
          by (simp add: Se')
        have "m  cc"
        proof
          assume m: "m  cc"
          with 𝒮 e' n = 𝒮 e n r have "𝒮 e n = cc"
            by (simp add: Se')
          with S_reflexive[OF wf, of n] have "n  cc"
            by simp
          with n  cc show "False" ..
        qed
        with r 𝒮 e' n = 𝒮 e n have "𝒮 e m = 𝒮 e n"
          by (simp add: Se')
        with S_reflexive[OF wf, of m] have "m  𝒮 e n"
          by simp
        with 𝒮 e' n = 𝒮 e n show ?thesis
          by simp
      qed
    qed
  qed

  moreover
  have "n. n  visited e'  𝒮 e' n = {n}"
  proof (clarify)
    fix n
    assume "n  visited e'"
    hence "n  visited e"
      by (simp add: e'_def unite_def)
    moreover have "n  cc"
    proof
      assume "n  cc"
      then obtain m where "m  set pfx  {hd (stack e')}" "n  𝒮 e m"
        by (auto simp: cc_def)
      with stack e = pfx @ stack e' stack e'  []
      have "m  set (stack e)"
        by auto
      with stack_class[OF wf this n  𝒮 e m] n  visited e
      show "False"
        by simp
    qed
    ultimately show "𝒮 e' n = {n}"
      using wf by (auto simp: wf_env_def Se')
  qed

  moreover
  have "n  set (stack e'). m  set (stack e'). n  m  𝒮 e' n  𝒮 e' m = {}"
  proof (clarify)
    fix n m
    assume "n  set (stack e')" "m  set (stack e')" "n  m"
    show "𝒮 e' n  𝒮 e' m = {}"
    proof (cases "n = hd (stack e')")
      case True
      with m  set (stack e') n  m stack e'  []
      have "m  set (tl (stack e'))"
        by (metis hd_Cons_tl set_ConsD)
      with True hd_cc tl_cc unite_S_tl[OF wf w]
      show ?thesis
        by (auto simp: e'_def)
    next
      case False
      with n  set (stack e') stack e'  []
      have "n  set (tl (stack e'))"
        by (metis hd_Cons_tl set_ConsD)
      show ?thesis
      proof (cases "m = hd (stack e')")
        case True
        with n  set (tl (stack e')) hd_cc tl_cc unite_S_tl[OF wf w]
        show ?thesis
          by (auto simp: e'_def)
      next
        case False
        with m  set (stack e') stack e'  []
        have "m  set (tl (stack e'))"
          by (metis hd_Cons_tl set_ConsD)
        with n  set (tl (stack e'))
        have "𝒮 e' m = 𝒮 e m  𝒮 e' n = 𝒮 e n"
          by (auto simp: e'_def unite_S_tl[OF wf w])
        moreover
        from m  set (stack e') n  set (stack e') stack e = pfx @ stack e'
        have "m  set (stack e)  n  set (stack e)"
          by auto
        ultimately show ?thesis
          using wf n  m by (auto simp: wf_env_def)
      qed
    qed
  qed

  moreover
  {
    from unite_S_equal[OF wf w]
    have " {𝒮 e' n | n. n  set (stack e')} =  {𝒮 e n | n. n  set (stack e)}"
      by (simp add: e'_def)
    with wf
    have " {𝒮 e' n | n. n  set (stack e')} = visited e - explored e"
      by (simp add: wf_env_def)
  }
  hence " {𝒮 e' n | n. n  set (stack e')} = visited e' - explored e'"
    by (simp add: e'_def unite_def)

  moreover
  have "n  set (stack e'). m  𝒮 e' n. 
           m  set (cstack e')  m  n in cstack e'"
  proof (clarify)
    fix n m
    assume "n  set (stack e')" "m  𝒮 e' n" "m  set (cstack e')"
    from m  set (cstack e') have "m  set (cstack e)"
      by (simp add: e'_def unite_def)
    have "m  n in cstack e"
    proof (cases "n = hd (stack e')")
      case True
      with m  𝒮 e' n have "m  cc"
        by (simp add: hd_cc)
      then obtain l where 
        "l  set pfx  {hd (stack e')}" "m  𝒮 e l"
        by (auto simp: cc_def)
      with stack e = pfx @ stack e' stack e'  []
      have "l  set (stack e)"
        by auto
      with m  𝒮 e l m  set (cstack e) wf
      have "m  l in cstack e"
        by (auto simp: wf_env_def)
      moreover
      from l  set pfx  {hd (stack e')} True 
           stack e = pfx @ stack e' stack e'  []
      have "l  n in stack e"
        by (metis List.set_insert empty_set hd_Cons_tl insert_Nil set_append split_list_precedes)
      with wf have "l  n in cstack e"
        by (auto simp: wf_env_def)
      ultimately show ?thesis
        using wf unfolding wf_env_def
        by (meson precedes_trans)
    next
      case False
      with n  set (stack e') stack e'  []
      have "n  set (tl (stack e'))"
        by (metis list.collapse set_ConsD)
      with unite_S_tl[OF wf w] m  𝒮 e' n
      have "m  𝒮 e n"
        by (simp add: e'_def)
      with n  set (stack e') stack e = pfx @ stack e'
           m  set (cstack e) wf
      show ?thesis
        by (auto simp: wf_env_def)
    qed
    thus "m  n in cstack e'"
      by (simp add: e'_def unite_def)
  qed

  moreover
  have "n m. n  m in stack e'  n  m 
        (u  𝒮 e' n. ¬ reachable_avoiding u m (unvisited e' n))"
  proof (clarify)
    fix x y u
    assume xy: "x  y in stack e'" "x  y"
       and u: "u  𝒮 e' x" "reachable_avoiding u y (unvisited e' x)"
    show "False"
    proof (cases "x = hd (stack e')")
      case True
      hence "𝒮 e' x = cc" 
        by (simp add: hd_cc)
      with u  𝒮 e' x obtain x' where
        x': "x'  set pfx  {hd (stack e')}" "u  𝒮 e x'"
        by (auto simp: cc_def)
      from stack e = pfx @ stack e' stack e'  []
      have "stack e = pfx @ (hd (stack e') # tl (stack e'))"
        by auto
      with x' True have "x'  x in stack e"
        by (simp add: split_list_precedes)
      moreover
      from xy stack e = pfx @ stack e' have "x  y in stack e"
        by (simp add: precedes_append_left)
      ultimately have "x'  y in stack e"
        using wf by (auto simp: wf_env_def elim: precedes_trans)
      from x'  x in stack e x  y in stack e wf x  y
      have "x'  y"
        by (auto simp: wf_env_def dest: precedes_antisym)
      let ?unv = " {unvisited e y | y. y  set pfx  {hd (stack e')}}"
      from 𝒮 e' x = cc have "?unv = unvisited e' x"
        by (auto simp: unvisited_def cc_def e'_def unite_def)
      with reachable_avoiding u y (unvisited e' x)
      have "reachable_avoiding u y ?unv"
        by simp
      with x' have "reachable_avoiding u y (unvisited e x')"
        by (blast intro: ra_mono)
      with x'  y in stack e x'  y u  𝒮 e x' wf
      show ?thesis
        by (auto simp: wf_env_def)
    next
      case False
      with x  y in stack e' stack e'  []
      have "x  set (tl (stack e'))"
        by (metis list.exhaust_sel precedes_mem(1) set_ConsD)
      with u  𝒮 e' x have "u  𝒮 e x"
        by (auto simp add: unite_S_tl[OF wf w] e'_def)
      moreover
      from x  y in stack e' stack e = pfx @ stack e'
      have "x  y in stack e"
        by (simp add: precedes_append_left)
      moreover
      from unite_S_tl[OF wf w] x  set (tl (stack e'))
      have "unvisited e' x = unvisited e x"
        by (auto simp: unvisited_def e'_def unite_def)
      ultimately show ?thesis
        using x  y reachable_avoiding u y (unvisited e' x) wf
        by (auto simp: wf_env_def)
    qed
  qed


  moreover
  have "n. is_subscc (𝒮 e' n)"
  proof
    fix n
    show "is_subscc (𝒮 e' n)"
    proof (cases "n  cc")
      case True
      hence "𝒮 e' n = cc"
        by (simp add: Se')
      with unite_subscc[OF pre w] hd_cc
      show ?thesis
        by (auto simp: e'_def)
    next
      case False
      with wf show ?thesis
        by (simp add: Se' wf_env_def)
    qed
  qed

  ultimately show ?thesis
    unfolding wf_env_def by blast
qed

subsection ‹Lemmas establishing the pre-conditions›

text ‹
  The precondition of function @{text dfs} ensures the precondition
  of @{text dfss} at the call of that function.
›
lemma pre_dfs_pre_dfss:
  assumes "pre_dfs v e"
  shows "pre_dfss v (evisited := visited e  {v}, 
                       stack := v # stack e, 
                       cstack := v # cstack e)"
        (is "pre_dfss v ?e'")
proof -
  from assms have wf: "wf_env e"
    by (simp add: pre_dfs_def)

  from assms have v: "v  visited e"
    by (simp add: pre_dfs_def)

  from assms stack_visited[OF wf]
  have "n  visited ?e'. reachable (root ?e') n"
       "distinct (stack ?e')"
       "distinct (cstack ?e')"
       "explored ?e'  visited ?e'"
       "set (cstack ?e')  visited ?e'"
       "n  explored ?e'. m. reachable n m  m  explored ?e'"
       "n. vsuccs ?e' n  successors n"
       "n  explored ?e'. vsuccs ?e' n = successors n"
       "n  visited ?e' - set(cstack ?e'). vsuccs ?e' n = successors n"
       "n. n  visited ?e'  vsuccs ?e' n = {}"
       "(n m. m  𝒮 ?e' n  (𝒮 ?e' n = 𝒮 ?e' m))"
       "(n. n  visited ?e'  𝒮 ?e' n = {n})"
       "n. is_subscc (𝒮 ?e' n)"
       "S  sccs ?e'. is_scc S"
       " (sccs ?e') = explored ?e'"
    by (auto simp: pre_dfs_def wf_env_def)

  moreover
  have "n m. n  m in stack ?e'  reachable m n"
  proof (clarify)
    fix x y
    assume "x  y in stack ?e'"
    show "reachable y x"
    proof (cases "x=v")
      assume "x=v"
      with x  y in stack ?e' assms show ?thesis
        apply (simp add: pre_dfs_def)
        by (metis insert_iff list.simps(15) precedes_mem(2) reachable_refl)
    next
      assume "x  v"
      with x  y in stack ?e' wf show ?thesis
        by (simp add: pre_dfs_def wf_env_def precedes_in_tail)
    qed
  qed

  moreover
  from wf v have "n. vsuccs ?e' n  visited ?e'"
    by (auto simp: wf_env_def)

  moreover
  from wf v 
  have "(n  set (stack ?e').  m  set (stack ?e'). n  m  𝒮 ?e' n  𝒮 ?e' m = {})"
    apply (simp add: wf_env_def)
    by (metis singletonD)

  moreover
  have " {𝒮 ?e' v | v . v  set (stack ?e')} = visited ?e' - explored ?e'"
  proof -
    have " {𝒮 ?e' v | v . v  set (stack ?e')} = 
          ( {𝒮 e v | v . v  set (stack e)})  𝒮 e v"
      by auto
    also from wf v have " = visited ?e' - explored ?e'"
      by (auto simp: wf_env_def)
    finally show ?thesis .
  qed

  moreover
  have "n m. n  m in stack ?e'  n  m 
           (u  𝒮 ?e' n. ¬ reachable_avoiding u m (unvisited ?e' n))"
  proof (clarify)
    fix x y u
    assume asm: "x  y in stack ?e'" "x  y" "u  𝒮 ?e' x"
                "reachable_avoiding u y (unvisited ?e' x)"
    show "False"
    proof (cases "x = v")
      case True
      with wf v u  𝒮 ?e' x have "u = v" "vsuccs ?e' v = {}"
        by (auto simp: wf_env_def)
      with reachable_avoiding u y (unvisited ?e' x)[THEN ra_cases]
           True x  y wf
      show ?thesis
        by (auto simp: wf_env_def unvisited_def)
    next
      case False
      with asm wf show ?thesis
        by (auto simp: precedes_in_tail wf_env_def unvisited_def)
    qed
  qed

  moreover 
  have "n m. n  m in stack ?e'  n  m in cstack ?e'"
  proof (clarsimp)
    fix n m
    assume "n  m in (v # stack e)"
    with assms show "n  m in (v # cstack e)"
      unfolding pre_dfs_def wf_env_def
      by (metis head_precedes insertI1 list.simps(15) precedes_in_tail precedes_mem(2) precedes_refl)
  qed

  moreover 
  have "n  set (stack ?e'). m  𝒮 ?e' n. m  set (cstack ?e')  m  n in cstack ?e'"
  proof (clarify)
    fix n m
    assume "n  set (stack ?e')" "m  𝒮 ?e' n" "m  set (cstack ?e')"
    show "m  n in cstack ?e'"
    proof (cases "n = v")
      case True
      with wf v m  𝒮 ?e' n show ?thesis
        by (auto simp: wf_env_def)
    next
      case False
      with n  set (stack ?e') m  𝒮 ?e' n
      have "n  set (stack e)" "m  𝒮 e n"
        by auto
      with wf v False m  𝒮 e n m  set (cstack ?e')
      show ?thesis
        apply (simp add: wf_env_def)
        by (metis (mono_tags, lifting) precedes_in_tail singletonD)
    qed
  qed

  ultimately have "wf_env ?e'"
    unfolding wf_env_def by (meson le_inf_iff)

  moreover
  from assms 
  have "w  vsuccs ?e' v. w  explored ?e'  𝒮 ?e' (hd (stack ?e'))"
    by (simp add: pre_dfs_def)

  moreover
  from n m. n  m in stack ?e'  reachable m n
  have "n  set (stack ?e'). reachable n v"
    by (simp add: head_precedes)

  moreover
  from wf v have "𝒮 ?e' (hd (stack ?e')) = {v}"
    by (simp add: pre_dfs_def wf_env_def)

  ultimately show ?thesis
    by (auto simp: pre_dfss_def)
qed

text ‹
  Similarly, we now show that the pre-conditions of the different
  function calls in the body of function @{text dfss} are satisfied.
  First, it is very easy to see that the pre-condition of @{text dfs}
  holds at the call of that function.
›
lemma pre_dfss_pre_dfs:
  assumes "pre_dfss v e" and "w  visited e" and "w  successors v"
  shows "pre_dfs w e" 
  using assms unfolding pre_dfss_def pre_dfs_def wf_env_def
  by (meson succ_reachable)

text ‹
  The pre-condition of @{text dfss} holds when the successor
  considered in the current iteration has already been explored.
›
lemma pre_dfss_explored_pre_dfss:
  fixes e v w
  defines "e''  evsuccs := (λx. if x=v then vsuccs e v  {w} else vsuccs e x)"
  assumes 1: "pre_dfss v e" and 2: "w  successors v" and 3: "w  explored e"
  shows "pre_dfss v e''"
proof -
  from 1 have v: "v  visited e"
    by (simp add: pre_dfss_def)
  have "wf_env e''"
  proof -
    from 1 have wf: "wf_env e"
      by (simp add: pre_dfss_def)
    hence "v  visited e''. reachable (root e'') v"
          "distinct (stack e'')"
          "distinct (cstack e'')"
          "n m. n  m in stack e''  n  m in cstack e''"
          "n m. n  m in stack e''  reachable m n"
          "explored e''  visited e''"
          "set (cstack e'')  visited e''"
          "n  explored e''. m. reachable n m  m  explored e''"
          "n m. m  𝒮 e'' n  (𝒮 e'' n = 𝒮 e'' m)"
          "n. n  visited e''  𝒮 e'' n = {n}"
          "n  set (stack e'').  m  set (stack e''). 
                n  m  𝒮 e'' n  𝒮 e'' m = {}"
          " {𝒮 e'' n | n. n  set (stack e'')} = visited e'' - explored e''"
          "n  set (stack e''). m  𝒮 e'' n. 
              m  set (cstack e'')  m  n in cstack e''"
          "n. is_subscc (𝒮 e'' n)"
          "S  sccs e''. is_scc S"
          " (sccs e'') = explored e''"
      by (auto simp: wf_env_def e''_def)
    moreover
    from wf 2 3 have "v. vsuccs e'' v  successors v  visited e''"
      by (auto simp: wf_env_def e''_def)
    moreover
    from wf v have "n. n  visited e''  vsuccs e'' n = {}"
      by (auto simp: wf_env_def e''_def)
    moreover
    from wf 2
    have "v. v  explored e''  vsuccs e'' v = successors v"
      by (auto simp: wf_env_def e''_def)
    moreover
    have "x y. x  y in stack e''  x  y 
             (u  𝒮 e'' x. ¬ reachable_avoiding u y (unvisited e'' x))"
    proof (clarify)
      fix x y u
      assume "x  y in stack e''" "x  y" 
             "u  𝒮 e'' x"
             "reachable_avoiding u y (unvisited e'' x)"
      hence prec: "x  y in stack e" "u  𝒮 e x"
        by (auto simp: e''_def)
      with stack_unexplored[OF wf] have "y  explored e"
        by (blast dest: precedes_mem)
      have "(unvisited e x = unvisited e'' x)
           (unvisited e x = unvisited e'' x  {(v,w)})"
        by (auto simp: e''_def unvisited_def split: if_splits)
      thus "False"
      proof
        assume "unvisited e x = unvisited e'' x"
        with prec x  y reachable_avoiding u y (unvisited e'' x) wf
        show ?thesis
          unfolding wf_env_def by metis
      next
        assume "unvisited e x = unvisited e'' x  {(v,w)}"
        with wf reachable_avoiding u y (unvisited e'' x)
             y  explored e w  explored e prec x  y
        show ?thesis
          using avoiding_explored[OF wf] unfolding wf_env_def
          by (metis (no_types, lifting))
      qed
    qed
    moreover 
    from wf 2
    have "n  visited e'' - set (cstack e''). vsuccs e'' n = successors n"
      by (auto simp: e''_def wf_env_def)
    ultimately show ?thesis
      unfolding wf_env_def by meson
  qed
  with 1 3 show ?thesis
    by (auto simp: pre_dfss_def e''_def)
qed

text ‹
  The call to @{text dfs} establishes the pre-condition for the
  recursive call to @{text dfss} in the body of @{text dfss}.
›
lemma pre_dfss_post_dfs_pre_dfss:
  fixes e v w
  defines "e'  dfs w e"
  defines "e''  e'vsuccs := (λx. if x=v then vsuccs e' v  {w} else vsuccs e' x)"
  assumes pre: "pre_dfss v e"
      and w: "w  successors v" "w  visited e"
      and post: "post_dfs w e e'"
  shows "pre_dfss v e''"
proof -
  from pre 
  have "wf_env e" "v  visited e" "stack e  []" "v  𝒮 e (hd (stack e))"
    by (auto simp: pre_dfss_def)
  with post have "stack e'  []" "v  𝒮 e' (hd (stack e'))"
    by (auto dest: dfs_S_hd_stack)
  
  from post have "w  visited e'"
    by (simp add: post_dfs_def)

  have "wf_env e''"
  proof -
    from post have wf': "wf_env e'"
      by (simp add: post_dfs_def)
    hence "n  visited e''. reachable (root e'') n"
          "distinct (stack e'')"
          "distinct (cstack e'')"
          "n m. n  m in stack e''  n  m in cstack e''"
          "n m. n  m in stack e''  reachable m n"
          "explored e''  visited e''"
          "set (cstack e'')  visited e''"
          "n  explored e''. m. reachable n m  m  explored e''"
          "n m. m  𝒮 e'' n  (𝒮 e'' n = 𝒮 e'' m)"
          "n. n  visited e''  𝒮 e'' n = {n}"
          "n  set (stack e'').  m  set (stack e''). 
              n  m  𝒮 e'' n  𝒮 e'' m = {}"
          " {𝒮 e'' n | n. n  set (stack e'')} = visited e'' - explored e''"
          "n  set (stack e'').  m  𝒮 e'' n. m  set (cstack e'')  m  n in cstack e''"
          "n. is_subscc (𝒮 e'' n)"
          "S  sccs e''. is_scc S"
          " (sccs e'') = explored e''"
      by (auto simp: wf_env_def e''_def)
    moreover
    from wf' w have "n. vsuccs e'' n  successors n"
      by (auto simp: wf_env_def e''_def)
    moreover 
    from wf' w  visited e' have "n. vsuccs e'' n  visited e''"
      by (auto simp: wf_env_def e''_def)
    moreover
    from post v  visited e
    have "n. n  visited e''  vsuccs e'' n = {}"
      apply (simp add: post_dfs_def wf_env_def sub_env_def e''_def)
      by (meson subsetD)
    moreover
    from wf' w
    have "n  explored e''. vsuccs e'' n = successors n"
      by (auto simp: wf_env_def e''_def)
    moreover
    have "n m. n  m in stack e''  n  m 
             (u  𝒮 e'' n. ¬ reachable_avoiding u m (unvisited e'' n))"
    proof (clarify)
      fix x y u
      assume "x  y in stack e''" "x  y"
             "u  𝒮 e'' x"
             "reachable_avoiding u y (unvisited e'' x)"
      hence 1: "x  y in stack e'" "u  𝒮 e' x"
        by (auto simp: e''_def)
      with stack_unexplored[OF wf'] have "y  explored e'"
        by (auto dest: precedes_mem)
      have "(unvisited e' x = unvisited e'' x)
           (unvisited e' x = unvisited e'' x  {(v,w)})"
        by (auto simp: e''_def unvisited_def split: if_splits)
      thus "False"
      proof
        assume "unvisited e' x = unvisited e'' x"
        with 1 x  y reachable_avoiding u y (unvisited e'' x) wf'
        show ?thesis
          unfolding wf_env_def by metis
      next
        assume unv: "unvisited e' x = unvisited e'' x  {(v,w)}"
        from post
        have "w  explored e' 
            (w  𝒮 e' (hd (stack e'))  (n  set (tl (stack e')). 𝒮 e' n = 𝒮 e n))"
          by (auto simp: post_dfs_def)
        thus ?thesis
        proof
          assume "w  explored e'"
          with wf' unv reachable_avoiding u y (unvisited e'' x)
               y  explored e' 1 x  y
          show ?thesis
            using avoiding_explored[OF wf'] unfolding wf_env_def
            by (metis (no_types, lifting))
        next
          assume w: "w  𝒮 e' (hd (stack e'))
                   (n  set (tl (stack e')). 𝒮 e' n = 𝒮 e n)"
          from reachable_avoiding u y (unvisited e'' x)[THEN ra_add_edge]
               unv
          have "reachable_avoiding u y (unvisited e' x)
               reachable_avoiding w y (unvisited e' x)"
            by auto
          thus ?thesis
          proof
            assume "reachable_avoiding u y (unvisited e' x)"
            with x  y in stack e'' x  y u  𝒮 e'' x wf'
            show ?thesis
              by (auto simp: e''_def wf_env_def)
          next
            assume "reachable_avoiding w y (unvisited e' x)"
            from unv have "v  𝒮 e' x"
              by (auto simp: unvisited_def)
            from x  y in stack e'' have "x  set (stack e')"
              by (simp add: e''_def precedes_mem)
            have "x = hd (stack e')"
            proof (rule ccontr)
              assume "x  hd (stack e')"
              with x  set (stack e') stack e'  []
              have "x  set (tl (stack e'))"
                by (metis hd_Cons_tl set_ConsD)
              with w v  𝒮 e' x have "v  𝒮 e x"
                by auto
              moreover
              from post stack e'  [] x  set (stack e') x  set (tl (stack e'))
              have "x  set (tl (stack e))"
                unfolding post_dfs_def
                by (metis Un_iff self_append_conv2 set_append tl_append2)
              moreover
              from pre have "wf_env e" "stack e  []" "v  𝒮 e (hd (stack e))"
                by (auto simp: pre_dfss_def)
              ultimately show "False"
                unfolding wf_env_def
                by (metis (no_types, lifting) distinct.simps(2) hd_Cons_tl insert_disjoint(2) 
                          list.set_sel(1) list.set_sel(2) mk_disjoint_insert)
            qed
            with reachable_avoiding w y (unvisited e' x)
                 x  y in stack e'' x  y w wf'
            show ?thesis
              by (auto simp add: e''_def wf_env_def)
          qed
        qed
      qed
    qed

    moreover
    from wf' n. vsuccs e'' n  successors n
    have "n  visited e'' - set (cstack e''). vsuccs e'' n = successors n"
      by (auto simp: wf_env_def e''_def split: if_splits)
    ultimately show ?thesis
      unfolding wf_env_def by (meson le_inf_iff)
  qed

  show "pre_dfss v e''"
  proof -
    from pre post
    have "v  visited e''"
      by (auto simp: pre_dfss_def post_dfs_def sub_env_def e''_def)
    moreover
    {
      fix u
      assume u: "u  vsuccs e'' v"
      have "u  explored e''  𝒮 e'' (hd (stack e''))"
      proof (cases "u = w")
        case True
        with post show ?thesis
          by (auto simp: post_dfs_def e''_def)
      next
        case False
        with u pre post
        have "u  explored e  𝒮 e (hd (stack e))"
          by (auto simp: pre_dfss_def post_dfs_def e''_def)
        then show ?thesis
        proof
          assume "u  explored e"
          with post show ?thesis
            by (auto simp: post_dfs_def sub_env_def e''_def)
        next
          assume "u  𝒮 e (hd (stack e))"
          with wf_env e post stack e  []
          show ?thesis
            by (auto simp: e''_def dest: dfs_S_hd_stack)
        qed
      qed
    }
    moreover
    from pre post
    have "n  set (stack e''). reachable n v"
      unfolding pre_dfss_def post_dfs_def
      using e''_def by force
    moreover
    from stack e'  [] have "stack e''  []"
      by (simp add: e''_def)
    moreover
    from v  𝒮 e' (hd (stack e')) have "v  𝒮 e'' (hd (stack e''))"
      by (simp add: e''_def)
    moreover
    from pre post have "ns. cstack e'' = v # ns"
      by (auto simp: pre_dfss_def post_dfs_def e''_def)
    ultimately show ?thesis
      using wf_env e'' unfolding pre_dfss_def by blast
  qed
qed

text ‹
  Finally, the pre-condition for the recursive call to @{text dfss}
  at the end of the body of function @{text dfss} also holds if
  @{text unite} was applied.
›

lemma pre_dfss_unite_pre_dfss:
  fixes e v w
  defines "e'  unite v w e"
  defines "e''  e'vsuccs := (λx. if x=v then vsuccs e' v  {w} else vsuccs e' x)"
  assumes pre: "pre_dfss v e"
      and w: "w  successors v" "w  vsuccs e v" "w  visited e" "w  explored e"
  shows "pre_dfss v e''"
proof -
  from pre have wf: "wf_env e"
    by (simp add: pre_dfss_def)
  from pre have "v  visited e"
    by (simp add: pre_dfss_def)
  from pre w have "v  explored e"
    unfolding pre_dfss_def wf_env_def
    by (meson reachable_edge)

  from unite_stack[OF wf w] obtain pfx where
    pfx: "stack e = pfx @ stack e'" "stack e'  []"
         "let cc =  {𝒮 e n |n. n  set pfx  {hd (stack e')}}
          in 𝒮 e' = (λx. if x  cc then cc else 𝒮 e x)"
         "w  𝒮 e' (hd (stack e'))"
    by (auto simp: e'_def)
  define cc where "cc   {𝒮 e n |n. n  set pfx  {hd (stack e')}}"

  from unite_wf_env[OF pre w] have wf': "wf_env e'"
    by (simp add: e'_def)

  from stack e = pfx @ stack e' stack e'  []
  have "hd (stack e)  set pfx  {hd (stack e')}"
    by (simp add: hd_append)
  with pre have "v  cc"
    by (auto simp: pre_dfss_def cc_def)
  from S_reflexive[OF wf, of "hd (stack e')"]
  have "hd (stack e')  cc"
    by (auto simp: cc_def)
  with pfx v  cc have "v  𝒮 e' (hd (stack e'))"
    by (auto simp: cc_def)

  from unite_sub_env[OF pre w] have "sub_env e e'"
    by (simp add: e'_def)

  have "wf_env e''"
  proof -
    from wf'
    have "n  visited e''. reachable (root e'') n"
         "distinct (stack e'')" 
         "distinct (cstack e'')"
         "n m. n  m in stack e''  n  m in cstack e''"
         "n m. n  m in stack e''  reachable m n"
         "explored e''  visited e''"
         "set (cstack e'')  visited e''"
         "n  explored e''. m. reachable n m  m  explored e''"
         "n m. m  𝒮 e'' n  (𝒮 e'' n = 𝒮 e'' m)"
         "n. n  visited e''  𝒮 e'' n = {n}"
         "n  set (stack e''). m  set (stack e''). 
             n  m  𝒮 e'' n  𝒮 e'' m = {}"
         " {𝒮 e'' n | n. n  set (stack e'')} = visited e'' - explored e''"
         "n  set (stack e''). m  𝒮 e'' n. 
             m  set (cstack e'')  m  n in cstack e''"
         "n. is_subscc (𝒮 e'' n)"
         "S  sccs e''. is_scc S"
         " (sccs e'') = explored e''"
      by (auto simp: wf_env_def e''_def)

    moreover
    from wf' w sub_env e e'
    have "n. vsuccs e'' n  successors n  visited e''"
      by (auto simp: wf_env_def sub_env_def e''_def)

    moreover
    from wf' v  visited e sub_env e e'
    have "n. n  visited e''  vsuccs e'' n = {}"
      by (auto simp: wf_env_def sub_env_def e''_def)

    moreover
    from wf' v  explored e
    have "n  explored e''. vsuccs e'' n = successors n"
      by (auto simp: wf_env_def e''_def e'_def unite_def)

    moreover
    from wf' w  successors v
    have "n  visited e'' - set (cstack e''). vsuccs e'' n = successors n"
      by (auto simp: wf_env_def e''_def e'_def unite_def)

    moreover
    have "x y. x  y in stack e''  x  y 
                (u  𝒮 e'' x. ¬ reachable_avoiding u y (unvisited e'' x))"
    proof (clarify)
      fix x y u
      assume xy: "x  y in stack e''" "x  y"
         and u: "u  𝒮 e'' x" "reachable_avoiding u y (unvisited e'' x)"
      hence prec: "x  y in stack e'" "u  𝒮 e' x"
        by (simp add: e''_def)+
      show "False"
      proof (cases "x = hd (stack e')")
        case True
        with v  𝒮 e' (hd (stack e'))
        have "unvisited e' x = unvisited e'' x
             (unvisited e' x = unvisited e'' x  {(v,w)})"
          by (auto simp: e''_def unvisited_def split: if_splits)
        thus "False"
        proof
          assume "unvisited e' x = unvisited e'' x"
          with prec x  y reachable_avoiding u y (unvisited e'' x) wf'
          show ?thesis
            unfolding wf_env_def by metis
        next
          assume "unvisited e' x = unvisited e'' x  {(v,w)}"
          with reachable_avoiding u y (unvisited e'' x)[THEN ra_add_edge]
          have "reachable_avoiding u y (unvisited e' x)
               reachable_avoiding w y (unvisited e' x)"
            by auto
          thus ?thesis
          proof
            assume "reachable_avoiding u y (unvisited e' x)"
            with prec x  y wf' show ?thesis
              by (auto simp: wf_env_def)
          next
            assume "reachable_avoiding w y (unvisited e' x)"
            with x = hd (stack e') w  𝒮 e' (hd (stack e')) 
                 x  y in stack e' x  y wf'
            show ?thesis
              by (auto simp: wf_env_def)
          qed
        qed
      next
        case False
        with x  y in stack e' stack e'  []
        have "x  set (tl (stack e'))"
          by (metis list.exhaust_sel precedes_mem(1) set_ConsD)
        with unite_S_tl[OF wf w] u  𝒮 e' x
        have "u  𝒮 e x"
          by (simp add: e'_def)
        moreover
        from x  y in stack e' stack e = pfx @ stack e'
        have "x  y in stack e"
          by (simp add: precedes_append_left)
        moreover
        from v  𝒮 e' (hd (stack e')) x  set (tl (stack e'))
             stack e'  [] wf'
        have "v  𝒮 e' x"
          unfolding wf_env_def
          by (metis (no_types, lifting) Diff_cancel Diff_triv distinct.simps(2) insert_not_empty 
                    list.exhaust_sel list.set_sel(1) list.set_sel(2) mk_disjoint_insert)
        hence "unvisited e'' x = unvisited e' x"
          by (auto simp: unvisited_def e''_def split: if_splits)
        moreover
        from x  set (tl (stack e')) unite_S_tl[OF wf w]
        have "unvisited e' x = unvisited e x"
          by (simp add: unvisited_def e'_def unite_def)
        ultimately show ?thesis
          using x  y reachable_avoiding u y (unvisited e'' x) wf
          by (auto simp: wf_env_def)
      qed
    qed
      
    ultimately show ?thesis
      unfolding wf_env_def by meson
  qed

  show "pre_dfss v e''"
  proof -
    from pre have "v  visited e''"
      by (simp add: pre_dfss_def e''_def e'_def unite_def)

    moreover
    {
      fix u
      assume u: "u  vsuccs e'' v"
      have "u  explored e''  𝒮 e'' (hd (stack e''))"
      proof (cases "u = w")
        case True
        with w  𝒮 e' (hd (stack e')) show ?thesis 
          by (simp add: e''_def)
      next
        case False
        with u have "u  vsuccs e v"
          by (simp add: e''_def e'_def unite_def)
        with pre have "u  explored e  𝒮 e (hd (stack e))"
          by (auto simp: pre_dfss_def)
        then show ?thesis
        proof
          assume "u  explored e"
          thus ?thesis
            by (simp add: e''_def e'_def unite_def)
        next
          assume "u  𝒮 e (hd (stack e))"
          with hd (stack e)  set pfx  {hd (stack e')}
          have "u  cc"
            by (auto simp: cc_def)
          moreover
          from S_reflexive[OF wf, of "hd (stack e')"] pfx
          have "𝒮 e' (hd (stack e')) = cc"
            by (auto simp: cc_def)
          ultimately show ?thesis
            by (simp add: e''_def)
        qed
      qed
    }
    hence "w  vsuccs e'' v. w  explored e''  𝒮 e'' (hd (stack e''))"
      by blast

    moreover
    from pre stack e = pfx @ stack e'
    have "n  set (stack e''). reachable n v"
      by (auto simp: pre_dfss_def e''_def)

    moreover
    from stack e'  [] have "stack e''  []"
      by (simp add: e''_def)

    moreover
    from v  𝒮 e' (hd (stack e')) have "v  𝒮 e'' (hd (stack e''))"
      by (simp add: e''_def)

    moreover
    from pre have "ns. cstack e'' = v # ns"
      by (auto simp: pre_dfss_def e''_def e'_def unite_def)

    ultimately show ?thesis
      using wf_env e'' unfolding pre_dfss_def by blast
  qed
qed

subsection ‹Lemmas establishing the post-conditions›

text ‹
  Assuming the pre-condition of function @{text dfs} and the post-condition of
  the call to @{text dfss} in the body of that function, the post-condition of
  @{text dfs} is established.
›
lemma pre_dfs_implies_post_dfs:
  fixes v e
  defines "e1  evisited := visited e  {v}, 
                  stack := (v # stack e), 
                  cstack:=(v # cstack e)"
  defines "e'  dfss v e1"
  defines "e''  e' cstack := tl(cstack e')"
  assumes 1: "pre_dfs v e"
      and 2: "dfs_dfss_dom (Inl(v, e))"
      and 3: "post_dfss v e1 e'"
  shows "post_dfs v e (dfs v e)"
proof -
  from 1 have wf: "wf_env e"
    by (simp add: pre_dfs_def)
  from 1 have v: "v  visited e"
    by (simp add: pre_dfs_def)
  from 3 have wf': "wf_env e'"
    by (simp add: post_dfss_def)
  from 3 have cst': "cstack e' = v # cstack e"
    by (simp add: post_dfss_def e1_def)
  show ?thesis
  proof (cases "v = hd(stack e')")
    case True
    have notempty: "stack e' = v # stack e"
    proof -
      from 3 obtain ns where 
        ns: "stack e1 = ns @ (stack e')" "stack e'  []"
        by (auto simp: post_dfss_def)
      have "ns = []"
      proof (rule ccontr)
        assume "ns  []"
        with ns have "hd ns = v"
          apply (simp add: e1_def)
          by (metis hd_append2 list.sel(1))
        with True ns ns  [] have "¬ distinct (stack e1)"
          by (metis disjoint_iff_not_equal distinct_append hd_in_set) 
        with wf v stack_visited[OF wf] show False
          by (auto simp: wf_env_def e1_def)
      qed
      with ns show ?thesis
        by (simp add: e1_def)
    qed
    have e2: "dfs v e = e'sccs := sccs e'  {𝒮 e' v}, 
                           explored := explored e'  (𝒮 e' v), 
                           stack := tl (stack e'),
                           cstack := tl (cstack e')" (is "_ = ?e2")
      using True 2 dfs.psimps[of v e] unfolding e1_def e'_def 
      by (fastforce simp: e1_def e'_def)

    have sub: "sub_env e e1"
      by (auto simp: sub_env_def e1_def)

    from notempty have stack2: "stack ?e2 = stack e"
      by (simp add: e1_def)

    moreover from 3 have "v  visited ?e2"
      by (auto simp: post_dfss_def sub_env_def e1_def)

    moreover
    from sub 3 have "sub_env e e'"
      unfolding post_dfss_def by (auto elim: sub_env_trans)
    with stack2 have subenv: "sub_env e ?e2"
      by (fastforce simp: sub_env_def)

    moreover have "wf_env ?e2"
    proof -
      from wf' 
      have "n  visited ?e2. reachable (root ?e2) n"
           "distinct (stack ?e2)"
           "n. vsuccs ?e2 n  successors n  visited ?e2"
           "n. n  visited ?e2  vsuccs ?e2 n = {}"
           "n m. m  𝒮 ?e2 n  (𝒮 ?e2 n = 𝒮 ?e2 m)"
           "n. n  visited ?e2  𝒮 ?e2 n = {n}"
           "n. is_subscc (𝒮 ?e2 n)"
           " (sccs ?e2) = explored ?e2"
        by (auto simp: wf_env_def distinct_tl)

      moreover
      from 1 cst' have "distinct (cstack ?e2)"
        by (auto simp: pre_dfs_def wf_env_def)

      moreover
      from 1 stack2 have "n m. n  m in stack ?e2  reachable m n"
        by (auto simp: pre_dfs_def wf_env_def)

      moreover
      from 1 stack2 cst'
      have "n m. n  m in stack ?e2  n  m in cstack ?e2"
        by (auto simp: pre_dfs_def wf_env_def)

      moreover 
      from notempty wf' have "explored ?e2  visited ?e2"
        apply (simp add: wf_env_def)
        using stack_class[OF wf']
        by (smt (verit, del_insts) Diff_iff insert_subset list.simps(15) subset_eq)
  
      moreover
      from 3 cst' have "set (cstack ?e2)  visited ?e2"
        by (simp add: post_dfss_def wf_env_def e1_def)

      moreover
      {
        fix u
        assume "u  explored ?e2"
        have "vsuccs ?e2 u = successors u"
        proof (cases "u  explored e'")
          case True
          with wf' show ?thesis
            by (auto simp: wf_env_def)
        next
          case False
          with u  explored ?e2 have "u  𝒮 e' v"
            by simp
          show ?thesis
          proof (cases "u = v")
            case True
            with 3 show ?thesis
              by (auto simp: post_dfss_def)
          next
            case False
            have "u  visited e' - set (cstack e')"
            proof
              from notempty u  𝒮 e' v stack_class[OF wf'] False
              show "u  visited e'"
                by auto
            next
              show "u  set (cstack e')"
              proof
                assume u: "u  set (cstack e')"
                with notempty u  𝒮 e' v wf_env e' have "u  v in cstack e'"
                  by (auto simp: wf_env_def)
                with cst' u False wf' show "False"
                  unfolding wf_env_def
                  by (metis head_precedes precedes_antisym)
              qed
            qed
            with 3 show ?thesis
              by (auto simp: post_dfss_def wf_env_def)
          qed
        qed
      }
      note explored_vsuccs = this

      moreover have "n  explored ?e2. m. reachable n m  m  explored ?e2"
      proof (clarify)
        fix x y
        assume asm: "x  explored ?e2" "reachable x y"
        show "y  explored ?e2"
        proof (cases "x  explored e'")
          case True
          with wf_env e' reachable x y show ?thesis
            by (simp add: wf_env_def)
        next
          case False
          with asm have "x  𝒮 e' v"
             by simp
          with explored ?e2  visited ?e2 have "x  visited e'"
            by auto
          from x  𝒮 e' v wf' have "reachable v x"
            by (auto simp: wf_env_def is_subscc_def)
          have "y  visited e'"
          proof (rule ccontr)
            assume "y  visited e'"
            with reachable_visited[OF wf' x  visited e' reachable x y]
            obtain n m where
              "n  visited e'" "m  successors n - vsuccs e' n"
              "reachable x n" "reachable m y"
              by blast
            from wf' m  successors n - vsuccs e' n
            have "n  explored e'"
              by (auto simp: wf_env_def)
            obtain n' where
              "n'  set (stack e')" "n  𝒮 e' n'"
              by (rule visited_unexplored[OF wf' n  visited e' n  explored e'])
            have "n' = v"
            proof (rule ccontr)
              assume "n'  v"
              with n'  set (stack e') v = hd (stack e')
              have "n'  set (tl (stack e'))"
                by (metis emptyE hd_Cons_tl set_ConsD set_empty)
              moreover
              from n  𝒮 e' n' wf_env e' have "reachable n n'"
                by (auto simp: wf_env_def is_subscc_def)
              with reachable v x reachable x n reachable_trans 
              have "reachable v n'"
                by blast
              ultimately show "False"
                using 3 v = hd (stack e') 
                by (auto simp: post_dfss_def)
            qed
            with n  𝒮 e' n' m  successors n - vsuccs e' n explored_vsuccs
            show "False"
              by auto
          qed
          show ?thesis
          proof (cases "y  explored e'")
            case True
            then show ?thesis
              by simp
          next
            case False
            obtain n where ndef: "n  set (stack e')" "(y  𝒮 e' n)"
              by (rule visited_unexplored[OF wf' y  visited e' False])
            show ?thesis
            proof (cases "n = v")
              case True
              with ndef show ?thesis by simp
            next
              case False
              with ndef notempty have "n  set (tl (stack e'))"
                by simp
              moreover
              from wf' ndef have "reachable y n"
                by (auto simp: wf_env_def is_subscc_def)
              with reachable v x reachable x y
              have "reachable v n"
                by (meson reachable_trans)
              ultimately show ?thesis
                using v = hd (stack e') 3 
                by (simp add: post_dfss_def)
            qed
          qed
        qed
      qed

      moreover
      from 3 cst'
      have "n  visited ?e2 - set (cstack ?e2). vsuccs ?e2 n = successors n"
        apply (simp add: post_dfss_def wf_env_def)
        by (metis (no_types, lifting) Diff_empty Diff_iff empty_set insertE 
                  list.exhaust_sel list.sel(1) list.simps(15))

      moreover
      from wf' notempty
      have "n m. n  set (stack ?e2)  m  set (stack ?e2)  n  m 
                   (𝒮 ?e2 n  𝒮 ?e2 m = {})"
        by (simp add: wf_env_def)

      moreover 
      have " {𝒮 ?e2 n | n . n  set (stack ?e2)} = visited ?e2 - explored ?e2"
      proof -
        from wf' notempty
        have "( {𝒮 ?e2 n | n . n  set (stack ?e2)})  𝒮 e' v = {}"
          by (auto simp: wf_env_def)
        with notempty
        have " {𝒮 ?e2 n | n . n  set (stack ?e2)} =
              ( {𝒮 e' n | n . n  set (stack e')}) - 𝒮 e' v"
          by auto
        also from wf' 
        have " = (visited e' - explored e') - 𝒮 e' v"
          by (simp add: wf_env_def)
        finally show ?thesis
          by auto
      qed

      moreover
      have "n  set (stack ?e2). m  𝒮 ?e2 n. m  set (cstack ?e2)  m  n in cstack ?e2"
      proof (clarsimp simp: cst')
        fix n m
        assume "n  set (tl (stack e'))"
               "m  𝒮 e' n" "m  set (cstack e)"
        with 3 have "m  𝒮 e n"
          by (auto simp: post_dfss_def e1_def)
        with wf notempty n  set (tl (stack e')) m  set (cstack e)
        show "m  n in cstack e"
          by (auto simp: wf_env_def)
      qed

      moreover
      {
        fix x y u
        assume xy: "x  y in stack ?e2" "x  y"
           and u: "u  𝒮 ?e2 x" "reachable_avoiding u y (unvisited ?e2 x)"
        from xy notempty stack2
        have "x  y in stack e'"
          by (metis head_precedes insert_iff list.simps(15) precedes_in_tail precedes_mem(2))
        with wf' x  y u have "False"
          by (auto simp: wf_env_def unvisited_def)
      }

      moreover have "S  sccs ?e2. is_scc S"
      proof (clarify)
        fix S
        assume asm: "S  sccs ?e2"
        show "is_scc S"
        proof (cases "S = 𝒮 e' v")
          case True
          with S_reflexive[OF wf'] have "S  {}"
            by blast
          from wf' True have subscc: "is_subscc S"
            by (simp add: wf_env_def)
          {
            assume "¬ is_scc S"
            with S  {} is_subscc S obtain S' where
              S'_def: "S'  S" "S  S'" "is_subscc S'" 
              unfolding is_scc_def by blast
            then obtain x where "x  S'  x  S"
              by blast
            with True S'_def wf'
            have xv: "reachable v x  reachable x v"
              unfolding wf_env_def is_subscc_def by (metis in_mono)
            from v w. w  𝒮 ?e2 v  (𝒮 ?e2 v = 𝒮 ?e2 w)
            have "v  explored ?e2"
              by auto
            with x  explored ?e2. y. reachable x y  y  explored ?e2
                 xv S = 𝒮 e' v x  S'  x  S
            have "x  explored e'"
              by auto
            with wf' xv have "v  explored e'"
              by (auto simp: wf_env_def)
            with notempty have "False"
              by (auto intro: stack_unexplored[OF wf'])
          }
          then show ?thesis
            by blast
        next
          case False
          with asm wf' show ?thesis
            by (auto simp: wf_env_def)
        qed
      qed

      ultimately show ?thesis
        unfolding wf_env_def by meson
    qed

    moreover 
    from wf_env ?e2 have "v  explored ?e2"
      by (auto simp: wf_env_def)

    moreover
    from 3 have "vsuccs ?e2 v = successors v"
      by (simp add: post_dfss_def)

    moreover
    from 1 3 have "w  visited e. vsuccs ?e2 w = vsuccs e w"
      by (auto simp: pre_dfs_def post_dfss_def e1_def)

    moreover 
    from stack2 1 
    have "n  set (stack ?e2). reachable n v"
      by (simp add: pre_dfs_def)

    moreover 
    from stack2 have "ns. stack e = ns @ (stack ?e2)"
      by auto

    moreover
    from 3 have "n  set (stack ?e2). 𝒮 ?e2 n = 𝒮 e n"
      by (auto simp: post_dfss_def e1_def)

    moreover
    from cst' have "cstack ?e2 = cstack e"
      by simp

    ultimately show ?thesis
      unfolding post_dfs_def using e2 by simp
  next
    case False
    with 2 have e': "dfs v e = e''"
      by (simp add: dfs.psimps e''_def e'_def e1_def)

    moreover have "wf_env e''"
    proof -
      from wf'
      have "n  visited e''. reachable (root e'') n"
           "distinct (stack e'')"
           "distinct (cstack e'')"
           "n m. n  m in stack e''  reachable m n"
           "explored e''  visited e''"
           "n  explored e''. m. reachable n m  m  explored e''"
           "n. vsuccs e'' n  successors n  visited e''"
           "n. n  visited e''  vsuccs e'' n = {}"
           "n  explored e''. vsuccs e'' n = successors n"
           "n m. m  𝒮 e'' n  (𝒮 e'' n = 𝒮 e'' m)"
           "n. n  visited e''  𝒮 e'' n = {n}"
           "n  set (stack e''). m  set (stack e''). 
               n  m  𝒮 e'' n  𝒮 e'' m = {}"
           " {𝒮 e'' n | n. n  set (stack e'')} = visited e'' - explored e''"
           "n. is_subscc (𝒮 e'' n)"
           "S  sccs e''. is_scc S"
           " (sccs e'') = explored e''"
        by (auto simp: e''_def wf_env_def distinct_tl)

      moreover have "n m. n  m in stack e''  n  m in cstack e''"
      proof (clarsimp simp add: e''_def)
        fix n m
        assume nm: "n  m in stack e'"
        with 3 have "n  m in cstack e'"
          unfolding post_dfss_def wf_env_def 
          by meson
        moreover
        have "n  v"
        proof
          assume "n = v"
          with nm have "n  set (stack e')"
            by (simp add: precedes_mem)
          with 3 n = v have "v = hd (stack e')"
            unfolding post_dfss_def wf_env_def
            by (metis (no_types, opaque_lifting) IntI equals0D list.set_sel(1))
          with v  hd (stack e') show "False"
            by simp
        qed
        ultimately show "n  m in tl (cstack e')"
          by (simp add: cst' precedes_in_tail)
      qed

      moreover
      from 3 have "set (cstack e'')  visited e''"
        by (simp add: post_dfss_def wf_env_def e''_def e1_def subset_eq)

      moreover 
      from 3
      have "n  visited e'' - set (cstack e''). vsuccs e'' n = successors n"
        apply (simp add: post_dfss_def wf_env_def e''_def e1_def)
        by (metis (no_types, opaque_lifting) DiffE DiffI set_ConsD)

      moreover 
      have "n  set (stack e''). m  𝒮 e'' n. 
               m  set (cstack e'')  m  n in cstack e''"
      proof (clarsimp simp: e''_def)
        fix m n
        assume asm: "n  set (stack e')" "m  𝒮 e' n" 
                    "m  set (tl (cstack e'))"
        with wf' cst' have "m  v" "m  n in cstack e'"
          by (auto simp: wf_env_def)
        with cst' show "m  n in tl (cstack e')"
          by (simp add: precedes_in_tail)
      qed

      moreover
      from wf'
      have "(x y. x  y in stack e''  x  y 
                (u  𝒮 e'' x. ¬ reachable_avoiding u y (unvisited e'' x)))"
        by (force simp: e''_def wf_env_def unvisited_def)

      ultimately show ?thesis
        unfolding wf_env_def by blast
    qed

    moreover
    from 3 have "v  visited e''"
      by (auto simp: post_dfss_def sub_env_def e''_def e1_def)

    moreover
    have subenv: "sub_env e e''"
    proof -
      have "sub_env e e1"
        by (auto simp: sub_env_def e1_def)
      with 3 have "sub_env e e'"
        by (auto simp: post_dfss_def elim: sub_env_trans)
      thus ?thesis
        by (auto simp add: sub_env_def e''_def)
    qed

    moreover
    from 3 have "vsuccs e'' v = successors v"
      by (simp add: post_dfss_def e''_def)

    moreover
    from 1 3 have "w  visited e. vsuccs e'' w = vsuccs e w"
      by (auto simp: pre_dfs_def post_dfss_def e1_def e''_def)

    moreover 
    from 3 have "n  set (stack e''). reachable n v"
      by (auto simp: e''_def post_dfss_def)

    moreover 
    from 3 v  hd (stack e')
    have "ns. stack e = ns @ (stack e'')"
      apply (simp add: post_dfss_def e''_def e1_def)
      by (metis append_Nil list.sel(1) list.sel(3) tl_append2)

    moreover
    from 3
    have "stack e''  []" "v  𝒮 e'' (hd (stack e''))"
         "n  set (tl (stack e'')). 𝒮 e'' n = 𝒮 e n"
      by (auto simp: post_dfss_def e1_def e''_def)

    moreover 
    from cst' have "cstack e'' = cstack e"
      by (simp add: e''_def)

    ultimately show ?thesis unfolding post_dfs_def
      by blast
  qed
qed

text ‹
  The following lemma is central for proving
  partial correctness: assuming termination (represented by
  the predicate @{text dfs_dfss_dom}) and the pre-condition
  of the functions, both @{text dfs} and @{text dfss}
  establish their post-conditions. The first part of the
  theorem follows directly from the preceding lemma and the
  computational induction rule generated by Isabelle, the
  second part is proved directly, distinguishing the different
  cases in the definition of function @{text dfss}.
›
lemma pre_post:
  shows
  "dfs_dfss_dom (Inl(v,e)); pre_dfs v e  post_dfs v e (dfs v e)"
  "dfs_dfss_dom (Inr(v,e)); pre_dfss v e  post_dfss v e (dfss v e)"
proof (induct rule: dfs_dfss.pinduct)
  fix v e
  assume dom: "dfs_dfss_dom (Inl(v,e))"
     and predfs: "pre_dfs v e"
     and prepostdfss: "e1.  e1 = e visited := visited e  {v}, stack := v # stack e,
                                      cstack := v # cstack e; pre_dfss v e1 
                             post_dfss v e1 (dfss v e1)"
  then show "post_dfs v e (dfs v e)"
    using pre_dfs_implies_post_dfs pre_dfs_pre_dfss by auto
next
  fix v e
  assume dom: "dfs_dfss_dom (Inr(v,e))"
     and predfss: "pre_dfss v e"
     and prepostdfs: 
           "vs w. 
              vs = successors v - vsuccs e v; vs  {}; w = (SOME x. x  vs); 
               w  explored e; w  visited e; pre_dfs w e 
              post_dfs w e (dfs w e)"
     and prepostdfss:
           "vs w e' e''. 
              vs = successors v - vsuccs e v; vs  {}; w = (SOME x. x  vs); 
               e' = (if w  explored e then e 
                     else if w  visited e then dfs w e 
                     else unite v w e); 
               e'' = e'vsuccs := λx. if x = v then vsuccs e' v  {w} 
                                      else vsuccs e' x ;
               pre_dfss v e'' 
              post_dfss v e'' (dfss v e'')"
  show "post_dfss v e (dfss v e)"
  proof -
    let ?vs = "successors v - vsuccs e v"
    from predfss have wf: "wf_env e"
      by (simp add: pre_dfss_def)
    from predfss have "v  visited e"
      by (simp add: pre_dfss_def)
    from predfss have "v  explored e"
      by (meson DiffD2 list.set_sel(1) pre_dfss_def stack_class)

    show ?thesis
    proof (cases "?vs = {}") 
      case True
      with dom have "dfss v e = e" 
        by (simp add: dfss.psimps)
      moreover 
      from True wf have "vsuccs e v = successors v"
        unfolding wf_env_def
        by (meson Diff_eq_empty_iff le_infE subset_antisym)
      moreover 
      have "sub_env e e"
        by (simp add: sub_env_def)
      moreover 
      from predfss vsuccs e v = successors v
      have "w  successors v. w  explored e  𝒮 e (hd (stack e))"
           "n  set (stack e). reachable n v"
           "stack e  []"
           "v  𝒮 e (hd (stack e))"
        by (auto simp: pre_dfss_def)
      moreover have "ns. stack e = ns @ (stack e)"
        by simp
      moreover
      {
        fix n
        assume asm: "hd (stack e) = v"
                    "n  set (tl (stack e))"
                    "reachable v n"
        with stack e  [] have "v  n in stack e"
          by (metis head_precedes hd_Cons_tl list.set_sel(2))
        moreover
        from wf stack e  [] asm have "v  n"
          unfolding wf_env_def
          by (metis distinct.simps(2) list.exhaust_sel)
        moreover
        from wf have "v  𝒮 e v"
          by (rule S_reflexive)
        moreover
        {
          fix a b
          assume "a  𝒮 e v" "b  successors a - vsuccs e a"
          with vsuccs e v = successors v have "a  v"
            by auto
          from stack e  [] hd (stack e) = v 
          have "v  set (stack e)"
            by auto
          with a  v a  𝒮 e v wf have "a  visited e"
            unfolding wf_env_def by (metis singletonD)
          have "False"
          proof (cases "a  set (cstack e)")
            case True
            with v  set (stack e) a  𝒮 e v wf_env e
            have "a  v in cstack e"
              by (auto simp: wf_env_def)
            moreover
            from predfss obtain ns where "cstack e = v # ns"
              by (auto simp: pre_dfss_def)
            moreover
            from wf have "distinct (cstack e)"
              by (simp add: wf_env_def)
            ultimately have "a = v"
              using tail_not_precedes by force 
            with a  v show ?thesis ..
          next
            case False
            with a  visited e wf have "vsuccs e a = successors a"
              by (auto simp: wf_env_def)
            with b  successors a - vsuccs e a show ?thesis
              by simp
          qed
        }
        hence "unvisited e v = {}"
          by (auto simp: unvisited_def)

        ultimately have "¬ reachable_avoiding v n {}"
          using wf unfolding wf_env_def by metis
        with reachable v n have "False"
          by (simp add: ra_empty)
      }
      ultimately show ?thesis
        using wf by (auto simp: post_dfss_def)
    next
      case vs_case: False
      define w where "w = (SOME x. x  ?vs)"
      define e' where "e' = (if w  explored e then e 
                             else if w  visited e then dfs w e 
                             else unite v w e)"
      define e'' where "e'' = (e'vsuccs := λx. if x=v then vsuccs e' v  {w} else vsuccs e' x)"

      from dom vs_case have dfss: "dfss v e = dfss v e''"
        apply (simp add: dfss.psimps e''_def)
        using e'_def w_def by auto

      from vs_case have wvs: "w  ?vs"
        unfolding w_def by (metis some_in_eq)
      show ?thesis
      proof (cases "w  explored e")
        case True
        hence e': "e' = e"
          by (simp add: e'_def)
        with predfss wvs True
        have "pre_dfss v e''"
          by (auto simp: e''_def pre_dfss_explored_pre_dfss)
        with prepostdfss vs_case 
        have post'': "post_dfss v e'' (dfss v e'')"
          by (auto simp: w_def e'_def e''_def)

        moreover
        from post''
        have "u  visited e - {v}. vsuccs (dfss v e'') u = vsuccs e u"
          by (auto simp: post_dfss_def e' e''_def)

        moreover
        have "sub_env e e''"
          by (auto simp: sub_env_def e' e''_def)
        with post'' have "sub_env e (dfss v e'')"
          by (auto simp: post_dfss_def elim: sub_env_trans)

        moreover
        from e' have "stack e'' = stack e" "𝒮 e'' = 𝒮 e"
          by (auto simp add: e''_def)

        moreover
        have "cstack e'' = cstack e"
          by (simp add: e''_def e')

        ultimately show ?thesis
          by (auto simp: dfss post_dfss_def)
      next
        case notexplored: False
        then show ?thesis
        proof (cases "w  visited e")
          case True
          with e'_def notexplored have "e' = dfs w e"
            by auto
          with True notexplored pre_dfss_pre_dfs predfss 
               prepostdfs vs_case w_def
          have postdfsw: "post_dfs w e e'"
            by (metis DiffD1 some_in_eq)
          with predfss wvs True e' = dfs w e
          have "pre_dfss v e''"
            by (auto simp: e''_def pre_dfss_post_dfs_pre_dfss)
          with prepostdfss vs_case 
          have post'': "post_dfss v e'' (dfss v e'')"
            by (auto simp: w_def e'_def e''_def)

          moreover
          have "u  visited e - {v}. vsuccs (dfss v e'') u = vsuccs e u"
          proof
            fix u
            assume "u  visited e - {v}"
            with postdfsw 
            have u: "vsuccs e' u = vsuccs e u" "u  visited e'' - {v}"
              by (auto simp: post_dfs_def sub_env_def e''_def)
            with post'' have "vsuccs (dfss v e'') u = vsuccs e'' u"
              by (auto simp: post_dfss_def)
            with u show "vsuccs (dfss v e'') u = vsuccs e u"
              by (simp add: e''_def)
          qed

          moreover
          have "sub_env e (dfss v e'')"
          proof -
            from postdfsw have "sub_env e e'"
              by (simp add: post_dfs_def)
            moreover
            have "sub_env e' e''"
              by (auto simp: sub_env_def e''_def)
            moreover
            from post'' have "sub_env e'' (dfss v e'')"
              by (simp add: post_dfss_def)
            ultimately show ?thesis
              by (metis sub_env_trans)
          qed

          moreover
          from postdfsw post''
          have "ns. stack e = ns @ (stack (dfss v e''))"
            by (auto simp: post_dfs_def post_dfss_def e''_def)

          moreover
          {
            fix n
            assume n: "n  set (tl (stack (dfss v e'')))"
            with post'' have "𝒮 (dfss v e'') n = 𝒮 e' n"
              by (simp add: post_dfss_def e''_def)
            moreover
            from pre_dfss v e'' n post'' 
            have "stack e'  []  n  set (tl (stack e''))"
              apply (simp add: pre_dfss_def post_dfss_def e''_def)
              by (metis (no_types, lifting) Un_iff list.set_sel(2) self_append_conv2 set_append tl_append2)
            with postdfsw have "𝒮 e' n = 𝒮 e n"
              apply (simp add: post_dfs_def e''_def)
              by (metis list.set_sel(2))
            ultimately have "𝒮 (dfss v e'') n = 𝒮 e n"
              by simp
          }

          moreover
          from postdfsw have "cstack e'' = cstack e"
            by (auto simp: post_dfs_def e''_def)

          ultimately show ?thesis
            by (auto simp: dfss post_dfss_def)

        next
          case False
          hence e': "e' = unite v w e" using notexplored
            using e'_def by simp
          from False have "w  visited e"
            by simp
          from wf wvs notexplored False obtain pfx where
            pfx: "stack e = pfx @ (stack e')" "stack e'  []"
            unfolding e' by (blast dest: unite_stack)

          from predfss wvs notexplored False e' = unite v w e
          have "pre_dfss v e''"
            by (auto simp: e''_def pre_dfss_unite_pre_dfss)

          with prepostdfss vs_case e' = unite v w e  w  explored e w  visited e
          have post'': "post_dfss v e'' (dfss v e'')"
            by (auto simp: w_def e''_def)

          moreover
          from post''
          have "u  visited e - {v}. vsuccs (dfss v e'') u = vsuccs e u"
            by (auto simp: post_dfss_def e''_def e' unite_def)

          moreover
          have "sub_env e (dfss v e'')"
          proof -
            from predfss wvs w  visited e notexplored
            have "sub_env e e'"
              unfolding e' by (blast dest: unite_sub_env)
            moreover
            have "sub_env e' e''"
              by (auto simp: sub_env_def e''_def)
            moreover
            from post'' have "sub_env e'' (dfss v e'')"
              by (simp add: post_dfss_def)
            ultimately show ?thesis
              by (metis sub_env_trans)
          qed

          moreover
          from post'' stack e = pfx @ stack e'
          have "ns. stack e = ns @ (stack (dfss v e''))"
            by (auto simp: post_dfss_def e''_def)

          moreover
          {
            fix n
            assume n: "n  set (tl (stack (dfss v e'')))"
            with post'' have "𝒮 (dfss v e'') n = 𝒮 e'' n"
              by (simp add: post_dfss_def)
            moreover
            from n post'' stack e'  [] 
            have "n  set (tl (stack e''))"
              apply (simp add: post_dfss_def e''_def)
              by (metis (no_types, lifting) Un_iff list.set_sel(2) self_append_conv2 set_append tl_append2)
            with wf wvs w  visited e notexplored
            have "𝒮 e'' n = 𝒮 e n"
              by (auto simp: e''_def e' dest: unite_S_tl)
            ultimately have "𝒮 (dfss v e'') n = 𝒮 e n"
              by simp
          }

          moreover
          from post'' have "cstack (dfss v e'') = cstack e"
            by (simp add: post_dfss_def e''_def e' unite_def)

          ultimately show ?thesis
            by (simp add: dfss post_dfss_def)
        qed
      qed
    qed
  qed
qed

text ‹
  We can now show partial correctness of the algorithm:
  applied to some node @{text "v"} and the empty environment,
  it computes the set of strongly connected components in
  the subgraph reachable from node @{text "v"}. In particular,
  if @{text "v"} is a root of the graph, the algorithm computes
  the set of SCCs of the graph.
›

theorem partial_correctness:
  fixes v
  defines "e  dfs v (init_env v)"
  assumes "dfs_dfss_dom (Inl (v, init_env v))"
  shows "sccs e = {S . is_scc S  (nS. reachable v n)}"
    (is "_ = ?rhs")
proof -
  from assms init_env_pre_dfs[of v]
  have post: "post_dfs v (init_env v) e"
    by (auto dest: pre_post)
  hence wf: "wf_env e"
    by (simp add: post_dfs_def)
  from post have "cstack e = []"
    by (auto simp: post_dfs_def init_env_def)
  have "stack e = []"
  proof (rule ccontr)
    assume "stack e  []"
    hence "hd (stack e)  hd (stack e) in stack e"
      by simp
    with wf cstack e = [] show "False"
      unfolding wf_env_def
      by (metis empty_iff empty_set precedes_mem(2))
  qed
  with post have vexp: "v  explored e"
    by (simp add: post_dfs_def)
  from wf stack e = [] have "explored e = visited e"
    by (auto simp: wf_env_def)
  have "sccs e  ?rhs"
  proof
    fix S
    assume S: "S  sccs e"
    with wf have "is_scc S"
      by (simp add: wf_env_def)
    moreover
    from S wf have "S  explored e"
      unfolding wf_env_def
      by blast
    with post explored e = visited e have "nS. reachable v n"
      by (auto simp: post_dfs_def wf_env_def sub_env_def init_env_def)
    ultimately show "S  ?rhs"
      by auto
  qed
  moreover
  {
    fix S
    assume "is_scc S" "nS. reachable v n"
    from nS. reachable v n vexp wf
    have "S   (sccs e)"
      unfolding wf_env_def by (metis subset_eq)
    with is_scc S obtain S' where S': "S'  sccs e  S  S'  {}"
      unfolding is_scc_def
      by (metis Union_disjoint inf.absorb_iff2 inf_commute)
    with wf have "is_scc S'"
      by (simp add: wf_env_def)
    with S' is_scc S have "S  sccs e"
      by (auto dest: scc_partition)
  }
  ultimately show ?thesis by blast
qed

section ‹Proof of termination and total correctness›

text ‹
  We define a binary relation on the arguments of functions @{text dfs} and @{text dfss},
  and prove that this relation is well-founded and that all calls within
  the function bodies respect the relation, assuming that the pre-conditions
  of the initial function call are satisfied. By well-founded induction,
  we conclude that the pre-conditions of the functions are sufficient to
  ensure termination.

  Following the internal representation of the two mutually recursive
  functions in Isabelle as a single function on the disjoint sum of the
  types of arguments, our relation is defined as a set of argument pairs
  injected into the sum type. The left injection @{text Inl} takes
  arguments of function @{text dfs}, the right injection @{text Inr}
  takes arguments of function @{text dfss}.\footnote{Note that the
    types of the arguments of @{text dfs} and @{text dfss} are actually
    identical. We nevertheless use the sum type in order to remember
    the function that was called.}
  The conditions on the arguments in the definition of the relation
  overapproximate the arguments in the actual calls.
›

definition dfs_dfss_term::"(('v × 'v env + 'v × 'v env) × ('v × 'v env + 'v × 'v env)) set" where
  "dfs_dfss_term 
    { (Inr(v, e1), Inl(v, e)) | v e e1. 
      v  vertices - visited e  visited e1 = visited e  {v} }
   { (Inl(w, e), Inr(v, e)) | v w e. v  vertices}
   { (Inr(v, e''), Inr(v, e)) | v e e''. 
       v  vertices  sub_env e e'' 
      (w  vertices. w  vsuccs e v  w  vsuccs e'' v)}"

text ‹
  Informally, termination is ensured because at each call,
  either a new vertex is visited (hence the complement of
  the set of visited nodes w.r.t. the finite set of vertices
  decreases) or a new successor is added to the set 
  @{text "vsuccs e v"} of some vertex @{text v}.

  In order to make this argument formal, we inject the argument
  tuples that appear in our relation into tuples consisting of
  the sets mentioned in the informal argument. However, there is
  one added complication because the call of @{text dfs} from
  @{text dfss} does not immediately add the vertex to the set
  of visited nodes (this happens only at the beginning of
  function @{text dfs}). We therefore add a third component of
  $0$ or $1$ to these tuples, reflecting the fact that there
  can only be one call of @{text dfs} from @{text dfss} for a
  given vertex @{text v}.
›

fun dfs_dfss_to_tuple where
  "dfs_dfss_to_tuple (Inl(v::'v, e::'v env)) = 
   (vertices - visited e, vertices × vertices - {(u,u') | u u'. u'  vsuccs e u}, 0)"
| "dfs_dfss_to_tuple (Inr(v::'v, e::'v env)) = 
   (vertices - visited e, vertices × vertices - {(u,u') | u u'. u'  vsuccs e u}, 1::nat)"


text ‹
  The triples defined in this way can be ordered lexicographically
  (with the first two components ordered as finite subsets and the
  third one following the predecessor relation on natural numbers).
  We prove that the injection of the above relation into sets
  of triples respects the lexicographic ordering and conclude that
  our relation is well-founded.
›

lemma wf_term: "wf dfs_dfss_term"
proof -
  let ?r = "(finite_psubset :: ('v set × 'v set) set)
            <*lex*> (finite_psubset :: ((('v × 'v) set) × ('v × 'v) set) set)
            <*lex*> pred_nat"
  have "wf (finite_psubset :: ('v set × 'v set) set)"
    by (rule wf_finite_psubset)
  moreover
  have "wf (finite_psubset :: ((('v × 'v) set) × ('v × 'v) set) set)"
    by (rule wf_finite_psubset)
  ultimately have "wf ?r"
    using wf_pred_nat by blast
  moreover
  have "dfs_dfss_term  inv_image ?r dfs_dfss_to_tuple"
  proof (clarify)
    fix a b
    assume "(a,b)  dfs_dfss_term"
    hence "(v w e e''. a = Inr(v,e'')  b = Inr(v,e)  v  vertices  sub_env e e''
                       w  vertices  w  vsuccs e v  w  vsuccs e'' v)
          (v e e1. a = Inr(v,e1)  b = Inl(v,e)  v  vertices - visited e 
                    visited e1 = visited e  {v})
          (v w e. a = Inl(w,e)  b = Inr(v,e))"
         (is "?c1  ?c2  ?c3")
      by (auto simp: dfs_dfss_term_def)
    then show "(a,b)  inv_image ?r dfs_dfss_to_tuple"
    proof
      assume "?c1"
      then obtain v w e e'' where
        ab: "a = Inr(v, e'')" "b = Inr(v,e)" and
        vw: "v  vertices" "w  vertices" "w  vsuccs e'' v" "w  vsuccs e v" and
        sub: "sub_env e e''"
        by blast
      from sub have "vertices - visited e''  vertices - visited e"
        by (auto simp: sub_env_def)
      moreover
      from sub vw
      have "(vertices × vertices - {(u,u') | u u'. u'  vsuccs e'' u})
           (vertices × vertices - {(u,u') | u u'. u'  vsuccs e u})"
        by (auto simp: sub_env_def)
      ultimately show ?thesis
        using vfin ab by auto
    next
      assume "?c2  ?c3"
      with vfin show ?thesis
        by (auto simp: pred_nat_def)
    qed
  qed
  ultimately show ?thesis
    using wf_inv_image wf_subset by blast
qed

text ‹
  The following theorem establishes sufficient conditions that ensure
  termination of the two functions @{text dfs} and @{text dfss}. 
  The proof proceeds by well-founded induction using the relation 
  @{text dfs_dfss_term}. Isabelle represents the termination domains
  of the functions by the predicate @{text dfs_dfss_dom} and
  generates a theorem @{text dfs_dfss.domintros} for proving
  membership of arguments in the termination domains. The
  actual formulation is a litte technical because the mutual
  induction must again be encoded in a single induction argument
  over the sum type representing the arguments of both functions.
›

theorem dfs_dfss_termination:
  "v  vertices ; pre_dfs v e  dfs_dfss_dom(Inl(v, e))"
  "v  vertices ; pre_dfss v e  dfs_dfss_dom(Inr(v, e))"
proof -
  { fix args
    have "(case args
          of Inl(v,e)  
            v  vertices  pre_dfs v e
          |  Inr(v,e)  
             v  vertices  pre_dfss v e)
         dfs_dfss_dom args" (is "?P args  ?Q args")
    proof (rule wf_induct[OF wf_term])
      fix arg :: "('v × 'v env) + ('v × 'v env)"
      assume ih: " arg'. (arg', arg)  dfs_dfss_term  (?P arg'  ?Q arg')"
      show "?P arg  ?Q arg"
      proof
        assume P: "?P arg"
        show "?Q arg"
        proof (cases arg)
          case (Inl a)
          then obtain v e where a: "arg = Inl(v, e)" 
            using dfs.cases by metis
          with P have pre: "v  vertices  pre_dfs v e"
            by simp
          let ?e1 = "evisited := visited e  {v}, stack := v # stack e, cstack := v # cstack e"
          let ?recarg = "Inr(v, ?e1)"

          from a pre
          have "(?recarg, arg)  dfs_dfss_term"
            by (auto simp: pre_dfs_def dfs_dfss_term_def)
          moreover
          from pre have "?P ?recarg"
            by (auto dest: pre_dfs_pre_dfss)
          ultimately have "?Q ?recarg"
            using ih a by auto
          then have "?Q (Inl(v, e))"
            by (auto intro: dfs_dfss.domintros)
          then show ?thesis
            by (simp add: a)
        next
          case (Inr b)
          then obtain v e where b: "arg = Inr(v, e)" 
            using dfs.cases by metis
          with P have pre: "v  vertices  pre_dfss v e"
            by simp
          let ?sw = "SOME w. w  successors v  w  vsuccs e v"
          have "?Q (Inr(v, e))"
          proof (rule dfs_dfss.domintros)
            fix w
            assume "w  successors v"
                   "?sw  explored e"
                   "?sw  visited e"
                   "¬ dfs_dfss_dom (Inl (?sw, e))"
            show "w  vsuccs e v"
            proof (rule ccontr)
              assume "w  vsuccs e v"
              with w  successors v have sw: "?sw  successors v - vsuccs e v"
                by (metis (mono_tags, lifting) Diff_iff some_eq_imp)
              with pre ?sw  visited e have "pre_dfs ?sw e"
                by (blast intro: pre_dfss_pre_dfs)
              moreover
              from pre sw sclosed have "?sw  vertices"
                by blast
              moreover
              from pre have "(Inl(?sw,e), Inr(v,e))  dfs_dfss_term"
                by (simp add: dfs_dfss_term_def)
              ultimately have "dfs_dfss_dom (Inl(?sw,e))"
                using ih b by auto
              with ¬ dfs_dfss_dom (Inl (?sw, e)) 
              show "False" ..
            qed
          next
            let ?e' = "dfs ?sw e"
            let ?e''= "?e'vsuccs := λx. if x = v then vsuccs ?e' v  {?sw} 
                                         else vsuccs ?e' x"
            fix w
            assume asm: "w  successors v" "w  vsuccs e v"
                        "?sw  visited e" "?sw  explored e"
            from w  successors v w  vsuccs e v
            have sw: "?sw  successors v - vsuccs e v"
              by (metis (no_types, lifting) Diff_iff some_eq_imp)
            with pre ?sw  visited e have "pre_dfs ?sw e"
              by (blast intro: pre_dfss_pre_dfs)
            moreover
            from pre sw sclosed have "?sw  vertices"
              by blast
            moreover
            from pre have "(Inl(?sw, e), Inr(v,e))  dfs_dfss_term"
              by (simp add: dfs_dfss_term_def)
            ultimately have "dfs_dfss_dom (Inl(?sw, e))"
              using ih b by auto
            from this pre_dfs ?sw e have post: "post_dfs ?sw e ?e'"
              by (rule pre_post)
            hence "sub_env e ?e'"
              by (simp add: post_dfs_def)
            moreover
            have "sub_env ?e' ?e''"
              by (auto simp: sub_env_def)
            ultimately have "sub_env e ?e''"
              by (rule sub_env_trans)
            with pre ?sw  vertices sw
            have "(Inr(v, ?e''), Inr(v, e))  dfs_dfss_term"
              by (auto simp: dfs_dfss_term_def)
            moreover
            from pre post sw ?sw  visited e have "pre_dfss v ?e''"
              by (blast intro: pre_dfss_post_dfs_pre_dfss)
            ultimately show "dfs_dfss_dom(Inr(v, ?e''))"
              using pre ih b by auto
          next
            let ?e'' =  "evsuccs := λx. if x = v then vsuccs e v  {?sw} else vsuccs e x"
            fix w
            assume "w  successors v" "w  vsuccs e v"
                   "?sw  visited e" "?sw  explored e"
            with pre have "False"
              unfolding pre_dfss_def wf_env_def
              by (meson subsetD)
            thus "?Q (Inr(v, ?e''))"
              by simp 
          next
            fix w
            assume asm: "w  successors v" "w  vsuccs e v"
                        "?sw  visited e" "?sw  explored e"
            let ?e'' =  "evsuccs := λx. if x = v then vsuccs e v  {?sw} else vsuccs e x"
            let ?recarg = "Inr(v, ?e'')"

            from w  successors v w  vsuccs e v
            have sw: "?sw  successors v - vsuccs e v"
              by (metis (no_types, lifting) Diff_iff some_eq_imp)

            have "(?recarg, arg)  dfs_dfss_term"
            proof -
              have "sub_env e ?e''"
                by (auto simp: sub_env_def)
              moreover 
              from sw pre sclosed
              have "u  vertices. u  vsuccs e v  u  vsuccs ?e'' v"
                by auto
              ultimately show ?thesis
                using pre b unfolding dfs_dfss_term_def by blast
            qed
            
            moreover 
            from pre sw ?sw  explored e have "?P ?recarg"
              by (auto dest: pre_dfss_explored_pre_dfss)

            ultimately show "?Q ?recarg"
              using ih b by blast
          next
            fix w
            assume asm: "w  successors v" "w  vsuccs e v"
                        "?sw  visited e" "?sw  explored e"
            let ?eu = "unite v ?sw e"
            let ?e'' = "?euvsuccs := λx. if x = v then vsuccs ?eu v  {?sw} else vsuccs ?eu x"
            let ?recarg = "Inr(v, ?e'')"

            from w  successors v w  vsuccs e v
            have sw: "?sw  successors v - vsuccs e v"
              by (metis (no_types, lifting) Diff_iff some_eq_imp)

            have "(?recarg, arg)  dfs_dfss_term"
            proof -
              from pre asm sw have "sub_env e ?eu"
                by (blast dest: unite_sub_env)
              hence "sub_env e ?e''"
                by (auto simp: sub_env_def)
              moreover 
              from sw pre sclosed
              have "u  vertices. u  vsuccs e v  u  vsuccs ?e'' v"
                by auto
              ultimately show ?thesis
                using pre b unfolding dfs_dfss_term_def by blast
            qed

            moreover 
            from pre sw ?sw  visited e ?sw  explored e have "?P ?recarg"
              by (auto dest: pre_dfss_unite_pre_dfss)

            ultimately show "?Q ?recarg"
              using ih b by auto
          qed
          then show ?thesis
            by (simp add: b)
        qed
      qed
    qed
  }
  note dom=this
  from dom
  show " v  vertices ; pre_dfs v e  dfs_dfss_dom(Inl(v, e))"
    by auto
  from dom
  show " v  vertices ; pre_dfss v e  dfs_dfss_dom(Inr(v, e))"
    by auto
qed

text ‹
  Putting everything together, we prove the total correctness of
  the algorithm when applied to some (root) vertex.
›
theorem correctness:
  assumes "v  vertices"
  shows "sccs (dfs v (init_env v)) = {S . is_scc S  (nS. reachable v n)}"
  using assms init_env_pre_dfs[of v]
  by (simp add: dfs_dfss_termination partial_correctness)


end
end