Library examples

Load seplog_header.

Open Local Scope Z_scope.

Module GCD.

Require Import Znumtheory.

Definition gcd0 a b x y :=
  x <- int_e a;
  y <- int_e b;
  while (var_e x =/= var_e y) (
    ifte (var_e y >> var_e x) thendo
    (y <- var_e y -e var_e x)
    elsedo
    (x <- var_e x -e var_e y)
  ).

Definition x := 0%nat.
Definition y := 1%nat.

Lemma gcd0_verif : forall a b,
  {{ fun s h => 0 < a /\ 0 < b }}
  gcd0 a b x y
  {{ fun s h => exists d, store.lookup x s = d /\ Zis_gcd a b d }}.

Definition a := 0%nat.
Definition b := 1%nat.
Definition r := 2%nat.

Definition gcd1 a b r :=
  while (var_e b =/= int_e 0) (
    r <- (var_e a) mode (var_e b);
    a <- var_e b;
    b <- var_e r
  ).

Lemma gcd1_verif : forall va vb, 0 <= va -> 0 <= vb ->
  {{ fun s h => store.lookup a s = va /\ store.lookup b s = vb }}
  gcd1 a b r
  {{ fun s h => exists wa, exists wb, exists d,
    store.lookup a s = wa /\ store.lookup b s = wb /\
    Zis_gcd wa wb d /\ Zis_gcd va vb d }}.

End GCD.

Module Factorial.

Definition Zfact (z:Z) : Z :=
  match z with
    Zpos p => Z_of_nat (fact (nat_of_P p))
    | _ => 1
  end.

Lemma Zfact_neg : forall z, z < 0 -> Zfact z = 1.

Lemma Zfact_zero : Zfact 0 = 1.

Lemma fact_lemma : forall n, (0 < n)%nat -> (fact n = n * fact (n - 1))%nat.

Lemma Zfact_pos : forall z, z > 0 -> Zfact z = z * Zfact (z - 1).

Lemma factorial : forall n, 0 <= n ->
  forall x m, var.set (x::m::nil) ->
    {{ fun s h => store.lookup m s = n /\ store.lookup x s = 1 }}
    while (var_e m =/= int_e 0%Z) (
      x <- var_e x *e var_e m;
      m <- var_e m -e int_e 1
    )
    {{ fun s h => store.lookup x s = Zfact n }}.

Open Local Scope vc_scope.

Lemma vc_factorial : forall n, n >= 0 ->
  forall x m, var.set (x::m::nil) ->
    {{ fun s (_:heap.h) => store.lookup m s = n /\ store.lookup x s = 1 }} proj_cmd
    (while' ((var_e m) =/= int_e 0)
      (fun s (_:heap.h) => store.lookup x s * Zfact (store.lookup m s) = Zfact n /\ store.lookup m s >= 0)
      (x <- var_e x *e var_e m;
        m <- var_e m -e int_e 1))
    {{ fun s (_:heap.h) => store.lookup x s = Zfact n }}.

Close Local Scope vc_scope.

End Factorial.

Module Swap.

Lemma swap: forall x y z v a b, var.set (x::y::z::v::nil) ->
  {{ (var_e x |-> int_e a) ** (var_e y |-> int_e b) }}
  z <-* var_e x;
  v <-* var_e y;
  var_e x *<- var_e v;
  var_e y *<- var_e z
  {{ (var_e x |-> int_e b) ** (var_e y |-> int_e a) }}.

Open Local Scope vc_scope.

Lemma vc_swap: forall x y z v a b, var.set (x::y::z::v::nil) ->
  {{ (var_e x |-> int_e a) ** (var_e y |-> int_e b) }}
  proj_cmd
  (z <-*var_e x;
   v <-* var_e y;
   var_e x *<- var_e v;
   var_e y *<- var_e z)
  {{ (var_e x |-> int_e b) ** (var_e y |-> int_e a) }}.

Close Local Scope vc_scope.

End Swap.

Module StringCopy.

Definition StringCopy c1 c2 buf str str_tmp :=
  c1 <- var_e buf;
  c2 <- var_e str;
  str_tmp <-* var_e c2;
  while (var_e str_tmp =/= int_e 0) (
    var_e c1 *<- var_e str_tmp;
    c1 <- var_e c1 +e int_e 1;
    c2 <- var_e c2 +e int_e 1;
    str_tmp <-* var_e c2
    );
  var_e c1 *<- var_e str_tmp.


Lemma fold_right_app_map_nil : forall (A:Set) (lst:list A) (B:Set),
  fold_right (@app B) nil (map (fun _ => nil) lst) = nil.

