Library frag_list_vcg

Load seplog_header.

Require Import Omega.
Require Import Bool.
Require Import frag_list_entail.
Require Import frag_list_triple.

Inductive cmd' : Type :=
| skip' : cmd'
| assign' : var.v -> expr -> cmd'
| lookup' : var.v -> expr -> cmd'
| mutation' : expr -> expr -> cmd'
| malloc' : var.v -> expr -> cmd'
| free' : expr -> cmd'
| while' : expr_b -> Assrt -> cmd' -> cmd'
| seq' : cmd' -> cmd' -> cmd'
| ifte' : expr_b -> cmd' -> cmd' -> cmd'.

Notation "x <- e" := (assign' x e) (at level 80) : frag2_scope.
Notation "x '<-*' e" := (lookup' x e) (at level 80) : frag2_scope.
Notation "e '*<-' f" := (mutation' e f) (at level 80) : frag2_scope.
Notation "x '<-malloc' e" := (malloc' x e) (at level 80) : frag2_scope.
Notation "c ; d" := (seq' c d) (at level 81, right associativity) : frag2_scope.
Notation "'ifte' a 'thendo' x 'elsedo' y" := (ifte' a x y) (at level 80) : frag2_scope.

Fixpoint proj_cmd (c':cmd') : cmd :=
  match c' with
    skip' => skip
    | assign' x e => x <- e
    | lookup' x e => x <-* e
    | mutation' e f => e *<- f
    | malloc' x e => x <-malloc e
    | free' e => free e
    | while' b Q c'' => while b (proj_cmd c'')
    | seq' c1 c2 => (seq (proj_cmd c1) (proj_cmd c2))
    | ifte' b c1 c2 => (ifte b thendo (proj_cmd c1) elsedo (proj_cmd c2))
  end.

Fixpoint Assrt_and_expr_b (A: Assrt) (b: expr_b) {struct A} : Assrt :=
  match A with
    | nil => nil
    | (pi,sig)::tl => (pi &&& b, sig) :: (Assrt_and_expr_b tl b)
  end.

Lemma Assrt_and_expr_b_sem: forall A b s h,
  (Assrt_interp (Assrt_and_expr_b A b) s h) -> (Assrt_interp A s h /\ eval_b b s = true).

Lemma Assrt_and_expr_b_sem': forall A b s h,
  (Assrt_interp A s h /\ eval_b b s = true) -> (Assrt_interp (Assrt_and_expr_b A b) s h).

Fixpoint vcg (c: cmd') (Q: wpAssrt) {struct c} : option (wpAssrt * (list (Assrt * wpAssrt))) :=
  match c with
    skip' => Some (Q, nil)
    | assign' x e => Some (wpSubst ((x,e)::nil) Q, nil)
    | lookup' x e => Some (wpLookup x e Q, nil)
    | mutation' e f => Some (wpMutation e f Q, nil)
    | seq' c1 c2 =>
      match vcg c2 Q with
        | None => None
        | Some (Q2,l2) =>
          match vcg c1 Q2 with
            | None => None
            | Some (Q1, l1) =>
              Some (Q1, l1 ++ l2)
          end
      end
    | ifte' b c1 c2 =>
      match vcg c2 Q with
        | None => None
        | Some (Q2,l2) =>
          match vcg c1 Q with
            | None => None
            | Some (Q1,l1) => Some (wpIf b Q1 Q2, (l1 ++ l2))
          end
      end
    | while' b Inv c' =>
      match vcg c' (wpElt Inv) with
        | None => None
        | Some (Q',l') => Some (wpElt Inv,           (Assrt_and_expr_b Inv (neg_b b), Q)           :: (Assrt_and_expr_b Inv b, Q')           :: l')
      end
    | _ => None
  end.

Lemma vcg_None_is_None: forall c, wp_frag None c = None.

Lemma vcg_correct : forall c Q Q' l,
  vcg c Q = Some (Q', l) ->
  (forall A L,
    In (A,L) l ->
    ((Assrt_interp A) ==> (wpAssrt_interp L))
  ) ->
  {{ wpAssrt_interp Q' }} (proj_cmd c) {{ wpAssrt_interp Q }}.









Fixpoint triple_transformations (l: list (Assrt * wpAssrt)) : option (list ((Pi * Sigma) * wpAssrt)) :=
  match l with
    | nil => Some nil
    | (A,L)::tl =>
      match triple_transformation A L with
        | Some l =>
          match triple_transformations tl with
            | Some l' => Some (l ++ l')
            | None => None
          end
        | None => None
      end
  end.

Lemma triple_transformations_correct: forall l,
  triple_transformations l = Some nil ->
  forall A L, In (A,L) l -> (Assrt_interp A) ==> (wpAssrt_interp L).

Definition bigtoe_fct (c: cmd') (P Q: Assrt): option (list ((Pi * Sigma) * wpAssrt)) :=
  match vcg c (wpElt Q) with
    | None => None
    | Some (Q', l) =>
      match triple_transformation P Q' with
        | Some l' =>
          match triple_transformations l with
            | Some l'' => Some (l' ++ l'')
            | None => None
          end
        | None => None
      end
  end.

Lemma bigtoe_fct_correct: forall P Q c,
  bigtoe_fct c P Q = Some nil ->
  {{ Assrt_interp P }} proj_cmd c {{ Assrt_interp Q }}.

Fixpoint resolve_list_Assrt_wpAssrt2 (l: list (Assrt * wpAssrt)) : bool :=
  match l with
    | nil => true
    | (A,L)::tl => andb (triple_transformation2 A L) (resolve_list_Assrt_wpAssrt2 tl)
  end.

Lemma resolve_list_Assrt_wpAssrt2_correct: forall l,
  resolve_list_Assrt_wpAssrt2 l = true ->
  forall A L, In (A,L) l -> Assrt_interp A ==> wpAssrt_interp L.

Definition frag2_hoare (c: cmd') (P Q: Assrt): bool :=
  match vcg c (wpElt Q) with
    | None => false
    | Some (Q',l) =>
      andb (triple_transformation2 P Q') (resolve_list_Assrt_wpAssrt2 l)
  end.

Lemma frag2_hoare_correct: forall P Q c,
  frag2_hoare c P Q = true ->
  {{ Assrt_interp P }} (proj_cmd c) {{ Assrt_interp Q }}.