Library bipl

Require Import List.
Require Import ZArith.
Require Import List.
Require Import EqNat.
Require Import Classical.
Require Import Bool.

Require Import util.
Require Import heap.

Module Type VAR.
  Parameter v : Set.
  Parameter eqdec : forall (x y : v), {x=y} + {x <> y}.
  Parameter fresh : v -> list v -> Prop.
  Parameter set : list v -> Prop.
End VAR.

Module var <: VAR with Definition v := nat.
  Definition v := nat.
  Definition eqdec := eq_nat_dec.
  Fixpoint fresh (x:v) (he:list v) {struct he} : Prop :=
    match he with
      | nil => True
      | hd::tl => x <> hd /\ fresh x tl
    end.

  Fixpoint set (he:list v) : Prop :=
    match he with
      nil => True
      | hd::tl => fresh hd tl /\ set tl
    end.
End var.

Ltac Var_uneq :=
   match goal with
      | id: ?v <> ?v' |- ?v <> ?v' => auto
      | id: ?v' <> ?v |- ?v <> ?v' => auto
      | id: var.fresh ?v ?l |- ?v <> ?v' => let x := fresh in
                    (inversion_clear id as [x X]; decompose [and] X; clear X); Var_uneq
      | id: var.fresh ?v' ?l |- ?v <> ?v' => let x := fresh in
                    (inversion_clear id as [x X]; decompose [and] X; clear X); Var_uneq
      | id: var.set ?l |- ?v <> ?v' =>
             let x := fresh in
                    (inversion_clear id as [x X]; decompose [and] X; clear X); Var_uneq
      | _ => fail
   end.

