aboutsummaryrefslogtreecommitdiffstats
path: root/powerpc/NeedOp.v
diff options
context:
space:
mode:
Diffstat (limited to 'powerpc/NeedOp.v')
-rw-r--r--powerpc/NeedOp.v5
1 files changed, 5 insertions, 0 deletions
diff --git a/powerpc/NeedOp.v b/powerpc/NeedOp.v
index 9a579cc5..5ea09bd8 100644
--- a/powerpc/NeedOp.v
+++ b/powerpc/NeedOp.v
@@ -65,6 +65,7 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| Ofloatofwords | 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 :=
@@ -147,6 +148,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: