Library topsy_hm_old

Load seplog_header.

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 status := 0.
Definition next := 1.

Definition hmStart : var.v := 0%nat.
Definition hmEnd : var.v := 1%nat.

Close Local Scope Z_scope.

Inductive Heap_List : list (loc * nat * expr) -> loc -> loc -> store.s -> heap.h -> Prop :=

  Heap_List_last: forall s next startl endl h, endl = 0 -> next = 0 -> startl > 0 ->
    (nat_e startl |--> Allocated::nat_e next::nil) s h ->
    Heap_List nil startl endl s h

  | Heap_List_trans: forall s startl endl h, endl > 0 -> startl = endl ->
    sep.emp s h ->
    Heap_List nil startl endl s h

  | Heap_List_suiv_Free: forall s h next startl endl h1 h2 hd_adr hd_size hd_expr tl ,
    h1 # h2 -> h = h1 +++ h2 ->
    hd_expr = Free ->
    next = startl + 2 + hd_size ->
    startl = hd_adr ->
    startl > 0 ->
    ((nat_e startl |--> hd_expr::nat_e next::nil) ** (Array (startl+2) hd_size)) s h1 ->
    Heap_List tl next endl s h2 ->
    Heap_List ((hd_adr, hd_size,hd_expr)::tl) startl endl s h

  | Heap_List_suiv_Allocated: forall s h next startl endl h1 h2 hd_adr hd_size hd_expr tl ,
    h1 # h2 -> h = h1 +++ h2 ->
    hd_expr = Allocated ->
    next = startl + 2 + hd_size ->
    startl = hd_adr ->
    startl > 0 ->
    (nat_e startl |--> hd_expr::nat_e next::nil) s h1 ->
    Heap_List tl next endl s h2 ->
    Heap_List ((hd_adr, hd_size,hd_expr)::tl) startl endl s h.

Definition hl_nil : list (loc * nat * expr) := nil.

Lemma Heap_List_start_pos: forall l startl endl s h,
   (Heap_List l startl endl) s h ->
        startl > 0.

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

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

Lemma Heap_List_equal_start: forall x sizex status startl endl l s h,
        Heap_List ((x,sizex,status)::l) startl endl s h ->
        startl = x.

Lemma Heap_List_suiv_ge: forall l x sizex status a b c startl endl s h,
        Heap_List ((a,b,c) ::l) startl endl s h ->
        In (x,sizex,status) l ->
        x > a.

Lemma Heap_List_block_adr_pos: forall l x sizex status startl endl s h,
        Heap_List l startl endl s h ->
        In (x,sizex,status) l ->
        x > 0.

Lemma Heap_List_app : forall a b startl s h,
     startl > 0 ->
     (Heap_List (a ++ b) startl 0 s h
     <->
     (exists j, ((Heap_List a startl j) ** (Heap_List b j 0)) s h)).







Ltac THeap_List_app_hyp id :=
              match goal with
                  | id : Heap_List (?l1 ++ ?l2) ?startl 0 ?s ?h |- _ =>
                             assert (THeap_List_app_H1: startl > 0); [apply (Heap_List_start_pos (l1 ++ l2) startl 0 s h id) |
                             generalize (Heap_List_app l1 l2 startl s h THeap_List_app_H1); clear THeap_List_app_H1; intros THeap_List_app_H1;
                             inversion_clear THeap_List_app_H1 as [THeap_List_app_H11 THeap_List_app_H12]; clear THeap_List_app_H12;
                             generalize (THeap_List_app_H11 id); clear THeap_List_app_H11; intro THeap_List_app_H1;
                             inversion_clear THeap_List_app_H1
                             ]
              end.

Ltac THeap_List_app_goal:=
              match goal with
                  | |- Heap_List (?l1 ++ ?l2) ?startl 0 ?s ?h =>
                             assert (THeap_List_app_H1: startl > 0); [(omega || tauto) |
                             generalize (Heap_List_app l1 l2 startl s h THeap_List_app_H1); clear THeap_List_app_H1; intros THeap_List_app_H1;
                             inversion_clear THeap_List_app_H1 as [THeap_List_app_H11 THeap_List_app_H12]; clear THeap_List_app_H11;
                             eapply THeap_List_app_H12; clear THeap_List_app_H12
                             ]
              end.

