Library RTL
The RTL intermediate language: abstract syntax and semantics.
RTL stands for "Register Transfer Language". This is the first
intermediate language after Cminor and CminorSel.
Require Import Coqlib.
Require Import Maps.
Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Events.
Require Import Memory.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Op.
Require Import Registers.
Abstract syntax
Definition node := positive.
Inductive instruction: Type :=
| Inop: node -> instruction
No operation -- just branch to the successor.
| Iop: operation -> list reg -> reg -> node -> instruction
Iop op args dest succ performs the arithmetic operation op
over the values of registers args, stores the result in dest,
and branches to succ.
| Iload: memory_chunk -> addressing -> list reg -> reg -> node -> instruction
Iload chunk addr args dest succ loads a chunk quantity from
the address determined by the addressing mode addr and the
values of the args registers, stores the quantity just read
into dest, and branches to succ.
| Istore: memory_chunk -> addressing -> list reg -> reg -> node -> instruction
Istore chunk addr args src succ stores the value of register
src in the chunk quantity at the
the address determined by the addressing mode addr and the
values of the args registers, then branches to succ.
| Icall: signature -> reg + ident -> list reg -> reg -> node -> instruction
Icall sig fn args dest succ invokes the function determined by
fn (either a function pointer found in a register or a
function name), giving it the values of registers args
as arguments. It stores the return value in dest and branches
to succ.
| Itailcall: signature -> reg + ident -> list reg -> instruction
Itailcall sig fn args performs a function invocation
in tail-call position.
| Ibuiltin: pseudo_external_function -> list reg -> reg -> node -> instruction
Ibuiltin ef args dest succ calls the built-in function
identified by ef, giving it the values of args as arguments.
It stores the return value in dest and branches to succ.
| Icond: condition -> list reg -> node -> node -> instruction
Icond cond args ifso ifnot evaluates the boolean condition
cond over the values of registers args. If the condition
is true, it transitions to ifso. If the condition is false,
it transitions to ifnot.
| Ijumptable: reg -> list node -> instruction
Ijumptable arg tbl transitions to the node that is the n-th
element of the list tbl, where n is the signed integer
value of register arg.
| Ireturn: option reg -> instruction.
Ireturn terminates the execution of the current function
(it has no successor). It returns the value of the given
register, or Vundef if none is given.
Definition code: Type := PTree.t instruction.
Record function: Type := mkfunction {
fn_sig: signature;
fn_params: list reg;
fn_stacksize: Z;
fn_code: code;
fn_entrypoint: node
}.
A function description comprises a control-flow graph (CFG) fn_code
(a partial finite mapping from nodes to instructions). As in Cminor,
fn_sig is the function signature and fn_stacksize the number of bytes
for its stack-allocated activation record. fn_params is the list
of registers that are bound to the values of arguments at call time.
fn_entrypoint is the node of the first instruction of the function
in the CFG. fn_code_wf asserts that all instructions of the function
have nodes no greater than fn_nextpc.
Definition fundef := AST.fundef external_function function.
Definition funsig (fd: fundef) :=
match fd with
| Internal f => fn_sig f
| External ef => ef_sig ef
end.
Definition pfundef := AST.fundef pseudo_external_function function.
Definition pfunsig (fd: pfundef) :=
match fd with
| Internal f => fn_sig f
| External ef => pef_sig ef
end.
Definition genv := Genv.t (ident * signature) fundef unit.
Definition regset := Regmap.t val.
Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset :=
match rl, vl with
| r1 :: rs, v1 :: vs => Regmap.set r1 v1 (init_regs vs rs)
| _, _ => Regmap.init Vundef
end.
The dynamic semantics of RTL is given in small-step style, as a
set of transitions between states. A state captures the current
point in the execution. Three kinds of states appear in the transitions:
- State cs f sp pc rs m describes an execution point within a function. f is the current function. sp is the pointer to the stack block for its current activation (as in Cminor). pc is the current program point (CFG node) within the code c. rs gives the current values for the pseudo-registers. m is the current memory state.
- Callstate cs f args m is an intermediate state that appears during function calls. f is the function definition that we are calling. args (a list of values) are the arguments for this call. m is the current memory state.
- Returnstate cs v m is an intermediate state that appears when a function terminates and returns to its caller. v is the return value and m the current memory state.
Inductive stackframe : Type :=
| Stackframe:
forall (res: reg)
(f: function)
(sp: val)
(pc: node)
(rs: regset),
stackframe.
Inductive state : Type :=
| State:
forall (stack: list stackframe)
(f: function)
(sp: val)
(pc: node)
(rs: regset)
(m: mem),
state
| Callstate:
forall (stack: list stackframe)
(f: fundef)
(args: list val)
(m: mem),
state
| Returnstate:
forall (stack: list stackframe)
(v: val)
(m: mem),
state.
Section RELSEM.
Variable ge: genv.
Definition find_function
(ros: reg + ident) (rs: regset) : option (ident * signature) :=
match ros with
| inl r => Genv.find_funct ge rs#r
| inr symb =>
match Genv.find_symbol ge symb with
| None => None
| Some b => Genv.find_funct_ptr ge b
end
end.
The transitions are presented as an inductive predicate
step ge st1 t st2, where ge is the global environment,
st1 the initial state, st2 the final state, and t the trace
of system calls performed during this transition.
Inductive step: state -> trace -> state -> Prop :=
| exec_Inop:
forall s f sp pc rs m pc',
(fn_code f)!pc = Some(Inop pc') ->
step (State s f sp pc rs m)
E0 (State s f sp pc' rs m)
| exec_Iop:
forall s f sp pc rs m op args res pc' v,
(fn_code f)!pc = Some(Iop op args res pc') ->
eval_operation ge sp op rs##args m = Some v ->
step (State s f sp pc rs m)
E0 (State s f sp pc' (rs#res <- v) m)
| exec_Iload:
forall s f sp pc rs m chunk addr args dst pc' a v,
(fn_code f)!pc = Some(Iload chunk addr args dst pc') ->
eval_addressing ge sp addr rs##args = Some a ->
Mem.loadv chunk m a = Some v ->
step (State s f sp pc rs m)
E0 (State s f sp pc' (rs#dst <- v) m)
| exec_Istore:
forall s f sp pc rs m chunk addr args src pc' a m',
(fn_code f)!pc = Some(Istore chunk addr args src pc') ->
eval_addressing ge sp addr rs##args = Some a ->
Mem.storev chunk m a rs#src = Some m' ->
step (State s f sp pc rs m)
E0 (State s f sp pc' rs m')
| exec_Icall:
forall s f sp pc rs m sig ros args res pc' fd,
(fn_code f)!pc = Some(Icall sig ros args res pc') ->
find_function ros rs = Some fd ->
funsig (Genv.genv_funs_defs ge fd) = sig ->
step (State s f sp pc rs m)
E0 (Callstate (Stackframe res f sp pc' rs :: s) (Genv.genv_funs_defs ge fd) rs##args m)
| exec_Ibuiltin:
forall s f sp pc rs m ef args res pc' t v m',
(fn_code f)!pc = Some(Ibuiltin ef args res pc') ->
pseudo_external_call ef ge rs##args m t v m' ->
step (State s f sp pc rs m)
t (State s f sp pc' (rs#res <- v) m')
| exec_Icond:
forall s f sp pc rs m cond args ifso ifnot b pc',
(fn_code f)!pc = Some(Icond cond args ifso ifnot) ->
eval_condition cond rs##args m = Some b ->
pc' = (if b then ifso else ifnot) ->
step (State s f sp pc rs m)
E0 (State s f sp pc' rs m)
| exec_Ijumptable:
forall s f sp pc rs m arg tbl n pc',
(fn_code f)!pc = Some(Ijumptable arg tbl) ->
rs#arg = Vint n ->
list_nth_z tbl (Int.unsigned n) = Some pc' ->
step (State s f sp pc rs m)
E0 (State s f sp pc' rs m)
| exec_Ireturn:
forall s f stk pc rs m or m',
(fn_code f)!pc = Some(Ireturn or) ->
Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
step (State s f (Vptr stk Int.zero) pc rs m)
E0 (Returnstate s (regmap_optget or Vundef rs) m')
| exec_function_internal:
forall s f args m m' stk,
Mem.alloc m 0 f.(fn_stacksize) = (m', stk) ->
step (Callstate s (Internal f) args m)
E0 (State s
( f)
(Vptr stk Int.zero)
( f).(fn_entrypoint)
(init_regs args ( f).(fn_params))
m')
| exec_function_external:
forall s ef args res t m m',
external_call ef ge args m t res m' ->
step (Callstate s (External ef) args m)
t (Returnstate s res m')
| exec_return:
forall res f sp pc rs s vres m,
step (Returnstate (Stackframe res f sp pc rs :: s) vres m)
E0 (State s f sp pc (rs#res <- vres) m).
Lemma exec_Iop':
forall s f sp pc rs m op args res pc' rs' v,
(fn_code f)!pc = Some(Iop op args res pc') ->
eval_operation ge sp op rs##args m = Some v ->
rs' = (rs#res <- v) ->
step (State s f sp pc rs m)
E0 (State s f sp pc' rs' m).
Proof.
intros. subst rs'. eapply exec_Iop; eauto.
Qed.
Lemma exec_Iload':
forall s f sp pc rs m chunk addr args dst pc' rs' a v,
(fn_code f)!pc = Some(Iload chunk addr args dst pc') ->
eval_addressing ge sp addr rs##args = Some a ->
Mem.loadv chunk m a = Some v ->
rs' = (rs#dst <- v) ->
step (State s f sp pc rs m)
E0 (State s f sp pc' rs' m).
Proof.
intros. subst rs'. eapply exec_Iload; eauto.
Qed.
End RELSEM.
Execution of whole programs are described as sequences of transitions
from an initial state to a final state. An initial state is a Callstate
corresponding to the invocation of the ``main'' function of the program
without arguments and with an empty call stack.
(To be changed, parameterize with the function name and signature)
Inductive initial_state (ge: genv) : ((ident * signature) * (list val * mem)) -> state -> Prop :=
| initial_state_intro: forall f s args m,
funsig (Genv.genv_funs_defs ge (f,s)) = s ->
initial_state ge ((f, s), (args, m)) (Callstate nil (Genv.genv_funs_defs ge (f,s)) args m)
.
A final state is a Returnstate with an empty call stack.
(To be changed, the result can be an arbitrary value, not necessarily an integer)
Inductive final_state: state -> (mem * val) -> Prop :=
| final_state_intro: forall r m,
final_state (Returnstate nil r m) (m,r).
The small-step semantics for a program.
Definition semantics (ge: genv) :=
Semantics step (initial_state ge) final_state ge.
This semantics is receptive to changes in events.
Lemma semantics_single_events : forall ge, single_events (semantics ge).
Proof.
intros.
red; intros; inv H; simpl; try omega.
eapply ec_trace_length; eauto using pseudo_external_call_spec.
eapply external_call_trace_length; eauto.
Qed.
Lemma semantics_receptive:
forall ge, receptive (semantics ge).
Proof.
intros. constructor; simpl; intros; eauto using semantics_single_events.
assert (t1 = E0 -> exists s2, step ge s t2 s2).
intros. subst. inv H0. exists s1; auto.
inversion H; subst; auto.
generalize (pseudo_external_call_spec ef). intro.
exploit ec_receptive; eauto. intros [vres2 [m2 EC2]].
exists (State s0 f sp pc' (rs#res <- vres2) m2). eapply exec_Ibuiltin; eauto.
exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
exists (Returnstate s0 vres2 m2). econstructor; eauto.
Qed.
Lemma semantics_determinate:
forall ge, determinate (semantics ge).
Proof.
intros. constructor; try (intros; eauto using semantics_single_events; fail).
inversion 1; subst; inversion 1; subst; try (exfalso; congruence );
(split ; [try (constructor; fail) | try (intros; congruence; fail) ] ).
rewrite H11 in H0. injection H0; intros; subst.
destruct (ec_determ (pseudo_external_call_spec _) _ _ _ _ _ _ _ _ _ H1 H12).
assumption.
rewrite H11 in H0. injection H0; intros; subst.
destruct (ec_determ (pseudo_external_call_spec _) _ _ _ _ _ _ _ _ _ H1 H12).
destruct (H4 (refl_equal _)).
intros; subst; reflexivity.
rewrite H11 in H0; injection H0; intros; subst.
rewrite H12 in H1; injection H1; intros; subst.
reflexivity.
destruct (external_call_determ ef _ _ _ _ _ _ _ _ _ H0 H8).
assumption.
destruct (external_call_determ ef _ _ _ _ _ _ _ _ _ H0 H8).
intros; subst.
destruct (H3 (refl_equal _)). subst. reflexivity.
inversion 1; inversion 1; reflexivity.
inversion 1; subst.
intro.
intro.
intro.
inversion H0.
inversion 1; subst; inversion 1; subst; reflexivity.
Qed.
Semantics of a module.
The global environment, defining symbols and their memory blocks, is not part of a module. A module only contains the definitions of its functions. It is a finite set of such functions including pseudo-external functions, but not external functions.
For now let's say that there are finitely many functions in the module.
Definition module := list (ident * pfundef).
Definition idsig_eq_dec : forall i1 i2: ident * signature, {i1 = i2} + {i1 <> i2}.
Proof.
decide equality.
decide equality.
apply opt_typ_eq.
decide equality.
apply typ_eq.
apply peq.
Qed.
Fixpoint module_pseudofuns (m: module) (idsig: ident * signature) : option (AST.fundef pseudo_external_function function) :=
match m with
| nil => None
| (id', Internal f) :: q =>
if idsig_eq_dec (id', fn_sig f) idsig then Some (Internal f) else module_pseudofuns q idsig
| (id', External pef) :: q =>
if idsig_eq_dec (id', pef_sig pef) idsig then Some (External pef) else module_pseudofuns q idsig
end.
Definition module_funs (m: module) (idsig: ident * signature): fundef :=
match module_pseudofuns m idsig with
| Some (Internal f) => Internal f
| Some (External pef) => External (EF_pseudo pef)
| None => match idsig with (id, sig) => External (EF_true_external id sig) end
end.
Lemma module_pseudofuns_sig: forall m id sig f,
module_pseudofuns m (id, sig) = Some f ->
pfunsig f = sig.
Proof.
induction m; simpl.
discriminate.
intros.
destruct a.
destruct p.
destruct (idsig_eq_dec (i, fn_sig f0) (id, sig)); eauto.
injection H; intros; subst; unfold pfunsig; congruence.
destruct (idsig_eq_dec (i, pef_sig p) (id, sig)); eauto.
injection H; intros; subst; unfold pfunsig; congruence.
Qed.
Lemma module_funs_sig : forall m id sig,
funsig (module_funs m (id, sig)) = sig.
Proof.
unfold module_funs.
intros.
case_eq (module_pseudofuns m (id, sig)).
intros.
generalize (module_pseudofuns_sig _ _ _ _ H).
destruct f; tauto.
reflexivity.
Qed.
Definition module_genv (m: module) (FN: Type) (ge: Genv.t (ident * signature) FN unit) : genv :=
Genv.replace_fundefs (module_funs m) ge.
Definition module_sem (m: module) (FN: Type) (ge: Genv.t (ident * signature) FN unit) :=
semantics (module_genv m _ ge).
Operations on RTL abstract syntax
Definition successors_instr (i: instruction) : list node :=
match i with
| Inop s => s :: nil
| Iop op args res s => s :: nil
| Iload chunk addr args dst s => s :: nil
| Istore chunk addr args src s => s :: nil
| Icall sig ros args res s => s :: nil
| Itailcall sig ros args => nil
| Ibuiltin ef args res s => s :: nil
| Icond cond args ifso ifnot => ifso :: ifnot :: nil
| Ijumptable arg tbl => tbl
| Ireturn optarg => nil
end.
Definition successors (f: function) : PTree.t (list node) :=
PTree.map1 successors_instr f.(fn_code).
The registers used by an instruction
Definition instr_uses (i: instruction) : list reg :=
match i with
| Inop s => nil
| Iop op args res s => args
| Iload chunk addr args dst s => args
| Istore chunk addr args src s => src :: args
| Icall sig (inl r) args res s => r :: args
| Icall sig (inr id) args res s => args
| Itailcall sig (inl r) args => r :: args
| Itailcall sig (inr id) args => args
| Ibuiltin ef args res s => args
| Icond cond args ifso ifnot => args
| Ijumptable arg tbl => arg :: nil
| Ireturn None => nil
| Ireturn (Some arg) => arg :: nil
end.
The register defined by an instruction, if any
Definition instr_defs (i: instruction) : option reg :=
match i with
| Inop s => None
| Iop op args res s => Some res
| Iload chunk addr args dst s => Some dst
| Istore chunk addr args src s => None
| Icall sig ros args res s => Some res
| Itailcall sig ros args => None
| Ibuiltin ef args res s => Some res
| Icond cond args ifso ifnot => None
| Ijumptable arg tbl => None
| Ireturn optarg => None
end.
Transformation of a RTL function instruction by instruction.
This applies a given transformation function to all instructions
of a function and constructs a transformed function from that.
Section TRANSF.
Variable transf: node -> instruction -> instruction.
Definition transf_function (f: function) : function :=
mkfunction
f.(fn_sig)
f.(fn_params)
f.(fn_stacksize)
(PTree.map transf f.(fn_code))
f.(fn_entrypoint).
End TRANSF.
Require Import Errors.
Section TRANSF_PARTIAL.
Variable transf: function -> res function.
Hypothesis transf_sig: forall f f', transf f = OK f' ->
fn_sig f = fn_sig f'.
Lemma transf_partial_fundef_sig : forall f f',
AST.transf_partial_fundef transf f = OK f' ->
pfunsig f = pfunsig f'.
Proof.
unfold AST.transf_partial_fundef.
destruct f.
unfold bind.
intro.
case_eq (transf f); try discriminate.
injection 2; intros; subst.
eapply transf_sig.
assumption.
congruence.
Qed.
Fixpoint transf_module (m: module): res module :=
match m with
| nil => OK nil
| (i, f) :: m' => match AST.transf_partial_fundef transf f with
| OK f' => match transf_module m' with
| OK m'' => OK ((i, f') :: m'')
| e => e
end
| Error e => Error e
end
end.
Lemma transf_module_pseudofuns_some : forall m m',
transf_module m = OK m' ->
forall idsig f, module_pseudofuns m idsig = Some f ->
exists f', AST.transf_partial_fundef transf f = OK f' /\ module_pseudofuns m' idsig = Some f'.
Proof.
induction m; simpl.
injection 1; intro; subst. discriminate.
intro.
destruct a.
case_eq (transf_partial_fundef transf p); try discriminate.
intros until 1.
case_eq (transf_module m); try discriminate.
injection 2; intro; subst.
simpl.
intros until f0.
destruct p.
destruct (idsig_eq_dec (i, fn_sig f1) idsig).
injection 1; intro; subst.
esplit.
split.
eassumption.
revert H.
simpl.
unfold bind.
case_eq (transf f1); try discriminate.
injection 2; intros; subst.
destruct (idsig_eq_dec (i, fn_sig f0) (i, fn_sig f1)).
trivial.
destruct n.
f_equal.
symmetry.
apply transf_sig.
assumption.
intros.
revert H.
simpl.
unfold bind.
case_eq (transf f1); try discriminate.
injection 2; intros; subst.
rewrite <- (transf_sig _ _ H).
destruct (idsig_eq_dec (i, fn_sig f1) idsig).
contradiction.
eauto.
destruct (idsig_eq_dec (i, pef_sig p) idsig).
injection 1; intros; subst.
revert H.
simpl.
injection 1; intros; subst.
destruct (idsig_eq_dec (i, pef_sig p) (i, pef_sig p)).
eauto.
congruence.
intros.
revert H.
simpl.
injection 1; intros; subst.
destruct (idsig_eq_dec (i, pef_sig p) idsig).
contradiction.
eauto.
Qed.
Lemma transf_module_pseudofuns_none : forall m m',
transf_module m = OK m' ->
forall idsig, module_pseudofuns m idsig = None ->
module_pseudofuns m' idsig = None.
Proof.
induction m; simpl.
injection 1; intros; subst; reflexivity.
destruct a.
intro.
case_eq (transf_partial_fundef transf p); try discriminate.
case_eq (transf_module m); try discriminate.
injection 3; intro; subst.
intro.
revert H0.
simpl.
destruct p; simpl; unfold bind.
case_eq (transf f0); try discriminate.
injection 2; intro; subst.
rewrite (transf_sig _ _ H0).
destruct (idsig_eq_dec (i, fn_sig f1) idsig).
discriminate.
eauto.
injection 1; intro; subst.
destruct (idsig_eq_dec (i, pef_sig p) idsig).
tauto.
eauto.
Qed.
Lemma transf_module_funs : forall m m',
transf_module m = OK m' ->
forall idsig,
AST.transf_partial_fundef transf (module_funs m idsig) = OK (module_funs m' idsig).
Proof.
intros.
unfold module_funs.
case_eq (module_pseudofuns m idsig).
intros.
destruct (transf_module_pseudofuns_some _ _ H _ _ H0).
destruct H1.
revert H1.
destruct f.
simpl.
unfold bind.
case_eq (transf f); try discriminate.
injection 2; intros until 1; subst.
rewrite H2.
trivial.
simpl.
injection 1; intro; subst.
rewrite H2.
trivial.
destruct idsig.
simpl.
intro.
rewrite (transf_module_pseudofuns_none _ _ H _ H0).
reflexivity.
Qed.
End TRANSF_PARTIAL.