aboutsummaryrefslogtreecommitdiffstats
path: root/mppa_k1c/NeedOp.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-06-04 21:42:30 +0200
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-06-04 21:42:30 +0200
commit6064bac57701ba0a12031d43acbe25cb0140730c (patch)
tree58577a651e670b9a39b94d644f5da039def73864 /mppa_k1c/NeedOp.v
parentac366a59308ae85a0cbfefb8b9be79763d5c5f91 (diff)
downloadcompcert-kvx-6064bac57701ba0a12031d43acbe25cb0140730c.tar.gz
compcert-kvx-6064bac57701ba0a12031d43acbe25cb0140730c.zip
begin osel imm
Diffstat (limited to 'mppa_k1c/NeedOp.v')
-rw-r--r--mppa_k1c/NeedOp.v12
1 files changed, 12 insertions, 0 deletions
diff --git a/mppa_k1c/NeedOp.v b/mppa_k1c/NeedOp.v
index 047c180a..4748f38b 100644
--- a/mppa_k1c/NeedOp.v
+++ b/mppa_k1c/NeedOp.v
@@ -133,6 +133,8 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| Oextfz _ _ | Oextfs _ _ | Oextfzl _ _ | Oextfsl _ _ => op1 (default nv)
| Oinsf _ _ | Oinsfl _ _ => op2 (default nv)
| Osel c ty => nv :: nv :: needs_of_condition0 c
+ | Oselimm c imm
+ | Osellimm c imm => nv :: needs_of_condition0 c
end.
Definition operation_is_redundant (op: operation) (nv: nval): bool :=
@@ -347,6 +349,16 @@ Proof.
erewrite needs_of_condition0_sound by eauto.
apply select_sound; auto.
simpl; auto with na.
+ (* select imm *)
+- destruct (eval_condition0 _ _ _) as [b|] eqn:EC.
+ { erewrite needs_of_condition0_sound by eauto.
+ apply select_sound; auto with na. }
+ simpl; auto with na.
+ (* select long imm *)
+- destruct (eval_condition0 _ _ _) as [b|] eqn:EC.
+ { erewrite needs_of_condition0_sound by eauto.
+ apply select_sound; auto with na. }
+ simpl; auto with na.
Qed.
Lemma operation_is_redundant_sound: