aboutsummaryrefslogtreecommitdiffstats
path: root/riscV/NeedOp.v
diff options
context:
space:
mode:
authorLéo Gourdin <leo.gourdin@univ-grenoble-alpes.fr>2021-03-02 10:29:23 +0100
committerLéo Gourdin <leo.gourdin@univ-grenoble-alpes.fr>2021-03-02 10:29:23 +0100
commitbe4dcbd9fcd3c859a0fae7a37cd226493a8abefb (patch)
tree3d69513c3f082e7db01bd9ade327c0ebc8a2f441 /riscV/NeedOp.v
parentfd1a0395e540ee9fcd91d8f09161b34a22d9c51e (diff)
parenta9763cd4327e12316d62e80648122f122581cca4 (diff)
downloadcompcert-kvx-be4dcbd9fcd3c859a0fae7a37cd226493a8abefb.tar.gz
compcert-kvx-be4dcbd9fcd3c859a0fae7a37cd226493a8abefb.zip
Merge remote-tracking branch 'origin/riscV-cmov' into riscv-work
Diffstat (limited to 'riscV/NeedOp.v')
-rw-r--r--riscV/NeedOp.v15
1 files changed, 15 insertions, 0 deletions
diff --git a/riscV/NeedOp.v b/riscV/NeedOp.v
index d90c1b66..4b309f5b 100644
--- a/riscV/NeedOp.v
+++ b/riscV/NeedOp.v
@@ -116,6 +116,11 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| OEfeqs => op2 (default nv)
| OEflts => op2 (default nv)
| OEfles => op2 (default nv)
+ | Obits_of_single => op1 (default nv)
+ | 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 :=
@@ -184,6 +189,16 @@ Proof.
- apply shrimm_sound; auto.
- apply shruimm_sound; auto.
- apply xor_sound; auto with na.
+- (* 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: