aboutsummaryrefslogtreecommitdiffstats
path: root/x86/NeedOp.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2019-04-15 15:21:49 +0200
committerXavier Leroy <xavierleroy@users.noreply.github.com>2019-05-20 18:00:46 +0200
commit996f2e071baaf52105714283d141af2eac8ffbfb (patch)
tree0d3968e7cb130b2399fe53a30ac7b3b6405d2284 /x86/NeedOp.v
parentd002919334e83904447957f666f0d48704c5e55b (diff)
downloadcompcert-996f2e071baaf52105714283d141af2eac8ffbfb.tar.gz
compcert-996f2e071baaf52105714283d141af2eac8ffbfb.zip
Implement a `Osel` operation for x86
The operation compiles down to conditional moves.
Diffstat (limited to 'x86/NeedOp.v')
-rw-r--r--x86/NeedOp.v5
1 files changed, 5 insertions, 0 deletions
diff --git a/x86/NeedOp.v b/x86/NeedOp.v
index 68ecc745..d9a58fbb 100644
--- a/x86/NeedOp.v
+++ b/x86/NeedOp.v
@@ -120,6 +120,7 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| Ointoffloat | Ofloatofint | Ointofsingle | Osingleofint => op1 (default nv)
| Olongoffloat | Ofloatoflong | Olongofsingle | Osingleoflong => op1 (default nv)
| Ocmp c => needs_of_condition c
+ | Osel c ty => nv :: nv :: needs_of_condition c
end.
Definition operation_is_redundant (op: operation) (nv: nval): bool :=
@@ -231,6 +232,10 @@ Proof.
erewrite needs_of_condition_sound by eauto.
subst v; simpl. auto with na.
subst v; auto with na.
+- destruct (eval_condition c args m) as [b|] eqn:EC.
+ erewrite needs_of_condition_sound by eauto.
+ apply select_sound; auto.
+ simpl; auto with na.
Qed.
Lemma operation_is_redundant_sound: