Animating the CompCert C semantics
Require Import FunInd.
Require Import Axioms Classical.
Require Import String Coqlib Decidableplus.
Require Import Errors Maps Integers Floats.
Require Import AST Values Memory Events Globalenvs Builtins Determinism.
Require Import Ctypes Cop Csyntax Csem.
Require Cstrategy.
Local Open Scope string_scope.
Local Open Scope list_scope.
Error monad with options or lists
Declare Scope option_monad_scope.
Notation "'
do'
X <-
A ;
B" := (
match A with Some X =>
B |
None =>
None end)
(
at level 200,
X ident,
A at level 100,
B at level 200)
:
option_monad_scope.
Notation "'
do'
X ,
Y <-
A ;
B" := (
match A with Some (
X,
Y) =>
B |
None =>
None end)
(
at level 200,
X ident,
Y ident,
A at level 100,
B at level 200)
:
option_monad_scope.
Notation "'
do'
X ,
Y ,
Z <-
A ;
B" := (
match A with Some (
X,
Y,
Z) =>
B |
None =>
None end)
(
at level 200,
X ident,
Y ident,
Z ident,
A at level 100,
B at level 200)
:
option_monad_scope.
Notation "'
do'
X ,
Y ,
Z ,
W <-
A ;
B" := (
match A with Some (
X,
Y,
Z,
W) =>
B |
None =>
None end)
(
at level 200,
X ident,
Y ident,
Z ident,
W ident,
A at level 100,
B at level 200)
:
option_monad_scope.
Notation " '
check'
A ;
B" := (
if A then B else None)
(
at level 200,
A at level 100,
B at level 200)
:
option_monad_scope.
Declare Scope list_monad_scope.
Notation "'
do'
X <-
A ;
B" := (
match A with Some X =>
B |
None =>
nil end)
(
at level 200,
X ident,
A at level 100,
B at level 200)
:
list_monad_scope.
Notation " '
check'
A ;
B" := (
if A then B else nil)
(
at level 200,
A at level 100,
B at level 200)
:
list_monad_scope.
Definition is_val (
a:
expr) :
option (
val *
type) :=
match a with
|
Eval v ty =>
Some(
v,
ty)
|
_ =>
None
end.
Lemma is_val_inv:
forall a v ty,
is_val a =
Some(
v,
ty) ->
a =
Eval v ty.
Proof.
intros until ty. destruct a; simpl; congruence.
Qed.
Definition is_loc (
a:
expr) :
option (
block *
ptrofs *
bitfield *
type) :=
match a with
|
Eloc b ofs bf ty =>
Some(
b,
ofs,
bf,
ty)
|
_ =>
None
end.
Lemma is_loc_inv:
forall a b ofs bf ty,
is_loc a =
Some(
b,
ofs,
bf,
ty) ->
a =
Eloc b ofs bf ty.
Proof.
intros until ty. destruct a; simpl; congruence.
Qed.
Local Open Scope option_monad_scope.
Fixpoint is_val_list (
al:
exprlist) :
option (
list (
val *
type)) :=
match al with
|
Enil =>
Some nil
|
Econs a1 al =>
do vt1 <-
is_val a1;
do vtl <-
is_val_list al;
Some(
vt1::
vtl)
end.
Definition is_skip (
s:
statement) : {
s =
Sskip} + {
s <>
Sskip}.
Proof.
destruct s; (left; congruence) || (right; congruence).
Defined.
Events, volatile memory accesses, and external functions.
Section EXEC.
Variable ge:
genv.
Definition eventval_of_val (
v:
val) (
t:
typ) :
option eventval :=
match v with
|
Vint i =>
check (
typ_eq t AST.Tint);
Some (
EVint i)
|
Vfloat f =>
check (
typ_eq t AST.Tfloat);
Some (
EVfloat f)
|
Vsingle f =>
check (
typ_eq t AST.Tsingle);
Some (
EVsingle f)
|
Vlong n =>
check (
typ_eq t AST.Tlong);
Some (
EVlong n)
|
Vptr b ofs =>
do id <-
Genv.invert_symbol ge b;
check (
Genv.public_symbol ge id);
check (
typ_eq t AST.Tptr);
Some (
EVptr_global id ofs)
|
_ =>
None
end.
Fixpoint list_eventval_of_val (
vl:
list val) (
tl:
list typ) :
option (
list eventval) :=
match vl,
tl with
|
nil,
nil =>
Some nil
|
v1::
vl,
t1::
tl =>
do ev1 <-
eventval_of_val v1 t1;
do evl <-
list_eventval_of_val vl tl;
Some (
ev1 ::
evl)
|
_,
_ =>
None
end.
Definition val_of_eventval (
ev:
eventval) (
t:
typ) :
option val :=
match ev with
|
EVint i =>
check (
typ_eq t AST.Tint);
Some (
Vint i)
|
EVfloat f =>
check (
typ_eq t AST.Tfloat);
Some (
Vfloat f)
|
EVsingle f =>
check (
typ_eq t AST.Tsingle);
Some (
Vsingle f)
|
EVlong n =>
check (
typ_eq t AST.Tlong);
Some (
Vlong n)
|
EVptr_global id ofs =>
check (
Genv.public_symbol ge id);
check (
typ_eq t AST.Tptr);
do b <-
Genv.find_symbol ge id;
Some (
Vptr b ofs)
end.
Ltac mydestr :=
match goal with
| [ |-
None =
Some _ ->
_ ] =>
let X :=
fresh "
X"
in intro X;
discriminate
| [ |-
Some _ =
Some _ ->
_ ] =>
let X :=
fresh "
X"
in intro X;
inv X
| [ |-
match ?
x with Some _ =>
_ |
None =>
_ end =
Some _ ->
_ ] =>
destruct x eqn:?;
mydestr
| [ |-
match ?
x with true =>
_ |
false =>
_ end =
Some _ ->
_ ] =>
destruct x eqn:?;
mydestr
| [ |-
match ?
x with left _ =>
_ |
right _ =>
_ end =
Some _ ->
_ ] =>
destruct x;
mydestr
|
_ =>
idtac
end.
Lemma eventval_of_val_sound:
forall v t ev,
eventval_of_val v t =
Some ev ->
eventval_match ge ev t v.
Proof.
Lemma eventval_of_val_complete:
forall ev t v,
eventval_match ge ev t v ->
eventval_of_val v t =
Some ev.
Proof.
Lemma list_eventval_of_val_sound:
forall vl tl evl,
list_eventval_of_val vl tl =
Some evl ->
eventval_list_match ge evl tl vl.
Proof with
Lemma list_eventval_of_val_complete:
forall evl tl vl,
eventval_list_match ge evl tl vl ->
list_eventval_of_val vl tl =
Some evl.
Proof.
Lemma val_of_eventval_sound:
forall ev t v,
val_of_eventval ev t =
Some v ->
eventval_match ge ev t v.
Proof.
intros until v. destruct ev; simpl; mydestr; constructor; auto.
Qed.
Lemma val_of_eventval_complete:
forall ev t v,
eventval_match ge ev t v ->
val_of_eventval ev t =
Some v.
Proof.
induction 1;
simpl.
-
auto.
-
auto.
-
auto.
-
auto.
-
simpl in *.
rewrite H,
H0.
rewrite dec_eq_true.
auto.
Qed.
Volatile memory accesses.
Definition do_volatile_load (
w:
world) (
chunk:
memory_chunk) (
m:
mem) (
b:
block) (
ofs:
ptrofs)
:
option (
world *
trace *
val) :=
if Genv.block_is_volatile ge b then
do id <-
Genv.invert_symbol ge b;
match nextworld_vload w chunk id ofs with
|
None =>
None
|
Some(
res,
w') =>
do vres <-
val_of_eventval res (
type_of_chunk chunk);
Some(
w',
Event_vload chunk id ofs res ::
nil,
Val.load_result chunk vres)
end
else
do v <-
Mem.load chunk m b (
Ptrofs.unsigned ofs);
Some(
w,
E0,
v).
Definition do_volatile_store (
w:
world) (
chunk:
memory_chunk) (
m:
mem) (
b:
block) (
ofs:
ptrofs) (
v:
val)
:
option (
world *
trace *
mem *
val) :=
if Genv.block_is_volatile ge b then
do id <-
Genv.invert_symbol ge b;
do ev <-
eventval_of_val (
Val.load_result chunk v) (
type_of_chunk chunk);
do w' <-
nextworld_vstore w chunk id ofs ev;
Some(
w',
Event_vstore chunk id ofs ev ::
nil,
m,
v)
else
do m' <-
Mem.store chunk m b (
Ptrofs.unsigned ofs)
v;
Some(
w,
E0,
m',
v).
Lemma do_volatile_load_sound:
forall w chunk m b ofs w'
t v,
do_volatile_load w chunk m b ofs =
Some(
w',
t,
v) ->
volatile_load ge chunk m b ofs t v /\
possible_trace w t w'.
Proof.
Lemma do_volatile_load_complete:
forall w chunk m b ofs w'
t v,
volatile_load ge chunk m b ofs t v ->
possible_trace w t w' ->
do_volatile_load w chunk m b ofs =
Some(
w',
t,
v).
Proof.
Lemma do_volatile_store_sound:
forall w chunk m b ofs v w'
t m'
v',
do_volatile_store w chunk m b ofs v =
Some(
w',
t,
m',
v') ->
volatile_store ge chunk m b ofs v t m' /\
possible_trace w t w' /\
v' =
v.
Proof.
Lemma do_volatile_store_complete:
forall w chunk m b ofs v w'
t m',
volatile_store ge chunk m b ofs v t m' ->
possible_trace w t w' ->
do_volatile_store w chunk m b ofs v =
Some(
w',
t,
m',
v).
Proof.
Accessing locations
Definition do_deref_loc (
w:
world) (
ty:
type) (
m:
mem) (
b:
block) (
ofs:
ptrofs) (
bf:
bitfield) :
option (
world *
trace *
val) :=
match bf with
|
Full =>
match access_mode ty with
|
By_value chunk =>
match type_is_volatile ty with
|
false =>
do v <-
Mem.loadv chunk m (
Vptr b ofs);
Some(
w,
E0,
v)
|
true =>
do_volatile_load w chunk m b ofs
end
|
By_reference =>
Some(
w,
E0,
Vptr b ofs)
|
By_copy =>
Some(
w,
E0,
Vptr b ofs)
|
_ =>
None
end
|
Bits sz sg pos width =>
match ty with
|
Tint sz1 sg1 _ =>
check (
intsize_eq sz1 sz &&
signedness_eq sg1 (
if zlt width (
bitsize_intsize sz)
then Signed else sg) &&
zle 0
pos &&
zlt 0
width &&
zle width (
bitsize_intsize sz) &&
zle (
pos +
width) (
bitsize_carrier sz));
match Mem.loadv (
chunk_for_carrier sz)
m (
Vptr b ofs)
with
|
Some (
Vint c) =>
Some (
w,
E0,
Vint (
bitfield_extract sz sg pos width c))
|
_ =>
None
end
|
_ =>
None
end
end.
Definition assign_copy_ok (
ty:
type) (
b:
block) (
ofs:
ptrofs) (
b':
block) (
ofs':
ptrofs) :
Prop :=
(
alignof_blockcopy ge ty |
Ptrofs.unsigned ofs') /\ (
alignof_blockcopy ge ty |
Ptrofs.unsigned ofs) /\
(
b' <>
b \/
Ptrofs.unsigned ofs' =
Ptrofs.unsigned ofs
\/
Ptrofs.unsigned ofs' +
sizeof ge ty <=
Ptrofs.unsigned ofs
\/
Ptrofs.unsigned ofs +
sizeof ge ty <=
Ptrofs.unsigned ofs').
Remark check_assign_copy:
forall (
ty:
type) (
b:
block) (
ofs:
ptrofs) (
b':
block) (
ofs':
ptrofs),
{
assign_copy_ok ty b ofs b'
ofs' } + {~
assign_copy_ok ty b ofs b'
ofs' }.
Proof with
Definition do_assign_loc (
w:
world) (
ty:
type) (
m:
mem) (
b:
block) (
ofs:
ptrofs) (
bf:
bitfield) (
v:
val):
option (
world *
trace *
mem *
val) :=
match bf with
|
Full =>
match access_mode ty with
|
By_value chunk =>
match type_is_volatile ty with
|
false =>
do m' <-
Mem.storev chunk m (
Vptr b ofs)
v;
Some(
w,
E0,
m',
v)
|
true =>
do_volatile_store w chunk m b ofs v
end
|
By_copy =>
match v with
|
Vptr b'
ofs' =>
if check_assign_copy ty b ofs b'
ofs'
then
do bytes <-
Mem.loadbytes m b' (
Ptrofs.unsigned ofs') (
sizeof ge ty);
do m' <-
Mem.storebytes m b (
Ptrofs.unsigned ofs)
bytes;
Some(
w,
E0,
m',
v)
else None
|
_ =>
None
end
|
_ =>
None
end
|
Bits sz sg pos width =>
check (
zle 0
pos &&
zlt 0
width &&
zle width (
bitsize_intsize sz) &&
zle (
pos +
width) (
bitsize_carrier sz));
match ty,
v,
Mem.loadv (
chunk_for_carrier sz)
m (
Vptr b ofs)
with
|
Tint sz1 sg1 _,
Vint n,
Some (
Vint c) =>
check (
intsize_eq sz1 sz &&
signedness_eq sg1 (
if zlt width (
bitsize_intsize sz)
then Signed else sg));
do m' <-
Mem.storev (
chunk_for_carrier sz)
m (
Vptr b ofs)
(
Vint ((
Int.bitfield_insert (
first_bit sz pos width)
width c n)));
Some (
w,
E0,
m',
Vint (
bitfield_normalize sz sg width n))
|
_,
_,
_ =>
None
end
end.
Lemma do_deref_loc_sound:
forall w ty m b ofs bf w'
t v,
do_deref_loc w ty m b ofs bf =
Some(
w',
t,
v) ->
deref_loc ge ty m b ofs bf t v /\
possible_trace w t w'.
Proof.
Lemma do_deref_loc_complete:
forall w ty m b ofs bf w'
t v,
deref_loc ge ty m b ofs bf t v ->
possible_trace w t w' ->
do_deref_loc w ty m b ofs bf =
Some(
w',
t,
v).
Proof.
Lemma do_assign_loc_sound:
forall w ty m b ofs bf v w'
t m'
v',
do_assign_loc w ty m b ofs bf v =
Some(
w',
t,
m',
v') ->
assign_loc ge ty m b ofs bf v t m'
v' /\
possible_trace w t w'.
Proof.
Lemma do_assign_loc_complete:
forall w ty m b ofs bf v w'
t m'
v',
assign_loc ge ty m b ofs bf v t m'
v' ->
possible_trace w t w' ->
do_assign_loc w ty m b ofs bf v =
Some(
w',
t,
m',
v').
Proof.
External calls
Variable do_external_function:
string ->
signature ->
Senv.t ->
world ->
list val ->
mem ->
option (
world *
trace *
val *
mem).
Hypothesis do_external_function_sound:
forall id sg ge vargs m t vres m'
w w',
do_external_function id sg ge w vargs m =
Some(
w',
t,
vres,
m') ->
external_functions_sem id sg ge vargs m t vres m' /\
possible_trace w t w'.
Hypothesis do_external_function_complete:
forall id sg ge vargs m t vres m'
w w',
external_functions_sem id sg ge vargs m t vres m' ->
possible_trace w t w' ->
do_external_function id sg ge w vargs m =
Some(
w',
t,
vres,
m').
Variable do_inline_assembly:
string ->
signature ->
Senv.t ->
world ->
list val ->
mem ->
option (
world *
trace *
val *
mem).
Hypothesis do_inline_assembly_sound:
forall txt sg ge vargs m t vres m'
w w',
do_inline_assembly txt sg ge w vargs m =
Some(
w',
t,
vres,
m') ->
inline_assembly_sem txt sg ge vargs m t vres m' /\
possible_trace w t w'.
Hypothesis do_inline_assembly_complete:
forall txt sg ge vargs m t vres m'
w w',
inline_assembly_sem txt sg ge vargs m t vres m' ->
possible_trace w t w' ->
do_inline_assembly txt sg ge w vargs m =
Some(
w',
t,
vres,
m').
Definition do_ef_volatile_load (
chunk:
memory_chunk)
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
match vargs with
|
Vptr b ofs ::
nil =>
do w',
t,
v <-
do_volatile_load w chunk m b ofs;
Some(
w',
t,
v,
m)
|
_ =>
None
end.
Definition do_ef_volatile_store (
chunk:
memory_chunk)
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
match vargs with
|
Vptr b ofs ::
v ::
nil =>
do w',
t,
m',
v' <-
do_volatile_store w chunk m b ofs v;
Some(
w',
t,
Vundef,
m')
|
_ =>
None
end.
Definition do_ef_volatile_load_global (
chunk:
memory_chunk) (
id:
ident) (
ofs:
ptrofs)
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
do b <-
Genv.find_symbol ge id;
do_ef_volatile_load chunk w (
Vptr b ofs ::
vargs)
m.
Definition do_ef_volatile_store_global (
chunk:
memory_chunk) (
id:
ident) (
ofs:
ptrofs)
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
do b <-
Genv.find_symbol ge id;
do_ef_volatile_store chunk w (
Vptr b ofs ::
vargs)
m.
Definition do_alloc_size (
v:
val) :
option ptrofs :=
match v with
|
Vint n =>
if Archi.ptr64 then None else Some (
Ptrofs.of_int n)
|
Vlong n =>
if Archi.ptr64 then Some (
Ptrofs.of_int64 n)
else None
|
_ =>
None
end.
Definition do_ef_malloc
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
match vargs with
|
v ::
nil =>
do sz <-
do_alloc_size v;
let (
m',
b) :=
Mem.alloc m (-
size_chunk Mptr) (
Ptrofs.unsigned sz)
in
do m'' <-
Mem.store Mptr m'
b (-
size_chunk Mptr)
v;
Some(
w,
E0,
Vptr b Ptrofs.zero,
m'')
|
_ =>
None
end.
Definition do_ef_free
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
match vargs with
|
Vptr b lo ::
nil =>
do vsz <-
Mem.load Mptr m b (
Ptrofs.unsigned lo -
size_chunk Mptr);
do sz <-
do_alloc_size vsz;
do m' <-
Mem.free m b (
Ptrofs.unsigned lo -
size_chunk Mptr) (
Ptrofs.unsigned lo +
Ptrofs.unsigned sz);
Some(
w,
E0,
Vundef,
m')
|
Vint n ::
nil =>
if Int.eq_dec n Int.zero &&
negb Archi.ptr64
then Some(
w,
E0,
Vundef,
m)
else None
|
Vlong n ::
nil =>
if Int64.eq_dec n Int64.zero &&
Archi.ptr64
then Some(
w,
E0,
Vundef,
m)
else None
|
_ =>
None
end.
Definition memcpy_args_ok
(
sz al:
Z) (
bdst:
block) (
odst:
Z) (
bsrc:
block) (
osrc:
Z) :
Prop :=
(
al = 1 \/
al = 2 \/
al = 4 \/
al = 8)
/\
sz >= 0 /\ (
al |
sz)
/\ (
sz > 0 -> (
al |
osrc))
/\ (
sz > 0 -> (
al |
odst))
/\ (
bsrc <>
bdst \/
osrc =
odst \/
osrc +
sz <=
odst \/
odst +
sz <=
osrc).
Definition do_ef_memcpy (
sz al:
Z)
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
match vargs with
|
Vptr bdst odst ::
Vptr bsrc osrc ::
nil =>
if decide (
memcpy_args_ok sz al bdst (
Ptrofs.unsigned odst)
bsrc (
Ptrofs.unsigned osrc))
then
do bytes <-
Mem.loadbytes m bsrc (
Ptrofs.unsigned osrc)
sz;
do m' <-
Mem.storebytes m bdst (
Ptrofs.unsigned odst)
bytes;
Some(
w,
E0,
Vundef,
m')
else None
|
_ =>
None
end.
Definition do_ef_annot (
text:
string) (
targs:
list typ)
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
do args <-
list_eventval_of_val vargs targs;
Some(
w,
Event_annot text args ::
E0,
Vundef,
m).
Definition do_ef_annot_val (
text:
string) (
targ:
typ)
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
match vargs with
|
varg ::
nil =>
do arg <-
eventval_of_val varg targ;
Some(
w,
Event_annot text (
arg ::
nil) ::
E0,
varg,
m)
|
_ =>
None
end.
Definition do_ef_debug (
kind:
positive) (
text:
ident) (
targs:
list typ)
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
Some(
w,
E0,
Vundef,
m).
Definition do_builtin_or_external (
name:
string) (
sg:
signature)
(
w:
world) (
vargs:
list val) (
m:
mem) :
option (
world *
trace *
val *
mem) :=
match lookup_builtin_function name sg with
|
Some bf =>
match builtin_function_sem bf vargs with Some v =>
Some(
w,
E0,
v,
m) |
None =>
None end
|
None =>
do_external_function name sg ge w vargs m
end.
Definition do_external (
ef:
external_function):
world ->
list val ->
mem ->
option (
world *
trace *
val *
mem) :=
match ef with
|
EF_external name sg =>
do_external_function name sg ge
|
EF_builtin name sg =>
do_builtin_or_external name sg
|
EF_runtime name sg =>
do_builtin_or_external name sg
|
EF_vload chunk =>
do_ef_volatile_load chunk
|
EF_vstore chunk =>
do_ef_volatile_store chunk
|
EF_malloc =>
do_ef_malloc
|
EF_free =>
do_ef_free
|
EF_memcpy sz al =>
do_ef_memcpy sz al
|
EF_annot kind text targs =>
do_ef_annot text targs
|
EF_annot_val kind text targ =>
do_ef_annot_val text targ
|
EF_inline_asm text sg clob =>
do_inline_assembly text sg ge
|
EF_debug kind text targs =>
do_ef_debug kind text targs
end.
Lemma do_ef_external_sound:
forall ef w vargs m w'
t vres m',
do_external ef w vargs m =
Some(
w',
t,
vres,
m') ->
external_call ef ge vargs m t vres m' /\
possible_trace w t w'.
Proof with
Lemma do_ef_external_complete:
forall ef w vargs m w'
t vres m',
external_call ef ge vargs m t vres m' ->
possible_trace w t w' ->
do_external ef w vargs m =
Some(
w',
t,
vres,
m').
Proof.
Reduction of expressions
Inductive reduction:
Type :=
|
Lred (
rule:
string) (
l':
expr) (
m':
mem)
|
Rred (
rule:
string) (
r':
expr) (
m':
mem) (
t:
trace)
|
Callred (
rule:
string) (
fd:
fundef) (
args:
list val) (
tyres:
type) (
m':
mem)
|
Stuckred.
Section EXPRS.
Variable e:
env.
Variable w:
world.
Fixpoint sem_cast_arguments (
vtl:
list (
val *
type)) (
tl:
typelist) (
m:
mem) :
option (
list val) :=
match vtl,
tl with
|
nil,
Tnil =>
Some nil
| (
v1,
t1)::
vtl,
Tcons t1'
tl =>
do v <-
sem_cast v1 t1 t1'
m;
do vl <-
sem_cast_arguments vtl tl m;
Some(
v::
vl)
|
_,
_ =>
None
end.
The result of stepping an expression is a list ll of possible reducts.
Each element of ll is a pair of a context and the result of reducing
inside this context (see type reduction above).
The list ll is empty if the expression is fully reduced
(it's Eval for a r-value and Eloc for a l-value).
Definition reducts (
A:
Type):
Type :=
list ((
expr ->
A) *
reduction).
Definition topred (
r:
reduction) :
reducts expr :=
((
fun (
x:
expr) =>
x),
r) ::
nil.
Definition stuck :
reducts expr :=
((
fun (
x:
expr) =>
x),
Stuckred) ::
nil.
Definition incontext {
A B:
Type} (
ctx:
A ->
B) (
ll:
reducts A) :
reducts B :=
map (
fun z => ((
fun (
x:
expr) =>
ctx(
fst z x)),
snd z))
ll.
Definition incontext2 {
A1 A2 B:
Type}
(
ctx1:
A1 ->
B) (
ll1:
reducts A1)
(
ctx2:
A2 ->
B) (
ll2:
reducts A2) :
reducts B :=
incontext ctx1 ll1 ++
incontext ctx2 ll2.
Declare Scope reducts_monad_scope.
Notation "'
do'
X <-
A ;
B" := (
match A with Some X =>
B |
None =>
stuck end)
(
at level 200,
X ident,
A at level 100,
B at level 200)
:
reducts_monad_scope.
Notation "'
do'
X ,
Y <-
A ;
B" := (
match A with Some (
X,
Y) =>
B |
None =>
stuck end)
(
at level 200,
X ident,
Y ident,
A at level 100,
B at level 200)
:
reducts_monad_scope.
Notation "'
do'
X ,
Y ,
Z <-
A ;
B" := (
match A with Some (
X,
Y,
Z) =>
B |
None =>
stuck end)
(
at level 200,
X ident,
Y ident,
Z ident,
A at level 100,
B at level 200)
:
reducts_monad_scope.
Notation "'
do'
X ,
Y ,
Z ,
W <-
A ;
B" := (
match A with Some (
X,
Y,
Z,
W) =>
B |
None =>
stuck end)
(
at level 200,
X ident,
Y ident,
Z ident,
W ident,
A at level 100,
B at level 200)
:
reducts_monad_scope.
Notation " '
check'
A ;
B" := (
if A then B else stuck)
(
at level 200,
A at level 100,
B at level 200)
:
reducts_monad_scope.
Local Open Scope reducts_monad_scope.
Fixpoint step_expr (
k:
kind) (
a:
expr) (
m:
mem):
reducts expr :=
match k,
a with
|
LV,
Eloc b ofs bf ty =>
nil
|
LV,
Evar x ty =>
match e!
x with
|
Some(
b,
ty') =>
check type_eq ty ty';
topred (
Lred "
red_var_local" (
Eloc b Ptrofs.zero Full ty)
m)
|
None =>
do b <-
Genv.find_symbol ge x;
topred (
Lred "
red_var_global" (
Eloc b Ptrofs.zero Full ty)
m)
end
|
LV,
Ederef r ty =>
match is_val r with
|
Some(
Vptr b ofs,
ty') =>
topred (
Lred "
red_deref" (
Eloc b ofs Full ty)
m)
|
Some _ =>
stuck
|
None =>
incontext (
fun x =>
Ederef x ty) (
step_expr RV r m)
end
|
LV,
Efield r f ty =>
match is_val r with
|
Some(
Vptr b ofs,
ty') =>
match ty'
with
|
Tstruct id _ =>
do co <-
ge.(
genv_cenv)!
id;
match field_offset ge f (
co_members co)
with
|
Error _ =>
stuck
|
OK (
delta,
bf) =>
topred (
Lred "
red_field_struct" (
Eloc b (
Ptrofs.add ofs (
Ptrofs.repr delta))
bf ty)
m)
end
|
Tunion id _ =>
do co <-
ge.(
genv_cenv)!
id;
match union_field_offset ge f (
co_members co)
with
|
Error _ =>
stuck
|
OK (
delta,
bf) =>
topred (
Lred "
red_field_union" (
Eloc b (
Ptrofs.add ofs (
Ptrofs.repr delta))
bf ty)
m)
end
|
_ =>
stuck
end
|
Some _ =>
stuck
|
None =>
incontext (
fun x =>
Efield x f ty) (
step_expr RV r m)
end
|
RV,
Eval v ty =>
nil
|
RV,
Evalof l ty =>
match is_loc l with
|
Some(
b,
ofs,
bf,
ty') =>
check type_eq ty ty';
do w',
t,
v <-
do_deref_loc w ty m b ofs bf;
topred (
Rred "
red_rvalof" (
Eval v ty)
m t)
|
None =>
incontext (
fun x =>
Evalof x ty) (
step_expr LV l m)
end
|
RV,
Eaddrof l ty =>
match is_loc l with
|
Some(
b,
ofs,
bf,
ty') =>
match bf with Full =>
topred (
Rred "
red_addrof" (
Eval (
Vptr b ofs)
ty)
m E0)
|
Bits _ _ _ _ =>
stuck
end
|
None =>
incontext (
fun x =>
Eaddrof x ty) (
step_expr LV l m)
end
|
RV,
Eunop op r1 ty =>
match is_val r1 with
|
Some(
v1,
ty1) =>
do v <-
sem_unary_operation op v1 ty1 m;
topred (
Rred "
red_unop" (
Eval v ty)
m E0)
|
None =>
incontext (
fun x =>
Eunop op x ty) (
step_expr RV r1 m)
end
|
RV,
Ebinop op r1 r2 ty =>
match is_val r1,
is_val r2 with
|
Some(
v1,
ty1),
Some(
v2,
ty2) =>
do v <-
sem_binary_operation ge op v1 ty1 v2 ty2 m;
topred (
Rred "
red_binop" (
Eval v ty)
m E0)
|
_,
_ =>
incontext2 (
fun x =>
Ebinop op x r2 ty) (
step_expr RV r1 m)
(
fun x =>
Ebinop op r1 x ty) (
step_expr RV r2 m)
end
|
RV,
Ecast r1 ty =>
match is_val r1 with
|
Some(
v1,
ty1) =>
do v <-
sem_cast v1 ty1 ty m;
topred (
Rred "
red_cast" (
Eval v ty)
m E0)
|
None =>
incontext (
fun x =>
Ecast x ty) (
step_expr RV r1 m)
end
|
RV,
Eseqand r1 r2 ty =>
match is_val r1 with
|
Some(
v1,
ty1) =>
do b <-
bool_val v1 ty1 m;
if b then topred (
Rred "
red_seqand_true" (
Eparen r2 type_bool ty)
m E0)
else topred (
Rred "
red_seqand_false" (
Eval (
Vint Int.zero)
ty)
m E0)
|
None =>
incontext (
fun x =>
Eseqand x r2 ty) (
step_expr RV r1 m)
end
|
RV,
Eseqor r1 r2 ty =>
match is_val r1 with
|
Some(
v1,
ty1) =>
do b <-
bool_val v1 ty1 m;
if b then topred (
Rred "
red_seqor_true" (
Eval (
Vint Int.one)
ty)
m E0)
else topred (
Rred "
red_seqor_false" (
Eparen r2 type_bool ty)
m E0)
|
None =>
incontext (
fun x =>
Eseqor x r2 ty) (
step_expr RV r1 m)
end
|
RV,
Econdition r1 r2 r3 ty =>
match is_val r1 with
|
Some(
v1,
ty1) =>
do b <-
bool_val v1 ty1 m;
topred (
Rred "
red_condition" (
Eparen (
if b then r2 else r3)
ty ty)
m E0)
|
None =>
incontext (
fun x =>
Econdition x r2 r3 ty) (
step_expr RV r1 m)
end
|
RV,
Esizeof ty'
ty =>
topred (
Rred "
red_sizeof" (
Eval (
Vptrofs (
Ptrofs.repr (
sizeof ge ty')))
ty)
m E0)
|
RV,
Ealignof ty'
ty =>
topred (
Rred "
red_alignof" (
Eval (
Vptrofs (
Ptrofs.repr (
alignof ge ty')))
ty)
m E0)
|
RV,
Eassign l1 r2 ty =>
match is_loc l1,
is_val r2 with
|
Some(
b,
ofs,
bf,
ty1),
Some(
v2,
ty2) =>
check type_eq ty1 ty;
do v <-
sem_cast v2 ty2 ty1 m;
do w',
t,
m',
v' <-
do_assign_loc w ty1 m b ofs bf v;
topred (
Rred "
red_assign" (
Eval v'
ty)
m'
t)
|
_,
_ =>
incontext2 (
fun x =>
Eassign x r2 ty) (
step_expr LV l1 m)
(
fun x =>
Eassign l1 x ty) (
step_expr RV r2 m)
end
|
RV,
Eassignop op l1 r2 tyres ty =>
match is_loc l1,
is_val r2 with
|
Some(
b,
ofs,
bf,
ty1),
Some(
v2,
ty2) =>
check type_eq ty1 ty;
do w',
t,
v1 <-
do_deref_loc w ty1 m b ofs bf;
let r' :=
Eassign (
Eloc b ofs bf ty1)
(
Ebinop op (
Eval v1 ty1) (
Eval v2 ty2)
tyres)
ty1 in
topred (
Rred "
red_assignop"
r'
m t)
|
_,
_ =>
incontext2 (
fun x =>
Eassignop op x r2 tyres ty) (
step_expr LV l1 m)
(
fun x =>
Eassignop op l1 x tyres ty) (
step_expr RV r2 m)
end
|
RV,
Epostincr id l ty =>
match is_loc l with
|
Some(
b,
ofs,
bf,
ty1) =>
check type_eq ty1 ty;
do w',
t,
v1 <-
do_deref_loc w ty m b ofs bf;
let op :=
match id with Incr =>
Oadd |
Decr =>
Osub end in
let r' :=
Ecomma (
Eassign (
Eloc b ofs bf ty)
(
Ebinop op (
Eval v1 ty) (
Eval (
Vint Int.one)
type_int32s) (
incrdecr_type ty))
ty)
(
Eval v1 ty)
ty in
topred (
Rred "
red_postincr"
r'
m t)
|
None =>
incontext (
fun x =>
Epostincr id x ty) (
step_expr LV l m)
end
|
RV,
Ecomma r1 r2 ty =>
match is_val r1 with
|
Some _ =>
check type_eq (
typeof r2)
ty;
topred (
Rred "
red_comma"
r2 m E0)
|
None =>
incontext (
fun x =>
Ecomma x r2 ty) (
step_expr RV r1 m)
end
|
RV,
Eparen r1 tycast ty =>
match is_val r1 with
|
Some (
v1,
ty1) =>
do v <-
sem_cast v1 ty1 tycast m;
topred (
Rred "
red_paren" (
Eval v ty)
m E0)
|
None =>
incontext (
fun x =>
Eparen x tycast ty) (
step_expr RV r1 m)
end
|
RV,
Ecall r1 rargs ty =>
match is_val r1,
is_val_list rargs with
|
Some(
vf,
tyf),
Some vtl =>
match classify_fun tyf with
|
fun_case_f tyargs tyres cconv =>
do fd <-
Genv.find_funct ge vf;
do vargs <-
sem_cast_arguments vtl tyargs m;
check type_eq (
type_of_fundef fd) (
Tfunction tyargs tyres cconv);
topred (
Callred "
red_call"
fd vargs ty m)
|
_ =>
stuck
end
|
_,
_ =>
incontext2 (
fun x =>
Ecall x rargs ty) (
step_expr RV r1 m)
(
fun x =>
Ecall r1 x ty) (
step_exprlist rargs m)
end
|
RV,
Ebuiltin ef tyargs rargs ty =>
match is_val_list rargs with
|
Some vtl =>
do vargs <-
sem_cast_arguments vtl tyargs m;
match do_external ef w vargs m with
|
None =>
stuck
|
Some(
w',
t,
v,
m') =>
topred (
Rred "
red_builtin" (
Eval v ty)
m'
t)
end
|
_ =>
incontext (
fun x =>
Ebuiltin ef tyargs x ty) (
step_exprlist rargs m)
end
|
_,
_ =>
stuck
end
with step_exprlist (
rl:
exprlist) (
m:
mem):
reducts exprlist :=
match rl with
|
Enil =>
nil
|
Econs r1 rs =>
incontext2 (
fun x =>
Econs x rs) (
step_expr RV r1 m)
(
fun x =>
Econs r1 x) (
step_exprlist rs m)
end.
Technical properties on safe expressions.
Inductive imm_safe_t:
kind ->
expr ->
mem ->
Prop :=
|
imm_safe_t_val:
forall v ty m,
imm_safe_t RV (
Eval v ty)
m
|
imm_safe_t_loc:
forall b ofs ty bf m,
imm_safe_t LV (
Eloc b ofs bf ty)
m
|
imm_safe_t_lred:
forall to C l m l'
m',
lred ge e l m l'
m' ->
context LV to C ->
imm_safe_t to (
C l)
m
|
imm_safe_t_rred:
forall to C r m t r'
m'
w',
rred ge r m t r'
m' ->
possible_trace w t w' ->
context RV to C ->
imm_safe_t to (
C r)
m
|
imm_safe_t_callred:
forall to C r m fd args ty,
callred ge r m fd args ty ->
context RV to C ->
imm_safe_t to (
C r)
m.
Remark imm_safe_t_imm_safe:
forall k a m,
imm_safe_t k a m ->
imm_safe ge e k a m.
Proof.
Fixpoint exprlist_all_values (
rl:
exprlist) :
Prop :=
match rl with
|
Enil =>
True
|
Econs (
Eval v ty)
rl' =>
exprlist_all_values rl'
|
Econs _ _ =>
False
end.
Definition invert_expr_prop (
a:
expr) (
m:
mem) :
Prop :=
match a with
|
Eloc b ofs bf ty =>
False
|
Evar x ty =>
exists b,
e!
x =
Some(
b,
ty)
\/ (
e!
x =
None /\
Genv.find_symbol ge x =
Some b)
|
Ederef (
Eval v ty1)
ty =>
exists b,
exists ofs,
v =
Vptr b ofs
|
Eaddrof (
Eloc b ofs bf ty1)
ty =>
bf =
Full
|
Efield (
Eval v ty1)
f ty =>
exists b,
exists ofs,
v =
Vptr b ofs /\
match ty1 with
|
Tstruct id _ =>
exists co delta bf,
ge.(
genv_cenv)!
id =
Some co /\
field_offset ge f (
co_members co) =
OK (
delta,
bf)
|
Tunion id _ =>
exists co delta bf,
ge.(
genv_cenv)!
id =
Some co /\
union_field_offset ge f (
co_members co) =
OK (
delta,
bf)
|
_ =>
False
end
|
Eval v ty =>
False
|
Evalof (
Eloc b ofs bf ty')
ty =>
ty' =
ty /\
exists t,
exists v,
exists w',
deref_loc ge ty m b ofs bf t v /\
possible_trace w t w'
|
Eunop op (
Eval v1 ty1)
ty =>
exists v,
sem_unary_operation op v1 ty1 m =
Some v
|
Ebinop op (
Eval v1 ty1) (
Eval v2 ty2)
ty =>
exists v,
sem_binary_operation ge op v1 ty1 v2 ty2 m =
Some v
|
Ecast (
Eval v1 ty1)
ty =>
exists v,
sem_cast v1 ty1 ty m =
Some v
|
Eseqand (
Eval v1 ty1)
r2 ty =>
exists b,
bool_val v1 ty1 m =
Some b
|
Eseqor (
Eval v1 ty1)
r2 ty =>
exists b,
bool_val v1 ty1 m =
Some b
|
Econdition (
Eval v1 ty1)
r1 r2 ty =>
exists b,
bool_val v1 ty1 m =
Some b
|
Eassign (
Eloc b ofs bf ty1) (
Eval v2 ty2)
ty =>
exists v,
exists m',
exists v',
exists t,
exists w',
ty =
ty1 /\
sem_cast v2 ty2 ty1 m =
Some v /\
assign_loc ge ty1 m b ofs bf v t m'
v' /\
possible_trace w t w'
|
Eassignop op (
Eloc b ofs bf ty1) (
Eval v2 ty2)
tyres ty =>
exists t,
exists v1,
exists w',
ty =
ty1 /\
deref_loc ge ty1 m b ofs bf t v1 /\
possible_trace w t w'
|
Epostincr id (
Eloc b ofs bf ty1)
ty =>
exists t,
exists v1,
exists w',
ty =
ty1 /\
deref_loc ge ty m b ofs bf t v1 /\
possible_trace w t w'
|
Ecomma (
Eval v ty1)
r2 ty =>
typeof r2 =
ty
|
Eparen (
Eval v1 ty1)
tycast ty =>
exists v,
sem_cast v1 ty1 tycast m =
Some v
|
Ecall (
Eval vf tyf)
rargs ty =>
exprlist_all_values rargs ->
exists tyargs tyres cconv fd vl,
classify_fun tyf =
fun_case_f tyargs tyres cconv
/\
Genv.find_funct ge vf =
Some fd
/\
cast_arguments m rargs tyargs vl
/\
type_of_fundef fd =
Tfunction tyargs tyres cconv
|
Ebuiltin ef tyargs rargs ty =>
exprlist_all_values rargs ->
exists vargs t vres m'
w',
cast_arguments m rargs tyargs vargs
/\
external_call ef ge vargs m t vres m'
/\
possible_trace w t w'
|
_ =>
True
end.
Lemma lred_invert:
forall l m l'
m',
lred ge e l m l'
m' ->
invert_expr_prop l m.
Proof.
induction 1; red; auto.
exists b; auto.
exists b; auto.
exists b; exists ofs; auto.
exists b; exists ofs; split; auto. exists co, delta, bf; auto.
exists b; exists ofs; split; auto. exists co, delta, bf; auto.
Qed.
Lemma rred_invert:
forall w'
r m t r'
m',
rred ge r m t r'
m' ->
possible_trace w t w' ->
invert_expr_prop r m.
Proof.
induction 1;
intros;
red;
auto.
split;
auto;
exists t;
exists v;
exists w';
auto.
exists v;
auto.
exists v;
auto.
exists v;
auto.
exists true;
auto.
exists false;
auto.
exists true;
auto.
exists false;
auto.
exists b;
auto.
exists v;
exists m';
exists v';
exists t;
exists w';
auto.
exists t;
exists v1;
exists w';
auto.
exists t;
exists v1;
exists w';
auto.
exists v;
auto.
intros;
exists vargs;
exists t;
exists vres;
exists m';
exists w';
auto.
Qed.
Lemma callred_invert:
forall r fd args ty m,
callred ge r m fd args ty ->
invert_expr_prop r m.
Proof.
intros. inv H. simpl.
intros. exists tyargs, tyres, cconv, fd, args; auto.
Qed.
Scheme context_ind2 :=
Minimality for context Sort Prop
with contextlist_ind2 :=
Minimality for contextlist Sort Prop.
Combined Scheme context_contextlist_ind from context_ind2,
contextlist_ind2.
Lemma invert_expr_context:
(
forall from to C,
context from to C ->
forall a m,
invert_expr_prop a m ->
invert_expr_prop (
C a)
m)
/\(
forall from C,
contextlist from C ->
forall a m,
invert_expr_prop a m ->
~
exprlist_all_values (
C a)).
Proof.
apply context_contextlist_ind;
intros;
try (
exploit H0; [
eauto|
intros]);
simpl.
auto.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct e1;
auto;
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct e1;
auto;
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct e1;
auto;
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
destruct e1;
auto.
intros.
elim (
H0 a m);
auto.
intros.
elim (
H0 a m);
auto.
destruct (
C a);
auto;
contradiction.
destruct (
C a);
auto;
contradiction.
red;
intros.
destruct (
C a);
auto.
red;
intros.
destruct e1;
auto.
elim (
H0 a m);
auto.
Qed.
Lemma imm_safe_t_inv:
forall k a m,
imm_safe_t k a m ->
match a with
|
Eloc _ _ _ _ =>
True
|
Eval _ _ =>
True
|
_ =>
invert_expr_prop a m
end.
Proof.
Soundness: if step_expr returns Some ll, then every element
of ll is a reduct.
Lemma context_compose:
forall k2 k3 C2,
context k2 k3 C2 ->
forall k1 C1,
context k1 k2 C1 ->
context k1 k3 (
fun x =>
C2(
C1 x))
with contextlist_compose:
forall k2 C2,
contextlist k2 C2 ->
forall k1 C1,
context k1 k2 C1 ->
contextlist k1 (
fun x =>
C2(
C1 x)).
Proof.
induction 1; intros; try (constructor; eauto).
replace (fun x => C1 x) with C1. auto. apply extensionality; auto.
induction 1; intros; constructor; eauto.
Qed.
Local Hint Constructors context contextlist :
core.
Local Hint Resolve context_compose contextlist_compose :
core.
Definition reduction_ok (
k:
kind) (
a:
expr) (
m:
mem) (
rd:
reduction) :
Prop :=
match k,
rd with
|
LV,
Lred _ l'
m' =>
lred ge e a m l'
m'
|
RV,
Rred _ r'
m'
t =>
rred ge a m t r'
m' /\
exists w',
possible_trace w t w'
|
RV,
Callred _ fd args tyres m' =>
callred ge a m fd args tyres /\
m' =
m
|
LV,
Stuckred => ~
imm_safe_t k a m
|
RV,
Stuckred => ~
imm_safe_t k a m
|
_,
_ =>
False
end.
Definition reducts_ok (
k:
kind) (
a:
expr) (
m:
mem) (
ll:
reducts expr) :
Prop :=
(
forall C rd,
In (
C,
rd)
ll ->
exists a',
exists k',
context k'
k C /\
a =
C a' /\
reduction_ok k'
a'
m rd)
/\ (
ll =
nil ->
match k with LV =>
is_loc a <>
None |
RV =>
is_val a <>
None end).
Definition list_reducts_ok (
al:
exprlist) (
m:
mem) (
ll:
reducts exprlist) :
Prop :=
(
forall C rd,
In (
C,
rd)
ll ->
exists a',
exists k',
contextlist k'
C /\
al =
C a' /\
reduction_ok k'
a'
m rd)
/\ (
ll =
nil ->
is_val_list al <>
None).
Ltac monadInv :=
match goal with
| [
H:
match ?
x with Some _ =>
_ |
None =>
None end =
Some ?
y |-
_ ] =>
destruct x eqn:?; [
monadInv|
discriminate]
| [
H:
match ?
x with left _ =>
_ |
right _ =>
None end =
Some ?
y |-
_ ] =>
destruct x; [
monadInv|
discriminate]
|
_ =>
idtac
end.
Lemma sem_cast_arguments_sound:
forall m rargs vtl tyargs vargs,
is_val_list rargs =
Some vtl ->
sem_cast_arguments vtl tyargs m =
Some vargs ->
cast_arguments m rargs tyargs vargs.
Proof.
induction rargs;
simpl;
intros.
inv H.
destruct tyargs;
simpl in H0;
inv H0.
constructor.
monadInv.
inv H.
simpl in H0.
destruct p as [
v1 t1].
destruct tyargs;
try congruence.
monadInv.
inv H0.
rewrite (
is_val_inv _ _ _ Heqo).
constructor.
auto.
eauto.
Qed.
Lemma sem_cast_arguments_complete:
forall m al tyl vl,
cast_arguments m al tyl vl ->
exists vtl,
is_val_list al =
Some vtl /\
sem_cast_arguments vtl tyl m =
Some vl.
Proof.
induction 1.
exists (@
nil (
val *
type));
auto.
destruct IHcast_arguments as [
vtl [
A B]].
exists ((
v,
ty) ::
vtl);
simpl.
rewrite A;
rewrite B;
rewrite H.
auto.
Qed.
Lemma topred_ok:
forall k a m rd,
reduction_ok k a m rd ->
reducts_ok k a m (
topred rd).
Proof.
intros.
unfold topred;
split;
simpl;
intros.
destruct H0;
try contradiction.
inv H0.
exists a;
exists k;
auto.
congruence.
Qed.
Lemma stuck_ok:
forall k a m,
~
imm_safe_t k a m ->
reducts_ok k a m stuck.
Proof.
intros.
unfold stuck;
split;
simpl;
intros.
destruct H0;
try contradiction.
inv H0.
exists a;
exists k;
intuition.
red.
destruct k;
auto.
congruence.
Qed.
Lemma wrong_kind_ok:
forall k a m,
k <>
Cstrategy.expr_kind a ->
reducts_ok k a m stuck.
Proof.
Lemma not_invert_ok:
forall k a m,
match a with
|
Eloc _ _ _ _ =>
False
|
Eval _ _ =>
False
|
_ =>
invert_expr_prop a m ->
False
end ->
reducts_ok k a m stuck.
Proof.
Lemma incontext_ok:
forall k a m C res k'
a',
reducts_ok k'
a'
m res ->
a =
C a' ->
context k'
k C ->
match k'
with LV =>
is_loc a' =
None |
RV =>
is_val a' =
None end ->
reducts_ok k a m (
incontext C res).
Proof.
unfold reducts_ok,
incontext;
intros.
destruct H.
split;
intros.
exploit list_in_map_inv;
eauto.
intros [[
C1 rd1] [
P Q]].
inv P.
exploit H;
eauto.
intros [
a'' [
k'' [
U [
V W]]]].
exists a'';
exists k''.
split.
eapply context_compose;
eauto.
rewrite V;
auto.
destruct res;
simpl in H4;
try congruence.
destruct k';
intuition congruence.
Qed.
Lemma incontext2_ok:
forall k a m k1 a1 res1 k2 a2 res2 C1 C2,
reducts_ok k1 a1 m res1 ->
reducts_ok k2 a2 m res2 ->
a =
C1 a1 ->
a =
C2 a2 ->
context k1 k C1 ->
context k2 k C2 ->
match k1 with LV =>
is_loc a1 =
None |
RV =>
is_val a1 =
None end
\/
match k2 with LV =>
is_loc a2 =
None |
RV =>
is_val a2 =
None end ->
reducts_ok k a m (
incontext2 C1 res1 C2 res2).
Proof.
unfold reducts_ok,
incontext2,
incontext;
intros.
destruct H;
destruct H0;
split;
intros.
destruct (
in_app_or _ _ _ H8).
exploit list_in_map_inv;
eauto.
intros [[
C'
rd'] [
P Q]].
inv P.
exploit H;
eauto.
intros [
a'' [
k'' [
U [
V W]]]].
exists a'';
exists k''.
split.
eapply context_compose;
eauto.
rewrite V;
auto.
exploit list_in_map_inv;
eauto.
intros [[
C'
rd'] [
P Q]].
inv P.
exploit H0;
eauto.
intros [
a'' [
k'' [
U [
V W]]]].
exists a'';
exists k''.
split.
eapply context_compose;
eauto.
rewrite H2;
rewrite V;
auto.
destruct res1;
simpl in H8;
try congruence.
destruct res2;
simpl in H8;
try congruence.
destruct H5.
destruct k1;
intuition congruence.
destruct k2;
intuition congruence.
Qed.
Lemma incontext_list_ok:
forall ef tyargs al ty m res,
list_reducts_ok al m res ->
is_val_list al =
None ->
reducts_ok RV (
Ebuiltin ef tyargs al ty)
m
(
incontext (
fun x =>
Ebuiltin ef tyargs x ty)
res).
Proof.
unfold reducts_ok,
incontext;
intros.
destruct H.
split;
intros.
exploit list_in_map_inv;
eauto.
intros [[
C1 rd1] [
P Q]].
inv P.
exploit H;
eauto.
intros [
a'' [
k'' [
U [
V W]]]].
exists a'';
exists k''.
split.
eauto.
rewrite V;
auto.
destruct res;
simpl in H2.
elim H1;
auto.
congruence.
Qed.
Lemma incontext2_list_ok:
forall a1 a2 ty m res1 res2,
reducts_ok RV a1 m res1 ->
list_reducts_ok a2 m res2 ->
is_val a1 =
None \/
is_val_list a2 =
None ->
reducts_ok RV (
Ecall a1 a2 ty)
m
(
incontext2 (
fun x =>
Ecall x a2 ty)
res1
(
fun x =>
Ecall a1 x ty)
res2).
Proof.
unfold reducts_ok,
incontext2,
incontext;
intros.
destruct H;
destruct H0;
split;
intros.
destruct (
in_app_or _ _ _ H4).
exploit list_in_map_inv;
eauto.
intros [[
C'
rd'] [
P Q]].
inv P.
exploit H;
eauto.
intros [
a'' [
k'' [
U [
V W]]]].
exists a'';
exists k''.
split.
eauto.
rewrite V;
auto.
exploit list_in_map_inv;
eauto.
intros [[
C'
rd'] [
P Q]].
inv P.
exploit H0;
eauto.
intros [
a'' [
k'' [
U [
V W]]]].
exists a'';
exists k''.
split.
eauto.
rewrite V;
auto.
destruct res1;
simpl in H4;
try congruence.
destruct res2;
simpl in H4;
try congruence.
tauto.
Qed.
Lemma incontext2_list_ok':
forall a1 a2 m res1 res2,
reducts_ok RV a1 m res1 ->
list_reducts_ok a2 m res2 ->
list_reducts_ok (
Econs a1 a2)
m
(
incontext2 (
fun x =>
Econs x a2)
res1
(
fun x =>
Econs a1 x)
res2).
Proof.
unfold reducts_ok,
list_reducts_ok,
incontext2,
incontext;
intros.
destruct H;
destruct H0.
split;
intros.
destruct (
in_app_or _ _ _ H3).
exploit list_in_map_inv;
eauto.
intros [[
C'
rd'] [
P Q]].
inv P.
exploit H;
eauto.
intros [
a'' [
k'' [
U [
V W]]]].
exists a'';
exists k''.
split.
eauto.
rewrite V;
auto.
exploit list_in_map_inv;
eauto.
intros [[
C'
rd'] [
P Q]].
inv P.
exploit H0;
eauto.
intros [
a'' [
k'' [
U [
V W]]]].
exists a'';
exists k''.
split.
eauto.
rewrite V;
auto.
destruct res1;
simpl in H3;
try congruence.
destruct res2;
simpl in H3;
try congruence.
simpl.
destruct (
is_val a1).
destruct (
is_val_list a2).
congruence.
intuition congruence.
intuition congruence.
Qed.
Lemma is_val_list_all_values:
forall al vtl,
is_val_list al =
Some vtl ->
exprlist_all_values al.
Proof.
induction al;
simpl;
intros.
auto.
destruct (
is_val r1)
as [[
v ty]|]
eqn:?;
try discriminate.
destruct (
is_val_list al)
as [
vtl'|]
eqn:?;
try discriminate.
rewrite (
is_val_inv _ _ _ Heqo).
eauto.
Qed.
Ltac myinv :=
match goal with
| [
H:
False |-
_ ] =>
destruct H
| [
H:
_ /\
_ |-
_ ] =>
destruct H;
myinv
| [
H:
exists _,
_ |-
_ ] =>
destruct H;
myinv
|
_ =>
idtac
end.
Theorem step_expr_sound:
forall a k m,
reducts_ok k a m (
step_expr k a m)
with step_exprlist_sound:
forall al m,
list_reducts_ok al m (
step_exprlist al m).
Proof with
Lemma step_exprlist_val_list:
forall m al,
is_val_list al <>
None ->
step_exprlist al m =
nil.
Proof.
induction al;
simpl;
intros.
auto.
destruct (
is_val r1)
as [[
v1 ty1]|]
eqn:?;
try congruence.
destruct (
is_val_list al)
eqn:?;
try congruence.
rewrite (
is_val_inv _ _ _ Heqo).
rewrite IHal.
auto.
congruence.
Qed.
Completeness part 1: step_expr contains all possible non-error reducts.
Lemma lred_topred:
forall l1 m1 l2 m2,
lred ge e l1 m1 l2 m2 ->
exists rule,
step_expr LV l1 m1 =
topred (
Lred rule l2 m2).
Proof.
induction 1;
simpl.
rewrite H.
rewrite dec_eq_true.
econstructor;
eauto.
rewrite H;
rewrite H0.
econstructor;
eauto.
econstructor;
eauto.
rewrite H,
H0;
econstructor;
eauto.
rewrite H,
H0;
econstructor;
eauto.
Qed.
Lemma rred_topred:
forall w'
r1 m1 t r2 m2,
rred ge r1 m1 t r2 m2 ->
possible_trace w t w' ->
exists rule,
step_expr RV r1 m1 =
topred (
Rred rule r2 m2 t).
Proof.
induction 1;
simpl;
intros.
rewrite dec_eq_true.
rewrite (
do_deref_loc_complete _ _ _ _ _ _ _ _ _ H H0).
econstructor;
eauto.
inv H.
econstructor;
eauto.
inv H0.
rewrite H;
econstructor;
eauto.
inv H0.
rewrite H;
econstructor;
eauto.
inv H0.
rewrite H;
econstructor;
eauto.
inv H0.
rewrite H;
econstructor;
eauto.
inv H0.
rewrite H;
econstructor;
eauto.
inv H0.
rewrite H;
econstructor;
eauto.
inv H0.
rewrite H;
econstructor;
eauto.
inv H0.
rewrite H;
econstructor;
eauto.
inv H.
econstructor;
eauto.
inv H.
econstructor;
eauto.
rewrite dec_eq_true.
rewrite H.
rewrite (
do_assign_loc_complete _ _ _ _ _ _ _ _ _ _ _ H0 H1).
econstructor;
eauto.
rewrite dec_eq_true.
rewrite (
do_deref_loc_complete _ _ _ _ _ _ _ _ _ H H0).
econstructor;
eauto.
rewrite dec_eq_true.
subst.
rewrite (
do_deref_loc_complete _ _ _ _ _ _ _ _ _ H H1).
econstructor;
eauto.
inv H0.
rewrite dec_eq_true.
econstructor;
eauto.
inv H0.
rewrite H;
econstructor;
eauto.
exploit sem_cast_arguments_complete;
eauto.
intros [
vtl [
A B]].
exploit do_ef_external_complete;
eauto.
intros C.
rewrite A.
rewrite B.
rewrite C.
econstructor;
eauto.
Qed.
Lemma callred_topred:
forall a fd args ty m,
callred ge a m fd args ty ->
exists rule,
step_expr RV a m =
topred (
Callred rule fd args ty m).
Proof.
Definition reducts_incl {
A B:
Type} (
C:
A ->
B) (
res1:
reducts A) (
res2:
reducts B) :
Prop :=
forall C1 rd,
In (
C1,
rd)
res1 ->
In ((
fun x =>
C(
C1 x)),
rd)
res2.
Lemma reducts_incl_trans:
forall (
A1 A2:
Type) (
C:
A1 ->
A2)
res1 res2,
reducts_incl C res1 res2 ->
forall (
A3:
Type) (
C':
A2 ->
A3)
res3,
reducts_incl C'
res2 res3 ->
reducts_incl (
fun x =>
C'(
C x))
res1 res3.
Proof.
Lemma reducts_incl_nil:
forall (
A B:
Type) (
C:
A ->
B)
res,
reducts_incl C nil res.
Proof.
intros; red. intros; contradiction.
Qed.
Lemma reducts_incl_val:
forall (
A:
Type)
a m v ty (
C:
expr ->
A)
res,
is_val a =
Some(
v,
ty) ->
reducts_incl C (
step_expr RV a m)
res.
Proof.
Lemma reducts_incl_loc:
forall (
A:
Type)
a m b ofs ty bf (
C:
expr ->
A)
res,
is_loc a =
Some(
b,
ofs,
bf,
ty) ->
reducts_incl C (
step_expr LV a m)
res.
Proof.
Lemma reducts_incl_listval:
forall (
A:
Type)
a m vtl (
C:
exprlist ->
A)
res,
is_val_list a =
Some vtl ->
reducts_incl C (
step_exprlist a m)
res.
Proof.
Lemma reducts_incl_incontext:
forall (
A B:
Type) (
C:
A ->
B)
res,
reducts_incl C res (
incontext C res).
Proof.
Lemma reducts_incl_incontext2_left:
forall (
A1 A2 B:
Type) (
C1:
A1 ->
B)
res1 (
C2:
A2 ->
B)
res2,
reducts_incl C1 res1 (
incontext2 C1 res1 C2 res2).
Proof.
Lemma reducts_incl_incontext2_right:
forall (
A1 A2 B:
Type) (
C1:
A1 ->
B)
res1 (
C2:
A2 ->
B)
res2,
reducts_incl C2 res2 (
incontext2 C1 res1 C2 res2).
Proof.
Local Hint Resolve reducts_incl_val reducts_incl_loc reducts_incl_listval
reducts_incl_incontext reducts_incl_incontext2_left
reducts_incl_incontext2_right :
core.
Lemma step_expr_context:
forall from to C,
context from to C ->
forall a m,
reducts_incl C (
step_expr from a m) (
step_expr to (
C a)
m)
with step_exprlist_context:
forall from C,
contextlist from C ->
forall a m,
reducts_incl C (
step_expr from a m) (
step_exprlist (
C a)
m).
Proof.
Completeness part 2: if we can reduce to Stuckstate, step_expr
contains at least one Stuckred reduction.
Lemma not_stuckred_imm_safe:
forall m a k,
(
forall C, ~
In (
C,
Stuckred) (
step_expr k a m)) ->
imm_safe_t k a m.
Proof.
intros.
generalize (
step_expr_sound a k m).
intros [
A B].
destruct (
step_expr k a m)
as [|[
C rd]
res]
eqn:?.
specialize (
B (
eq_refl _)).
destruct k.
destruct a;
simpl in B;
try congruence.
constructor.
destruct a;
simpl in B;
try congruence.
constructor.
assert (
NOTSTUCK:
rd <>
Stuckred).
red;
intros.
elim (
H C);
subst rd;
auto with coqlib.
exploit A.
eauto with coqlib.
intros [
a' [
k' [
P [
Q R]]]].
destruct k';
destruct rd;
simpl in R;
intuition.
subst a.
eapply imm_safe_t_lred;
eauto.
subst a.
destruct H1 as [
w'
PT].
eapply imm_safe_t_rred;
eauto.
subst.
eapply imm_safe_t_callred;
eauto.
Qed.
Lemma not_imm_safe_stuck_red:
forall m a k C,
context k RV C ->
~
imm_safe_t k a m ->
exists C',
In (
C',
Stuckred) (
step_expr RV (
C a)
m).
Proof.
Connections between imm_safe_t and imm_safe
Lemma imm_safe_imm_safe_t:
forall k a m,
imm_safe ge e k a m ->
imm_safe_t k a m \/
exists C,
exists a1,
exists t,
exists a1',
exists m',
context RV k C /\
a =
C a1 /\
rred ge a1 m t a1'
m' /\
forall w', ~
possible_trace w t w'.
Proof.
A state can "crash the world" if it can make an observable transition
whose trace is not accepted by the external world.
Definition can_crash_world (
w:
world) (
S:
state) :
Prop :=
exists t,
exists S',
Csem.step ge S t S' /\
forall w', ~
possible_trace w t w'.
Theorem not_imm_safe_t:
forall K C a m f k,
context K RV C ->
~
imm_safe_t K a m ->
Csem.step ge (
ExprState f (
C a)
k e m)
E0 Stuckstate \/
can_crash_world w (
ExprState f (
C a)
k e m).
Proof.
intros.
destruct (
classic (
imm_safe ge e K a m)).
exploit imm_safe_imm_safe_t;
eauto.
intros [
A | [
C1 [
a1 [
t [
a1' [
m' [
A [
B [
D E]]]]]]]]].
contradiction.
right.
red.
exists t;
econstructor;
split;
auto.
left.
rewrite B.
eapply step_rred with (
C :=
fun x =>
C(
C1 x)).
eauto.
eauto.
left.
left.
eapply step_stuck;
eauto.
Qed.
End EXPRS.
Transitions over states.
Fixpoint do_alloc_variables (
e:
env) (
m:
mem) (
l:
list (
ident *
type)) {
struct l} :
env *
mem :=
match l with
|
nil => (
e,
m)
| (
id,
ty) ::
l' =>
let (
m1,
b1) :=
Mem.alloc m 0 (
sizeof ge ty)
in
do_alloc_variables (
PTree.set id (
b1,
ty)
e)
m1 l'
end.
Lemma do_alloc_variables_sound:
forall l e m,
alloc_variables ge e m l (
fst (
do_alloc_variables e m l)) (
snd (
do_alloc_variables e m l)).
Proof.
induction l;
intros;
simpl.
constructor.
destruct a as [
id ty].
destruct (
Mem.alloc m 0 (
sizeof ge ty))
as [
m1 b1]
eqn:?;
simpl.
econstructor;
eauto.
Qed.
Lemma do_alloc_variables_complete:
forall e1 m1 l e2 m2,
alloc_variables ge e1 m1 l e2 m2 ->
do_alloc_variables e1 m1 l = (
e2,
m2).
Proof.
induction 1; simpl.
auto.
rewrite H; rewrite IHalloc_variables; auto.
Qed.
Function sem_bind_parameters (
w:
world) (
e:
env) (
m:
mem) (
l:
list (
ident *
type)) (
lv:
list val)
{
struct l} :
option mem :=
match l,
lv with
|
nil,
nil =>
Some m
| (
id,
ty) ::
params,
v1::
lv =>
match PTree.get id e with
|
Some (
b,
ty') =>
check (
type_eq ty ty');
do w',
t,
m1,
v' <-
do_assign_loc w ty m b Ptrofs.zero Full v1;
match t with nil =>
sem_bind_parameters w e m1 params lv |
_ =>
None end
|
None =>
None
end
|
_,
_ =>
None
end.
Lemma sem_bind_parameters_sound :
forall w e m l lv m',
sem_bind_parameters w e m l lv =
Some m' ->
bind_parameters ge e m l lv m'.
Proof.
Lemma sem_bind_parameters_complete :
forall w e m l lv m',
bind_parameters ge e m l lv m' ->
sem_bind_parameters w e m l lv =
Some m'.
Proof.
Inductive transition :
Type :=
TR (
rule:
string) (
t:
trace) (
S':
state).
Definition expr_final_state (
f:
function) (
k:
cont) (
e:
env) (
C_rd: (
expr ->
expr) *
reduction) :=
match snd C_rd with
|
Lred rule a m =>
TR rule E0 (
ExprState f (
fst C_rd a)
k e m)
|
Rred rule a m t =>
TR rule t (
ExprState f (
fst C_rd a)
k e m)
|
Callred rule fd vargs ty m =>
TR rule E0 (
Callstate fd vargs (
Kcall f e (
fst C_rd)
ty k)
m)
|
Stuckred =>
TR "
step_stuck"
E0 Stuckstate
end.
Local Open Scope list_monad_scope.
Definition ret (
rule:
string) (
S:
state) :
list transition :=
TR rule E0 S ::
nil.
Definition do_step (
w:
world) (
s:
state) :
list transition :=
match s with
|
ExprState f a k e m =>
match is_val a with
|
Some(
v,
ty) =>
match k with
|
Kdo k =>
ret "
step_do_2" (
State f Sskip k e m )
|
Kifthenelse s1 s2 k =>
do b <-
bool_val v ty m;
ret "
step_ifthenelse_2" (
State f (
if b then s1 else s2)
k e m)
|
Kwhile1 x s k =>
do b <-
bool_val v ty m;
if b
then ret "
step_while_true" (
State f s (
Kwhile2 x s k)
e m)
else ret "
step_while_false" (
State f Sskip k e m)
|
Kdowhile2 x s k =>
do b <-
bool_val v ty m;
if b
then ret "
step_dowhile_true" (
State f (
Sdowhile x s)
k e m)
else ret "
step_dowhile_false" (
State f Sskip k e m)
|
Kfor2 a2 a3 s k =>
do b <-
bool_val v ty m;
if b
then ret "
step_for_true" (
State f s (
Kfor3 a2 a3 s k)
e m)
else ret "
step_for_false" (
State f Sskip k e m)
|
Kreturn k =>
do v' <-
sem_cast v ty f.(
fn_return)
m;
do m' <-
Mem.free_list m (
blocks_of_env ge e);
ret "
step_return_2" (
Returnstate v' (
call_cont k)
m')
|
Kswitch1 sl k =>
do n <-
sem_switch_arg v ty;
ret "
step_expr_switch" (
State f (
seq_of_labeled_statement (
select_switch n sl)) (
Kswitch2 k)
e m)
|
_ =>
nil
end
|
None =>
map (
expr_final_state f k e) (
step_expr e w RV a m)
end
|
State f (
Sdo x)
k e m =>
ret "
step_do_1" (
ExprState f x (
Kdo k)
e m)
|
State f (
Ssequence s1 s2)
k e m =>
ret "
step_seq" (
State f s1 (
Kseq s2 k)
e m)
|
State f Sskip (
Kseq s k)
e m =>
ret "
step_skip_seq" (
State f s k e m)
|
State f Scontinue (
Kseq s k)
e m =>
ret "
step_continue_seq" (
State f Scontinue k e m)
|
State f Sbreak (
Kseq s k)
e m =>
ret "
step_break_seq" (
State f Sbreak k e m)
|
State f (
Sifthenelse a s1 s2)
k e m =>
ret "
step_ifthenelse_1" (
ExprState f a (
Kifthenelse s1 s2 k)
e m)
|
State f (
Swhile x s)
k e m =>
ret "
step_while" (
ExprState f x (
Kwhile1 x s k)
e m)
|
State f (
Sskip|
Scontinue) (
Kwhile2 x s k)
e m =>
ret "
step_skip_or_continue_while" (
State f (
Swhile x s)
k e m)
|
State f Sbreak (
Kwhile2 x s k)
e m =>
ret "
step_break_while" (
State f Sskip k e m)
|
State f (
Sdowhile a s)
k e m =>
ret "
step_dowhile" (
State f s (
Kdowhile1 a s k)
e m)
|
State f (
Sskip|
Scontinue) (
Kdowhile1 x s k)
e m =>
ret "
step_skip_or_continue_dowhile" (
ExprState f x (
Kdowhile2 x s k)
e m)
|
State f Sbreak (
Kdowhile1 x s k)
e m =>
ret "
step_break_dowhile" (
State f Sskip k e m)
|
State f (
Sfor a1 a2 a3 s)
k e m =>
if is_skip a1
then ret "
step_for" (
ExprState f a2 (
Kfor2 a2 a3 s k)
e m)
else ret "
step_for_start" (
State f a1 (
Kseq (
Sfor Sskip a2 a3 s)
k)
e m)
|
State f (
Sskip|
Scontinue) (
Kfor3 a2 a3 s k)
e m =>
ret "
step_skip_or_continue_for3" (
State f a3 (
Kfor4 a2 a3 s k)
e m)
|
State f Sbreak (
Kfor3 a2 a3 s k)
e m =>
ret "
step_break_for3" (
State f Sskip k e m)
|
State f Sskip (
Kfor4 a2 a3 s k)
e m =>
ret "
step_skip_for4" (
State f (
Sfor Sskip a2 a3 s)
k e m)
|
State f (
Sreturn None)
k e m =>
do m' <-
Mem.free_list m (
blocks_of_env ge e);
ret "
step_return_0" (
Returnstate Vundef (
call_cont k)
m')
|
State f (
Sreturn (
Some x))
k e m =>
ret "
step_return_1" (
ExprState f x (
Kreturn k)
e m)
|
State f Sskip ((
Kstop |
Kcall _ _ _ _ _)
as k)
e m =>
do m' <-
Mem.free_list m (
blocks_of_env ge e);
ret "
step_skip_call" (
Returnstate Vundef k m')
|
State f (
Sswitch x sl)
k e m =>
ret "
step_switch" (
ExprState f x (
Kswitch1 sl k)
e m)
|
State f (
Sskip|
Sbreak) (
Kswitch2 k)
e m =>
ret "
step_skip_break_switch" (
State f Sskip k e m)
|
State f Scontinue (
Kswitch2 k)
e m =>
ret "
step_continue_switch" (
State f Scontinue k e m)
|
State f (
Slabel lbl s)
k e m =>
ret "
step_label" (
State f s k e m)
|
State f (
Sgoto lbl)
k e m =>
match find_label lbl f.(
fn_body) (
call_cont k)
with
|
Some(
s',
k') =>
ret "
step_goto" (
State f s'
k'
e m)
|
None =>
nil
end
|
Callstate (
Internal f)
vargs k m =>
check (
list_norepet_dec ident_eq (
var_names (
fn_params f) ++
var_names (
fn_vars f)));
let (
e,
m1) :=
do_alloc_variables empty_env m (
f.(
fn_params) ++
f.(
fn_vars))
in
do m2 <-
sem_bind_parameters w e m1 f.(
fn_params)
vargs;
ret "
step_internal_function" (
State f f.(
fn_body)
k e m2)
|
Callstate (
External ef _ _ _)
vargs k m =>
match do_external ef w vargs m with
|
None =>
nil
|
Some(
w',
t,
v,
m') =>
TR "
step_external_function"
t (
Returnstate v k m') ::
nil
end
|
Returnstate v (
Kcall f e C ty k)
m =>
ret "
step_returnstate" (
ExprState f (
C (
Eval v ty))
k e m)
|
_ =>
nil
end.
Ltac myinv :=
match goal with
| [ |-
In _ nil ->
_ ] =>
let X :=
fresh "
X"
in intro X;
elim X
| [ |-
In _ (
ret _ _) ->
_ ] =>
let X :=
fresh "
X"
in
intro X;
elim X;
clear X;
[
let EQ :=
fresh "
EQ"
in intro EQ;
unfold ret in EQ;
inv EQ;
myinv |
myinv]
| [ |-
In _ (
_ ::
nil) ->
_ ] =>
let X :=
fresh "
X"
in
intro X;
elim X;
clear X; [
let EQ :=
fresh "
EQ"
in intro EQ;
inv EQ;
myinv |
myinv]
| [ |-
In _ (
match ?
x with Some _ =>
_ |
None =>
_ end) ->
_ ] =>
destruct x eqn:?;
myinv
| [ |-
In _ (
match ?
x with false =>
_ |
true =>
_ end) ->
_ ] =>
destruct x eqn:?;
myinv
| [ |-
In _ (
match ?
x with left _ =>
_ |
right _ =>
_ end) ->
_ ] =>
destruct x;
myinv
|
_ =>
idtac
end.
Local Hint Extern 3 =>
exact I :
core.
Theorem do_step_sound:
forall w S rule t S',
In (
TR rule t S') (
do_step w S) ->
Csem.step ge S t S' \/ (
t =
E0 /\
S' =
Stuckstate /\
can_crash_world w S).
Proof with
Remark estep_not_val:
forall f a k e m t S,
estep ge (
ExprState f a k e m)
t S ->
is_val a =
None.
Proof.
intros.
assert (
forall b from to C,
context from to C -> (
from =
to /\
C =
fun x =>
x) \/
is_val (
C b) =
None).
induction 1;
simpl;
auto.
inv H.
destruct (
H0 a0 _ _ _ H9)
as [[
A B] |
A].
subst.
inv H8;
auto.
auto.
destruct (
H0 a0 _ _ _ H9)
as [[
A B] |
A].
subst.
inv H8;
auto.
auto.
destruct (
H0 a0 _ _ _ H9)
as [[
A B] |
A].
subst.
inv H8;
auto.
auto.
destruct (
H0 a0 _ _ _ H8)
as [[
A B] |
A].
subst.
destruct a0;
auto.
elim H9.
constructor.
auto.
Qed.
Theorem do_step_complete:
forall w S t S'
w',
possible_trace w t w' ->
Csem.step ge S t S' ->
exists rule,
In (
TR rule t S') (
do_step w S).
Proof with
End EXEC.
Local Open Scope option_monad_scope.
Definition do_initial_state (
p:
program):
option (
genv *
state) :=
let ge :=
globalenv p in
do m0 <-
Genv.init_mem p;
do b <-
Genv.find_symbol ge p.(
prog_main);
do f <-
Genv.find_funct_ptr ge b;
check (
type_eq (
type_of_fundef f) (
Tfunction Tnil type_int32s cc_default));
Some (
ge,
Callstate f nil Kstop m0).
Definition at_final_state (
S:
state):
option int :=
match S with
|
Returnstate (
Vint r)
Kstop m =>
Some r
|
_ =>
None
end.