Module Type STORE.
  Parameter s : Set.
  Parameter emp : s.
  Parameter lookup : var.v -> s -> val.
  Parameter update : var.v -> val -> s -> s.
  Parameter lookup_emp : forall x,
    lookup x emp = 0%Z.
  Parameter lookup_update : forall x y z st,
    x <> y -> lookup x st = lookup x (update y z st).
  Parameter lookup_update' : forall x z st,
    lookup x (update x z st) = z.
  Parameter extensionality: forall st1 st2, (forall x, lookup x st1 = lookup x st2) -> st1 = st2.
  Parameter update_update: forall s x x' v v',
        x <> x' ->
        update x v (update x' v' s) = update x' v' (update x v s).
  Parameter update_update': forall s x v v',
        update x v (update x v' s) = update x v s.
  Parameter update_lookup : forall s x,
    update x (lookup x s) s = s.
  Parameter update_lookup_update: forall x v s,
    update x (lookup x s) (update x v s) = s.
End STORE.

Module non_null_integer.
  Inductive negpos : Set :=
    neg : positive -> negpos
    | pos : positive -> negpos.
  Definition elem := negpos.
End non_null_integer.

Module store_heap := map Loc non_null_integer.

Module store : STORE.
  Definition s := store_heap.h.

  Definition emp : s := store_heap.emp.

  Import non_null_integer.


  Definition lookup (w:var.v) (st:s) : Z :=
    match store_heap.lookup w st with
      Some (pos x) => Zpos x
      | Some (neg x) => Zneg x
      | None => Z0
    end.

  Definition update (i:var.v) (w:Z) (st:s) : s :=
    match w with
      Z0 => store_heap.dif st i
      | Zneg p => store_heap.union (store_heap.singleton i (neg p)) st
      | Zpos p => store_heap.union (store_heap.singleton i (pos p)) st
    end.

  Lemma lookup_emp : forall x,
    lookup x emp = 0%Z.

  Lemma lookup_update : forall x y z st,
    x <> y -> lookup x st = lookup x (update y z st).

  Lemma lookup_update' : forall w z st,
    lookup w (update w z st) = z.

  Lemma extensionality: forall st1 st2,
    (forall x, lookup x st1= lookup x st2) -> st1 = st2.

  Lemma update_update: forall s x x' v v', x <> x' ->
    update x v (update x' v' s) = update x' v' (update x v s).

  Lemma update_update': forall s x v v',
    update x v (update x v' s) = update x v s.

  Lemma update_lookup : forall s x,
    update x (lookup x s) s = s.

  Lemma update_lookup_update: forall x v s,
    update x (lookup x s) (update x v s) = s.

End store.

Inductive expr : Set :=
  var_e : var.v -> expr
| int_e : val -> expr
| add_e : expr -> expr -> expr
| min_e : expr -> expr -> expr
| mul_e : expr -> expr -> expr
| div_e : expr -> expr -> expr
| mod_e : expr -> expr -> expr.

Definition nat_e x := int_e (Z_of_nat x).

Definition null := int_e 0%Z.

Notation "e1 '+e' e2" := (add_e e1 e2) (at level 79) : sep_scope.
Notation "e1 '-e' e2" := (min_e e1 e2) (at level 79) : sep_scope.
Notation "e1 '*e' e2" := (mul_e e1 e2) (at level 79) : sep_scope.
Notation "e1 '/e' e2" := (div_e e1 e2) (at level 79) : sep_scope.
Notation "e1 'mode' e2" := (mod_e e1 e2) (at level 79) : sep_scope.

Open Local Scope sep_scope.

Fixpoint expr_var (e:expr) {struct e} : list var.v :=
  match e with
    var_e x => x::nil
    | int_e z => nil
    | add_e e1 e2 => expr_var e1 ++ expr_var e2
    | min_e e1 e2 => expr_var e1 ++ expr_var e2
    | mul_e e1 e2 => expr_var e1 ++ expr_var e2
    | div_e e1 e2 => expr_var e1 ++ expr_var e2
    | mod_e e1 e2 => expr_var e1 ++ expr_var e2
  end.

Fixpoint expr_eq (e1:expr) (e2:expr) {struct e1} : bool :=
  match e1 with
    var_e w1 => match e2 with
                  var_e w2 => beq_nat w1 w2
                  | _ => false
                end
    | int_e i1 => match e2 with
                    int_e i2 => Zeq_bool i1 i2
                    | _ => false
                  end
    | add_e e11 e12 => match e2 with
                         add_e e21 e22 => andb (expr_eq e11 e21) (expr_eq e12 e22)
                         | _ => false
                       end
    | min_e e11 e12 => match e2 with
                         min_e e21 e22 => andb (expr_eq e11 e21) (expr_eq e12 e22)
                         | _ => false
                       end
    | mul_e e11 e12 => match e2 with
                         mul_e e21 e22 => andb (expr_eq e11 e21) (expr_eq e12 e22)
                         | _ => false
                       end
    | div_e e11 e12 => match e2 with
                         div_e e21 e22 => andb (expr_eq e11 e21) (expr_eq e12 e22)
                         | _ => false
                       end
    | mod_e e11 e12 => match e2 with
                         mod_e e21 e22 => andb (expr_eq e11 e21) (expr_eq e12 e22)
                         | _ => false
                       end
  end.

Fixpoint expr_rewrite (e:expr) (patern:expr) (rep:expr) {struct e} : expr :=
  match e with
    var_e w => match expr_eq e patern with
                 true => rep
                 | false => e
               end
    | int_e i => match expr_eq e patern with
                   true => rep
                   | false => e
                 end
    | add_e e1 e2 => match expr_eq e patern with
                       true => rep
                       | false => add_e (expr_rewrite e1 patern rep) (expr_rewrite e2 patern rep)
                     end
    | min_e e1 e2 => match expr_eq e patern with
                       true => rep
                       | false => min_e (expr_rewrite e1 patern rep) (expr_rewrite e2 patern rep)
                     end
    | mul_e e1 e2 => match expr_eq e patern with
                       true => rep
                       | false => mul_e (expr_rewrite e1 patern rep) (expr_rewrite e2 patern rep)
                     end
    | div_e e1 e2 => match expr_eq e patern with
                       true => rep
                       | false => div_e (expr_rewrite e1 patern rep) (expr_rewrite e2 patern rep)
                     end
    | mod_e e1 e2 => match expr_eq e patern with
                       true => rep
                       | false => mod_e (expr_rewrite e1 patern rep) (expr_rewrite e2 patern rep)
                     end
end.
Fixpoint eval (e:expr) (s:store.s) {struct e} : Z :=
match e with
  var_e w => store.lookup w s
  | int_e i => i
  | add_e e1 e2 => Zplus (eval e1 s) (eval e2 s)
  | min_e e1 e2 => Zminus (eval e1 s) (eval e2 s)
  | mul_e e1 e2 => Zmult (eval e1 s) (eval e2 s)
  | div_e e1 e2 => Zdiv (eval e1 s) (eval e2 s)
  | mod_e e1 e2 => Zmod (eval e1 s) (eval e2 s)
end.

Lemma eval_store_update: forall e x v s,
  eval e (store.update x (eval v s) s) = eval (expr_rewrite e (var_e x) v) s.

Lemma inde_expr : forall e l,
  inter (expr_var e) l nil ->
  forall x, In x l ->
    forall s z,
      eval e s = eval e (store.update x z s).

Lemma expr_inde_var: forall e s v x,
  ~In x (expr_var e) ->
  eval e s = eval e (store.update x v s).

Definition field x f := var_e x +e int_e f.
Notation "x '-.>' f " := (field x f) (at level 79) : sep_scope.

Inductive expr_b : Set :=
  true_b: expr_b
| eq_b : expr -> expr -> expr_b
| neq_b : expr -> expr -> expr_b
| ge_b : expr -> expr -> expr_b
| gt_b : expr -> expr -> expr_b
| neg_b : expr_b -> expr_b
| and_b : expr_b -> expr_b -> expr_b
| or_b : expr_b -> expr_b -> expr_b.

Notation "e == e'" := (eq_b e e') (at level 78) : sep_scope.
Notation "e =/= e'" := (neq_b e e') (at level 78) : sep_scope.
Notation "e &&& e'" := (and_b e e') (at level 78) : sep_scope.
Notation "e ||| e'" := (or_b e e') (at level 78) : sep_scope.
Notation "e >>= e'" := (ge_b e e') (at level 78) : sep_scope.
Notation "e >> e'" := (gt_b e e') (at level 78) : sep_scope.

Fixpoint expr_b_var (e:expr_b) : list var.v :=
  match e with
    true_b => nil
    | eq_b e1 e2 => expr_var e1 ++ expr_var e2
    | neq_b e1 e2 => expr_var e1 ++ expr_var e2
    | ge_b e1 e2 => expr_var e1 ++ expr_var e2
    | gt_b e1 e2 => expr_var e1 ++ expr_var e2
    | and_b e1 e2 => expr_b_var e1 ++ expr_b_var e2
    | or_b e1 e2 => expr_b_var e1 ++ expr_b_var e2
    | neg_b e => expr_b_var e
  end.

Fixpoint expr_b_eq (e1: expr_b) (e2: expr_b) {struct e1} : bool :=
match e1 with
  true_b => match e2 with
              true_b => true
              | _ => false
            end
  | f == g => match e2 with
                f' == g' => andb (expr_eq f f') (expr_eq g g')
                | _ => false
              end
  | f =/= g => match e2 with
                 f' =/= g' => andb (expr_eq f f') (expr_eq g g')
                 | _ => false
               end
  | f >>= g => match e2 with
                 f' >>= g' => andb (expr_eq f f') (expr_eq g g')
                 | _ => false
               end
  | f >> g => match e2 with
                f' >> g' => andb (expr_eq f f') (expr_eq g g')
                | _ => false
              end
  | f &&& g => match e2 with
                 f' &&& g' => andb (expr_b_eq f f') (expr_b_eq g g')
                 | _ => false
               end
  | f ||| g => match e2 with
                 f' ||| g' => andb (expr_b_eq f f') (expr_b_eq g g')
                 | _ => false
               end
  | neg_b e => match e2 with
                  neg_b e' => (expr_b_eq e e')
                  | _ => false
                end
end.

Fixpoint expr_b_rewrite (e: expr_b) (patern: expr) (rep: expr) {struct e}: expr_b :=
  match e with
    true_b => true_b
    | f == g => expr_rewrite f patern rep == expr_rewrite g patern rep
    | f =/= g => expr_rewrite f patern rep =/= expr_rewrite g patern rep
    | f >>= g => expr_rewrite f patern rep >>= expr_rewrite g patern rep
    | f >> g => expr_rewrite f patern rep >> expr_rewrite g patern rep
    | f &&& g => expr_b_rewrite f patern rep &&& expr_b_rewrite g patern rep
    | f ||| g => expr_b_rewrite f patern rep ||| expr_b_rewrite g patern rep
    | neg_b e => neg_b (expr_b_rewrite e patern rep)
  end.

Fixpoint eval_b (e:expr_b) (s:store.s) {struct e} : bool :=
match e with
  true_b => true
  | f == g => Zeq_bool (eval f s) (eval g s)
  | f =/= g => negb (Zeq_bool (eval f s) (eval g s))
  | f >>= g => Zge_bool (eval f s) (eval g s)
  | f >> g => Zgt_bool (eval f s) (eval g s)
  | f &&& g => andb (eval_b f s) (eval_b g s)
  | f ||| g => orb (eval_b f s) (eval_b g s)
  | neg_b e => negb (eval_b e s)
end.

Lemma eval_b_store_update: forall b x v s,
  eval_b b (store.update x (eval v s) s) = eval_b (expr_b_rewrite b (var_e x) v) s.

Fixpoint expr_b_sem (e: expr_b) (s: store.s) {struct e} : Prop :=
match e with
  true_b => True
  | f == g => eval f s = eval g s
  | f =/= g => (eval f s <> eval g s)
  | f >>= g => (eval f s >= eval g s)%Z
  | f >> g => (eval f s > eval g s)%Z
  | f &&& g => expr_b_sem f s /\ expr_b_sem g s
  | f ||| g => expr_b_sem f s \/ expr_b_sem g s
  | neg_b e => ~ expr_b_sem e s
end.

Lemma expr_b_inde_var: forall b s v x,
   ~In x (expr_b_var b) ->
   eval_b b s = eval_b b (store.update x v s).

Lemma expr_b_semantic_good : forall e s,
  eval_b e s = true <-> expr_b_sem e s.









Lemma expr_b_false_negb_true: forall b s,
    eval_b b s = false -> eval_b (neg_b b) s = true.

Lemma expr_b_true_negb_false: forall b s,
    eval_b (neg_b b) s = true -> eval_b b s = false.

Open Local Scope heap_scope.

Definition assert := store.s -> heap.h -> Prop.

Definition TT : assert := fun s h => True.

Definition FF : assert := fun s h => False.

Definition And (P Q:assert) : assert := fun s h => P s h /\ Q s h.

Module sep.

Definition con (P Q:assert) : assert := fun s h =>
  exists h1, exists h2, h1 # h2 /\ h = h1 +++ h2 /\ P s h1 /\ Q s h2.

Notation "P ** Q" := (con P Q) (at level 80).

Lemma con_inde_store: forall P Q s s' h,
  ((P ** Q) s h) ->
  (forall s s' h, P s h -> P s' h) ->
  (forall s s' h, Q s h -> Q s' h) ->
  ((P ** Q) s' h).

Definition imp (P Q:assert) : assert := fun s h =>
  forall h', h # h' /\ P s h' ->
    forall h'', h'' = h +++ h' -> Q s h''.

Notation "P -* Q" := (imp P Q) (at level 80).

Assertions lemma

Definition entails (P Q:assert) : Prop :=
  forall s h, P s h -> Q s h.

Notation "P ==> Q" := (entails P Q) (at level 90, no associativity).

Definition equiv (P Q:assert) : Prop :=
  forall s h, P s h <-> Q s h.

Notation "P <==> Q" := (equiv P Q) (at level 90, no associativity).

Lemma con_TT : forall P, P ==> (P ** TT).

Lemma con_com : forall P Q, P ** Q ==> Q ** P.

Axiom assert_ext: forall P Q, (P <==> Q) -> P = Q.

Lemma con_com_equiv : forall P Q, (P ** Q) = (Q ** P).

Lemma con_assoc: forall P Q R, (P ** Q) ** R ==> P ** (Q ** R).

Lemma con_assoc_equiv : forall P Q R, ((Q ** P) ** R) = (Q ** (P ** R)).

Lemma mp : forall P Q, Q ** (Q -* P) ==> P.

Lemma monotony : forall (P1 P2 Q1 Q2:assert),
  forall s h,
    ((forall h', P1 s h' -> P2 s h') /\ (forall h'', Q1 s h'' -> Q2 s h'')) ->
    ((P1 ** Q1) s h -> (P2 ** Q2) s h).

Lemma monotony': forall P1 P2 P3 P4,
 (P1 ==> P2) ->
 (P3 ==> P4) ->
 (P1 ** P3 ==> P2 ** P4).

Lemma monotony'': forall P1 P2 P3 P4,
 (P2 ==> P1) ->
 (P3 ==> P4) ->
 (P1 -* P3 ==> P2 -* P4).

Lemma adjunction : forall (P1 P2 P3:assert),
  forall s,
    (forall h, (P1 ** P2) s h -> P3 s h) ->
    (forall h, P1 s h -> (P2 -* P3) s h).

Lemma adjunction' : forall (P1 P2 P3:assert),
  forall s,
    (forall h, P1 s h -> (P2 -* P3) s h) ->
       (forall h, (P1 ** P2) s h -> P3 s h).

Lemma imp_reg : forall P Q, P ==> Q -* (P ** Q).

Definition emp : assert := fun s h => h = heap.emp.

Lemma con_emp: forall P, (P ** sep.emp) ==> P .

Lemma con_emp': forall P, P ==> (P ** sep.emp).

Definition mapsto e e' s h := exists p,
    val2loc (eval e s) = p /\ h = heap.singleton p (eval e' s).

Notation "e1 '|->' e2" := (mapsto e1 e2) (at level 79, no associativity).

Lemma mapsto_con_inversion : forall l e P s h,
  ((l |-> e) ** P) s h ->
  exists n,
    val2loc (eval l s) = n /\
    exists h',
      h = (heap.singleton n (eval e s)) +++ h' /\
      P s h'.

Definition cell_exists (e1: expr) : assert :=
 fun s h => exists y, (e1 |-> int_e y) s h.

Notation " e '|->_' " := (cell_exists e) (right associativity, at level 80).

Fixpoint mapstos (e:expr) (l:list expr) {struct l} : assert :=
  match l with
    nil => sep.emp
    | e1::tl => e |-> e1 ** mapstos (e +e int_e 1%Z) tl
  end.

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

End sep.

Notation "P ** Q" := (sep.con P Q) (at level 80) : sep_scope.
Notation "P -* Q" := (sep.imp P Q) (at level 80) : sep_scope.
Notation "P ==> Q" := (sep.entails P Q) (at level 90, no associativity) : sep_scope.
Notation "P <==> Q" := (sep.equiv P Q) (at level 90, no associativity) : sep_scope.
Notation "e1 '|->' e2" := (sep.mapsto e1 e2) (at level 79, no associativity) : sep_scope.
Notation "x '|-->' l" := (sep.mapstos x l) (at level 80) : sep_scope.

Lemma mapsto_equiv' : forall s s' h e1 e2 e3 e4,
  (e1|->e2) s' h ->
  eval e1 s' = eval e3 s ->
  eval e2 s' = eval e4 s ->
  (e3|->e4) s h.

Lemma mapsto_equiv : forall s h e1 e2 e3 e4,
    (e1|->e2) s h ->
    eval e1 s = eval e3 s ->
    eval e2 s = eval e4 s ->
    (e3|->e4) s h.

Definition inde (l:list var.v) (P:assert) :=
  forall s h,
    (forall x v, In x l -> (P s h <-> P (store.update x v s) h)).

Lemma inde_update_store : forall (P:assert) x z s h,
  inde (x::nil) P ->
  P s h ->
  P (store.update x z s) h.

Lemma inde_update_store' : forall (P:assert) x z s h,
  inde (x::nil) P ->
  P (store.update x z s) h ->
  P s h.

Lemma inde_TT : forall l, inde l TT.

Lemma inde_sep_con : forall l (P Q:assert),
  inde l P -> inde l Q ->
  inde l (P ** Q).

Lemma inde_sep_imp : forall l (P Q:assert),
  inde l P -> inde l Q ->
  inde l (P -* Q).

Lemma inde_mapsto : forall lst e e',
  inter (expr_var e) lst nil ->
  inter (expr_var e') lst nil ->
  inde lst (e |-> e').

Lemma inde_mapsto' : forall e x,
  inde x (fun s h => exists z, (int_e e |-> int_e z) s h).

Lemma inde_mapsto'' : forall e1 e2 x,
  inde x (int_e e1 |-> int_e e2).

Lemma inde_mapstos : forall lst l e,
  inter (fold_right (@app var.v) nil (map expr_var lst)) l nil ->
  inter (expr_var e) l nil ->
  inde l (e |--> lst).