Correctness proof for x86-64 generation: auxiliary results.
Require Import Coqlib.
Require Import AST Errors Integers Floats Values Memory Globalenvs.
Require Import Op Locations Conventions Mach Asm.
Require Import Asmgen Asmgenproof0.
Local Open Scope error_monad_scope.
Correspondence between Mach registers and x86 registers
Lemma agree_nextinstr_nf:
forall ms sp rs sz,
agree ms sp rs ->
agree ms sp (
nextinstr_nf rs sz).
Proof.
Useful properties of the PC register.
Lemma nextinstr_nf_inv:
forall r rs sz,
match r with PC =>
False |
CR _ =>
False |
_ =>
True end ->
(
nextinstr_nf rs sz)#
r =
rs#
r.
Proof.
Lemma nextinstr_nf_inv1:
forall r rs sz,
data_preg r =
true -> (
nextinstr_nf rs sz)#
r =
rs#
r.
Proof.
Lemma nextinstr_nf_set_preg:
forall rs m v sz,
(
nextinstr_nf (
rs#(
preg_of m) <-
v)
sz)#
PC =
Val.offset_ptr rs#
PC sz.
Proof.
Useful simplification tactic
Ltac Simplif :=
match goal with
| [ |-
nextinstr_nf _ _ _ =
_ ] =>
((
rewrite nextinstr_nf_inv by auto with asmgen)
|| (
rewrite nextinstr_nf_inv1 by auto with asmgen));
auto
| [ |-
nextinstr _ _ _ =
_ ] =>
((
rewrite nextinstr_inv by auto with asmgen)
|| (
rewrite nextinstr_inv1 by auto with asmgen));
auto
| [ |-
Pregmap.get ?
x (
Pregmap.set ?
x _ _) =
_ ] =>
rewrite Pregmap.gss;
auto
| [ |-
Pregmap.set ?
x _ _ ?
x =
_ ] =>
rewrite Pregmap.gss;
auto
| [ |-
Pregmap.get _ (
Pregmap.set _ _ _) =
_ ] =>
rewrite Pregmap.gso by (
auto with asmgen);
auto
| [ |-
Pregmap.set _ _ _ _ =
_ ] =>
rewrite Pregmap.gso by (
auto with asmgen);
auto
end.
Ltac Simplifs :=
repeat Simplif.
Correctness of x86-64 constructor functions
Local Existing Instance mem_accessors_default.
Section CONSTRUCTORS.
Context `{
memory_model_prf:
Mem.MemoryModel}.
Variable init_stk:
stack.
Variable ge:
genv.
Variable fn:
function.
Definition instr_size_in_ptrofs (
i:
instruction) :
ptrofs :=
Ptrofs.repr (
instr_size i).
Smart constructor for moves.
Lemma mk_mov_correct:
forall rd rs k c rs1 m,
mk_mov rd rs k =
OK c ->
exists rs2,
exec_straight init_stk ge fn c rs1 m k rs2 m
/\
rs2#
rd =
rs1#
rs
/\
forall r,
data_preg r =
true ->
r <>
rd ->
rs2#
r =
rs1#
r.
Proof.
unfold mk_mov;
intros.
destruct rd;
try (
monadInv H);
destruct rs;
monadInv H.
mov *)
econstructor.
split.
apply exec_straight_one.
unfold exec_instr;
simpl.
eauto.
auto.
split.
Simplifs.
intros;
Simplifs.
movsd *)
econstructor.
split.
apply exec_straight_one.
unfold exec_instr;
simpl.
eauto.
auto.
split.
Simplifs.
intros;
Simplifs.
Qed.
Properties of division
Lemma divu_modu_exists:
forall v1 v2,
Val.divu v1 v2 <>
None \/
Val.modu v1 v2 <>
None ->
exists n d q r,
v1 =
Vint n /\
v2 =
Vint d
/\
Int.divmodu2 Int.zero n d =
Some(
q,
r)
/\
Val.divu v1 v2 =
Some (
Vint q) /\
Val.modu v1 v2 =
Some (
Vint r).
Proof.
Lemma divs_mods_exists:
forall v1 v2,
Val.divs v1 v2 <>
None \/
Val.mods v1 v2 <>
None ->
exists nh nl d q r,
Val.shr v1 (
Vint (
Int.repr 31)) =
Vint nh /\
v1 =
Vint nl /\
v2 =
Vint d
/\
Int.divmods2 nh nl d =
Some(
q,
r)
/\
Val.divs v1 v2 =
Some (
Vint q) /\
Val.mods v1 v2 =
Some (
Vint r).
Proof.
Lemma divlu_modlu_exists:
forall v1 v2,
Val.divlu v1 v2 <>
None \/
Val.modlu v1 v2 <>
None ->
exists n d q r,
v1 =
Vlong n /\
v2 =
Vlong d
/\
Int64.divmodu2 Int64.zero n d =
Some(
q,
r)
/\
Val.divlu v1 v2 =
Some (
Vlong q) /\
Val.modlu v1 v2 =
Some (
Vlong r).
Proof.
Lemma divls_modls_exists:
forall v1 v2,
Val.divls v1 v2 <>
None \/
Val.modls v1 v2 <>
None ->
exists nh nl d q r,
Val.shrl v1 (
Vint (
Int.repr 63)) =
Vlong nh /\
v1 =
Vlong nl /\
v2 =
Vlong d
/\
Int64.divmods2 nh nl d =
Some(
q,
r)
/\
Val.divls v1 v2 =
Some (
Vlong q) /\
Val.modls v1 v2 =
Some (
Vlong r).
Proof.
Smart constructor for shrx
Lemma mk_shrximm_correct:
forall n k c (
rs1:
regset)
v m,
mk_shrximm n k =
OK c ->
Val.shrx (
rs1#
RAX) (
Vint n) =
Some v ->
exists rs2,
exec_straight init_stk ge fn c rs1 m k rs2 m
/\
rs2#
RAX =
v
/\
forall r,
data_preg r =
true ->
r <>
RAX ->
r <>
RCX ->
rs2#
r =
rs1#
r.
Proof.
Smart constructor for shrxl
Lemma mk_shrxlimm_correct:
forall n k c (
rs1:
regset)
v m,
mk_shrxlimm n k =
OK c ->
Val.shrxl (
rs1#
RAX) (
Vint n) =
Some v ->
exists rs2,
exec_straight init_stk ge fn c rs1 m k rs2 m
/\
rs2#
RAX =
v
/\
forall r,
data_preg r =
true ->
r <>
RAX ->
r <>
RDX ->
rs2#
r =
rs1#
r.
Proof.
Smart constructor for integer conversions
Lemma mk_intconv_correct:
forall mk sem rd rs k c rs1 m,
mk_intconv mk rd rs k =
OK c ->
(
forall c rd rs r m,
exec_instr init_stk ge c (
mk rd rs)
r m =
Next (
nextinstr (
r#
rd <- (
sem r#
rs)) (
instr_size_in_ptrofs(
mk rd rs)))
m) ->
exists rs2,
exec_straight init_stk ge fn c rs1 m k rs2 m
/\
rs2#
rd =
sem rs1#
rs
/\
forall r,
data_preg r =
true ->
r <>
rd ->
r <>
RAX ->
rs2#
r =
rs1#
r.
Proof.
Smart constructor for small stores
Lemma addressing_mentions_correct:
forall a r (
rs1 rs2:
regset),
(
forall (
r':
ireg),
r' <>
r ->
rs1 r' =
rs2 r') ->
addressing_mentions a r =
false ->
eval_addrmode32 ge a rs1 =
eval_addrmode32 ge a rs2.
Proof.
Lemma mk_storebyte_correct:
forall addr r k c rs1 m1 m2,
mk_storebyte addr r k =
OK c ->
Mem.storev Mint8unsigned m1 (
eval_addrmode ge addr rs1) (
rs1 r) =
Some m2 ->
exists rs2,
exec_straight init_stk ge fn c rs1 m1 k rs2 m2
/\
forall r,
data_preg r =
true ->
preg_notin r (
if Archi.ptr64 then nil else AX ::
CX ::
nil) ->
rs2#
r =
rs1#
r.
Proof.
Accessing slots in the stack frame
Remark eval_addrmode_indexed:
forall (
base:
ireg)
ofs (
rs:
regset),
match rs#
base with Vptr _ _ =>
True |
_ =>
False end ->
eval_addrmode ge (
Addrmode (
Some base)
None (
inl _ (
Ptrofs.unsigned ofs)))
rs =
Val.offset_ptr rs#
base ofs.
Proof.
Ltac loadind_correct_solve :=
match goal with
|
H:
Error _ =
OK _ |-
_ =>
discriminate H
|
H:
OK _ =
OK _ |-
_ =>
inv H
|
H:
match ?
x with _ =>
_ end =
OK _ |-
_ =>
destruct x eqn:?;
loadind_correct_solve
|
_ =>
idtac
end.
Lemma loadind_correct:
forall (
base:
ireg)
ofs ty dst k (
rs:
regset)
c m v,
loadind base ofs ty dst k =
OK c ->
Mem.loadv (
chunk_of_type ty)
m (
Val.offset_ptr rs#
base ofs) =
Some v ->
exists rs',
exec_straight init_stk ge fn c rs m k rs'
m
/\
rs'#(
preg_of dst) =
v
/\
forall r,
data_preg r =
true ->
r <>
preg_of dst ->
rs'#
r =
rs#
r.
Proof.
Lemma storeind_correct:
forall (
base:
ireg)
ofs ty src k (
rs:
regset)
c m m',
storeind src base ofs ty k =
OK c ->
Mem.storev (
chunk_of_type ty)
m (
Val.offset_ptr rs#
base ofs) (
rs#(
preg_of src)) =
Some m' ->
exists rs',
exec_straight init_stk ge fn c rs m k rs'
m'
/\
forall r,
data_preg r =
true ->
preg_notin r (
destroyed_by_setstack ty) ->
rs'#
r =
rs#
r.
Proof.
Translation of addressing modes
Lemma transl_addressing_mode_32_correct:
forall addr args am (
rs:
regset)
v,
transl_addressing addr args =
OK am ->
eval_addressing32 ge (
rs RSP)
addr (
List.map rs (
List.map preg_of args)) =
Some v ->
Val.lessdef v (
eval_addrmode32 ge am rs).
Proof.
Lemma transl_addressing_mode_64_correct:
forall addr args am (
rs:
regset)
v,
transl_addressing addr args =
OK am ->
eval_addressing64 ge (
rs RSP)
addr (
List.map rs (
List.map preg_of args)) =
Some v ->
Val.lessdef v (
eval_addrmode64 ge am rs).
Proof.
Lemma transl_addressing_mode_correct:
forall addr args am (
rs:
regset)
v,
transl_addressing addr args =
OK am ->
eval_addressing ge (
rs RSP)
addr (
List.map rs (
List.map preg_of args)) =
Some v ->
Val.lessdef v (
eval_addrmode ge am rs).
Proof.
Lemma normalize_addrmode_32_correct:
forall am rs,
eval_addrmode32 ge (
normalize_addrmode_32 am)
rs =
eval_addrmode32 ge am rs.
Proof.
intros;
destruct am as [
base ofs [
n|
r]];
simpl;
auto.
rewrite Int.repr_signed.
auto.
Qed.
Lemma normalize_addrmode_64_correct:
forall am rs,
eval_addrmode64 ge am rs =
match normalize_addrmode_64 am with
| (
am',
None) =>
eval_addrmode64 ge am'
rs
| (
am',
Some delta) =>
Val.addl (
eval_addrmode64 ge am'
rs) (
Vlong delta)
end.
Proof.
Processor conditions and comparisons
Lemma compare_ints_spec:
forall rs v1 v2 m sz,
let rs' :=
nextinstr (
compare_ints v1 v2 rs m)
sz in
rs'#
ZF =
Val.cmpu (
Mem.valid_pointer m)
Ceq v1 v2
/\
rs'#
CF =
Val.cmpu (
Mem.valid_pointer m)
Clt v1 v2
/\
rs'#
SF =
Val.negative (
Val.sub v1 v2)
/\
rs'#
OF =
Val.sub_overflow v1 v2
/\ (
forall r,
data_preg r =
true ->
rs'#
r =
rs#
r).
Proof.
intros.
unfold rs';
unfold compare_ints.
split.
auto.
split.
auto.
split.
auto.
split.
auto.
intros.
Simplifs.
Qed.
Lemma testcond_for_signed_comparison_32_correct:
forall c v1 v2 rs m b sz,
Val.cmp_bool c v1 v2 =
Some b ->
eval_testcond (
testcond_for_signed_comparison c)
(
nextinstr (
compare_ints v1 v2 rs m)
sz) =
Some b.
Proof.
Lemma testcond_for_unsigned_comparison_32_correct:
forall c v1 v2 rs m b sz,
Val.cmpu_bool (
Mem.valid_pointer m)
c v1 v2 =
Some b ->
eval_testcond (
testcond_for_unsigned_comparison c)
(
nextinstr (
compare_ints v1 v2 rs m)
sz) =
Some b.
Proof.
Lemma compare_longs_spec:
forall rs v1 v2 m sz,
let rs' :=
nextinstr (
compare_longs v1 v2 rs m)
sz in
rs'#
ZF =
Val.maketotal (
Val.cmplu (
Mem.valid_pointer m)
Ceq v1 v2)
/\
rs'#
CF =
Val.maketotal (
Val.cmplu (
Mem.valid_pointer m)
Clt v1 v2)
/\
rs'#
SF =
Val.negativel (
Val.subl v1 v2)
/\
rs'#
OF =
Val.subl_overflow v1 v2
/\ (
forall r,
data_preg r =
true ->
rs'#
r =
rs#
r).
Proof.
intros.
unfold rs';
unfold compare_longs.
split.
auto.
split.
auto.
split.
auto.
split.
auto.
intros.
Simplifs.
Qed.
Lemma int64_sub_overflow:
forall x y,
Int.xor (
Int.repr (
Int64.unsigned (
Int64.sub_overflow x y Int64.zero)))
(
Int.repr (
Int64.unsigned (
Int64.negative (
Int64.sub x y)))) =
(
if Int64.lt x y then Int.one else Int.zero).
Proof.
Lemma testcond_for_signed_comparison_64_correct:
forall c v1 v2 rs m b sz,
Val.cmpl_bool c v1 v2 =
Some b ->
eval_testcond (
testcond_for_signed_comparison c)
(
nextinstr (
compare_longs v1 v2 rs m)
sz) =
Some b.
Proof.
Lemma testcond_for_unsigned_comparison_64_correct:
forall c v1 v2 rs m b sz,
Val.cmplu_bool (
Mem.valid_pointer m)
c v1 v2 =
Some b ->
eval_testcond (
testcond_for_unsigned_comparison c)
(
nextinstr (
compare_longs v1 v2 rs m)
sz) =
Some b.
Proof.
Lemma compare_floats_spec:
forall rs n1 n2 sz,
let rs' :=
nextinstr (
compare_floats (
Vfloat n1) (
Vfloat n2)
rs)
sz in
rs'#
ZF =
Val.of_bool (
negb (
Float.cmp Cne n1 n2))
/\
rs'#
CF =
Val.of_bool (
negb (
Float.cmp Cge n1 n2))
/\
rs'#
PF =
Val.of_bool (
negb (
Float.cmp Ceq n1 n2 ||
Float.cmp Clt n1 n2 ||
Float.cmp Cgt n1 n2))
/\ (
forall r,
data_preg r =
true ->
rs'#
r =
rs#
r).
Proof.
intros.
unfold rs';
unfold compare_floats.
split.
auto.
split.
auto.
split.
auto.
intros.
Simplifs.
Qed.
Lemma compare_floats32_spec:
forall rs n1 n2 sz,
let rs' :=
nextinstr (
compare_floats32 (
Vsingle n1) (
Vsingle n2)
rs)
sz in
rs'#
ZF =
Val.of_bool (
negb (
Float32.cmp Cne n1 n2))
/\
rs'#
CF =
Val.of_bool (
negb (
Float32.cmp Cge n1 n2))
/\
rs'#
PF =
Val.of_bool (
negb (
Float32.cmp Ceq n1 n2 ||
Float32.cmp Clt n1 n2 ||
Float32.cmp Cgt n1 n2))
/\ (
forall r,
data_preg r =
true ->
rs'#
r =
rs#
r).
Proof.
intros.
unfold rs';
unfold compare_floats32.
split.
auto.
split.
auto.
split.
auto.
intros.
Simplifs.
Qed.
Definition eval_extcond (
xc:
extcond) (
rs:
regset) :
option bool :=
match xc with
|
Cond_base c =>
eval_testcond c rs
|
Cond_and c1 c2 =>
match eval_testcond c1 rs,
eval_testcond c2 rs with
|
Some b1,
Some b2 =>
Some (
b1 &&
b2)
|
_,
_ =>
None
end
|
Cond_or c1 c2 =>
match eval_testcond c1 rs,
eval_testcond c2 rs with
|
Some b1,
Some b2 =>
Some (
b1 ||
b2)
|
_,
_ =>
None
end
end.
Definition swap_floats {
A:
Type} (
c:
comparison) (
n1 n2:
A) :
A :=
match c with
|
Clt |
Cle =>
n2
|
Ceq |
Cne |
Cgt |
Cge =>
n1
end.
Lemma testcond_for_float_comparison_correct:
forall c n1 n2 rs sz,
eval_extcond (
testcond_for_condition (
Ccompf c))
(
nextinstr (
compare_floats (
Vfloat (
swap_floats c n1 n2))
(
Vfloat (
swap_floats c n2 n1))
rs)
sz) =
Some(
Float.cmp c n1 n2).
Proof.
Lemma testcond_for_neg_float_comparison_correct:
forall c n1 n2 rs sz,
eval_extcond (
testcond_for_condition (
Cnotcompf c))
(
nextinstr (
compare_floats (
Vfloat (
swap_floats c n1 n2))
(
Vfloat (
swap_floats c n2 n1))
rs)
sz) =
Some(
negb(
Float.cmp c n1 n2)).
Proof.
Lemma testcond_for_float32_comparison_correct:
forall c n1 n2 rs sz,
eval_extcond (
testcond_for_condition (
Ccompfs c))
(
nextinstr (
compare_floats32 (
Vsingle (
swap_floats c n1 n2))
(
Vsingle (
swap_floats c n2 n1))
rs)
sz) =
Some(
Float32.cmp c n1 n2).
Proof.
Lemma testcond_for_neg_float32_comparison_correct:
forall c n1 n2 rs sz,
eval_extcond (
testcond_for_condition (
Cnotcompfs c))
(
nextinstr (
compare_floats32 (
Vsingle (
swap_floats c n1 n2))
(
Vsingle (
swap_floats c n2 n1))
rs)
sz) =
Some(
negb(
Float32.cmp c n1 n2)).
Proof.
Remark swap_floats_commut:
forall (
A B:
Type)
c (
f:
A ->
B)
x y,
swap_floats c (
f x) (
f y) =
f (
swap_floats c x y).
Proof.
intros. destruct c; auto.
Qed.
Remark compare_floats_inv:
forall vx vy rs r,
r <>
CR ZF ->
r <>
CR CF ->
r <>
CR PF ->
r <>
CR SF ->
r <>
CR OF ->
compare_floats vx vy rs r =
rs r.
Proof.
Remark compare_floats32_inv:
forall vx vy rs r,
r <>
CR ZF ->
r <>
CR CF ->
r <>
CR PF ->
r <>
CR SF ->
r <>
CR OF ->
compare_floats32 vx vy rs r =
rs r.
Proof.
Lemma transl_cond_correct:
forall cond args k c rs m,
transl_cond cond args k =
OK c ->
exists rs',
exec_straight init_stk ge fn c rs m k rs'
m
/\
match eval_condition cond (
map rs (
map preg_of args))
m with
|
None =>
True
|
Some b =>
eval_extcond (
testcond_for_condition cond)
rs' =
Some b
end
/\
forall r,
data_preg r =
true ->
rs'#
r =
rs r.
Proof.
Remark eval_testcond_nextinstr:
forall c rs sz,
eval_testcond c (
nextinstr rs sz) =
eval_testcond c rs.
Proof.
Remark eval_testcond_set_ireg:
forall c rs r v,
eval_testcond c (
rs#(
IR r) <-
v) =
eval_testcond c rs.
Proof.
Lemma mk_setcc_base_correct:
forall cond rd k rs1 m,
exists rs2,
exec_straight init_stk ge fn (
mk_setcc_base cond rd k)
rs1 m k rs2 m
/\
rs2#
rd =
Val.of_optbool(
eval_extcond cond rs1)
/\
forall r,
data_preg r =
true ->
r <>
RAX /\
r <>
RCX ->
r <>
rd ->
rs2#
r =
rs1#
r.
Proof.
Lemma mk_setcc_correct:
forall cond rd k rs1 m,
exists rs2,
exec_straight init_stk ge fn (
mk_setcc cond rd k)
rs1 m k rs2 m
/\
rs2#
rd =
Val.of_optbool(
eval_extcond cond rs1)
/\
forall r,
data_preg r =
true ->
r <>
RAX /\
r <>
RCX ->
r <>
rd ->
rs2#
r =
rs1#
r.
Proof.
Translation of arithmetic operations.
Ltac ArgsInv :=
match goal with
| [
H:
Error _ =
OK _ |-
_ ] =>
discriminate
| [
H:
match ?
args with nil =>
_ |
_ ::
_ =>
_ end =
OK _ |-
_ ] =>
destruct args;
ArgsInv
| [
H:
bind _ _ =
OK _ |-
_ ] =>
monadInv H;
ArgsInv
| [
H:
match _ with left _ =>
_ |
right _ =>
assertion_failed end =
OK _ |-
_ ] =>
monadInv H;
ArgsInv
| [
H:
match _ with true =>
_ |
false =>
assertion_failed end =
OK _ |-
_ ] =>
monadInv H;
ArgsInv
| [
H:
ireg_of _ =
OK _ |-
_ ] =>
simpl in *;
rewrite (
ireg_of_eq _ _ H)
in *;
clear H;
ArgsInv
| [
H:
freg_of _ =
OK _ |-
_ ] =>
simpl in *;
rewrite (
freg_of_eq _ _ H)
in *;
clear H;
ArgsInv
|
_ =>
idtac
end.
Ltac TranslOp :=
econstructor;
split;
[
apply exec_straight_one; [
unfold exec_instr;
simpl;
eauto |
auto ]
|
split; [
Simplifs |
intros;
Simplifs ]].
Lemma transl_op_correct:
forall op args res k c (
rs:
regset)
m v,
transl_op op args res k =
OK c ->
eval_operation ge (
rs#
RSP)
op (
map rs (
map preg_of args))
m =
Some v ->
exists rs',
exec_straight init_stk ge fn c rs m k rs'
m
/\
Val.lessdef v rs'#(
preg_of res)
/\
forall r,
data_preg r =
true ->
r <>
preg_of res ->
preg_notin r (
destroyed_by_op op) ->
rs'
r =
rs r.
Proof.
Translation of memory loads.
Lemma transl_load_correct:
forall chunk addr args dest k c (
rs:
regset)
m a v,
transl_load chunk addr args dest k =
OK c ->
eval_addressing ge (
rs#
RSP)
addr (
map rs (
map preg_of args)) =
Some a ->
Mem.loadv chunk m a =
Some v ->
exists rs',
exec_straight init_stk ge fn c rs m k rs'
m
/\
rs'#(
preg_of dest) =
v
/\
forall r,
data_preg r =
true ->
r <>
preg_of dest ->
rs'#
r =
rs#
r.
Proof.
Lemma transl_store_correct:
forall chunk addr args src k c (
rs:
regset)
m a m',
transl_store chunk addr args src k =
OK c ->
eval_addressing ge (
rs#
RSP)
addr (
map rs (
map preg_of args)) =
Some a ->
Mem.storev chunk m a (
rs (
preg_of src)) =
Some m' ->
exists rs',
exec_straight init_stk ge fn c rs m k rs'
m'
/\
forall r,
data_preg r =
true ->
preg_notin r (
destroyed_by_store chunk addr) ->
rs'#
r =
rs#
r.
Proof.
End CONSTRUCTORS.