Lemma Heap_List_block_status: forall l x sizex status startl s h,
        Heap_List l startl 0 s h ->
        In (x,sizex,status) l ->
        status = Allocated \/ status = Free.

Lemma Heap_List_block_status': forall l x sizex status startl endl s h,
        Heap_List l startl endl s h ->
        In (x,sizex,status) l ->
        status = Allocated \/ status = Free.

Lemma Heap_List_app' : forall a b startl endl s h,
     startl > 0 ->
     endl > 0 ->
     (Heap_List (a ++ b) startl endl s h
     <->
     (exists j, ((Heap_List a startl j) ** (Heap_List b j endl)) s h)).







Lemma Heap_List_last_bloc: forall adr size status l s h startl,
        Heap_List (l ++ ((adr,size,status)::nil)) startl 0 s h ->
        (Heap_List (l ++ ((adr,size,status)::nil)) startl (adr + 2 + size) ** Heap_List nil (adr + 2 + size) 0) s h.



Lemma Heap_List_next_bloc: forall adr size status adr' size' status' l1 l2 s h startl,
        Heap_List (l1 ++ ((adr,size,status)::(adr',size',status')::nil) ++ l2) startl 0 s h ->
        adr' = (adr + 2 + size).

Lemma Compaction: forall l1 l2 x1 sizex1 sizex2 startl s h,
   startl > 0 ->
   (Heap_List (l1 ++ ((x1,sizex1,Free)::(x1 + 2 + sizex1,sizex2,Free)::nil) ++ l2) startl 0) s h ->
   (((nat_e x1 +e (int_e 1%Z))|->(nat_e (x1 + 2 + sizex1))) **
         (((nat_e x1 +e (int_e 1%Z))|-> nat_e (x1 + sizex1 + 4 + sizex2)) -*
                (Heap_List (l1 ++ ((x1, sizex1 + 2 + sizex2, Free)::nil) ++ l2 ) startl 0))) s h.


Lemma change_status_Free_to_Alloc: forall l1 l2 startl s h x sizex,
   startl > 0 ->
   (Heap_List (l1 ++ ((x, sizex, Free)::nil) ++ l2) startl 0) s h ->
   exists e,
   (((nat_e x) |-> e) **
   (((nat_e x) |-> Allocated) -* ((Heap_List (l1 ++ ((x, sizex, Allocated)::nil) ++ l2) startl 0) ** (Array (x+2) sizex)))) s h.

Lemma change_status_Alloc_to_Free: forall l1 l2 startl s h x sizex,
   startl > 0 ->
   ((Heap_List (l1 ++ ((x, sizex, Allocated)::nil) ++ l2) startl 0) ** (Array (x + 2) (sizex))) s h ->
   exists e,
   (((nat_e x) |-> e) **
   (((nat_e x) |-> Free) -* (Heap_List (l1 ++ ((x, sizex, Free)::nil) ++ l2) startl 0))) s h.

Lemma change_status_Free_to_Free: forall l1 l2 startl s h x sizex,
   startl > 0 ->
   (Heap_List (l1 ++ ((x, sizex, Free)::nil) ++ l2) startl 0 s h) ->
   exists e,
   (((nat_e x) |-> e) **
   (((nat_e x) |-> Free) -* (Heap_List (l1 ++ ((x, sizex, Free)::nil) ++ l2) startl 0))) s h.

Lemma Splitting: forall l1 l2 x size1 size2 s h startl,
   startl > 0 ->
   Heap_List (l1 ++ ((x, size1 + 2 + size2, Free)::nil) ++ l2) startl 0 s h ->
   exists e'' : expr,
     (((nat_e (x + 3 + size1)) |-> e'') **
      (((nat_e (x + 3 + size1)) |-> (nat_e (x + 4 + size1 + size2))) -*
       (fun (s1 : store.s) (h1 : heap.h) =>
        exists e''0 : expr,
          (((nat_e (x + 2 + size1)) |-> e''0) **
           (((nat_e (x + 2 + size1)) |-> Free) -*
            (fun (s2 : store.s) (h2 : heap.h) =>
             exists e''1 : expr,
               (((nat_e (x + 1)) |-> e''1) **
                (((nat_e (x + 1)) |-> (nat_e (x + 2 + size1))) -*
                 (fun (s : store.s) (h : heap.h) =>
                  Heap_List (l1 ++ ((x, size1,Free)::(x + 2 + size1, size2,Free)::nil) ++ l2) startl 0 s h))) s2 h2))) s1 h1))) s h.

Lemma Heap_List_header: forall x sizex statusx startl endl l s h,
        Heap_List ((x, sizex, statusx)::l) startl endl s h ->
        x = startl /\
        (Heap_List ((x, sizex, statusx)::nil) x (x + 2 + sizex) ** Heap_List l (x + 2 + sizex) endl) s h.

Ltac Heap_List_Trans :=
eapply Heap_List_trans; [auto | (Omega_exprb || auto) | red; Equal_heap].

Lemma Heap_List_next: forall l1 l2 bloc_adr bloc_size bloc_status startl s h e x P,
(Heap_List (l1 ++ ((bloc_adr, bloc_size, bloc_status) :: nil) ++ l2) startl 0 ** TT) s h ->
eval e s = eval (nat_e (bloc_adr + 1)) s ->
((update_store2 x (nat_e (bloc_adr + 2 +bloc_size)) P) s h) ->
exists e0, ((e |-> e0) ** ((e |-> e0) -* update_store2 x e0 P)) s h.






Lemma Heap_List_next': forall l bloc_adr bloc_size bloc_status startl s h e x P,
In (bloc_adr, bloc_size, bloc_status) l ->
(Heap_List l startl 0 ** TT) s h ->
eval e s = eval (nat_e (bloc_adr + 1)) s ->
((update_store2 x (nat_e (bloc_adr + 2 +bloc_size)) P) s h) ->
exists e0, ((e |-> e0) ** ((e |-> e0) -* update_store2 x e0 P)) s h.

Lemma Heap_List_next_last: forall bloc_adr s h e x P,
bloc_adr > 0 ->
(Heap_List nil bloc_adr 0 ** TT) s h ->
eval e s = eval (nat_e (bloc_adr + 1)) s ->
((update_store2 x (nat_e 0) P) s h) ->
exists e0, ((e |-> e0) ** ((e |-> e0) -* update_store2 x e0 P)) s h.

Lemma Heap_List_status: forall l adr size status e x P startl s h ,
In (adr,size,status) l ->
(Heap_List l startl 0 ** TT) s h ->
eval e s = eval (nat_e adr) s ->
(update_store2 x status P) s h ->
exists e0, ((e |-> e0) ** ((e |-> e0) -* update_store2 x e0 P)) s h.


Lemma Heap_List_status_last: forall bloc_adr s h e x P,
bloc_adr > 0 ->
(Heap_List nil bloc_adr 0 ** TT) s h ->
eval e s = eval (nat_e bloc_adr) s ->
(update_store2 x Allocated P) s h ->
exists e0, ((e |-> e0) ** ((e |-> e0) -* update_store2 x e0 P)) s h.

Lemma Heap_List_bloc_unique: forall l startl s h adr size size' status status',
    Heap_List l startl 0 s h ->
    In (adr, size, status) l ->
    In (adr, size', status') l ->
    status = status' /\ size = size'.

Lemma Heap_List_list_split_cont: forall l adr size stat size' stat' startl s h,
    Heap_List l startl 0 s h ->
    In (adr, size, stat) l ->
    In (adr + 2 + size, size', stat') l ->
    (exists l1, exists l2, l = l1 ++ (adr, size, stat)::(adr + 2 + size, size', stat')::l2 ).






Ltac Resolve_topsy:=
      (simpl; red; intuition; repeat (Store_update || Heap_List_equiv)).

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'
          | |- {{?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.