aboutsummaryrefslogtreecommitdiffstats
path: root/arm/NeedOp.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2019-04-15 17:50:30 +0200
committerXavier Leroy <xavierleroy@users.noreply.github.com>2019-05-20 18:00:46 +0200
commit49342474c19558709c8cea6d70eaba9a4dd7a150 (patch)
tree86fb6e9fc2081a71a2d770811164bb2f72b867e6 /arm/NeedOp.v
parent996f2e071baaf52105714283d141af2eac8ffbfb (diff)
downloadcompcert-kvx-49342474c19558709c8cea6d70eaba9a4dd7a150.tar.gz
compcert-kvx-49342474c19558709c8cea6d70eaba9a4dd7a150.zip
Implement a `Osel` operation for ARM
The operation comples down to conditional moves. Both integer and floating-point conditional moves are supported.
Diffstat (limited to 'arm/NeedOp.v')
-rw-r--r--arm/NeedOp.v5
1 files changed, 5 insertions, 0 deletions
diff --git a/arm/NeedOp.v b/arm/NeedOp.v
index dee7cae1..c70c7e40 100644
--- a/arm/NeedOp.v
+++ b/arm/NeedOp.v
@@ -83,6 +83,7 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| Omakelong => op2 (default nv)
| Olowlong | Ohighlong => 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 :=
@@ -183,6 +184,10 @@ Proof.
- apply notint_sound; auto.
- apply notint_sound. apply needs_of_shift_sound; auto.
- apply needs_of_shift_sound; auto.
+- 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: