Library topsy_hm

Load seplog_header.

Require Import Bool.

Ltac Step R :=
  match goal with
    | id: {{?P'}} ?c {{?Q'}} |- {{?P}} ?c {{?Q}} => eapply apply_triple; [apply id | idtac | idtac]
    | id: {{?P'}} ?c {{?Q'}} |- {{?P}} ?c; ?c' {{?Q}} => eapply apply_triple'; [apply id | idtac | idtac]
    | |- {{?P}} ?x <- ?e {{?Q}} => eapply semax_assign'; red
    | |- {{?P}} ?x <- ?e ; ?c {{?Q}} => eapply semax_assign'' with R; [red; intros | idtac]
    | |- {{?P}} ?x <-* ?e {{?Q}} => eapply semax_lookup_backwards'
    | |- {{?P}} ?x <-* ?e ; ?c {{?Q}} => eapply semax_lookup_backwards'' with R; [red; intros | idtac]
    | |- {{?P}} ?e1 *<- ?e2 {{?Q}} => eapply semax_mutation_backwards'
    | |- {{?P}} ?e1 *<- ?e2 ; ?c {{?Q}} => eapply semax_mutation_backwards'' with R; [red; intros | idtac]
    | |- {{?P}} while ?b ?c1 {{?Q}} => eapply semax_while' with R
    | |- {{?P}} while ?b ?c1 ; ?c2 {{?Q}} => eapply semax_while'' with R; [red; intros | idtac | idtac]
    | |- {{?P}} skip {{?Q}} => eapply semax_skip'
    | |- {{?P}} ifte ?b thendo ?c1 elsedo ?c2 {{?Q}} => eapply semax_ifte
    | |- {{?P}} (ifte ?b thendo ?c1 elsedo ?c2); ?c' {{?Q}} => apply semax_seq with R; [eapply semax_ifte; [idtac| idtac] | idtac]
    | |- {{?P}} ?c1; ?c2 {{?Q}} => apply semax_seq with R; [idtac| idtac]
  end.

Notation " s '|b=' b " := (eval_b b s = true) (right associativity, at level 80).

Notation " s '|b/=' b " := (eval_b b s = false) (right associativity, at level 80).

Ltac Decompose_hyp :=
  match goal with
    | id: ?P1 /\ ?P2 |- _ => (decompose [and] id; clear id); Decompose_hyp
    | id: (?P1 ** ?P2) ?s ?h |- _ =>
       (let x:=fresh in (
           inversion_clear id as [x X];
           let y:= fresh in (
              inversion_clear X as [y Y];
              decompose [and] Y; clear Y
           )
        )); Decompose_hyp
    | id: ex ?P |- _ => (inversion_clear id); Decompose_hyp
    | _ => idtac
  end.

Ltac Rewrite_heap h :=
  (match goal with
    | id: h = (?h1 +++ ?h2) |- _ => rewrite id; (Rewrite_heap h1); (Rewrite_heap h2)
    | _ => idtac
   end) .

Ltac Resolve_assert1 :=
  match goal with
    | |- (?P1 ** ?P2) ?s ?h => Rewrite_heap h; repeat rewrite sep.con_assoc_equiv
    | |- (?P1 -* ?P2) ?s ?h => red; intro; intro X; inversion_clear X; intros
    | _ => idtac
  end.

Lemma expr_store_update_rewrite: forall e x p s,
  eval (expr_rewrite e (var_e x) (int_e p)) s = eval e (store.update x p s).


Lemma expr_b_store_update_rewrite: forall b x p s,
  eval_b (expr_b_rewrite b (var_e x) (int_e p)) s = eval_b b (store.update x p s).

Lemma mapsto_store_update_rewrite: forall e1 e2 x p s h,
  ((expr_rewrite e1 (var_e x) (int_e p)) |-> (expr_rewrite e2 (var_e x) (int_e p))) s h ->
  (e1 |-> e2) (store.update x p s) h.

Definition bool_eq (a b: bool) :=
  match (a,b) with
     | (true, true) => true
     | (false, false) => true
     | _ => false
  end.

