why3doc index index
The graph is represented by a pair (vertices
, successor
)
vertices
: this constant is the set of graph vertices
successor
: this function gives for each vertex the set of vertices directly joinable from it
Notice that this proof uses paths for the soundness part, but relies on the non-black-to-white property for its completeness.
This proof is close to the one in Cormen et al.; but is simpler since
there is no need for the so-called Parenthesis theorem, nor for a
time variable. The present proof follows from the structure of the
recursion.
The formula "A && B" means "A /\ A -> B" .
The formula "A by B" means "B /\ B -> A"
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 whiteaccess (r b v: set vertex) = forall z. mem z b -> exists x l. mem x r /\ wpath x l z v predicate nbtw (b g: set vertex) = forall x x'. edge x x' -> mem x b -> mem x' (union b g) lemma whiteaccess_var: forall r r' s v. subset r r' -> whiteaccess r s v -> whiteaccess r' s v lemma whiteaccess_covar1: forall r s s' v. subset s s' -> whiteaccess r s' v -> whiteaccess r s 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 s v v'. subset v v' -> whiteaccess r s v' -> whiteaccess r s 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' s v. whiteaccess r r' v -> whiteaccess r' s v -> whiteaccess r s 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 lemma whiteaccess_minus: forall r s v. whiteaccess r s v -> whiteaccess (diff r v) s v lemma nbtw_path: forall r v. nbtw (diff r v) v -> forall x l z. mem x (diff r v) -> wpath x l z v -> mem z (diff r v) lemma nbtw_whiteaccess: forall r s v. nbtw (diff r v) v -> whiteaccess r s v -> subset s (diff r v) lemma diff_empty: forall b g r: set vertex. inter r g == empty -> diff r (union b g) == diff r b lemma whiteaccess_roots_result: forall r r' s v b g. v = union b g -> subset (diff r g) r' -> whiteaccess r s v -> whiteaccess (diff r' v) s v by whiteaccess r' s v by whiteaccess (diff r g) s v by whiteaccess (diff r v) s v
let rec dfs r g b variant {(cardinal vertices - cardinal g), (cardinal r)} = requires {subset r vertices } requires {subset b vertices } requires {subset g vertices } requires {inter b g == empty} ensures {subset result vertices } ensures {subset b result } ensures {inter result g == empty} ensures {subset (diff r g) result && nbtw (diff result (union b g)) (union b g) && forall s. whiteaccess r s (union b g) <-> subset s (diff result b) } if is_empty r then b else let x = choose r in let r' = remove x r in if mem x (union b g) then dfs r' g b else let v' = dfs (successors x) (add x g) b in let ghost v = union b g in assert {whiteaccess (add x empty) (diff v' (add x b)) v by whiteaccess (successors x) (diff v' (add x b)) (union b (add x g)) by subset b (add x b) }; assert {wpath x Nil x v }; assert {whiteaccess (add x empty) (diff (add x v') b) v by diff (add x v') b == add x (diff v' (add x b)) }; let v'' = dfs r' g (add x v') in assert {whiteaccess r' (diff v'' (add x v')) v by whiteaccess r' (diff v'' (add x v')) (union (add x v') g) by subset v (union (add x v') g) }; assert {diff v'' b == union (diff v'' (add x v')) (diff (add x v') b) }; v'' end
Generated by why3doc 0.88.3