aboutsummaryrefslogtreecommitdiffstats
path: root/riscV/NeedOp.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2021-02-02 10:21:17 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2021-02-02 10:21:17 +0100
commitd2159e300b2d5e017a3144c747d34949b2ff2769 (patch)
tree2186636bc3f2e85b5effd19f9e7c23f4183545aa /riscV/NeedOp.v
parente0f1a90c2dcf7c43137064470ce4b12368b8435d (diff)
downloadcompcert-kvx-d2159e300b2d5e017a3144c747d34949b2ff2769.tar.gz
compcert-kvx-d2159e300b2d5e017a3144c747d34949b2ff2769.zip
begin implementing select
Diffstat (limited to 'riscV/NeedOp.v')
-rw-r--r--riscV/NeedOp.v11
1 files changed, 11 insertions, 0 deletions
diff --git a/riscV/NeedOp.v b/riscV/NeedOp.v
index 136c157f..9e1ad004 100644
--- a/riscV/NeedOp.v
+++ b/riscV/NeedOp.v
@@ -91,6 +91,7 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| Obits_of_float => op1 (default nv)
| Osingle_of_bits => op1 (default nv)
| Ofloat_of_bits => op1 (default nv)
+ | Oselectl => All :: nv :: nv :: nil
end.
Definition operation_is_redundant (op: operation) (nv: nval): bool :=
@@ -158,6 +159,16 @@ Proof.
- apply shlimm_sound; auto.
- apply shrimm_sound; auto.
- apply shruimm_sound; auto.
+- (* selectl *)
+ unfold ExtValues.select01_long.
+ destruct v0; auto with na.
+ assert (Val.lessdef (Vint i) v4) as LESSDEF by auto with na.
+ inv LESSDEF.
+ destruct (Int.eq i Int.one).
+ { apply normalize_sound; auto. }
+ destruct (Int.eq i Int.zero).
+ { apply normalize_sound; auto. }
+ cbn. auto with na.
Qed.
Lemma operation_is_redundant_sound: