summaryrefslogtreecommitdiff
path: root/powerpc
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2014-01-03 10:54:41 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2014-01-03 10:54:41 +0000
commit089c6c6dc139a0c32f8566d028702d39d0748077 (patch)
tree24b5b61fb63d94836118d276636e6285a690094f /powerpc
parent40653a485d26e94b57a01250623182f0f5358e10 (diff)
Update for the multiple-input-needs case.
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2401 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'powerpc')
-rw-r--r--powerpc/NeedOp.v85
1 files changed, 45 insertions, 40 deletions
diff --git a/powerpc/NeedOp.v b/powerpc/NeedOp.v
index 63323eb..d43af3c 100644
--- a/powerpc/NeedOp.v
+++ b/powerpc/NeedOp.v
@@ -11,42 +11,52 @@ Require Import RTL.
(** Neededness analysis for PowerPC operators *)
-Definition needs_of_condition (cond: condition): nval :=
+Definition op1 (nv: nval) := nv :: nil.
+Definition op2 (nv: nval) := nv :: nv :: nil.
+
+Definition needs_of_condition (cond: condition): list nval :=
match cond with
- | Cmaskzero n | Cmasknotzero n => maskzero n
- | _ => All
+ | Cmaskzero n | Cmasknotzero n => op1 (maskzero n)
+ | _ => nil
end.
-Definition needs_of_operation (op: operation) (nv: nval): nval :=
+Definition needs_of_operation (op: operation) (nv: nval): list nval :=
match op with
- | Omove => nv
- | Ointconst n => Nothing
- | Ofloatconst n => Nothing
- | Oaddrsymbol id ofs => Nothing
- | Oaddrstack ofs => Nothing
- | Ocast8signed => sign_ext 8 nv
- | Ocast16signed => sign_ext 16 nv
- | Oadd => modarith nv
- | Oaddimm n => modarith nv
- | Omul => modarith nv
- | Omulimm n => modarith nv
- | Oand => bitwise nv
- | Oandimm n => andimm nv n
- | Oor => bitwise nv
- | Oorimm n => orimm nv n
- | Oxor => bitwise nv
- | Oxorimm n => bitwise nv
- | Onot => bitwise nv
- | Onand => bitwise nv
- | Onor => bitwise nv
- | Onxor => bitwise nv
- | Oandc => bitwise nv
- | Oorc => bitwise nv
- | Oshrimm n => shrimm nv n
- | Orolm amount mask => rolm nv amount mask
- | Osingleoffloat => singleoffloat nv
+ | Omove => op1 nv
+ | Ointconst n => nil
+ | Ofloatconst 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)
+ | Oaddsymbol id ofs => op1 (modarith nv)
+ | Osub => op2 (default nv)
+ | Osubimm n => op1 (default nv)
+ | Omul => op2 (modarith nv)
+ | Omulimm n => op1 (modarith nv)
+ | Omulhs | Omulhu | Odiv | Odivu => op2 (default nv)
+ | Oand => op2 (bitwise nv)
+ | Oandimm n => op1 (andimm nv n)
+ | Oor => op2 (bitwise nv)
+ | Oorimm n => op1 (orimm nv n)
+ | Oxor => op2 (bitwise nv)
+ | Oxorimm n => op1 (bitwise nv)
+ | Onot => op1 (bitwise nv)
+ | Onand | Onor | Onxor | Oandc | Oorc => op2 (bitwise nv)
+ | Oshl | Oshr | Oshru => op2 (default nv)
+ | Oshrimm n => op1 (shrimm nv n)
+ | Oshrximm n => op1 (default nv)
+ | Orolm amount mask => op1 (rolm nv amount mask)
+ | Oroli amount mask => op1 (default nv)
+ | Onegf | Oabsf => op1 (default nv)
+ | Oaddf | Osubf | Omulf | Odivf => op2 (default nv)
+ | Osingleoffloat => op1 (singleoffloat nv)
+ | Ointoffloat => op1 (default nv)
+ | Ofloatofwords | Omakelong => op2 (default nv)
+ | Olowlong | Ohighlong => op1 (default nv)
| Ocmp c => needs_of_condition c
- | _ => default nv
end.
Definition operation_is_redundant (op: operation) (nv: nval): bool :=
@@ -64,8 +74,6 @@ Ltac InvAgree :=
match goal with
| [H: vagree_list nil _ _ |- _ ] => inv H; InvAgree
| [H: vagree_list (_::_) _ _ |- _ ] => inv H; InvAgree
- | [H: list_forall2 _ nil _ |- _ ] => inv H; InvAgree
- | [H: list_forall2 _ (_::_) _ |- _ ] => inv H; InvAgree
| _ => idtac
end.
@@ -107,14 +115,11 @@ Lemma needs_of_operation_sound:
Proof.
unfold needs_of_operation; intros; destruct op; try (eapply default_needs_of_operation_sound; eauto; fail);
simpl in *; FuncInv; InvAgree; TrivialExists.
-- auto with na.
-- auto with na.
-- auto with na.
-- auto with na.
- apply sign_ext_sound; auto. compute; auto.
- apply sign_ext_sound; auto. compute; auto.
- apply add_sound; auto.
- apply add_sound; auto with na.
+- apply add_sound; auto with na.
- apply mul_sound; auto.
- apply mul_sound; auto with na.
- apply and_sound; auto.
@@ -139,13 +144,13 @@ Proof.
Qed.
Lemma operation_is_redundant_sound:
- forall op nv arg1 args v arg1',
+ forall op nv arg1 args v arg1' args',
operation_is_redundant op nv = true ->
eval_operation ge (Vptr sp Int.zero) op (arg1 :: args) m = Some v ->
- vagree arg1 arg1' (needs_of_operation op nv) ->
+ vagree_list (arg1 :: args) (arg1' :: args') (needs_of_operation op nv) ->
vagree v arg1' nv.
Proof.
- intros. destruct op; simpl in *; try discriminate; FuncInv; subst.
+ intros. destruct op; simpl in *; try discriminate; inv H1; FuncInv; subst.
- apply sign_ext_redundant_sound; auto. omega.
- apply sign_ext_redundant_sound; auto. omega.
- apply andimm_redundant_sound; auto.