summaryrefslogtreecommitdiff
path: root/arm/ConstpropOp.v
diff options
context:
space:
mode:
Diffstat (limited to 'arm/ConstpropOp.v')
-rw-r--r--arm/ConstpropOp.v14
1 files changed, 14 insertions, 0 deletions
diff --git a/arm/ConstpropOp.v b/arm/ConstpropOp.v
index fa97c6c..86b6d66 100644
--- a/arm/ConstpropOp.v
+++ b/arm/ConstpropOp.v
@@ -345,6 +345,12 @@ Inductive eval_static_operation_cases: forall (op: operation) (vl: list approx),
| eval_static_operation_case49:
forall n1,
eval_static_operation_cases (Ofloatofint) (I n1 :: nil)
+ | eval_static_operation_case50:
+ forall n1,
+ eval_static_operation_cases (Ointuoffloat) (F n1 :: nil)
+ | eval_static_operation_case53:
+ forall n1,
+ eval_static_operation_cases (Ofloatofintu) (I n1 :: nil)
| eval_static_operation_case51:
forall c vl,
eval_static_operation_cases (Ocmp c) (vl)
@@ -455,6 +461,10 @@ Definition eval_static_operation_match (op: operation) (vl: list approx) :=
eval_static_operation_case48 n1
| Ofloatofint, I n1 :: nil =>
eval_static_operation_case49 n1
+ | Ointuoffloat, F n1 :: nil =>
+ eval_static_operation_case50 n1
+ | Ofloatofintu, I n1 :: nil =>
+ eval_static_operation_case53 n1
| Ocmp c, vl =>
eval_static_operation_case51 c vl
| Oshrximm n, I n1 :: nil =>
@@ -563,6 +573,10 @@ Definition eval_static_operation (op: operation) (vl: list approx) :=
match Float.intoffloat n1 with Some x => I x | None => Unknown end
| eval_static_operation_case49 n1 =>
F(Float.floatofint n1)
+ | eval_static_operation_case50 n1 =>
+ match Float.intuoffloat n1 with Some x => I x | None => Unknown end
+ | eval_static_operation_case53 n1 =>
+ F(Float.floatofintu n1)
| eval_static_operation_case51 c vl =>
match eval_static_condition c vl with
| None => Unknown