Require Import Coqlib.
Require Import AST Integers Floats.
Require Import Values Memory Globalenvs.
Require Import Op RTL.
Require Import NeedDomain.
Require Import Lia.
Neededness analysis for RISC-V operators
Definition op1 (
nv:
nval) :=
nv ::
nil.
Definition op2 (
nv:
nval) :=
nv ::
nv ::
nil.
Definition op3 (
nv:
nval) :=
nv ::
nv ::
nv ::
nil.
Definition needs_of_condition (
cond:
condition):
list nval :=
nil.
Definition needs_of_condition0 (
cond0:
condition0):
list nval :=
nil.
Definition needs_of_operation (
op:
operation) (
nv:
nval):
list nval :=
match op with
|
Omove =>
op1 nv
|
Ocopy =>
op2 (
default nv)
|
Ocopyimm _ =>
op1 (
default nv)
|
Ointconst n =>
nil
|
Olongconst n =>
nil
|
Ofloatconst n =>
nil
|
Osingleconst n =>
nil
|
Oaddrsymbol id ofs =>
nil
|
Oaddrstack ofs =>
nil
|
Ocast8signed =>
op1 (
sign_ext 8
nv)
|
Ocast16signed =>
op1 (
sign_ext 16
nv)
|
Oadd =>
op2 (
modarith nv)
|
Oaddimm n =>
op1 (
modarith nv)
|
Oaddx _ =>
op2 (
default nv)
|
Oaddximm _ _ =>
op1 (
default nv)
|
Oneg =>
op1 (
modarith nv)
|
Osub =>
op2 (
default nv)
|
Orevsubimm _ =>
op1 (
default nv)
|
Orevsubx _ =>
op2 (
default nv)
|
Orevsubximm _ _ =>
op1 (
default nv)
|
Omul =>
op2 (
modarith nv)
|
Omulimm _ =>
op1 (
modarith nv)
|
Odiv |
Odivu |
Omod |
Omodu =>
op2 (
default nv)
|
Oand =>
op2 (
bitwise nv)
|
Oandimm n =>
op1 (
andimm nv n)
|
Onand =>
op2 (
bitwise nv)
|
Onandimm n =>
op1 (
andimm nv n)
|
Oor =>
op2 (
bitwise nv)
|
Oorimm n =>
op1 (
orimm nv n)
|
Onor =>
op2 (
bitwise nv)
|
Onorimm n =>
op1 (
orimm nv n)
|
Oxor =>
op2 (
bitwise nv)
|
Oxorimm n =>
op1 (
bitwise nv)
|
Onxor =>
op2 (
bitwise nv)
|
Onxorimm n =>
op1 (
bitwise nv)
|
Onot =>
op1 (
bitwise nv)
|
Oandn =>
op2 (
bitwise nv)
|
Oandnimm n =>
op1 (
andimm nv n)
|
Oorn =>
op2 (
bitwise nv)
|
Oornimm n =>
op1 (
orimm nv n)
|
Oshl |
Oshr |
Oshru =>
op2 (
default nv)
|
Oshlimm n =>
op1 (
shlimm nv n)
|
Oshrimm n =>
op1 (
shrimm nv n)
|
Ororimm n =>
op1 (
ror nv n)
|
Oshruimm n =>
op1 (
shruimm nv n)
|
Oshrximm n =>
op1 (
default nv)
|
Omadd =>
op3 (
modarith nv)
|
Omaddimm n =>
op2 (
modarith nv)
|
Omsub =>
op3 (
modarith nv)
|
Omakelong =>
op2 (
default nv)
|
Olowlong |
Ohighlong =>
op1 (
default nv)
|
Ocast32signed =>
op1 (
default nv)
|
Ocast32unsigned =>
op1 (
default nv)
|
Oaddl =>
op2 (
default nv)
|
Oaddlimm n =>
op1 (
default nv)
|
Oaddxl _ =>
op2 (
default nv)
|
Oaddxlimm _ _ =>
op1 (
default nv)
|
Orevsublimm _ =>
op1 (
default nv)
|
Orevsubxl _ =>
op2 (
default nv)
|
Orevsubxlimm _ _ =>
op1 (
default nv)
|
Onegl =>
op1 (
default nv)
|
Osubl =>
op2 (
default nv)
|
Omull =>
op2 (
default nv)
|
Omullimm _ =>
op1 (
default nv)
|
Omullhs |
Omullhu |
Odivl |
Odivlu |
Omodl |
Omodlu =>
op2 (
default nv)
|
Oandl =>
op2 (
default nv)
|
Oandlimm n =>
op1 (
default nv)
|
Onandl =>
op2 (
default nv)
|
Onandlimm n =>
op1 (
default nv)
|
Oorl =>
op2 (
default nv)
|
Oorlimm n =>
op1 (
default nv)
|
Onorl =>
op2 (
default nv)
|
Onorlimm n =>
op1 (
default nv)
|
Oxorl =>
op2 (
default nv)
|
Oxorlimm n =>
op1 (
default nv)
|
Onxorl =>
op2 (
default nv)
|
Onxorlimm n =>
op1 (
default nv)
|
Onotl =>
op1 (
default nv)
|
Oandnl =>
op2 (
default nv)
|
Oandnlimm n =>
op1 (
default nv)
|
Oornl =>
op2 (
default nv)
|
Oornlimm n =>
op1 (
default nv)
|
Oshll |
Oshrl |
Oshrlu =>
op2 (
default nv)
|
Oshllimm n =>
op1 (
default nv)
|
Oshrlimm n =>
op1 (
default nv)
|
Oshrluimm n =>
op1 (
default nv)
|
Oshrxlimm n =>
op1 (
default nv)
|
Omaddl =>
op3 (
default nv)
|
Omaddlimm n =>
op2 (
default nv)
|
Omsubl =>
op3 (
default nv)
|
Onegf |
Oabsf =>
op1 (
default nv)
|
Oaddf |
Osubf |
Omulf |
Odivf |
Ominf |
Omaxf =>
op2 (
default nv)
|
Ofmaddf |
Ofmsubf =>
op3 (
default nv)
|
Onegfs |
Oabsfs =>
op1 (
default nv)
|
Oaddfs |
Osubfs |
Omulfs |
Odivfs |
Ominfs |
Omaxfs =>
op2 (
default nv)
|
Oinvfs =>
op1 (
default nv)
|
Ofmaddfs |
Ofmsubfs =>
op3 (
default nv)
|
Ofloatofsingle |
Osingleoffloat =>
op1 (
default nv)
|
Ointoffloat |
Ointuoffloat =>
op1 (
default nv)
|
Olongoffloat |
Olonguoffloat |
Ofloatoflong |
Ofloatoflongu =>
op1 (
default nv)
|
Ointofsingle |
Ointuofsingle |
Osingleofint |
Osingleofintu =>
op1 (
default nv)
|
Olongofsingle |
Olonguofsingle |
Osingleoflong |
Osingleoflongu =>
op1 (
default nv)
|
Ointofsingle_ne |
Ointuofsingle_ne =>
op1 (
default nv)
|
Olongoffloat_ne |
Olonguoffloat_ne =>
op1 (
default nv)
|
Ocmp c =>
needs_of_condition c
|
Oextfz _ _ |
Oextfs _ _ |
Oextfzl _ _ |
Oextfsl _ _ =>
op1 (
default nv)
|
Oinsf _ _ |
Oinsfl _ _ =>
op2 (
default nv)
|
Osel c =>
nv ::
nv ::
needs_of_condition0 c
|
Oselimm c imm
|
Osellimm c imm =>
nv ::
needs_of_condition0 c
|
Oabsdiff =>
op2 (
default nv)
|
Oabsdiffimm _ =>
op1 (
default nv)
|
Oabsdiffl =>
op2 (
default nv)
|
Oabsdifflimm _ =>
op1 (
default nv)
end.
Definition operation_is_redundant (
op:
operation) (
nv:
nval):
bool :=
match op with
|
Ocast8signed =>
sign_ext_redundant 8
nv
|
Ocast16signed =>
sign_ext_redundant 16
nv
|
Oandimm n =>
andimm_redundant nv n
|
Oorimm n =>
orimm_redundant nv n
| _ =>
false
end.
Ltac InvAgree :=
match goal with
| [
H:
vagree_list nil _ _ |- _ ] =>
inv H;
InvAgree
| [
H:
vagree_list (_::_) _ _ |- _ ] =>
inv H;
InvAgree
| _ =>
idtac
end.
Ltac TrivialExists :=
match goal with
| [ |-
exists v,
Some ?
x =
Some v /\ _ ] =>
exists x;
split;
auto
| _ =>
idtac
end.
Section SOUNDNESS.
Variable ge:
genv.
Variable sp:
block.
Variables m1 m2:
mem.
Hypothesis PERM:
forall b ofs k p,
Mem.perm m1 b ofs k p ->
Mem.perm m2 b ofs k p.
Lemma needs_of_condition_sound:
forall cond args b args',
eval_condition cond args m1 =
Some b ->
vagree_list args args' (
needs_of_condition cond) ->
eval_condition cond args' m2 =
Some b.
Proof.
Let valid_pointer_inj:
forall b1 ofs b2 delta,
inject_id b1 =
Some(
b2,
delta) ->
Mem.valid_pointer m1 b1 (
Ptrofs.unsigned ofs) =
true ->
Mem.valid_pointer m2 b2 (
Ptrofs.unsigned (
Ptrofs.add ofs (
Ptrofs.repr delta))) =
true.
Proof.
Let weak_valid_pointer_inj:
forall b1 ofs b2 delta,
inject_id b1 =
Some(
b2,
delta) ->
Mem.weak_valid_pointer m1 b1 (
Ptrofs.unsigned ofs) =
true ->
Mem.weak_valid_pointer m2 b2 (
Ptrofs.unsigned (
Ptrofs.add ofs (
Ptrofs.repr delta))) =
true.
Proof.
Let weak_valid_pointer_no_overflow:
forall b1 ofs b2 delta,
inject_id b1 =
Some(
b2,
delta) ->
Mem.weak_valid_pointer m1 b1 (
Ptrofs.unsigned ofs) =
true ->
0 <=
Ptrofs.unsigned ofs +
Ptrofs.unsigned (
Ptrofs.repr delta) <=
Ptrofs.max_unsigned.
Proof.
Let valid_different_pointers_inj:
forall b1 ofs1 b2 ofs2 b1' delta1 b2' delta2,
b1 <>
b2 ->
Mem.valid_pointer m1 b1 (
Ptrofs.unsigned ofs1) =
true ->
Mem.valid_pointer m1 b2 (
Ptrofs.unsigned ofs2) =
true ->
inject_id b1 =
Some (
b1',
delta1) ->
inject_id b2 =
Some (
b2',
delta2) ->
b1' <>
b2' \/
Ptrofs.unsigned (
Ptrofs.add ofs1 (
Ptrofs.repr delta1)) <>
Ptrofs.unsigned (
Ptrofs.add ofs2 (
Ptrofs.repr delta2)).
Proof.
unfold inject_id;
intros.
left;
congruence.
Qed.
Lemma needs_of_condition0_sound:
forall cond arg1 b arg2,
eval_condition0 cond arg1 m1 =
Some b ->
vagree arg1 arg2 All ->
eval_condition0 cond arg2 m2 =
Some b.
Proof.
Lemma subl_lessdef:
forall v1 v1' v2 v2',
Val.lessdef v1 v1' ->
Val.lessdef v2 v2' ->
Val.lessdef (
Val.subl v1 v2) (
Val.subl v1' v2').
Proof.
intros. inv H. inv H0. auto. destruct v1'; cbn; auto. cbn; auto.
Qed.
Remark default_idem:
forall nv,
default (
default nv) =
default nv.
Proof.
destruct nv; cbn; trivial.
Qed.
Lemma lessdef_triple_op_float:
forall f a b c x y z
(
Hax :
Val.lessdef (
Vfloat a)
x)
(
Hby :
Val.lessdef (
Vfloat b)
y)
(
Hcz :
Val.lessdef (
Vfloat c)
z),
Val.lessdef (
Vfloat (
f a b c)) (
ExtValues.triple_op_float f x y z).
Proof.
intros. inv Hax. inv Hby. inv Hcz. constructor.
Qed.
Lemma vagree_triple_op_float :
forall f a b c x y z nv,
(
vagree a x (
default nv)) ->
(
vagree b y (
default nv)) ->
(
vagree c z (
default nv)) ->
(
vagree (
ExtValues.triple_op_float f a b c)
(
ExtValues.triple_op_float f x y z)
nv).
Proof.
induction nv;
intros Hax Hby Hcz.
-
trivial.
-
cbn in *.
destruct a;
cbn;
trivial.
destruct b;
cbn;
trivial.
destruct c;
cbn;
trivial.
-
cbn in *.
destruct a;
cbn;
trivial.
destruct b;
cbn;
trivial.
destruct c;
cbn;
trivial.
-
cbn in *.
destruct a;
cbn;
trivial.
destruct b;
cbn;
trivial.
destruct c;
cbn;
trivial.
apply lessdef_triple_op_float;
auto.
Qed.
Lemma lessdef_triple_op_single:
forall f a b c x y z
(
Hax :
Val.lessdef (
Vsingle a)
x)
(
Hby :
Val.lessdef (
Vsingle b)
y)
(
Hcz :
Val.lessdef (
Vsingle c)
z),
Val.lessdef (
Vsingle (
f a b c)) (
ExtValues.triple_op_single f x y z).
Proof.
intros. inv Hax. inv Hby. inv Hcz. constructor.
Qed.
Lemma vagree_triple_op_single :
forall f a b c x y z nv,
(
vagree a x (
default nv)) ->
(
vagree b y (
default nv)) ->
(
vagree c z (
default nv)) ->
(
vagree (
ExtValues.triple_op_single f a b c)
(
ExtValues.triple_op_single f x y z)
nv).
Proof.
induction nv;
intros Hax Hby Hcz.
-
trivial.
-
cbn in *.
destruct a;
cbn;
trivial.
destruct b;
cbn;
trivial.
destruct c;
cbn;
trivial.
-
cbn in *.
destruct a;
cbn;
trivial.
destruct b;
cbn;
trivial.
destruct c;
cbn;
trivial.
-
cbn in *.
destruct a;
cbn;
trivial.
destruct b;
cbn;
trivial.
destruct c;
cbn;
trivial.
apply lessdef_triple_op_single;
auto.
Qed.
Hint Resolve vagree_triple_op_float vagree_triple_op_single :
na.
Lemma needs_of_operation_sound:
forall op args v nv args',
eval_operation ge (
Vptr sp Ptrofs.zero)
op args m1 =
Some v ->
vagree_list args args' (
needs_of_operation op nv) ->
nv <>
Nothing ->
exists v',
eval_operation ge (
Vptr sp Ptrofs.zero)
op args' m2 =
Some v'
/\
vagree v v' nv.
Proof.
Lemma operation_is_redundant_sound:
forall op nv arg1 args v arg1' args',
operation_is_redundant op nv =
true ->
eval_operation ge (
Vptr sp Ptrofs.zero)
op (
arg1 ::
args)
m1 =
Some v ->
vagree_list (
arg1 ::
args) (
arg1' ::
args') (
needs_of_operation op nv) ->
vagree v arg1' nv.
Proof.
End SOUNDNESS.