Library compcert.ia32.Asm
Abstract syntax and semantics for IA32 assembly language
Require Import Coqlib.
Require Import Maps.
Require Import AST.
Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Memory.
Require Import Events.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Locations.
Require Import Stacklayout.
Require Import Conventions.
Inductive ireg: Type :=
| EAX: ireg | EBX: ireg | ECX: ireg | EDX: ireg
| ESI: ireg | EDI: ireg | EBP: ireg | ESP: ireg.
Floating-point registers, i.e. SSE2 registers
Inductive freg: Type :=
| XMM0: freg | XMM1: freg | XMM2: freg | XMM3: freg
| XMM4: freg | XMM5: freg | XMM6: freg | XMM7: freg.
Lemma ireg_eq: forall (x y: ireg), {x=y} + {x<>y}.
Proof. decide equality. Defined.
Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}.
Proof. decide equality. Defined.
Bits of the flags register.
All registers modeled here.
Inductive preg: Type :=
| PC: preg
| IR: ireg -> preg
| FR: freg -> preg
| ST0: preg
| CR: crbit -> preg
| RA: preg.
Coercion IR: ireg >-> preg.
Coercion FR: freg >-> preg.
Coercion CR: crbit >-> preg.
Conventional names for stack pointer (SP) and return address (RA)
General form of an addressing mode.
Inductive addrmode: Type :=
| Addrmode (base: option ireg)
(ofs: option (ireg * int))
(const: int + ident * int).
Testable conditions (for conditional jumps and more).
Inductive testcond: Type :=
| Cond_e | Cond_ne
| Cond_b | Cond_be | Cond_ae | Cond_a
| Cond_l | Cond_le | Cond_ge | Cond_g
| Cond_p | Cond_np.
Instructions. IA32 instructions accept many combinations of
registers, memory references and immediate constants as arguments.
Here, we list only the combinations that we actually use.
Naming conventions:
- r: integer register operand
- f: XMM register operand
- m: memory operand
- i: immediate integer operand
- s: immediate symbol operand
- l: immediate label operand
- cl: the CL register
Moves
| Pmov_rr (rd: ireg) (r1: ireg)
| Pmov_ri (rd: ireg) (n: int)
| Pmov_ra (rd: ireg) (id: ident)
| Pmov_rm (rd: ireg) (a: addrmode)
| Pmov_mr (a: addrmode) (rs: ireg)
| Pmovsd_ff (rd: freg) (r1: freg)
| Pmovsd_fi (rd: freg) (n: float)
| Pmovsd_fm (rd: freg) (a: addrmode)
| Pmovsd_mf (a: addrmode) (r1: freg)
| Pfld_f (r1: freg)
| Pfld_m (a: addrmode)
| Pfstp_f (rd: freg)
| Pfstp_m (a: addrmode)
| Pxchg_rr (r1: ireg) (r2: ireg)
| Pmov_ri (rd: ireg) (n: int)
| Pmov_ra (rd: ireg) (id: ident)
| Pmov_rm (rd: ireg) (a: addrmode)
| Pmov_mr (a: addrmode) (rs: ireg)
| Pmovsd_ff (rd: freg) (r1: freg)
| Pmovsd_fi (rd: freg) (n: float)
| Pmovsd_fm (rd: freg) (a: addrmode)
| Pmovsd_mf (a: addrmode) (r1: freg)
| Pfld_f (r1: freg)
| Pfld_m (a: addrmode)
| Pfstp_f (rd: freg)
| Pfstp_m (a: addrmode)
| Pxchg_rr (r1: ireg) (r2: ireg)
Moves with conversion
| Pmovb_mr (a: addrmode) (rs: ireg)
| Pmovw_mr (a: addrmode) (rs: ireg)
| Pmovzb_rr (rd: ireg) (rs: ireg)
| Pmovzb_rm (rd: ireg) (a: addrmode)
| Pmovsb_rr (rd: ireg) (rs: ireg)
| Pmovsb_rm (rd: ireg) (a: addrmode)
| Pmovzw_rr (rd: ireg) (rs: ireg)
| Pmovzw_rm (rd: ireg) (a: addrmode)
| Pmovsw_rr (rd: ireg) (rs: ireg)
| Pmovsw_rm (rd: ireg) (a: addrmode)
| Pcvtss2sd_fm (rd: freg) (a: addrmode)
| Pcvtsd2ss_ff (rd: freg) (r1: freg)
| Pcvtsd2ss_mf (a: addrmode) (r1: freg)
| Pcvttsd2si_rf (rd: ireg) (r1: freg)
| Pcvtsi2sd_fr (rd: freg) (r1: ireg)
| Pmovw_mr (a: addrmode) (rs: ireg)
| Pmovzb_rr (rd: ireg) (rs: ireg)
| Pmovzb_rm (rd: ireg) (a: addrmode)
| Pmovsb_rr (rd: ireg) (rs: ireg)
| Pmovsb_rm (rd: ireg) (a: addrmode)
| Pmovzw_rr (rd: ireg) (rs: ireg)
| Pmovzw_rm (rd: ireg) (a: addrmode)
| Pmovsw_rr (rd: ireg) (rs: ireg)
| Pmovsw_rm (rd: ireg) (a: addrmode)
| Pcvtss2sd_fm (rd: freg) (a: addrmode)
| Pcvtsd2ss_ff (rd: freg) (r1: freg)
| Pcvtsd2ss_mf (a: addrmode) (r1: freg)
| Pcvttsd2si_rf (rd: ireg) (r1: freg)
| Pcvtsi2sd_fr (rd: freg) (r1: ireg)
Integer arithmetic
| Plea (rd: ireg) (a: addrmode)
| Pneg (rd: ireg)
| Psub_rr (rd: ireg) (r1: ireg)
| Pimul_rr (rd: ireg) (r1: ireg)
| Pimul_ri (rd: ireg) (n: int)
| Pimul_r (r1: ireg)
| Pmul_r (r1: ireg)
| Pdiv (r1: ireg)
| Pidiv (r1: ireg)
| Pand_rr (rd: ireg) (r1: ireg)
| Pand_ri (rd: ireg) (n: int)
| Por_rr (rd: ireg) (r1: ireg)
| Por_ri (rd: ireg) (n: int)
| Pxor_r (rd: ireg)
| Pxor_rr (rd: ireg) (r1: ireg)
| Pxor_ri (rd: ireg) (n: int)
| Pnot (rd: ireg)
| Psal_rcl (rd: ireg)
| Psal_ri (rd: ireg) (n: int)
| Pshr_rcl (rd: ireg)
| Pshr_ri (rd: ireg) (n: int)
| Psar_rcl (rd: ireg)
| Psar_ri (rd: ireg) (n: int)
| Pshld_ri (rd: ireg) (r1: ireg) (n: int)
| Pror_ri (rd: ireg) (n: int)
| Pcmp_rr (r1 r2: ireg)
| Pcmp_ri (r1: ireg) (n: int)
| Ptest_rr (r1 r2: ireg)
| Ptest_ri (r1: ireg) (n: int)
| Pcmov (c: testcond) (rd: ireg) (r1: ireg)
| Psetcc (c: testcond) (rd: ireg)
| Pneg (rd: ireg)
| Psub_rr (rd: ireg) (r1: ireg)
| Pimul_rr (rd: ireg) (r1: ireg)
| Pimul_ri (rd: ireg) (n: int)
| Pimul_r (r1: ireg)
| Pmul_r (r1: ireg)
| Pdiv (r1: ireg)
| Pidiv (r1: ireg)
| Pand_rr (rd: ireg) (r1: ireg)
| Pand_ri (rd: ireg) (n: int)
| Por_rr (rd: ireg) (r1: ireg)
| Por_ri (rd: ireg) (n: int)
| Pxor_r (rd: ireg)
| Pxor_rr (rd: ireg) (r1: ireg)
| Pxor_ri (rd: ireg) (n: int)
| Pnot (rd: ireg)
| Psal_rcl (rd: ireg)
| Psal_ri (rd: ireg) (n: int)
| Pshr_rcl (rd: ireg)
| Pshr_ri (rd: ireg) (n: int)
| Psar_rcl (rd: ireg)
| Psar_ri (rd: ireg) (n: int)
| Pshld_ri (rd: ireg) (r1: ireg) (n: int)
| Pror_ri (rd: ireg) (n: int)
| Pcmp_rr (r1 r2: ireg)
| Pcmp_ri (r1: ireg) (n: int)
| Ptest_rr (r1 r2: ireg)
| Ptest_ri (r1: ireg) (n: int)
| Pcmov (c: testcond) (rd: ireg) (r1: ireg)
| Psetcc (c: testcond) (rd: ireg)
Floating-point arithmetic
| Paddd_ff (rd: freg) (r1: freg)
| Psubd_ff (rd: freg) (r1: freg)
| Pmuld_ff (rd: freg) (r1: freg)
| Pdivd_ff (rd: freg) (r1: freg)
| Pnegd (rd: freg)
| Pabsd (rd: freg)
| Pcomisd_ff (r1 r2: freg)
| Pxorpd_f (rd: freg)
| Psubd_ff (rd: freg) (r1: freg)
| Pmuld_ff (rd: freg) (r1: freg)
| Pdivd_ff (rd: freg) (r1: freg)
| Pnegd (rd: freg)
| Pabsd (rd: freg)
| Pcomisd_ff (r1 r2: freg)
| Pxorpd_f (rd: freg)
Branches and calls
| Pjmp_l (l: label)
| Pjmp_s (symb: ident) (sg: signature)
| Pjmp_r (r: ireg) (sg: signature)
| Pjcc (c: testcond)(l: label)
| Pjcc2 (c1 c2: testcond)(l: label)
| Pjmptbl (r: ireg) (tbl: list label)
| Pcall_s (symb: ident) (sg: signature)
| Pcall_r (r: ireg) (sg: signature)
| Pret
| Pjmp_s (symb: ident) (sg: signature)
| Pjmp_r (r: ireg) (sg: signature)
| Pjcc (c: testcond)(l: label)
| Pjcc2 (c1 c2: testcond)(l: label)
| Pjmptbl (r: ireg) (tbl: list label)
| Pcall_s (symb: ident) (sg: signature)
| Pcall_r (r: ireg) (sg: signature)
| Pret
Pseudo-instructions
| Plabel(l: label)
| Pallocframe(sz: Z)(ofs_ra ofs_link: int)
| Pfreeframe(sz: Z)(ofs_ra ofs_link: int)
| Pbuiltin(ef: external_function)(args: list preg)(res: list preg)
| Pannot(ef: external_function)(args: list annot_param)
with annot_param : Type :=
| APreg: preg -> annot_param
| APstack: memory_chunk -> Z -> annot_param.
Definition code := list instruction.
Record function : Type := mkfunction { fn_sig: signature; fn_code: code }.
Definition fundef := AST.fundef function.
Definition program := AST.program fundef unit.
| Pallocframe(sz: Z)(ofs_ra ofs_link: int)
| Pfreeframe(sz: Z)(ofs_ra ofs_link: int)
| Pbuiltin(ef: external_function)(args: list preg)(res: list preg)
| Pannot(ef: external_function)(args: list annot_param)
with annot_param : Type :=
| APreg: preg -> annot_param
| APstack: memory_chunk -> Z -> annot_param.
Definition code := list instruction.
Record function : Type := mkfunction { fn_sig: signature; fn_code: code }.
Definition fundef := AST.fundef function.
Definition program := AST.program fundef unit.
Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}.
Proof. decide equality. apply ireg_eq. apply freg_eq. decide equality. Defined.
Module PregEq.
Definition t := preg.
Definition eq := preg_eq.
End PregEq.
Module Pregmap := EMap(PregEq).
Definition regset := Pregmap.t val.
Definition genv := Genv.t fundef unit.
Notation "a # b" := (a b) (at level 1, only parsing).
Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level).
Undefining some registers
Fixpoint undef_regs (l: list preg) (rs: regset) : regset :=
match l with
| nil => rs
| r :: l´ => undef_regs l´ (rs#r <- Vundef)
end.
Assigning multiple registers
Fixpoint set_regs (rl: list preg) (vl: list val) (rs: regset) : regset :=
match rl, vl with
| r1 :: rl´, v1 :: vl´ => set_regs rl´ vl´ (rs#r1 <- v1)
| _, _ => rs
end.
CompCertX:test-compcert-param-memory We create section WITHMEM and associated
contexts to parameterize the proof over the memory model. CompCertX:test-compcert-param-extcall Actually, we also need to parameterize
over external functions. To this end, we created a CompilerConfiguration class
(cf. Events) which is designed to be the single class on which the whole CompCert is to be
parameterized. It includes all operations and properties on which CompCert depends:
memory model, semantics of external functions and their preservation through
compilation.
Section WITHCONFIG.
Context `{compiler_config: CompilerConfiguration}.
Section RELSEM.
Class FindLabels {function instructionx}
(is_label : label -> instructionx -> bool)
(fn_code : function -> list instructionx).
Context `{compiler_config: CompilerConfiguration}.
Section RELSEM.
Class FindLabels {function instructionx}
(is_label : label -> instructionx -> bool)
(fn_code : function -> list instructionx).
Looking up instructions in a code sequence by position.
Fixpoint find_instr `{Hfl: FindLabels} (pos: Z) (c: list instructionx) {struct c} : option instructionx :=
match c with
| nil => None
| i :: il => if zeq pos 0 then Some i else find_instr (pos - 1) il
end.
Position corresponding to a label
Definition is_label (lbl: label) (instr: instruction) : bool :=
match instr with
| Plabel lbl´ => if peq lbl lbl´ then true else false
| _ => false
end.
Global Instance: FindLabels is_label fn_code.
Lemma is_label_correct:
forall lbl instr,
if is_label lbl instr then instr = Plabel lbl else instr <> Plabel lbl.
Proof.
intros. destruct instr; simpl; try discriminate.
case (peq lbl l); intro; congruence.
Qed.
Section WITH_FIND_LABELS.
Context {function instructionx is_label fn_code}
`{Hfl: FindLabels function instructionx is_label fn_code}.
Fixpoint label_pos `{Hfl: FindLabels function instructionx is_label fn_code}
(lbl: label) (pos: Z) (c: list instructionx) {struct c} : option Z :=
match c with
| nil => None
| instr :: c´ =>
if is_label lbl instr then Some (pos + 1) else label_pos lbl (pos + 1) c´
end.
Section WITHGE.
Context {F V: Type}.
Variable ge: Genv.t F V.
Definition symbol_offset (id: ident) (ofs: int) : val :=
match Genv.find_symbol ge id with
| Some b => Vptr b ofs
| None => Vundef
end.
Evaluating an addressing mode
Definition eval_addrmode (a: addrmode) (rs: regset) : val :=
match a with Addrmode base ofs const =>
Val.add (match base with
| None => Vzero
| Some r => rs r
end)
(Val.add (match ofs with
| None => Vzero
| Some(r, sc) =>
if Int.eq sc Int.one then rs r else Val.mul (rs r) (Vint sc)
end)
(match const with
| inl ofs => Vint ofs
| inr(id, ofs) => symbol_offset id ofs
end))
end.
End WITHGE.
Performing a comparison
Integer comparison between x and y:
- ZF = 1 if x = y, 0 if x != y
- CF = 1 if x <u y, 0 if x >=u y
- SF = 1 if x - y is negative, 0 if x - y is positive
- OF = 1 if x - y overflows (signed), 0 if not
- PF is undefined
Definition compare_ints (x y: val) (rs: regset) (m: mem): regset :=
rs #ZF <- (Val.cmpu (Mem.valid_pointer m) Ceq x y)
#CF <- (Val.cmpu (Mem.valid_pointer m) Clt x y)
#SF <- (Val.negative (Val.sub x y))
#OF <- (Val.sub_overflow x y)
#PF <- Vundef.
Floating-point comparison between x and y:
- ZF = 1 if x=y or unordered, 0 if x<>y
- CF = 1 if x<y or unordered, 0 if x>=y
- PF = 1 if unordered, 0 if ordered.
- SF and 0F are undefined
Definition compare_floats (vx vy: val) (rs: regset) : regset :=
match vx, vy with
| Vfloat x, Vfloat y =>
rs #ZF <- (Val.of_bool (negb (Float.cmp Cne x y)))
#CF <- (Val.of_bool (negb (Float.cmp Cge x y)))
#PF <- (Val.of_bool (negb (Float.cmp Ceq x y || Float.cmp Clt x y || Float.cmp Cgt x y)))
#SF <- Vundef
#OF <- Vundef
| _, _ =>
undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs
end.
Testing a condition
Definition eval_testcond (c: testcond) (rs: regset) : option bool :=
match c with
| Cond_e =>
match rs ZF with
| Vint n => Some (Int.eq n Int.one)
| _ => None
end
| Cond_ne =>
match rs ZF with
| Vint n => Some (Int.eq n Int.zero)
| _ => None
end
| Cond_b =>
match rs CF with
| Vint n => Some (Int.eq n Int.one)
| _ => None
end
| Cond_be =>
match rs CF, rs ZF with
| Vint c, Vint z => Some (Int.eq c Int.one || Int.eq z Int.one)
| _, _ => None
end
| Cond_ae =>
match rs CF with
| Vint n => Some (Int.eq n Int.zero)
| _ => None
end
| Cond_a =>
match rs CF, rs ZF with
| Vint c, Vint z => Some (Int.eq c Int.zero && Int.eq z Int.zero)
| _, _ => None
end
| Cond_l =>
match rs OF, rs SF with
| Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.one)
| _, _ => None
end
| Cond_le =>
match rs OF, rs SF, rs ZF with
| Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.one || Int.eq z Int.one)
| _, _, _ => None
end
| Cond_ge =>
match rs OF, rs SF with
| Vint o, Vint s => Some (Int.eq (Int.xor o s) Int.zero)
| _, _ => None
end
| Cond_g =>
match rs OF, rs SF, rs ZF with
| Vint o, Vint s, Vint z => Some (Int.eq (Int.xor o s) Int.zero && Int.eq z Int.zero)
| _, _, _ => None
end
| Cond_p =>
match rs PF with
| Vint n => Some (Int.eq n Int.one)
| _ => None
end
| Cond_np =>
match rs PF with
| Vint n => Some (Int.eq n Int.zero)
| _ => None
end
end.
The semantics is purely small-step and defined as a function
from the current state (a register set + a memory state)
to either Next rs´ m´ where rs´ and m´ are the updated register
set and memory state after execution of the instruction at rs#PC,
or Stuck if the processor is stuck.
CompCertX:test-compcert-param-memory The outcome now depends on the type mem for
memory states, which is an implicit argument. To have Coq
guess the right one, we make outcome also depend on memory operations.
But we have to declare it explicitly, as we cannot have
this definition reuse the existing context.
Inductive outcome `{memory_model_ops: Mem.MemoryModelOps mem}: Type :=
| Next: regset -> mem -> outcome
| Stuck: outcome.
| Next: regset -> mem -> outcome
| Stuck: outcome.
Manipulations over the PC register: continuing with the next
instruction (nextinstr) or branching to a label (goto_label).
nextinstr_nf is a variant of nextinstr that sets condition flags
to Vundef in addition to incrementing the PC.
Definition nextinstr (rs: regset) :=
rs#PC <- (Val.add rs#PC Vone).
Definition nextinstr_nf (rs: regset) : regset :=
nextinstr (undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) rs).
Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) :=
match label_pos lbl 0 (fn_code f) with
| None => Stuck
| Some pos =>
match rs#PC with
| Vptr b ofs => Next (rs#PC <- (Vptr b (Int.repr pos))) m
| _ => Stuck
end
end.
Auxiliaries for memory accesses.
CompCertiKOS:test-compcert-param-mem-accessors For CertiKOS, we
need to parameterize over exec_load and exec_store, which will be
defined differently depending on whether we are in kernel or user
mode.
Class MemAccessors
`{!Mem.MemoryModelOps mem}
(exec_load: forall F V: Type, Genv.t F V -> memory_chunk -> mem -> addrmode -> regset -> preg -> outcome)
(exec_store: forall F V: Type, Genv.t F V -> memory_chunk -> mem -> addrmode -> regset -> preg -> list preg -> outcome)
: Prop := {}.
Section MEM_ACCESSORS_DEFAULT.
CompCertiKOS:test-compcert-param-mem-accessors Compcert does not
care about kernel vs. user mode, and uses its memory model to define
its memory accessors.
Definition exec_load {F V} (ge: Genv.t F V) (chunk: memory_chunk) (m: mem)
(a: addrmode) (rs: regset) (rd: preg) :=
match Mem.loadv chunk m (eval_addrmode ge a rs) with
| Some v => Next (nextinstr_nf (rs#rd <- v)) m
| None => Stuck
end.
Definition exec_store {F V} (ge: Genv.t F V) (chunk: memory_chunk) (m: mem)
(a: addrmode) (rs: regset) (r1: preg)
(destroyed: list preg) :=
match Mem.storev chunk m (eval_addrmode ge a rs) (rs r1) with
| Some m´ => Next (nextinstr_nf (undef_regs destroyed rs)) m´
| None => Stuck
end.
Local Instance mem_accessors_default: MemAccessors (@exec_load) (@exec_store).
End MEM_ACCESSORS_DEFAULT.
Execution of a single instruction i in initial state
rs and m. Return updated state. For instructions
that correspond to actual IA32 instructions, the cases are
straightforward transliterations of the informal descriptions
given in the IA32 reference manuals. For pseudo-instructions,
refer to the informal descriptions given above.
Note that we set to Vundef the registers used as temporaries by
the expansions of the pseudo-instructions, so that the IA32 code
we generate cannot use those registers to hold values that must
survive the execution of the pseudo-instruction.
Concerning condition flags, the comparison instructions set them
accurately; other instructions (moves, lea) preserve them;
and all other instruction set those flags to Vundef, to reflect
the fact that the processor updates some or all of those flags,
but we do not need to model this precisely.
Definition exec_instr {exec_load exec_store} `{!MemAccessors exec_load exec_store} {F V} (ge: Genv.t F V) (f: function) (i: instruction) (rs: regset) (m: mem) : outcome :=
match i with
Moves
| Pmov_rr rd r1 =>
Next (nextinstr (rs#rd <- (rs r1))) m
| Pmov_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Vint n))) m
| Pmov_ra rd id =>
Next (nextinstr_nf (rs#rd <- (symbol_offset ge id Int.zero))) m
| Pmov_rm rd a =>
exec_load _ _ ge Mint32 m a rs rd
| Pmov_mr a r1 =>
exec_store _ _ ge Mint32 m a rs r1 nil
| Pmovsd_ff rd r1 =>
Next (nextinstr (rs#rd <- (rs r1))) m
| Pmovsd_fi rd n =>
Next (nextinstr (rs#rd <- (Vfloat n))) m
| Pmovsd_fm rd a =>
exec_load _ _ ge Mfloat64 m a rs rd
| Pmovsd_mf a r1 =>
exec_store _ _ ge Mfloat64 m a rs r1 nil
| Pfld_f r1 =>
Next (nextinstr (rs#ST0 <- (rs r1))) m
| Pfld_m a =>
exec_load _ _ ge Mfloat64 m a rs ST0
| Pfstp_f rd =>
Next (nextinstr (rs#rd <- (rs ST0) #ST0 <- Vundef)) m
| Pfstp_m a =>
exec_store _ _ ge Mfloat64 m a rs ST0 (ST0 :: nil)
| Pxchg_rr r1 r2 =>
Next (nextinstr (rs#r1 <- (rs r2) #r2 <- (rs r1))) m
Next (nextinstr (rs#rd <- (rs r1))) m
| Pmov_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Vint n))) m
| Pmov_ra rd id =>
Next (nextinstr_nf (rs#rd <- (symbol_offset ge id Int.zero))) m
| Pmov_rm rd a =>
exec_load _ _ ge Mint32 m a rs rd
| Pmov_mr a r1 =>
exec_store _ _ ge Mint32 m a rs r1 nil
| Pmovsd_ff rd r1 =>
Next (nextinstr (rs#rd <- (rs r1))) m
| Pmovsd_fi rd n =>
Next (nextinstr (rs#rd <- (Vfloat n))) m
| Pmovsd_fm rd a =>
exec_load _ _ ge Mfloat64 m a rs rd
| Pmovsd_mf a r1 =>
exec_store _ _ ge Mfloat64 m a rs r1 nil
| Pfld_f r1 =>
Next (nextinstr (rs#ST0 <- (rs r1))) m
| Pfld_m a =>
exec_load _ _ ge Mfloat64 m a rs ST0
| Pfstp_f rd =>
Next (nextinstr (rs#rd <- (rs ST0) #ST0 <- Vundef)) m
| Pfstp_m a =>
exec_store _ _ ge Mfloat64 m a rs ST0 (ST0 :: nil)
| Pxchg_rr r1 r2 =>
Next (nextinstr (rs#r1 <- (rs r2) #r2 <- (rs r1))) m
Moves with conversion
| Pmovb_mr a r1 =>
exec_store _ _ ge Mint8unsigned m a rs r1 nil
| Pmovw_mr a r1 =>
exec_store _ _ ge Mint16unsigned m a rs r1 nil
| Pmovzb_rr rd r1 =>
Next (nextinstr (rs#rd <- (Val.zero_ext 8 rs#r1))) m
| Pmovzb_rm rd a =>
exec_load _ _ ge Mint8unsigned m a rs rd
| Pmovsb_rr rd r1 =>
Next (nextinstr (rs#rd <- (Val.sign_ext 8 rs#r1))) m
| Pmovsb_rm rd a =>
exec_load _ _ ge Mint8signed m a rs rd
| Pmovzw_rr rd r1 =>
Next (nextinstr (rs#rd <- (Val.zero_ext 16 rs#r1))) m
| Pmovzw_rm rd a =>
exec_load _ _ ge Mint16unsigned m a rs rd
| Pmovsw_rr rd r1 =>
Next (nextinstr (rs#rd <- (Val.sign_ext 16 rs#r1))) m
| Pmovsw_rm rd a =>
exec_load _ _ ge Mint16signed m a rs rd
| Pcvtss2sd_fm rd a =>
exec_load _ _ ge Mfloat32 m a rs rd
| Pcvtsd2ss_ff rd r1 =>
Next (nextinstr (rs#rd <- (Val.singleoffloat rs#r1))) m
| Pcvtsd2ss_mf a r1 =>
exec_store _ _ ge Mfloat32 m a rs r1 (FR XMM7 :: nil)
| Pcvttsd2si_rf rd r1 =>
Next (nextinstr (rs#rd <- (Val.maketotal (Val.intoffloat rs#r1)))) m
| Pcvtsi2sd_fr rd r1 =>
Next (nextinstr (rs#rd <- (Val.maketotal (Val.floatofint rs#r1)))) m
exec_store _ _ ge Mint8unsigned m a rs r1 nil
| Pmovw_mr a r1 =>
exec_store _ _ ge Mint16unsigned m a rs r1 nil
| Pmovzb_rr rd r1 =>
Next (nextinstr (rs#rd <- (Val.zero_ext 8 rs#r1))) m
| Pmovzb_rm rd a =>
exec_load _ _ ge Mint8unsigned m a rs rd
| Pmovsb_rr rd r1 =>
Next (nextinstr (rs#rd <- (Val.sign_ext 8 rs#r1))) m
| Pmovsb_rm rd a =>
exec_load _ _ ge Mint8signed m a rs rd
| Pmovzw_rr rd r1 =>
Next (nextinstr (rs#rd <- (Val.zero_ext 16 rs#r1))) m
| Pmovzw_rm rd a =>
exec_load _ _ ge Mint16unsigned m a rs rd
| Pmovsw_rr rd r1 =>
Next (nextinstr (rs#rd <- (Val.sign_ext 16 rs#r1))) m
| Pmovsw_rm rd a =>
exec_load _ _ ge Mint16signed m a rs rd
| Pcvtss2sd_fm rd a =>
exec_load _ _ ge Mfloat32 m a rs rd
| Pcvtsd2ss_ff rd r1 =>
Next (nextinstr (rs#rd <- (Val.singleoffloat rs#r1))) m
| Pcvtsd2ss_mf a r1 =>
exec_store _ _ ge Mfloat32 m a rs r1 (FR XMM7 :: nil)
| Pcvttsd2si_rf rd r1 =>
Next (nextinstr (rs#rd <- (Val.maketotal (Val.intoffloat rs#r1)))) m
| Pcvtsi2sd_fr rd r1 =>
Next (nextinstr (rs#rd <- (Val.maketotal (Val.floatofint rs#r1)))) m
Integer arithmetic
| Plea rd a =>
Next (nextinstr (rs#rd <- (eval_addrmode ge a rs))) m
| Pneg rd =>
Next (nextinstr_nf (rs#rd <- (Val.neg rs#rd))) m
| Psub_rr rd r1 =>
Next (nextinstr_nf (rs#rd <- (Val.sub rs#rd rs#r1))) m
| Pimul_rr rd r1 =>
Next (nextinstr_nf (rs#rd <- (Val.mul rs#rd rs#r1))) m
| Pimul_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.mul rs#rd (Vint n)))) m
| Pimul_r r1 =>
Next (nextinstr_nf (rs#EAX <- (Val.mul rs#EAX rs#r1)
#EDX <- (Val.mulhs rs#EAX rs#r1))) m
| Pmul_r r1 =>
Next (nextinstr_nf (rs#EAX <- (Val.mul rs#EAX rs#r1)
#EDX <- (Val.mulhu rs#EAX rs#r1))) m
| Pdiv r1 =>
let vn := rs#EAX in let vd := (rs#EDX <- Vundef)#r1 in
match Val.divu vn vd, Val.modu vn vd with
| Some vq, Some vr => Next (nextinstr_nf (rs#EAX <- vq #EDX <- vr)) m
| _, _ => Stuck
end
| Pidiv r1 =>
let vn := rs#EAX in let vd := (rs#EDX <- Vundef)#r1 in
match Val.divs vn vd, Val.mods vn vd with
| Some vq, Some vr => Next (nextinstr_nf (rs#EAX <- vq #EDX <- vr)) m
| _, _ => Stuck
end
| Pand_rr rd r1 =>
Next (nextinstr_nf (rs#rd <- (Val.and rs#rd rs#r1))) m
| Pand_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.and rs#rd (Vint n)))) m
| Por_rr rd r1 =>
Next (nextinstr_nf (rs#rd <- (Val.or rs#rd rs#r1))) m
| Por_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.or rs#rd (Vint n)))) m
| Pxor_r rd =>
Next (nextinstr_nf (rs#rd <- Vzero)) m
| Pxor_rr rd r1 =>
Next (nextinstr_nf (rs#rd <- (Val.xor rs#rd rs#r1))) m
| Pxor_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.xor rs#rd (Vint n)))) m
| Pnot rd =>
Next (nextinstr_nf (rs#rd <- (Val.notint rs#rd))) m
| Psal_rcl rd =>
Next (nextinstr_nf (rs#rd <- (Val.shl rs#rd rs#ECX))) m
| Psal_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.shl rs#rd (Vint n)))) m
| Pshr_rcl rd =>
Next (nextinstr_nf (rs#rd <- (Val.shru rs#rd rs#ECX))) m
| Pshr_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.shru rs#rd (Vint n)))) m
| Psar_rcl rd =>
Next (nextinstr_nf (rs#rd <- (Val.shr rs#rd rs#ECX))) m
| Psar_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.shr rs#rd (Vint n)))) m
| Pshld_ri rd r1 n =>
Next (nextinstr_nf
(rs#rd <- (Val.or (Val.shl rs#rd (Vint n))
(Val.shru rs#r1 (Vint (Int.sub Int.iwordsize n)))))) m
| Pror_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.ror rs#rd (Vint n)))) m
| Pcmp_rr r1 r2 =>
Next (nextinstr (compare_ints (rs r1) (rs r2) rs m)) m
| Pcmp_ri r1 n =>
Next (nextinstr (compare_ints (rs r1) (Vint n) rs m)) m
| Ptest_rr r1 r2 =>
Next (nextinstr (compare_ints (Val.and (rs r1) (rs r2)) Vzero rs m)) m
| Ptest_ri r1 n =>
Next (nextinstr (compare_ints (Val.and (rs r1) (Vint n)) Vzero rs m)) m
| Pcmov c rd r1 =>
match eval_testcond c rs with
| Some true => Next (nextinstr (rs#rd <- (rs#r1))) m
| Some false => Next (nextinstr rs) m
| None => Next (nextinstr (rs#rd <- Vundef)) m
end
| Psetcc c rd =>
Next (nextinstr (rs#rd <- (Val.of_optbool (eval_testcond c rs)))) m
Next (nextinstr (rs#rd <- (eval_addrmode ge a rs))) m
| Pneg rd =>
Next (nextinstr_nf (rs#rd <- (Val.neg rs#rd))) m
| Psub_rr rd r1 =>
Next (nextinstr_nf (rs#rd <- (Val.sub rs#rd rs#r1))) m
| Pimul_rr rd r1 =>
Next (nextinstr_nf (rs#rd <- (Val.mul rs#rd rs#r1))) m
| Pimul_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.mul rs#rd (Vint n)))) m
| Pimul_r r1 =>
Next (nextinstr_nf (rs#EAX <- (Val.mul rs#EAX rs#r1)
#EDX <- (Val.mulhs rs#EAX rs#r1))) m
| Pmul_r r1 =>
Next (nextinstr_nf (rs#EAX <- (Val.mul rs#EAX rs#r1)
#EDX <- (Val.mulhu rs#EAX rs#r1))) m
| Pdiv r1 =>
let vn := rs#EAX in let vd := (rs#EDX <- Vundef)#r1 in
match Val.divu vn vd, Val.modu vn vd with
| Some vq, Some vr => Next (nextinstr_nf (rs#EAX <- vq #EDX <- vr)) m
| _, _ => Stuck
end
| Pidiv r1 =>
let vn := rs#EAX in let vd := (rs#EDX <- Vundef)#r1 in
match Val.divs vn vd, Val.mods vn vd with
| Some vq, Some vr => Next (nextinstr_nf (rs#EAX <- vq #EDX <- vr)) m
| _, _ => Stuck
end
| Pand_rr rd r1 =>
Next (nextinstr_nf (rs#rd <- (Val.and rs#rd rs#r1))) m
| Pand_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.and rs#rd (Vint n)))) m
| Por_rr rd r1 =>
Next (nextinstr_nf (rs#rd <- (Val.or rs#rd rs#r1))) m
| Por_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.or rs#rd (Vint n)))) m
| Pxor_r rd =>
Next (nextinstr_nf (rs#rd <- Vzero)) m
| Pxor_rr rd r1 =>
Next (nextinstr_nf (rs#rd <- (Val.xor rs#rd rs#r1))) m
| Pxor_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.xor rs#rd (Vint n)))) m
| Pnot rd =>
Next (nextinstr_nf (rs#rd <- (Val.notint rs#rd))) m
| Psal_rcl rd =>
Next (nextinstr_nf (rs#rd <- (Val.shl rs#rd rs#ECX))) m
| Psal_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.shl rs#rd (Vint n)))) m
| Pshr_rcl rd =>
Next (nextinstr_nf (rs#rd <- (Val.shru rs#rd rs#ECX))) m
| Pshr_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.shru rs#rd (Vint n)))) m
| Psar_rcl rd =>
Next (nextinstr_nf (rs#rd <- (Val.shr rs#rd rs#ECX))) m
| Psar_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.shr rs#rd (Vint n)))) m
| Pshld_ri rd r1 n =>
Next (nextinstr_nf
(rs#rd <- (Val.or (Val.shl rs#rd (Vint n))
(Val.shru rs#r1 (Vint (Int.sub Int.iwordsize n)))))) m
| Pror_ri rd n =>
Next (nextinstr_nf (rs#rd <- (Val.ror rs#rd (Vint n)))) m
| Pcmp_rr r1 r2 =>
Next (nextinstr (compare_ints (rs r1) (rs r2) rs m)) m
| Pcmp_ri r1 n =>
Next (nextinstr (compare_ints (rs r1) (Vint n) rs m)) m
| Ptest_rr r1 r2 =>
Next (nextinstr (compare_ints (Val.and (rs r1) (rs r2)) Vzero rs m)) m
| Ptest_ri r1 n =>
Next (nextinstr (compare_ints (Val.and (rs r1) (Vint n)) Vzero rs m)) m
| Pcmov c rd r1 =>
match eval_testcond c rs with
| Some true => Next (nextinstr (rs#rd <- (rs#r1))) m
| Some false => Next (nextinstr rs) m
| None => Next (nextinstr (rs#rd <- Vundef)) m
end
| Psetcc c rd =>
Next (nextinstr (rs#rd <- (Val.of_optbool (eval_testcond c rs)))) m
Arithmetic operations over floats
| Paddd_ff rd r1 =>
Next (nextinstr (rs#rd <- (Val.addf rs#rd rs#r1))) m
| Psubd_ff rd r1 =>
Next (nextinstr (rs#rd <- (Val.subf rs#rd rs#r1))) m
| Pmuld_ff rd r1 =>
Next (nextinstr (rs#rd <- (Val.mulf rs#rd rs#r1))) m
| Pdivd_ff rd r1 =>
Next (nextinstr (rs#rd <- (Val.divf rs#rd rs#r1))) m
| Pnegd rd =>
Next (nextinstr (rs#rd <- (Val.negf rs#rd))) m
| Pabsd rd =>
Next (nextinstr (rs#rd <- (Val.absf rs#rd))) m
| Pcomisd_ff r1 r2 =>
Next (nextinstr (compare_floats (rs r1) (rs r2) rs)) m
| Pxorpd_f rd =>
Next (nextinstr_nf (rs#rd <- (Vfloat Float.zero))) m
Next (nextinstr (rs#rd <- (Val.addf rs#rd rs#r1))) m
| Psubd_ff rd r1 =>
Next (nextinstr (rs#rd <- (Val.subf rs#rd rs#r1))) m
| Pmuld_ff rd r1 =>
Next (nextinstr (rs#rd <- (Val.mulf rs#rd rs#r1))) m
| Pdivd_ff rd r1 =>
Next (nextinstr (rs#rd <- (Val.divf rs#rd rs#r1))) m
| Pnegd rd =>
Next (nextinstr (rs#rd <- (Val.negf rs#rd))) m
| Pabsd rd =>
Next (nextinstr (rs#rd <- (Val.absf rs#rd))) m
| Pcomisd_ff r1 r2 =>
Next (nextinstr (compare_floats (rs r1) (rs r2) rs)) m
| Pxorpd_f rd =>
Next (nextinstr_nf (rs#rd <- (Vfloat Float.zero))) m
Branches and calls
| Pjmp_l lbl =>
goto_label f lbl rs m
| Pjmp_s id sg =>
Next (rs#PC <- (symbol_offset ge id Int.zero)) m
| Pjmp_r r sg =>
Next (rs#PC <- (rs r)) m
| Pjcc cond lbl =>
match eval_testcond cond rs with
| Some true => goto_label f lbl rs m
| Some false => Next (nextinstr rs) m
| None => Stuck
end
| Pjcc2 cond1 cond2 lbl =>
match eval_testcond cond1 rs, eval_testcond cond2 rs with
| Some true, Some true => goto_label f lbl rs m
| Some _, Some _ => Next (nextinstr rs) m
| _, _ => Stuck
end
| Pjmptbl r tbl =>
match rs#r with
| Vint n =>
match list_nth_z tbl (Int.unsigned n) with
| None => Stuck
| Some lbl => goto_label f lbl rs m
end
| _ => Stuck
end
| Pcall_s id sg =>
Next (rs#RA <- (Val.add rs#PC Vone) #PC <- (symbol_offset ge id Int.zero)) m
| Pcall_r r sg =>
Next (rs#RA <- (Val.add rs#PC Vone) #PC <- (rs r)) m
| Pret =>
goto_label f lbl rs m
| Pjmp_s id sg =>
Next (rs#PC <- (symbol_offset ge id Int.zero)) m
| Pjmp_r r sg =>
Next (rs#PC <- (rs r)) m
| Pjcc cond lbl =>
match eval_testcond cond rs with
| Some true => goto_label f lbl rs m
| Some false => Next (nextinstr rs) m
| None => Stuck
end
| Pjcc2 cond1 cond2 lbl =>
match eval_testcond cond1 rs, eval_testcond cond2 rs with
| Some true, Some true => goto_label f lbl rs m
| Some _, Some _ => Next (nextinstr rs) m
| _, _ => Stuck
end
| Pjmptbl r tbl =>
match rs#r with
| Vint n =>
match list_nth_z tbl (Int.unsigned n) with
| None => Stuck
| Some lbl => goto_label f lbl rs m
end
| _ => Stuck
end
| Pcall_s id sg =>
Next (rs#RA <- (Val.add rs#PC Vone) #PC <- (symbol_offset ge id Int.zero)) m
| Pcall_r r sg =>
Next (rs#RA <- (Val.add rs#PC Vone) #PC <- (rs r)) m
| Pret =>
CompCertX:test-compcert-ra-vundef We need to erase the value of RA,
which is actually popped away from the stack in reality.
Pseudo-instructions
| Plabel lbl =>
Next (nextinstr rs) m
| Pallocframe sz ofs_ra ofs_link =>
let (m1, stk) := Mem.alloc m 0 sz in
let sp := Vptr stk Int.zero in
match Mem.storev Mint32 m1 (Val.add sp (Vint ofs_link)) rs#ESP with
| None => Stuck
| Some m2 =>
match Mem.storev Mint32 m2 (Val.add sp (Vint ofs_ra)) rs#RA with
| None => Stuck
| Some m3 => Next (nextinstr (rs #EDX <- (rs#ESP) #ESP <- sp)) m3
end
end
| Pfreeframe sz ofs_ra ofs_link =>
match Mem.loadv Mint32 m (Val.add rs#ESP (Vint ofs_ra)) with
| None => Stuck
| Some ra =>
match Mem.loadv Mint32 m (Val.add rs#ESP (Vint ofs_link)) with
| None => Stuck
| Some sp =>
match rs#ESP with
| Vptr stk ofs =>
match Mem.free m stk 0 sz with
| None => Stuck
| Some m´ => Next (nextinstr (rs#ESP <- sp #RA <- ra)) m´
end
| _ => Stuck
end
end
end
| Pbuiltin ef args res =>
Stuck
| Pannot ef args =>
Stuck
end.
End WITH_FIND_LABELS.
Next (nextinstr rs) m
| Pallocframe sz ofs_ra ofs_link =>
let (m1, stk) := Mem.alloc m 0 sz in
let sp := Vptr stk Int.zero in
match Mem.storev Mint32 m1 (Val.add sp (Vint ofs_link)) rs#ESP with
| None => Stuck
| Some m2 =>
match Mem.storev Mint32 m2 (Val.add sp (Vint ofs_ra)) rs#RA with
| None => Stuck
| Some m3 => Next (nextinstr (rs #EDX <- (rs#ESP) #ESP <- sp)) m3
end
end
| Pfreeframe sz ofs_ra ofs_link =>
match Mem.loadv Mint32 m (Val.add rs#ESP (Vint ofs_ra)) with
| None => Stuck
| Some ra =>
match Mem.loadv Mint32 m (Val.add rs#ESP (Vint ofs_link)) with
| None => Stuck
| Some sp =>
match rs#ESP with
| Vptr stk ofs =>
match Mem.free m stk 0 sz with
| None => Stuck
| Some m´ => Next (nextinstr (rs#ESP <- sp #RA <- ra)) m´
end
| _ => Stuck
end
end
end
| Pbuiltin ef args res =>
Stuck
| Pannot ef args =>
Stuck
end.
End WITH_FIND_LABELS.
Translation of the LTL/Linear/Mach view of machine registers
to the Asm view.
Definition preg_of (r: mreg) : preg :=
match r with
| AX => IR EAX
| BX => IR EBX
| CX => IR ECX
| DX => IR EDX
| SI => IR ESI
| DI => IR EDI
| BP => IR EBP
| X0 => FR XMM0
| X1 => FR XMM1
| X2 => FR XMM2
| X3 => FR XMM3
| X4 => FR XMM4
| X5 => FR XMM5
| X6 => FR XMM6
| X7 => FR XMM7
| FP0 => ST0
end.
Extract the values of the arguments of an external call.
We exploit the calling conventions from module Conventions, except that
we use machine registers instead of locations.
Inductive extcall_arg (rs: regset) (m: mem): loc -> val -> Prop :=
| extcall_arg_reg: forall r,
extcall_arg rs m (R r) (rs (preg_of r))
| extcall_arg_stack: forall ofs ty bofs v,
bofs = Stacklayout.fe_ofs_arg + 4 * ofs ->
Mem.loadv (chunk_of_type ty) m
(Val.add (rs (IR ESP)) (Vint (Int.repr bofs))) = Some v ->
extcall_arg rs m (S Outgoing ofs ty) v.
Definition extcall_arguments
(rs: regset) (m: mem) (sg: signature) (args: list val) : Prop :=
list_forall2 (extcall_arg rs m) (loc_arguments sg) args.
Definition loc_external_result (sg: signature) : list preg :=
map preg_of (loc_result sg).
Extract the values of the arguments of an annotation.
Inductive annot_arg (rs: regset) (m: mem): annot_param -> val -> Prop :=
| annot_arg_reg: forall r,
annot_arg rs m (APreg r) (rs r)
| annot_arg_stack: forall chunk ofs stk base v,
rs (IR ESP) = Vptr stk base ->
Mem.load chunk m stk (Int.unsigned base + ofs) = Some v ->
annot_arg rs m (APstack chunk ofs) v.
Definition annot_arguments
(rs: regset) (m: mem) (params: list annot_param) (args: list val) : Prop :=
list_forall2 (annot_arg rs m) params args.
Execution of the instruction at rs#PC.
CompCertX:test-compcert-param-memory The state now depends on the type mem for
memory states, which is an implicit argument. To have Coq
guess the right one, we make state also depend on memory operations.
But we have to declare it explicitly, as we cannot have
this definition reuse the existing context.
Inductive state `{memory_model_ops: Mem.MemoryModelOps mem}: Type :=
| State: regset -> mem -> state.
Function typ_of_preg (r: preg): typ :=
match r with
| FR _ => Tfloat
| ST0 => Tfloat
| _ => Tint
end.
Inductive step {exec_load exec_store} `{!MemAccessors exec_load exec_store} (ge: genv): state -> trace -> state -> Prop :=
| exec_step_internal:
forall b ofs f i rs m rs´ m´,
rs PC = Vptr b ofs ->
Genv.find_funct_ptr ge b = Some (Internal f) ->
find_instr (Int.unsigned ofs) f.(fn_code) = Some i ->
exec_instr ge f i rs m = Next rs´ m´ ->
step ge (State rs m) E0 (State rs´ m´)
| exec_step_builtin:
forall b ofs f ef args res rs m t vl rs´ m´,
rs PC = Vptr b ofs ->
Genv.find_funct_ptr ge b = Some (Internal f) ->
find_instr (Int.unsigned ofs) f.(fn_code) = Some (Pbuiltin ef args res) ->
external_call´ (fun _ => True) ef ge (map rs args) m t vl m´ ->
rs´ = nextinstr_nf
(set_regs res vl
(undef_regs (map preg_of (destroyed_by_builtin ef)) rs)) ->
| State: regset -> mem -> state.
Function typ_of_preg (r: preg): typ :=
match r with
| FR _ => Tfloat
| ST0 => Tfloat
| _ => Tint
end.
Inductive step {exec_load exec_store} `{!MemAccessors exec_load exec_store} (ge: genv): state -> trace -> state -> Prop :=
| exec_step_internal:
forall b ofs f i rs m rs´ m´,
rs PC = Vptr b ofs ->
Genv.find_funct_ptr ge b = Some (Internal f) ->
find_instr (Int.unsigned ofs) f.(fn_code) = Some i ->
exec_instr ge f i rs m = Next rs´ m´ ->
step ge (State rs m) E0 (State rs´ m´)
| exec_step_builtin:
forall b ofs f ef args res rs m t vl rs´ m´,
rs PC = Vptr b ofs ->
Genv.find_funct_ptr ge b = Some (Internal f) ->
find_instr (Int.unsigned ofs) f.(fn_code) = Some (Pbuiltin ef args res) ->
external_call´ (fun _ => True) ef ge (map rs args) m t vl m´ ->
rs´ = nextinstr_nf
(set_regs res vl
(undef_regs (map preg_of (destroyed_by_builtin ef)) rs)) ->
CompCertX:test-compcert-disable-extcall-as-builtin We may need
to disallow the use of external function calls (EF_external) as
builtins. This is already the case in assembly generation
(PrintAsm.ml), but not in the semantics of languages, which we propose
to fix through providing a switch in the compiler configuration, hence
the CompilerConfigOps class, and this new clause in the operational
semantics.
CompCertX:test-compcert-wt-builtin We need to prove that registers updated by builtins are
of the same type as the return type of the builtin.
forall BUILTIN_WT:
subtype_list (proj_sig_res´ (ef_sig ef)) (map typ_of_preg res) = true,
step ge (State rs m) t (State rs´ m´)
| exec_step_annot:
forall b ofs f ef args rs m vargs t v m´,
rs PC = Vptr b ofs ->
Genv.find_funct_ptr ge b = Some (Internal f) ->
find_instr (Int.unsigned ofs) f.(fn_code) = Some (Pannot ef args) ->
annot_arguments rs m args vargs ->
external_call´ (fun _ => True) ef ge vargs m t v m´ ->
forall BUILTIN_ENABLED: builtin_enabled ef,
step ge (State rs m) t
(State (nextinstr rs) m´)
| exec_step_external:
forall b ef args res rs m t rs´ m´,
rs PC = Vptr b Int.zero ->
Genv.find_funct_ptr ge b = Some (External ef) ->
extcall_arguments rs m (ef_sig ef) args ->
external_call´ (fun _ => True) ef ge args m t res m´ ->
subtype_list (proj_sig_res´ (ef_sig ef)) (map typ_of_preg res) = true,
step ge (State rs m) t (State rs´ m´)
| exec_step_annot:
forall b ofs f ef args rs m vargs t v m´,
rs PC = Vptr b ofs ->
Genv.find_funct_ptr ge b = Some (Internal f) ->
find_instr (Int.unsigned ofs) f.(fn_code) = Some (Pannot ef args) ->
annot_arguments rs m args vargs ->
external_call´ (fun _ => True) ef ge vargs m t v m´ ->
forall BUILTIN_ENABLED: builtin_enabled ef,
step ge (State rs m) t
(State (nextinstr rs) m´)
| exec_step_external:
forall b ef args res rs m t rs´ m´,
rs PC = Vptr b Int.zero ->
Genv.find_funct_ptr ge b = Some (External ef) ->
extcall_arguments rs m (ef_sig ef) args ->
external_call´ (fun _ => True) ef ge args m t res m´ ->
CompCertX:test-compcert-undef-destroyed-by-call We erase non-callee-save registers. CompCertX:test-compcert-ra-vundef We need to erase the value of RA,
which is actually popped away from the stack in reality.
rs´ = (set_regs (loc_external_result (ef_sig ef)) res (undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF :: nil) (undef_regs (map preg_of destroyed_at_call) rs))) #PC <- (rs RA) #RA <- Vundef ->
CompCertX:test-compcert-protect-stack-arg The following STACK condition
allows to protect the locations of the caller's stack corresponding to the
arguments passed to the callee, for non-whole-program settings. This
STACK condition will be required at function entry for code compiled from
Clight replacing the external function. Then, later, this Clight code will
need to prove that no such memory locations are ever overwritten.
forall STACK:
forall b o, rs ESP = Vptr b o ->
(Ple (Genv.genv_next ge) b /\ Plt b (Mem.nextblock m)),
forall b o, rs ESP = Vptr b o ->
(Ple (Genv.genv_next ge) b /\ Plt b (Mem.nextblock m)),
CompCertX:test-compcert-protect-stack-arg The following
NOT_VUNDEF conditions will allow to later parameterize the
per-function/module semantics of back-end languages over the stack
pointer and the return address.
forall SP_NOT_VUNDEF: rs ESP <> Vundef,
forall RA_NOT_VUNDEF: rs RA <> Vundef,
step ge (State rs m) t (State rs´ m´).
End RELSEM.
forall RA_NOT_VUNDEF: rs RA <> Vundef,
step ge (State rs m) t (State rs´ m´).
End RELSEM.
Execution of whole programs.
Inductive initial_state {F V} (p: AST.program F V): state -> Prop :=
| initial_state_intro: forall m0,
Genv.init_mem p = Some m0 ->
let ge := Genv.globalenv p in
let rs0 :=
(Pregmap.init Vundef)
# PC <- (symbol_offset ge p.(prog_main) Int.zero)
# RA <- Vzero
# ESP <- Vzero in
initial_state p (State rs0 m0).
Inductive final_state: state -> int -> Prop :=
| final_state_intro: forall rs m r,
rs#PC = Vzero ->
rs#EAX = Vint r ->
final_state (State rs m) r.
CompCertiKOS:test-compcert-param-mem-accessors Compcert does not
care about kernel vs. user mode, and uses its memory model to define
its memory accessors.
Local Existing Instance mem_accessors_default.
Definition semantics (p: program) :=
Semantics step (initial_state p) final_state (Genv.globalenv p).
Definition semantics (p: program) :=
Semantics step (initial_state p) final_state (Genv.globalenv p).
Determinacy of the Asm semantics.
Remark extcall_arguments_determ:
forall rs m sg args1 args2,
extcall_arguments rs m sg args1 -> extcall_arguments rs m sg args2 -> args1 = args2.
Proof.
intros until m.
assert (forall ll vl1, list_forall2 (extcall_arg rs m) ll vl1 ->
forall vl2, list_forall2 (extcall_arg rs m) ll vl2 -> vl1 = vl2).
induction 1; intros vl2 EA; inv EA.
auto.
f_equal; auto.
inv H; inv H3; congruence.
intros. red in H0; red in H1. eauto.
Qed.
Remark annot_arguments_determ:
forall rs m params args1 args2,
annot_arguments rs m params args1 -> annot_arguments rs m params args2 -> args1 = args2.
Proof.
unfold annot_arguments. intros. revert params args1 H args2 H0.
induction 1; intros.
inv H0; auto.
inv H1. decEq; eauto. inv H; inv H4. auto. congruence.
Qed.
Lemma semantics_determinate: forall p, determinate (semantics p).
Proof.
Ltac Equalities :=
match goal with
| [ H1: ?a = ?b, H2: ?a = ?c |- _ ] =>
rewrite H1 in H2; inv H2; Equalities
| _ => idtac
end.
intros; constructor; simpl; intros.
inv H; inv H0; Equalities.
split. constructor. auto.
discriminate.
discriminate.
inv H11.
exploit external_call_determ´. eexact H4. eexact H9. intros [A B].
split. auto. intros. destruct B; auto. subst. auto.
inv H12.
assert (vargs0 = vargs) by (eapply annot_arguments_determ; eauto). subst vargs0.
exploit external_call_determ´. eexact H5. eexact H13. intros [A B].
split. auto. intros. destruct B; auto. subst. auto.
assert (args0 = args) by (eapply extcall_arguments_determ; eauto). subst args0.
exploit external_call_determ´. eexact H4. eexact H9. intros [A B].
split. auto. intros. destruct B; auto. subst. auto.
red; intros; inv H; simpl.
omega.
inv H3. eapply external_call_trace_length; eauto.
inv H4. eapply external_call_trace_length; eauto.
inv H3. eapply external_call_trace_length; eauto.
inv H; inv H0. f_equal. congruence.
inv H. unfold Vzero in H0. red; intros; red; intros. inv H; congruence.
inv H; inv H0. congruence.
Qed.
Classification functions for processor registers (used in Asmgenproof).