Lemma in_map' : forall (A:Set) (lst:list A) (B:Set) x (f:A->B),
  In x (map f lst) -> exists x', In x' lst /\ x = f x'.

Lemma map_length : forall (A:Set) (lst:list A) (B:Set) (f:A->B),
  length (map f lst) = length lst.

Lemma del_heads_destruct: forall (A: Set) (l: list A) i a,
  (length l > i)%nat ->
  del_heads l i = nth i l a :: del_heads l (i+1)%nat.
Implicit Arguments del_heads_destruct [A].

Lemma heads_destruct: forall (A: Set) (l: list A) i a,
  (length l > i)%nat ->
  heads l (i+1) = heads l i ++ (nth i l a)::nil.
Implicit Arguments heads_destruct [A].

Lemma nth_last: forall (A:Set) (l: list A) n a,
  n = length l ->
  nth n l a = a.

Lemma heads_last : forall l i,
  l <> nil ->
  i = (length l - 1)%nat ->
  l = heads l i ++ (nth i l (-1))::nil.

Lemma map_expr_var_list_Z : forall (lst:list expr),
  (forall e, In e lst -> expr_var e = nil) ->
  map expr_var lst = map (fun _ => nil) lst.

Lemma mapstos_equiv : forall l s h e1 e3,
  (e1 |--> l) s h ->
  eval e1 s = eval e3 s ->
  (e3 |--> l) s h.

Definition mapstos' (e:expr) (lst:list Z) :=
  sep.mapstos e (map (fun x => int_e x) lst).

Notation "x '|--->' l" := (mapstos' x l) (at level 80).

Lemma inde_mapstos' : forall lst l p,
  inter (expr_var p) l nil ->
  inde l (p |---> lst).

Lemma mapstos'_app_sepcon: forall l1 l2 st s h,
  (st |---> l1 ++ l2) s h ->
  ((st |---> l1) ** ((st +e (nat_e (length l1))) |---> l2)) s h.

Lemma mapstos'_cons_sepcon: forall a l st s h,
  (st |---> a::l) s h ->
  ((st |-> int_e a) ** ((st +e (nat_e 1)) |---> l)) s h.

Lemma mapstos'_sepcon_app: forall l1 l2 st s h,
  ((st |---> l1) ** ((st +e (nat_e (length l1))) |---> l2)) s h ->
  (st |---> l1 ++ l2) s h.

Definition string' (lst:list nat) := ~ In O lst.

Definition string (lst:list Z) :=
  exists lst',
    string' lst' /\
    lst = (map (fun x => Z_of_nat x) lst') ++ (0::nil).

Lemma string_nil : ~ string nil.

Lemma string'_sub : forall lst,
  string' lst ->
  string' (tail lst).

Lemma string_sub : forall lst,
  (2 <= length lst)%nat ->
  string lst ->
  string (tail lst).

Lemma string_sup : forall lst,
  string lst ->
  forall lst',
    string' lst' ->
    string (map (fun x => Z_of_nat x) lst' ++ lst).

Lemma string_last : forall lst, string lst ->
  forall i, nth i lst (-1) = 0 ->
    (i = length lst - 1)%nat.

Lemma string_hd_ge0: forall a l,
  string (a::l) ->
  a >= 0.

Lemma string_last' : forall i lst,
  string lst ->
  (i = length lst - 1)%nat ->
  nth i lst (-1) = 0.

Lemma string_last'' : forall i lst,
  string lst ->
  (i < length lst)%nat ->
  nth i lst (-1) >= 0.

Definition c1 := O.
Definition c2 := 1%nat.
Definition buf := 2%nat.
Definition str := 3%nat.
Definition str_tmp := 4%nat.

Hint Unfold c1 c2 buf str str_tmp.

Lemma StringCopy_specif : forall buf_lst str_lst
  (Hbuf: (0 < length buf_lst)%nat)
  (Hbuf2: (length str_lst <= length buf_lst)%nat)
  (Hstr: string str_lst),
  {{ (var_e buf |---> buf_lst) ** (var_e str |---> str_lst) }}
  StringCopy c1 c2 buf str str_tmp
  {{ (var_e buf |---> str_lst) ** TT ** (var_e str |---> str_lst) }}.



















  Eval_b_hyp H1.
  Eval_b_hyp H1.

  Opaque nth.
  Transparent nth.












  Opaque nth.
  Transparent nth.
  Opaque nth.
  Transparent nth.



  Eval_b_hyp_clean.






  Opaque nth.
  Transparent nth.




End StringCopy.