why3doc index index



DFS in graph - the whitepath theorem

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

The algorithm is depth-first-search in the graph. 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 exactly the vertices reachable by a white path w.r.t the old visited set.
Fully automatic proof, with inductive definition of white paths and heavy use of the 'by' why3 connector.

Notice that this proof avoids usage of paths and relies on the non-black-to-white predicate.
This proof is close to the one of Cormen et al.; it is simpler since there is no need for the well-parenthesis lemma, since it just follows from the structure of the recursion.


module DfsWhitePathThm
  use import int.Int
  use import list.List
  use import list.Append
  use import list.Mem as L
  use import list.Elements as E
  use import init_graph.GraphSetSucc

  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 whitepath (x: vertex) (l: list vertex) (z: vertex) (v: set vertex) =
    path x l z /\ (forall y. L.mem y l -> white_vertex y v) /\ white_vertex 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

  lemma whiteaccess_var:
    forall r r' b v. subset r r' -> whiteaccess r b v -> whiteaccess r' b v

  lemma whiteaccess_covar1:
    forall r b b' v. subset b b' -> whiteaccess r b' v -> whiteaccess r b v

  lemma wpath_covar2:
    forall x l z v v'. subset v v' -> wpath x l z v' -> wpath x l z v

  lemma whiteaccess_covar2:
    forall r b v v'. subset v v' -> whiteaccess r b v' -> whiteaccess r b v

  lemma wpath_trans:
    forall x l y l' z v. wpath x l y v -> wpath y l' z v -> wpath x (l ++ l') z v

  lemma whiteaccess_trans:
    forall r r' b v. whiteaccess r r' v -> whiteaccess r' b v -> whiteaccess r b v

  lemma whiteaccess_cons:
    forall x s v. mem x vertices ->
      white_vertex x v -> whiteaccess (successors x) s v -> whiteaccess (add x empty) s v

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

  lemma nbtw_path:
    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)

program

	  
let rec dfs r v 
  variant {(cardinal vertices - cardinal v), (cardinal r)} =
  requires {subset r vertices }
  requires {subset v vertices }
  ensures {subset result vertices}
  ensures {subset v result }
  ensures {subset r result }
  ensures {nbtw (diff result v) result &&
           forall s. whiteaccess r s v <-> subset s (diff result v)}
  if is_empty r then v else
  let x = choose r in
  let r' = remove x r in
  if mem x v then dfs r' v else
  let v' = dfs (successors x) (add x v) in
  assert {whiteaccess (add x empty) (diff v' (add x v)) v
    by whiteaccess (successors x) (diff v' (add x v)) v
    by subset (diff v' (add x v)) (diff v' v)
    by subset v (add x v) };
  assert {wpath x Nil x v};
  assert {whiteaccess (add x empty) (diff v' v) v
    by diff v' v == add x (diff v' (add x v))};
  let v'' = dfs r' v' in
  assert {diff v'' v == union (diff v'' v') (diff v' v) };
  assert {whiteaccess r' (diff v'' v') v
    by subset v v'};
  v''

end

[session]



Generated by why3doc 0.88.3