Lemma bool_eq_classic: forall a b,
  bool_eq a b = true \/ bool_eq a b = false.

Lemma bool_eq_true: forall a b,
  bool_eq a b = true -> a = b.

Lemma bool_eq_false: forall a b,
  bool_eq a b = false -> a <> b.

Lemma bool_eq_dif: forall a b,
  a <> b -> bool_eq a b = false.

Lemma bool_eq_refl: forall a,
  bool_eq a a = true.

Fixpoint get_endl (l:list (nat * bool)) (startl: nat) {struct l} : nat :=
   match l with
     nil => startl
     | (b,c)::tl => get_endl tl (startl + 2 + b)
   end.

Lemma get_endl_app: forall a b startl,
  get_endl (a++b) startl = get_endl b (get_endl a startl).

Lemma get_endl_gt: forall l startl,
  startl <= get_endl l startl.

Fixpoint In_hl (l:list (nat*bool)) (block:nat*nat*bool) (startp: nat) {struct l}: Prop :=
  match l with
    nil => False
    | (size , stat)::tl => match block with
                             (adr,sz,st) =>
                                if (andb (andb (beq_nat startp adr) (bool_eq st stat)) (beq_nat size sz)) then
                                    True
                                else
                                    In_hl tl block (get_endl ((size,stat)::nil) startp)
                           end
  end.

Lemma In_hl_app_or: forall l1 l2 a b c startp,
  In_hl (l1 ++ l2) (a,b,c) startp ->
  (In_hl l1 (a,b,c) startp \/ In_hl l2 (a,b,c) (get_endl l1 startp)).

Lemma In_hl_or_app: forall l1 l2 a b c startp,
  (In_hl l1 (a,b,c) startp \/ In_hl l2 (a,b,c) (get_endl l1 startp)) ->
  In_hl (l1 ++ l2) (a,b,c) startp.

Open Local Scope Z_scope.

Definition HM_FREEFAILED := int_e 0.
Definition HM_FREEOK := int_e 1.

Definition Allocated := int_e 0.
Definition Free := int_e 1.

Definition alloc := false.
Definition free := true.

Definition status := 0.
Definition next := 1.

Close Local Scope Z_scope.

Definition hmStart := 0.
Definition hmEnd := 1.
Definition entry := 3.
Definition cptr := 4.
Definition nptr := 5.
Definition result := 6.
Definition fnd := 7.
Definition stts := 8.
Definition sz := 9.

Hint Unfold fnd.
Hint Unfold stts.
Hint Unfold sz.
Hint Unfold entry.
Hint Unfold cptr.
Hint Unfold nptr.
Hint Unfold result.
Hint Unfold HM_FREEFAILED.
Hint Unfold HM_FREEOK.
Hint Unfold Allocated.
Hint Unfold Free.
Hint Unfold status.
Hint Unfold next.
Hint Unfold hmStart.
Hint Unfold hmEnd.

Inductive hl : loc -> list (nat * bool) -> assert :=

  | hl_last: forall s p h,
    sep.emp s h ->
    hl p nil s h

  | hl_Free: forall s h next p h1 h2 size flag tl,
    h1 # h2 -> h = h1 +++ h2 ->
    flag = true ->
    next = p + 2 + size ->
    (nat_e p |--> Free::nat_e next::nil ** Array (p+2) size) s h1 ->
    hl next tl s h2 ->
    hl p ((size,flag)::tl) s h

  | hl_Allocated: forall s h next p h1 h2 size flag tl,
    h1 # h2 -> h = h1 +++ h2 ->
    flag = false ->
    next = p + 2 + size ->
    (nat_e p |--> Allocated::nat_e next::nil) s h1 ->
    hl next tl s h2 ->
    hl p ((size,flag)::tl) s h.

Definition hlstat_bool2expr (b: bool) : expr :=
   match b with
      true => Free
      | false => Allocated
   end.

Definition Heap_List (l:list (nat * bool)) (p:nat) : assert :=
  (hl p l) ** (nat_e (get_endl l p) |--> Allocated::null::nil).

Lemma Heap_List_inde_store: forall l startl s h,
  Heap_List startl l s h -> forall s', Heap_List startl l s' h.



