why3doc index index



DFS in graph - soundness of the whitepath theorem

The graph is represented by a pair (vertices, successors)

The algorithm is depth-first-search in the graph as described in Sedgewick (Algorithms). It picks randomly the son on which recursive call is done.
This theorem refers to the whitepath theorem in Cormen et al.
The new visited vertices are reachable by a white path w.r.t the old visited set.
Fully automatic proof, with inductive definition of white paths, mutable set of visited vertices, black/white color. The proof is similar to the functional one.

Notice that this proof uses paths.


module DfsWhitePathSoundness
  use import int.Int
  use import list.List
  use import list.Append
  use import list.Mem as L
  use import list.Length
  use import list.HdTlNoOpt
  use import list.Reverse
  use import list.Elements as E
  use import array.Array
  use import ref.Ref
  use import init_graph.GraphListArraySucc

  predicate white_vertex (x: vertex) (v: set vertex) =
    not (mem x v)

  inductive wpath vertex (list vertex) vertex (set vertex) =
  | WPath_empty:
      forall x v. white_vertex x v -> wpath x Nil x v
  | WPath_cons:
      forall x y l z v.
      white_vertex x v -> edge x y -> wpath y l z v -> wpath x (Cons x l) z v

  predicate whiteaccess (roots b v: set vertex) =
    forall z. mem z b -> exists x l. mem x roots /\ wpath x l z v

  predicate nbtw (b v: set vertex) =
    forall x x'. edge x x' -> mem x b -> mem x' (union b v)

  lemma nbtw_wpath:
    forall v v'. nbtw (diff v' v) v' -> 
      forall x l z. mem x (diff v' v) -> wpath x l z v -> mem z (diff v' v)

  lemma path_wpathempty:
    forall y l z. wpath y l z empty -> path y l z

  lemma nbtw_path:
    forall v. nbtw v v -> 
      forall x l z. mem x v -> path x l z -> mem z v


program

  type color = WHITE | GRAY | BLACK

let rec dfs1 x col graph (ghost v) =
  requires {Array.length graph = Array.length col = cardinal vertices}
  requires {forall x. graph[x] = successors x}
  requires {mem x vertices /\ col[x] = WHITE}
  requires {forall x. mem x !v <-> mem x vertices /\ col[x] <> WHITE}
  ensures {forall x. mem x !v <-> mem x vertices /\ col[x] <> WHITE}
  ensures {subset !(old v) !v }
  ensures {mem x !v}
  ensures {nbtw (diff !v !(old v)) !v &&
           forall s. whiteaccess (add x empty) s !(old v) -> subset s (diff !v !(old v)) }
'L0:
  let ghost v0 = !v in
  assert {not mem x v0};
  let ghost col0 = Array.copy col in
  col[x] <- GRAY; v := add x !v; 
  let sons = ref graph[x] in
  let ghost dejavu = ref Nil in
  while !sons <> Nil do
    invariant {subset (elements !sons) vertices}
    invariant {subset !v vertices}
    invariant {forall z. mem z !v <-> mem z vertices /\ col[z] <> WHITE}
    invariant {subset (add x v0) !v}
    invariant {subset (elements !dejavu) !v}
    invariant {nbtw (diff !v (add x v0)) !v}
    invariant {successors x = (reverse !dejavu) ++ !sons}
    let y = hd !sons in
      if col[y] = WHITE then dfs1 y col graph v;
    sons := tl !sons;
    dejavu := Cons y !dejavu;
  done;
'L1:
  let ghost v' = !v in
  assert {diff v' v0 == add x (diff v' (add x v0))};
  col[x] <- BLACK
    
let dfs_main roots graph =
  requires {Array.length graph = cardinal vertices}
  requires {forall x. graph[x] = successors x}
  requires {subset (elements roots) vertices}
  ensures {subset (elements roots) (elements result)}
  ensures {nbtw (elements result) (elements result) 
   so forall x l z. mem x (elements roots) -> path x l z -> mem z (elements result)}
  let n = Array.length graph in
  let col = Array.make n WHITE in
  let ghost v = ref empty in
  let xs = ref roots in
  let ghost dejavu = ref Nil in
  while !xs <> Nil do
    invariant {subset (elements !xs) vertices}
    invariant {subset !v vertices}
    invariant {forall z. mem z !v <-> mem z vertices /\ col[z] <> WHITE}
    invariant {nbtw !v !v}
    invariant {subset (elements !dejavu) !v}
    invariant {roots = (reverse !dejavu) ++ !xs}
    let y = hd !xs in
       if col[y] = WHITE then dfs1 y col graph v;
    xs := tl !xs;
    dejavu := Cons y !dejavu;
  done;
  assert {subset (elements roots) !v};
  let res = ref Nil in
  for x = 0 to n - 1  do
    invariant {forall y. L.mem y !res <-> 0 <= y < x /\ col[y] <> WHITE}
    if col[x] <> WHITE then res := Cons x !res
  done;
  assert {elements !res == !v};
  !res

[session]   [zip]


extracted program


(* let rec dfs1 x col graph =
  col[x] <- GRAY; 
  let sons = ref graph[x] in
  while !sons <> Nil do
    let y = hd !sons in
      if col[y] = WHITE then dfs1 y col;
    sons := tl !sons;
  done;
  col[x] <- BLACK
    
let dfs_main roots graph =
  let n = Array.length graph in
  let col = Array.make n WHITE in
  let xs = ref roots in
  while !xs <> Nil do
    let y = hd !xs in
      if col[y] = WHITE then dfs1 y col;
    xs := tl !xs;
  done;
  let res = ref Nil in
  for x = 0 to n - 1  do
    if col[x] <> WHITE then res := Cons x !res
  done;
  !res
*)

end

Generated by why3doc 0.88.3