why3doc index index



Random search in graph

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

The algorithm is a ministep of a random search in the graph. It is generic to any search strategy.
This soundness proof is adapted from Dowek and Munoz and is fully automatic.


module RandomSearch
  use import int.Int
  use import list.List
  use import list.Mem as L
  use import init_graph.GraphSetSucc

  predicate white_vertex (x: vertex) (v: set vertex) =
    not (mem x 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

  lemma whitepath_id:
    forall x v. not mem x v -> whitepath x Nil x v

  lemma whitepath_covar:
    forall v v' x l y. subset v v' -> whitepath x l y v' -> whitepath x l y v

  lemma whitepath_cons:
    forall x x' y l v. not mem x v -> edge x x' -> whitepath x' l y v -> whitepath x (Cons x l) y v

program


let rec random_search r v
  variant {(cardinal vertices - cardinal v), (cardinal r)} = 
  requires {subset r vertices }
  requires {subset v vertices }
  ensures {forall z. mem z (diff result v) -> exists y l. mem y r /\ whitepath y l z v}

  if is_empty r then v else
  let x = choose r in
  let r' = remove x r in
  if mem x v then random_search r' v else
  let v' = random_search (union r' (successors x)) (add x v) in
  assert {forall z. mem z (diff v' v) -> z <> x ->
      (exists y l. mem y r' /\ (whitepath y l z v by whitepath y l z (add x v)))
   \/ ((exists l. whitepath x l z v)
       by  exists x' l'. edge x x' /\ whitepath x' l' z v)};
   v'
      

let random_search_main (roots: set vertex) =
   requires {subset roots vertices}
   ensures {forall z. mem z result -> exists y l. mem y roots /\ path y l z}
   random_search roots empty 

end

[session]



Generated by why3doc 0.88.3