Library hl_abstraction
Load seplog_header.
Require Import topsy_hm.
Lemma eval_store_update: forall e x v s,
eval e (store.update x v s) = eval (expr_rewrite e (var_e x) (int_e v)) s.
Lemma eval_b_store_update: forall b x v s,
eval_b b (store.update x v s) = eval_b (expr_b_rewrite b (var_e x) (int_e v)) s.
Inductive abstract_hl: list (loc * expr) -> loc -> nat -> store.s -> heap.h -> Prop :=
abstract_hl_last: forall l start size s h,
size = 0 ->
l = nil ->
sep.emp s h ->
abstract_hl l start size s h
| abstract_hl_free: forall l l' adr start size s h h1 h2,
start = adr ->
size > 0 ->
l = (adr, Free)::l' ->
h1 # h2 -> h = (h1 +++ h2) ->
((nat_e start) |-> Free) s h1 ->
abstract_hl l' (start + 1) (size - 1) s h2 ->
abstract_hl l start size s h
| abstract_hl_alloc: forall l l' adr start size s h h1 h2,
start = adr ->
size > 0 ->
l = (adr, Allocated)::l' ->
h1 # h2 -> h = (h1 +++ h2) ->
((nat_e start) |-> Allocated) s h1 ->
abstract_hl l' (start + 1) (size - 1) s h2 ->
abstract_hl l start size s h.
Lemma abstract_hl_inde_store: forall l start size s s' h,
abstract_hl l start size s h -> abstract_hl l start size s' h.
Lemma abstract_hl_start: forall l a b start size s h,
abstract_hl ((a,b)::l) start size s h -> a = start.
Lemma abstract_hl_size: forall l start size s h,
abstract_hl l start size s h -> size = length l.
Lemma abstract_hl_free_alloc: forall l a b start size s h,
abstract_hl l start size s h -> In (a,b) l ->
b = Free \/ b =Allocated.
Lemma abstract_hl_app: forall l1 l2 start size s h,
abstract_hl (l1 ++ l2) start size s h ->
(abstract_hl l1 start (size - length l2) ** abstract_hl l2 (start + length l1) (size - length l1)) s h.
Lemma abstract_hl_app': forall l1 l2 start size s h,
(abstract_hl l1 start (size - length l2) ** abstract_hl l2 (start + length l1) (size - length l1)) s h ->
abstract_hl (l1 ++ l2) start size s h.
Lemma abstract_hl_cons: forall l a b start size s h,
abstract_hl ((a,b)::l) start size s h ->
((nat_e a |-> b) ** abstract_hl l (start + 1) (size - 1)) s h.
Lemma abstract_hl_cons': forall l a b start size s h,
b = Free \/ b =Allocated ->
a = start ->
size > 0 ->
((nat_e a |-> b) ** abstract_hl l (start + 1) (size - 1)) s h ->
abstract_hl ((a,b)::l) start size s h.
Axiom coucou: False.
Ltac Coucou := generalize coucou; contradiction.
Lemma abstract_hl_cell_gt: forall l a b x y start size s h,
abstract_hl ((x, y) :: l) start size s h ->
In (a, b) l ->
a > x.
Lemma abstract_hl_notin_list: forall (l: list (nat * expr)) a b x y start size s h,
abstract_hl l start size s h ->
In (a,b) l ->
In (x,y) l ->
b <> y ->
a <> x.
Definition abstract_hmAlloc (start: loc) (size: nat) (result i j: var.v) : cmd :=
i <- (nat_e 0);
j <-* ((nat_e start) +e (var_e i));
while (((nat_e size) >> (var_e i)) &&& ((var_e j) =/= (Free))) (
i <- ((var_e i) +e (nat_e 1));
ifte ((nat_e size) >> (var_e i)) thendo
j <-* ((nat_e start) +e (var_e i))
elsedo
skip
);
(ifte ((nat_e size) >> (var_e i)) thendo (((nat_e start) +e (var_e i)) *<- Allocated; result <- ((nat_e start) +e (var_e i))) elsedo (result <- (int_e (-1)%Z))).
Definition abstract_hmAlloc_specif1 := forall start size result i j x,
(var.set (result::i::j::nil)) ->
size > 0 ->
{{fun s h => exists l,
abstract_hl l start size s h /\
In (x,Allocated) l}}
abstract_hmAlloc start size result i j
{{ fun s h =>
(exists l, exists y,
eval (var_e result) s = Z_of_nat (y) /\
abstract_hl l start size s h /\
In (x,Allocated) l /\
In (y,Allocated) l /\
x <> y)
\/
(exists l,
eval (var_e result) s = (-1)%Z /\
abstract_hl l start size s h /\
In (x,Allocated) l) }}.
Lemma abstract_hmAlloc_verif1: abstract_hmAlloc_specif1.