Ltac Heap_List_equiv :=
  match goal with
    | id: Heap_List ?l ?start1 ?s1 ?h |- Heap_List ?l ?start2 ?s2 ?h =>
      assert (Heap_List_equivA1: start2 = start1); [omega | ((rewrite Heap_List_equivA1) || idtac); eapply (Heap_List_inde_store); apply id ]
  end.

Lemma hl_app : forall a b startl s h,
     (hl startl (a ++ b) s h)
     <->
     ((hl startl a ** hl (get_endl a startl) b) s h).







Lemma hl_compaction: forall l1 l2 size size' p s h,
  Heap_List (l1++(size,free)::(size',free)::nil++l2) p s h ->
  exists y, (nat_e (get_endl l1 p + 1) |-> y **
    (nat_e (get_endl l1 p + 1) |-> nat_e (get_endl l1 p + size + size' + 4) -*
      Heap_List (l1 ++ (size+size'+2,free)::nil ++ l2) p)) s h.

Lemma compaction_example: forall p,
 {{ Heap_List ((8,free)::(10,free)::nil) p }}
 nat_e p +e int_e 1%Z *<- nat_e (p + 22)
 {{ Heap_List ((20,free)::nil) p }}.

Lemma hl_splitting: forall l1 l2 a b startp s h,
   Heap_List (l1 ++ (a + 2 + b,true)::l2) startp s h ->
   (exists y, ((nat_e (get_endl l1 startp + a + 3) |-> y) **
     ((nat_e (get_endl l1 startp + a + 3) |-> (nat_e (get_endl l1 startp + a + b + 4))) -*
        (fun s h => exists y, ((nat_e (get_endl l1 startp + a + 2) |-> y) **
          ((nat_e (get_endl l1 startp + a + 2) |-> Free) -*
             (fun s h => exists y, ((nat_e (get_endl l1 startp + 1) |-> y) **
                 ((nat_e (get_endl l1 startp + 1) |-> (nat_e (get_endl l1 startp + a + 2))) -*
                    Heap_List (l1 ++ (a,true)::(b,true)::l2) startp
                 )) s h
             )
          )) s h
        )
     )) s h
   ) .


Lemma hl_free2alloc: forall l1 l2 a startp s h,
   Heap_List (l1 ++ (a,true)::l2) startp s h ->
   exists y, ((nat_e (get_endl l1 startp) |-> y) **
   ((nat_e (get_endl l1 startp) |-> Allocated) -*
       (Heap_List (l1 ++ (a,false)::l2) startp ** Array (get_endl l1 startp + 2) a))) s h.

Lemma hl_alloc2free: forall l1 l2 a startp s h,
   (Heap_List (l1 ++ (a,false)::l2) startp ** Array (get_endl l1 startp + 2) a) s h ->
   exists y, ((nat_e (get_endl l1 startp) |-> y) **
   ( (nat_e (get_endl l1 startp) |-> Free) -*
       (Heap_List (l1 ++ (a,true)::l2) startp))) s h.

Lemma hl_free2free: forall l1 l2 a startp s h,
   Heap_List (l1 ++ (a,true)::l2) startp s h ->
   exists y, ((nat_e (get_endl l1 startp) |-> y) **
   ( (nat_e (get_endl l1 startp) |-> Free) -*
       (Heap_List (l1 ++ (a,true)::l2) startp))) s h.

Lemma hl_getnext: forall l1 l2 a b startp s h,
   Heap_List (l1 ++ (a,b)::l2) startp s h ->
   ((nat_e (get_endl l1 startp + 1) |-> (nat_e (get_endl l1 startp + 2 + a))) **
   ((nat_e (get_endl l1 startp + 1) |-> (nat_e (get_endl l1 startp + 2 + a))) -*
       (Heap_List (l1 ++ (a,b)::l2) startp)
     )) s h.

Lemma hl_getstatus: forall l1 l2 a b startp s h,
   Heap_List (l1 ++ (a,b)::l2) startp s h ->
   ((nat_e (get_endl l1 startp) |-> (hlstat_bool2expr b)) **
   ((nat_e (get_endl l1 startp) |-> (hlstat_bool2expr b)) -*
       (Heap_List (l1 ++ (a,b)::l2) startp)
     )) s h.

Lemma hl_getnext': forall l startp s h,
   Heap_List l startp s h ->
   ((nat_e (get_endl l startp + 1) |-> (nat_e 0)) **
   ((nat_e (get_endl l startp + 1) |-> (nat_e 0)) -*
       (Heap_List l startp)
     )) s h.

Lemma hl_getstat': forall l startp s h,
   Heap_List l startp s h ->
   ((nat_e (get_endl l startp) |-> Allocated) **
   ((nat_e (get_endl l startp) |-> Allocated) -*
       (Heap_List l startp)
     )) s h.

Lemma In_hl_lt: forall l a b c startp,
  In_hl l (a,b,c) startp ->
  a < get_endl l startp.

Lemma In_hl_destruct: forall l a b c adr,
   In_hl l (a,b,c) adr ->
   exists l1, exists l2, l = l1 ++ (b,c)::l2 /\ get_endl l1 adr = a.

Lemma In_hl_ge_start: forall l x y z adr,
 In_hl l (x,y,z) adr -> x >= adr.

Lemma In_hl_dif: forall l x y a b adr,
 In_hl l (x,y,alloc) adr -> In_hl l (a,b,free) adr -> a <> x.

Lemma In_hl_match: forall l1 l2 x y size stat startp,
  In_hl (l1 ++ (x,y)::l2) (get_endl l1 startp, size, stat) startp ->
  (size = x /\ stat = y).

Lemma In_hl_next: forall l x y z b c adr,
  In_hl l (x,y,z) adr ->
  In_hl l (x + 2 + y,b,c) adr ->
  exists l1,
    exists l2,
      l = l1 ++ (y,z)::(b,c)::l2 /\
      x = get_endl l1 adr.

Lemma In_hl_last: forall l a b c adr,
  In_hl l (a,b,c) adr ->
  get_endl l adr = a + 2 + b ->
  exists l', l = l'++ (b,c)::nil.

Lemma In_hl_first: forall l adr sz st,
  In_hl l (adr,sz,st) adr ->
  exists l', l = (sz,st)::l'.

Ltac Print x := assert (x=x); [auto | idtac].

Ltac Cutrewrite_hyp H A :=
  match A with
    | ?x = ?x => idtac
    | _ =>
      let x:= fresh in
        (
          assert (x: A); [idtac | rewrite x in H; clear x]
        )
  end.

Ltac Omega_exprb_contradiction_assert P :=
  assert (P); Omega_exprb.

Ltac Omega_exprb_contradiction :=
  match goal with
    | id: ?s |b= ?P |- _ => assert (False); [Omega_exprb_contradiction_assert (s |b/= P) | contradiction]
    | id: ?s |b/= ?P |- _ => assert (False); [Omega_exprb_contradiction_assert (s |b= P) | contradiction]
  end.

Ltac In_hl_destruct_hyp H :=
  match goal with
    | H: In_hl ?l (?adr, ?size, ?status) ?start |- _ =>
      generalize (In_hl_destruct l adr size status start H); intros; Decompose_hyp; subst l
  end.

  Ltac HLList_head l :=
    match l with
      | ?hd::?tl => constr:(hd::nil)
      | ?l1 ++ ?l2 =>
        let x:=(HLList_head l1) in (
          match x with
            | nil => l1
            | _ => x
          end
        )
      | _ => l
    end.

  Ltac HLList_tail l :=
    match l with
      | ?hd::?tl => tl
      | ?l1 ++ ?l2 =>
        let x:=(HLList_tail l1) in (
          match x with
            | nil => l2
            | _ => constr:(x ++ l2)
          end
        )
      | _ => constr:(@nil (nat*bool))
    end.

  Ltac Normalize_HLList l :=
    match l with
      | nil => constr:(@nil (nat*bool))
      | _ =>
        let hd := (HLList_head l) in (
          let tl := (HLList_tail l) in (
            let tl' := (Normalize_HLList tl) in (
              match tl' with
                | nil => hd
                | _ => constr:(hd ++ tl')
              end
            )
          )
        )
    end.

  Ltac Format_HLList_head l elt :=
    match l with
      | ((elt,?st)::nil) => constr:(@nil (nat*bool))
      | ((elt,?st)::nil) ++ ?l2 => constr:(@nil (nat*bool))
      | ?l1 ++ ?l2 =>
        let hd:= (Format_HLList_head l2 elt) in (
          match hd with
            | nil => l1
            | _ => constr:(l1 ++ hd)
          end
          )
    end.

  Ltac Format_HLList_tail l elt :=
    match l with
      | ((elt,?st)::nil) => l
      | ((elt,?st)::nil) ++ ?l2 => constr:((elt,st)::l2)
      | ?l1 ++ ?l2 =>
        let tl := (Format_HLList_tail l1 elt) in (
          match tl with
            | nil => (Format_HLList_tail l2 elt)
            | _ => constr:(tl ++ l2)
          end
        )
      | _ => constr:(@nil (nat*bool))
    end.

  Ltac Format_HLList l elt :=
    let y:= (Normalize_HLList l) in (
      let hd:= (Format_HLList_head y elt) in (
        let tl:= (Format_HLList_tail y elt) in (
          constr:(hd ++ tl)
        )
      )
    ).


  Ltac List_eq :=
    repeat rewrite <- ass_app; repeat rewrite <- app_nil_end; repeat rewrite <- app_comm_cons; auto.

  Ltac Hl_getstatus H elt :=
    match goal with
      | H: Heap_List ?l ?start ?s ?h |- _ =>
        let l' := (Format_HLList l elt) in (
          match l with
            | l' =>
              let h1:= fresh in(
                let h2:= fresh in (
                  let x:= fresh in (
                    generalize (hl_getstatus _ _ _ _ _ _ _ H); intro x; Decompose_sepcon x h1 h2;
                      Compose_sepcon h1 h2; [idtac | red; auto]
                  )
                )
              )
            | _ =>
                  Cutrewrite_hyp H (l = l'); [List_eq |
                    (let h1:= fresh in(
                      let h2:= fresh in (
                        let x:= fresh in (
                          generalize (hl_getstatus _ _ _ _ _ _ _ H); intro x; Decompose_sepcon x h1 h2;
                            Compose_sepcon h1 h2; [idtac | red; auto]
                        )
                      )
                    )
                    )
                  ]
          end
          )
          
    end.

  Ltac Hl_getstat' H :=
    let h1:= fresh in(
      let h2:= fresh in (
        let x:= fresh in (
          generalize (hl_getstat' _ _ _ _ H); intro x; Decompose_sepcon x h1 h2;
            Compose_sepcon h1 h2; [idtac | red; auto]
        )
      )
    ).

  Ltac Hl_getnext H elt :=
    match goal with
      | H: Heap_List ?l ?start ?s ?h |- _ =>
        let l' := (Format_HLList l elt) in (
          match l with
            | l' =>
              let h1:= fresh in(
                let h2:= fresh in (
                  let x:= fresh in (
                    generalize (hl_getnext _ _ _ _ _ _ _ H); intro x; Decompose_sepcon x h1 h2;
                      Compose_sepcon h1 h2; [idtac | red; auto]
                  )
                )
              )
            | _ =>
                  Cutrewrite_hyp H (l = l'); [List_eq |
                    (let h1:= fresh in(
                      let h2:= fresh in (
                        let x:= fresh in (
                          generalize (hl_getnext _ _ _ _ _ _ _ H); intro x; Decompose_sepcon x h1 h2;
                            Compose_sepcon h1 h2; [idtac | red; auto]
                        )
                      )
                    )
                    )
                  ]
          end
          )
          
    end.

  Ltac Hl_getnext' H :=
    let h1:= fresh in(
      let h2:= fresh in (
        let x:= fresh in (
          generalize (hl_getnext' _ _ _ _ H); intro x; Decompose_sepcon x h1 h2;
            Compose_sepcon h1 h2; [idtac | red; auto]
        )
      )
    ).